--- a/src/HOL/Analysis/Arcwise_Connected.thy Sun Jan 13 20:25:41 2019 +0100
+++ b/src/HOL/Analysis/Arcwise_Connected.thy Mon Jan 14 11:59:19 2019 +0000
@@ -10,12 +10,12 @@
subsection%important \<open>The Brouwer reduction theorem\<close>
-theorem%important Brouwer_reduction_theorem_gen:
+theorem Brouwer_reduction_theorem_gen:
fixes S :: "'a::euclidean_space set"
assumes "closed S" "\<phi> S"
and \<phi>: "\<And>F. \<lbrakk>\<And>n. closed(F n); \<And>n. \<phi>(F n); \<And>n. F(Suc n) \<subseteq> F n\<rbrakk> \<Longrightarrow> \<phi>(\<Inter>range F)"
obtains T where "T \<subseteq> S" "closed T" "\<phi> T" "\<And>U. \<lbrakk>U \<subseteq> S; closed U; \<phi> U\<rbrakk> \<Longrightarrow> \<not> (U \<subset> T)"
-proof%unimportant -
+proof -
obtain B :: "nat \<Rightarrow> 'a set"
where "inj B" "\<And>n. open(B n)" and open_cov: "\<And>S. open S \<Longrightarrow> \<exists>K. S = \<Union>(B ` K)"
by (metis Setcompr_eq_image that univ_second_countable_sequence)
@@ -78,13 +78,13 @@
qed
qed
-corollary%important Brouwer_reduction_theorem:
+corollary Brouwer_reduction_theorem:
fixes S :: "'a::euclidean_space set"
assumes "compact S" "\<phi> S" "S \<noteq> {}"
and \<phi>: "\<And>F. \<lbrakk>\<And>n. compact(F n); \<And>n. F n \<noteq> {}; \<And>n. \<phi>(F n); \<And>n. F(Suc n) \<subseteq> F n\<rbrakk> \<Longrightarrow> \<phi>(\<Inter>range F)"
obtains T where "T \<subseteq> S" "compact T" "T \<noteq> {}" "\<phi> T"
"\<And>U. \<lbrakk>U \<subseteq> S; closed U; U \<noteq> {}; \<phi> U\<rbrakk> \<Longrightarrow> \<not> (U \<subset> T)"
-proof%unimportant (rule Brouwer_reduction_theorem_gen [of S "\<lambda>T. T \<noteq> {} \<and> T \<subseteq> S \<and> \<phi> T"])
+proof (rule Brouwer_reduction_theorem_gen [of S "\<lambda>T. T \<noteq> {} \<and> T \<subseteq> S \<and> \<phi> T"])
fix F
assume cloF: "\<And>n. closed (F n)"
and F: "\<And>n. F n \<noteq> {} \<and> F n \<subseteq> S \<and> \<phi> (F n)" and Fsub: "\<And>n. F (Suc n) \<subseteq> F n"
@@ -106,14 +106,14 @@
qed (meson assms compact_imp_closed seq_compact_closed_subset seq_compact_eq_compact)+
-subsection%important\<open>Arcwise Connections\<close>(*FIX ME this subsection is empty(?) *)
+subsection%unimportant\<open>Arcwise Connections\<close>(*FIX ME this subsection is empty(?) *)
subsection%important\<open>Density of points with dyadic rational coordinates\<close>
-proposition%important closure_dyadic_rationals:
+proposition closure_dyadic_rationals:
"closure (\<Union>k. \<Union>f \<in> Basis \<rightarrow> \<int>.
{ \<Sum>i :: 'a :: euclidean_space \<in> Basis. (f i / 2^k) *\<^sub>R i }) = UNIV"
-proof%unimportant -
+proof -
have "x \<in> closure (\<Union>k. \<Union>f \<in> Basis \<rightarrow> \<int>. {\<Sum>i \<in> Basis. (f i / 2^k) *\<^sub>R i})" for x::'a
proof (clarsimp simp: closure_approachable)
fix e::real
@@ -151,9 +151,9 @@
then show ?thesis by auto
qed
-corollary%important closure_rational_coordinates:
+corollary closure_rational_coordinates:
"closure (\<Union>f \<in> Basis \<rightarrow> \<rat>. { \<Sum>i :: 'a :: euclidean_space \<in> Basis. f i *\<^sub>R i }) = UNIV"
-proof%unimportant -
+proof -
have *: "(\<Union>k. \<Union>f \<in> Basis \<rightarrow> \<int>. { \<Sum>i::'a \<in> Basis. (f i / 2^k) *\<^sub>R i })
\<subseteq> (\<Union>f \<in> Basis \<rightarrow> \<rat>. { \<Sum>i \<in> Basis. f i *\<^sub>R i })"
proof clarsimp
@@ -167,7 +167,7 @@
using closure_dyadic_rationals closure_mono [OF *] by blast
qed
-lemma%unimportant closure_dyadic_rationals_in_convex_set:
+lemma closure_dyadic_rationals_in_convex_set:
"\<lbrakk>convex S; interior S \<noteq> {}\<rbrakk>
\<Longrightarrow> closure(S \<inter>
(\<Union>k. \<Union>f \<in> Basis \<rightarrow> \<int>.
@@ -175,7 +175,7 @@
closure S"
by (simp add: closure_dyadic_rationals closure_convex_Int_superset)
-lemma%unimportant closure_rationals_in_convex_set:
+lemma closure_rationals_in_convex_set:
"\<lbrakk>convex S; interior S \<noteq> {}\<rbrakk>
\<Longrightarrow> closure(S \<inter> (\<Union>f \<in> Basis \<rightarrow> \<rat>. { \<Sum>i :: 'a :: euclidean_space \<in> Basis. f i *\<^sub>R i })) =
closure S"
@@ -186,11 +186,11 @@
path connection is equivalent to arcwise connection for distinct points.
The proof is based on Whyburn's "Topological Analysis".\<close>
-lemma%important closure_dyadic_rationals_in_convex_set_pos_1:
+lemma closure_dyadic_rationals_in_convex_set_pos_1:
fixes S :: "real set"
assumes "convex S" and intnz: "interior S \<noteq> {}" and pos: "\<And>x. x \<in> S \<Longrightarrow> 0 \<le> x"
shows "closure(S \<inter> (\<Union>k m. {of_nat m / 2^k})) = closure S"
-proof%unimportant -
+proof -
have "\<exists>m. f 1/2^k = real m / 2^k" if "(f 1) / 2^k \<in> S" "f 1 \<in> \<int>" for k and f :: "real \<Rightarrow> real"
using that by (force simp: Ints_def zero_le_divide_iff power_le_zero_eq dest: pos zero_le_imp_eq_int)
then have "S \<inter> (\<Union>k m. {real m / 2^k}) = S \<inter>
@@ -201,13 +201,13 @@
qed
-definition%important dyadics :: "'a::field_char_0 set" where "dyadics \<equiv> \<Union>k m. {of_nat m / 2^k}"
+definition%unimportant dyadics :: "'a::field_char_0 set" where "dyadics \<equiv> \<Union>k m. {of_nat m / 2^k}"
-lemma%unimportant real_in_dyadics [simp]: "real m \<in> dyadics"
+lemma real_in_dyadics [simp]: "real m \<in> dyadics"
apply (simp add: dyadics_def)
by (metis divide_numeral_1 numeral_One power_0)
-lemma%unimportant nat_neq_4k1: "of_nat m \<noteq> (4 * of_nat k + 1) / (2 * 2^n :: 'a::field_char_0)"
+lemma nat_neq_4k1: "of_nat m \<noteq> (4 * of_nat k + 1) / (2 * 2^n :: 'a::field_char_0)"
proof
assume "of_nat m = (4 * of_nat k + 1) / (2 * 2^n :: 'a)"
then have "of_nat (m * (2 * 2^n)) = (of_nat (Suc (4 * k)) :: 'a)"
@@ -220,8 +220,8 @@
by simp
qed
-lemma%important nat_neq_4k3: "of_nat m \<noteq> (4 * of_nat k + 3) / (2 * 2^n :: 'a::field_char_0)"
-proof%unimportant
+lemma nat_neq_4k3: "of_nat m \<noteq> (4 * of_nat k + 3) / (2 * 2^n :: 'a::field_char_0)"
+proof
assume "of_nat m = (4 * of_nat k + 3) / (2 * 2^n :: 'a)"
then have "of_nat (m * (2 * 2^n)) = (of_nat (4 * k + 3) :: 'a)"
by (simp add: divide_simps)
@@ -233,10 +233,10 @@
by simp
qed
-lemma%important iff_4k:
+lemma iff_4k:
assumes "r = real k" "odd k"
shows "(4 * real m + r) / (2 * 2^n) = (4 * real m' + r) / (2 * 2 ^ n') \<longleftrightarrow> m=m' \<and> n=n'"
-proof%unimportant -
+proof -
{ assume "(4 * real m + r) / (2 * 2^n) = (4 * real m' + r) / (2 * 2 ^ n')"
then have "real ((4 * m + k) * (2 * 2 ^ n')) = real ((4 * m' + k) * (2 * 2^n))"
using assms by (auto simp: field_simps)
@@ -253,8 +253,8 @@
then show ?thesis by blast
qed
-lemma%important neq_4k1_k43: "(4 * real m + 1) / (2 * 2^n) \<noteq> (4 * real m' + 3) / (2 * 2 ^ n')"
-proof%unimportant
+lemma neq_4k1_k43: "(4 * real m + 1) / (2 * 2^n) \<noteq> (4 * real m' + 3) / (2 * 2 ^ n')"
+proof
assume "(4 * real m + 1) / (2 * 2^n) = (4 * real m' + 3) / (2 * 2 ^ n')"
then have "real (Suc (4 * m) * (2 * 2 ^ n')) = real ((4 * m' + 3) * (2 * 2^n))"
by (auto simp: field_simps)
@@ -270,11 +270,11 @@
using even_Suc by presburger
qed
-lemma%important dyadic_413_cases:
+lemma dyadic_413_cases:
obtains "(of_nat m::'a::field_char_0) / 2^k \<in> Nats"
| m' k' where "k' < k" "(of_nat m:: 'a) / 2^k = of_nat (4*m' + 1) / 2^Suc k'"
| m' k' where "k' < k" "(of_nat m:: 'a) / 2^k = of_nat (4*m' + 3) / 2^Suc k'"
-proof%unimportant (cases "m>0")
+proof (cases "m>0")
case False
then have "m=0" by simp
with that show ?thesis by auto
@@ -323,11 +323,11 @@
qed
-lemma%important dyadics_iff:
+lemma dyadics_iff:
"(dyadics :: 'a::field_char_0 set) =
Nats \<union> (\<Union>k m. {of_nat (4*m + 1) / 2^Suc k}) \<union> (\<Union>k m. {of_nat (4*m + 3) / 2^Suc k})"
(is "_ = ?rhs")
-proof%unimportant
+proof
show "dyadics \<subseteq> ?rhs"
unfolding dyadics_def
apply clarify
@@ -344,7 +344,7 @@
qed
-function (domintros) dyad_rec :: "[nat \<Rightarrow> 'a, 'a\<Rightarrow>'a, 'a\<Rightarrow>'a, real] \<Rightarrow> 'a" where
+function%unimportant (domintros) dyad_rec :: "[nat \<Rightarrow> 'a, 'a\<Rightarrow>'a, 'a\<Rightarrow>'a, real] \<Rightarrow> 'a" where
"dyad_rec b l r (real m) = b m"
| "dyad_rec b l r ((4 * real m + 1) / 2 ^ (Suc n)) = l (dyad_rec b l r ((2*m + 1) / 2^n))"
| "dyad_rec b l r ((4 * real m + 3) / 2 ^ (Suc n)) = r (dyad_rec b l r ((2*m + 1) / 2^n))"
@@ -354,14 +354,14 @@
apply (fastforce simp add: dyadics_iff Nats_def field_simps)+
done
-lemma%unimportant dyadics_levels: "dyadics = (\<Union>K. \<Union>k<K. \<Union> m. {of_nat m / 2^k})"
+lemma dyadics_levels: "dyadics = (\<Union>K. \<Union>k<K. \<Union> m. {of_nat m / 2^k})"
unfolding dyadics_def by auto
-lemma%important dyad_rec_level_termination:
+lemma dyad_rec_level_termination:
assumes "k < K"
shows "dyad_rec_dom(b, l, r, real m / 2^k)"
using assms
-proof%unimportant (induction K arbitrary: k m)
+proof (induction K arbitrary: k m)
case 0
then show ?case by auto
next
@@ -426,22 +426,22 @@
qed
-lemma%unimportant dyad_rec_termination: "x \<in> dyadics \<Longrightarrow> dyad_rec_dom(b,l,r,x)"
+lemma dyad_rec_termination: "x \<in> dyadics \<Longrightarrow> dyad_rec_dom(b,l,r,x)"
by (auto simp: dyadics_levels intro: dyad_rec_level_termination)
-lemma%unimportant dyad_rec_of_nat [simp]: "dyad_rec b l r (real m) = b m"
+lemma dyad_rec_of_nat [simp]: "dyad_rec b l r (real m) = b m"
by (simp add: dyad_rec.psimps dyad_rec_termination)
-lemma%unimportant dyad_rec_41 [simp]: "dyad_rec b l r ((4 * real m + 1) / 2 ^ (Suc n)) = l (dyad_rec b l r ((2*m + 1) / 2^n))"
+lemma dyad_rec_41 [simp]: "dyad_rec b l r ((4 * real m + 1) / 2 ^ (Suc n)) = l (dyad_rec b l r ((2*m + 1) / 2^n))"
apply (rule dyad_rec.psimps)
by (metis dyad_rec_level_termination lessI add.commute of_nat_Suc of_nat_mult of_nat_numeral)
-lemma%unimportant dyad_rec_43 [simp]: "dyad_rec b l r ((4 * real m + 3) / 2 ^ (Suc n)) = r (dyad_rec b l r ((2*m + 1) / 2^n))"
+lemma dyad_rec_43 [simp]: "dyad_rec b l r ((4 * real m + 3) / 2 ^ (Suc n)) = r (dyad_rec b l r ((2*m + 1) / 2^n))"
apply (rule dyad_rec.psimps)
by (metis dyad_rec_level_termination lessI of_nat_add of_nat_mult of_nat_numeral)
-lemma%unimportant dyad_rec_41_times2:
+lemma dyad_rec_41_times2:
assumes "n > 0"
shows "dyad_rec b l r (2 * ((4 * real m + 1) / 2^Suc n)) = l (dyad_rec b l r (2 * (2 * real m + 1) / 2^n))"
proof -
@@ -460,10 +460,10 @@
finally show ?thesis .
qed
-lemma%important dyad_rec_43_times2:
+lemma dyad_rec_43_times2:
assumes "n > 0"
shows "dyad_rec b l r (2 * ((4 * real m + 3) / 2^Suc n)) = r (dyad_rec b l r (2 * (2 * real m + 1) / 2^n))"
-proof%unimportant -
+proof -
obtain n' where n': "n = Suc n'"
using assms not0_implies_Suc by blast
have "dyad_rec b l r (2 * ((4 * real m + 3) / 2^Suc n)) = dyad_rec b l r ((2 * (4 * real m + 3)) / (2 * 2^n))"
@@ -479,22 +479,22 @@
finally show ?thesis .
qed
-definition%important dyad_rec2
+definition%unimportant dyad_rec2
where "dyad_rec2 u v lc rc x =
dyad_rec (\<lambda>z. (u,v)) (\<lambda>(a,b). (a, lc a b (midpoint a b))) (\<lambda>(a,b). (rc a b (midpoint a b), b)) (2*x)"
-abbreviation leftrec where "leftrec u v lc rc x \<equiv> fst (dyad_rec2 u v lc rc x)"
-abbreviation rightrec where "rightrec u v lc rc x \<equiv> snd (dyad_rec2 u v lc rc x)"
+abbreviation%unimportant leftrec where "leftrec u v lc rc x \<equiv> fst (dyad_rec2 u v lc rc x)"
+abbreviation%unimportant rightrec where "rightrec u v lc rc x \<equiv> snd (dyad_rec2 u v lc rc x)"
-lemma%unimportant leftrec_base: "leftrec u v lc rc (real m / 2) = u"
+lemma leftrec_base: "leftrec u v lc rc (real m / 2) = u"
by (simp add: dyad_rec2_def)
-lemma%unimportant leftrec_41: "n > 0 \<Longrightarrow> leftrec u v lc rc ((4 * real m + 1) / 2 ^ (Suc n)) = leftrec u v lc rc ((2 * real m + 1) / 2^n)"
+lemma leftrec_41: "n > 0 \<Longrightarrow> leftrec u v lc rc ((4 * real m + 1) / 2 ^ (Suc n)) = leftrec u v lc rc ((2 * real m + 1) / 2^n)"
apply (simp only: dyad_rec2_def dyad_rec_41_times2)
apply (simp add: case_prod_beta)
done
-lemma%unimportant leftrec_43: "n > 0 \<Longrightarrow>
+lemma leftrec_43: "n > 0 \<Longrightarrow>
leftrec u v lc rc ((4 * real m + 3) / 2 ^ (Suc n)) =
rc (leftrec u v lc rc ((2 * real m + 1) / 2^n)) (rightrec u v lc rc ((2 * real m + 1) / 2^n))
(midpoint (leftrec u v lc rc ((2 * real m + 1) / 2^n)) (rightrec u v lc rc ((2 * real m + 1) / 2^n)))"
@@ -502,10 +502,10 @@
apply (simp add: case_prod_beta)
done
-lemma%unimportant rightrec_base: "rightrec u v lc rc (real m / 2) = v"
+lemma rightrec_base: "rightrec u v lc rc (real m / 2) = v"
by (simp add: dyad_rec2_def)
-lemma%unimportant rightrec_41: "n > 0 \<Longrightarrow>
+lemma rightrec_41: "n > 0 \<Longrightarrow>
rightrec u v lc rc ((4 * real m + 1) / 2 ^ (Suc n)) =
lc (leftrec u v lc rc ((2 * real m + 1) / 2^n)) (rightrec u v lc rc ((2 * real m + 1) / 2^n))
(midpoint (leftrec u v lc rc ((2 * real m + 1) / 2^n)) (rightrec u v lc rc ((2 * real m + 1) / 2^n)))"
@@ -513,24 +513,24 @@
apply (simp add: case_prod_beta)
done
-lemma%unimportant rightrec_43: "n > 0 \<Longrightarrow> rightrec u v lc rc ((4 * real m + 3) / 2 ^ (Suc n)) = rightrec u v lc rc ((2 * real m + 1) / 2^n)"
+lemma rightrec_43: "n > 0 \<Longrightarrow> rightrec u v lc rc ((4 * real m + 3) / 2 ^ (Suc n)) = rightrec u v lc rc ((2 * real m + 1) / 2^n)"
apply (simp only: dyad_rec2_def dyad_rec_43_times2)
apply (simp add: case_prod_beta)
done
-lemma%unimportant dyadics_in_open_unit_interval:
+lemma dyadics_in_open_unit_interval:
"{0<..<1} \<inter> (\<Union>k m. {real m / 2^k}) = (\<Union>k. \<Union>m \<in> {0<..<2^k}. {real m / 2^k})"
by (auto simp: divide_simps)
-theorem%important homeomorphic_monotone_image_interval:
+theorem homeomorphic_monotone_image_interval:
fixes f :: "real \<Rightarrow> 'a::{real_normed_vector,complete_space}"
assumes cont_f: "continuous_on {0..1} f"
and conn: "\<And>y. connected ({0..1} \<inter> f -` {y})"
and f_1not0: "f 1 \<noteq> f 0"
shows "(f ` {0..1}) homeomorphic {0..1::real}"
-proof%unimportant -
+proof -
have "\<exists>c d. a \<le> c \<and> c \<le> m \<and> m \<le> d \<and> d \<le> b \<and>
(\<forall>x \<in> {c..d}. f x = f m) \<and>
(\<forall>x \<in> {a..<c}. (f x \<noteq> f m)) \<and>
@@ -1663,11 +1663,11 @@
qed
-theorem%important path_contains_arc:
+theorem path_contains_arc:
fixes p :: "real \<Rightarrow> 'a::{complete_space,real_normed_vector}"
assumes "path p" and a: "pathstart p = a" and b: "pathfinish p = b" and "a \<noteq> b"
obtains q where "arc q" "path_image q \<subseteq> path_image p" "pathstart q = a" "pathfinish q = b"
-proof%unimportant -
+proof -
have ucont_p: "uniformly_continuous_on {0..1} p"
using \<open>path p\<close> unfolding path_def
by (metis compact_Icc compact_uniformly_continuous)
@@ -1960,12 +1960,12 @@
qed
-corollary%important path_connected_arcwise:
+corollary path_connected_arcwise:
fixes S :: "'a::{complete_space,real_normed_vector} set"
shows "path_connected S \<longleftrightarrow>
(\<forall>x \<in> S. \<forall>y \<in> S. x \<noteq> y \<longrightarrow> (\<exists>g. arc g \<and> path_image g \<subseteq> S \<and> pathstart g = x \<and> pathfinish g = y))"
(is "?lhs = ?rhs")
-proof%unimportant (intro iffI impI ballI)
+proof (intro iffI impI ballI)
fix x y
assume "path_connected S" "x \<in> S" "y \<in> S" "x \<noteq> y"
then obtain p where p: "path p" "path_image p \<subseteq> S" "pathstart p = x" "pathfinish p = y"
@@ -1991,23 +1991,23 @@
qed
-corollary%important arc_connected_trans:
+corollary arc_connected_trans:
fixes g :: "real \<Rightarrow> 'a::{complete_space,real_normed_vector}"
assumes "arc g" "arc h" "pathfinish g = pathstart h" "pathstart g \<noteq> pathfinish h"
obtains i where "arc i" "path_image i \<subseteq> path_image g \<union> path_image h"
"pathstart i = pathstart g" "pathfinish i = pathfinish h"
- by%unimportant (metis (no_types, hide_lams) arc_imp_path assms path_contains_arc path_image_join path_join pathfinish_join pathstart_join)
+ by (metis (no_types, hide_lams) arc_imp_path assms path_contains_arc path_image_join path_join pathfinish_join pathstart_join)
subsection%important\<open>Accessibility of frontier points\<close>
-lemma%important dense_accessible_frontier_points:
+lemma dense_accessible_frontier_points:
fixes S :: "'a::{complete_space,real_normed_vector} set"
assumes "open S" and opeSV: "openin (subtopology euclidean (frontier S)) V" and "V \<noteq> {}"
obtains g where "arc g" "g ` {0..<1} \<subseteq> S" "pathstart g \<in> S" "pathfinish g \<in> V"
-proof%unimportant -
+proof -
obtain z where "z \<in> V"
using \<open>V \<noteq> {}\<close> by auto
then obtain r where "r > 0" and r: "ball z r \<inter> frontier S \<subseteq> V"
@@ -2098,12 +2098,12 @@
qed
-lemma%important dense_accessible_frontier_points_connected:
+lemma dense_accessible_frontier_points_connected:
fixes S :: "'a::{complete_space,real_normed_vector} set"
assumes "open S" "connected S" "x \<in> S" "V \<noteq> {}"
and ope: "openin (subtopology euclidean (frontier S)) V"
obtains g where "arc g" "g ` {0..<1} \<subseteq> S" "pathstart g = x" "pathfinish g \<in> V"
-proof%unimportant -
+proof -
have "V \<subseteq> frontier S"
using ope openin_imp_subset by blast
with \<open>open S\<close> \<open>x \<in> S\<close> have "x \<notin> V"
@@ -2133,14 +2133,14 @@
using h \<open>pathfinish g \<in> V\<close> by auto
qed
-lemma%important dense_access_fp_aux:
+lemma dense_access_fp_aux:
fixes S :: "'a::{complete_space,real_normed_vector} set"
assumes S: "open S" "connected S"
and opeSU: "openin (subtopology euclidean (frontier S)) U"
and opeSV: "openin (subtopology euclidean (frontier S)) V"
and "V \<noteq> {}" "\<not> U \<subseteq> V"
obtains g where "arc g" "pathstart g \<in> U" "pathfinish g \<in> V" "g ` {0<..<1} \<subseteq> S"
-proof%unimportant -
+proof -
have "S \<noteq> {}"
using opeSV \<open>V \<noteq> {}\<close> by (metis frontier_empty openin_subtopology_empty)
then obtain x where "x \<in> S" by auto
@@ -2180,14 +2180,14 @@
qed
qed
-lemma%important dense_accessible_frontier_point_pairs:
+lemma dense_accessible_frontier_point_pairs:
fixes S :: "'a::{complete_space,real_normed_vector} set"
assumes S: "open S" "connected S"
and opeSU: "openin (subtopology euclidean (frontier S)) U"
and opeSV: "openin (subtopology euclidean (frontier S)) V"
and "U \<noteq> {}" "V \<noteq> {}" "U \<noteq> V"
obtains g where "arc g" "pathstart g \<in> U" "pathfinish g \<in> V" "g ` {0<..<1} \<subseteq> S"
-proof%unimportant -
+proof -
consider "\<not> U \<subseteq> V" | "\<not> V \<subseteq> U"
using \<open>U \<noteq> V\<close> by blast
then show ?thesis
--- a/src/HOL/Analysis/Binary_Product_Measure.thy Sun Jan 13 20:25:41 2019 +0100
+++ b/src/HOL/Analysis/Binary_Product_Measure.thy Mon Jan 14 11:59:19 2019 +0000
@@ -8,10 +8,10 @@
imports Nonnegative_Lebesgue_Integration
begin
-lemma%unimportant Pair_vimage_times[simp]: "Pair x -` (A \<times> B) = (if x \<in> A then B else {})"
+lemma Pair_vimage_times[simp]: "Pair x -` (A \<times> B) = (if x \<in> A then B else {})"
by auto
-lemma%unimportant rev_Pair_vimage_times[simp]: "(\<lambda>x. (x, y)) -` (A \<times> B) = (if y \<in> B then A else {})"
+lemma rev_Pair_vimage_times[simp]: "(\<lambda>x. (x, y)) -` (A \<times> B) = (if y \<in> B then A else {})"
by auto
subsection%important "Binary products"
@@ -21,48 +21,48 @@
{a \<times> b | a b. a \<in> sets A \<and> b \<in> sets B}
(\<lambda>X. \<integral>\<^sup>+x. (\<integral>\<^sup>+y. indicator X (x,y) \<partial>B) \<partial>A)"
-lemma%important pair_measure_closed: "{a \<times> b | a b. a \<in> sets A \<and> b \<in> sets B} \<subseteq> Pow (space A \<times> space B)"
- using%unimportant sets.space_closed[of A] sets.space_closed[of B] by auto
+lemma pair_measure_closed: "{a \<times> b | a b. a \<in> sets A \<and> b \<in> sets B} \<subseteq> Pow (space A \<times> space B)"
+ using sets.space_closed[of A] sets.space_closed[of B] by auto
-lemma%important space_pair_measure:
+lemma space_pair_measure:
"space (A \<Otimes>\<^sub>M B) = space A \<times> space B"
unfolding pair_measure_def using pair_measure_closed[of A B]
- by%unimportant (rule space_measure_of)
+ by (rule space_measure_of)
-lemma%unimportant SIGMA_Collect_eq: "(SIGMA x:space M. {y\<in>space N. P x y}) = {x\<in>space (M \<Otimes>\<^sub>M N). P (fst x) (snd x)}"
+lemma SIGMA_Collect_eq: "(SIGMA x:space M. {y\<in>space N. P x y}) = {x\<in>space (M \<Otimes>\<^sub>M N). P (fst x) (snd x)}"
by (auto simp: space_pair_measure)
-lemma%unimportant sets_pair_measure:
+lemma sets_pair_measure:
"sets (A \<Otimes>\<^sub>M B) = sigma_sets (space A \<times> space B) {a \<times> b | a b. a \<in> sets A \<and> b \<in> sets B}"
unfolding pair_measure_def using pair_measure_closed[of A B]
by (rule sets_measure_of)
-lemma%unimportant sets_pair_measure_cong[measurable_cong, cong]:
+lemma sets_pair_measure_cong[measurable_cong, cong]:
"sets M1 = sets M1' \<Longrightarrow> sets M2 = sets M2' \<Longrightarrow> sets (M1 \<Otimes>\<^sub>M M2) = sets (M1' \<Otimes>\<^sub>M M2')"
unfolding sets_pair_measure by (simp cong: sets_eq_imp_space_eq)
-lemma%unimportant pair_measureI[intro, simp, measurable]:
+lemma pair_measureI[intro, simp, measurable]:
"x \<in> sets A \<Longrightarrow> y \<in> sets B \<Longrightarrow> x \<times> y \<in> sets (A \<Otimes>\<^sub>M B)"
by (auto simp: sets_pair_measure)
-lemma%unimportant sets_Pair: "{x} \<in> sets M1 \<Longrightarrow> {y} \<in> sets M2 \<Longrightarrow> {(x, y)} \<in> sets (M1 \<Otimes>\<^sub>M M2)"
+lemma sets_Pair: "{x} \<in> sets M1 \<Longrightarrow> {y} \<in> sets M2 \<Longrightarrow> {(x, y)} \<in> sets (M1 \<Otimes>\<^sub>M M2)"
using pair_measureI[of "{x}" M1 "{y}" M2] by simp
-lemma%unimportant measurable_pair_measureI:
+lemma measurable_pair_measureI:
assumes 1: "f \<in> space M \<rightarrow> space M1 \<times> space M2"
assumes 2: "\<And>A B. A \<in> sets M1 \<Longrightarrow> B \<in> sets M2 \<Longrightarrow> f -` (A \<times> B) \<inter> space M \<in> sets M"
shows "f \<in> measurable M (M1 \<Otimes>\<^sub>M M2)"
unfolding pair_measure_def using 1 2
by (intro measurable_measure_of) (auto dest: sets.sets_into_space)
-lemma%unimportant measurable_split_replace[measurable (raw)]:
+lemma measurable_split_replace[measurable (raw)]:
"(\<lambda>x. f x (fst (g x)) (snd (g x))) \<in> measurable M N \<Longrightarrow> (\<lambda>x. case_prod (f x) (g x)) \<in> measurable M N"
unfolding split_beta' .
-lemma%important measurable_Pair[measurable (raw)]:
+lemma measurable_Pair[measurable (raw)]:
assumes f: "f \<in> measurable M M1" and g: "g \<in> measurable M M2"
shows "(\<lambda>x. (f x, g x)) \<in> measurable M (M1 \<Otimes>\<^sub>M M2)"
-proof%unimportant (rule measurable_pair_measureI)
+proof (rule measurable_pair_measureI)
show "(\<lambda>x. (f x, g x)) \<in> space M \<rightarrow> space M1 \<times> space M2"
using f g by (auto simp: measurable_def)
fix A B assume *: "A \<in> sets M1" "B \<in> sets M2"
@@ -73,59 +73,59 @@
finally show "(\<lambda>x. (f x, g x)) -` (A \<times> B) \<inter> space M \<in> sets M" .
qed
-lemma%unimportant measurable_fst[intro!, simp, measurable]: "fst \<in> measurable (M1 \<Otimes>\<^sub>M M2) M1"
+lemma measurable_fst[intro!, simp, measurable]: "fst \<in> measurable (M1 \<Otimes>\<^sub>M M2) M1"
by (auto simp: fst_vimage_eq_Times space_pair_measure sets.sets_into_space times_Int_times
measurable_def)
-lemma%unimportant measurable_snd[intro!, simp, measurable]: "snd \<in> measurable (M1 \<Otimes>\<^sub>M M2) M2"
+lemma measurable_snd[intro!, simp, measurable]: "snd \<in> measurable (M1 \<Otimes>\<^sub>M M2) M2"
by (auto simp: snd_vimage_eq_Times space_pair_measure sets.sets_into_space times_Int_times
measurable_def)
-lemma%unimportant measurable_Pair_compose_split[measurable_dest]:
+lemma measurable_Pair_compose_split[measurable_dest]:
assumes f: "case_prod f \<in> measurable (M1 \<Otimes>\<^sub>M M2) N"
assumes g: "g \<in> measurable M M1" and h: "h \<in> measurable M M2"
shows "(\<lambda>x. f (g x) (h x)) \<in> measurable M N"
using measurable_compose[OF measurable_Pair f, OF g h] by simp
-lemma%unimportant measurable_Pair1_compose[measurable_dest]:
+lemma measurable_Pair1_compose[measurable_dest]:
assumes f: "(\<lambda>x. (f x, g x)) \<in> measurable M (M1 \<Otimes>\<^sub>M M2)"
assumes [measurable]: "h \<in> measurable N M"
shows "(\<lambda>x. f (h x)) \<in> measurable N M1"
using measurable_compose[OF f measurable_fst] by simp
-lemma%unimportant measurable_Pair2_compose[measurable_dest]:
+lemma measurable_Pair2_compose[measurable_dest]:
assumes f: "(\<lambda>x. (f x, g x)) \<in> measurable M (M1 \<Otimes>\<^sub>M M2)"
assumes [measurable]: "h \<in> measurable N M"
shows "(\<lambda>x. g (h x)) \<in> measurable N M2"
using measurable_compose[OF f measurable_snd] by simp
-lemma%unimportant measurable_pair:
+lemma measurable_pair:
assumes "(fst \<circ> f) \<in> measurable M M1" "(snd \<circ> f) \<in> measurable M M2"
shows "f \<in> measurable M (M1 \<Otimes>\<^sub>M M2)"
using measurable_Pair[OF assms] by simp
-lemma%unimportant (*FIX ME needs a name *)
+lemma (*FIX ME needs a name *)
assumes f[measurable]: "f \<in> measurable M (N \<Otimes>\<^sub>M P)"
shows measurable_fst': "(\<lambda>x. fst (f x)) \<in> measurable M N"
and measurable_snd': "(\<lambda>x. snd (f x)) \<in> measurable M P"
by simp_all
-lemma%unimportant (*FIX ME needs a name *)
+lemma (*FIX ME needs a name *)
assumes f[measurable]: "f \<in> measurable M N"
shows measurable_fst'': "(\<lambda>x. f (fst x)) \<in> measurable (M \<Otimes>\<^sub>M P) N"
and measurable_snd'': "(\<lambda>x. f (snd x)) \<in> measurable (P \<Otimes>\<^sub>M M) N"
by simp_all
-lemma%unimportant sets_pair_in_sets:
+lemma sets_pair_in_sets:
assumes "\<And>a b. a \<in> sets A \<Longrightarrow> b \<in> sets B \<Longrightarrow> a \<times> b \<in> sets N"
shows "sets (A \<Otimes>\<^sub>M B) \<subseteq> sets N"
unfolding sets_pair_measure
by (intro sets.sigma_sets_subset') (auto intro!: assms)
-lemma%important sets_pair_eq_sets_fst_snd:
+lemma sets_pair_eq_sets_fst_snd:
"sets (A \<Otimes>\<^sub>M B) = sets (Sup {vimage_algebra (space A \<times> space B) fst A, vimage_algebra (space A \<times> space B) snd B})"
(is "?P = sets (Sup {?fst, ?snd})")
-proof%unimportant -
+proof -
{ fix a b assume ab: "a \<in> sets A" "b \<in> sets B"
then have "a \<times> b = (fst -` a \<inter> (space A \<times> space B)) \<inter> (snd -` b \<inter> (space A \<times> space B))"
by (auto dest: sets.sets_into_space)
@@ -157,29 +157,29 @@
done
qed
-lemma%unimportant measurable_pair_iff:
+lemma measurable_pair_iff:
"f \<in> measurable M (M1 \<Otimes>\<^sub>M M2) \<longleftrightarrow> (fst \<circ> f) \<in> measurable M M1 \<and> (snd \<circ> f) \<in> measurable M M2"
by (auto intro: measurable_pair[of f M M1 M2])
-lemma%unimportant measurable_split_conv:
+lemma measurable_split_conv:
"(\<lambda>(x, y). f x y) \<in> measurable A B \<longleftrightarrow> (\<lambda>x. f (fst x) (snd x)) \<in> measurable A B"
by (intro arg_cong2[where f="(\<in>)"]) auto
-lemma%unimportant measurable_pair_swap': "(\<lambda>(x,y). (y, x)) \<in> measurable (M1 \<Otimes>\<^sub>M M2) (M2 \<Otimes>\<^sub>M M1)"
+lemma measurable_pair_swap': "(\<lambda>(x,y). (y, x)) \<in> measurable (M1 \<Otimes>\<^sub>M M2) (M2 \<Otimes>\<^sub>M M1)"
by (auto intro!: measurable_Pair simp: measurable_split_conv)
-lemma%unimportant measurable_pair_swap:
+lemma measurable_pair_swap:
assumes f: "f \<in> measurable (M1 \<Otimes>\<^sub>M M2) M" shows "(\<lambda>(x,y). f (y, x)) \<in> measurable (M2 \<Otimes>\<^sub>M M1) M"
using measurable_comp[OF measurable_Pair f] by (auto simp: measurable_split_conv comp_def)
-lemma%unimportant measurable_pair_swap_iff:
+lemma measurable_pair_swap_iff:
"f \<in> measurable (M2 \<Otimes>\<^sub>M M1) M \<longleftrightarrow> (\<lambda>(x,y). f (y,x)) \<in> measurable (M1 \<Otimes>\<^sub>M M2) M"
by (auto dest: measurable_pair_swap)
-lemma%unimportant measurable_Pair1': "x \<in> space M1 \<Longrightarrow> Pair x \<in> measurable M2 (M1 \<Otimes>\<^sub>M M2)"
+lemma measurable_Pair1': "x \<in> space M1 \<Longrightarrow> Pair x \<in> measurable M2 (M1 \<Otimes>\<^sub>M M2)"
by simp
-lemma%unimportant sets_Pair1[measurable (raw)]:
+lemma sets_Pair1[measurable (raw)]:
assumes A: "A \<in> sets (M1 \<Otimes>\<^sub>M M2)" shows "Pair x -` A \<in> sets M2"
proof -
have "Pair x -` A = (if x \<in> space M1 then Pair x -` A \<inter> space M2 else {})"
@@ -189,10 +189,10 @@
finally show ?thesis .
qed
-lemma%unimportant measurable_Pair2': "y \<in> space M2 \<Longrightarrow> (\<lambda>x. (x, y)) \<in> measurable M1 (M1 \<Otimes>\<^sub>M M2)"
+lemma measurable_Pair2': "y \<in> space M2 \<Longrightarrow> (\<lambda>x. (x, y)) \<in> measurable M1 (M1 \<Otimes>\<^sub>M M2)"
by (auto intro!: measurable_Pair)
-lemma%unimportant sets_Pair2: assumes A: "A \<in> sets (M1 \<Otimes>\<^sub>M M2)" shows "(\<lambda>x. (x, y)) -` A \<in> sets M1"
+lemma sets_Pair2: assumes A: "A \<in> sets (M1 \<Otimes>\<^sub>M M2)" shows "(\<lambda>x. (x, y)) -` A \<in> sets M1"
proof -
have "(\<lambda>x. (x, y)) -` A = (if y \<in> space M2 then (\<lambda>x. (x, y)) -` A \<inter> space M1 else {})"
using A[THEN sets.sets_into_space] by (auto simp: space_pair_measure)
@@ -201,23 +201,23 @@
finally show ?thesis .
qed
-lemma%unimportant measurable_Pair2:
+lemma measurable_Pair2:
assumes f: "f \<in> measurable (M1 \<Otimes>\<^sub>M M2) M" and x: "x \<in> space M1"
shows "(\<lambda>y. f (x, y)) \<in> measurable M2 M"
using measurable_comp[OF measurable_Pair1' f, OF x]
by (simp add: comp_def)
-lemma%unimportant measurable_Pair1:
+lemma measurable_Pair1:
assumes f: "f \<in> measurable (M1 \<Otimes>\<^sub>M M2) M" and y: "y \<in> space M2"
shows "(\<lambda>x. f (x, y)) \<in> measurable M1 M"
using measurable_comp[OF measurable_Pair2' f, OF y]
by (simp add: comp_def)
-lemma%unimportant Int_stable_pair_measure_generator: "Int_stable {a \<times> b | a b. a \<in> sets A \<and> b \<in> sets B}"
+lemma Int_stable_pair_measure_generator: "Int_stable {a \<times> b | a b. a \<in> sets A \<and> b \<in> sets B}"
unfolding Int_stable_def
by safe (auto simp add: times_Int_times)
-lemma%unimportant (in finite_measure) finite_measure_cut_measurable:
+lemma (in finite_measure) finite_measure_cut_measurable:
assumes [measurable]: "Q \<in> sets (N \<Otimes>\<^sub>M M)"
shows "(\<lambda>x. emeasure M (Pair x -` Q)) \<in> borel_measurable N"
(is "?s Q \<in> _")
@@ -239,7 +239,7 @@
unfolding sets_pair_measure[symmetric] by simp
qed (auto simp add: if_distrib Int_def[symmetric] intro!: measurable_If)
-lemma%unimportant (in sigma_finite_measure) measurable_emeasure_Pair:
+lemma (in sigma_finite_measure) measurable_emeasure_Pair:
assumes Q: "Q \<in> sets (N \<Otimes>\<^sub>M M)" shows "(\<lambda>x. emeasure M (Pair x -` Q)) \<in> borel_measurable N" (is "?s Q \<in> _")
proof -
from sigma_finite_disjoint guess F . note F = this
@@ -279,7 +279,7 @@
by auto
qed
-lemma%unimportant (in sigma_finite_measure) measurable_emeasure[measurable (raw)]:
+lemma (in sigma_finite_measure) measurable_emeasure[measurable (raw)]:
assumes space: "\<And>x. x \<in> space N \<Longrightarrow> A x \<subseteq> space M"
assumes A: "{x\<in>space (N \<Otimes>\<^sub>M M). snd x \<in> A (fst x)} \<in> sets (N \<Otimes>\<^sub>M M)"
shows "(\<lambda>x. emeasure M (A x)) \<in> borel_measurable N"
@@ -290,7 +290,7 @@
by (auto cong: measurable_cong)
qed
-lemma%unimportant (in sigma_finite_measure) emeasure_pair_measure:
+lemma (in sigma_finite_measure) emeasure_pair_measure:
assumes "X \<in> sets (N \<Otimes>\<^sub>M M)"
shows "emeasure (N \<Otimes>\<^sub>M M) X = (\<integral>\<^sup>+ x. \<integral>\<^sup>+ y. indicator X (x, y) \<partial>M \<partial>N)" (is "_ = ?\<mu> X")
proof (rule emeasure_measure_of[OF pair_measure_def])
@@ -314,7 +314,7 @@
using sets.space_closed[of N] sets.space_closed[of M] by auto
qed fact
-lemma%unimportant (in sigma_finite_measure) emeasure_pair_measure_alt:
+lemma (in sigma_finite_measure) emeasure_pair_measure_alt:
assumes X: "X \<in> sets (N \<Otimes>\<^sub>M M)"
shows "emeasure (N \<Otimes>\<^sub>M M) X = (\<integral>\<^sup>+x. emeasure M (Pair x -` X) \<partial>N)"
proof -
@@ -324,10 +324,10 @@
using X by (auto intro!: nn_integral_cong simp: emeasure_pair_measure sets_Pair1)
qed
-lemma%important (in sigma_finite_measure) emeasure_pair_measure_Times:
+proposition (in sigma_finite_measure) emeasure_pair_measure_Times:
assumes A: "A \<in> sets N" and B: "B \<in> sets M"
shows "emeasure (N \<Otimes>\<^sub>M M) (A \<times> B) = emeasure N A * emeasure M B"
-proof%unimportant -
+proof -
have "emeasure (N \<Otimes>\<^sub>M M) (A \<times> B) = (\<integral>\<^sup>+x. emeasure M B * indicator A x \<partial>N)"
using A B by (auto intro!: nn_integral_cong simp: emeasure_pair_measure_alt)
also have "\<dots> = emeasure M B * emeasure N A"
@@ -338,16 +338,16 @@
subsection%important \<open>Binary products of \<open>\<sigma>\<close>-finite emeasure spaces\<close>
-locale%important pair_sigma_finite = M1?: sigma_finite_measure M1 + M2?: sigma_finite_measure M2
+locale%unimportant pair_sigma_finite = M1?: sigma_finite_measure M1 + M2?: sigma_finite_measure M2
for M1 :: "'a measure" and M2 :: "'b measure"
-lemma%unimportant (in pair_sigma_finite) measurable_emeasure_Pair1:
+lemma (in pair_sigma_finite) measurable_emeasure_Pair1:
"Q \<in> sets (M1 \<Otimes>\<^sub>M M2) \<Longrightarrow> (\<lambda>x. emeasure M2 (Pair x -` Q)) \<in> borel_measurable M1"
using M2.measurable_emeasure_Pair .
-lemma%important (in pair_sigma_finite) measurable_emeasure_Pair2:
+lemma (in pair_sigma_finite) measurable_emeasure_Pair2:
assumes Q: "Q \<in> sets (M1 \<Otimes>\<^sub>M M2)" shows "(\<lambda>y. emeasure M1 ((\<lambda>x. (x, y)) -` Q)) \<in> borel_measurable M2"
-proof%unimportant -
+proof -
have "(\<lambda>(x, y). (y, x)) -` Q \<inter> space (M2 \<Otimes>\<^sub>M M1) \<in> sets (M2 \<Otimes>\<^sub>M M1)"
using Q measurable_pair_swap' by (auto intro: measurable_sets)
note M1.measurable_emeasure_Pair[OF this]
@@ -356,11 +356,11 @@
ultimately show ?thesis by simp
qed
-lemma%important (in pair_sigma_finite) sigma_finite_up_in_pair_measure_generator:
+proposition (in pair_sigma_finite) sigma_finite_up_in_pair_measure_generator:
defines "E \<equiv> {A \<times> B | A B. A \<in> sets M1 \<and> B \<in> sets M2}"
shows "\<exists>F::nat \<Rightarrow> ('a \<times> 'b) set. range F \<subseteq> E \<and> incseq F \<and> (\<Union>i. F i) = space M1 \<times> space M2 \<and>
(\<forall>i. emeasure (M1 \<Otimes>\<^sub>M M2) (F i) \<noteq> \<infinity>)"
-proof%unimportant -
+proof -
from M1.sigma_finite_incseq guess F1 . note F1 = this
from M2.sigma_finite_incseq guess F2 . note F2 = this
from F1 F2 have space: "space M1 = (\<Union>i. F1 i)" "space M2 = (\<Union>i. F2 i)" by auto
@@ -394,7 +394,7 @@
qed
qed
-sublocale%important pair_sigma_finite \<subseteq> P?: sigma_finite_measure "M1 \<Otimes>\<^sub>M M2"
+sublocale%unimportant pair_sigma_finite \<subseteq> P?: sigma_finite_measure "M1 \<Otimes>\<^sub>M M2"
proof
from M1.sigma_finite_countable guess F1 ..
moreover from M2.sigma_finite_countable guess F2 ..
@@ -404,7 +404,7 @@
(auto simp: M2.emeasure_pair_measure_Times space_pair_measure set_eq_iff subset_eq ennreal_mult_eq_top_iff)
qed
-lemma%unimportant sigma_finite_pair_measure:
+lemma sigma_finite_pair_measure:
assumes A: "sigma_finite_measure A" and B: "sigma_finite_measure B"
shows "sigma_finite_measure (A \<Otimes>\<^sub>M B)"
proof -
@@ -414,14 +414,14 @@
show ?thesis ..
qed
-lemma%unimportant sets_pair_swap:
+lemma sets_pair_swap:
assumes "A \<in> sets (M1 \<Otimes>\<^sub>M M2)"
shows "(\<lambda>(x, y). (y, x)) -` A \<inter> space (M2 \<Otimes>\<^sub>M M1) \<in> sets (M2 \<Otimes>\<^sub>M M1)"
using measurable_pair_swap' assms by (rule measurable_sets)
-lemma%important (in pair_sigma_finite) distr_pair_swap:
+lemma (in pair_sigma_finite) distr_pair_swap:
"M1 \<Otimes>\<^sub>M M2 = distr (M2 \<Otimes>\<^sub>M M1) (M1 \<Otimes>\<^sub>M M2) (\<lambda>(x, y). (y, x))" (is "?P = ?D")
-proof%unimportant -
+proof -
from sigma_finite_up_in_pair_measure_generator guess F :: "nat \<Rightarrow> ('a \<times> 'b) set" .. note F = this
let ?E = "{a \<times> b |a b. a \<in> sets M1 \<and> b \<in> sets M2}"
show ?thesis
@@ -446,11 +446,11 @@
qed
qed
-lemma%unimportant (in pair_sigma_finite) emeasure_pair_measure_alt2:
+lemma (in pair_sigma_finite) emeasure_pair_measure_alt2:
assumes A: "A \<in> sets (M1 \<Otimes>\<^sub>M M2)"
shows "emeasure (M1 \<Otimes>\<^sub>M M2) A = (\<integral>\<^sup>+y. emeasure M1 ((\<lambda>x. (x, y)) -` A) \<partial>M2)"
(is "_ = ?\<nu> A")
-proof%unimportant -
+proof -
have [simp]: "\<And>y. (Pair y -` ((\<lambda>(x, y). (y, x)) -` A \<inter> space (M2 \<Otimes>\<^sub>M M1))) = (\<lambda>x. (x, y)) -` A"
using sets.sets_into_space[OF A] by (auto simp: space_pair_measure)
show ?thesis using A
@@ -459,7 +459,7 @@
M1.emeasure_pair_measure_alt emeasure_distr[OF measurable_pair_swap' A])
qed
-lemma%unimportant (in pair_sigma_finite) AE_pair:
+lemma (in pair_sigma_finite) AE_pair:
assumes "AE x in (M1 \<Otimes>\<^sub>M M2). Q x"
shows "AE x in M1. (AE y in M2. Q (x, y))"
proof -
@@ -485,11 +485,11 @@
qed
qed
-lemma%important (in pair_sigma_finite) AE_pair_measure:
+lemma (in pair_sigma_finite) AE_pair_measure:
assumes "{x\<in>space (M1 \<Otimes>\<^sub>M M2). P x} \<in> sets (M1 \<Otimes>\<^sub>M M2)"
assumes ae: "AE x in M1. AE y in M2. P (x, y)"
shows "AE x in M1 \<Otimes>\<^sub>M M2. P x"
-proof%unimportant (subst AE_iff_measurable[OF _ refl])
+proof (subst AE_iff_measurable[OF _ refl])
show "{x\<in>space (M1 \<Otimes>\<^sub>M M2). \<not> P x} \<in> sets (M1 \<Otimes>\<^sub>M M2)"
by (rule sets.sets_Collect) fact
then have "emeasure (M1 \<Otimes>\<^sub>M M2) {x \<in> space (M1 \<Otimes>\<^sub>M M2). \<not> P x} =
@@ -505,12 +505,12 @@
finally show "emeasure (M1 \<Otimes>\<^sub>M M2) {x \<in> space (M1 \<Otimes>\<^sub>M M2). \<not> P x} = 0" by simp
qed
-lemma%unimportant (in pair_sigma_finite) AE_pair_iff:
+lemma (in pair_sigma_finite) AE_pair_iff:
"{x\<in>space (M1 \<Otimes>\<^sub>M M2). P (fst x) (snd x)} \<in> sets (M1 \<Otimes>\<^sub>M M2) \<Longrightarrow>
(AE x in M1. AE y in M2. P x y) \<longleftrightarrow> (AE x in (M1 \<Otimes>\<^sub>M M2). P (fst x) (snd x))"
using AE_pair[of "\<lambda>x. P (fst x) (snd x)"] AE_pair_measure[of "\<lambda>x. P (fst x) (snd x)"] by auto
-lemma%unimportant (in pair_sigma_finite) AE_commute:
+lemma (in pair_sigma_finite) AE_commute:
assumes P: "{x\<in>space (M1 \<Otimes>\<^sub>M M2). P (fst x) (snd x)} \<in> sets (M1 \<Otimes>\<^sub>M M2)"
shows "(AE x in M1. AE y in M2. P x y) \<longleftrightarrow> (AE y in M2. AE x in M1. P x y)"
proof -
@@ -533,14 +533,14 @@
subsection%important "Fubinis theorem"
-lemma%unimportant measurable_compose_Pair1:
+lemma measurable_compose_Pair1:
"x \<in> space M1 \<Longrightarrow> g \<in> measurable (M1 \<Otimes>\<^sub>M M2) L \<Longrightarrow> (\<lambda>y. g (x, y)) \<in> measurable M2 L"
by simp
-lemma%unimportant (in sigma_finite_measure) borel_measurable_nn_integral_fst:
+lemma (in sigma_finite_measure) borel_measurable_nn_integral_fst:
assumes f: "f \<in> borel_measurable (M1 \<Otimes>\<^sub>M M)"
shows "(\<lambda>x. \<integral>\<^sup>+ y. f (x, y) \<partial>M) \<in> borel_measurable M1"
-using f proof%unimportant induct
+using f proof induct
case (cong u v)
then have "\<And>w x. w \<in> space M1 \<Longrightarrow> x \<in> space M \<Longrightarrow> u (w, x) = v (w, x)"
by (auto simp: space_pair_measure)
@@ -561,7 +561,7 @@
nn_integral_monotone_convergence_SUP incseq_def le_fun_def
cong: measurable_cong)
-lemma%unimportant (in sigma_finite_measure) nn_integral_fst:
+lemma (in sigma_finite_measure) nn_integral_fst:
assumes f: "f \<in> borel_measurable (M1 \<Otimes>\<^sub>M M)"
shows "(\<integral>\<^sup>+ x. \<integral>\<^sup>+ y. f (x, y) \<partial>M \<partial>M1) = integral\<^sup>N (M1 \<Otimes>\<^sub>M M) f" (is "?I f = _")
using f proof induct
@@ -575,14 +575,14 @@
borel_measurable_nn_integral_fst nn_integral_mono incseq_def le_fun_def
cong: nn_integral_cong)
-lemma%unimportant (in sigma_finite_measure) borel_measurable_nn_integral[measurable (raw)]:
+lemma (in sigma_finite_measure) borel_measurable_nn_integral[measurable (raw)]:
"case_prod f \<in> borel_measurable (N \<Otimes>\<^sub>M M) \<Longrightarrow> (\<lambda>x. \<integral>\<^sup>+ y. f x y \<partial>M) \<in> borel_measurable N"
using borel_measurable_nn_integral_fst[of "case_prod f" N] by simp
-lemma%important (in pair_sigma_finite) nn_integral_snd:
+proposition (in pair_sigma_finite) nn_integral_snd:
assumes f[measurable]: "f \<in> borel_measurable (M1 \<Otimes>\<^sub>M M2)"
shows "(\<integral>\<^sup>+ y. (\<integral>\<^sup>+ x. f (x, y) \<partial>M1) \<partial>M2) = integral\<^sup>N (M1 \<Otimes>\<^sub>M M2) f"
-proof%unimportant -
+proof -
note measurable_pair_swap[OF f]
from M1.nn_integral_fst[OF this]
have "(\<integral>\<^sup>+ y. (\<integral>\<^sup>+ x. f (x, y) \<partial>M1) \<partial>M2) = (\<integral>\<^sup>+ (x, y). f (y, x) \<partial>(M2 \<Otimes>\<^sub>M M1))"
@@ -592,24 +592,24 @@
finally show ?thesis .
qed
-lemma%important (in pair_sigma_finite) Fubini:
+theorem (in pair_sigma_finite) Fubini:
assumes f: "f \<in> borel_measurable (M1 \<Otimes>\<^sub>M M2)"
shows "(\<integral>\<^sup>+ y. (\<integral>\<^sup>+ x. f (x, y) \<partial>M1) \<partial>M2) = (\<integral>\<^sup>+ x. (\<integral>\<^sup>+ y. f (x, y) \<partial>M2) \<partial>M1)"
unfolding nn_integral_snd[OF assms] M2.nn_integral_fst[OF assms] ..
-lemma%important (in pair_sigma_finite) Fubini':
+theorem (in pair_sigma_finite) Fubini':
assumes f: "case_prod f \<in> borel_measurable (M1 \<Otimes>\<^sub>M M2)"
shows "(\<integral>\<^sup>+ y. (\<integral>\<^sup>+ x. f x y \<partial>M1) \<partial>M2) = (\<integral>\<^sup>+ x. (\<integral>\<^sup>+ y. f x y \<partial>M2) \<partial>M1)"
using Fubini[OF f] by simp
subsection%important \<open>Products on counting spaces, densities and distributions\<close>
-lemma%important sigma_prod:
+proposition sigma_prod:
assumes X_cover: "\<exists>E\<subseteq>A. countable E \<and> X = \<Union>E" and A: "A \<subseteq> Pow X"
assumes Y_cover: "\<exists>E\<subseteq>B. countable E \<and> Y = \<Union>E" and B: "B \<subseteq> Pow Y"
shows "sigma X A \<Otimes>\<^sub>M sigma Y B = sigma (X \<times> Y) {a \<times> b | a b. a \<in> A \<and> b \<in> B}"
(is "?P = ?S")
-proof%unimportant (rule measure_eqI)
+proof (rule measure_eqI)
have [simp]: "snd \<in> X \<times> Y \<rightarrow> Y" "fst \<in> X \<times> Y \<rightarrow> X"
by auto
let ?XY = "{{fst -` a \<inter> X \<times> Y | a. a \<in> A}, {snd -` b \<inter> X \<times> Y | b. b \<in> B}}"
@@ -662,7 +662,7 @@
by (simp add: emeasure_pair_measure_alt emeasure_sigma)
qed
-lemma%unimportant sigma_sets_pair_measure_generator_finite:
+lemma sigma_sets_pair_measure_generator_finite:
assumes "finite A" and "finite B"
shows "sigma_sets (A \<times> B) { a \<times> b | a b. a \<subseteq> A \<and> b \<subseteq> B} = Pow (A \<times> B)"
(is "sigma_sets ?prod ?sets = _")
@@ -686,14 +686,14 @@
show "a \<in> A" and "b \<in> B" by auto
qed
-lemma%important sets_pair_eq:
+proposition sets_pair_eq:
assumes Ea: "Ea \<subseteq> Pow (space A)" "sets A = sigma_sets (space A) Ea"
and Ca: "countable Ca" "Ca \<subseteq> Ea" "\<Union>Ca = space A"
and Eb: "Eb \<subseteq> Pow (space B)" "sets B = sigma_sets (space B) Eb"
and Cb: "countable Cb" "Cb \<subseteq> Eb" "\<Union>Cb = space B"
shows "sets (A \<Otimes>\<^sub>M B) = sets (sigma (space A \<times> space B) { a \<times> b | a b. a \<in> Ea \<and> b \<in> Eb })"
(is "_ = sets (sigma ?\<Omega> ?E)")
-proof%unimportant
+proof
show "sets (sigma ?\<Omega> ?E) \<subseteq> sets (A \<Otimes>\<^sub>M B)"
using Ea(1) Eb(1) by (subst sigma_le_sets) (auto simp: Ea(2) Eb(2))
have "?E \<subseteq> Pow ?\<Omega>"
@@ -733,10 +733,10 @@
finally show "sets (A \<Otimes>\<^sub>M B) \<subseteq> sets (sigma ?\<Omega> ?E)" .
qed
-lemma%important borel_prod:
+proposition borel_prod:
"(borel \<Otimes>\<^sub>M borel) = (borel :: ('a::second_countable_topology \<times> 'b::second_countable_topology) measure)"
(is "?P = ?B")
-proof%unimportant -
+proof -
have "?B = sigma UNIV {A \<times> B | A B. open A \<and> open B}"
by (rule second_countable_borel_measurable[OF open_prod_generated])
also have "\<dots> = ?P"
@@ -745,10 +745,10 @@
finally show ?thesis ..
qed
-lemma%important pair_measure_count_space:
+proposition pair_measure_count_space:
assumes A: "finite A" and B: "finite B"
shows "count_space A \<Otimes>\<^sub>M count_space B = count_space (A \<times> B)" (is "?P = ?C")
-proof%unimportant (rule measure_eqI)
+proof (rule measure_eqI)
interpret A: finite_measure "count_space A" by (rule finite_measure_count_space) fact
interpret B: finite_measure "count_space B" by (rule finite_measure_count_space) fact
interpret P: pair_sigma_finite "count_space A" "count_space B" ..
@@ -776,14 +776,14 @@
qed
-lemma%unimportant emeasure_prod_count_space:
+lemma emeasure_prod_count_space:
assumes A: "A \<in> sets (count_space UNIV \<Otimes>\<^sub>M M)" (is "A \<in> sets (?A \<Otimes>\<^sub>M ?B)")
shows "emeasure (?A \<Otimes>\<^sub>M ?B) A = (\<integral>\<^sup>+ x. \<integral>\<^sup>+ y. indicator A (x, y) \<partial>?B \<partial>?A)"
by (rule emeasure_measure_of[OF pair_measure_def])
(auto simp: countably_additive_def positive_def suminf_indicator A
nn_integral_suminf[symmetric] dest: sets.sets_into_space)
-lemma%unimportant emeasure_prod_count_space_single[simp]: "emeasure (count_space UNIV \<Otimes>\<^sub>M count_space UNIV) {x} = 1"
+lemma emeasure_prod_count_space_single[simp]: "emeasure (count_space UNIV \<Otimes>\<^sub>M count_space UNIV) {x} = 1"
proof -
have [simp]: "\<And>a b x y. indicator {(a, b)} (x, y) = (indicator {a} x * indicator {b} y::ennreal)"
by (auto split: split_indicator)
@@ -791,11 +791,11 @@
by (cases x) (auto simp: emeasure_prod_count_space nn_integral_cmult sets_Pair)
qed
-lemma%important emeasure_count_space_prod_eq:
+lemma emeasure_count_space_prod_eq:
fixes A :: "('a \<times> 'b) set"
assumes A: "A \<in> sets (count_space UNIV \<Otimes>\<^sub>M count_space UNIV)" (is "A \<in> sets (?A \<Otimes>\<^sub>M ?B)")
shows "emeasure (?A \<Otimes>\<^sub>M ?B) A = emeasure (count_space UNIV) A"
-proof%unimportant -
+proof -
{ fix A :: "('a \<times> 'b) set" assume "countable A"
then have "emeasure (?A \<Otimes>\<^sub>M ?B) (\<Union>a\<in>A. {a}) = (\<integral>\<^sup>+a. emeasure (?A \<Otimes>\<^sub>M ?B) {a} \<partial>count_space A)"
by (intro emeasure_UN_countable) (auto simp: sets_Pair disjoint_family_on_def)
@@ -822,7 +822,7 @@
qed
qed
-lemma%unimportant nn_integral_count_space_prod_eq:
+lemma nn_integral_count_space_prod_eq:
"nn_integral (count_space UNIV \<Otimes>\<^sub>M count_space UNIV) f = nn_integral (count_space UNIV) f"
(is "nn_integral ?P f = _")
proof cases
@@ -874,12 +874,12 @@
by (simp add: top_unique)
qed
-lemma%important pair_measure_density:
+theorem pair_measure_density:
assumes f: "f \<in> borel_measurable M1"
assumes g: "g \<in> borel_measurable M2"
assumes "sigma_finite_measure M2" "sigma_finite_measure (density M2 g)"
shows "density M1 f \<Otimes>\<^sub>M density M2 g = density (M1 \<Otimes>\<^sub>M M2) (\<lambda>(x,y). f x * g y)" (is "?L = ?R")
-proof%unimportant (rule measure_eqI)
+proof (rule measure_eqI)
interpret M2: sigma_finite_measure M2 by fact
interpret D2: sigma_finite_measure "density M2 g" by fact
@@ -894,7 +894,7 @@
cong: nn_integral_cong)
qed simp
-lemma%unimportant sigma_finite_measure_distr:
+lemma sigma_finite_measure_distr:
assumes "sigma_finite_measure (distr M N f)" and f: "f \<in> measurable M N"
shows "sigma_finite_measure M"
proof -
@@ -909,7 +909,7 @@
qed
qed
-lemma%unimportant pair_measure_distr:
+lemma pair_measure_distr:
assumes f: "f \<in> measurable M S" and g: "g \<in> measurable N T"
assumes "sigma_finite_measure (distr N T g)"
shows "distr M S f \<Otimes>\<^sub>M distr N T g = distr (M \<Otimes>\<^sub>M N) (S \<Otimes>\<^sub>M T) (\<lambda>(x, y). (f x, g y))" (is "?P = ?D")
@@ -924,12 +924,12 @@
intro!: nn_integral_cong arg_cong[where f="emeasure N"])
qed simp
-lemma%important pair_measure_eqI:
+lemma pair_measure_eqI:
assumes "sigma_finite_measure M1" "sigma_finite_measure M2"
assumes sets: "sets (M1 \<Otimes>\<^sub>M M2) = sets M"
assumes emeasure: "\<And>A B. A \<in> sets M1 \<Longrightarrow> B \<in> sets M2 \<Longrightarrow> emeasure M1 A * emeasure M2 B = emeasure M (A \<times> B)"
shows "M1 \<Otimes>\<^sub>M M2 = M"
-proof%unimportant -
+proof -
interpret M1: sigma_finite_measure M1 by fact
interpret M2: sigma_finite_measure M2 by fact
interpret pair_sigma_finite M1 M2 ..
@@ -959,11 +959,11 @@
qed
qed
-lemma%important sets_pair_countable:
+lemma sets_pair_countable:
assumes "countable S1" "countable S2"
assumes M: "sets M = Pow S1" and N: "sets N = Pow S2"
shows "sets (M \<Otimes>\<^sub>M N) = Pow (S1 \<times> S2)"
-proof%unimportant auto
+proof auto
fix x a b assume x: "x \<in> sets (M \<Otimes>\<^sub>M N)" "(a, b) \<in> x"
from sets.sets_into_space[OF x(1)] x(2)
sets_eq_imp_space_eq[of N "count_space S2"] sets_eq_imp_space_eq[of M "count_space S1"] M N
@@ -980,10 +980,10 @@
finally show "X \<in> sets (M \<Otimes>\<^sub>M N)" .
qed
-lemma%important pair_measure_countable:
+lemma pair_measure_countable:
assumes "countable S1" "countable S2"
shows "count_space S1 \<Otimes>\<^sub>M count_space S2 = count_space (S1 \<times> S2)"
-proof%unimportant (rule pair_measure_eqI)
+proof (rule pair_measure_eqI)
show "sigma_finite_measure (count_space S1)" "sigma_finite_measure (count_space S2)"
using assms by (auto intro!: sigma_finite_measure_count_space_countable)
show "sets (count_space S1 \<Otimes>\<^sub>M count_space S2) = sets (count_space (S1 \<times> S2))"
@@ -995,10 +995,10 @@
by (subst (1 2 3) emeasure_count_space) (auto simp: finite_cartesian_product_iff ennreal_mult_top ennreal_top_mult)
qed
-lemma%important nn_integral_fst_count_space:
+proposition nn_integral_fst_count_space:
"(\<integral>\<^sup>+ x. \<integral>\<^sup>+ y. f (x, y) \<partial>count_space UNIV \<partial>count_space UNIV) = integral\<^sup>N (count_space UNIV) f"
(is "?lhs = ?rhs")
-proof%unimportant(cases)
+proof(cases)
assume *: "countable {xy. f xy \<noteq> 0}"
let ?A = "fst ` {xy. f xy \<noteq> 0}"
let ?B = "snd ` {xy. f xy \<noteq> 0}"
@@ -1073,7 +1073,7 @@
ultimately show ?thesis by simp
qed
-lemma nn_integral_snd_count_space:
+proposition nn_integral_snd_count_space:
"(\<integral>\<^sup>+ y. \<integral>\<^sup>+ x. f (x, y) \<partial>count_space UNIV \<partial>count_space UNIV) = integral\<^sup>N (count_space UNIV) f"
(is "?lhs = ?rhs")
proof -
@@ -1088,7 +1088,7 @@
finally show ?thesis .
qed
-lemma%unimportant measurable_pair_measure_countable1:
+lemma measurable_pair_measure_countable1:
assumes "countable A"
and [measurable]: "\<And>x. x \<in> A \<Longrightarrow> (\<lambda>y. f (x, y)) \<in> measurable N K"
shows "f \<in> measurable (count_space A \<Otimes>\<^sub>M N) K"
@@ -1097,11 +1097,11 @@
subsection%important \<open>Product of Borel spaces\<close>
-lemma%important borel_Times:
+theorem borel_Times:
fixes A :: "'a::topological_space set" and B :: "'b::topological_space set"
assumes A: "A \<in> sets borel" and B: "B \<in> sets borel"
shows "A \<times> B \<in> sets borel"
-proof%unimportant -
+proof -
have "A \<times> B = (A\<times>UNIV) \<inter> (UNIV \<times> B)"
by auto
moreover
@@ -1146,7 +1146,7 @@
by auto
qed
-lemma%unimportant finite_measure_pair_measure:
+lemma finite_measure_pair_measure:
assumes "finite_measure M" "finite_measure N"
shows "finite_measure (N \<Otimes>\<^sub>M M)"
proof (rule finite_measureI)
--- a/src/HOL/Analysis/Bochner_Integration.thy Sun Jan 13 20:25:41 2019 +0100
+++ b/src/HOL/Analysis/Bochner_Integration.thy Mon Jan 14 11:59:19 2019 +0000
@@ -15,12 +15,12 @@
\<close>
-lemma%important borel_measurable_implies_sequence_metric:
+proposition borel_measurable_implies_sequence_metric:
fixes f :: "'a \<Rightarrow> 'b :: {metric_space, second_countable_topology}"
assumes [measurable]: "f \<in> borel_measurable M"
shows "\<exists>F. (\<forall>i. simple_function M (F i)) \<and> (\<forall>x\<in>space M. (\<lambda>i. F i x) \<longlonglongrightarrow> f x) \<and>
(\<forall>i. \<forall>x\<in>space M. dist (F i x) z \<le> 2 * dist (f x) z)"
-proof%unimportant -
+proof -
obtain D :: "'b set" where "countable D" and D: "\<And>X. open X \<Longrightarrow> X \<noteq> {} \<Longrightarrow> \<exists>d\<in>D. d \<in> X"
by (erule countable_dense_setE)
@@ -155,14 +155,14 @@
qed
qed
-lemma%unimportant
+lemma
fixes f :: "'a \<Rightarrow> 'b::semiring_1" assumes "finite A"
shows sum_mult_indicator[simp]: "(\<Sum>x \<in> A. f x * indicator (B x) (g x)) = (\<Sum>x\<in>{x\<in>A. g x \<in> B x}. f x)"
and sum_indicator_mult[simp]: "(\<Sum>x \<in> A. indicator (B x) (g x) * f x) = (\<Sum>x\<in>{x\<in>A. g x \<in> B x}. f x)"
unfolding indicator_def
using assms by (auto intro!: sum.mono_neutral_cong_right split: if_split_asm)
-lemma%unimportant borel_measurable_induct_real[consumes 2, case_names set mult add seq]:
+lemma borel_measurable_induct_real[consumes 2, case_names set mult add seq]:
fixes P :: "('a \<Rightarrow> real) \<Rightarrow> bool"
assumes u: "u \<in> borel_measurable M" "\<And>x. 0 \<le> u x"
assumes set: "\<And>A. A \<in> sets M \<Longrightarrow> P (indicator A)"
@@ -227,7 +227,7 @@
qed
qed
-lemma%unimportant scaleR_cong_right:
+lemma scaleR_cong_right:
fixes x :: "'a :: real_vector"
shows "(x \<noteq> 0 \<Longrightarrow> r = p) \<Longrightarrow> r *\<^sub>R x = p *\<^sub>R x"
by (cases "x = 0") auto
@@ -236,7 +236,7 @@
"simple_function M f \<Longrightarrow> emeasure M {y\<in>space M. f y \<noteq> 0} \<noteq> \<infinity> \<Longrightarrow>
simple_bochner_integrable M f"
-lemma%unimportant simple_bochner_integrable_compose2:
+lemma simple_bochner_integrable_compose2:
assumes p_0: "p 0 0 = 0"
shows "simple_bochner_integrable M f \<Longrightarrow> simple_bochner_integrable M g \<Longrightarrow>
simple_bochner_integrable M (\<lambda>x. p (f x) (g x))"
@@ -261,7 +261,7 @@
using fin by (auto simp: top_unique)
qed
-lemma%unimportant simple_function_finite_support:
+lemma simple_function_finite_support:
assumes f: "simple_function M f" and fin: "(\<integral>\<^sup>+x. f x \<partial>M) < \<infinity>" and nn: "\<And>x. 0 \<le> f x"
shows "emeasure M {x\<in>space M. f x \<noteq> 0} \<noteq> \<infinity>"
proof cases
@@ -296,7 +296,7 @@
show ?thesis unfolding * by simp
qed
-lemma%unimportant simple_bochner_integrableI_bounded:
+lemma simple_bochner_integrableI_bounded:
assumes f: "simple_function M f" and fin: "(\<integral>\<^sup>+x. norm (f x) \<partial>M) < \<infinity>"
shows "simple_bochner_integrable M f"
proof
@@ -312,13 +312,13 @@
definition%important simple_bochner_integral :: "'a measure \<Rightarrow> ('a \<Rightarrow> 'b::real_vector) \<Rightarrow> 'b" where
"simple_bochner_integral M f = (\<Sum>y\<in>f`space M. measure M {x\<in>space M. f x = y} *\<^sub>R y)"
-lemma%important simple_bochner_integral_partition:
+proposition simple_bochner_integral_partition:
assumes f: "simple_bochner_integrable M f" and g: "simple_function M g"
assumes sub: "\<And>x y. x \<in> space M \<Longrightarrow> y \<in> space M \<Longrightarrow> g x = g y \<Longrightarrow> f x = f y"
assumes v: "\<And>x. x \<in> space M \<Longrightarrow> f x = v (g x)"
shows "simple_bochner_integral M f = (\<Sum>y\<in>g ` space M. measure M {x\<in>space M. g x = y} *\<^sub>R v y)"
(is "_ = ?r")
-proof%unimportant -
+proof -
from f g have [simp]: "finite (f`space M)" "finite (g`space M)"
by (auto simp: simple_function_def elim: simple_bochner_integrable.cases)
@@ -394,7 +394,7 @@
by (simp add: sum.distrib[symmetric] scaleR_add_right)
qed
-lemma%unimportant simple_bochner_integral_linear:
+lemma simple_bochner_integral_linear:
assumes "linear f"
assumes g: "simple_bochner_integrable M g"
shows "simple_bochner_integral M (\<lambda>x. f (g x)) = f (simple_bochner_integral M g)"
@@ -484,14 +484,14 @@
finally show ?thesis .
qed
-lemma%important simple_bochner_integral_bounded:
+lemma simple_bochner_integral_bounded:
fixes f :: "'a \<Rightarrow> 'b::{real_normed_vector, second_countable_topology}"
assumes f[measurable]: "f \<in> borel_measurable M"
assumes s: "simple_bochner_integrable M s" and t: "simple_bochner_integrable M t"
shows "ennreal (norm (simple_bochner_integral M s - simple_bochner_integral M t)) \<le>
(\<integral>\<^sup>+ x. norm (f x - s x) \<partial>M) + (\<integral>\<^sup>+ x. norm (f x - t x) \<partial>M)"
(is "ennreal (norm (?s - ?t)) \<le> ?S + ?T")
-proof%unimportant -
+proof -
have [measurable]: "s \<in> borel_measurable M" "t \<in> borel_measurable M"
using s t by (auto intro: borel_measurable_simple_function elim: simple_bochner_integrable.cases)
@@ -520,13 +520,13 @@
(\<lambda>i. simple_bochner_integral M (s i)) \<longlonglongrightarrow> x \<Longrightarrow>
has_bochner_integral M f x"
-lemma%unimportant has_bochner_integral_cong:
+lemma has_bochner_integral_cong:
assumes "M = N" "\<And>x. x \<in> space N \<Longrightarrow> f x = g x" "x = y"
shows "has_bochner_integral M f x \<longleftrightarrow> has_bochner_integral N g y"
unfolding has_bochner_integral.simps assms(1,3)
using assms(2) by (simp cong: measurable_cong_simp nn_integral_cong_simp)
-lemma%unimportant has_bochner_integral_cong_AE:
+lemma has_bochner_integral_cong_AE:
"f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> (AE x in M. f x = g x) \<Longrightarrow>
has_bochner_integral M f x \<longleftrightarrow> has_bochner_integral M g x"
unfolding has_bochner_integral.simps
@@ -534,22 +534,22 @@
nn_integral_cong_AE)
auto
-lemma%unimportant borel_measurable_has_bochner_integral:
+lemma borel_measurable_has_bochner_integral:
"has_bochner_integral M f x \<Longrightarrow> f \<in> borel_measurable M"
by (rule has_bochner_integral.cases)
-lemma%unimportant borel_measurable_has_bochner_integral'[measurable_dest]:
+lemma borel_measurable_has_bochner_integral'[measurable_dest]:
"has_bochner_integral M f x \<Longrightarrow> g \<in> measurable N M \<Longrightarrow> (\<lambda>x. f (g x)) \<in> borel_measurable N"
using borel_measurable_has_bochner_integral[measurable] by measurable
-lemma%unimportant has_bochner_integral_simple_bochner_integrable:
+lemma has_bochner_integral_simple_bochner_integrable:
"simple_bochner_integrable M f \<Longrightarrow> has_bochner_integral M f (simple_bochner_integral M f)"
by (rule has_bochner_integral.intros[where s="\<lambda>_. f"])
(auto intro: borel_measurable_simple_function
elim: simple_bochner_integrable.cases
simp: zero_ennreal_def[symmetric])
-lemma%unimportant has_bochner_integral_real_indicator:
+lemma has_bochner_integral_real_indicator:
assumes [measurable]: "A \<in> sets M" and A: "emeasure M A < \<infinity>"
shows "has_bochner_integral M (indicator A) (measure M A)"
proof -
@@ -567,7 +567,7 @@
by (metis has_bochner_integral_simple_bochner_integrable)
qed
-lemma%unimportant has_bochner_integral_add[intro]:
+lemma has_bochner_integral_add[intro]:
"has_bochner_integral M f x \<Longrightarrow> has_bochner_integral M g y \<Longrightarrow>
has_bochner_integral M (\<lambda>x. f x + g x) (x + y)"
proof (safe intro!: has_bochner_integral.intros elim!: has_bochner_integral.cases)
@@ -604,7 +604,7 @@
qed
qed (auto simp: simple_bochner_integral_add tendsto_add)
-lemma%unimportant has_bochner_integral_bounded_linear:
+lemma has_bochner_integral_bounded_linear:
assumes "bounded_linear T"
shows "has_bochner_integral M f x \<Longrightarrow> has_bochner_integral M (\<lambda>x. T (f x)) (T x)"
proof (safe intro!: has_bochner_integral.intros elim!: has_bochner_integral.cases)
@@ -650,79 +650,79 @@
by (auto intro!: T.tendsto simp: simple_bochner_integral_linear T.linear_axioms)
qed
-lemma%unimportant has_bochner_integral_zero[intro]: "has_bochner_integral M (\<lambda>x. 0) 0"
+lemma has_bochner_integral_zero[intro]: "has_bochner_integral M (\<lambda>x. 0) 0"
by (auto intro!: has_bochner_integral.intros[where s="\<lambda>_ _. 0"]
simp: zero_ennreal_def[symmetric] simple_bochner_integrable.simps
simple_bochner_integral_def image_constant_conv)
-lemma%unimportant has_bochner_integral_scaleR_left[intro]:
+lemma has_bochner_integral_scaleR_left[intro]:
"(c \<noteq> 0 \<Longrightarrow> has_bochner_integral M f x) \<Longrightarrow> has_bochner_integral M (\<lambda>x. f x *\<^sub>R c) (x *\<^sub>R c)"
by (cases "c = 0") (auto simp add: has_bochner_integral_bounded_linear[OF bounded_linear_scaleR_left])
-lemma%unimportant has_bochner_integral_scaleR_right[intro]:
+lemma has_bochner_integral_scaleR_right[intro]:
"(c \<noteq> 0 \<Longrightarrow> has_bochner_integral M f x) \<Longrightarrow> has_bochner_integral M (\<lambda>x. c *\<^sub>R f x) (c *\<^sub>R x)"
by (cases "c = 0") (auto simp add: has_bochner_integral_bounded_linear[OF bounded_linear_scaleR_right])
-lemma%unimportant has_bochner_integral_mult_left[intro]:
+lemma has_bochner_integral_mult_left[intro]:
fixes c :: "_::{real_normed_algebra,second_countable_topology}"
shows "(c \<noteq> 0 \<Longrightarrow> has_bochner_integral M f x) \<Longrightarrow> has_bochner_integral M (\<lambda>x. f x * c) (x * c)"
by (cases "c = 0") (auto simp add: has_bochner_integral_bounded_linear[OF bounded_linear_mult_left])
-lemma%unimportant has_bochner_integral_mult_right[intro]:
+lemma has_bochner_integral_mult_right[intro]:
fixes c :: "_::{real_normed_algebra,second_countable_topology}"
shows "(c \<noteq> 0 \<Longrightarrow> has_bochner_integral M f x) \<Longrightarrow> has_bochner_integral M (\<lambda>x. c * f x) (c * x)"
by (cases "c = 0") (auto simp add: has_bochner_integral_bounded_linear[OF bounded_linear_mult_right])
-lemmas%unimportant has_bochner_integral_divide =
+lemmas has_bochner_integral_divide =
has_bochner_integral_bounded_linear[OF bounded_linear_divide]
-lemma%unimportant has_bochner_integral_divide_zero[intro]:
+lemma has_bochner_integral_divide_zero[intro]:
fixes c :: "_::{real_normed_field, field, second_countable_topology}"
shows "(c \<noteq> 0 \<Longrightarrow> has_bochner_integral M f x) \<Longrightarrow> has_bochner_integral M (\<lambda>x. f x / c) (x / c)"
using has_bochner_integral_divide by (cases "c = 0") auto
-lemma%unimportant has_bochner_integral_inner_left[intro]:
+lemma has_bochner_integral_inner_left[intro]:
"(c \<noteq> 0 \<Longrightarrow> has_bochner_integral M f x) \<Longrightarrow> has_bochner_integral M (\<lambda>x. f x \<bullet> c) (x \<bullet> c)"
by (cases "c = 0") (auto simp add: has_bochner_integral_bounded_linear[OF bounded_linear_inner_left])
-lemma%unimportant has_bochner_integral_inner_right[intro]:
+lemma has_bochner_integral_inner_right[intro]:
"(c \<noteq> 0 \<Longrightarrow> has_bochner_integral M f x) \<Longrightarrow> has_bochner_integral M (\<lambda>x. c \<bullet> f x) (c \<bullet> x)"
by (cases "c = 0") (auto simp add: has_bochner_integral_bounded_linear[OF bounded_linear_inner_right])
-lemmas%unimportant has_bochner_integral_minus =
+lemmas has_bochner_integral_minus =
has_bochner_integral_bounded_linear[OF bounded_linear_minus[OF bounded_linear_ident]]
-lemmas%unimportant has_bochner_integral_Re =
+lemmas has_bochner_integral_Re =
has_bochner_integral_bounded_linear[OF bounded_linear_Re]
-lemmas%unimportant has_bochner_integral_Im =
+lemmas has_bochner_integral_Im =
has_bochner_integral_bounded_linear[OF bounded_linear_Im]
-lemmas%unimportant has_bochner_integral_cnj =
+lemmas has_bochner_integral_cnj =
has_bochner_integral_bounded_linear[OF bounded_linear_cnj]
-lemmas%unimportant has_bochner_integral_of_real =
+lemmas has_bochner_integral_of_real =
has_bochner_integral_bounded_linear[OF bounded_linear_of_real]
-lemmas%unimportant has_bochner_integral_fst =
+lemmas has_bochner_integral_fst =
has_bochner_integral_bounded_linear[OF bounded_linear_fst]
-lemmas%unimportant has_bochner_integral_snd =
+lemmas has_bochner_integral_snd =
has_bochner_integral_bounded_linear[OF bounded_linear_snd]
-lemma%unimportant has_bochner_integral_indicator:
+lemma has_bochner_integral_indicator:
"A \<in> sets M \<Longrightarrow> emeasure M A < \<infinity> \<Longrightarrow>
has_bochner_integral M (\<lambda>x. indicator A x *\<^sub>R c) (measure M A *\<^sub>R c)"
by (intro has_bochner_integral_scaleR_left has_bochner_integral_real_indicator)
-lemma%unimportant has_bochner_integral_diff:
+lemma has_bochner_integral_diff:
"has_bochner_integral M f x \<Longrightarrow> has_bochner_integral M g y \<Longrightarrow>
has_bochner_integral M (\<lambda>x. f x - g x) (x - y)"
unfolding diff_conv_add_uminus
by (intro has_bochner_integral_add has_bochner_integral_minus)
-lemma%unimportant has_bochner_integral_sum:
+lemma has_bochner_integral_sum:
"(\<And>i. i \<in> I \<Longrightarrow> has_bochner_integral M (f i) (x i)) \<Longrightarrow>
has_bochner_integral M (\<lambda>x. \<Sum>i\<in>I. f i x) (\<Sum>i\<in>I. x i)"
by (induct I rule: infinite_finite_induct) auto
-lemma%important has_bochner_integral_implies_finite_norm:
+proposition has_bochner_integral_implies_finite_norm:
"has_bochner_integral M f x \<Longrightarrow> (\<integral>\<^sup>+x. norm (f x) \<partial>M) < \<infinity>"
-proof%unimportant (elim has_bochner_integral.cases)
+proof (elim has_bochner_integral.cases)
fix s v
assume [measurable]: "f \<in> borel_measurable M" and s: "\<And>i. simple_bochner_integrable M (s i)" and
lim_0: "(\<lambda>i. \<integral>\<^sup>+ x. ennreal (norm (f x - s i x)) \<partial>M) \<longlonglongrightarrow> 0"
@@ -756,10 +756,10 @@
finally show "(\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M) < \<infinity>" .
qed
-lemma%important has_bochner_integral_norm_bound:
+proposition has_bochner_integral_norm_bound:
assumes i: "has_bochner_integral M f x"
shows "norm x \<le> (\<integral>\<^sup>+x. norm (f x) \<partial>M)"
-using assms proof%unimportant
+using assms proof
fix s assume
x: "(\<lambda>i. simple_bochner_integral M (s i)) \<longlonglongrightarrow> x" (is "?s \<longlonglongrightarrow> x") and
s[simp]: "\<And>i. simple_bochner_integrable M (s i)" and
@@ -797,9 +797,9 @@
qed
qed
-lemma%important has_bochner_integral_eq:
+lemma has_bochner_integral_eq:
"has_bochner_integral M f x \<Longrightarrow> has_bochner_integral M f y \<Longrightarrow> x = y"
-proof%unimportant (elim has_bochner_integral.cases)
+proof (elim has_bochner_integral.cases)
assume f[measurable]: "f \<in> borel_measurable M"
fix s t
@@ -834,7 +834,7 @@
then show "x = y" by simp
qed
-lemma%unimportant has_bochner_integralI_AE:
+lemma has_bochner_integralI_AE:
assumes f: "has_bochner_integral M f x"
and g: "g \<in> borel_measurable M"
and ae: "AE x in M. f x = g x"
@@ -848,7 +848,7 @@
finally show "(\<lambda>i. \<integral>\<^sup>+ x. ennreal (norm (g x - s i x)) \<partial>M) \<longlonglongrightarrow> 0" .
qed (auto intro: g)
-lemma%unimportant has_bochner_integral_eq_AE:
+lemma has_bochner_integral_eq_AE:
assumes f: "has_bochner_integral M f x"
and g: "has_bochner_integral M g y"
and ae: "AE x in M. f x = g x"
@@ -860,7 +860,7 @@
by (rule has_bochner_integral_eq)
qed
-lemma%unimportant simple_bochner_integrable_restrict_space:
+lemma simple_bochner_integrable_restrict_space:
fixes f :: "_ \<Rightarrow> 'b::real_normed_vector"
assumes \<Omega>: "\<Omega> \<inter> space M \<in> sets M"
shows "simple_bochner_integrable (restrict_space M \<Omega>) f \<longleftrightarrow>
@@ -869,13 +869,13 @@
simple_function_restrict_space[OF \<Omega>] emeasure_restrict_space[OF \<Omega>] Collect_restrict
indicator_eq_0_iff conj_left_commute)
-lemma%important simple_bochner_integral_restrict_space:
+lemma simple_bochner_integral_restrict_space:
fixes f :: "_ \<Rightarrow> 'b::real_normed_vector"
assumes \<Omega>: "\<Omega> \<inter> space M \<in> sets M"
assumes f: "simple_bochner_integrable (restrict_space M \<Omega>) f"
shows "simple_bochner_integral (restrict_space M \<Omega>) f =
simple_bochner_integral M (\<lambda>x. indicator \<Omega> x *\<^sub>R f x)"
-proof%unimportant -
+proof -
have "finite ((\<lambda>x. indicator \<Omega> x *\<^sub>R f x)`space M)"
using f simple_bochner_integrable_restrict_space[OF \<Omega>, of f]
by (simp add: simple_bochner_integrable.simps simple_function_def)
@@ -910,155 +910,155 @@
translations
"LINT x|M. f" == "CONST lebesgue_integral M (\<lambda>x. f)"
-lemma%unimportant has_bochner_integral_integral_eq: "has_bochner_integral M f x \<Longrightarrow> integral\<^sup>L M f = x"
+lemma has_bochner_integral_integral_eq: "has_bochner_integral M f x \<Longrightarrow> integral\<^sup>L M f = x"
by (metis the_equality has_bochner_integral_eq lebesgue_integral_def)
-lemma%unimportant has_bochner_integral_integrable:
+lemma has_bochner_integral_integrable:
"integrable M f \<Longrightarrow> has_bochner_integral M f (integral\<^sup>L M f)"
by (auto simp: has_bochner_integral_integral_eq integrable.simps)
-lemma%unimportant has_bochner_integral_iff:
+lemma has_bochner_integral_iff:
"has_bochner_integral M f x \<longleftrightarrow> integrable M f \<and> integral\<^sup>L M f = x"
by (metis has_bochner_integral_integrable has_bochner_integral_integral_eq integrable.intros)
-lemma%unimportant simple_bochner_integrable_eq_integral:
+lemma simple_bochner_integrable_eq_integral:
"simple_bochner_integrable M f \<Longrightarrow> simple_bochner_integral M f = integral\<^sup>L M f"
using has_bochner_integral_simple_bochner_integrable[of M f]
by (simp add: has_bochner_integral_integral_eq)
-lemma%unimportant not_integrable_integral_eq: "\<not> integrable M f \<Longrightarrow> integral\<^sup>L M f = 0"
+lemma not_integrable_integral_eq: "\<not> integrable M f \<Longrightarrow> integral\<^sup>L M f = 0"
unfolding integrable.simps lebesgue_integral_def by (auto intro!: arg_cong[where f=The])
-lemma%unimportant integral_eq_cases:
+lemma integral_eq_cases:
"integrable M f \<longleftrightarrow> integrable N g \<Longrightarrow>
(integrable M f \<Longrightarrow> integrable N g \<Longrightarrow> integral\<^sup>L M f = integral\<^sup>L N g) \<Longrightarrow>
integral\<^sup>L M f = integral\<^sup>L N g"
by (metis not_integrable_integral_eq)
-lemma%unimportant borel_measurable_integrable[measurable_dest]: "integrable M f \<Longrightarrow> f \<in> borel_measurable M"
+lemma borel_measurable_integrable[measurable_dest]: "integrable M f \<Longrightarrow> f \<in> borel_measurable M"
by (auto elim: integrable.cases has_bochner_integral.cases)
-lemma%unimportant borel_measurable_integrable'[measurable_dest]:
+lemma borel_measurable_integrable'[measurable_dest]:
"integrable M f \<Longrightarrow> g \<in> measurable N M \<Longrightarrow> (\<lambda>x. f (g x)) \<in> borel_measurable N"
using borel_measurable_integrable[measurable] by measurable
-lemma%unimportant integrable_cong:
+lemma integrable_cong:
"M = N \<Longrightarrow> (\<And>x. x \<in> space N \<Longrightarrow> f x = g x) \<Longrightarrow> integrable M f \<longleftrightarrow> integrable N g"
by (simp cong: has_bochner_integral_cong add: integrable.simps)
-lemma%unimportant integrable_cong_AE:
+lemma integrable_cong_AE:
"f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> AE x in M. f x = g x \<Longrightarrow>
integrable M f \<longleftrightarrow> integrable M g"
unfolding integrable.simps
by (intro has_bochner_integral_cong_AE arg_cong[where f=Ex] ext)
-lemma%unimportant integrable_cong_AE_imp:
+lemma integrable_cong_AE_imp:
"integrable M g \<Longrightarrow> f \<in> borel_measurable M \<Longrightarrow> (AE x in M. g x = f x) \<Longrightarrow> integrable M f"
using integrable_cong_AE[of f M g] by (auto simp: eq_commute)
-lemma%unimportant integral_cong:
+lemma integral_cong:
"M = N \<Longrightarrow> (\<And>x. x \<in> space N \<Longrightarrow> f x = g x) \<Longrightarrow> integral\<^sup>L M f = integral\<^sup>L N g"
by (simp cong: has_bochner_integral_cong cong del: if_weak_cong add: lebesgue_integral_def)
-lemma%unimportant integral_cong_AE:
+lemma integral_cong_AE:
"f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> AE x in M. f x = g x \<Longrightarrow>
integral\<^sup>L M f = integral\<^sup>L M g"
unfolding lebesgue_integral_def
by (rule arg_cong[where x="has_bochner_integral M f"]) (intro has_bochner_integral_cong_AE ext)
-lemma%unimportant integrable_add[simp, intro]: "integrable M f \<Longrightarrow> integrable M g \<Longrightarrow> integrable M (\<lambda>x. f x + g x)"
+lemma integrable_add[simp, intro]: "integrable M f \<Longrightarrow> integrable M g \<Longrightarrow> integrable M (\<lambda>x. f x + g x)"
by (auto simp: integrable.simps)
-lemma%unimportant integrable_zero[simp, intro]: "integrable M (\<lambda>x. 0)"
+lemma integrable_zero[simp, intro]: "integrable M (\<lambda>x. 0)"
by (metis has_bochner_integral_zero integrable.simps)
-lemma%unimportant integrable_sum[simp, intro]: "(\<And>i. i \<in> I \<Longrightarrow> integrable M (f i)) \<Longrightarrow> integrable M (\<lambda>x. \<Sum>i\<in>I. f i x)"
+lemma integrable_sum[simp, intro]: "(\<And>i. i \<in> I \<Longrightarrow> integrable M (f i)) \<Longrightarrow> integrable M (\<lambda>x. \<Sum>i\<in>I. f i x)"
by (metis has_bochner_integral_sum integrable.simps)
-lemma%unimportant integrable_indicator[simp, intro]: "A \<in> sets M \<Longrightarrow> emeasure M A < \<infinity> \<Longrightarrow>
+lemma integrable_indicator[simp, intro]: "A \<in> sets M \<Longrightarrow> emeasure M A < \<infinity> \<Longrightarrow>
integrable M (\<lambda>x. indicator A x *\<^sub>R c)"
by (metis has_bochner_integral_indicator integrable.simps)
-lemma%unimportant integrable_real_indicator[simp, intro]: "A \<in> sets M \<Longrightarrow> emeasure M A < \<infinity> \<Longrightarrow>
+lemma integrable_real_indicator[simp, intro]: "A \<in> sets M \<Longrightarrow> emeasure M A < \<infinity> \<Longrightarrow>
integrable M (indicator A :: 'a \<Rightarrow> real)"
by (metis has_bochner_integral_real_indicator integrable.simps)
-lemma%unimportant integrable_diff[simp, intro]: "integrable M f \<Longrightarrow> integrable M g \<Longrightarrow> integrable M (\<lambda>x. f x - g x)"
+lemma integrable_diff[simp, intro]: "integrable M f \<Longrightarrow> integrable M g \<Longrightarrow> integrable M (\<lambda>x. f x - g x)"
by (auto simp: integrable.simps intro: has_bochner_integral_diff)
-lemma%unimportant integrable_bounded_linear: "bounded_linear T \<Longrightarrow> integrable M f \<Longrightarrow> integrable M (\<lambda>x. T (f x))"
+lemma integrable_bounded_linear: "bounded_linear T \<Longrightarrow> integrable M f \<Longrightarrow> integrable M (\<lambda>x. T (f x))"
by (auto simp: integrable.simps intro: has_bochner_integral_bounded_linear)
-lemma%unimportant integrable_scaleR_left[simp, intro]: "(c \<noteq> 0 \<Longrightarrow> integrable M f) \<Longrightarrow> integrable M (\<lambda>x. f x *\<^sub>R c)"
+lemma integrable_scaleR_left[simp, intro]: "(c \<noteq> 0 \<Longrightarrow> integrable M f) \<Longrightarrow> integrable M (\<lambda>x. f x *\<^sub>R c)"
unfolding integrable.simps by fastforce
-lemma%unimportant integrable_scaleR_right[simp, intro]: "(c \<noteq> 0 \<Longrightarrow> integrable M f) \<Longrightarrow> integrable M (\<lambda>x. c *\<^sub>R f x)"
+lemma integrable_scaleR_right[simp, intro]: "(c \<noteq> 0 \<Longrightarrow> integrable M f) \<Longrightarrow> integrable M (\<lambda>x. c *\<^sub>R f x)"
unfolding integrable.simps by fastforce
-lemma%unimportant integrable_mult_left[simp, intro]:
+lemma integrable_mult_left[simp, intro]:
fixes c :: "_::{real_normed_algebra,second_countable_topology}"
shows "(c \<noteq> 0 \<Longrightarrow> integrable M f) \<Longrightarrow> integrable M (\<lambda>x. f x * c)"
unfolding integrable.simps by fastforce
-lemma%unimportant integrable_mult_right[simp, intro]:
+lemma integrable_mult_right[simp, intro]:
fixes c :: "_::{real_normed_algebra,second_countable_topology}"
shows "(c \<noteq> 0 \<Longrightarrow> integrable M f) \<Longrightarrow> integrable M (\<lambda>x. c * f x)"
unfolding integrable.simps by fastforce
-lemma%unimportant integrable_divide_zero[simp, intro]:
+lemma integrable_divide_zero[simp, intro]:
fixes c :: "_::{real_normed_field, field, second_countable_topology}"
shows "(c \<noteq> 0 \<Longrightarrow> integrable M f) \<Longrightarrow> integrable M (\<lambda>x. f x / c)"
unfolding integrable.simps by fastforce
-lemma%unimportant integrable_inner_left[simp, intro]:
+lemma integrable_inner_left[simp, intro]:
"(c \<noteq> 0 \<Longrightarrow> integrable M f) \<Longrightarrow> integrable M (\<lambda>x. f x \<bullet> c)"
unfolding integrable.simps by fastforce
-lemma%unimportant integrable_inner_right[simp, intro]:
+lemma integrable_inner_right[simp, intro]:
"(c \<noteq> 0 \<Longrightarrow> integrable M f) \<Longrightarrow> integrable M (\<lambda>x. c \<bullet> f x)"
unfolding integrable.simps by fastforce
-lemmas%unimportant integrable_minus[simp, intro] =
+lemmas integrable_minus[simp, intro] =
integrable_bounded_linear[OF bounded_linear_minus[OF bounded_linear_ident]]
-lemmas%unimportant integrable_divide[simp, intro] =
+lemmas integrable_divide[simp, intro] =
integrable_bounded_linear[OF bounded_linear_divide]
-lemmas%unimportant integrable_Re[simp, intro] =
+lemmas integrable_Re[simp, intro] =
integrable_bounded_linear[OF bounded_linear_Re]
-lemmas%unimportant integrable_Im[simp, intro] =
+lemmas integrable_Im[simp, intro] =
integrable_bounded_linear[OF bounded_linear_Im]
-lemmas%unimportant integrable_cnj[simp, intro] =
+lemmas integrable_cnj[simp, intro] =
integrable_bounded_linear[OF bounded_linear_cnj]
-lemmas%unimportant integrable_of_real[simp, intro] =
+lemmas integrable_of_real[simp, intro] =
integrable_bounded_linear[OF bounded_linear_of_real]
-lemmas%unimportant integrable_fst[simp, intro] =
+lemmas integrable_fst[simp, intro] =
integrable_bounded_linear[OF bounded_linear_fst]
-lemmas%unimportant integrable_snd[simp, intro] =
+lemmas integrable_snd[simp, intro] =
integrable_bounded_linear[OF bounded_linear_snd]
-lemma%unimportant integral_zero[simp]: "integral\<^sup>L M (\<lambda>x. 0) = 0"
+lemma integral_zero[simp]: "integral\<^sup>L M (\<lambda>x. 0) = 0"
by (intro has_bochner_integral_integral_eq has_bochner_integral_zero)
-lemma%unimportant integral_add[simp]: "integrable M f \<Longrightarrow> integrable M g \<Longrightarrow>
+lemma integral_add[simp]: "integrable M f \<Longrightarrow> integrable M g \<Longrightarrow>
integral\<^sup>L M (\<lambda>x. f x + g x) = integral\<^sup>L M f + integral\<^sup>L M g"
by (intro has_bochner_integral_integral_eq has_bochner_integral_add has_bochner_integral_integrable)
-lemma%unimportant integral_diff[simp]: "integrable M f \<Longrightarrow> integrable M g \<Longrightarrow>
+lemma integral_diff[simp]: "integrable M f \<Longrightarrow> integrable M g \<Longrightarrow>
integral\<^sup>L M (\<lambda>x. f x - g x) = integral\<^sup>L M f - integral\<^sup>L M g"
by (intro has_bochner_integral_integral_eq has_bochner_integral_diff has_bochner_integral_integrable)
-lemma%unimportant integral_sum: "(\<And>i. i \<in> I \<Longrightarrow> integrable M (f i)) \<Longrightarrow>
+lemma integral_sum: "(\<And>i. i \<in> I \<Longrightarrow> integrable M (f i)) \<Longrightarrow>
integral\<^sup>L M (\<lambda>x. \<Sum>i\<in>I. f i x) = (\<Sum>i\<in>I. integral\<^sup>L M (f i))"
by (intro has_bochner_integral_integral_eq has_bochner_integral_sum has_bochner_integral_integrable)
-lemma%unimportant integral_sum'[simp]: "(\<And>i. i \<in> I =simp=> integrable M (f i)) \<Longrightarrow>
+lemma integral_sum'[simp]: "(\<And>i. i \<in> I =simp=> integrable M (f i)) \<Longrightarrow>
integral\<^sup>L M (\<lambda>x. \<Sum>i\<in>I. f i x) = (\<Sum>i\<in>I. integral\<^sup>L M (f i))"
unfolding simp_implies_def by (rule integral_sum)
-lemma%unimportant integral_bounded_linear: "bounded_linear T \<Longrightarrow> integrable M f \<Longrightarrow>
+lemma integral_bounded_linear: "bounded_linear T \<Longrightarrow> integrable M f \<Longrightarrow>
integral\<^sup>L M (\<lambda>x. T (f x)) = T (integral\<^sup>L M f)"
by (metis has_bochner_integral_bounded_linear has_bochner_integral_integrable has_bochner_integral_integral_eq)
-lemma%unimportant integral_bounded_linear':
+lemma integral_bounded_linear':
assumes T: "bounded_linear T" and T': "bounded_linear T'"
assumes *: "\<not> (\<forall>x. T x = 0) \<Longrightarrow> (\<forall>x. T' (T x) = x)"
shows "integral\<^sup>L M (\<lambda>x. T (f x)) = T (integral\<^sup>L M f)"
@@ -1085,76 +1085,76 @@
qed
qed
-lemma%unimportant integral_scaleR_left[simp]: "(c \<noteq> 0 \<Longrightarrow> integrable M f) \<Longrightarrow> (\<integral> x. f x *\<^sub>R c \<partial>M) = integral\<^sup>L M f *\<^sub>R c"
+lemma integral_scaleR_left[simp]: "(c \<noteq> 0 \<Longrightarrow> integrable M f) \<Longrightarrow> (\<integral> x. f x *\<^sub>R c \<partial>M) = integral\<^sup>L M f *\<^sub>R c"
by (intro has_bochner_integral_integral_eq has_bochner_integral_integrable has_bochner_integral_scaleR_left)
-lemma%unimportant integral_scaleR_right[simp]: "(\<integral> x. c *\<^sub>R f x \<partial>M) = c *\<^sub>R integral\<^sup>L M f"
+lemma integral_scaleR_right[simp]: "(\<integral> x. c *\<^sub>R f x \<partial>M) = c *\<^sub>R integral\<^sup>L M f"
by (rule integral_bounded_linear'[OF bounded_linear_scaleR_right bounded_linear_scaleR_right[of "1 / c"]]) simp
-lemma%unimportant integral_mult_left[simp]:
+lemma integral_mult_left[simp]:
fixes c :: "_::{real_normed_algebra,second_countable_topology}"
shows "(c \<noteq> 0 \<Longrightarrow> integrable M f) \<Longrightarrow> (\<integral> x. f x * c \<partial>M) = integral\<^sup>L M f * c"
by (intro has_bochner_integral_integral_eq has_bochner_integral_integrable has_bochner_integral_mult_left)
-lemma%unimportant integral_mult_right[simp]:
+lemma integral_mult_right[simp]:
fixes c :: "_::{real_normed_algebra,second_countable_topology}"
shows "(c \<noteq> 0 \<Longrightarrow> integrable M f) \<Longrightarrow> (\<integral> x. c * f x \<partial>M) = c * integral\<^sup>L M f"
by (intro has_bochner_integral_integral_eq has_bochner_integral_integrable has_bochner_integral_mult_right)
-lemma%unimportant integral_mult_left_zero[simp]:
+lemma integral_mult_left_zero[simp]:
fixes c :: "_::{real_normed_field,second_countable_topology}"
shows "(\<integral> x. f x * c \<partial>M) = integral\<^sup>L M f * c"
by (rule integral_bounded_linear'[OF bounded_linear_mult_left bounded_linear_mult_left[of "1 / c"]]) simp
-lemma%unimportant integral_mult_right_zero[simp]:
+lemma integral_mult_right_zero[simp]:
fixes c :: "_::{real_normed_field,second_countable_topology}"
shows "(\<integral> x. c * f x \<partial>M) = c * integral\<^sup>L M f"
by (rule integral_bounded_linear'[OF bounded_linear_mult_right bounded_linear_mult_right[of "1 / c"]]) simp
-lemma%unimportant integral_inner_left[simp]: "(c \<noteq> 0 \<Longrightarrow> integrable M f) \<Longrightarrow> (\<integral> x. f x \<bullet> c \<partial>M) = integral\<^sup>L M f \<bullet> c"
+lemma integral_inner_left[simp]: "(c \<noteq> 0 \<Longrightarrow> integrable M f) \<Longrightarrow> (\<integral> x. f x \<bullet> c \<partial>M) = integral\<^sup>L M f \<bullet> c"
by (intro has_bochner_integral_integral_eq has_bochner_integral_integrable has_bochner_integral_inner_left)
-lemma%unimportant integral_inner_right[simp]: "(c \<noteq> 0 \<Longrightarrow> integrable M f) \<Longrightarrow> (\<integral> x. c \<bullet> f x \<partial>M) = c \<bullet> integral\<^sup>L M f"
+lemma integral_inner_right[simp]: "(c \<noteq> 0 \<Longrightarrow> integrable M f) \<Longrightarrow> (\<integral> x. c \<bullet> f x \<partial>M) = c \<bullet> integral\<^sup>L M f"
by (intro has_bochner_integral_integral_eq has_bochner_integral_integrable has_bochner_integral_inner_right)
-lemma%unimportant integral_divide_zero[simp]:
+lemma integral_divide_zero[simp]:
fixes c :: "_::{real_normed_field, field, second_countable_topology}"
shows "integral\<^sup>L M (\<lambda>x. f x / c) = integral\<^sup>L M f / c"
by (rule integral_bounded_linear'[OF bounded_linear_divide bounded_linear_mult_left[of c]]) simp
-lemma%unimportant integral_minus[simp]: "integral\<^sup>L M (\<lambda>x. - f x) = - integral\<^sup>L M f"
+lemma integral_minus[simp]: "integral\<^sup>L M (\<lambda>x. - f x) = - integral\<^sup>L M f"
by (rule integral_bounded_linear'[OF bounded_linear_minus[OF bounded_linear_ident] bounded_linear_minus[OF bounded_linear_ident]]) simp
-lemma%unimportant integral_complex_of_real[simp]: "integral\<^sup>L M (\<lambda>x. complex_of_real (f x)) = of_real (integral\<^sup>L M f)"
+lemma integral_complex_of_real[simp]: "integral\<^sup>L M (\<lambda>x. complex_of_real (f x)) = of_real (integral\<^sup>L M f)"
by (rule integral_bounded_linear'[OF bounded_linear_of_real bounded_linear_Re]) simp
-lemma%unimportant integral_cnj[simp]: "integral\<^sup>L M (\<lambda>x. cnj (f x)) = cnj (integral\<^sup>L M f)"
+lemma integral_cnj[simp]: "integral\<^sup>L M (\<lambda>x. cnj (f x)) = cnj (integral\<^sup>L M f)"
by (rule integral_bounded_linear'[OF bounded_linear_cnj bounded_linear_cnj]) simp
-lemmas%unimportant integral_divide[simp] =
+lemmas integral_divide[simp] =
integral_bounded_linear[OF bounded_linear_divide]
-lemmas%unimportant integral_Re[simp] =
+lemmas integral_Re[simp] =
integral_bounded_linear[OF bounded_linear_Re]
-lemmas%unimportant integral_Im[simp] =
+lemmas integral_Im[simp] =
integral_bounded_linear[OF bounded_linear_Im]
-lemmas%unimportant integral_of_real[simp] =
+lemmas integral_of_real[simp] =
integral_bounded_linear[OF bounded_linear_of_real]
-lemmas%unimportant integral_fst[simp] =
+lemmas integral_fst[simp] =
integral_bounded_linear[OF bounded_linear_fst]
-lemmas%unimportant integral_snd[simp] =
+lemmas integral_snd[simp] =
integral_bounded_linear[OF bounded_linear_snd]
-lemma%unimportant integral_norm_bound_ennreal:
+lemma integral_norm_bound_ennreal:
"integrable M f \<Longrightarrow> norm (integral\<^sup>L M f) \<le> (\<integral>\<^sup>+x. norm (f x) \<partial>M)"
by (metis has_bochner_integral_integrable has_bochner_integral_norm_bound)
-lemma%important integrableI_sequence:
+lemma integrableI_sequence:
fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
assumes f[measurable]: "f \<in> borel_measurable M"
assumes s: "\<And>i. simple_bochner_integrable M (s i)"
assumes lim: "(\<lambda>i. \<integral>\<^sup>+x. norm (f x - s i x) \<partial>M) \<longlonglongrightarrow> 0" (is "?S \<longlonglongrightarrow> 0")
shows "integrable M f"
-proof%unimportant -
+proof -
let ?s = "\<lambda>n. simple_bochner_integral M (s n)"
have "\<exists>x. ?s \<longlonglongrightarrow> x"
@@ -1184,7 +1184,7 @@
by (rule, rule) fact+
qed
-lemma%important nn_integral_dominated_convergence_norm:
+proposition nn_integral_dominated_convergence_norm:
fixes u' :: "_ \<Rightarrow> _::{real_normed_vector, second_countable_topology}"
assumes [measurable]:
"\<And>i. u i \<in> borel_measurable M" "u' \<in> borel_measurable M" "w \<in> borel_measurable M"
@@ -1192,7 +1192,7 @@
and w: "(\<integral>\<^sup>+x. w x \<partial>M) < \<infinity>"
and u': "AE x in M. (\<lambda>i. u i x) \<longlonglongrightarrow> u' x"
shows "(\<lambda>i. (\<integral>\<^sup>+x. norm (u' x - u i x) \<partial>M)) \<longlonglongrightarrow> 0"
-proof%unimportant -
+proof -
have "AE x in M. \<forall>j. norm (u j x) \<le> w x"
unfolding AE_all_countable by rule fact
with u' have bnd: "AE x in M. \<forall>j. norm (u' x - u j x) \<le> 2 * w x"
@@ -1225,11 +1225,11 @@
then show ?thesis by simp
qed
-lemma%important integrableI_bounded:
+proposition integrableI_bounded:
fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
assumes f[measurable]: "f \<in> borel_measurable M" and fin: "(\<integral>\<^sup>+x. norm (f x) \<partial>M) < \<infinity>"
shows "integrable M f"
-proof%unimportant -
+proof -
from borel_measurable_implies_sequence_metric[OF f, of 0] obtain s where
s: "\<And>i. simple_function M (s i)" and
pointwise: "\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>i. s i x) \<longlonglongrightarrow> f x" and
@@ -1266,14 +1266,14 @@
qed fact
qed
-lemma%important integrableI_bounded_set:
+lemma integrableI_bounded_set:
fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
assumes [measurable]: "A \<in> sets M" "f \<in> borel_measurable M"
assumes finite: "emeasure M A < \<infinity>"
and bnd: "AE x in M. x \<in> A \<longrightarrow> norm (f x) \<le> B"
and null: "AE x in M. x \<notin> A \<longrightarrow> f x = 0"
shows "integrable M f"
-proof%unimportant (rule integrableI_bounded)
+proof (rule integrableI_bounded)
{ fix x :: 'b have "norm x \<le> B \<Longrightarrow> 0 \<le> B"
using norm_ge_zero[of x] by arith }
with bnd null have "(\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M) \<le> (\<integral>\<^sup>+ x. ennreal (max 0 B) * indicator A x \<partial>M)"
@@ -1283,37 +1283,37 @@
finally show "(\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M) < \<infinity>" .
qed simp
-lemma%unimportant integrableI_bounded_set_indicator:
+lemma integrableI_bounded_set_indicator:
fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
shows "A \<in> sets M \<Longrightarrow> f \<in> borel_measurable M \<Longrightarrow>
emeasure M A < \<infinity> \<Longrightarrow> (AE x in M. x \<in> A \<longrightarrow> norm (f x) \<le> B) \<Longrightarrow>
integrable M (\<lambda>x. indicator A x *\<^sub>R f x)"
by (rule integrableI_bounded_set[where A=A]) auto
-lemma%important integrableI_nonneg:
+lemma integrableI_nonneg:
fixes f :: "'a \<Rightarrow> real"
assumes "f \<in> borel_measurable M" "AE x in M. 0 \<le> f x" "(\<integral>\<^sup>+x. f x \<partial>M) < \<infinity>"
shows "integrable M f"
-proof%unimportant -
+proof -
have "(\<integral>\<^sup>+x. norm (f x) \<partial>M) = (\<integral>\<^sup>+x. f x \<partial>M)"
using assms by (intro nn_integral_cong_AE) auto
then show ?thesis
using assms by (intro integrableI_bounded) auto
qed
-lemma%important integrable_iff_bounded:
+lemma integrable_iff_bounded:
fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
shows "integrable M f \<longleftrightarrow> f \<in> borel_measurable M \<and> (\<integral>\<^sup>+x. norm (f x) \<partial>M) < \<infinity>"
using integrableI_bounded[of f M] has_bochner_integral_implies_finite_norm[of M f]
unfolding integrable.simps has_bochner_integral.simps[abs_def] by auto
-lemma%important integrable_bound:
+lemma integrable_bound:
fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
and g :: "'a \<Rightarrow> 'c::{banach, second_countable_topology}"
shows "integrable M f \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> (AE x in M. norm (g x) \<le> norm (f x)) \<Longrightarrow>
integrable M g"
unfolding integrable_iff_bounded
-proof%unimportant safe
+proof safe
assume "f \<in> borel_measurable M" "g \<in> borel_measurable M"
assume "AE x in M. norm (g x) \<le> norm (f x)"
then have "(\<integral>\<^sup>+ x. ennreal (norm (g x)) \<partial>M) \<le> (\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M)"
@@ -1322,71 +1322,71 @@
finally show "(\<integral>\<^sup>+ x. ennreal (norm (g x)) \<partial>M) < \<infinity>" .
qed
-lemma%unimportant integrable_mult_indicator:
+lemma integrable_mult_indicator:
fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
shows "A \<in> sets M \<Longrightarrow> integrable M f \<Longrightarrow> integrable M (\<lambda>x. indicator A x *\<^sub>R f x)"
by (rule integrable_bound[of M f]) (auto split: split_indicator)
-lemma%unimportant integrable_real_mult_indicator:
+lemma integrable_real_mult_indicator:
fixes f :: "'a \<Rightarrow> real"
shows "A \<in> sets M \<Longrightarrow> integrable M f \<Longrightarrow> integrable M (\<lambda>x. f x * indicator A x)"
using integrable_mult_indicator[of A M f] by (simp add: mult_ac)
-lemma%unimportant integrable_abs[simp, intro]:
+lemma integrable_abs[simp, intro]:
fixes f :: "'a \<Rightarrow> real"
assumes [measurable]: "integrable M f" shows "integrable M (\<lambda>x. \<bar>f x\<bar>)"
using assms by (rule integrable_bound) auto
-lemma%unimportant integrable_norm[simp, intro]:
+lemma integrable_norm[simp, intro]:
fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
assumes [measurable]: "integrable M f" shows "integrable M (\<lambda>x. norm (f x))"
using assms by (rule integrable_bound) auto
-lemma%unimportant integrable_norm_cancel:
+lemma integrable_norm_cancel:
fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
assumes [measurable]: "integrable M (\<lambda>x. norm (f x))" "f \<in> borel_measurable M" shows "integrable M f"
using assms by (rule integrable_bound) auto
-lemma%unimportant integrable_norm_iff:
+lemma integrable_norm_iff:
fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
shows "f \<in> borel_measurable M \<Longrightarrow> integrable M (\<lambda>x. norm (f x)) \<longleftrightarrow> integrable M f"
by (auto intro: integrable_norm_cancel)
-lemma%unimportant integrable_abs_cancel:
+lemma integrable_abs_cancel:
fixes f :: "'a \<Rightarrow> real"
assumes [measurable]: "integrable M (\<lambda>x. \<bar>f x\<bar>)" "f \<in> borel_measurable M" shows "integrable M f"
using assms by (rule integrable_bound) auto
-lemma%unimportant integrable_abs_iff:
+lemma integrable_abs_iff:
fixes f :: "'a \<Rightarrow> real"
shows "f \<in> borel_measurable M \<Longrightarrow> integrable M (\<lambda>x. \<bar>f x\<bar>) \<longleftrightarrow> integrable M f"
by (auto intro: integrable_abs_cancel)
-lemma%unimportant integrable_max[simp, intro]:
+lemma integrable_max[simp, intro]:
fixes f :: "'a \<Rightarrow> real"
assumes fg[measurable]: "integrable M f" "integrable M g"
shows "integrable M (\<lambda>x. max (f x) (g x))"
using integrable_add[OF integrable_norm[OF fg(1)] integrable_norm[OF fg(2)]]
by (rule integrable_bound) auto
-lemma%unimportant integrable_min[simp, intro]:
+lemma integrable_min[simp, intro]:
fixes f :: "'a \<Rightarrow> real"
assumes fg[measurable]: "integrable M f" "integrable M g"
shows "integrable M (\<lambda>x. min (f x) (g x))"
using integrable_add[OF integrable_norm[OF fg(1)] integrable_norm[OF fg(2)]]
by (rule integrable_bound) auto
-lemma%unimportant integral_minus_iff[simp]:
+lemma integral_minus_iff[simp]:
"integrable M (\<lambda>x. - f x ::'a::{banach, second_countable_topology}) \<longleftrightarrow> integrable M f"
unfolding integrable_iff_bounded
by (auto intro: borel_measurable_uminus[of "\<lambda>x. - f x" M, simplified])
-lemma%unimportant integrable_indicator_iff:
+lemma integrable_indicator_iff:
"integrable M (indicator A::_ \<Rightarrow> real) \<longleftrightarrow> A \<inter> space M \<in> sets M \<and> emeasure M (A \<inter> space M) < \<infinity>"
by (simp add: integrable_iff_bounded borel_measurable_indicator_iff ennreal_indicator nn_integral_indicator'
cong: conj_cong)
-lemma%unimportant integral_indicator[simp]: "integral\<^sup>L M (indicator A) = measure M (A \<inter> space M)"
+lemma integral_indicator[simp]: "integral\<^sup>L M (indicator A) = measure M (A \<inter> space M)"
proof cases
assume *: "A \<inter> space M \<in> sets M \<and> emeasure M (A \<inter> space M) < \<infinity>"
have "integral\<^sup>L M (indicator A) = integral\<^sup>L M (indicator (A \<inter> space M))"
@@ -1405,7 +1405,7 @@
finally show ?thesis .
qed
-lemma%unimportant integrable_discrete_difference:
+lemma integrable_discrete_difference:
fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
assumes X: "countable X"
assumes null: "\<And>x. x \<in> X \<Longrightarrow> emeasure M {x} = 0"
@@ -1429,7 +1429,7 @@
by simp
qed
-lemma%unimportant integral_discrete_difference:
+lemma integral_discrete_difference:
fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
assumes X: "countable X"
assumes null: "\<And>x. x \<in> X \<Longrightarrow> emeasure M {x} = 0"
@@ -1453,7 +1453,7 @@
qed
qed
-lemma%unimportant has_bochner_integral_discrete_difference:
+lemma has_bochner_integral_discrete_difference:
fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
assumes X: "countable X"
assumes null: "\<And>x. x \<in> X \<Longrightarrow> emeasure M {x} = 0"
@@ -1464,7 +1464,7 @@
using integral_discrete_difference[of X M f g, OF assms]
by (metis has_bochner_integral_iff)
-lemma%important (*FIX ME needs name *)
+lemma (*FIX ME needs name *)
fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}" and w :: "'a \<Rightarrow> real"
assumes "f \<in> borel_measurable M" "\<And>i. s i \<in> borel_measurable M" "integrable M w"
assumes lim: "AE x in M. (\<lambda>i. s i x) \<longlonglongrightarrow> f x"
@@ -1472,7 +1472,7 @@
shows integrable_dominated_convergence: "integrable M f"
and integrable_dominated_convergence2: "\<And>i. integrable M (s i)"
and integral_dominated_convergence: "(\<lambda>i. integral\<^sup>L M (s i)) \<longlonglongrightarrow> integral\<^sup>L M f"
-proof%unimportant -
+proof -
have w_nonneg: "AE x in M. 0 \<le> w x"
using bound[of 0] by eventually_elim (auto intro: norm_ge_zero order_trans)
then have "(\<integral>\<^sup>+x. w x \<partial>M) = (\<integral>\<^sup>+x. norm (w x) \<partial>M)"
@@ -1539,7 +1539,7 @@
assumes bound: "\<forall>\<^sub>F i in at_top. AE x in M. norm (s i x) \<le> w x"
begin
-lemma%unimportant integral_dominated_convergence_at_top: "((\<lambda>t. integral\<^sup>L M (s t)) \<longlongrightarrow> integral\<^sup>L M f) at_top"
+lemma integral_dominated_convergence_at_top: "((\<lambda>t. integral\<^sup>L M (s t)) \<longlongrightarrow> integral\<^sup>L M f) at_top"
proof (rule tendsto_at_topI_sequentially)
fix X :: "nat \<Rightarrow> real" assume X: "filterlim X at_top sequentially"
from filterlim_iff[THEN iffD1, OF this, rule_format, OF bound]
@@ -1560,7 +1560,7 @@
qed fact+
qed
-lemma%unimportant integrable_dominated_convergence_at_top: "integrable M f"
+lemma integrable_dominated_convergence_at_top: "integrable M f"
proof -
from bound obtain N where w: "\<And>n. N \<le> n \<Longrightarrow> AE x in M. norm (s n x) \<le> w x"
by (auto simp: eventually_at_top_linorder)
@@ -1581,13 +1581,13 @@
end
-lemma%unimportant integrable_mult_left_iff:
+lemma integrable_mult_left_iff:
fixes f :: "'a \<Rightarrow> real"
shows "integrable M (\<lambda>x. c * f x) \<longleftrightarrow> c = 0 \<or> integrable M f"
using integrable_mult_left[of c M f] integrable_mult_left[of "1 / c" M "\<lambda>x. c * f x"]
by (cases "c = 0") auto
-lemma%unimportant integrableI_nn_integral_finite:
+lemma integrableI_nn_integral_finite:
assumes [measurable]: "f \<in> borel_measurable M"
and nonneg: "AE x in M. 0 \<le> f x"
and finite: "(\<integral>\<^sup>+x. f x \<partial>M) = ennreal x"
@@ -1599,7 +1599,7 @@
by auto
qed simp
-lemma%unimportant integral_nonneg_AE:
+lemma integral_nonneg_AE:
fixes f :: "'a \<Rightarrow> real"
assumes nonneg: "AE x in M. 0 \<le> f x"
shows "0 \<le> integral\<^sup>L M f"
@@ -1635,16 +1635,16 @@
finally show ?thesis .
qed (simp add: not_integrable_integral_eq)
-lemma%unimportant integral_nonneg[simp]:
+lemma integral_nonneg[simp]:
fixes f :: "'a \<Rightarrow> real"
shows "(\<And>x. x \<in> space M \<Longrightarrow> 0 \<le> f x) \<Longrightarrow> 0 \<le> integral\<^sup>L M f"
by (intro integral_nonneg_AE) auto
-lemma%important nn_integral_eq_integral:
+proposition nn_integral_eq_integral:
assumes f: "integrable M f"
assumes nonneg: "AE x in M. 0 \<le> f x"
shows "(\<integral>\<^sup>+ x. f x \<partial>M) = integral\<^sup>L M f"
-proof%unimportant -
+proof -
{ fix f :: "'a \<Rightarrow> real" assume f: "f \<in> borel_measurable M" "\<And>x. 0 \<le> f x" "integrable M f"
then have "(\<integral>\<^sup>+ x. f x \<partial>M) = integral\<^sup>L M f"
proof (induct rule: borel_measurable_induct_real)
@@ -1687,7 +1687,7 @@
finally show ?thesis .
qed
-lemma%unimportant nn_integral_eq_integrable:
+lemma nn_integral_eq_integrable:
assumes f: "f \<in> M \<rightarrow>\<^sub>M borel" "AE x in M. 0 \<le> f x" and "0 \<le> x"
shows "(\<integral>\<^sup>+x. f x \<partial>M) = ennreal x \<longleftrightarrow> (integrable M f \<and> integral\<^sup>L M f = x)"
proof (safe intro!: nn_integral_eq_integral assms)
@@ -1697,7 +1697,7 @@
by (simp_all add: * assms integral_nonneg_AE)
qed
-lemma%unimportant (* FIX ME needs name*)
+lemma (* FIX ME needs name*)
fixes f :: "_ \<Rightarrow> _ \<Rightarrow> 'a :: {banach, second_countable_topology}"
assumes integrable[measurable]: "\<And>i. integrable M (f i)"
and summable: "AE x in M. summable (\<lambda>i. norm (f i x))"
@@ -1748,10 +1748,10 @@
unfolding sums_iff by auto
qed
-lemma%important integral_norm_bound [simp]:
+proposition integral_norm_bound [simp]:
fixes f :: "_ \<Rightarrow> 'a :: {banach, second_countable_topology}"
shows "norm (integral\<^sup>L M f) \<le> (\<integral>x. norm (f x) \<partial>M)"
-proof%unimportant (cases "integrable M f")
+proof (cases "integrable M f")
case True then show ?thesis
using nn_integral_eq_integral[of M "\<lambda>x. norm (f x)"] integral_norm_bound_ennreal[of M f]
by (simp add: integral_nonneg_AE)
@@ -1762,11 +1762,11 @@
ultimately show ?thesis by simp
qed
-lemma%unimportant integral_abs_bound [simp]:
+proposition integral_abs_bound [simp]:
fixes f :: "'a \<Rightarrow> real" shows "abs (\<integral>x. f x \<partial>M) \<le> (\<integral>x. \<bar>f x\<bar> \<partial>M)"
using integral_norm_bound[of M f] by auto
-lemma%unimportant integral_eq_nn_integral:
+lemma integral_eq_nn_integral:
assumes [measurable]: "f \<in> borel_measurable M"
assumes nonneg: "AE x in M. 0 \<le> f x"
shows "integral\<^sup>L M f = enn2real (\<integral>\<^sup>+ x. ennreal (f x) \<partial>M)"
@@ -1787,7 +1787,7 @@
by (simp add: integral_nonneg_AE)
qed
-lemma%unimportant enn2real_nn_integral_eq_integral:
+lemma enn2real_nn_integral_eq_integral:
assumes eq: "AE x in M. f x = ennreal (g x)" and nn: "AE x in M. 0 \<le> g x"
and fin: "(\<integral>\<^sup>+x. f x \<partial>M) < top"
and [measurable]: "g \<in> M \<rightarrow>\<^sub>M borel"
@@ -1812,20 +1812,20 @@
using nn by (simp add: integral_nonneg_AE)
qed
-lemma%unimportant has_bochner_integral_nn_integral:
+lemma has_bochner_integral_nn_integral:
assumes "f \<in> borel_measurable M" "AE x in M. 0 \<le> f x" "0 \<le> x"
assumes "(\<integral>\<^sup>+x. f x \<partial>M) = ennreal x"
shows "has_bochner_integral M f x"
unfolding has_bochner_integral_iff
using assms by (auto simp: assms integral_eq_nn_integral intro: integrableI_nn_integral_finite)
-lemma%unimportant integrableI_simple_bochner_integrable:
+lemma integrableI_simple_bochner_integrable:
fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
shows "simple_bochner_integrable M f \<Longrightarrow> integrable M f"
by (intro integrableI_sequence[where s="\<lambda>_. f"] borel_measurable_simple_function)
(auto simp: zero_ennreal_def[symmetric] simple_bochner_integrable.simps)
-lemma%important integrable_induct[consumes 1, case_names base add lim, induct pred: integrable]:
+proposition integrable_induct[consumes 1, case_names base add lim, induct pred: integrable]:
fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
assumes "integrable M f"
assumes base: "\<And>A c. A \<in> sets M \<Longrightarrow> emeasure M A < \<infinity> \<Longrightarrow> P (\<lambda>x. indicator A x *\<^sub>R c)"
@@ -1834,7 +1834,7 @@
(\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>i. s i x) \<longlonglongrightarrow> f x) \<Longrightarrow>
(\<And>i x. x \<in> space M \<Longrightarrow> norm (s i x) \<le> 2 * norm (f x)) \<Longrightarrow> integrable M f \<Longrightarrow> P f"
shows "P f"
-proof%unimportant -
+proof -
from \<open>integrable M f\<close> have f: "f \<in> borel_measurable M" "(\<integral>\<^sup>+x. norm (f x) \<partial>M) < \<infinity>"
unfolding integrable_iff_bounded by auto
from borel_measurable_implies_sequence_metric[OF f(1)]
@@ -1895,12 +1895,12 @@
qed fact
qed
-lemma%unimportant integral_eq_zero_AE:
+lemma integral_eq_zero_AE:
"(AE x in M. f x = 0) \<Longrightarrow> integral\<^sup>L M f = 0"
using integral_cong_AE[of f M "\<lambda>_. 0"]
by (cases "integrable M f") (simp_all add: not_integrable_integral_eq)
-lemma%unimportant integral_nonneg_eq_0_iff_AE:
+lemma integral_nonneg_eq_0_iff_AE:
fixes f :: "_ \<Rightarrow> real"
assumes f[measurable]: "integrable M f" and nonneg: "AE x in M. 0 \<le> f x"
shows "integral\<^sup>L M f = 0 \<longleftrightarrow> (AE x in M. f x = 0)"
@@ -1914,7 +1914,7 @@
by auto
qed (auto simp add: integral_eq_zero_AE)
-lemma%unimportant integral_mono_AE:
+lemma integral_mono_AE:
fixes f :: "'a \<Rightarrow> real"
assumes "integrable M f" "integrable M g" "AE x in M. f x \<le> g x"
shows "integral\<^sup>L M f \<le> integral\<^sup>L M g"
@@ -1926,7 +1926,7 @@
finally show ?thesis by simp
qed
-lemma%unimportant integral_mono:
+lemma integral_mono:
fixes f :: "'a \<Rightarrow> real"
shows "integrable M f \<Longrightarrow> integrable M g \<Longrightarrow> (\<And>x. x \<in> space M \<Longrightarrow> f x \<le> g x) \<Longrightarrow>
integral\<^sup>L M f \<le> integral\<^sup>L M g"
@@ -1936,11 +1936,11 @@
integrability assumption. The price to pay is that the upper function has to be nonnegative,
but this is often true and easy to check in computations.\<close>
-lemma%important integral_mono_AE':
+lemma integral_mono_AE':
fixes f::"_ \<Rightarrow> real"
assumes "integrable M f" "AE x in M. g x \<le> f x" "AE x in M. 0 \<le> f x"
shows "(\<integral>x. g x \<partial>M) \<le> (\<integral>x. f x \<partial>M)"
-proof%unimportant (cases "integrable M g")
+proof (cases "integrable M g")
case True
show ?thesis by (rule integral_mono_AE, auto simp add: assms True)
next
@@ -1950,16 +1950,16 @@
finally show ?thesis by simp
qed
-lemma%important integral_mono':
+lemma integral_mono':
fixes f::"_ \<Rightarrow> real"
assumes "integrable M f" "\<And>x. x \<in> space M \<Longrightarrow> g x \<le> f x" "\<And>x. x \<in> space M \<Longrightarrow> 0 \<le> f x"
shows "(\<integral>x. g x \<partial>M) \<le> (\<integral>x. f x \<partial>M)"
by (rule integral_mono_AE', insert assms, auto)
-lemma%important (in finite_measure) integrable_measure:
+lemma (in finite_measure) integrable_measure:
assumes I: "disjoint_family_on X I" "countable I"
shows "integrable (count_space I) (\<lambda>i. measure M (X i))"
-proof%unimportant -
+proof -
have "(\<integral>\<^sup>+i. measure M (X i) \<partial>count_space I) = (\<integral>\<^sup>+i. measure M (if X i \<in> sets M then X i else {}) \<partial>count_space I)"
by (auto intro!: nn_integral_cong measure_notin_sets)
also have "\<dots> = measure M (\<Union>i\<in>I. if X i \<in> sets M then X i else {})"
@@ -1969,7 +1969,7 @@
by (auto intro!: integrableI_bounded)
qed
-lemma%unimportant integrableI_real_bounded:
+lemma integrableI_real_bounded:
assumes f: "f \<in> borel_measurable M" and ae: "AE x in M. 0 \<le> f x" and fin: "integral\<^sup>N M f < \<infinity>"
shows "integrable M f"
proof (rule integrableI_bounded)
@@ -1979,13 +1979,13 @@
finally show "(\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M) < \<infinity>" .
qed fact
-lemma%unimportant nn_integral_nonneg_infinite:
+lemma nn_integral_nonneg_infinite:
fixes f::"'a \<Rightarrow> real"
assumes "f \<in> borel_measurable M" "\<not> integrable M f" "AE x in M. f x \<ge> 0"
shows "(\<integral>\<^sup>+x. f x \<partial>M) = \<infinity>"
using assms integrableI_real_bounded less_top by auto
-lemma%unimportant integral_real_bounded:
+lemma integral_real_bounded:
assumes "0 \<le> r" "integral\<^sup>N M f \<le> ennreal r"
shows "integral\<^sup>L M f \<le> r"
proof cases
@@ -2009,22 +2009,22 @@
using \<open>0 \<le> r\<close> by (simp add: not_integrable_integral_eq)
qed
-lemma%unimportant integrable_MIN:
+lemma integrable_MIN:
fixes f:: "_ \<Rightarrow> _ \<Rightarrow> real"
shows "\<lbrakk> finite I; I \<noteq> {}; \<And>i. i \<in> I \<Longrightarrow> integrable M (f i) \<rbrakk>
\<Longrightarrow> integrable M (\<lambda>x. MIN i\<in>I. f i x)"
by (induct rule: finite_ne_induct) simp+
-lemma%unimportant integrable_MAX:
+lemma integrable_MAX:
fixes f:: "_ \<Rightarrow> _ \<Rightarrow> real"
shows "\<lbrakk> finite I; I \<noteq> {}; \<And>i. i \<in> I \<Longrightarrow> integrable M (f i) \<rbrakk>
\<Longrightarrow> integrable M (\<lambda>x. MAX i\<in>I. f i x)"
by (induct rule: finite_ne_induct) simp+
-lemma%important integral_Markov_inequality:
+theorem integral_Markov_inequality:
assumes [measurable]: "integrable M u" and "AE x in M. 0 \<le> u x" "0 < (c::real)"
shows "(emeasure M) {x\<in>space M. u x \<ge> c} \<le> (1/c) * (\<integral>x. u x \<partial>M)"
-proof%unimportant -
+proof -
have "(\<integral>\<^sup>+ x. ennreal(u x) * indicator (space M) x \<partial>M) \<le> (\<integral>\<^sup>+ x. u x \<partial>M)"
by (rule nn_integral_mono_AE, auto simp add: \<open>c>0\<close> less_eq_real_def)
also have "... = (\<integral>x. u x \<partial>M)"
@@ -2044,7 +2044,7 @@
using \<open>0<c\<close> by (simp add: ennreal_mult'[symmetric])
qed
-lemma%unimportant integral_ineq_eq_0_then_AE:
+lemma integral_ineq_eq_0_then_AE:
fixes f::"_ \<Rightarrow> real"
assumes "AE x in M. f x \<le> g x" "integrable M f" "integrable M g"
"(\<integral>x. f x \<partial>M) = (\<integral>x. g x \<partial>M)"
@@ -2057,7 +2057,7 @@
then show ?thesis unfolding h_def by auto
qed
-lemma%unimportant not_AE_zero_int_E:
+lemma not_AE_zero_int_E:
fixes f::"'a \<Rightarrow> real"
assumes "AE x in M. f x \<ge> 0" "(\<integral>x. f x \<partial>M) > 0"
and [measurable]: "f \<in> borel_measurable M"
@@ -2069,12 +2069,12 @@
then show False using assms(2) by simp
qed
-lemma%important tendsto_L1_int:
+proposition tendsto_L1_int:
fixes u :: "_ \<Rightarrow> _ \<Rightarrow> 'b::{banach, second_countable_topology}"
assumes [measurable]: "\<And>n. integrable M (u n)" "integrable M f"
and "((\<lambda>n. (\<integral>\<^sup>+x. norm(u n x - f x) \<partial>M)) \<longlongrightarrow> 0) F"
shows "((\<lambda>n. (\<integral>x. u n x \<partial>M)) \<longlongrightarrow> (\<integral>x. f x \<partial>M)) F"
-proof%unimportant -
+proof -
have "((\<lambda>n. norm((\<integral>x. u n x \<partial>M) - (\<integral>x. f x \<partial>M))) \<longlongrightarrow> (0::ennreal)) F"
proof (rule tendsto_sandwich[of "\<lambda>_. 0", where ?h = "\<lambda>n. (\<integral>\<^sup>+x. norm(u n x - f x) \<partial>M)"], auto simp add: assms)
{
@@ -2103,12 +2103,12 @@
text \<open>The next lemma asserts that, if a sequence of functions converges in \<open>L\<^sup>1\<close>, then
it admits a subsequence that converges almost everywhere.\<close>
-lemma%important tendsto_L1_AE_subseq:
+proposition tendsto_L1_AE_subseq:
fixes u :: "nat \<Rightarrow> 'a \<Rightarrow> 'b::{banach, second_countable_topology}"
assumes [measurable]: "\<And>n. integrable M (u n)"
and "(\<lambda>n. (\<integral>x. norm(u n x) \<partial>M)) \<longlonglongrightarrow> 0"
shows "\<exists>r::nat\<Rightarrow>nat. strict_mono r \<and> (AE x in M. (\<lambda>n. u (r n) x) \<longlonglongrightarrow> 0)"
-proof%unimportant -
+proof -
{
fix k
have "eventually (\<lambda>n. (\<integral>x. norm(u n x) \<partial>M) < (1/2)^k) sequentially"
@@ -2203,7 +2203,7 @@
subsection%important \<open>Restricted measure spaces\<close>
-lemma%unimportant integrable_restrict_space:
+lemma integrable_restrict_space:
fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
assumes \<Omega>[simp]: "\<Omega> \<inter> space M \<in> sets M"
shows "integrable (restrict_space M \<Omega>) f \<longleftrightarrow> integrable M (\<lambda>x. indicator \<Omega> x *\<^sub>R f x)"
@@ -2212,11 +2212,11 @@
nn_integral_restrict_space[OF \<Omega>]
by (simp add: ac_simps ennreal_indicator ennreal_mult)
-lemma%important integral_restrict_space:
+lemma integral_restrict_space:
fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
assumes \<Omega>[simp]: "\<Omega> \<inter> space M \<in> sets M"
shows "integral\<^sup>L (restrict_space M \<Omega>) f = integral\<^sup>L M (\<lambda>x. indicator \<Omega> x *\<^sub>R f x)"
-proof%unimportant (rule integral_eq_cases)
+proof (rule integral_eq_cases)
assume "integrable (restrict_space M \<Omega>) f"
then show ?thesis
proof induct
@@ -2243,7 +2243,7 @@
qed
qed (simp add: integrable_restrict_space)
-lemma%unimportant integral_empty:
+lemma integral_empty:
assumes "space M = {}"
shows "integral\<^sup>L M f = 0"
proof -
@@ -2254,7 +2254,7 @@
subsection%important \<open>Measure spaces with an associated density\<close>
-lemma%unimportant integrable_density:
+lemma integrable_density:
fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}" and g :: "'a \<Rightarrow> real"
assumes [measurable]: "f \<in> borel_measurable M" "g \<in> borel_measurable M"
and nn: "AE x in M. 0 \<le> g x"
@@ -2265,12 +2265,12 @@
apply (auto simp: ennreal_mult)
done
-lemma%important integral_density:
+lemma integral_density:
fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}" and g :: "'a \<Rightarrow> real"
assumes f: "f \<in> borel_measurable M"
and g[measurable]: "g \<in> borel_measurable M" "AE x in M. 0 \<le> g x"
shows "integral\<^sup>L (density M g) f = integral\<^sup>L M (\<lambda>x. g x *\<^sub>R f x)"
-proof%unimportant (rule integral_eq_cases)
+proof (rule integral_eq_cases)
assume "integrable (density M g) f"
then show ?thesis
proof induct
@@ -2325,14 +2325,14 @@
qed
qed (simp add: f g integrable_density)
-lemma%unimportant (*FIX ME needs name *)
+lemma (*FIX ME needs name *)
fixes g :: "'a \<Rightarrow> real"
assumes "f \<in> borel_measurable M" "AE x in M. 0 \<le> f x" "g \<in> borel_measurable M"
shows integral_real_density: "integral\<^sup>L (density M f) g = (\<integral> x. f x * g x \<partial>M)"
and integrable_real_density: "integrable (density M f) g \<longleftrightarrow> integrable M (\<lambda>x. f x * g x)"
using assms integral_density[of g M f] integrable_density[of g M f] by auto
-lemma%important has_bochner_integral_density:
+lemma has_bochner_integral_density:
fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}" and g :: "'a \<Rightarrow> real"
shows "f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> (AE x in M. 0 \<le> g x) \<Longrightarrow>
has_bochner_integral M (\<lambda>x. g x *\<^sub>R f x) x \<Longrightarrow> has_bochner_integral (density M g) f x"
@@ -2340,23 +2340,23 @@
subsection%important \<open>Distributions\<close>
-lemma%unimportant integrable_distr_eq:
+lemma integrable_distr_eq:
fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
assumes [measurable]: "g \<in> measurable M N" "f \<in> borel_measurable N"
shows "integrable (distr M N g) f \<longleftrightarrow> integrable M (\<lambda>x. f (g x))"
unfolding integrable_iff_bounded by (simp_all add: nn_integral_distr)
-lemma%unimportant integrable_distr:
+lemma integrable_distr:
fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
shows "T \<in> measurable M M' \<Longrightarrow> integrable (distr M M' T) f \<Longrightarrow> integrable M (\<lambda>x. f (T x))"
by (subst integrable_distr_eq[symmetric, where g=T])
(auto dest: borel_measurable_integrable)
-lemma%important integral_distr:
+lemma integral_distr:
fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
assumes g[measurable]: "g \<in> measurable M N" and f: "f \<in> borel_measurable N"
shows "integral\<^sup>L (distr M N g) f = integral\<^sup>L M (\<lambda>x. f (g x))"
-proof%unimportant (rule integral_eq_cases)
+proof (rule integral_eq_cases)
assume "integrable (distr M N g) f"
then show ?thesis
proof induct
@@ -2404,27 +2404,27 @@
qed
qed (simp add: f g integrable_distr_eq)
-lemma%important has_bochner_integral_distr:
+lemma has_bochner_integral_distr:
fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
shows "f \<in> borel_measurable N \<Longrightarrow> g \<in> measurable M N \<Longrightarrow>
has_bochner_integral M (\<lambda>x. f (g x)) x \<Longrightarrow> has_bochner_integral (distr M N g) f x"
- by%unimportant (simp add: has_bochner_integral_iff integrable_distr_eq integral_distr)
+ by (simp add: has_bochner_integral_iff integrable_distr_eq integral_distr)
subsection%important \<open>Lebesgue integration on \<^const>\<open>count_space\<close>\<close>
-lemma%unimportant integrable_count_space:
+lemma integrable_count_space:
fixes f :: "'a \<Rightarrow> 'b::{banach,second_countable_topology}"
shows "finite X \<Longrightarrow> integrable (count_space X) f"
by (auto simp: nn_integral_count_space integrable_iff_bounded)
-lemma%unimportant measure_count_space[simp]:
+lemma measure_count_space[simp]:
"B \<subseteq> A \<Longrightarrow> finite B \<Longrightarrow> measure (count_space A) B = card B"
unfolding measure_def by (subst emeasure_count_space ) auto
-lemma%important lebesgue_integral_count_space_finite_support:
+lemma lebesgue_integral_count_space_finite_support:
assumes f: "finite {a\<in>A. f a \<noteq> 0}"
shows "(\<integral>x. f x \<partial>count_space A) = (\<Sum>a | a \<in> A \<and> f a \<noteq> 0. f a)"
-proof%unimportant -
+proof -
have eq: "\<And>x. x \<in> A \<Longrightarrow> (\<Sum>a | x = a \<and> a \<in> A \<and> f a \<noteq> 0. f a) = (\<Sum>x\<in>{x}. f x)"
by (intro sum.mono_neutral_cong_left) auto
@@ -2436,17 +2436,17 @@
by auto
qed
-lemma%unimportant lebesgue_integral_count_space_finite: "finite A \<Longrightarrow> (\<integral>x. f x \<partial>count_space A) = (\<Sum>a\<in>A. f a)"
+lemma lebesgue_integral_count_space_finite: "finite A \<Longrightarrow> (\<integral>x. f x \<partial>count_space A) = (\<Sum>a\<in>A. f a)"
by (subst lebesgue_integral_count_space_finite_support)
(auto intro!: sum.mono_neutral_cong_left)
-lemma%unimportant integrable_count_space_nat_iff:
+lemma integrable_count_space_nat_iff:
fixes f :: "nat \<Rightarrow> _::{banach,second_countable_topology}"
shows "integrable (count_space UNIV) f \<longleftrightarrow> summable (\<lambda>x. norm (f x))"
by (auto simp add: integrable_iff_bounded nn_integral_count_space_nat ennreal_suminf_neq_top
intro: summable_suminf_not_top)
-lemma%unimportant sums_integral_count_space_nat:
+lemma sums_integral_count_space_nat:
fixes f :: "nat \<Rightarrow> _::{banach,second_countable_topology}"
assumes *: "integrable (count_space UNIV) f"
shows "f sums (integral\<^sup>L (count_space UNIV) f)"
@@ -2471,18 +2471,18 @@
finally show ?thesis .
qed
-lemma%unimportant integral_count_space_nat:
+lemma integral_count_space_nat:
fixes f :: "nat \<Rightarrow> _::{banach,second_countable_topology}"
shows "integrable (count_space UNIV) f \<Longrightarrow> integral\<^sup>L (count_space UNIV) f = (\<Sum>x. f x)"
using sums_integral_count_space_nat by (rule sums_unique)
-lemma%unimportant integrable_bij_count_space:
+lemma integrable_bij_count_space:
fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
assumes g: "bij_betw g A B"
shows "integrable (count_space A) (\<lambda>x. f (g x)) \<longleftrightarrow> integrable (count_space B) f"
unfolding integrable_iff_bounded by (subst nn_integral_bij_count_space[OF g]) auto
-lemma%important integral_bij_count_space:
+lemma integral_bij_count_space:
fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
assumes g: "bij_betw g A B"
shows "integral\<^sup>L (count_space A) (\<lambda>x. f (g x)) = integral\<^sup>L (count_space B) f"
@@ -2492,23 +2492,23 @@
apply auto
done
-lemma%important has_bochner_integral_count_space_nat:
+lemma has_bochner_integral_count_space_nat:
fixes f :: "nat \<Rightarrow> _::{banach,second_countable_topology}"
shows "has_bochner_integral (count_space UNIV) f x \<Longrightarrow> f sums x"
unfolding has_bochner_integral_iff by (auto intro!: sums_integral_count_space_nat)
subsection%important \<open>Point measure\<close>
-lemma%unimportant lebesgue_integral_point_measure_finite:
+lemma lebesgue_integral_point_measure_finite:
fixes g :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
shows "finite A \<Longrightarrow> (\<And>a. a \<in> A \<Longrightarrow> 0 \<le> f a) \<Longrightarrow>
integral\<^sup>L (point_measure A f) g = (\<Sum>a\<in>A. f a *\<^sub>R g a)"
by (simp add: lebesgue_integral_count_space_finite AE_count_space integral_density point_measure_def)
-lemma%important integrable_point_measure_finite:
+proposition integrable_point_measure_finite:
fixes g :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}" and f :: "'a \<Rightarrow> real"
shows "finite A \<Longrightarrow> integrable (point_measure A f) g"
- unfolding%unimportant point_measure_def
+ unfolding point_measure_def
apply (subst density_cong[where f'="\<lambda>x. ennreal (max 0 (f x))"])
apply (auto split: split_max simp: ennreal_neg)
apply (subst integrable_density)
@@ -2517,24 +2517,23 @@
subsection%important \<open>Lebesgue integration on \<^const>\<open>null_measure\<close>\<close>
-lemma%unimportant has_bochner_integral_null_measure_iff[iff]:
+lemma has_bochner_integral_null_measure_iff[iff]:
"has_bochner_integral (null_measure M) f 0 \<longleftrightarrow> f \<in> borel_measurable M"
by (auto simp add: has_bochner_integral.simps simple_bochner_integral_def[abs_def]
intro!: exI[of _ "\<lambda>n x. 0"] simple_bochner_integrable.intros)
-lemma%important integrable_null_measure_iff[iff]: "integrable (null_measure M) f \<longleftrightarrow> f \<in> borel_measurable M"
- by%unimportant (auto simp add: integrable.simps)
-
-lemma%unimportant integral_null_measure[simp]: "integral\<^sup>L (null_measure M) f = 0"
+lemma integrable_null_measure_iff[iff]: "integrable (null_measure M) f \<longleftrightarrow> f \<in> borel_measurable M"
+ by (auto simp add: integrable.simps)
+
+lemma integral_null_measure[simp]: "integral\<^sup>L (null_measure M) f = 0"
by (cases "integrable (null_measure M) f")
(auto simp add: not_integrable_integral_eq has_bochner_integral_integral_eq)
subsection%important \<open>Legacy lemmas for the real-valued Lebesgue integral\<close>
-
-lemma%important real_lebesgue_integral_def:
+theorem real_lebesgue_integral_def:
assumes f[measurable]: "integrable M f"
shows "integral\<^sup>L M f = enn2real (\<integral>\<^sup>+x. f x \<partial>M) - enn2real (\<integral>\<^sup>+x. ennreal (- f x) \<partial>M)"
-proof%unimportant -
+proof -
have "integral\<^sup>L M f = integral\<^sup>L M (\<lambda>x. max 0 (f x) - max 0 (- f x))"
by (auto intro!: arg_cong[where f="integral\<^sup>L M"])
also have "\<dots> = integral\<^sup>L M (\<lambda>x. max 0 (f x)) - integral\<^sup>L M (\<lambda>x. max 0 (- f x))"
@@ -2546,11 +2545,11 @@
finally show ?thesis .
qed
-lemma%important real_integrable_def:
+theorem real_integrable_def:
"integrable M f \<longleftrightarrow> f \<in> borel_measurable M \<and>
(\<integral>\<^sup>+ x. ennreal (f x) \<partial>M) \<noteq> \<infinity> \<and> (\<integral>\<^sup>+ x. ennreal (- f x) \<partial>M) \<noteq> \<infinity>"
unfolding integrable_iff_bounded
-proof%unimportant (safe del: notI)
+proof (safe del: notI)
assume *: "(\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M) < \<infinity>"
have "(\<integral>\<^sup>+ x. ennreal (f x) \<partial>M) \<le> (\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M)"
by (intro nn_integral_mono) auto
@@ -2574,12 +2573,12 @@
finally show "(\<integral>\<^sup>+ x. norm (f x) \<partial>M) < \<infinity>" .
qed
-lemma%unimportant integrableD[dest]:
+lemma integrableD[dest]:
assumes "integrable M f"
shows "f \<in> borel_measurable M" "(\<integral>\<^sup>+ x. ennreal (f x) \<partial>M) \<noteq> \<infinity>" "(\<integral>\<^sup>+ x. ennreal (- f x) \<partial>M) \<noteq> \<infinity>"
using assms unfolding real_integrable_def by auto
-lemma%unimportant integrableE:
+lemma integrableE:
assumes "integrable M f"
obtains r q where "0 \<le> r" "0 \<le> q"
"(\<integral>\<^sup>+x. ennreal (f x)\<partial>M) = ennreal r"
@@ -2588,7 +2587,7 @@
using assms unfolding real_integrable_def real_lebesgue_integral_def[OF assms]
by (cases rule: ennreal2_cases[of "(\<integral>\<^sup>+x. ennreal (-f x)\<partial>M)" "(\<integral>\<^sup>+x. ennreal (f x)\<partial>M)"]) auto
-lemma%important integral_monotone_convergence_nonneg:
+lemma integral_monotone_convergence_nonneg:
fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> real"
assumes i: "\<And>i. integrable M (f i)" and mono: "AE x in M. mono (\<lambda>n. f n x)"
and pos: "\<And>i. AE x in M. 0 \<le> f i x"
@@ -2597,7 +2596,7 @@
and u: "u \<in> borel_measurable M"
shows "integrable M u"
and "integral\<^sup>L M u = x"
-proof%unimportant -
+proof -
have nn: "AE x in M. \<forall>i. 0 \<le> f i x"
using pos unfolding AE_all_countable by auto
with lim have u_nn: "AE x in M. 0 \<le> u x"
@@ -2629,7 +2628,7 @@
by (auto simp: real_integrable_def real_lebesgue_integral_def u)
qed
-lemma%unimportant (*FIX ME needs name *)
+lemma (*FIX ME needs name *)
fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> real"
assumes f: "\<And>i. integrable M (f i)" and mono: "AE x in M. mono (\<lambda>n. f n x)"
and lim: "AE x in M. (\<lambda>i. f i x) \<longlonglongrightarrow> u x"
@@ -2660,7 +2659,7 @@
by (metis has_bochner_integral_integrable)
qed
-lemma%unimportant integral_norm_eq_0_iff:
+lemma integral_norm_eq_0_iff:
fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
assumes f[measurable]: "integrable M f"
shows "(\<integral>x. norm (f x) \<partial>M) = 0 \<longleftrightarrow> emeasure M {x\<in>space M. f x \<noteq> 0} = 0"
@@ -2675,18 +2674,18 @@
by simp
qed
-lemma%unimportant integral_0_iff:
+lemma integral_0_iff:
fixes f :: "'a \<Rightarrow> real"
shows "integrable M f \<Longrightarrow> (\<integral>x. \<bar>f x\<bar> \<partial>M) = 0 \<longleftrightarrow> emeasure M {x\<in>space M. f x \<noteq> 0} = 0"
using integral_norm_eq_0_iff[of M f] by simp
-lemma%unimportant (in finite_measure) integrable_const[intro!, simp]: "integrable M (\<lambda>x. a)"
+lemma (in finite_measure) integrable_const[intro!, simp]: "integrable M (\<lambda>x. a)"
using integrable_indicator[of "space M" M a] by (simp cong: integrable_cong add: less_top[symmetric])
-lemma%important lebesgue_integral_const[simp]:
+lemma lebesgue_integral_const[simp]:
fixes a :: "'a :: {banach, second_countable_topology}"
shows "(\<integral>x. a \<partial>M) = measure M (space M) *\<^sub>R a"
-proof%unimportant -
+proof -
{ assume "emeasure M (space M) = \<infinity>" "a \<noteq> 0"
then have ?thesis
by (auto simp add: not_integrable_integral_eq ennreal_mult_less_top measure_def integrable_iff_bounded) }
@@ -2704,7 +2703,7 @@
ultimately show ?thesis by blast
qed
-lemma%unimportant (in finite_measure) integrable_const_bound:
+lemma (in finite_measure) integrable_const_bound:
fixes f :: "'a \<Rightarrow> 'b::{banach,second_countable_topology}"
shows "AE x in M. norm (f x) \<le> B \<Longrightarrow> f \<in> borel_measurable M \<Longrightarrow> integrable M f"
apply (rule integrable_bound[OF integrable_const[of B], of f])
@@ -2713,19 +2712,19 @@
apply auto
done
-lemma%unimportant (in finite_measure) integral_bounded_eq_bound_then_AE:
+lemma (in finite_measure) integral_bounded_eq_bound_then_AE:
assumes "AE x in M. f x \<le> (c::real)"
"integrable M f" "(\<integral>x. f x \<partial>M) = c * measure M (space M)"
shows "AE x in M. f x = c"
apply (rule integral_ineq_eq_0_then_AE) using assms by auto
-lemma%important integral_indicator_finite_real:
+lemma integral_indicator_finite_real:
fixes f :: "'a \<Rightarrow> real"
assumes [simp]: "finite A"
assumes [measurable]: "\<And>a. a \<in> A \<Longrightarrow> {a} \<in> sets M"
assumes finite: "\<And>a. a \<in> A \<Longrightarrow> emeasure M {a} < \<infinity>"
shows "(\<integral>x. f x * indicator A x \<partial>M) = (\<Sum>a\<in>A. f a * measure M {a})"
-proof%unimportant -
+proof -
have "(\<integral>x. f x * indicator A x \<partial>M) = (\<integral>x. (\<Sum>a\<in>A. f a * indicator {a} x) \<partial>M)"
proof (intro integral_cong refl)
fix x show "f x * indicator A x = (\<Sum>a\<in>A. f a * indicator {a} x)"
@@ -2736,7 +2735,7 @@
finally show ?thesis .
qed
-lemma%unimportant (in finite_measure) ennreal_integral_real:
+lemma (in finite_measure) ennreal_integral_real:
assumes [measurable]: "f \<in> borel_measurable M"
assumes ae: "AE x in M. f x \<le> ennreal B" "0 \<le> B"
shows "ennreal (\<integral>x. enn2real (f x) \<partial>M) = (\<integral>\<^sup>+x. f x \<partial>M)"
@@ -2747,7 +2746,7 @@
using ae by (intro nn_integral_cong_AE) (auto simp: le_less_trans[OF _ ennreal_less_top])
qed auto
-lemma%unimportant (in finite_measure) integral_less_AE:
+lemma (in finite_measure) integral_less_AE:
fixes X Y :: "'a \<Rightarrow> real"
assumes int: "integrable M X" "integrable M Y"
assumes A: "(emeasure M) A \<noteq> 0" "A \<in> sets M" "AE x in M. x \<in> A \<longrightarrow> X x \<noteq> Y x"
@@ -2778,14 +2777,14 @@
ultimately show ?thesis by auto
qed
-lemma%unimportant (in finite_measure) integral_less_AE_space:
+lemma (in finite_measure) integral_less_AE_space:
fixes X Y :: "'a \<Rightarrow> real"
assumes int: "integrable M X" "integrable M Y"
assumes gt: "AE x in M. X x < Y x" "emeasure M (space M) \<noteq> 0"
shows "integral\<^sup>L M X < integral\<^sup>L M Y"
using gt by (intro integral_less_AE[OF int, where A="space M"]) auto
-lemma%unimportant tendsto_integral_at_top:
+lemma tendsto_integral_at_top:
fixes f :: "real \<Rightarrow> 'a::{banach, second_countable_topology}"
assumes [measurable_cong]: "sets M = sets borel" and f[measurable]: "integrable M f"
shows "((\<lambda>y. \<integral> x. indicator {.. y} x *\<^sub>R f x \<partial>M) \<longlongrightarrow> \<integral> x. f x \<partial>M) at_top"
@@ -2809,7 +2808,7 @@
qed auto
qed
-lemma%unimportant (*FIX ME needs name *)
+lemma (*FIX ME needs name *)
fixes f :: "real \<Rightarrow> real"
assumes M: "sets M = sets borel"
assumes nonneg: "AE x in M. 0 \<le> f x"
@@ -2842,28 +2841,28 @@
subsection%important \<open>Product measure\<close>
-lemma%important (in sigma_finite_measure) borel_measurable_lebesgue_integrable[measurable (raw)]:
+lemma (in sigma_finite_measure) borel_measurable_lebesgue_integrable[measurable (raw)]:
fixes f :: "_ \<Rightarrow> _ \<Rightarrow> _::{banach, second_countable_topology}"
assumes [measurable]: "case_prod f \<in> borel_measurable (N \<Otimes>\<^sub>M M)"
shows "Measurable.pred N (\<lambda>x. integrable M (f x))"
-proof%unimportant -
+proof -
have [simp]: "\<And>x. x \<in> space N \<Longrightarrow> integrable M (f x) \<longleftrightarrow> (\<integral>\<^sup>+y. norm (f x y) \<partial>M) < \<infinity>"
unfolding integrable_iff_bounded by simp
show ?thesis
by (simp cong: measurable_cong)
qed
-lemma%unimportant (in sigma_finite_measure) measurable_measure[measurable (raw)]:
+lemma (in sigma_finite_measure) measurable_measure[measurable (raw)]:
"(\<And>x. x \<in> space N \<Longrightarrow> A x \<subseteq> space M) \<Longrightarrow>
{x \<in> space (N \<Otimes>\<^sub>M M). snd x \<in> A (fst x)} \<in> sets (N \<Otimes>\<^sub>M M) \<Longrightarrow>
(\<lambda>x. measure M (A x)) \<in> borel_measurable N"
unfolding measure_def by (intro measurable_emeasure borel_measurable_enn2real) auto
-lemma%important (in sigma_finite_measure) borel_measurable_lebesgue_integral[measurable (raw)]:
+proposition (in sigma_finite_measure) borel_measurable_lebesgue_integral[measurable (raw)]:
fixes f :: "_ \<Rightarrow> _ \<Rightarrow> _::{banach, second_countable_topology}"
assumes f[measurable]: "case_prod f \<in> borel_measurable (N \<Otimes>\<^sub>M M)"
shows "(\<lambda>x. \<integral>y. f x y \<partial>M) \<in> borel_measurable N"
-proof%unimportant -
+proof -
from borel_measurable_implies_sequence_metric[OF f, of 0] guess s ..
then have s: "\<And>i. simple_function (N \<Otimes>\<^sub>M M) (s i)"
"\<And>x y. x \<in> space N \<Longrightarrow> y \<in> space M \<Longrightarrow> (\<lambda>i. s i (x, y)) \<longlonglongrightarrow> f x y"
@@ -2934,7 +2933,7 @@
qed
qed
-lemma%unimportant (in pair_sigma_finite) integrable_product_swap:
+lemma (in pair_sigma_finite) integrable_product_swap:
fixes f :: "_ \<Rightarrow> _::{banach, second_countable_topology}"
assumes "integrable (M1 \<Otimes>\<^sub>M M2) f"
shows "integrable (M2 \<Otimes>\<^sub>M M1) (\<lambda>(x,y). f (y,x))"
@@ -2946,7 +2945,7 @@
(simp add: distr_pair_swap[symmetric] assms)
qed
-lemma%unimportant (in pair_sigma_finite) integrable_product_swap_iff:
+lemma (in pair_sigma_finite) integrable_product_swap_iff:
fixes f :: "_ \<Rightarrow> _::{banach, second_countable_topology}"
shows "integrable (M2 \<Otimes>\<^sub>M M1) (\<lambda>(x,y). f (y,x)) \<longleftrightarrow> integrable (M1 \<Otimes>\<^sub>M M2) f"
proof -
@@ -2955,7 +2954,7 @@
show ?thesis by auto
qed
-lemma%unimportant (in pair_sigma_finite) integral_product_swap:
+lemma (in pair_sigma_finite) integral_product_swap:
fixes f :: "_ \<Rightarrow> _::{banach, second_countable_topology}"
assumes f: "f \<in> borel_measurable (M1 \<Otimes>\<^sub>M M2)"
shows "(\<integral>(x,y). f (y,x) \<partial>(M2 \<Otimes>\<^sub>M M1)) = integral\<^sup>L (M1 \<Otimes>\<^sub>M M2) f"
@@ -2965,13 +2964,13 @@
by (simp add: integral_distr[symmetric, OF measurable_pair_swap' f] distr_pair_swap[symmetric])
qed
-lemma%important (in pair_sigma_finite) Fubini_integrable:
+theorem (in pair_sigma_finite) Fubini_integrable:
fixes f :: "_ \<Rightarrow> _::{banach, second_countable_topology}"
assumes f[measurable]: "f \<in> borel_measurable (M1 \<Otimes>\<^sub>M M2)"
and integ1: "integrable M1 (\<lambda>x. \<integral> y. norm (f (x, y)) \<partial>M2)"
and integ2: "AE x in M1. integrable M2 (\<lambda>y. f (x, y))"
shows "integrable (M1 \<Otimes>\<^sub>M M2) f"
-proof%unimportant (rule integrableI_bounded)
+proof (rule integrableI_bounded)
have "(\<integral>\<^sup>+ p. norm (f p) \<partial>(M1 \<Otimes>\<^sub>M M2)) = (\<integral>\<^sup>+ x. (\<integral>\<^sup>+ y. norm (f (x, y)) \<partial>M2) \<partial>M1)"
by (simp add: M2.nn_integral_fst [symmetric])
also have "\<dots> = (\<integral>\<^sup>+ x. \<bar>\<integral>y. norm (f (x, y)) \<partial>M2\<bar> \<partial>M1)"
@@ -2992,7 +2991,7 @@
finally show "(\<integral>\<^sup>+ p. norm (f p) \<partial>(M1 \<Otimes>\<^sub>M M2)) < \<infinity>" .
qed fact
-lemma%unimportant (in pair_sigma_finite) emeasure_pair_measure_finite:
+lemma (in pair_sigma_finite) emeasure_pair_measure_finite:
assumes A: "A \<in> sets (M1 \<Otimes>\<^sub>M M2)" and finite: "emeasure (M1 \<Otimes>\<^sub>M M2) A < \<infinity>"
shows "AE x in M1. emeasure M2 {y\<in>space M2. (x, y) \<in> A} < \<infinity>"
proof -
@@ -3006,7 +3005,7 @@
ultimately show ?thesis by (auto simp: less_top)
qed
-lemma%unimportant (in pair_sigma_finite) AE_integrable_fst':
+lemma (in pair_sigma_finite) AE_integrable_fst':
fixes f :: "_ \<Rightarrow> _::{banach, second_countable_topology}"
assumes f[measurable]: "integrable (M1 \<Otimes>\<^sub>M M2) f"
shows "AE x in M1. integrable M2 (\<lambda>y. f (x, y))"
@@ -3023,7 +3022,7 @@
(auto simp: integrable_iff_bounded measurable_compose[OF _ borel_measurable_integrable[OF f]] less_top)
qed
-lemma%unimportant (in pair_sigma_finite) integrable_fst':
+lemma (in pair_sigma_finite) integrable_fst':
fixes f :: "_ \<Rightarrow> _::{banach, second_countable_topology}"
assumes f[measurable]: "integrable (M1 \<Otimes>\<^sub>M M2) f"
shows "integrable M1 (\<lambda>x. \<integral>y. f (x, y) \<partial>M2)"
@@ -3040,11 +3039,11 @@
finally show "(\<integral>\<^sup>+ x. ennreal (norm (\<integral> y. f (x, y) \<partial>M2)) \<partial>M1) < \<infinity>" .
qed
-lemma%important (in pair_sigma_finite) integral_fst':
+proposition (in pair_sigma_finite) integral_fst':
fixes f :: "_ \<Rightarrow> _::{banach, second_countable_topology}"
assumes f: "integrable (M1 \<Otimes>\<^sub>M M2) f"
shows "(\<integral>x. (\<integral>y. f (x, y) \<partial>M2) \<partial>M1) = integral\<^sup>L (M1 \<Otimes>\<^sub>M M2) f"
-using f proof%unimportant induct
+using f proof induct
case (base A c)
have A[measurable]: "A \<in> sets (M1 \<Otimes>\<^sub>M M2)" by fact
@@ -3153,7 +3152,7 @@
qed
qed
-lemma%unimportant (in pair_sigma_finite) (* FIX ME needs name *)
+lemma (in pair_sigma_finite) (* FIX ME needs name *)
fixes f :: "_ \<Rightarrow> _ \<Rightarrow> _::{banach, second_countable_topology}"
assumes f: "integrable (M1 \<Otimes>\<^sub>M M2) (case_prod f)"
shows AE_integrable_fst: "AE x in M1. integrable M2 (\<lambda>y. f x y)" (is "?AE")
@@ -3161,7 +3160,7 @@
and integral_fst: "(\<integral>x. (\<integral>y. f x y \<partial>M2) \<partial>M1) = integral\<^sup>L (M1 \<Otimes>\<^sub>M M2) (\<lambda>(x, y). f x y)" (is "?EQ")
using AE_integrable_fst'[OF f] integrable_fst'[OF f] integral_fst'[OF f] by auto
-lemma%unimportant (in pair_sigma_finite) (* FIX ME needs name *)
+lemma (in pair_sigma_finite) (* FIX ME needs name *)
fixes f :: "_ \<Rightarrow> _ \<Rightarrow> _::{banach, second_countable_topology}"
assumes f[measurable]: "integrable (M1 \<Otimes>\<^sub>M M2) (case_prod f)"
shows AE_integrable_snd: "AE y in M2. integrable M1 (\<lambda>x. f x y)" (is "?AE")
@@ -3177,13 +3176,13 @@
using integral_product_swap[of "case_prod f"] by simp
qed
-lemma%unimportant (in pair_sigma_finite) Fubini_integral:
+proposition (in pair_sigma_finite) Fubini_integral:
fixes f :: "_ \<Rightarrow> _ \<Rightarrow> _ :: {banach, second_countable_topology}"
assumes f: "integrable (M1 \<Otimes>\<^sub>M M2) (case_prod f)"
shows "(\<integral>y. (\<integral>x. f x y \<partial>M1) \<partial>M2) = (\<integral>x. (\<integral>y. f x y \<partial>M2) \<partial>M1)"
unfolding integral_snd[OF assms] integral_fst[OF assms] ..
-lemma%unimportant (in product_sigma_finite) product_integral_singleton:
+lemma (in product_sigma_finite) product_integral_singleton:
fixes f :: "_ \<Rightarrow> _::{banach, second_countable_topology}"
shows "f \<in> borel_measurable (M i) \<Longrightarrow> (\<integral>x. f (x i) \<partial>Pi\<^sub>M {i} M) = integral\<^sup>L (M i) f"
apply (subst distr_singleton[symmetric])
@@ -3191,7 +3190,7 @@
apply simp_all
done
-lemma%unimportant (in product_sigma_finite) product_integral_fold:
+lemma (in product_sigma_finite) product_integral_fold:
fixes f :: "_ \<Rightarrow> _::{banach, second_countable_topology}"
assumes IJ[simp]: "I \<inter> J = {}" and fin: "finite I" "finite J"
and f: "integrable (Pi\<^sub>M (I \<union> J) M) f"
@@ -3218,12 +3217,12 @@
done
qed
-lemma%important (in product_sigma_finite) product_integral_insert:
+lemma (in product_sigma_finite) product_integral_insert:
fixes f :: "_ \<Rightarrow> _::{banach, second_countable_topology}"
assumes I: "finite I" "i \<notin> I"
and f: "integrable (Pi\<^sub>M (insert i I) M) f"
shows "integral\<^sup>L (Pi\<^sub>M (insert i I) M) f = (\<integral>x. (\<integral>y. f (x(i:=y)) \<partial>M i) \<partial>Pi\<^sub>M I M)"
-proof%unimportant -
+proof -
have "integral\<^sup>L (Pi\<^sub>M (insert i I) M) f = integral\<^sup>L (Pi\<^sub>M (I \<union> {i}) M) f"
by simp
also have "\<dots> = (\<integral>x. (\<integral>y. f (merge I {i} (x,y)) \<partial>Pi\<^sub>M {i} M) \<partial>Pi\<^sub>M I M)"
@@ -3242,11 +3241,11 @@
finally show ?thesis .
qed
-lemma%important (in product_sigma_finite) product_integrable_prod:
+lemma (in product_sigma_finite) product_integrable_prod:
fixes f :: "'i \<Rightarrow> 'a \<Rightarrow> _::{real_normed_field,banach,second_countable_topology}"
assumes [simp]: "finite I" and integrable: "\<And>i. i \<in> I \<Longrightarrow> integrable (M i) (f i)"
shows "integrable (Pi\<^sub>M I M) (\<lambda>x. (\<Prod>i\<in>I. f i (x i)))" (is "integrable _ ?f")
-proof%unimportant (unfold integrable_iff_bounded, intro conjI)
+proof (unfold integrable_iff_bounded, intro conjI)
interpret finite_product_sigma_finite M I by standard fact
show "?f \<in> borel_measurable (Pi\<^sub>M I M)"
@@ -3261,11 +3260,11 @@
finally show "(\<integral>\<^sup>+ x. ennreal (norm (\<Prod>i\<in>I. f i (x i))) \<partial>Pi\<^sub>M I M) < \<infinity>" .
qed
-lemma%important (in product_sigma_finite) product_integral_prod:
+lemma (in product_sigma_finite) product_integral_prod:
fixes f :: "'i \<Rightarrow> 'a \<Rightarrow> _::{real_normed_field,banach,second_countable_topology}"
assumes "finite I" and integrable: "\<And>i. i \<in> I \<Longrightarrow> integrable (M i) (f i)"
shows "(\<integral>x. (\<Prod>i\<in>I. f i (x i)) \<partial>Pi\<^sub>M I M) = (\<Prod>i\<in>I. integral\<^sup>L (M i) (f i))"
-using assms proof%unimportant induct
+using assms proof induct
case empty
interpret finite_measure "Pi\<^sub>M {} M"
by rule (simp add: space_PiM)
@@ -3284,7 +3283,7 @@
by (simp add: * insert prod subset_insertI)
qed
-lemma%unimportant integrable_subalgebra:
+lemma integrable_subalgebra:
fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
assumes borel: "f \<in> borel_measurable N"
and N: "sets N \<subseteq> sets M" "space N = space M" "\<And>A. A \<in> sets N \<Longrightarrow> emeasure N A = emeasure M A"
@@ -3296,12 +3295,12 @@
using assms by (auto simp: integrable_iff_bounded nn_integral_subalgebra)
qed
-lemma%important integral_subalgebra:
+lemma integral_subalgebra:
fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
assumes borel: "f \<in> borel_measurable N"
and N: "sets N \<subseteq> sets M" "space N = space M" "\<And>A. A \<in> sets N \<Longrightarrow> emeasure N A = emeasure M A"
shows "integral\<^sup>L N f = integral\<^sup>L M f"
-proof%unimportant cases
+proof cases
assume "integrable N f"
then show ?thesis
proof induct
--- a/src/HOL/Analysis/Borel_Space.thy Sun Jan 13 20:25:41 2019 +0100
+++ b/src/HOL/Analysis/Borel_Space.thy Mon Jan 14 11:59:19 2019 +0000
@@ -10,15 +10,15 @@
Measurable Derivative Ordered_Euclidean_Space Extended_Real_Limits
begin
-lemma%unimportant sets_Collect_eventually_sequentially[measurable]:
+lemma 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%unimportant topological_basis_trivial: "topological_basis {A. open A}"
+lemma topological_basis_trivial: "topological_basis {A. open A}"
by (auto simp: topological_basis_def)
lemma%important open_prod_generated: "open = generate_topology {A \<times> B | A B. open A \<and> open B}"
-proof%unimportant -
+proof -
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
@@ -27,31 +27,31 @@
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%unimportant mono_onI:
+lemma 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%unimportant mono_onD:
+lemma 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%unimportant mono_imp_mono_on: "mono f \<Longrightarrow> mono_on f A"
+lemma mono_imp_mono_on: "mono f \<Longrightarrow> mono_on f A"
unfolding mono_def mono_on_def by auto
-lemma%unimportant mono_on_subset: "mono_on f A \<Longrightarrow> B \<subseteq> A \<Longrightarrow> mono_on f B"
+lemma mono_on_subset: "mono_on f A \<Longrightarrow> B \<subseteq> A \<Longrightarrow> mono_on f B"
unfolding mono_on_def by auto
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%unimportant strict_mono_onI:
+lemma 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%unimportant strict_mono_onD:
+lemma 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%unimportant mono_on_greaterD:
+lemma 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%unimportant strict_mono_inv:
+lemma 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%unimportant strict_mono_on_imp_inj_on:
+lemma 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%unimportant strict_mono_on_leD:
+lemma 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%unimportant strict_mono_on_eqD:
+lemma 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%important mono_on_imp_deriv_nonneg:
+proposition 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%unimportant (rule tendsto_lowerbound)
+proof (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%unimportant strict_mono_on_imp_mono_on:
+lemma 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%important mono_on_ctble_discont:
+proposition 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%unimportant -
+proof -
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%important mono_on_ctble_discont_open:
+lemma 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%unimportant -
+proof -
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,23 +209,23 @@
by (rule mono_on_ctble_discont, rule assms)
qed
-lemma%important mono_ctble_discont:
+lemma mono_ctble_discont:
fixes f :: "real \<Rightarrow> real"
assumes "mono f"
shows "countable {a. \<not> isCont f a}"
-using%unimportant assms mono_on_ctble_discont [of f UNIV] unfolding mono_on_def mono_def by auto
+ using assms mono_on_ctble_discont [of f UNIV] unfolding mono_on_def mono_def by auto
-lemma%important has_real_derivative_imp_continuous_on:
+lemma 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)
using assms differentiable_at_withinI real_differentiable_def by blast
-lemma%important closure_contains_Sup:
+lemma closure_contains_Sup:
fixes S :: "real set"
assumes "S \<noteq> {}" "bdd_above S"
shows "Sup S \<in> closure S"
-proof%unimportant -
+proof -
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)
@@ -234,22 +234,22 @@
finally show ?thesis by auto
qed
-lemma%unimportant closed_contains_Sup:
+lemma 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%unimportant closed_subset_contains_Sup:
+lemma 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%important deriv_nonneg_imp_mono:
+proposition 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%unimportant (cases "a < b")
+proof (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
@@ -258,11 +258,11 @@
with g_ab show ?thesis by simp
qed (insert ab, simp)
-lemma%important continuous_interval_vimage_Int:
+lemma 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%unimportant-
+proof-
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
@@ -301,41 +301,41 @@
abbreviation "borel_measurable M \<equiv> measurable M borel"
-lemma%important in_borel_measurable:
+lemma 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%unimportant (auto simp add: measurable_def borel_def)
+ by (auto simp add: measurable_def borel_def)
-lemma%important in_borel_measurable_borel:
+lemma in_borel_measurable_borel:
"f \<in> borel_measurable M \<longleftrightarrow>
(\<forall>S \<in> sets borel.
f -` S \<inter> space M \<in> sets M)"
- by%unimportant (auto simp add: measurable_def borel_def)
+ by (auto simp add: measurable_def borel_def)
-lemma%unimportant space_borel[simp]: "space borel = UNIV"
+lemma space_borel[simp]: "space borel = UNIV"
unfolding borel_def by auto
-lemma%unimportant space_in_borel[measurable]: "UNIV \<in> sets borel"
+lemma space_in_borel[measurable]: "UNIV \<in> sets borel"
unfolding borel_def by auto
-lemma%unimportant sets_borel: "sets borel = sigma_sets UNIV {S. open S}"
+lemma sets_borel: "sets borel = sigma_sets UNIV {S. open S}"
unfolding borel_def by (rule sets_measure_of) simp
-lemma%unimportant measurable_sets_borel:
+lemma 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%unimportant pred_Collect_borel[measurable (raw)]: "Measurable.pred borel P \<Longrightarrow> {x. P x} \<in> sets borel"
+lemma pred_Collect_borel[measurable (raw)]: "Measurable.pred borel P \<Longrightarrow> {x. P x} \<in> sets borel"
unfolding borel_def pred_def by auto
-lemma%unimportant borel_open[measurable (raw generic)]:
+lemma 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%unimportant borel_closed[measurable (raw generic)]:
+lemma borel_closed[measurable (raw generic)]:
assumes "closed A" shows "A \<in> sets borel"
proof -
have "space borel - (- A) \<in> sets borel"
@@ -343,11 +343,11 @@
thus ?thesis by simp
qed
-lemma%unimportant borel_singleton[measurable]:
+lemma 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%unimportant sets_borel_eq_count_space: "sets (borel :: 'a::{countable, t2_space} measure) = count_space UNIV"
+lemma 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
@@ -355,16 +355,16 @@
by auto
qed
-lemma%unimportant borel_comp[measurable]: "A \<in> sets borel \<Longrightarrow> - A \<in> sets borel"
+lemma borel_comp[measurable]: "A \<in> sets borel \<Longrightarrow> - A \<in> sets borel"
unfolding Compl_eq_Diff_UNIV by simp
-lemma%unimportant borel_measurable_vimage:
+lemma 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%unimportant borel_measurableI:
+lemma 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"
@@ -374,30 +374,30 @@
using assms[of S] by simp
qed
-lemma%unimportant borel_measurable_const:
+lemma borel_measurable_const:
"(\<lambda>x. c) \<in> borel_measurable M"
by auto
-lemma%unimportant borel_measurable_indicator:
+lemma 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%unimportant borel_measurable_count_space[measurable (raw)]:
+lemma borel_measurable_count_space[measurable (raw)]:
"f \<in> borel_measurable (count_space S)"
unfolding measurable_def by auto
-lemma%unimportant borel_measurable_indicator'[measurable (raw)]:
+lemma 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%important borel_measurable_indicator_iff:
+lemma 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%unimportant
+proof
assume "?I \<in> borel_measurable M"
then have "?I -` {1} \<inter> space M \<in> sets M"
unfolding measurable_def by auto
@@ -412,12 +412,12 @@
ultimately show "?I \<in> borel_measurable M" by auto
qed
-lemma%unimportant borel_measurable_subalgebra:
+lemma 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%unimportant borel_measurable_restrict_space_iff_ereal:
+lemma 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>
@@ -425,7 +425,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%unimportant borel_measurable_restrict_space_iff_ennreal:
+lemma 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>
@@ -433,7 +433,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%unimportant borel_measurable_restrict_space_iff:
+lemma 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>
@@ -442,27 +442,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%unimportant cbox_borel[measurable]: "cbox a b \<in> sets borel"
+lemma cbox_borel[measurable]: "cbox a b \<in> sets borel"
by (auto intro: borel_closed)
-lemma%unimportant box_borel[measurable]: "box a b \<in> sets borel"
+lemma box_borel[measurable]: "box a b \<in> sets borel"
by (auto intro: borel_open)
-lemma%unimportant borel_compact: "compact (A::'a::t2_space set) \<Longrightarrow> A \<in> sets borel"
+lemma borel_compact: "compact (A::'a::t2_space set) \<Longrightarrow> A \<in> sets borel"
by (auto intro: borel_closed dest!: compact_imp_closed)
-lemma%unimportant borel_sigma_sets_subset:
+lemma 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%important borel_eq_sigmaI1:
+lemma 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%unimportant (intro sigma_eqI antisym)
+proof (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"
@@ -474,7 +474,7 @@
unfolding borel_rev_eq using F by (intro borel_sigma_sets_subset) auto
qed auto
-lemma%unimportant borel_eq_sigmaI2:
+lemma 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)"
@@ -484,7 +484,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%unimportant borel_eq_sigmaI3:
+lemma 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))"
@@ -492,7 +492,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%unimportant borel_eq_sigmaI4:
+lemma 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)"
@@ -501,7 +501,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%unimportant borel_eq_sigmaI5:
+lemma 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)))"
@@ -509,12 +509,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%important second_countable_borel_measurable:
+lemma 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%unimportant (intro sigma_eqI sigma_sets_eqI)
+proof (intro sigma_eqI sigma_sets_eqI)
interpret X: sigma_algebra UNIV "sigma_sets UNIV X"
by (rule sigma_algebra_sigma_sets) simp
@@ -553,7 +553,7 @@
qed auto
qed (auto simp: eq intro: generate_topology.Basis)
-lemma%unimportant borel_eq_closed: "borel = sigma UNIV (Collect closed)"
+lemma 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"
@@ -569,13 +569,13 @@
finally show "x \<in> sigma_sets UNIV (Collect open)" by simp
qed simp_all
-lemma%important borel_eq_countable_basis:
+lemma 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%unimportant (intro sigma_eqI sigma_sets_eqI, safe)
+proof (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
@@ -587,7 +587,7 @@
thus "b \<in> sigma_sets UNIV (Collect open)" by auto
qed simp_all
-lemma%unimportant borel_measurable_continuous_on_restrict:
+lemma 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)"
@@ -599,16 +599,16 @@
by (force simp add: sets_restrict_space space_restrict_space)
qed
-lemma%unimportant borel_measurable_continuous_on1: "continuous_on UNIV f \<Longrightarrow> f \<in> borel_measurable borel"
+lemma borel_measurable_continuous_on1: "continuous_on UNIV f \<Longrightarrow> f \<in> borel_measurable borel"
by (drule borel_measurable_continuous_on_restrict) simp
-lemma%unimportant borel_measurable_continuous_on_if:
+lemma 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%unimportant borel_measurable_continuous_countable_exceptions:
+lemma borel_measurable_continuous_countable_exceptions:
fixes f :: "'a::t1_space \<Rightarrow> 'b::topological_space"
assumes X: "countable X"
assumes "continuous_on (- X) f"
@@ -620,23 +620,23 @@
by (intro borel_measurable_continuous_on_if assms continuous_intros)
qed auto
-lemma%unimportant borel_measurable_continuous_on:
+lemma 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%unimportant borel_measurable_continuous_on_indicator:
+lemma 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%important borel_measurable_Pair[measurable (raw)]:
+lemma 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%unimportant (subst borel_eq_countable_basis)
+proof (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)"
@@ -657,13 +657,13 @@
qed auto
qed
-lemma%important borel_measurable_continuous_Pair:
+lemma 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%unimportant -
+proof -
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
@@ -671,7 +671,7 @@
subsection%important \<open>Borel spaces on order topologies\<close>
-lemma%unimportant [measurable]:
+lemma [measurable]:
fixes a b :: "'a::linorder_topology"
shows lessThan_borel: "{..< a} \<in> sets borel"
and greaterThan_borel: "{a <..} \<in> sets borel"
@@ -685,7 +685,7 @@
by (blast intro: borel_open borel_closed open_lessThan open_greaterThan open_greaterThanLessThan
closed_atMost closed_atLeast closed_atLeastAtMost)+
-lemma%unimportant borel_Iio:
+lemma 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)
@@ -723,7 +723,7 @@
qed auto
qed auto
-lemma%unimportant borel_Ioi:
+lemma 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)
@@ -761,29 +761,29 @@
qed auto
qed auto
-lemma%unimportant borel_measurableI_less:
+lemma 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%important borel_measurableI_greater:
+lemma 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%unimportant (rule measurable_measure_of) (auto simp: Int_def conj_commute)
+ by (rule measurable_measure_of) (auto simp: Int_def conj_commute)
-lemma%unimportant borel_measurableI_le:
+lemma 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%unimportant borel_measurableI_ge:
+lemma 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%unimportant borel_measurable_less[measurable]:
+lemma 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"
@@ -797,7 +797,7 @@
finally show ?thesis .
qed
-lemma%important
+lemma (* FIX ME needs name *)
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"
@@ -805,23 +805,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%unimportant measurable
+ by measurable
-lemma%important borel_measurable_SUP[measurable (raw)]:
+lemma 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\<in>I. F i x) \<in> borel_measurable M"
- by%unimportant (rule borel_measurableI_greater) (simp add: less_SUP_iff)
+ by (rule borel_measurableI_greater) (simp add: less_SUP_iff)
-lemma%unimportant borel_measurable_INF[measurable (raw)]:
+lemma 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\<in>I. F i x) \<in> borel_measurable M"
by (rule borel_measurableI_less) (simp add: INF_less_iff)
-lemma%unimportant borel_measurable_cSUP[measurable (raw)]:
+lemma 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"
@@ -843,13 +843,13 @@
qed
qed
-lemma%important borel_measurable_cINF[measurable (raw)]:
+lemma 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\<in>I. F i x) \<in> borel_measurable M"
-proof%unimportant cases
+proof cases
assume "I = {}" then show ?thesis
unfolding \<open>I = {}\<close> image_empty by simp
next
@@ -865,7 +865,7 @@
qed
qed
-lemma%unimportant borel_measurable_lfp[consumes 1, case_names continuity step]:
+lemma 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"
@@ -882,7 +882,7 @@
finally show ?thesis .
qed
-lemma%unimportant borel_measurable_gfp[consumes 1, case_names continuity step]:
+lemma 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"
@@ -899,56 +899,56 @@
finally show ?thesis .
qed
-lemma%unimportant borel_measurable_max[measurable (raw)]:
+lemma 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%unimportant borel_measurable_min[measurable (raw)]:
+lemma 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%unimportant borel_measurable_Min[measurable (raw)]:
+lemma 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%unimportant borel_measurable_Max[measurable (raw)]:
+lemma 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%unimportant borel_measurable_sup[measurable (raw)]:
+lemma 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%unimportant borel_measurable_inf[measurable (raw)]:
+lemma 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%unimportant [measurable (raw)]:
+lemma [measurable (raw)]: (*FIXME needs name *)
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%unimportant measurable_convergent[measurable (raw)]:
+lemma 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%unimportant sets_Collect_convergent[measurable]:
+lemma 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%unimportant borel_measurable_lim[measurable (raw)]:
+lemma 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"
@@ -959,7 +959,7 @@
by simp
qed
-lemma%unimportant borel_measurable_LIMSEQ_order:
+lemma 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"
@@ -972,14 +972,14 @@
subsection%important \<open>Borel spaces on topological monoids\<close>
-lemma%unimportant borel_measurable_add[measurable (raw)]:
+lemma 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%unimportant borel_measurable_sum[measurable (raw)]:
+lemma 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"
@@ -988,7 +988,7 @@
thus ?thesis using assms by induct auto
qed simp
-lemma%important borel_measurable_suminf_order[measurable (raw)]:
+lemma 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"
@@ -996,22 +996,22 @@
subsection%important \<open>Borel spaces on Euclidean spaces\<close>
-lemma%important borel_measurable_inner[measurable (raw)]:
+lemma 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%unimportant (rule borel_measurable_continuous_Pair) (intro continuous_intros)
+ by (rule borel_measurable_continuous_Pair) (intro continuous_intros)
notation
eucl_less (infix "<e" 50)
-lemma%important box_oc: "{x. a <e x \<and> x \<le> b} = {x. a <e x} \<inter> {..b}"
+lemma 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%important eucl_ivals[measurable]:
+lemma 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"
@@ -1023,7 +1023,7 @@
unfolding box_oc box_co
by (auto intro: borel_open borel_closed)
-lemma%unimportant (*FIX ME this has no name *)
+lemma (*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"
@@ -1031,7 +1031,7 @@
and hafspace_greater_eq_borel: "{x. x \<bullet> i \<le> a} \<in> sets borel"
by simp_all
-lemma%unimportant borel_eq_box:
+lemma 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])
@@ -1044,7 +1044,7 @@
done
qed (auto simp: box_def)
-lemma%unimportant halfspace_gt_in_halfspace:
+lemma 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))"
@@ -1070,7 +1070,7 @@
by (auto intro!: Diff sigma_sets_Inter i)
qed
-lemma%unimportant borel_eq_halfspace_less:
+lemma 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])
@@ -1083,7 +1083,7 @@
finally show "box a b \<in> sets ?SIGMA" .
qed auto
-lemma%unimportant borel_eq_halfspace_le:
+lemma 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])
@@ -1107,7 +1107,7 @@
by (intro sets.countable_UN) (auto intro: i)
qed auto
-lemma%unimportant borel_eq_halfspace_ge:
+lemma 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])
@@ -1117,10 +1117,10 @@
using i by (intro sets.compl_sets) auto
qed auto
-lemma%important borel_eq_halfspace_greater:
+lemma borel_eq_halfspace_greater:
"borel = sigma UNIV ((\<lambda> (a, i). {x::'a::euclidean_space. a < x \<bullet> i}) ` (UNIV \<times> Basis))"
(is "_ = ?SIGMA")
-proof%unimportant (rule borel_eq_sigmaI2[OF borel_eq_halfspace_le])
+proof (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
@@ -1128,7 +1128,7 @@
by (intro sets.compl_sets) (auto intro: i)
qed auto
-lemma%unimportant borel_eq_atMost:
+lemma 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])
@@ -1147,7 +1147,7 @@
by (intro sets.countable_UN) auto
qed auto
-lemma%unimportant borel_eq_greaterThan:
+lemma 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])
@@ -1174,7 +1174,7 @@
done
qed auto
-lemma%unimportant borel_eq_lessThan:
+lemma 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])
@@ -1200,10 +1200,10 @@
done
qed auto
-lemma%important borel_eq_atLeastAtMost:
+lemma borel_eq_atLeastAtMost:
"borel = sigma UNIV (range (\<lambda>(a,b). {a..b} ::'a::ordered_euclidean_space set))"
(is "_ = ?SIGMA")
-proof%unimportant (rule borel_eq_sigmaI5[OF borel_eq_atMost])
+proof (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])
@@ -1222,12 +1222,12 @@
(auto intro!: sigma_sets_top)
qed auto
-lemma%important borel_set_induct[consumes 1, case_names empty interval compl union]:
+lemma 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%unimportant -
+proof -
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)
@@ -1244,7 +1244,7 @@
qed (auto intro: empty compl simp: Compl_eq_Diff_UNIV[symmetric] borel_eq_atLeastAtMost)
qed
-lemma%unimportant borel_sigma_sets_Ioc: "borel = sigma UNIV (range (\<lambda>(a, b). {a <.. b::real}))"
+lemma 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})"
@@ -1254,10 +1254,10 @@
finally show "{..i} \<in> sets (sigma UNIV (range (\<lambda>(i, j). {i<..j})))" .
qed simp
-lemma%unimportant eucl_lessThan: "{x::real. x <e a} = lessThan a"
+lemma eucl_lessThan: "{x::real. x <e a} = lessThan a"
by (simp add: eucl_less_def lessThan_def)
-lemma%unimportant borel_eq_atLeastLessThan:
+lemma 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
@@ -1268,7 +1268,7 @@
by (auto intro: sigma_sets.intros(2-) simp: eucl_lessThan)
qed auto
-lemma%unimportant borel_measurable_halfspacesI:
+lemma 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"
@@ -1283,47 +1283,47 @@
by (auto intro!: measurable_measure_of simp: S_eq F)
qed
-lemma%unimportant borel_measurable_iff_halfspace_le:
+lemma 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%unimportant borel_measurable_iff_halfspace_less:
+lemma 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%unimportant borel_measurable_iff_halfspace_ge:
+lemma 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%unimportant borel_measurable_iff_halfspace_greater:
+lemma 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%unimportant borel_measurable_iff_le:
+lemma 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%unimportant borel_measurable_iff_less:
+lemma 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%unimportant borel_measurable_iff_ge:
+lemma 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%unimportant borel_measurable_iff_greater:
+lemma 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%important borel_measurable_euclidean_space:
+lemma 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%unimportant safe
+proof 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
@@ -1331,64 +1331,64 @@
subsection%important "Borel measurable operators"
-lemma%important borel_measurable_norm[measurable]: "norm \<in> borel_measurable borel"
- by%unimportant (intro borel_measurable_continuous_on1 continuous_intros)
+lemma borel_measurable_norm[measurable]: "norm \<in> borel_measurable borel"
+ by (intro borel_measurable_continuous_on1 continuous_intros)
-lemma%unimportant borel_measurable_sgn [measurable]: "(sgn::'a::real_normed_vector \<Rightarrow> 'a) \<in> borel_measurable borel"
+lemma 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%important borel_measurable_uminus[measurable (raw)]:
+lemma 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%unimportant (rule borel_measurable_continuous_on[OF _ g]) (intro continuous_intros)
+ by (rule borel_measurable_continuous_on[OF _ g]) (intro continuous_intros)
-lemma%important borel_measurable_diff[measurable (raw)]:
+lemma 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%unimportant borel_measurable_add [of f M "- g"] assms by (simp add: fun_Compl_def)
+ using borel_measurable_add [of f M "- g"] assms by (simp add: fun_Compl_def)
-lemma%unimportant borel_measurable_times[measurable (raw)]:
+lemma 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%important borel_measurable_prod[measurable (raw)]:
+lemma 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%unimportant cases
+proof cases
assume "finite S"
thus ?thesis using assms by induct auto
qed simp
-lemma%important borel_measurable_dist[measurable (raw)]:
+lemma 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%unimportant f g by (rule borel_measurable_continuous_Pair) (intro continuous_intros)
+ using f g by (rule borel_measurable_continuous_Pair) (intro continuous_intros)
-lemma%unimportant borel_measurable_scaleR[measurable (raw)]:
+lemma 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%unimportant borel_measurable_uminus_eq [simp]:
+lemma 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%unimportant affine_borel_measurable_vector:
+lemma 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"
@@ -1409,15 +1409,15 @@
qed simp
qed
-lemma%unimportant borel_measurable_const_scaleR[measurable (raw)]:
+lemma 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%unimportant borel_measurable_const_add[measurable (raw)]:
+lemma 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%unimportant borel_measurable_inverse[measurable (raw)]:
+lemma 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"
@@ -1426,27 +1426,27 @@
apply (auto intro!: continuous_on_inverse continuous_on_id)
done
-lemma%unimportant borel_measurable_divide[measurable (raw)]:
+lemma 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%unimportant borel_measurable_abs[measurable (raw)]:
+lemma 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%unimportant borel_measurable_nth[measurable (raw)]:
+lemma borel_measurable_nth[measurable (raw)]:
"(\<lambda>x::real^'n. x $ i) \<in> borel_measurable borel"
by (simp add: cart_eq_inner_axis)
-lemma%important convex_measurable:
+lemma 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%unimportant (rule measurable_compose[where f=X and N="restrict_space borel A"])
+ by (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%unimportant borel_measurable_ln[measurable (raw)]:
+lemma 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])
@@ -1454,15 +1454,15 @@
apply (auto intro!: continuous_on_ln continuous_on_id)
done
-lemma%unimportant borel_measurable_log[measurable (raw)]:
+lemma 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%unimportant borel_measurable_exp[measurable]:
+lemma 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%unimportant measurable_real_floor[measurable]:
+lemma 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))"
@@ -1471,41 +1471,41 @@
by (auto simp: vimage_def measurable_count_space_eq2_countable)
qed
-lemma%unimportant measurable_real_ceiling[measurable]:
+lemma measurable_real_ceiling[measurable]:
"(ceiling :: real \<Rightarrow> int) \<in> measurable borel (count_space UNIV)"
unfolding ceiling_def[abs_def] by simp
-lemma%unimportant borel_measurable_real_floor: "(\<lambda>x::real. real_of_int \<lfloor>x\<rfloor>) \<in> borel_measurable borel"
+lemma borel_measurable_real_floor: "(\<lambda>x::real. real_of_int \<lfloor>x\<rfloor>) \<in> borel_measurable borel"
by simp
-lemma%unimportant borel_measurable_root [measurable]: "root n \<in> borel_measurable borel"
+lemma borel_measurable_root [measurable]: "root n \<in> borel_measurable borel"
by (intro borel_measurable_continuous_on1 continuous_intros)
-lemma%unimportant borel_measurable_sqrt [measurable]: "sqrt \<in> borel_measurable borel"
+lemma borel_measurable_sqrt [measurable]: "sqrt \<in> borel_measurable borel"
by (intro borel_measurable_continuous_on1 continuous_intros)
-lemma%unimportant borel_measurable_power [measurable (raw)]:
+lemma 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%unimportant borel_measurable_Re [measurable]: "Re \<in> borel_measurable borel"
+lemma borel_measurable_Re [measurable]: "Re \<in> borel_measurable borel"
by (intro borel_measurable_continuous_on1 continuous_intros)
-lemma%unimportant borel_measurable_Im [measurable]: "Im \<in> borel_measurable borel"
+lemma borel_measurable_Im [measurable]: "Im \<in> borel_measurable borel"
by (intro borel_measurable_continuous_on1 continuous_intros)
-lemma%unimportant borel_measurable_of_real [measurable]: "(of_real :: _ \<Rightarrow> (_::real_normed_algebra)) \<in> borel_measurable borel"
+lemma borel_measurable_of_real [measurable]: "(of_real :: _ \<Rightarrow> (_::real_normed_algebra)) \<in> borel_measurable borel"
by (intro borel_measurable_continuous_on1 continuous_intros)
-lemma%unimportant borel_measurable_sin [measurable]: "(sin :: _ \<Rightarrow> (_::{real_normed_field,banach})) \<in> borel_measurable borel"
+lemma borel_measurable_sin [measurable]: "(sin :: _ \<Rightarrow> (_::{real_normed_field,banach})) \<in> borel_measurable borel"
by (intro borel_measurable_continuous_on1 continuous_intros)
-lemma%unimportant borel_measurable_cos [measurable]: "(cos :: _ \<Rightarrow> (_::{real_normed_field,banach})) \<in> borel_measurable borel"
+lemma borel_measurable_cos [measurable]: "(cos :: _ \<Rightarrow> (_::{real_normed_field,banach})) \<in> borel_measurable borel"
by (intro borel_measurable_continuous_on1 continuous_intros)
-lemma%unimportant borel_measurable_arctan [measurable]: "arctan \<in> borel_measurable borel"
+lemma borel_measurable_arctan [measurable]: "arctan \<in> borel_measurable borel"
by (intro borel_measurable_continuous_on1 continuous_intros)
lemma%important borel_measurable_complex_iff:
@@ -1517,21 +1517,21 @@
apply auto
done
-lemma%important powr_real_measurable [measurable]:
+lemma 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%unimportant assms by (simp_all add: powr_def)
+ using assms by (simp_all add: powr_def)
-lemma%unimportant measurable_of_bool[measurable]: "of_bool \<in> count_space UNIV \<rightarrow>\<^sub>M borel"
+lemma measurable_of_bool[measurable]: "of_bool \<in> count_space UNIV \<rightarrow>\<^sub>M borel"
by simp
subsection%important "Borel space on the extended reals"
-lemma%unimportant borel_measurable_ereal[measurable (raw)]:
+lemma 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%unimportant borel_measurable_real_of_ereal[measurable (raw)]:
+lemma 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"
@@ -1540,7 +1540,7 @@
apply (auto intro: continuous_on_real simp: Compl_eq_Diff_UNIV)
done
-lemma%unimportant borel_measurable_ereal_cases:
+lemma 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"
@@ -1551,20 +1551,20 @@
with f H show ?thesis by simp
qed
-lemma%unimportant (*FIX ME needs a name *)
+lemma (*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%unimportant borel_measurable_uminus_eq_ereal[simp]:
+lemma 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%important set_Collect_ereal2:
+lemma set_Collect_ereal2:
fixes f g :: "'a \<Rightarrow> ereal"
assumes f: "f \<in> borel_measurable M"
assumes g: "g \<in> borel_measurable M"
@@ -1574,7 +1574,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%unimportant -
+proof -
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 }
@@ -1583,7 +1583,7 @@
by (subst *) (simp del: space_borel split del: if_split)
qed
-lemma%unimportant borel_measurable_ereal_iff:
+lemma 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"
@@ -1591,15 +1591,15 @@
show "f \<in> borel_measurable M" by auto
qed auto
-lemma%unimportant borel_measurable_erealD[measurable_dest]:
+lemma 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%important borel_measurable_ereal_iff_real:
+theorem 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%unimportant safe
+proof 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
@@ -1609,15 +1609,15 @@
finally show "f \<in> borel_measurable M" .
qed simp_all
-lemma%unimportant borel_measurable_ereal_iff_Iio:
+lemma 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%unimportant borel_measurable_ereal_iff_Ioi:
+lemma 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%unimportant vimage_sets_compl_iff:
+lemma 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"
@@ -1627,15 +1627,15 @@
by (metis double_complement)
qed
-lemma%unimportant borel_measurable_iff_Iic_ereal:
+lemma 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%unimportant borel_measurable_iff_Ici_ereal:
+lemma 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%important borel_measurable_ereal2:
+lemma borel_measurable_ereal2:
fixes f g :: "'a \<Rightarrow> ereal"
assumes f: "f \<in> borel_measurable M"
assumes g: "g \<in> borel_measurable M"
@@ -1645,7 +1645,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%unimportant -
+proof -
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 }
@@ -1653,14 +1653,14 @@
from assms show ?thesis unfolding * by simp
qed
-lemma%unimportant [measurable(raw)]:
+lemma [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%unimportant [measurable(raw)]:
+lemma [measurable(raw)]:
fixes f g :: "'a \<Rightarrow> ereal"
assumes "f \<in> borel_measurable M"
assumes "g \<in> borel_measurable M"
@@ -1668,19 +1668,19 @@
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%unimportant borel_measurable_ereal_sum[measurable (raw)]:
+lemma 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%unimportant borel_measurable_ereal_prod[measurable (raw)]:
+lemma 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%unimportant borel_measurable_extreal_suminf[measurable (raw)]:
+lemma 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"
@@ -1691,19 +1691,19 @@
text \<open> \<^type>\<open>ennreal\<close> is a topological monoid, so no rules for plus are required, also all order
statements are usually done on type classes. \<close>
-lemma%unimportant measurable_enn2ereal[measurable]: "enn2ereal \<in> borel \<rightarrow>\<^sub>M borel"
+lemma measurable_enn2ereal[measurable]: "enn2ereal \<in> borel \<rightarrow>\<^sub>M borel"
by (intro borel_measurable_continuous_on1 continuous_on_enn2ereal)
-lemma%unimportant measurable_e2ennreal[measurable]: "e2ennreal \<in> borel \<rightarrow>\<^sub>M borel"
+lemma measurable_e2ennreal[measurable]: "e2ennreal \<in> borel \<rightarrow>\<^sub>M borel"
by (intro borel_measurable_continuous_on1 continuous_on_e2ennreal)
-lemma%unimportant borel_measurable_enn2real[measurable (raw)]:
+lemma 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%important [simp]: "is_borel f M \<longleftrightarrow> f \<in> borel_measurable M"
-lemma%unimportant is_borel_transfer[transfer_rule]: "rel_fun (rel_fun (=) pcr_ennreal) (=) is_borel is_borel"
+lemma 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"
@@ -1715,14 +1715,14 @@
includes ennreal.lifting
begin
-lemma%unimportant measurable_ennreal[measurable]: "ennreal \<in> borel \<rightarrow>\<^sub>M borel"
+lemma measurable_ennreal[measurable]: "ennreal \<in> borel \<rightarrow>\<^sub>M borel"
unfolding is_borel_def[symmetric]
by transfer simp
-lemma%important borel_measurable_ennreal_iff[simp]:
+lemma 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%unimportant safe
+proof 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
@@ -1730,31 +1730,31 @@
by (rule measurable_cong[THEN iffD1, rotated]) auto
qed measurable
-lemma%unimportant borel_measurable_times_ennreal[measurable (raw)]:
+lemma 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%unimportant borel_measurable_inverse_ennreal[measurable (raw)]:
+lemma 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%unimportant borel_measurable_divide_ennreal[measurable (raw)]:
+lemma 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%unimportant borel_measurable_minus_ennreal[measurable (raw)]:
+lemma 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%important borel_measurable_prod_ennreal[measurable (raw)]:
+lemma 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%unimportant assms by (induction S rule: infinite_finite_induct) auto
+ using assms by (induction S rule: infinite_finite_induct) auto
end
@@ -1762,12 +1762,12 @@
subsection%important \<open>LIMSEQ is borel measurable\<close>
-lemma%important borel_measurable_LIMSEQ_real:
+lemma 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%unimportant -
+proof -
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"
@@ -1775,13 +1775,13 @@
ultimately show ?thesis by (simp cong: measurable_cong add: borel_measurable_ereal_iff)
qed
-lemma%important borel_measurable_LIMSEQ_metric:
+lemma 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%unimportant (safe intro!: measurable_measure_of)
+proof (safe intro!: measurable_measure_of)
fix A :: "'b set" assume "closed A"
have [measurable]: "(\<lambda>x. infdist (g x) A) \<in> borel_measurable M"
@@ -1806,13 +1806,13 @@
qed simp
qed auto
-lemma%important sets_Collect_Cauchy[measurable]:
+lemma 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%unimportant borel_measurable_lim_metric[measurable (raw)]:
+lemma 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"
@@ -1834,17 +1834,17 @@
unfolding * by measurable
qed
-lemma%unimportant borel_measurable_suminf[measurable (raw)]:
+lemma 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%unimportant Collect_closed_imp_pred_borel: "closed {x. P x} \<Longrightarrow> Measurable.pred borel P"
+lemma 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%unimportant isCont_borel_pred[measurable]:
+lemma isCont_borel_pred[measurable]:
fixes f :: "'b::metric_space \<Rightarrow> 'a::metric_space"
shows "Measurable.pred borel (isCont f)"
proof (subst measurable_cong)
@@ -1895,21 +1895,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%unimportant isCont_borel:
+lemma isCont_borel:
fixes f :: "'b::metric_space \<Rightarrow> 'a::metric_space"
shows "{x. isCont f x} \<in> sets borel"
by simp
-lemma%important is_real_interval:
+lemma 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%important real_interval_borel_measurable:
+lemma real_interval_borel_measurable:
assumes "is_interval (S::real set)"
shows "S \<in> sets borel"
-proof%unimportant -
+proof -
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 ..
@@ -1921,7 +1921,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%important borel_measurable_mono_on_fnc:
+lemma 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)"
@@ -1932,18 +1932,18 @@
intro!: borel_measurable_continuous_on_restrict intro: continuous_within_subset)
done
-lemma%unimportant borel_measurable_piecewise_mono:
+lemma 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)
-lemma%unimportant borel_measurable_mono:
+lemma 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%unimportant measurable_bdd_below_real[measurable (raw)]:
+lemma 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))"
@@ -1954,12 +1954,12 @@
using countable_int by measurable
qed
-lemma%important borel_measurable_cINF_real[measurable (raw)]:
+lemma 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\<in>I. F i x) \<in> borel_measurable M"
-proof%unimportant (rule measurable_piecewise_restrict)
+proof (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
@@ -1979,7 +1979,7 @@
qed
qed
-lemma%unimportant borel_Ici: "borel = sigma UNIV (range (\<lambda>x::real. {x ..}))"
+lemma 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 ..}"
@@ -1988,7 +1988,7 @@
unfolding eq by (intro sets.compl_sets) auto
qed auto
-lemma%unimportant borel_measurable_pred_less[measurable (raw)]:
+lemma 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)
@@ -1996,19 +1996,19 @@
no_notation
eucl_less (infix "<e" 50)
-lemma%important borel_measurable_Max2[measurable (raw)]:
+lemma 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%unimportant (simp add: borel_measurable_Max[OF assms(1), where ?f=f and ?M=M] Setcompr_eq_image)
+ by (simp add: borel_measurable_Max[OF assms(1), where ?f=f and ?M=M] Setcompr_eq_image)
-lemma%unimportant measurable_compose_n [measurable (raw)]:
+lemma 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%unimportant measurable_real_imp_nat:
+lemma 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)"
@@ -2020,7 +2020,7 @@
then show ?thesis using measurable_count_space_eq2_countable by blast
qed
-lemma%unimportant measurable_equality_set [measurable]:
+lemma 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"
@@ -2035,7 +2035,7 @@
then show ?thesis unfolding A_def by simp
qed
-lemma%unimportant measurable_inequality_set [measurable]:
+lemma 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"
@@ -2063,7 +2063,7 @@
ultimately show "{x \<in> space M. f x > g x} \<in> sets M" using * by (metis (mono_tags, lifting) measurable_sets)
qed
-lemma%unimportant measurable_limit [measurable]:
+proposition 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)"
@@ -2103,11 +2103,11 @@
then show ?thesis by auto
qed
-lemma%important measurable_limit2 [measurable]:
+lemma 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%unimportant -
+proof -
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
@@ -2115,7 +2115,7 @@
then show ?thesis using measurable_limit by auto
qed
-lemma%unimportant measurable_P_restriction [measurable (raw)]:
+lemma 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 -
@@ -2124,7 +2124,7 @@
then show ?thesis by auto
qed
-lemma%unimportant measurable_sum_nat [measurable (raw)]:
+lemma 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)"
@@ -2134,7 +2134,7 @@
qed simp
-lemma%unimportant measurable_abs_powr [measurable]:
+lemma 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"
@@ -2142,7 +2142,7 @@
text \<open>The next one is a variation around \<open>measurable_restrict_space\<close>.\<close>
-lemma%unimportant measurable_restrict_space3:
+lemma 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)"
@@ -2154,12 +2154,12 @@
text \<open>The next one is a variation around \<open>measurable_piecewise_restrict\<close>.\<close>
-lemma%important measurable_piecewise_restrict2:
+lemma 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%unimportant (rule measurableI)
+proof (rule measurableI)
fix B assume [measurable]: "B \<in> sets N"
{
fix n::nat
--- a/src/HOL/Analysis/Caratheodory.thy Sun Jan 13 20:25:41 2019 +0100
+++ b/src/HOL/Analysis/Caratheodory.thy Mon Jan 14 11:59:19 2019 +0000
@@ -13,7 +13,7 @@
Originally from the Hurd/Coble measure theory development, translated by Lawrence Paulson.
\<close>
-lemma%unimportant suminf_ennreal_2dimen:
+lemma suminf_ennreal_2dimen:
fixes f:: "nat \<times> nat \<Rightarrow> ennreal"
assumes "\<And>m. g m = (\<Sum>n. f (m,n))"
shows "(\<Sum>i. f (prod_decode i)) = suminf g"
@@ -60,7 +60,7 @@
where
"lambda_system \<Omega> M f = {l \<in> M. \<forall>x \<in> M. f (l \<inter> x) + f ((\<Omega> - l) \<inter> x) = f x}"
-lemma%unimportant (in algebra) lambda_system_eq:
+lemma (in algebra) lambda_system_eq:
"lambda_system \<Omega> M f = {l \<in> M. \<forall>x \<in> M. f (x \<inter> l) + f (x - l) = f x}"
proof -
have [simp]: "\<And>l x. l \<in> M \<Longrightarrow> x \<in> M \<Longrightarrow> (\<Omega> - l) \<inter> x = x - l"
@@ -69,13 +69,13 @@
by (auto simp add: lambda_system_def) (metis Int_commute)+
qed
-lemma%unimportant (in algebra) lambda_system_empty: "positive M f \<Longrightarrow> {} \<in> lambda_system \<Omega> M f"
+lemma (in algebra) lambda_system_empty: "positive M f \<Longrightarrow> {} \<in> lambda_system \<Omega> M f"
by (auto simp add: positive_def lambda_system_eq)
-lemma%unimportant lambda_system_sets: "x \<in> lambda_system \<Omega> M f \<Longrightarrow> x \<in> M"
+lemma lambda_system_sets: "x \<in> lambda_system \<Omega> M f \<Longrightarrow> x \<in> M"
by (simp add: lambda_system_def)
-lemma%unimportant (in algebra) lambda_system_Compl:
+lemma (in algebra) lambda_system_Compl:
fixes f:: "'a set \<Rightarrow> ennreal"
assumes x: "x \<in> lambda_system \<Omega> M f"
shows "\<Omega> - x \<in> lambda_system \<Omega> M f"
@@ -88,7 +88,7 @@
by (force simp add: lambda_system_def ac_simps)
qed
-lemma%unimportant (in algebra) lambda_system_Int:
+lemma (in algebra) lambda_system_Int:
fixes f:: "'a set \<Rightarrow> ennreal"
assumes xl: "x \<in> lambda_system \<Omega> M f" and yl: "y \<in> lambda_system \<Omega> M f"
shows "x \<inter> y \<in> lambda_system \<Omega> M f"
@@ -122,7 +122,7 @@
qed
qed
-lemma%unimportant (in algebra) lambda_system_Un:
+lemma (in algebra) lambda_system_Un:
fixes f:: "'a set \<Rightarrow> ennreal"
assumes xl: "x \<in> lambda_system \<Omega> M f" and yl: "y \<in> lambda_system \<Omega> M f"
shows "x \<union> y \<in> lambda_system \<Omega> M f"
@@ -136,7 +136,7 @@
by (metis lambda_system_Compl lambda_system_Int xl yl)
qed
-lemma%unimportant (in algebra) lambda_system_algebra:
+lemma (in algebra) lambda_system_algebra:
"positive M f \<Longrightarrow> algebra \<Omega> (lambda_system \<Omega> M f)"
apply (auto simp add: algebra_iff_Un)
apply (metis lambda_system_sets set_mp sets_into_space)
@@ -145,7 +145,7 @@
apply (metis lambda_system_Un)
done
-lemma%unimportant (in algebra) lambda_system_strong_additive:
+lemma (in algebra) lambda_system_strong_additive:
assumes z: "z \<in> M" and disj: "x \<inter> y = {}"
and xl: "x \<in> lambda_system \<Omega> M f" and yl: "y \<in> lambda_system \<Omega> M f"
shows "f (z \<inter> (x \<union> y)) = f (z \<inter> x) + f (z \<inter> y)"
@@ -160,7 +160,7 @@
by (simp add: lambda_system_eq)
qed
-lemma%unimportant (in algebra) lambda_system_additive: "additive (lambda_system \<Omega> M f) f"
+lemma (in algebra) lambda_system_additive: "additive (lambda_system \<Omega> M f) f"
proof (auto simp add: additive_def)
fix x and y
assume disj: "x \<inter> y = {}"
@@ -171,13 +171,13 @@
by (simp add: Un)
qed
-lemma%unimportant lambda_system_increasing: "increasing M f \<Longrightarrow> increasing (lambda_system \<Omega> M f) f"
+lemma lambda_system_increasing: "increasing M f \<Longrightarrow> increasing (lambda_system \<Omega> M f) f"
by (simp add: increasing_def lambda_system_def)
-lemma%unimportant lambda_system_positive: "positive M f \<Longrightarrow> positive (lambda_system \<Omega> M f) f"
+lemma lambda_system_positive: "positive M f \<Longrightarrow> positive (lambda_system \<Omega> M f) f"
by (simp add: positive_def lambda_system_def)
-lemma%unimportant (in algebra) lambda_system_strong_sum:
+lemma (in algebra) lambda_system_strong_sum:
fixes A:: "nat \<Rightarrow> 'a set" and f :: "'a set \<Rightarrow> ennreal"
assumes f: "positive M f" and a: "a \<in> M"
and A: "range A \<subseteq> lambda_system \<Omega> M f"
@@ -199,12 +199,12 @@
by (simp add: atLeastLessThanSuc lambda_system_strong_additive [OF a 2 3 4])
qed
-lemma%important (in sigma_algebra) lambda_system_caratheodory:
+proposition (in sigma_algebra) lambda_system_caratheodory:
assumes oms: "outer_measure_space M f"
and A: "range A \<subseteq> lambda_system \<Omega> M f"
and disj: "disjoint_family A"
shows "(\<Union>i. A i) \<in> lambda_system \<Omega> M f \<and> (\<Sum>i. f (A i)) = f (\<Union>i. A i)"
-proof%unimportant -
+proof -
have pos: "positive M f" and inc: "increasing M f"
and csa: "countably_subadditive M f"
by (metis oms outer_measure_space_def)+
@@ -274,11 +274,11 @@
by (simp add: lambda_system_eq sums_iff U_eq U_in)
qed
-lemma%important (in sigma_algebra) caratheodory_lemma:
+proposition (in sigma_algebra) caratheodory_lemma:
assumes oms: "outer_measure_space M f"
defines "L \<equiv> lambda_system \<Omega> M f"
shows "measure_space \<Omega> L f"
-proof%unimportant -
+proof -
have pos: "positive M f"
by (metis oms outer_measure_space_def)
have alg: "algebra \<Omega> L"
@@ -301,7 +301,7 @@
"outer_measure M f X =
(INF A\<in>{A. range A \<subseteq> M \<and> disjoint_family A \<and> X \<subseteq> (\<Union>i. A i)}. \<Sum>i. f (A i))"
-lemma%unimportant (in ring_of_sets) outer_measure_agrees:
+lemma (in ring_of_sets) outer_measure_agrees:
assumes posf: "positive M f" and ca: "countably_additive M f" and s: "s \<in> M"
shows "outer_measure M f s = f s"
unfolding outer_measure_def
@@ -326,19 +326,19 @@
(auto simp: disjoint_family_on_def)
qed
-lemma%unimportant outer_measure_empty:
+lemma outer_measure_empty:
"positive M f \<Longrightarrow> {} \<in> M \<Longrightarrow> outer_measure M f {} = 0"
unfolding outer_measure_def
by (intro antisym INF_lower2[of "\<lambda>_. {}"]) (auto simp: disjoint_family_on_def positive_def)
-lemma%unimportant (in ring_of_sets) positive_outer_measure:
+lemma (in ring_of_sets) positive_outer_measure:
assumes "positive M f" shows "positive (Pow \<Omega>) (outer_measure M f)"
unfolding positive_def by (auto simp: assms outer_measure_empty)
-lemma%unimportant (in ring_of_sets) increasing_outer_measure: "increasing (Pow \<Omega>) (outer_measure M f)"
+lemma (in ring_of_sets) increasing_outer_measure: "increasing (Pow \<Omega>) (outer_measure M f)"
by (force simp: increasing_def outer_measure_def intro!: INF_greatest intro: INF_lower)
-lemma%unimportant (in ring_of_sets) outer_measure_le:
+lemma (in ring_of_sets) outer_measure_le:
assumes pos: "positive M f" and inc: "increasing M f" and A: "range A \<subseteq> M" and X: "X \<subseteq> (\<Union>i. A i)"
shows "outer_measure M f X \<le> (\<Sum>i. f (A i))"
unfolding outer_measure_def
@@ -351,11 +351,11 @@
by (blast intro!: suminf_le)
qed (auto simp: X UN_disjointed_eq disjoint_family_disjointed)
-lemma%unimportant (in ring_of_sets) outer_measure_close:
+lemma (in ring_of_sets) outer_measure_close:
"outer_measure M f X < e \<Longrightarrow> \<exists>A. range A \<subseteq> M \<and> disjoint_family A \<and> X \<subseteq> (\<Union>i. A i) \<and> (\<Sum>i. f (A i)) < e"
unfolding outer_measure_def INF_less_iff by auto
-lemma%unimportant (in ring_of_sets) countably_subadditive_outer_measure:
+lemma (in ring_of_sets) countably_subadditive_outer_measure:
assumes posf: "positive M f" and inc: "increasing M f"
shows "countably_subadditive (Pow \<Omega>) (outer_measure M f)"
proof (simp add: countably_subadditive_def, safe)
@@ -398,12 +398,12 @@
qed
qed
-lemma%unimportant (in ring_of_sets) outer_measure_space_outer_measure:
+lemma (in ring_of_sets) outer_measure_space_outer_measure:
"positive M f \<Longrightarrow> increasing M f \<Longrightarrow> outer_measure_space (Pow \<Omega>) (outer_measure M f)"
by (simp add: outer_measure_space_def
positive_outer_measure increasing_outer_measure countably_subadditive_outer_measure)
-lemma%unimportant (in ring_of_sets) algebra_subset_lambda_system:
+lemma (in ring_of_sets) algebra_subset_lambda_system:
assumes posf: "positive M f" and inc: "increasing M f"
and add: "additive M f"
shows "M \<subseteq> lambda_system \<Omega> (Pow \<Omega>) (outer_measure M f)"
@@ -457,15 +457,15 @@
by (rule order_antisym)
qed
-lemma%unimportant measure_down: "measure_space \<Omega> N \<mu> \<Longrightarrow> sigma_algebra \<Omega> M \<Longrightarrow> M \<subseteq> N \<Longrightarrow> measure_space \<Omega> M \<mu>"
+lemma measure_down: "measure_space \<Omega> N \<mu> \<Longrightarrow> sigma_algebra \<Omega> M \<Longrightarrow> M \<subseteq> N \<Longrightarrow> measure_space \<Omega> M \<mu>"
by (auto simp add: measure_space_def positive_def countably_additive_def subset_eq)
subsection%important \<open>Caratheodory's theorem\<close>
-theorem%important (in ring_of_sets) caratheodory':
+theorem (in ring_of_sets) caratheodory':
assumes posf: "positive M f" and ca: "countably_additive M f"
shows "\<exists>\<mu> :: 'a set \<Rightarrow> ennreal. (\<forall>s \<in> M. \<mu> s = f s) \<and> measure_space \<Omega> (sigma_sets \<Omega> M) \<mu>"
-proof%unimportant -
+proof -
have inc: "increasing M f"
by (metis additive_increasing ca countably_additive_additive posf)
let ?O = "outer_measure M f"
@@ -489,11 +489,11 @@
by (intro exI[of _ ?O]) auto
qed
-lemma%important (in ring_of_sets) caratheodory_empty_continuous:
+lemma (in ring_of_sets) caratheodory_empty_continuous:
assumes f: "positive M f" "additive M f" and fin: "\<And>A. A \<in> M \<Longrightarrow> f A \<noteq> \<infinity>"
assumes cont: "\<And>A. range A \<subseteq> M \<Longrightarrow> decseq A \<Longrightarrow> (\<Inter>i. A i) = {} \<Longrightarrow> (\<lambda>i. f (A i)) \<longlonglongrightarrow> 0"
shows "\<exists>\<mu> :: 'a set \<Rightarrow> ennreal. (\<forall>s \<in> M. \<mu> s = f s) \<and> measure_space \<Omega> (sigma_sets \<Omega> M) \<mu>"
-proof%unimportant (intro caratheodory' empty_continuous_imp_countably_additive f)
+proof (intro caratheodory' empty_continuous_imp_countably_additive f)
show "\<forall>A\<in>M. f A \<noteq> \<infinity>" using fin by auto
qed (rule cont)
@@ -504,22 +504,22 @@
(f {} = 0) \<and> (\<forall>a\<in>M. 0 \<le> f a) \<and>
(\<forall>C\<subseteq>M. disjoint C \<longrightarrow> finite C \<longrightarrow> \<Union>C \<in> M \<longrightarrow> f (\<Union>C) = (\<Sum>c\<in>C. f c))"
-lemma%unimportant volumeI:
+lemma volumeI:
assumes "f {} = 0"
assumes "\<And>a. a \<in> M \<Longrightarrow> 0 \<le> f a"
assumes "\<And>C. C \<subseteq> M \<Longrightarrow> disjoint C \<Longrightarrow> finite C \<Longrightarrow> \<Union>C \<in> M \<Longrightarrow> f (\<Union>C) = (\<Sum>c\<in>C. f c)"
shows "volume M f"
using assms by (auto simp: volume_def)
-lemma%unimportant volume_positive:
+lemma volume_positive:
"volume M f \<Longrightarrow> a \<in> M \<Longrightarrow> 0 \<le> f a"
by (auto simp: volume_def)
-lemma%unimportant volume_empty:
+lemma volume_empty:
"volume M f \<Longrightarrow> f {} = 0"
by (auto simp: volume_def)
-lemma%unimportant volume_finite_additive:
+proposition volume_finite_additive:
assumes "volume M f"
assumes A: "\<And>i. i \<in> I \<Longrightarrow> A i \<in> M" "disjoint_family_on A I" "finite I" "\<Union>(A ` I) \<in> M"
shows "f (\<Union>(A ` I)) = (\<Sum>i\<in>I. f (A i))"
@@ -540,7 +540,7 @@
by simp
qed
-lemma%unimportant (in ring_of_sets) volume_additiveI:
+lemma (in ring_of_sets) volume_additiveI:
assumes pos: "\<And>a. a \<in> M \<Longrightarrow> 0 \<le> \<mu> a"
assumes [simp]: "\<mu> {} = 0"
assumes add: "\<And>a b. a \<in> M \<Longrightarrow> b \<in> M \<Longrightarrow> a \<inter> b = {} \<Longrightarrow> \<mu> (a \<union> b) = \<mu> a + \<mu> b"
@@ -557,10 +557,10 @@
qed simp
qed fact+
-lemma%important (in semiring_of_sets) extend_volume:
+proposition (in semiring_of_sets) extend_volume:
assumes "volume M \<mu>"
shows "\<exists>\<mu>'. volume generated_ring \<mu>' \<and> (\<forall>a\<in>M. \<mu>' a = \<mu> a)"
-proof%unimportant -
+proof -
let ?R = generated_ring
have "\<forall>a\<in>?R. \<exists>m. \<exists>C\<subseteq>M. a = \<Union>C \<and> finite C \<and> disjoint C \<and> m = (\<Sum>c\<in>C. \<mu> c)"
by (auto simp: generated_ring_def)
@@ -637,10 +637,10 @@
subsubsection%important \<open>Caratheodory on semirings\<close>
-theorem%important (in semiring_of_sets) caratheodory:
+theorem (in semiring_of_sets) caratheodory:
assumes pos: "positive M \<mu>" and ca: "countably_additive M \<mu>"
shows "\<exists>\<mu>' :: 'a set \<Rightarrow> ennreal. (\<forall>s \<in> M. \<mu>' s = \<mu> s) \<and> measure_space \<Omega> (sigma_sets \<Omega> M) \<mu>'"
-proof%unimportant -
+proof -
have "volume M \<mu>"
proof (rule volumeI)
{ fix a assume "a \<in> M" then show "0 \<le> \<mu> a"
@@ -816,7 +816,7 @@
by (intro exI[of _ \<mu>']) (auto intro: generated_ringI_Basic)
qed
-lemma%important extend_measure_caratheodory:
+lemma extend_measure_caratheodory:
fixes G :: "'i \<Rightarrow> 'a set"
assumes M: "M = extend_measure \<Omega> I G \<mu>"
assumes "i \<in> I"
@@ -828,7 +828,7 @@
(\<Union>i. G (A i)) = G j \<Longrightarrow> (\<Sum>n. \<mu> (A n)) = \<mu> j"
shows "emeasure M (G i) = \<mu> i"
-proof%unimportant -
+proof -
interpret semiring_of_sets \<Omega> "G ` I"
by fact
have "\<forall>g\<in>G`I. \<exists>i\<in>I. g = G i"
@@ -861,7 +861,7 @@
qed fact
qed
-lemma%important extend_measure_caratheodory_pair:
+proposition extend_measure_caratheodory_pair:
fixes G :: "'i \<Rightarrow> 'j \<Rightarrow> 'a set"
assumes M: "M = extend_measure \<Omega> {(a, b). P a b} (\<lambda>(a, b). G a b) (\<lambda>(a, b). \<mu> a b)"
assumes "P i j"
@@ -873,7 +873,7 @@
(\<And>n. P (A n) (B n)) \<Longrightarrow> P j k \<Longrightarrow> disjoint_family (\<lambda>n. G (A n) (B n)) \<Longrightarrow>
(\<Union>i. G (A i) (B i)) = G j k \<Longrightarrow> (\<Sum>n. \<mu> (A n) (B n)) = \<mu> j k"
shows "emeasure M (G i j) = \<mu> i j"
-proof%unimportant -
+proof -
have "emeasure M ((\<lambda>(a, b). G a b) (i, j)) = (\<lambda>(a, b). \<mu> a b) (i, j)"
proof (rule extend_measure_caratheodory[OF M])
show "semiring_of_sets \<Omega> ((\<lambda>(a, b). G a b) ` {(a, b). P a b})"