--- a/src/HOL/Analysis/Arcwise_Connected.thy Mon Aug 27 22:58:36 2018 +0200
+++ b/src/HOL/Analysis/Arcwise_Connected.thy Tue Aug 28 13:28:39 2018 +0100
@@ -9,14 +9,14 @@
begin
-subsection\<open>The Brouwer reduction theorem\<close>
+subsection%important \<open>The Brouwer reduction theorem\<close>
-theorem Brouwer_reduction_theorem_gen:
+theorem%important 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 -
+proof%unimportant -
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)
@@ -79,13 +79,13 @@
qed
qed
-corollary Brouwer_reduction_theorem:
+corollary%important 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 (rule Brouwer_reduction_theorem_gen [of S "\<lambda>T. T \<noteq> {} \<and> T \<subseteq> S \<and> \<phi> T"])
+proof%unimportant (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"
@@ -107,14 +107,14 @@
qed (meson assms compact_imp_closed seq_compact_closed_subset seq_compact_eq_compact)+
-subsection\<open>Arcwise Connections\<close>
+subsection%important\<open>Arcwise Connections\<close>(*FIX ME this subsection is empty(?) *)
-subsection\<open>Density of points with dyadic rational coordinates\<close>
+subsection%important\<open>Density of points with dyadic rational coordinates\<close>
-proposition closure_dyadic_rationals:
+proposition%important 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 -
+proof%unimportant -
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
@@ -152,9 +152,9 @@
then show ?thesis by auto
qed
-corollary closure_rational_coordinates:
+corollary%important closure_rational_coordinates:
"closure (\<Union>f \<in> Basis \<rightarrow> \<rat>. { \<Sum>i :: 'a :: euclidean_space \<in> Basis. f i *\<^sub>R i }) = UNIV"
-proof -
+proof%unimportant -
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
@@ -168,7 +168,7 @@
using closure_dyadic_rationals closure_mono [OF *] by blast
qed
-lemma closure_dyadic_rationals_in_convex_set:
+lemma%unimportant 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>.
@@ -176,7 +176,7 @@
closure S"
by (simp add: closure_dyadic_rationals closure_convex_Int_superset)
-lemma closure_rationals_in_convex_set:
+lemma%unimportant 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"
@@ -187,11 +187,11 @@
path connection is equivalent to arcwise connection for distinct points.
The proof is based on Whyburn's "Topological Analysis".\<close>
-lemma closure_dyadic_rationals_in_convex_set_pos_1:
+lemma%important 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 -
+proof%unimportant -
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>
@@ -202,13 +202,13 @@
qed
-definition dyadics :: "'a::field_char_0 set" where "dyadics \<equiv> \<Union>k m. {of_nat m / 2^k}"
+definition%important dyadics :: "'a::field_char_0 set" where "dyadics \<equiv> \<Union>k m. {of_nat m / 2^k}"
-lemma real_in_dyadics [simp]: "real m \<in> dyadics"
+lemma%unimportant real_in_dyadics [simp]: "real m \<in> dyadics"
apply (simp add: dyadics_def)
by (metis divide_numeral_1 numeral_One power_0)
-lemma nat_neq_4k1: "of_nat m \<noteq> (4 * of_nat k + 1) / (2 * 2^n :: 'a::field_char_0)"
+lemma%unimportant 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)"
@@ -221,8 +221,8 @@
by simp
qed
-lemma nat_neq_4k3: "of_nat m \<noteq> (4 * of_nat k + 3) / (2 * 2^n :: 'a::field_char_0)"
-proof
+lemma%important nat_neq_4k3: "of_nat m \<noteq> (4 * of_nat k + 3) / (2 * 2^n :: 'a::field_char_0)"
+proof%unimportant
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)
@@ -234,10 +234,10 @@
by simp
qed
-lemma iff_4k:
+lemma%important 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 -
+proof%unimportant -
{ 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)
@@ -254,8 +254,8 @@
then show ?thesis by blast
qed
-lemma neq_4k1_k43: "(4 * real m + 1) / (2 * 2^n) \<noteq> (4 * real m' + 3) / (2 * 2 ^ n')"
-proof
+lemma%important neq_4k1_k43: "(4 * real m + 1) / (2 * 2^n) \<noteq> (4 * real m' + 3) / (2 * 2 ^ n')"
+proof%unimportant
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)
@@ -271,11 +271,11 @@
using even_Suc by presburger
qed
-lemma dyadic_413_cases:
+lemma%important 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 (cases "m>0")
+proof%unimportant (cases "m>0")
case False
then have "m=0" by simp
with that show ?thesis by auto
@@ -324,11 +324,11 @@
qed
-lemma dyadics_iff:
+lemma%important 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
+proof%unimportant
show "dyadics \<subseteq> ?rhs"
unfolding dyadics_def
apply clarify
@@ -355,14 +355,14 @@
apply (fastforce simp add: dyadics_iff Nats_def field_simps)+
done
-lemma dyadics_levels: "dyadics = (\<Union>K. \<Union>k<K. \<Union> m. {of_nat m / 2^k})"
+lemma%unimportant dyadics_levels: "dyadics = (\<Union>K. \<Union>k<K. \<Union> m. {of_nat m / 2^k})"
unfolding dyadics_def by auto
-lemma dyad_rec_level_termination:
+lemma%important dyad_rec_level_termination:
assumes "k < K"
shows "dyad_rec_dom(b, l, r, real m / 2^k)"
using assms
-proof (induction K arbitrary: k m)
+proof%unimportant (induction K arbitrary: k m)
case 0
then show ?case by auto
next
@@ -427,22 +427,22 @@
qed
-lemma dyad_rec_termination: "x \<in> dyadics \<Longrightarrow> dyad_rec_dom(b,l,r,x)"
+lemma%unimportant 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 dyad_rec_of_nat [simp]: "dyad_rec b l r (real m) = b m"
+lemma%unimportant dyad_rec_of_nat [simp]: "dyad_rec b l r (real m) = b m"
by (simp add: dyad_rec.psimps dyad_rec_termination)
-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))"
+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))"
apply (rule dyad_rec.psimps)
by (metis dyad_rec_level_termination lessI add.commute of_nat_Suc of_nat_mult of_nat_numeral)
-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))"
+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))"
apply (rule dyad_rec.psimps)
by (metis dyad_rec_level_termination lessI of_nat_add of_nat_mult of_nat_numeral)
-lemma dyad_rec_41_times2:
+lemma%unimportant 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 -
@@ -461,10 +461,10 @@
finally show ?thesis .
qed
-lemma dyad_rec_43_times2:
+lemma%important 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 -
+proof%unimportant -
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))"
@@ -480,22 +480,22 @@
finally show ?thesis .
qed
-definition dyad_rec2
+definition%important 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)"
-lemma leftrec_base: "leftrec u v lc rc (real m / 2) = u"
+lemma%unimportant leftrec_base: "leftrec u v lc rc (real m / 2) = u"
by (simp add: dyad_rec2_def)
-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)"
+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)"
apply (simp only: dyad_rec2_def dyad_rec_41_times2)
apply (simp add: case_prod_beta)
done
-lemma leftrec_43: "n > 0 \<Longrightarrow>
+lemma%unimportant 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)))"
@@ -503,10 +503,10 @@
apply (simp add: case_prod_beta)
done
-lemma rightrec_base: "rightrec u v lc rc (real m / 2) = v"
+lemma%unimportant rightrec_base: "rightrec u v lc rc (real m / 2) = v"
by (simp add: dyad_rec2_def)
-lemma rightrec_41: "n > 0 \<Longrightarrow>
+lemma%unimportant 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)))"
@@ -514,24 +514,24 @@
apply (simp add: case_prod_beta)
done
-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)"
+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)"
apply (simp only: dyad_rec2_def dyad_rec_43_times2)
apply (simp add: case_prod_beta)
done
-lemma dyadics_in_open_unit_interval:
+lemma%unimportant 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 homeomorphic_monotone_image_interval:
+theorem%important 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 -
+proof%unimportant -
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>
@@ -1664,11 +1664,11 @@
qed
-theorem path_contains_arc:
+theorem%important 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 -
+proof%unimportant -
have ucont_p: "uniformly_continuous_on {0..1} p"
using \<open>path p\<close> unfolding path_def
by (metis compact_Icc compact_uniformly_continuous)
@@ -1961,12 +1961,12 @@
qed
-corollary path_connected_arcwise:
+corollary%important 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 (intro iffI impI ballI)
+proof%unimportant (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"
@@ -1992,23 +1992,23 @@
qed
-corollary arc_connected_trans:
+corollary%important 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 (metis (no_types, hide_lams) arc_imp_path assms path_contains_arc path_image_join path_join pathfinish_join pathstart_join)
+ by%unimportant (metis (no_types, hide_lams) arc_imp_path assms path_contains_arc path_image_join path_join pathfinish_join pathstart_join)
-subsection\<open>Accessibility of frontier points\<close>
+subsection%important\<open>Accessibility of frontier points\<close>
-lemma dense_accessible_frontier_points:
+lemma%important 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 -
+proof%unimportant -
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"
@@ -2099,12 +2099,12 @@
qed
-lemma dense_accessible_frontier_points_connected:
+lemma%important 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 -
+proof%unimportant -
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"
@@ -2134,14 +2134,14 @@
using h \<open>pathfinish g \<in> V\<close> by auto
qed
-lemma dense_access_fp_aux:
+lemma%important 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 -
+proof%unimportant -
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
@@ -2181,14 +2181,14 @@
qed
qed
-lemma dense_accessible_frontier_point_pairs:
+lemma%important 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 -
+proof%unimportant -
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 Mon Aug 27 22:58:36 2018 +0200
+++ b/src/HOL/Analysis/Binary_Product_Measure.thy Tue Aug 28 13:28:39 2018 +0100
@@ -2,67 +2,67 @@
Author: Johannes Hölzl, TU München
*)
-section \<open>Binary product measures\<close>
+section%important \<open>Binary product measures\<close>
theory Binary_Product_Measure
imports Nonnegative_Lebesgue_Integration
begin
-lemma Pair_vimage_times[simp]: "Pair x -` (A \<times> B) = (if x \<in> A then B else {})"
+lemma%unimportant Pair_vimage_times[simp]: "Pair x -` (A \<times> B) = (if x \<in> A then B else {})"
by auto
-lemma rev_Pair_vimage_times[simp]: "(\<lambda>x. (x, y)) -` (A \<times> B) = (if y \<in> B then A else {})"
+lemma%unimportant rev_Pair_vimage_times[simp]: "(\<lambda>x. (x, y)) -` (A \<times> B) = (if y \<in> B then A else {})"
by auto
-subsection "Binary products"
+subsection%important "Binary products"
-definition pair_measure (infixr "\<Otimes>\<^sub>M" 80) where
+definition%important pair_measure (infixr "\<Otimes>\<^sub>M" 80) where
"A \<Otimes>\<^sub>M B = measure_of (space A \<times> space B)
{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 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 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 space_pair_measure:
+lemma%important 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 (rule space_measure_of)
+ by%unimportant (rule space_measure_of)
-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)}"
+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)}"
by (auto simp: space_pair_measure)
-lemma sets_pair_measure:
+lemma%unimportant 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 sets_pair_measure_cong[measurable_cong, cong]:
+lemma%unimportant 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 pair_measureI[intro, simp, measurable]:
+lemma%unimportant 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 sets_Pair: "{x} \<in> sets M1 \<Longrightarrow> {y} \<in> sets M2 \<Longrightarrow> {(x, y)} \<in> sets (M1 \<Otimes>\<^sub>M M2)"
+lemma%unimportant 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 measurable_pair_measureI:
+lemma%unimportant 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 measurable_split_replace[measurable (raw)]:
+lemma%unimportant 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 measurable_Pair[measurable (raw)]:
+lemma%important 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 (rule measurable_pair_measureI)
+proof%unimportant (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 measurable_fst[intro!, simp, measurable]: "fst \<in> measurable (M1 \<Otimes>\<^sub>M M2) M1"
+lemma%unimportant 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 measurable_snd[intro!, simp, measurable]: "snd \<in> measurable (M1 \<Otimes>\<^sub>M M2) M2"
+lemma%unimportant 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 measurable_Pair_compose_split[measurable_dest]:
+lemma%unimportant 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 measurable_Pair1_compose[measurable_dest]:
+lemma%unimportant 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 measurable_Pair2_compose[measurable_dest]:
+lemma%unimportant 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 measurable_pair:
+lemma%unimportant 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
+lemma%unimportant (*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
+lemma%unimportant (*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 sets_pair_in_sets:
+lemma%unimportant 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 sets_pair_eq_sets_fst_snd:
+lemma%important 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 -
+proof%unimportant -
{ 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 measurable_pair_iff:
+lemma%unimportant 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 measurable_split_conv:
+lemma%unimportant 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 measurable_pair_swap': "(\<lambda>(x,y). (y, x)) \<in> measurable (M1 \<Otimes>\<^sub>M M2) (M2 \<Otimes>\<^sub>M M1)"
+lemma%unimportant 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 measurable_pair_swap:
+lemma%unimportant 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 measurable_pair_swap_iff:
+lemma%unimportant 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 measurable_Pair1': "x \<in> space M1 \<Longrightarrow> Pair x \<in> measurable M2 (M1 \<Otimes>\<^sub>M M2)"
+lemma%unimportant measurable_Pair1': "x \<in> space M1 \<Longrightarrow> Pair x \<in> measurable M2 (M1 \<Otimes>\<^sub>M M2)"
by simp
-lemma sets_Pair1[measurable (raw)]:
+lemma%unimportant 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 measurable_Pair2': "y \<in> space M2 \<Longrightarrow> (\<lambda>x. (x, y)) \<in> measurable M1 (M1 \<Otimes>\<^sub>M M2)"
+lemma%unimportant 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 sets_Pair2: assumes A: "A \<in> sets (M1 \<Otimes>\<^sub>M M2)" shows "(\<lambda>x. (x, y)) -` A \<in> sets M1"
+lemma%unimportant 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 measurable_Pair2:
+lemma%unimportant 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 measurable_Pair1:
+lemma%unimportant 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 Int_stable_pair_measure_generator: "Int_stable {a \<times> b | a b. a \<in> sets A \<and> b \<in> sets B}"
+lemma%unimportant 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 (in finite_measure) finite_measure_cut_measurable:
+lemma%unimportant (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 (in sigma_finite_measure) measurable_emeasure_Pair:
+lemma%unimportant (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 (in sigma_finite_measure) measurable_emeasure[measurable (raw)]:
+lemma%unimportant (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 (in sigma_finite_measure) emeasure_pair_measure:
+lemma%unimportant (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 (in sigma_finite_measure) emeasure_pair_measure_alt:
+lemma%unimportant (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 (in sigma_finite_measure) emeasure_pair_measure_Times:
+lemma%important (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 -
+proof%unimportant -
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"
@@ -336,18 +336,18 @@
by (simp add: ac_simps)
qed
-subsection \<open>Binary products of $\sigma$-finite emeasure spaces\<close>
+subsection%important \<open>Binary products of $\sigma$-finite emeasure spaces\<close>
-locale pair_sigma_finite = M1?: sigma_finite_measure M1 + M2?: sigma_finite_measure M2
+locale%important pair_sigma_finite = M1?: sigma_finite_measure M1 + M2?: sigma_finite_measure M2
for M1 :: "'a measure" and M2 :: "'b measure"
-lemma (in pair_sigma_finite) measurable_emeasure_Pair1:
+lemma%unimportant (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 (in pair_sigma_finite) measurable_emeasure_Pair2:
+lemma%important (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 -
+proof%unimportant -
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 (in pair_sigma_finite) sigma_finite_up_in_pair_measure_generator:
+lemma%important (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 -
+proof%unimportant -
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 pair_sigma_finite \<subseteq> P?: sigma_finite_measure "M1 \<Otimes>\<^sub>M M2"
+sublocale%important 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 sigma_finite_pair_measure:
+lemma%unimportant 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 sets_pair_swap:
+lemma%unimportant 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 (in pair_sigma_finite) distr_pair_swap:
+lemma%important (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 -
+proof%unimportant -
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 (in pair_sigma_finite) emeasure_pair_measure_alt2:
+lemma%unimportant (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 -
+proof%unimportant -
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 (in pair_sigma_finite) AE_pair:
+lemma%unimportant (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 (in pair_sigma_finite) AE_pair_measure:
+lemma%important (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 (subst AE_iff_measurable[OF _ refl])
+proof%unimportant (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 (in pair_sigma_finite) AE_pair_iff:
+lemma%unimportant (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 (in pair_sigma_finite) AE_commute:
+lemma%unimportant (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 -
@@ -531,16 +531,16 @@
done
qed
-subsection "Fubinis theorem"
+subsection%important "Fubinis theorem"
-lemma measurable_compose_Pair1:
+lemma%unimportant 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 (in sigma_finite_measure) borel_measurable_nn_integral_fst:
+lemma%unimportant (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 induct
+using f proof%unimportant 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,10 +561,10 @@
nn_integral_monotone_convergence_SUP incseq_def le_fun_def
cong: measurable_cong)
-lemma (in sigma_finite_measure) nn_integral_fst:
+lemma%unimportant (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
+ using f proof induct
case (cong u v)
then have "?I u = ?I v"
by (intro nn_integral_cong) (auto simp: space_pair_measure)
@@ -575,14 +575,14 @@
borel_measurable_nn_integral_fst nn_integral_mono incseq_def le_fun_def
cong: nn_integral_cong)
-lemma (in sigma_finite_measure) borel_measurable_nn_integral[measurable (raw)]:
+lemma%unimportant (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 (in pair_sigma_finite) nn_integral_snd:
+lemma%important (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 -
+proof%unimportant -
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 (in pair_sigma_finite) Fubini:
+lemma%important (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 (in pair_sigma_finite) Fubini':
+lemma%important (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 \<open>Products on counting spaces, densities and distributions\<close>
+subsection%important \<open>Products on counting spaces, densities and distributions\<close>
-lemma sigma_prod:
+lemma%important 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 (rule measure_eqI)
+proof%unimportant (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 sigma_sets_pair_measure_generator_finite:
+lemma%unimportant 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 sets_pair_eq:
+lemma%important 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
+proof%unimportant
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 borel_prod:
+lemma%important borel_prod:
"(borel \<Otimes>\<^sub>M borel) = (borel :: ('a::second_countable_topology \<times> 'b::second_countable_topology) measure)"
(is "?P = ?B")
-proof -
+proof%unimportant -
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 pair_measure_count_space:
+lemma%important 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 (rule measure_eqI)
+proof%unimportant (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 emeasure_prod_count_space:
+lemma%unimportant 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 emeasure_prod_count_space_single[simp]: "emeasure (count_space UNIV \<Otimes>\<^sub>M count_space UNIV) {x} = 1"
+lemma%unimportant 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 emeasure_count_space_prod_eq:
+lemma%important 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 -
+proof%unimportant -
{ 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 nn_integral_count_space_prod_eq:
+lemma%unimportant 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 pair_measure_density:
+lemma%important 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 (rule measure_eqI)
+proof%unimportant (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 sigma_finite_measure_distr:
+lemma%unimportant 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 pair_measure_distr:
+lemma%unimportant 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 pair_measure_eqI:
+lemma%important 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 -
+proof%unimportant -
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 sets_pair_countable:
+lemma%important 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 auto
+proof%unimportant 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 pair_measure_countable:
+lemma%important pair_measure_countable:
assumes "countable S1" "countable S2"
shows "count_space S1 \<Otimes>\<^sub>M count_space S2 = count_space (S1 \<times> S2)"
-proof (rule pair_measure_eqI)
+proof%unimportant (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 nn_integral_fst_count_space:
+lemma%important 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(cases)
+proof%unimportant(cases)
assume *: "countable {xy. f xy \<noteq> 0}"
let ?A = "fst ` {xy. f xy \<noteq> 0}"
let ?B = "snd ` {xy. f xy \<noteq> 0}"
@@ -1088,20 +1088,20 @@
finally show ?thesis .
qed
-lemma measurable_pair_measure_countable1:
+lemma%unimportant 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"
using _ _ assms(1)
by(rule measurable_compose_countable'[where f="\<lambda>a b. f (a, snd b)" and g=fst and I=A, simplified])simp_all
-subsection \<open>Product of Borel spaces\<close>
+subsection%important \<open>Product of Borel spaces\<close>
-lemma borel_Times:
+lemma%important 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 -
+proof%unimportant -
have "A \<times> B = (A\<times>UNIV) \<inter> (UNIV \<times> B)"
by auto
moreover
@@ -1146,7 +1146,7 @@
by auto
qed
-lemma finite_measure_pair_measure:
+lemma%unimportant 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 Mon Aug 27 22:58:36 2018 +0200
+++ b/src/HOL/Analysis/Bochner_Integration.thy Tue Aug 28 13:28:39 2018 +0100
@@ -2,7 +2,7 @@
Author: Johannes Hölzl, TU München
*)
-section \<open>Bochner Integration for Vector-Valued Functions\<close>
+section%important \<open>Bochner Integration for Vector-Valued Functions\<close>
theory Bochner_Integration
imports Finite_Product_Measure
@@ -15,12 +15,12 @@
\<close>
-lemma borel_measurable_implies_sequence_metric:
+lemma%important 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 -
+proof%unimportant -
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
+lemma%unimportant
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 borel_measurable_induct_real[consumes 2, case_names set mult add seq]:
+lemma%unimportant 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 scaleR_cong_right:
+lemma%unimportant 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 simple_bochner_integrable_compose2:
+lemma%unimportant 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 simple_function_finite_support:
+lemma%unimportant 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 simple_bochner_integrableI_bounded:
+lemma%unimportant 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
@@ -309,16 +309,16 @@
then show "emeasure M {y \<in> space M. f y \<noteq> 0} \<noteq> \<infinity>" by simp
qed fact
-definition simple_bochner_integral :: "'a measure \<Rightarrow> ('a \<Rightarrow> 'b::real_vector) \<Rightarrow> 'b" where
+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 simple_bochner_integral_partition:
+lemma%important 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 -
+proof%unimportant -
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 simple_bochner_integral_linear:
+lemma%unimportant 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 simple_bochner_integral_bounded:
+lemma%important 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 -
+proof%unimportant -
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 has_bochner_integral_cong:
+lemma%unimportant 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_strong nn_integral_cong_strong)
-lemma has_bochner_integral_cong_AE:
+lemma%unimportant 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 borel_measurable_has_bochner_integral:
+lemma%unimportant borel_measurable_has_bochner_integral:
"has_bochner_integral M f x \<Longrightarrow> f \<in> borel_measurable M"
by (rule has_bochner_integral.cases)
-lemma borel_measurable_has_bochner_integral'[measurable_dest]:
+lemma%unimportant 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 has_bochner_integral_simple_bochner_integrable:
+lemma%unimportant 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 has_bochner_integral_real_indicator:
+lemma%unimportant 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 has_bochner_integral_add[intro]:
+lemma%unimportant 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 has_bochner_integral_bounded_linear:
+lemma%unimportant 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 has_bochner_integral_zero[intro]: "has_bochner_integral M (\<lambda>x. 0) 0"
+lemma%unimportant 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 has_bochner_integral_scaleR_left[intro]:
+lemma%unimportant 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 has_bochner_integral_scaleR_right[intro]:
+lemma%unimportant 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 has_bochner_integral_mult_left[intro]:
+lemma%unimportant 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 has_bochner_integral_mult_right[intro]:
+lemma%unimportant 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 has_bochner_integral_divide =
+lemmas%unimportant has_bochner_integral_divide =
has_bochner_integral_bounded_linear[OF bounded_linear_divide]
-lemma has_bochner_integral_divide_zero[intro]:
+lemma%unimportant 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 has_bochner_integral_inner_left[intro]:
+lemma%unimportant 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 has_bochner_integral_inner_right[intro]:
+lemma%unimportant 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 has_bochner_integral_minus =
+lemmas%unimportant has_bochner_integral_minus =
has_bochner_integral_bounded_linear[OF bounded_linear_minus[OF bounded_linear_ident]]
-lemmas has_bochner_integral_Re =
+lemmas%unimportant has_bochner_integral_Re =
has_bochner_integral_bounded_linear[OF bounded_linear_Re]
-lemmas has_bochner_integral_Im =
+lemmas%unimportant has_bochner_integral_Im =
has_bochner_integral_bounded_linear[OF bounded_linear_Im]
-lemmas has_bochner_integral_cnj =
+lemmas%unimportant has_bochner_integral_cnj =
has_bochner_integral_bounded_linear[OF bounded_linear_cnj]
-lemmas has_bochner_integral_of_real =
+lemmas%unimportant has_bochner_integral_of_real =
has_bochner_integral_bounded_linear[OF bounded_linear_of_real]
-lemmas has_bochner_integral_fst =
+lemmas%unimportant has_bochner_integral_fst =
has_bochner_integral_bounded_linear[OF bounded_linear_fst]
-lemmas has_bochner_integral_snd =
+lemmas%unimportant has_bochner_integral_snd =
has_bochner_integral_bounded_linear[OF bounded_linear_snd]
-lemma has_bochner_integral_indicator:
+lemma%unimportant 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 has_bochner_integral_diff:
+lemma%unimportant 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 has_bochner_integral_sum:
+lemma%unimportant 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 has_bochner_integral_implies_finite_norm:
+lemma%important has_bochner_integral_implies_finite_norm:
"has_bochner_integral M f x \<Longrightarrow> (\<integral>\<^sup>+x. norm (f x) \<partial>M) < \<infinity>"
-proof (elim has_bochner_integral.cases)
+proof%unimportant (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 has_bochner_integral_norm_bound:
+lemma%important 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
+using assms proof%unimportant
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 has_bochner_integral_eq:
+lemma%important has_bochner_integral_eq:
"has_bochner_integral M f x \<Longrightarrow> has_bochner_integral M f y \<Longrightarrow> x = y"
-proof (elim has_bochner_integral.cases)
+proof%unimportant (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 has_bochner_integralI_AE:
+lemma%unimportant 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 has_bochner_integral_eq_AE:
+lemma%unimportant 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 simple_bochner_integrable_restrict_space:
+lemma%unimportant 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 simple_bochner_integral_restrict_space:
+lemma%important 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 -
+proof%unimportant -
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)
@@ -895,7 +895,7 @@
end
-definition lebesgue_integral ("integral\<^sup>L") where
+definition%important lebesgue_integral ("integral\<^sup>L") where
"integral\<^sup>L M f = (if \<exists>x. has_bochner_integral M f x then THE x. has_bochner_integral M f x else 0)"
syntax
@@ -910,155 +910,155 @@
translations
"LINT x|M. f" == "CONST lebesgue_integral M (\<lambda>x. f)"
-lemma has_bochner_integral_integral_eq: "has_bochner_integral M f x \<Longrightarrow> integral\<^sup>L M f = x"
+lemma%unimportant 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 has_bochner_integral_integrable:
+lemma%unimportant 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 has_bochner_integral_iff:
+lemma%unimportant 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 simple_bochner_integrable_eq_integral:
+lemma%unimportant 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 not_integrable_integral_eq: "\<not> integrable M f \<Longrightarrow> integral\<^sup>L M f = 0"
+lemma%unimportant 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 integral_eq_cases:
+lemma%unimportant 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 borel_measurable_integrable[measurable_dest]: "integrable M f \<Longrightarrow> f \<in> borel_measurable M"
+lemma%unimportant borel_measurable_integrable[measurable_dest]: "integrable M f \<Longrightarrow> f \<in> borel_measurable M"
by (auto elim: integrable.cases has_bochner_integral.cases)
-lemma borel_measurable_integrable'[measurable_dest]:
+lemma%unimportant 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 integrable_cong:
+lemma%unimportant 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 integrable_cong_AE:
+lemma%unimportant 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 integrable_cong_AE_imp:
+lemma%unimportant 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 integral_cong:
+lemma%unimportant 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 integral_cong_AE:
+lemma%unimportant 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 integrable_add[simp, intro]: "integrable M f \<Longrightarrow> integrable M g \<Longrightarrow> integrable M (\<lambda>x. f x + g x)"
+lemma%unimportant 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 integrable_zero[simp, intro]: "integrable M (\<lambda>x. 0)"
+lemma%unimportant integrable_zero[simp, intro]: "integrable M (\<lambda>x. 0)"
by (metis has_bochner_integral_zero integrable.simps)
-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)"
+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)"
by (metis has_bochner_integral_sum integrable.simps)
-lemma integrable_indicator[simp, intro]: "A \<in> sets M \<Longrightarrow> emeasure M A < \<infinity> \<Longrightarrow>
+lemma%unimportant 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 integrable_real_indicator[simp, intro]: "A \<in> sets M \<Longrightarrow> emeasure M A < \<infinity> \<Longrightarrow>
+lemma%unimportant 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 integrable_diff[simp, intro]: "integrable M f \<Longrightarrow> integrable M g \<Longrightarrow> integrable M (\<lambda>x. f x - g x)"
+lemma%unimportant 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 integrable_bounded_linear: "bounded_linear T \<Longrightarrow> integrable M f \<Longrightarrow> integrable M (\<lambda>x. T (f x))"
+lemma%unimportant 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 integrable_scaleR_left[simp, intro]: "(c \<noteq> 0 \<Longrightarrow> integrable M f) \<Longrightarrow> integrable M (\<lambda>x. f x *\<^sub>R c)"
+lemma%unimportant 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 integrable_scaleR_right[simp, intro]: "(c \<noteq> 0 \<Longrightarrow> integrable M f) \<Longrightarrow> integrable M (\<lambda>x. c *\<^sub>R f x)"
+lemma%unimportant 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 integrable_mult_left[simp, intro]:
+lemma%unimportant 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 integrable_mult_right[simp, intro]:
+lemma%unimportant 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 integrable_divide_zero[simp, intro]:
+lemma%unimportant 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 integrable_inner_left[simp, intro]:
+lemma%unimportant 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 integrable_inner_right[simp, intro]:
+lemma%unimportant 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 integrable_minus[simp, intro] =
+lemmas%unimportant integrable_minus[simp, intro] =
integrable_bounded_linear[OF bounded_linear_minus[OF bounded_linear_ident]]
-lemmas integrable_divide[simp, intro] =
+lemmas%unimportant integrable_divide[simp, intro] =
integrable_bounded_linear[OF bounded_linear_divide]
-lemmas integrable_Re[simp, intro] =
+lemmas%unimportant integrable_Re[simp, intro] =
integrable_bounded_linear[OF bounded_linear_Re]
-lemmas integrable_Im[simp, intro] =
+lemmas%unimportant integrable_Im[simp, intro] =
integrable_bounded_linear[OF bounded_linear_Im]
-lemmas integrable_cnj[simp, intro] =
+lemmas%unimportant integrable_cnj[simp, intro] =
integrable_bounded_linear[OF bounded_linear_cnj]
-lemmas integrable_of_real[simp, intro] =
+lemmas%unimportant integrable_of_real[simp, intro] =
integrable_bounded_linear[OF bounded_linear_of_real]
-lemmas integrable_fst[simp, intro] =
+lemmas%unimportant integrable_fst[simp, intro] =
integrable_bounded_linear[OF bounded_linear_fst]
-lemmas integrable_snd[simp, intro] =
+lemmas%unimportant integrable_snd[simp, intro] =
integrable_bounded_linear[OF bounded_linear_snd]
-lemma integral_zero[simp]: "integral\<^sup>L M (\<lambda>x. 0) = 0"
+lemma%unimportant integral_zero[simp]: "integral\<^sup>L M (\<lambda>x. 0) = 0"
by (intro has_bochner_integral_integral_eq has_bochner_integral_zero)
-lemma integral_add[simp]: "integrable M f \<Longrightarrow> integrable M g \<Longrightarrow>
+lemma%unimportant 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 integral_diff[simp]: "integrable M f \<Longrightarrow> integrable M g \<Longrightarrow>
+lemma%unimportant 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 integral_sum: "(\<And>i. i \<in> I \<Longrightarrow> integrable M (f i)) \<Longrightarrow>
+lemma%unimportant 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 integral_sum'[simp]: "(\<And>i. i \<in> I =simp=> integrable M (f i)) \<Longrightarrow>
+lemma%unimportant 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 integral_bounded_linear: "bounded_linear T \<Longrightarrow> integrable M f \<Longrightarrow>
+lemma%unimportant 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 integral_bounded_linear':
+lemma%unimportant 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 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%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"
by (intro has_bochner_integral_integral_eq has_bochner_integral_integrable has_bochner_integral_scaleR_left)
-lemma integral_scaleR_right[simp]: "(\<integral> x. c *\<^sub>R f x \<partial>M) = c *\<^sub>R integral\<^sup>L M f"
+lemma%unimportant 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 integral_mult_left[simp]:
+lemma%unimportant 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 integral_mult_right[simp]:
+lemma%unimportant 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 integral_mult_left_zero[simp]:
+lemma%unimportant 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 integral_mult_right_zero[simp]:
+lemma%unimportant 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 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%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"
by (intro has_bochner_integral_integral_eq has_bochner_integral_integrable has_bochner_integral_inner_left)
-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"
+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"
by (intro has_bochner_integral_integral_eq has_bochner_integral_integrable has_bochner_integral_inner_right)
-lemma integral_divide_zero[simp]:
+lemma%unimportant 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 integral_minus[simp]: "integral\<^sup>L M (\<lambda>x. - f x) = - integral\<^sup>L M f"
+lemma%unimportant 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 integral_complex_of_real[simp]: "integral\<^sup>L M (\<lambda>x. complex_of_real (f x)) = of_real (integral\<^sup>L M f)"
+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)"
by (rule integral_bounded_linear'[OF bounded_linear_of_real bounded_linear_Re]) simp
-lemma integral_cnj[simp]: "integral\<^sup>L M (\<lambda>x. cnj (f x)) = cnj (integral\<^sup>L M f)"
+lemma%unimportant 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 integral_divide[simp] =
+lemmas%unimportant integral_divide[simp] =
integral_bounded_linear[OF bounded_linear_divide]
-lemmas integral_Re[simp] =
+lemmas%unimportant integral_Re[simp] =
integral_bounded_linear[OF bounded_linear_Re]
-lemmas integral_Im[simp] =
+lemmas%unimportant integral_Im[simp] =
integral_bounded_linear[OF bounded_linear_Im]
-lemmas integral_of_real[simp] =
+lemmas%unimportant integral_of_real[simp] =
integral_bounded_linear[OF bounded_linear_of_real]
-lemmas integral_fst[simp] =
+lemmas%unimportant integral_fst[simp] =
integral_bounded_linear[OF bounded_linear_fst]
-lemmas integral_snd[simp] =
+lemmas%unimportant integral_snd[simp] =
integral_bounded_linear[OF bounded_linear_snd]
-lemma integral_norm_bound_ennreal:
+lemma%unimportant 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 integrableI_sequence:
+lemma%important 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 -
+proof%unimportant -
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 nn_integral_dominated_convergence_norm:
+lemma%important 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 -
+proof%unimportant -
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 integrableI_bounded:
+lemma%important 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 -
+proof%unimportant -
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 integrableI_bounded_set:
+lemma%important 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 (rule integrableI_bounded)
+proof%unimportant (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 integrableI_bounded_set_indicator:
+lemma%unimportant 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 integrableI_nonneg:
+lemma%important 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 -
+proof%unimportant -
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 integrable_iff_bounded:
+lemma%important 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 integrable_bound:
+lemma%important 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 safe
+proof%unimportant 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 integrable_mult_indicator:
+lemma%unimportant 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 integrable_real_mult_indicator:
+lemma%unimportant 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 integrable_abs[simp, intro]:
+lemma%unimportant 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 integrable_norm[simp, intro]:
+lemma%unimportant 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 integrable_norm_cancel:
+lemma%unimportant 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 integrable_norm_iff:
+lemma%unimportant 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 integrable_abs_cancel:
+lemma%unimportant 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 integrable_abs_iff:
+lemma%unimportant 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 integrable_max[simp, intro]:
+lemma%unimportant 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 integrable_min[simp, intro]:
+lemma%unimportant 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 integral_minus_iff[simp]:
+lemma%unimportant 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 integrable_indicator_iff:
+lemma%unimportant 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 integral_indicator[simp]: "integral\<^sup>L M (indicator A) = measure M (A \<inter> space M)"
+lemma%unimportant 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 integrable_discrete_difference:
+lemma%unimportant 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 integral_discrete_difference:
+lemma%unimportant 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 has_bochner_integral_discrete_difference:
+lemma%unimportant 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
+lemma%important (*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 -
+proof%unimportant -
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 integral_dominated_convergence_at_top: "((\<lambda>t. integral\<^sup>L M (s t)) \<longlongrightarrow> integral\<^sup>L M f) at_top"
+lemma%unimportant 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 integrable_dominated_convergence_at_top: "integrable M f"
+lemma%unimportant 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 integrable_mult_left_iff:
+lemma%unimportant 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 integrableI_nn_integral_finite:
+lemma%unimportant 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 integral_nonneg_AE:
+lemma%unimportant 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 integral_nonneg[simp]:
+lemma%unimportant 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 nn_integral_eq_integral:
+lemma%important 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 -
+proof%unimportant -
{ 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 nn_integral_eq_integrable:
+lemma%unimportant 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
+lemma%unimportant (* 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 integral_norm_bound [simp]:
+lemma%important 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 (cases "integrable M f")
+proof%unimportant (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 integral_abs_bound [simp]:
+lemma%unimportant 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 integral_eq_nn_integral:
+lemma%unimportant 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 enn2real_nn_integral_eq_integral:
+lemma%unimportant 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 has_bochner_integral_nn_integral:
+lemma%unimportant 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 integrableI_simple_bochner_integrable:
+lemma%unimportant 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 integrable_induct[consumes 1, case_names base add lim, induct pred: integrable]:
+lemma%important 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 -
+proof%unimportant -
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 integral_eq_zero_AE:
+lemma%unimportant 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 integral_nonneg_eq_0_iff_AE:
+lemma%unimportant 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 integral_mono_AE:
+lemma%unimportant 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 integral_mono:
+lemma%unimportant 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 integral_mono_AE':
+lemma%important 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 (cases "integrable M g")
+proof%unimportant (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 integral_mono':
+lemma%important 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 (in finite_measure) integrable_measure:
+lemma%important (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 -
+proof%unimportant -
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 integrableI_real_bounded:
+lemma%unimportant 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 nn_integral_nonneg_infinite:
+lemma%unimportant 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 integral_real_bounded:
+lemma%unimportant 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 integrable_MIN:
+lemma%unimportant 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 integrable_MAX:
+lemma%unimportant 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 integral_Markov_inequality:
+lemma%important 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 -
+proof%unimportant -
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 integral_ineq_eq_0_then_AE:
+lemma%unimportant 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 not_AE_zero_int_E:
+lemma%unimportant 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 tendsto_L1_int:
+lemma%important 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 -
+proof%unimportant -
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 $L^1$, then
it admits a subsequence that converges almost everywhere.\<close>
-lemma tendsto_L1_AE_subseq:
+lemma%important 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 -
+proof%unimportant -
{
fix k
have "eventually (\<lambda>n. (\<integral>x. norm(u n x) \<partial>M) < (1/2)^k) sequentially"
@@ -2201,9 +2201,9 @@
then show ?thesis using \<open>strict_mono r\<close> by auto
qed
-subsection \<open>Restricted measure spaces\<close>
-
-lemma integrable_restrict_space:
+subsection%important \<open>Restricted measure spaces\<close>
+
+lemma%unimportant 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 integral_restrict_space:
+lemma%important 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 (rule integral_eq_cases)
+proof%unimportant (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 integral_empty:
+lemma%unimportant integral_empty:
assumes "space M = {}"
shows "integral\<^sup>L M f = 0"
proof -
@@ -2252,9 +2252,9 @@
thus ?thesis by simp
qed
-subsection \<open>Measure spaces with an associated density\<close>
-
-lemma integrable_density:
+subsection%important \<open>Measure spaces with an associated density\<close>
+
+lemma%unimportant 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 integral_density:
+lemma%important 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 (rule integral_eq_cases)
+proof%unimportant (rule integral_eq_cases)
assume "integrable (density M g) f"
then show ?thesis
proof induct
@@ -2325,38 +2325,38 @@
qed
qed (simp add: f g integrable_density)
-lemma
+lemma%unimportant (*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 has_bochner_integral_density:
+lemma%important 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"
by (simp add: has_bochner_integral_iff integrable_density integral_density)
-subsection \<open>Distributions\<close>
-
-lemma integrable_distr_eq:
+subsection%important \<open>Distributions\<close>
+
+lemma%unimportant 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 integrable_distr:
+lemma%unimportant 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 integral_distr:
+lemma%important 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 (rule integral_eq_cases)
+proof%unimportant (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 has_bochner_integral_distr:
+lemma%important 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 (simp add: has_bochner_integral_iff integrable_distr_eq integral_distr)
-
-subsection \<open>Lebesgue integration on @{const count_space}\<close>
-
-lemma integrable_count_space:
+ by%unimportant (simp add: has_bochner_integral_iff integrable_distr_eq integral_distr)
+
+subsection%important \<open>Lebesgue integration on @{const count_space}\<close>
+
+lemma%unimportant 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 measure_count_space[simp]:
+lemma%unimportant 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 lebesgue_integral_count_space_finite_support:
+lemma%important 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 -
+proof%unimportant -
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 lebesgue_integral_count_space_finite: "finite A \<Longrightarrow> (\<integral>x. f x \<partial>count_space A) = (\<Sum>a\<in>A. f a)"
+lemma%unimportant 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 integrable_count_space_nat_iff:
+lemma%unimportant 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 sums_integral_count_space_nat:
+lemma%unimportant 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 integral_count_space_nat:
+lemma%unimportant 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 integrable_bij_count_space:
+lemma%unimportant 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 integral_bij_count_space:
+lemma%important 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,49 +2492,49 @@
apply auto
done
-lemma has_bochner_integral_count_space_nat:
+lemma%important 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 \<open>Point measure\<close>
-
-lemma lebesgue_integral_point_measure_finite:
+subsection%important \<open>Point measure\<close>
+
+lemma%unimportant 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 integrable_point_measure_finite:
+lemma%important 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 point_measure_def
+ unfolding%unimportant 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)
apply (auto simp: AE_count_space integrable_count_space)
done
-subsection \<open>Lebesgue integration on @{const null_measure}\<close>
-
-lemma has_bochner_integral_null_measure_iff[iff]:
+subsection%important \<open>Lebesgue integration on @{const null_measure}\<close>
+
+lemma%unimportant 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 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"
+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"
by (cases "integrable (null_measure M) f")
(auto simp add: not_integrable_integral_eq has_bochner_integral_integral_eq)
-subsection \<open>Legacy lemmas for the real-valued Lebesgue integral\<close>
-
-lemma real_lebesgue_integral_def:
+subsection%important \<open>Legacy lemmas for the real-valued Lebesgue integral\<close>
+
+lemma%important 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 -
+proof%unimportant -
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 +2546,11 @@
finally show ?thesis .
qed
-lemma real_integrable_def:
+lemma%important 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 (safe del: notI)
+proof%unimportant (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 +2574,12 @@
finally show "(\<integral>\<^sup>+ x. norm (f x) \<partial>M) < \<infinity>" .
qed
-lemma integrableD[dest]:
+lemma%unimportant 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 integrableE:
+lemma%unimportant 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 +2588,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 integral_monotone_convergence_nonneg:
+lemma%important 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 +2597,7 @@
and u: "u \<in> borel_measurable M"
shows "integrable M u"
and "integral\<^sup>L M u = x"
-proof -
+proof%unimportant -
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 +2629,7 @@
by (auto simp: real_integrable_def real_lebesgue_integral_def u)
qed
-lemma
+lemma%unimportant (*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 +2660,7 @@
by (metis has_bochner_integral_integrable)
qed
-lemma integral_norm_eq_0_iff:
+lemma%unimportant 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 +2675,18 @@
by simp
qed
-lemma integral_0_iff:
+lemma%unimportant 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 (in finite_measure) integrable_const[intro!, simp]: "integrable M (\<lambda>x. a)"
+lemma%unimportant (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 lebesgue_integral_const[simp]:
+lemma%important 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 -
+proof%unimportant -
{ 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 +2704,7 @@
ultimately show ?thesis by blast
qed
-lemma (in finite_measure) integrable_const_bound:
+lemma%unimportant (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 +2713,19 @@
apply auto
done
-lemma (in finite_measure) integral_bounded_eq_bound_then_AE:
+lemma%unimportant (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 integral_indicator_finite_real:
+lemma%important 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 -
+proof%unimportant -
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 +2736,7 @@
finally show ?thesis .
qed
-lemma (in finite_measure) ennreal_integral_real:
+lemma%unimportant (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 +2747,7 @@
using ae by (intro nn_integral_cong_AE) (auto simp: le_less_trans[OF _ ennreal_less_top])
qed auto
-lemma (in finite_measure) integral_less_AE:
+lemma%unimportant (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 +2778,14 @@
ultimately show ?thesis by auto
qed
-lemma (in finite_measure) integral_less_AE_space:
+lemma%unimportant (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 tendsto_integral_at_top:
+lemma%unimportant 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 +2809,7 @@
qed auto
qed
-lemma
+lemma%unimportant (*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"
@@ -2840,32 +2840,32 @@
by (auto simp: _has_bochner_integral_iff)
qed
-subsection \<open>Product measure\<close>
-
-lemma (in sigma_finite_measure) borel_measurable_lebesgue_integrable[measurable (raw)]:
+subsection%important \<open>Product measure\<close>
+
+lemma%important (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 -
+proof%unimportant -
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 Collect_subset [simp]: "{x\<in>A. P x} \<subseteq> A" by auto
-
-lemma (in sigma_finite_measure) measurable_measure[measurable (raw)]:
+lemma%unimportant Collect_subset [simp]: "{x\<in>A. P x} \<subseteq> A" by auto
+
+lemma%unimportant (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 (in sigma_finite_measure) borel_measurable_lebesgue_integral[measurable (raw)]:
+lemma%important (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 -
+proof%unimportant -
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"
@@ -2936,7 +2936,7 @@
qed
qed
-lemma (in pair_sigma_finite) integrable_product_swap:
+lemma%unimportant (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))"
@@ -2948,7 +2948,7 @@
(simp add: distr_pair_swap[symmetric] assms)
qed
-lemma (in pair_sigma_finite) integrable_product_swap_iff:
+lemma%unimportant (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 -
@@ -2957,7 +2957,7 @@
show ?thesis by auto
qed
-lemma (in pair_sigma_finite) integral_product_swap:
+lemma%unimportant (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"
@@ -2967,13 +2967,13 @@
by (simp add: integral_distr[symmetric, OF measurable_pair_swap' f] distr_pair_swap[symmetric])
qed
-lemma (in pair_sigma_finite) Fubini_integrable:
+lemma%important (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 (rule integrableI_bounded)
+proof%unimportant (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)"
@@ -2994,7 +2994,7 @@
finally show "(\<integral>\<^sup>+ p. norm (f p) \<partial>(M1 \<Otimes>\<^sub>M M2)) < \<infinity>" .
qed fact
-lemma (in pair_sigma_finite) emeasure_pair_measure_finite:
+lemma%unimportant (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 -
@@ -3008,7 +3008,7 @@
ultimately show ?thesis by (auto simp: less_top)
qed
-lemma (in pair_sigma_finite) AE_integrable_fst':
+lemma%unimportant (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))"
@@ -3025,7 +3025,7 @@
(auto simp: integrable_iff_bounded measurable_compose[OF _ borel_measurable_integrable[OF f]] less_top)
qed
-lemma (in pair_sigma_finite) integrable_fst':
+lemma%unimportant (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)"
@@ -3042,11 +3042,11 @@
finally show "(\<integral>\<^sup>+ x. ennreal (norm (\<integral> y. f (x, y) \<partial>M2)) \<partial>M1) < \<infinity>" .
qed
-lemma (in pair_sigma_finite) integral_fst':
+lemma%important (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 induct
+using f proof%unimportant induct
case (base A c)
have A[measurable]: "A \<in> sets (M1 \<Otimes>\<^sub>M M2)" by fact
@@ -3155,7 +3155,7 @@
qed
qed
-lemma (in pair_sigma_finite)
+lemma%unimportant (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")
@@ -3163,7 +3163,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 (in pair_sigma_finite)
+lemma%unimportant (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")
@@ -3179,13 +3179,13 @@
using integral_product_swap[of "case_prod f"] by simp
qed
-lemma (in pair_sigma_finite) Fubini_integral:
+lemma%unimportant (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 (in product_sigma_finite) product_integral_singleton:
+lemma%unimportant (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])
@@ -3193,7 +3193,7 @@
apply simp_all
done
-lemma (in product_sigma_finite) product_integral_fold:
+lemma%unimportant (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"
@@ -3220,12 +3220,12 @@
done
qed
-lemma (in product_sigma_finite) product_integral_insert:
+lemma%important (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 -
+proof%unimportant -
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)"
@@ -3244,11 +3244,11 @@
finally show ?thesis .
qed
-lemma (in product_sigma_finite) product_integrable_prod:
+lemma%important (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 (unfold integrable_iff_bounded, intro conjI)
+proof%unimportant (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)"
@@ -3263,11 +3263,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 (in product_sigma_finite) product_integral_prod:
+lemma%important (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 induct
+using assms proof%unimportant induct
case empty
interpret finite_measure "Pi\<^sub>M {} M"
by rule (simp add: space_PiM)
@@ -3286,7 +3286,7 @@
by (simp add: * insert prod subset_insertI)
qed
-lemma integrable_subalgebra:
+lemma%unimportant 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"
@@ -3298,12 +3298,12 @@
using assms by (auto simp: integrable_iff_bounded nn_integral_subalgebra)
qed
-lemma integral_subalgebra:
+lemma%important 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 cases
+proof%unimportant cases
assume "integrable N f"
then show ?thesis
proof induct
--- a/src/HOL/Analysis/Borel_Space.thy Mon Aug 27 22:58:36 2018 +0200
+++ b/src/HOL/Analysis/Borel_Space.thy Tue Aug 28 13:28:39 2018 +0100
@@ -3,55 +3,55 @@
Author: Armin Heller, TU München
*)
-section \<open>Borel spaces\<close>
+section%important \<open>Borel spaces\<close>
theory Borel_Space
imports
Measurable Derivative Ordered_Euclidean_Space Extended_Real_Limits
begin
-lemma sets_Collect_eventually_sequentially[measurable]:
+lemma%unimportant sets_Collect_eventually_sequentially[measurable]:
"(\<And>i. {x\<in>space M. P x i} \<in> sets M) \<Longrightarrow> {x\<in>space M. eventually (P x) sequentially} \<in> sets M"
unfolding eventually_sequentially by simp
-lemma topological_basis_trivial: "topological_basis {A. open A}"
+lemma%unimportant topological_basis_trivial: "topological_basis {A. open A}"
by (auto simp: topological_basis_def)
-lemma open_prod_generated: "open = generate_topology {A \<times> B | A B. open A \<and> open B}"
-proof -
+lemma%important open_prod_generated: "open = generate_topology {A \<times> B | A B. open A \<and> open B}"
+proof%unimportant -
have "{A \<times> B :: ('a \<times> 'b) set | A B. open A \<and> open B} = ((\<lambda>(a, b). a \<times> b) ` ({A. open A} \<times> {A. open A}))"
by auto
then show ?thesis
by (auto intro: topological_basis_prod topological_basis_trivial topological_basis_imp_subbasis)
qed
-definition "mono_on f A \<equiv> \<forall>r s. r \<in> A \<and> s \<in> A \<and> r \<le> s \<longrightarrow> f r \<le> f s"
+definition%important "mono_on f A \<equiv> \<forall>r s. r \<in> A \<and> s \<in> A \<and> r \<le> s \<longrightarrow> f r \<le> f s"
-lemma mono_onI:
+lemma%unimportant mono_onI:
"(\<And>r s. r \<in> A \<Longrightarrow> s \<in> A \<Longrightarrow> r \<le> s \<Longrightarrow> f r \<le> f s) \<Longrightarrow> mono_on f A"
unfolding mono_on_def by simp
-lemma mono_onD:
+lemma%unimportant mono_onD:
"\<lbrakk>mono_on f A; r \<in> A; s \<in> A; r \<le> s\<rbrakk> \<Longrightarrow> f r \<le> f s"
unfolding mono_on_def by simp
-lemma mono_imp_mono_on: "mono f \<Longrightarrow> mono_on f A"
+lemma%unimportant mono_imp_mono_on: "mono f \<Longrightarrow> mono_on f A"
unfolding mono_def mono_on_def by auto
-lemma mono_on_subset: "mono_on f A \<Longrightarrow> B \<subseteq> A \<Longrightarrow> mono_on f B"
+lemma%unimportant mono_on_subset: "mono_on f A \<Longrightarrow> B \<subseteq> A \<Longrightarrow> mono_on f B"
unfolding mono_on_def by auto
-definition "strict_mono_on f A \<equiv> \<forall>r s. r \<in> A \<and> s \<in> A \<and> r < s \<longrightarrow> f r < f s"
+definition%important "strict_mono_on f A \<equiv> \<forall>r s. r \<in> A \<and> s \<in> A \<and> r < s \<longrightarrow> f r < f s"
-lemma strict_mono_onI:
+lemma%unimportant strict_mono_onI:
"(\<And>r s. r \<in> A \<Longrightarrow> s \<in> A \<Longrightarrow> r < s \<Longrightarrow> f r < f s) \<Longrightarrow> strict_mono_on f A"
unfolding strict_mono_on_def by simp
-lemma strict_mono_onD:
+lemma%unimportant strict_mono_onD:
"\<lbrakk>strict_mono_on f A; r \<in> A; s \<in> A; r < s\<rbrakk> \<Longrightarrow> f r < f s"
unfolding strict_mono_on_def by simp
-lemma mono_on_greaterD:
+lemma%unimportant mono_on_greaterD:
assumes "mono_on g A" "x \<in> A" "y \<in> A" "g x > (g (y::_::linorder) :: _ :: linorder)"
shows "x > y"
proof (rule ccontr)
@@ -61,7 +61,7 @@
with assms(4) show False by simp
qed
-lemma strict_mono_inv:
+lemma%unimportant strict_mono_inv:
fixes f :: "('a::linorder) \<Rightarrow> ('b::linorder)"
assumes "strict_mono f" and "surj f" and inv: "\<And>x. g (f x) = x"
shows "strict_mono g"
@@ -72,7 +72,7 @@
with inv show "g x < g y" by simp
qed
-lemma strict_mono_on_imp_inj_on:
+lemma%unimportant strict_mono_on_imp_inj_on:
assumes "strict_mono_on (f :: (_ :: linorder) \<Rightarrow> (_ :: preorder)) A"
shows "inj_on f A"
proof (rule inj_onI)
@@ -82,7 +82,7 @@
(auto dest: strict_mono_onD[OF assms, of x y] strict_mono_onD[OF assms, of y x])
qed
-lemma strict_mono_on_leD:
+lemma%unimportant strict_mono_on_leD:
assumes "strict_mono_on (f :: (_ :: linorder) \<Rightarrow> _ :: preorder) A" "x \<in> A" "y \<in> A" "x \<le> y"
shows "f x \<le> f y"
proof (insert le_less_linear[of y x], elim disjE)
@@ -91,17 +91,17 @@
thus ?thesis by (rule less_imp_le)
qed (insert assms, simp)
-lemma strict_mono_on_eqD:
+lemma%unimportant strict_mono_on_eqD:
fixes f :: "(_ :: linorder) \<Rightarrow> (_ :: preorder)"
assumes "strict_mono_on f A" "f x = f y" "x \<in> A" "y \<in> A"
shows "y = x"
using assms by (rule_tac linorder_cases[of x y]) (auto dest: strict_mono_onD)
-lemma mono_on_imp_deriv_nonneg:
+lemma%important mono_on_imp_deriv_nonneg:
assumes mono: "mono_on f A" and deriv: "(f has_real_derivative D) (at x)"
assumes "x \<in> interior A"
shows "D \<ge> 0"
-proof (rule tendsto_lowerbound)
+proof%unimportant (rule tendsto_lowerbound)
let ?A' = "(\<lambda>y. y - x) ` interior A"
from deriv show "((\<lambda>h. (f (x + h) - f x) / h) \<longlongrightarrow> D) (at 0)"
by (simp add: field_has_derivative_at has_field_derivative_def)
@@ -124,16 +124,16 @@
qed
qed simp
-lemma strict_mono_on_imp_mono_on:
+lemma%unimportant strict_mono_on_imp_mono_on:
"strict_mono_on (f :: (_ :: linorder) \<Rightarrow> _ :: preorder) A \<Longrightarrow> mono_on f A"
by (rule mono_onI, rule strict_mono_on_leD)
-lemma mono_on_ctble_discont:
+lemma%important mono_on_ctble_discont:
fixes f :: "real \<Rightarrow> real"
fixes A :: "real set"
assumes "mono_on f A"
shows "countable {a\<in>A. \<not> continuous (at a within A) f}"
-proof -
+proof%unimportant -
have mono: "\<And>x y. x \<in> A \<Longrightarrow> y \<in> A \<Longrightarrow> x \<le> y \<Longrightarrow> f x \<le> f y"
using \<open>mono_on f A\<close> by (simp add: mono_on_def)
have "\<forall>a \<in> {a\<in>A. \<not> continuous (at a within A) f}. \<exists>q :: nat \<times> rat.
@@ -196,12 +196,12 @@
by (rule countableI')
qed
-lemma mono_on_ctble_discont_open:
+lemma%important mono_on_ctble_discont_open:
fixes f :: "real \<Rightarrow> real"
fixes A :: "real set"
assumes "open A" "mono_on f A"
shows "countable {a\<in>A. \<not>isCont f a}"
-proof -
+proof%unimportant -
have "{a\<in>A. \<not>isCont f a} = {a\<in>A. \<not>(continuous (at a within A) f)}"
by (auto simp add: continuous_within_open [OF _ \<open>open A\<close>])
thus ?thesis
@@ -209,13 +209,13 @@
by (rule mono_on_ctble_discont, rule assms)
qed
-lemma mono_ctble_discont:
+lemma%important mono_ctble_discont:
fixes f :: "real \<Rightarrow> real"
assumes "mono f"
shows "countable {a. \<not> isCont f a}"
-using assms mono_on_ctble_discont [of f UNIV] unfolding mono_on_def mono_def by auto
+using%unimportant assms mono_on_ctble_discont [of f UNIV] unfolding mono_on_def mono_def by auto
-lemma has_real_derivative_imp_continuous_on:
+lemma%important has_real_derivative_imp_continuous_on:
assumes "\<And>x. x \<in> A \<Longrightarrow> (f has_real_derivative f' x) (at x)"
shows "continuous_on A f"
apply (intro differentiable_imp_continuous_on, unfold differentiable_on_def)
@@ -224,11 +224,11 @@
apply simp_all
done
-lemma closure_contains_Sup:
+lemma%important closure_contains_Sup:
fixes S :: "real set"
assumes "S \<noteq> {}" "bdd_above S"
shows "Sup S \<in> closure S"
-proof-
+proof%unimportant -
have "Inf (uminus ` S) \<in> closure (uminus ` S)"
using assms by (intro closure_contains_Inf) auto
also have "Inf (uminus ` S) = -Sup S" by (simp add: Inf_real_def)
@@ -237,22 +237,22 @@
finally show ?thesis by auto
qed
-lemma closed_contains_Sup:
+lemma%unimportant closed_contains_Sup:
fixes S :: "real set"
shows "S \<noteq> {} \<Longrightarrow> bdd_above S \<Longrightarrow> closed S \<Longrightarrow> Sup S \<in> S"
by (subst closure_closed[symmetric], assumption, rule closure_contains_Sup)
-lemma closed_subset_contains_Sup:
+lemma%unimportant closed_subset_contains_Sup:
fixes A C :: "real set"
shows "closed C \<Longrightarrow> A \<subseteq> C \<Longrightarrow> A \<noteq> {} \<Longrightarrow> bdd_above A \<Longrightarrow> Sup A \<in> C"
by (metis closure_contains_Sup closure_minimal subset_eq)
-lemma deriv_nonneg_imp_mono:
+lemma%important deriv_nonneg_imp_mono:
assumes deriv: "\<And>x. x \<in> {a..b} \<Longrightarrow> (g has_real_derivative g' x) (at x)"
assumes nonneg: "\<And>x. x \<in> {a..b} \<Longrightarrow> g' x \<ge> 0"
assumes ab: "a \<le> b"
shows "g a \<le> g b"
-proof (cases "a < b")
+proof%unimportant (cases "a < b")
assume "a < b"
from deriv have "\<And>x. \<lbrakk>x \<ge> a; x \<le> b\<rbrakk> \<Longrightarrow> (g has_real_derivative g' x) (at x)" by simp
with MVT2[OF \<open>a < b\<close>] and deriv
@@ -261,11 +261,11 @@
with g_ab show ?thesis by simp
qed (insert ab, simp)
-lemma continuous_interval_vimage_Int:
+lemma%important continuous_interval_vimage_Int:
assumes "continuous_on {a::real..b} g" and mono: "\<And>x y. a \<le> x \<Longrightarrow> x \<le> y \<Longrightarrow> y \<le> b \<Longrightarrow> g x \<le> g y"
assumes "a \<le> b" "(c::real) \<le> d" "{c..d} \<subseteq> {g a..g b}"
obtains c' d' where "{a..b} \<inter> g -` {c..d} = {c'..d'}" "c' \<le> d'" "g c' = c" "g d' = d"
-proof-
+proof%unimportant-
let ?A = "{a..b} \<inter> g -` {c..d}"
from IVT'[of g a c b, OF _ _ \<open>a \<le> b\<close> assms(1)] assms(4,5)
obtain c'' where c'': "c'' \<in> ?A" "g c'' = c" by auto
@@ -297,48 +297,48 @@
ultimately show ?thesis using that by blast
qed
-subsection \<open>Generic Borel spaces\<close>
+subsection%important \<open>Generic Borel spaces\<close>
-definition (in topological_space) borel :: "'a measure" where
+definition%important (in topological_space) borel :: "'a measure" where
"borel = sigma UNIV {S. open S}"
abbreviation "borel_measurable M \<equiv> measurable M borel"
-lemma in_borel_measurable:
+lemma%important in_borel_measurable:
"f \<in> borel_measurable M \<longleftrightarrow>
(\<forall>S \<in> sigma_sets UNIV {S. open S}. f -` S \<inter> space M \<in> sets M)"
- by (auto simp add: measurable_def borel_def)
+ by%unimportant (auto simp add: measurable_def borel_def)
-lemma in_borel_measurable_borel:
+lemma%important in_borel_measurable_borel:
"f \<in> borel_measurable M \<longleftrightarrow>
(\<forall>S \<in> sets borel.
f -` S \<inter> space M \<in> sets M)"
- by (auto simp add: measurable_def borel_def)
+ by%unimportant (auto simp add: measurable_def borel_def)
-lemma space_borel[simp]: "space borel = UNIV"
+lemma%unimportant space_borel[simp]: "space borel = UNIV"
unfolding borel_def by auto
-lemma space_in_borel[measurable]: "UNIV \<in> sets borel"
+lemma%unimportant space_in_borel[measurable]: "UNIV \<in> sets borel"
unfolding borel_def by auto
-lemma sets_borel: "sets borel = sigma_sets UNIV {S. open S}"
+lemma%unimportant sets_borel: "sets borel = sigma_sets UNIV {S. open S}"
unfolding borel_def by (rule sets_measure_of) simp
-lemma measurable_sets_borel:
+lemma%unimportant measurable_sets_borel:
"\<lbrakk>f \<in> measurable borel M; A \<in> sets M\<rbrakk> \<Longrightarrow> f -` A \<in> sets borel"
by (drule (1) measurable_sets) simp
-lemma pred_Collect_borel[measurable (raw)]: "Measurable.pred borel P \<Longrightarrow> {x. P x} \<in> sets borel"
+lemma%unimportant pred_Collect_borel[measurable (raw)]: "Measurable.pred borel P \<Longrightarrow> {x. P x} \<in> sets borel"
unfolding borel_def pred_def by auto
-lemma borel_open[measurable (raw generic)]:
+lemma%unimportant borel_open[measurable (raw generic)]:
assumes "open A" shows "A \<in> sets borel"
proof -
have "A \<in> {S. open S}" unfolding mem_Collect_eq using assms .
thus ?thesis unfolding borel_def by auto
qed
-lemma borel_closed[measurable (raw generic)]:
+lemma%unimportant borel_closed[measurable (raw generic)]:
assumes "closed A" shows "A \<in> sets borel"
proof -
have "space borel - (- A) \<in> sets borel"
@@ -346,11 +346,11 @@
thus ?thesis by simp
qed
-lemma borel_singleton[measurable]:
+lemma%unimportant borel_singleton[measurable]:
"A \<in> sets borel \<Longrightarrow> insert x A \<in> sets (borel :: 'a::t1_space measure)"
unfolding insert_def by (rule sets.Un) auto
-lemma sets_borel_eq_count_space: "sets (borel :: 'a::{countable, t2_space} measure) = count_space UNIV"
+lemma%unimportant sets_borel_eq_count_space: "sets (borel :: 'a::{countable, t2_space} measure) = count_space UNIV"
proof -
have "(\<Union>a\<in>A. {a}) \<in> sets borel" for A :: "'a set"
by (intro sets.countable_UN') auto
@@ -358,16 +358,16 @@
by auto
qed
-lemma borel_comp[measurable]: "A \<in> sets borel \<Longrightarrow> - A \<in> sets borel"
+lemma%unimportant borel_comp[measurable]: "A \<in> sets borel \<Longrightarrow> - A \<in> sets borel"
unfolding Compl_eq_Diff_UNIV by simp
-lemma borel_measurable_vimage:
+lemma%unimportant borel_measurable_vimage:
fixes f :: "'a \<Rightarrow> 'x::t2_space"
assumes borel[measurable]: "f \<in> borel_measurable M"
shows "f -` {x} \<inter> space M \<in> sets M"
by simp
-lemma borel_measurableI:
+lemma%unimportant borel_measurableI:
fixes f :: "'a \<Rightarrow> 'x::topological_space"
assumes "\<And>S. open S \<Longrightarrow> f -` S \<inter> space M \<in> sets M"
shows "f \<in> borel_measurable M"
@@ -377,30 +377,30 @@
using assms[of S] by simp
qed
-lemma borel_measurable_const:
+lemma%unimportant borel_measurable_const:
"(\<lambda>x. c) \<in> borel_measurable M"
by auto
-lemma borel_measurable_indicator:
+lemma%unimportant borel_measurable_indicator:
assumes A: "A \<in> sets M"
shows "indicator A \<in> borel_measurable M"
unfolding indicator_def [abs_def] using A
by (auto intro!: measurable_If_set)
-lemma borel_measurable_count_space[measurable (raw)]:
+lemma%unimportant borel_measurable_count_space[measurable (raw)]:
"f \<in> borel_measurable (count_space S)"
unfolding measurable_def by auto
-lemma borel_measurable_indicator'[measurable (raw)]:
+lemma%unimportant borel_measurable_indicator'[measurable (raw)]:
assumes [measurable]: "{x\<in>space M. f x \<in> A x} \<in> sets M"
shows "(\<lambda>x. indicator (A x) (f x)) \<in> borel_measurable M"
unfolding indicator_def[abs_def]
by (auto intro!: measurable_If)
-lemma borel_measurable_indicator_iff:
+lemma%important borel_measurable_indicator_iff:
"(indicator A :: 'a \<Rightarrow> 'x::{t1_space, zero_neq_one}) \<in> borel_measurable M \<longleftrightarrow> A \<inter> space M \<in> sets M"
(is "?I \<in> borel_measurable M \<longleftrightarrow> _")
-proof
+proof%unimportant
assume "?I \<in> borel_measurable M"
then have "?I -` {1} \<inter> space M \<in> sets M"
unfolding measurable_def by auto
@@ -415,12 +415,12 @@
ultimately show "?I \<in> borel_measurable M" by auto
qed
-lemma borel_measurable_subalgebra:
+lemma%unimportant borel_measurable_subalgebra:
assumes "sets N \<subseteq> sets M" "space N = space M" "f \<in> borel_measurable N"
shows "f \<in> borel_measurable M"
using assms unfolding measurable_def by auto
-lemma borel_measurable_restrict_space_iff_ereal:
+lemma%unimportant borel_measurable_restrict_space_iff_ereal:
fixes f :: "'a \<Rightarrow> ereal"
assumes \<Omega>[measurable, simp]: "\<Omega> \<inter> space M \<in> sets M"
shows "f \<in> borel_measurable (restrict_space M \<Omega>) \<longleftrightarrow>
@@ -428,7 +428,7 @@
by (subst measurable_restrict_space_iff)
(auto simp: indicator_def if_distrib[where f="\<lambda>x. a * x" for a] cong del: if_weak_cong)
-lemma borel_measurable_restrict_space_iff_ennreal:
+lemma%unimportant borel_measurable_restrict_space_iff_ennreal:
fixes f :: "'a \<Rightarrow> ennreal"
assumes \<Omega>[measurable, simp]: "\<Omega> \<inter> space M \<in> sets M"
shows "f \<in> borel_measurable (restrict_space M \<Omega>) \<longleftrightarrow>
@@ -436,7 +436,7 @@
by (subst measurable_restrict_space_iff)
(auto simp: indicator_def if_distrib[where f="\<lambda>x. a * x" for a] cong del: if_weak_cong)
-lemma borel_measurable_restrict_space_iff:
+lemma%unimportant borel_measurable_restrict_space_iff:
fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"
assumes \<Omega>[measurable, simp]: "\<Omega> \<inter> space M \<in> sets M"
shows "f \<in> borel_measurable (restrict_space M \<Omega>) \<longleftrightarrow>
@@ -445,27 +445,27 @@
(auto simp: indicator_def if_distrib[where f="\<lambda>x. x *\<^sub>R a" for a] ac_simps
cong del: if_weak_cong)
-lemma cbox_borel[measurable]: "cbox a b \<in> sets borel"
+lemma%unimportant cbox_borel[measurable]: "cbox a b \<in> sets borel"
by (auto intro: borel_closed)
-lemma box_borel[measurable]: "box a b \<in> sets borel"
+lemma%unimportant box_borel[measurable]: "box a b \<in> sets borel"
by (auto intro: borel_open)
-lemma borel_compact: "compact (A::'a::t2_space set) \<Longrightarrow> A \<in> sets borel"
+lemma%unimportant borel_compact: "compact (A::'a::t2_space set) \<Longrightarrow> A \<in> sets borel"
by (auto intro: borel_closed dest!: compact_imp_closed)
-lemma borel_sigma_sets_subset:
+lemma%unimportant borel_sigma_sets_subset:
"A \<subseteq> sets borel \<Longrightarrow> sigma_sets UNIV A \<subseteq> sets borel"
using sets.sigma_sets_subset[of A borel] by simp
-lemma borel_eq_sigmaI1:
+lemma%important borel_eq_sigmaI1:
fixes F :: "'i \<Rightarrow> 'a::topological_space set" and X :: "'a::topological_space set set"
assumes borel_eq: "borel = sigma UNIV X"
assumes X: "\<And>x. x \<in> X \<Longrightarrow> x \<in> sets (sigma UNIV (F ` A))"
assumes F: "\<And>i. i \<in> A \<Longrightarrow> F i \<in> sets borel"
shows "borel = sigma UNIV (F ` A)"
unfolding borel_def
-proof (intro sigma_eqI antisym)
+proof%unimportant (intro sigma_eqI antisym)
have borel_rev_eq: "sigma_sets UNIV {S::'a set. open S} = sets borel"
unfolding borel_def by simp
also have "\<dots> = sigma_sets UNIV X"
@@ -477,7 +477,7 @@
unfolding borel_rev_eq using F by (intro borel_sigma_sets_subset) auto
qed auto
-lemma borel_eq_sigmaI2:
+lemma%unimportant borel_eq_sigmaI2:
fixes F :: "'i \<Rightarrow> 'j \<Rightarrow> 'a::topological_space set"
and G :: "'l \<Rightarrow> 'k \<Rightarrow> 'a::topological_space set"
assumes borel_eq: "borel = sigma UNIV ((\<lambda>(i, j). G i j)`B)"
@@ -487,7 +487,7 @@
using assms
by (intro borel_eq_sigmaI1[where X="(\<lambda>(i, j). G i j) ` B" and F="(\<lambda>(i, j). F i j)"]) auto
-lemma borel_eq_sigmaI3:
+lemma%unimportant borel_eq_sigmaI3:
fixes F :: "'i \<Rightarrow> 'j \<Rightarrow> 'a::topological_space set" and X :: "'a::topological_space set set"
assumes borel_eq: "borel = sigma UNIV X"
assumes X: "\<And>x. x \<in> X \<Longrightarrow> x \<in> sets (sigma UNIV ((\<lambda>(i, j). F i j) ` A))"
@@ -495,7 +495,7 @@
shows "borel = sigma UNIV ((\<lambda>(i, j). F i j) ` A)"
using assms by (intro borel_eq_sigmaI1[where X=X and F="(\<lambda>(i, j). F i j)"]) auto
-lemma borel_eq_sigmaI4:
+lemma%unimportant borel_eq_sigmaI4:
fixes F :: "'i \<Rightarrow> 'a::topological_space set"
and G :: "'l \<Rightarrow> 'k \<Rightarrow> 'a::topological_space set"
assumes borel_eq: "borel = sigma UNIV ((\<lambda>(i, j). G i j)`A)"
@@ -504,7 +504,7 @@
shows "borel = sigma UNIV (range F)"
using assms by (intro borel_eq_sigmaI1[where X="(\<lambda>(i, j). G i j) ` A" and F=F]) auto
-lemma borel_eq_sigmaI5:
+lemma%unimportant borel_eq_sigmaI5:
fixes F :: "'i \<Rightarrow> 'j \<Rightarrow> 'a::topological_space set" and G :: "'l \<Rightarrow> 'a::topological_space set"
assumes borel_eq: "borel = sigma UNIV (range G)"
assumes X: "\<And>i. G i \<in> sets (sigma UNIV (range (\<lambda>(i, j). F i j)))"
@@ -512,12 +512,12 @@
shows "borel = sigma UNIV (range (\<lambda>(i, j). F i j))"
using assms by (intro borel_eq_sigmaI1[where X="range G" and F="(\<lambda>(i, j). F i j)"]) auto
-lemma second_countable_borel_measurable:
+lemma%important second_countable_borel_measurable:
fixes X :: "'a::second_countable_topology set set"
assumes eq: "open = generate_topology X"
shows "borel = sigma UNIV X"
unfolding borel_def
-proof (intro sigma_eqI sigma_sets_eqI)
+proof%unimportant (intro sigma_eqI sigma_sets_eqI)
interpret X: sigma_algebra UNIV "sigma_sets UNIV X"
by (rule sigma_algebra_sigma_sets) simp
@@ -556,7 +556,7 @@
qed auto
qed (auto simp: eq intro: generate_topology.Basis)
-lemma borel_eq_closed: "borel = sigma UNIV (Collect closed)"
+lemma%unimportant borel_eq_closed: "borel = sigma UNIV (Collect closed)"
unfolding borel_def
proof (intro sigma_eqI sigma_sets_eqI, safe)
fix x :: "'a set" assume "open x"
@@ -572,13 +572,13 @@
finally show "x \<in> sigma_sets UNIV (Collect open)" by simp
qed simp_all
-lemma borel_eq_countable_basis:
+lemma%important borel_eq_countable_basis:
fixes B::"'a::topological_space set set"
assumes "countable B"
assumes "topological_basis B"
shows "borel = sigma UNIV B"
unfolding borel_def
-proof (intro sigma_eqI sigma_sets_eqI, safe)
+proof%unimportant (intro sigma_eqI sigma_sets_eqI, safe)
interpret countable_basis using assms by unfold_locales
fix X::"'a set" assume "open X"
from open_countable_basisE[OF this] guess B' . note B' = this
@@ -590,7 +590,7 @@
thus "b \<in> sigma_sets UNIV (Collect open)" by auto
qed simp_all
-lemma borel_measurable_continuous_on_restrict:
+lemma%unimportant borel_measurable_continuous_on_restrict:
fixes f :: "'a::topological_space \<Rightarrow> 'b::topological_space"
assumes f: "continuous_on A f"
shows "f \<in> borel_measurable (restrict_space borel A)"
@@ -602,16 +602,16 @@
by (force simp add: sets_restrict_space space_restrict_space)
qed
-lemma borel_measurable_continuous_on1: "continuous_on UNIV f \<Longrightarrow> f \<in> borel_measurable borel"
+lemma%unimportant borel_measurable_continuous_on1: "continuous_on UNIV f \<Longrightarrow> f \<in> borel_measurable borel"
by (drule borel_measurable_continuous_on_restrict) simp
-lemma borel_measurable_continuous_on_if:
+lemma%unimportant borel_measurable_continuous_on_if:
"A \<in> sets borel \<Longrightarrow> continuous_on A f \<Longrightarrow> continuous_on (- A) g \<Longrightarrow>
(\<lambda>x. if x \<in> A then f x else g x) \<in> borel_measurable borel"
by (auto simp add: measurable_If_restrict_space_iff Collect_neg_eq
intro!: borel_measurable_continuous_on_restrict)
-lemma borel_measurable_continuous_countable_exceptions:
+lemma%unimportant borel_measurable_continuous_countable_exceptions:
fixes f :: "'a::t1_space \<Rightarrow> 'b::topological_space"
assumes X: "countable X"
assumes "continuous_on (- X) f"
@@ -623,23 +623,23 @@
by (intro borel_measurable_continuous_on_if assms continuous_intros)
qed auto
-lemma borel_measurable_continuous_on:
+lemma%unimportant borel_measurable_continuous_on:
assumes f: "continuous_on UNIV f" and g: "g \<in> borel_measurable M"
shows "(\<lambda>x. f (g x)) \<in> borel_measurable M"
using measurable_comp[OF g borel_measurable_continuous_on1[OF f]] by (simp add: comp_def)
-lemma borel_measurable_continuous_on_indicator:
+lemma%unimportant borel_measurable_continuous_on_indicator:
fixes f g :: "'a::topological_space \<Rightarrow> 'b::real_normed_vector"
shows "A \<in> sets borel \<Longrightarrow> continuous_on A f \<Longrightarrow> (\<lambda>x. indicator A x *\<^sub>R f x) \<in> borel_measurable borel"
by (subst borel_measurable_restrict_space_iff[symmetric])
(auto intro: borel_measurable_continuous_on_restrict)
-lemma borel_measurable_Pair[measurable (raw)]:
+lemma%important borel_measurable_Pair[measurable (raw)]:
fixes f :: "'a \<Rightarrow> 'b::second_countable_topology" and g :: "'a \<Rightarrow> 'c::second_countable_topology"
assumes f[measurable]: "f \<in> borel_measurable M"
assumes g[measurable]: "g \<in> borel_measurable M"
shows "(\<lambda>x. (f x, g x)) \<in> borel_measurable M"
-proof (subst borel_eq_countable_basis)
+proof%unimportant (subst borel_eq_countable_basis)
let ?B = "SOME B::'b set set. countable B \<and> topological_basis B"
let ?C = "SOME B::'c set set. countable B \<and> topological_basis B"
let ?P = "(\<lambda>(b, c). b \<times> c) ` (?B \<times> ?C)"
@@ -660,21 +660,21 @@
qed auto
qed
-lemma borel_measurable_continuous_Pair:
+lemma%important borel_measurable_continuous_Pair:
fixes f :: "'a \<Rightarrow> 'b::second_countable_topology" and g :: "'a \<Rightarrow> 'c::second_countable_topology"
assumes [measurable]: "f \<in> borel_measurable M"
assumes [measurable]: "g \<in> borel_measurable M"
assumes H: "continuous_on UNIV (\<lambda>x. H (fst x) (snd x))"
shows "(\<lambda>x. H (f x) (g x)) \<in> borel_measurable M"
-proof -
+proof%unimportant -
have eq: "(\<lambda>x. H (f x) (g x)) = (\<lambda>x. (\<lambda>x. H (fst x) (snd x)) (f x, g x))" by auto
show ?thesis
unfolding eq by (rule borel_measurable_continuous_on[OF H]) auto
qed
-subsection \<open>Borel spaces on order topologies\<close>
+subsection%important \<open>Borel spaces on order topologies\<close>
-lemma [measurable]:
+lemma%unimportant [measurable]:
fixes a b :: "'a::linorder_topology"
shows lessThan_borel: "{..< a} \<in> sets borel"
and greaterThan_borel: "{a <..} \<in> sets borel"
@@ -688,7 +688,7 @@
by (blast intro: borel_open borel_closed open_lessThan open_greaterThan open_greaterThanLessThan
closed_atMost closed_atLeast closed_atLeastAtMost)+
-lemma borel_Iio:
+lemma%unimportant borel_Iio:
"borel = sigma UNIV (range lessThan :: 'a::{linorder_topology, second_countable_topology} set set)"
unfolding second_countable_borel_measurable[OF open_generated_order]
proof (intro sigma_eqI sigma_sets_eqI)
@@ -726,7 +726,7 @@
qed auto
qed auto
-lemma borel_Ioi:
+lemma%unimportant borel_Ioi:
"borel = sigma UNIV (range greaterThan :: 'a::{linorder_topology, second_countable_topology} set set)"
unfolding second_countable_borel_measurable[OF open_generated_order]
proof (intro sigma_eqI sigma_sets_eqI)
@@ -764,29 +764,29 @@
qed auto
qed auto
-lemma borel_measurableI_less:
+lemma%unimportant borel_measurableI_less:
fixes f :: "'a \<Rightarrow> 'b::{linorder_topology, second_countable_topology}"
shows "(\<And>y. {x\<in>space M. f x < y} \<in> sets M) \<Longrightarrow> f \<in> borel_measurable M"
unfolding borel_Iio
by (rule measurable_measure_of) (auto simp: Int_def conj_commute)
-lemma borel_measurableI_greater:
+lemma%important borel_measurableI_greater:
fixes f :: "'a \<Rightarrow> 'b::{linorder_topology, second_countable_topology}"
shows "(\<And>y. {x\<in>space M. y < f x} \<in> sets M) \<Longrightarrow> f \<in> borel_measurable M"
unfolding borel_Ioi
- by (rule measurable_measure_of) (auto simp: Int_def conj_commute)
+ by%unimportant (rule measurable_measure_of) (auto simp: Int_def conj_commute)
-lemma borel_measurableI_le:
+lemma%unimportant borel_measurableI_le:
fixes f :: "'a \<Rightarrow> 'b::{linorder_topology, second_countable_topology}"
shows "(\<And>y. {x\<in>space M. f x \<le> y} \<in> sets M) \<Longrightarrow> f \<in> borel_measurable M"
by (rule borel_measurableI_greater) (auto simp: not_le[symmetric])
-lemma borel_measurableI_ge:
+lemma%unimportant borel_measurableI_ge:
fixes f :: "'a \<Rightarrow> 'b::{linorder_topology, second_countable_topology}"
shows "(\<And>y. {x\<in>space M. y \<le> f x} \<in> sets M) \<Longrightarrow> f \<in> borel_measurable M"
by (rule borel_measurableI_less) (auto simp: not_le[symmetric])
-lemma borel_measurable_less[measurable]:
+lemma%unimportant borel_measurable_less[measurable]:
fixes f :: "'a \<Rightarrow> 'b::{second_countable_topology, linorder_topology}"
assumes "f \<in> borel_measurable M"
assumes "g \<in> borel_measurable M"
@@ -800,7 +800,7 @@
finally show ?thesis .
qed
-lemma
+lemma%important
fixes f :: "'a \<Rightarrow> 'b::{second_countable_topology, linorder_topology}"
assumes f[measurable]: "f \<in> borel_measurable M"
assumes g[measurable]: "g \<in> borel_measurable M"
@@ -808,23 +808,23 @@
and borel_measurable_eq[measurable]: "{w \<in> space M. f w = g w} \<in> sets M"
and borel_measurable_neq: "{w \<in> space M. f w \<noteq> g w} \<in> sets M"
unfolding eq_iff not_less[symmetric]
- by measurable
+ by%unimportant measurable
-lemma borel_measurable_SUP[measurable (raw)]:
+lemma%important borel_measurable_SUP[measurable (raw)]:
fixes F :: "_ \<Rightarrow> _ \<Rightarrow> _::{complete_linorder, linorder_topology, second_countable_topology}"
assumes [simp]: "countable I"
assumes [measurable]: "\<And>i. i \<in> I \<Longrightarrow> F i \<in> borel_measurable M"
shows "(\<lambda>x. SUP i:I. F i x) \<in> borel_measurable M"
- by (rule borel_measurableI_greater) (simp add: less_SUP_iff)
+ by%unimportant (rule borel_measurableI_greater) (simp add: less_SUP_iff)
-lemma borel_measurable_INF[measurable (raw)]:
+lemma%unimportant borel_measurable_INF[measurable (raw)]:
fixes F :: "_ \<Rightarrow> _ \<Rightarrow> _::{complete_linorder, linorder_topology, second_countable_topology}"
assumes [simp]: "countable I"
assumes [measurable]: "\<And>i. i \<in> I \<Longrightarrow> F i \<in> borel_measurable M"
shows "(\<lambda>x. INF i:I. F i x) \<in> borel_measurable M"
by (rule borel_measurableI_less) (simp add: INF_less_iff)
-lemma borel_measurable_cSUP[measurable (raw)]:
+lemma%unimportant borel_measurable_cSUP[measurable (raw)]:
fixes F :: "_ \<Rightarrow> _ \<Rightarrow> 'a::{conditionally_complete_linorder, linorder_topology, second_countable_topology}"
assumes [simp]: "countable I"
assumes [measurable]: "\<And>i. i \<in> I \<Longrightarrow> F i \<in> borel_measurable M"
@@ -846,13 +846,13 @@
qed
qed
-lemma borel_measurable_cINF[measurable (raw)]:
+lemma%important borel_measurable_cINF[measurable (raw)]:
fixes F :: "_ \<Rightarrow> _ \<Rightarrow> 'a::{conditionally_complete_linorder, linorder_topology, second_countable_topology}"
assumes [simp]: "countable I"
assumes [measurable]: "\<And>i. i \<in> I \<Longrightarrow> F i \<in> borel_measurable M"
assumes bdd: "\<And>x. x \<in> space M \<Longrightarrow> bdd_below ((\<lambda>i. F i x) ` I)"
shows "(\<lambda>x. INF i:I. F i x) \<in> borel_measurable M"
-proof cases
+proof%unimportant cases
assume "I = {}" then show ?thesis
unfolding \<open>I = {}\<close> image_empty by simp
next
@@ -868,7 +868,7 @@
qed
qed
-lemma borel_measurable_lfp[consumes 1, case_names continuity step]:
+lemma%unimportant borel_measurable_lfp[consumes 1, case_names continuity step]:
fixes F :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'b::{complete_linorder, linorder_topology, second_countable_topology})"
assumes "sup_continuous F"
assumes *: "\<And>f. f \<in> borel_measurable M \<Longrightarrow> F f \<in> borel_measurable M"
@@ -885,7 +885,7 @@
finally show ?thesis .
qed
-lemma borel_measurable_gfp[consumes 1, case_names continuity step]:
+lemma%unimportant borel_measurable_gfp[consumes 1, case_names continuity step]:
fixes F :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'b::{complete_linorder, linorder_topology, second_countable_topology})"
assumes "inf_continuous F"
assumes *: "\<And>f. f \<in> borel_measurable M \<Longrightarrow> F f \<in> borel_measurable M"
@@ -902,56 +902,56 @@
finally show ?thesis .
qed
-lemma borel_measurable_max[measurable (raw)]:
+lemma%unimportant borel_measurable_max[measurable (raw)]:
"f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> (\<lambda>x. max (g x) (f x) :: 'b::{second_countable_topology, linorder_topology}) \<in> borel_measurable M"
by (rule borel_measurableI_less) simp
-lemma borel_measurable_min[measurable (raw)]:
+lemma%unimportant borel_measurable_min[measurable (raw)]:
"f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> (\<lambda>x. min (g x) (f x) :: 'b::{second_countable_topology, linorder_topology}) \<in> borel_measurable M"
by (rule borel_measurableI_greater) simp
-lemma borel_measurable_Min[measurable (raw)]:
+lemma%unimportant borel_measurable_Min[measurable (raw)]:
"finite I \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> f i \<in> borel_measurable M) \<Longrightarrow> (\<lambda>x. Min ((\<lambda>i. f i x)`I) :: 'b::{second_countable_topology, linorder_topology}) \<in> borel_measurable M"
proof (induct I rule: finite_induct)
case (insert i I) then show ?case
by (cases "I = {}") auto
qed auto
-lemma borel_measurable_Max[measurable (raw)]:
+lemma%unimportant borel_measurable_Max[measurable (raw)]:
"finite I \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> f i \<in> borel_measurable M) \<Longrightarrow> (\<lambda>x. Max ((\<lambda>i. f i x)`I) :: 'b::{second_countable_topology, linorder_topology}) \<in> borel_measurable M"
proof (induct I rule: finite_induct)
case (insert i I) then show ?case
by (cases "I = {}") auto
qed auto
-lemma borel_measurable_sup[measurable (raw)]:
+lemma%unimportant borel_measurable_sup[measurable (raw)]:
"f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> (\<lambda>x. sup (g x) (f x) :: 'b::{lattice, second_countable_topology, linorder_topology}) \<in> borel_measurable M"
unfolding sup_max by measurable
-lemma borel_measurable_inf[measurable (raw)]:
+lemma%unimportant borel_measurable_inf[measurable (raw)]:
"f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> (\<lambda>x. inf (g x) (f x) :: 'b::{lattice, second_countable_topology, linorder_topology}) \<in> borel_measurable M"
unfolding inf_min by measurable
-lemma [measurable (raw)]:
+lemma%unimportant [measurable (raw)]:
fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> 'b::{complete_linorder, second_countable_topology, linorder_topology}"
assumes "\<And>i. f i \<in> borel_measurable M"
shows borel_measurable_liminf: "(\<lambda>x. liminf (\<lambda>i. f i x)) \<in> borel_measurable M"
and borel_measurable_limsup: "(\<lambda>x. limsup (\<lambda>i. f i x)) \<in> borel_measurable M"
unfolding liminf_SUP_INF limsup_INF_SUP using assms by auto
-lemma measurable_convergent[measurable (raw)]:
+lemma%unimportant measurable_convergent[measurable (raw)]:
fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> 'b::{complete_linorder, second_countable_topology, linorder_topology}"
assumes [measurable]: "\<And>i. f i \<in> borel_measurable M"
shows "Measurable.pred M (\<lambda>x. convergent (\<lambda>i. f i x))"
unfolding convergent_ereal by measurable
-lemma sets_Collect_convergent[measurable]:
+lemma%unimportant sets_Collect_convergent[measurable]:
fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> 'b::{complete_linorder, second_countable_topology, linorder_topology}"
assumes f[measurable]: "\<And>i. f i \<in> borel_measurable M"
shows "{x\<in>space M. convergent (\<lambda>i. f i x)} \<in> sets M"
by measurable
-lemma borel_measurable_lim[measurable (raw)]:
+lemma%unimportant borel_measurable_lim[measurable (raw)]:
fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> 'b::{complete_linorder, second_countable_topology, linorder_topology}"
assumes [measurable]: "\<And>i. f i \<in> borel_measurable M"
shows "(\<lambda>x. lim (\<lambda>i. f i x)) \<in> borel_measurable M"
@@ -962,7 +962,7 @@
by simp
qed
-lemma borel_measurable_LIMSEQ_order:
+lemma%unimportant borel_measurable_LIMSEQ_order:
fixes u :: "nat \<Rightarrow> 'a \<Rightarrow> 'b::{complete_linorder, second_countable_topology, linorder_topology}"
assumes u': "\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>i. u i x) \<longlonglongrightarrow> u' x"
and u: "\<And>i. u i \<in> borel_measurable M"
@@ -973,16 +973,16 @@
with u show ?thesis by (simp cong: measurable_cong)
qed
-subsection \<open>Borel spaces on topological monoids\<close>
+subsection%important \<open>Borel spaces on topological monoids\<close>
-lemma borel_measurable_add[measurable (raw)]:
+lemma%unimportant borel_measurable_add[measurable (raw)]:
fixes f g :: "'a \<Rightarrow> 'b::{second_countable_topology, topological_monoid_add}"
assumes f: "f \<in> borel_measurable M"
assumes g: "g \<in> borel_measurable M"
shows "(\<lambda>x. f x + g x) \<in> borel_measurable M"
using f g by (rule borel_measurable_continuous_Pair) (intro continuous_intros)
-lemma borel_measurable_sum[measurable (raw)]:
+lemma%unimportant borel_measurable_sum[measurable (raw)]:
fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> 'b::{second_countable_topology, topological_comm_monoid_add}"
assumes "\<And>i. i \<in> S \<Longrightarrow> f i \<in> borel_measurable M"
shows "(\<lambda>x. \<Sum>i\<in>S. f i x) \<in> borel_measurable M"
@@ -991,30 +991,30 @@
thus ?thesis using assms by induct auto
qed simp
-lemma borel_measurable_suminf_order[measurable (raw)]:
+lemma%important borel_measurable_suminf_order[measurable (raw)]:
fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> 'b::{complete_linorder, second_countable_topology, linorder_topology, topological_comm_monoid_add}"
assumes f[measurable]: "\<And>i. f i \<in> borel_measurable M"
shows "(\<lambda>x. suminf (\<lambda>i. f i x)) \<in> borel_measurable M"
unfolding suminf_def sums_def[abs_def] lim_def[symmetric] by simp
-subsection \<open>Borel spaces on Euclidean spaces\<close>
+subsection%important \<open>Borel spaces on Euclidean spaces\<close>
-lemma borel_measurable_inner[measurable (raw)]:
+lemma%important borel_measurable_inner[measurable (raw)]:
fixes f g :: "'a \<Rightarrow> 'b::{second_countable_topology, real_inner}"
assumes "f \<in> borel_measurable M"
assumes "g \<in> borel_measurable M"
shows "(\<lambda>x. f x \<bullet> g x) \<in> borel_measurable M"
using assms
- by (rule borel_measurable_continuous_Pair) (intro continuous_intros)
+ by%unimportant (rule borel_measurable_continuous_Pair) (intro continuous_intros)
notation
eucl_less (infix "<e" 50)
-lemma box_oc: "{x. a <e x \<and> x \<le> b} = {x. a <e x} \<inter> {..b}"
+lemma%important box_oc: "{x. a <e x \<and> x \<le> b} = {x. a <e x} \<inter> {..b}"
and box_co: "{x. a \<le> x \<and> x <e b} = {a..} \<inter> {x. x <e b}"
by auto
-lemma eucl_ivals[measurable]:
+lemma%important eucl_ivals[measurable]:
fixes a b :: "'a::ordered_euclidean_space"
shows "{x. x <e a} \<in> sets borel"
and "{x. a <e x} \<in> sets borel"
@@ -1026,7 +1026,7 @@
unfolding box_oc box_co
by (auto intro: borel_open borel_closed)
-lemma
+lemma%unimportant (*FIX ME this has no name *)
fixes i :: "'a::{second_countable_topology, real_inner}"
shows hafspace_less_borel: "{x. a < x \<bullet> i} \<in> sets borel"
and hafspace_greater_borel: "{x. x \<bullet> i < a} \<in> sets borel"
@@ -1034,7 +1034,7 @@
and hafspace_greater_eq_borel: "{x. x \<bullet> i \<le> a} \<in> sets borel"
by simp_all
-lemma borel_eq_box:
+lemma%unimportant borel_eq_box:
"borel = sigma UNIV (range (\<lambda> (a, b). box a b :: 'a :: euclidean_space set))"
(is "_ = ?SIGMA")
proof (rule borel_eq_sigmaI1[OF borel_def])
@@ -1047,7 +1047,7 @@
done
qed (auto simp: box_def)
-lemma halfspace_gt_in_halfspace:
+lemma%unimportant halfspace_gt_in_halfspace:
assumes i: "i \<in> A"
shows "{x::'a. a < x \<bullet> i} \<in>
sigma_sets UNIV ((\<lambda> (a, i). {x::'a::euclidean_space. x \<bullet> i < a}) ` (UNIV \<times> A))"
@@ -1073,7 +1073,7 @@
by (auto intro!: Diff sigma_sets_Inter i)
qed
-lemma borel_eq_halfspace_less:
+lemma%unimportant borel_eq_halfspace_less:
"borel = sigma UNIV ((\<lambda>(a, i). {x::'a::euclidean_space. x \<bullet> i < a}) ` (UNIV \<times> Basis))"
(is "_ = ?SIGMA")
proof (rule borel_eq_sigmaI2[OF borel_eq_box])
@@ -1086,7 +1086,7 @@
finally show "box a b \<in> sets ?SIGMA" .
qed auto
-lemma borel_eq_halfspace_le:
+lemma%unimportant borel_eq_halfspace_le:
"borel = sigma UNIV ((\<lambda> (a, i). {x::'a::euclidean_space. x \<bullet> i \<le> a}) ` (UNIV \<times> Basis))"
(is "_ = ?SIGMA")
proof (rule borel_eq_sigmaI2[OF borel_eq_halfspace_less])
@@ -1110,7 +1110,7 @@
by (intro sets.countable_UN) (auto intro: i)
qed auto
-lemma borel_eq_halfspace_ge:
+lemma%unimportant borel_eq_halfspace_ge:
"borel = sigma UNIV ((\<lambda> (a, i). {x::'a::euclidean_space. a \<le> x \<bullet> i}) ` (UNIV \<times> Basis))"
(is "_ = ?SIGMA")
proof (rule borel_eq_sigmaI2[OF borel_eq_halfspace_less])
@@ -1120,10 +1120,10 @@
using i by (intro sets.compl_sets) auto
qed auto
-lemma borel_eq_halfspace_greater:
+lemma%important borel_eq_halfspace_greater:
"borel = sigma UNIV ((\<lambda> (a, i). {x::'a::euclidean_space. a < x \<bullet> i}) ` (UNIV \<times> Basis))"
(is "_ = ?SIGMA")
-proof (rule borel_eq_sigmaI2[OF borel_eq_halfspace_le])
+proof%unimportant (rule borel_eq_sigmaI2[OF borel_eq_halfspace_le])
fix a :: real and i :: 'a assume "(a, i) \<in> (UNIV \<times> Basis)"
then have i: "i \<in> Basis" by auto
have *: "{x::'a. x\<bullet>i \<le> a} = space ?SIGMA - {x::'a. a < x\<bullet>i}" by auto
@@ -1131,7 +1131,7 @@
by (intro sets.compl_sets) (auto intro: i)
qed auto
-lemma borel_eq_atMost:
+lemma%unimportant borel_eq_atMost:
"borel = sigma UNIV (range (\<lambda>a. {..a::'a::ordered_euclidean_space}))"
(is "_ = ?SIGMA")
proof (rule borel_eq_sigmaI4[OF borel_eq_halfspace_le])
@@ -1150,7 +1150,7 @@
by (intro sets.countable_UN) auto
qed auto
-lemma borel_eq_greaterThan:
+lemma%unimportant borel_eq_greaterThan:
"borel = sigma UNIV (range (\<lambda>a::'a::ordered_euclidean_space. {x. a <e x}))"
(is "_ = ?SIGMA")
proof (rule borel_eq_sigmaI4[OF borel_eq_halfspace_le])
@@ -1177,7 +1177,7 @@
done
qed auto
-lemma borel_eq_lessThan:
+lemma%unimportant borel_eq_lessThan:
"borel = sigma UNIV (range (\<lambda>a::'a::ordered_euclidean_space. {x. x <e a}))"
(is "_ = ?SIGMA")
proof (rule borel_eq_sigmaI4[OF borel_eq_halfspace_ge])
@@ -1203,10 +1203,10 @@
done
qed auto
-lemma borel_eq_atLeastAtMost:
+lemma%important borel_eq_atLeastAtMost:
"borel = sigma UNIV (range (\<lambda>(a,b). {a..b} ::'a::ordered_euclidean_space set))"
(is "_ = ?SIGMA")
-proof (rule borel_eq_sigmaI5[OF borel_eq_atMost])
+proof%unimportant (rule borel_eq_sigmaI5[OF borel_eq_atMost])
fix a::'a
have *: "{..a} = (\<Union>n::nat. {- real n *\<^sub>R One .. a})"
proof (safe, simp_all add: eucl_le[where 'a='a])
@@ -1225,12 +1225,12 @@
(auto intro!: sigma_sets_top)
qed auto
-lemma borel_set_induct[consumes 1, case_names empty interval compl union]:
+lemma%important borel_set_induct[consumes 1, case_names empty interval compl union]:
assumes "A \<in> sets borel"
assumes empty: "P {}" and int: "\<And>a b. a \<le> b \<Longrightarrow> P {a..b}" and compl: "\<And>A. A \<in> sets borel \<Longrightarrow> P A \<Longrightarrow> P (-A)" and
un: "\<And>f. disjoint_family f \<Longrightarrow> (\<And>i. f i \<in> sets borel) \<Longrightarrow> (\<And>i. P (f i)) \<Longrightarrow> P (\<Union>i::nat. f i)"
shows "P (A::real set)"
-proof-
+proof%unimportant -
let ?G = "range (\<lambda>(a,b). {a..b::real})"
have "Int_stable ?G" "?G \<subseteq> Pow UNIV" "A \<in> sigma_sets UNIV ?G"
using assms(1) by (auto simp add: borel_eq_atLeastAtMost Int_stable_def)
@@ -1247,7 +1247,7 @@
qed (auto intro: empty compl simp: Compl_eq_Diff_UNIV[symmetric] borel_eq_atLeastAtMost)
qed
-lemma borel_sigma_sets_Ioc: "borel = sigma UNIV (range (\<lambda>(a, b). {a <.. b::real}))"
+lemma%unimportant borel_sigma_sets_Ioc: "borel = sigma UNIV (range (\<lambda>(a, b). {a <.. b::real}))"
proof (rule borel_eq_sigmaI5[OF borel_eq_atMost])
fix i :: real
have "{..i} = (\<Union>j::nat. {-j <.. i})"
@@ -1257,10 +1257,10 @@
finally show "{..i} \<in> sets (sigma UNIV (range (\<lambda>(i, j). {i<..j})))" .
qed simp
-lemma eucl_lessThan: "{x::real. x <e a} = lessThan a"
+lemma%unimportant eucl_lessThan: "{x::real. x <e a} = lessThan a"
by (simp add: eucl_less_def lessThan_def)
-lemma borel_eq_atLeastLessThan:
+lemma%unimportant borel_eq_atLeastLessThan:
"borel = sigma UNIV (range (\<lambda>(a, b). {a ..< b :: real}))" (is "_ = ?SIGMA")
proof (rule borel_eq_sigmaI5[OF borel_eq_lessThan])
have move_uminus: "\<And>x y::real. -x \<le> y \<longleftrightarrow> -y \<le> x" by auto
@@ -1271,7 +1271,7 @@
by (auto intro: sigma_sets.intros(2-) simp: eucl_lessThan)
qed auto
-lemma borel_measurable_halfspacesI:
+lemma%unimportant borel_measurable_halfspacesI:
fixes f :: "'a \<Rightarrow> 'c::euclidean_space"
assumes F: "borel = sigma UNIV (F ` (UNIV \<times> Basis))"
and S_eq: "\<And>a i. S a i = f -` F (a,i) \<inter> space M"
@@ -1286,112 +1286,112 @@
by (auto intro!: measurable_measure_of simp: S_eq F)
qed
-lemma borel_measurable_iff_halfspace_le:
+lemma%unimportant borel_measurable_iff_halfspace_le:
fixes f :: "'a \<Rightarrow> 'c::euclidean_space"
shows "f \<in> borel_measurable M = (\<forall>i\<in>Basis. \<forall>a. {w \<in> space M. f w \<bullet> i \<le> a} \<in> sets M)"
by (rule borel_measurable_halfspacesI[OF borel_eq_halfspace_le]) auto
-lemma borel_measurable_iff_halfspace_less:
+lemma%unimportant borel_measurable_iff_halfspace_less:
fixes f :: "'a \<Rightarrow> 'c::euclidean_space"
shows "f \<in> borel_measurable M \<longleftrightarrow> (\<forall>i\<in>Basis. \<forall>a. {w \<in> space M. f w \<bullet> i < a} \<in> sets M)"
by (rule borel_measurable_halfspacesI[OF borel_eq_halfspace_less]) auto
-lemma borel_measurable_iff_halfspace_ge:
+lemma%unimportant borel_measurable_iff_halfspace_ge:
fixes f :: "'a \<Rightarrow> 'c::euclidean_space"
shows "f \<in> borel_measurable M = (\<forall>i\<in>Basis. \<forall>a. {w \<in> space M. a \<le> f w \<bullet> i} \<in> sets M)"
by (rule borel_measurable_halfspacesI[OF borel_eq_halfspace_ge]) auto
-lemma borel_measurable_iff_halfspace_greater:
+lemma%unimportant borel_measurable_iff_halfspace_greater:
fixes f :: "'a \<Rightarrow> 'c::euclidean_space"
shows "f \<in> borel_measurable M \<longleftrightarrow> (\<forall>i\<in>Basis. \<forall>a. {w \<in> space M. a < f w \<bullet> i} \<in> sets M)"
by (rule borel_measurable_halfspacesI[OF borel_eq_halfspace_greater]) auto
-lemma borel_measurable_iff_le:
+lemma%unimportant borel_measurable_iff_le:
"(f::'a \<Rightarrow> real) \<in> borel_measurable M = (\<forall>a. {w \<in> space M. f w \<le> a} \<in> sets M)"
using borel_measurable_iff_halfspace_le[where 'c=real] by simp
-lemma borel_measurable_iff_less:
+lemma%unimportant borel_measurable_iff_less:
"(f::'a \<Rightarrow> real) \<in> borel_measurable M = (\<forall>a. {w \<in> space M. f w < a} \<in> sets M)"
using borel_measurable_iff_halfspace_less[where 'c=real] by simp
-lemma borel_measurable_iff_ge:
+lemma%unimportant borel_measurable_iff_ge:
"(f::'a \<Rightarrow> real) \<in> borel_measurable M = (\<forall>a. {w \<in> space M. a \<le> f w} \<in> sets M)"
using borel_measurable_iff_halfspace_ge[where 'c=real]
by simp
-lemma borel_measurable_iff_greater:
+lemma%unimportant borel_measurable_iff_greater:
"(f::'a \<Rightarrow> real) \<in> borel_measurable M = (\<forall>a. {w \<in> space M. a < f w} \<in> sets M)"
using borel_measurable_iff_halfspace_greater[where 'c=real] by simp
-lemma borel_measurable_euclidean_space:
+lemma%important borel_measurable_euclidean_space:
fixes f :: "'a \<Rightarrow> 'c::euclidean_space"
shows "f \<in> borel_measurable M \<longleftrightarrow> (\<forall>i\<in>Basis. (\<lambda>x. f x \<bullet> i) \<in> borel_measurable M)"
-proof safe
+proof%unimportant safe
assume f: "\<forall>i\<in>Basis. (\<lambda>x. f x \<bullet> i) \<in> borel_measurable M"
then show "f \<in> borel_measurable M"
by (subst borel_measurable_iff_halfspace_le) auto
qed auto
-subsection "Borel measurable operators"
+subsection%important "Borel measurable operators"
-lemma borel_measurable_norm[measurable]: "norm \<in> borel_measurable borel"
- by (intro borel_measurable_continuous_on1 continuous_intros)
+lemma%important borel_measurable_norm[measurable]: "norm \<in> borel_measurable borel"
+ by%unimportant (intro borel_measurable_continuous_on1 continuous_intros)
-lemma borel_measurable_sgn [measurable]: "(sgn::'a::real_normed_vector \<Rightarrow> 'a) \<in> borel_measurable borel"
+lemma%unimportant borel_measurable_sgn [measurable]: "(sgn::'a::real_normed_vector \<Rightarrow> 'a) \<in> borel_measurable borel"
by (rule borel_measurable_continuous_countable_exceptions[where X="{0}"])
(auto intro!: continuous_on_sgn continuous_on_id)
-lemma borel_measurable_uminus[measurable (raw)]:
+lemma%important borel_measurable_uminus[measurable (raw)]:
fixes g :: "'a \<Rightarrow> 'b::{second_countable_topology, real_normed_vector}"
assumes g: "g \<in> borel_measurable M"
shows "(\<lambda>x. - g x) \<in> borel_measurable M"
- by (rule borel_measurable_continuous_on[OF _ g]) (intro continuous_intros)
+ by%unimportant (rule borel_measurable_continuous_on[OF _ g]) (intro continuous_intros)
-lemma borel_measurable_diff[measurable (raw)]:
+lemma%important borel_measurable_diff[measurable (raw)]:
fixes f :: "'a \<Rightarrow> 'b::{second_countable_topology, real_normed_vector}"
assumes f: "f \<in> borel_measurable M"
assumes g: "g \<in> borel_measurable M"
shows "(\<lambda>x. f x - g x) \<in> borel_measurable M"
- using borel_measurable_add [of f M "- g"] assms by (simp add: fun_Compl_def)
+ using%unimportant borel_measurable_add [of f M "- g"] assms by (simp add: fun_Compl_def)
-lemma borel_measurable_times[measurable (raw)]:
+lemma%unimportant borel_measurable_times[measurable (raw)]:
fixes f :: "'a \<Rightarrow> 'b::{second_countable_topology, real_normed_algebra}"
assumes f: "f \<in> borel_measurable M"
assumes g: "g \<in> borel_measurable M"
shows "(\<lambda>x. f x * g x) \<in> borel_measurable M"
using f g by (rule borel_measurable_continuous_Pair) (intro continuous_intros)
-lemma borel_measurable_prod[measurable (raw)]:
+lemma%important borel_measurable_prod[measurable (raw)]:
fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> 'b::{second_countable_topology, real_normed_field}"
assumes "\<And>i. i \<in> S \<Longrightarrow> f i \<in> borel_measurable M"
shows "(\<lambda>x. \<Prod>i\<in>S. f i x) \<in> borel_measurable M"
-proof cases
+proof%unimportant cases
assume "finite S"
thus ?thesis using assms by induct auto
qed simp
-lemma borel_measurable_dist[measurable (raw)]:
+lemma%important borel_measurable_dist[measurable (raw)]:
fixes g f :: "'a \<Rightarrow> 'b::{second_countable_topology, metric_space}"
assumes f: "f \<in> borel_measurable M"
assumes g: "g \<in> borel_measurable M"
shows "(\<lambda>x. dist (f x) (g x)) \<in> borel_measurable M"
- using f g by (rule borel_measurable_continuous_Pair) (intro continuous_intros)
+ using%unimportant f g by (rule borel_measurable_continuous_Pair) (intro continuous_intros)
-lemma borel_measurable_scaleR[measurable (raw)]:
+lemma%unimportant borel_measurable_scaleR[measurable (raw)]:
fixes g :: "'a \<Rightarrow> 'b::{second_countable_topology, real_normed_vector}"
assumes f: "f \<in> borel_measurable M"
assumes g: "g \<in> borel_measurable M"
shows "(\<lambda>x. f x *\<^sub>R g x) \<in> borel_measurable M"
using f g by (rule borel_measurable_continuous_Pair) (intro continuous_intros)
-lemma borel_measurable_uminus_eq [simp]:
+lemma%unimportant borel_measurable_uminus_eq [simp]:
fixes f :: "'a \<Rightarrow> 'b::{second_countable_topology, real_normed_vector}"
shows "(\<lambda>x. - f x) \<in> borel_measurable M \<longleftrightarrow> f \<in> borel_measurable M" (is "?l = ?r")
proof
assume ?l from borel_measurable_uminus[OF this] show ?r by simp
qed auto
-lemma affine_borel_measurable_vector:
+lemma%unimportant affine_borel_measurable_vector:
fixes f :: "'a \<Rightarrow> 'x::real_normed_vector"
assumes "f \<in> borel_measurable M"
shows "(\<lambda>x. a + b *\<^sub>R f x) \<in> borel_measurable M"
@@ -1412,15 +1412,15 @@
qed simp
qed
-lemma borel_measurable_const_scaleR[measurable (raw)]:
+lemma%unimportant borel_measurable_const_scaleR[measurable (raw)]:
"f \<in> borel_measurable M \<Longrightarrow> (\<lambda>x. b *\<^sub>R f x ::'a::real_normed_vector) \<in> borel_measurable M"
using affine_borel_measurable_vector[of f M 0 b] by simp
-lemma borel_measurable_const_add[measurable (raw)]:
+lemma%unimportant borel_measurable_const_add[measurable (raw)]:
"f \<in> borel_measurable M \<Longrightarrow> (\<lambda>x. a + f x ::'a::real_normed_vector) \<in> borel_measurable M"
using affine_borel_measurable_vector[of f M a 1] by simp
-lemma borel_measurable_inverse[measurable (raw)]:
+lemma%unimportant borel_measurable_inverse[measurable (raw)]:
fixes f :: "'a \<Rightarrow> 'b::real_normed_div_algebra"
assumes f: "f \<in> borel_measurable M"
shows "(\<lambda>x. inverse (f x)) \<in> borel_measurable M"
@@ -1429,27 +1429,27 @@
apply (auto intro!: continuous_on_inverse continuous_on_id)
done
-lemma borel_measurable_divide[measurable (raw)]:
+lemma%unimportant borel_measurable_divide[measurable (raw)]:
"f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow>
(\<lambda>x. f x / g x::'b::{second_countable_topology, real_normed_div_algebra}) \<in> borel_measurable M"
by (simp add: divide_inverse)
-lemma borel_measurable_abs[measurable (raw)]:
+lemma%unimportant borel_measurable_abs[measurable (raw)]:
"f \<in> borel_measurable M \<Longrightarrow> (\<lambda>x. \<bar>f x :: real\<bar>) \<in> borel_measurable M"
unfolding abs_real_def by simp
-lemma borel_measurable_nth[measurable (raw)]:
+lemma%unimportant borel_measurable_nth[measurable (raw)]:
"(\<lambda>x::real^'n. x $ i) \<in> borel_measurable borel"
by (simp add: cart_eq_inner_axis)
-lemma convex_measurable:
+lemma%important convex_measurable:
fixes A :: "'a :: euclidean_space set"
shows "X \<in> borel_measurable M \<Longrightarrow> X ` space M \<subseteq> A \<Longrightarrow> open A \<Longrightarrow> convex_on A q \<Longrightarrow>
(\<lambda>x. q (X x)) \<in> borel_measurable M"
- by (rule measurable_compose[where f=X and N="restrict_space borel A"])
+ by%unimportant (rule measurable_compose[where f=X and N="restrict_space borel A"])
(auto intro!: borel_measurable_continuous_on_restrict convex_on_continuous measurable_restrict_space2)
-lemma borel_measurable_ln[measurable (raw)]:
+lemma%unimportant borel_measurable_ln[measurable (raw)]:
assumes f: "f \<in> borel_measurable M"
shows "(\<lambda>x. ln (f x :: real)) \<in> borel_measurable M"
apply (rule measurable_compose[OF f])
@@ -1457,15 +1457,15 @@
apply (auto intro!: continuous_on_ln continuous_on_id)
done
-lemma borel_measurable_log[measurable (raw)]:
+lemma%unimportant borel_measurable_log[measurable (raw)]:
"f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> (\<lambda>x. log (g x) (f x)) \<in> borel_measurable M"
unfolding log_def by auto
-lemma borel_measurable_exp[measurable]:
+lemma%unimportant borel_measurable_exp[measurable]:
"(exp::'a::{real_normed_field,banach}\<Rightarrow>'a) \<in> borel_measurable borel"
by (intro borel_measurable_continuous_on1 continuous_at_imp_continuous_on ballI isCont_exp)
-lemma measurable_real_floor[measurable]:
+lemma%unimportant measurable_real_floor[measurable]:
"(floor :: real \<Rightarrow> int) \<in> measurable borel (count_space UNIV)"
proof -
have "\<And>a x. \<lfloor>x\<rfloor> = a \<longleftrightarrow> (real_of_int a \<le> x \<and> x < real_of_int (a + 1))"
@@ -1474,44 +1474,44 @@
by (auto simp: vimage_def measurable_count_space_eq2_countable)
qed
-lemma measurable_real_ceiling[measurable]:
+lemma%unimportant measurable_real_ceiling[measurable]:
"(ceiling :: real \<Rightarrow> int) \<in> measurable borel (count_space UNIV)"
unfolding ceiling_def[abs_def] by simp
-lemma borel_measurable_real_floor: "(\<lambda>x::real. real_of_int \<lfloor>x\<rfloor>) \<in> borel_measurable borel"
+lemma%unimportant borel_measurable_real_floor: "(\<lambda>x::real. real_of_int \<lfloor>x\<rfloor>) \<in> borel_measurable borel"
by simp
-lemma borel_measurable_root [measurable]: "root n \<in> borel_measurable borel"
+lemma%unimportant borel_measurable_root [measurable]: "root n \<in> borel_measurable borel"
by (intro borel_measurable_continuous_on1 continuous_intros)
-lemma borel_measurable_sqrt [measurable]: "sqrt \<in> borel_measurable borel"
+lemma%unimportant borel_measurable_sqrt [measurable]: "sqrt \<in> borel_measurable borel"
by (intro borel_measurable_continuous_on1 continuous_intros)
-lemma borel_measurable_power [measurable (raw)]:
+lemma%unimportant borel_measurable_power [measurable (raw)]:
fixes f :: "_ \<Rightarrow> 'b::{power,real_normed_algebra}"
assumes f: "f \<in> borel_measurable M"
shows "(\<lambda>x. (f x) ^ n) \<in> borel_measurable M"
by (intro borel_measurable_continuous_on [OF _ f] continuous_intros)
-lemma borel_measurable_Re [measurable]: "Re \<in> borel_measurable borel"
+lemma%unimportant borel_measurable_Re [measurable]: "Re \<in> borel_measurable borel"
by (intro borel_measurable_continuous_on1 continuous_intros)
-lemma borel_measurable_Im [measurable]: "Im \<in> borel_measurable borel"
+lemma%unimportant borel_measurable_Im [measurable]: "Im \<in> borel_measurable borel"
by (intro borel_measurable_continuous_on1 continuous_intros)
-lemma borel_measurable_of_real [measurable]: "(of_real :: _ \<Rightarrow> (_::real_normed_algebra)) \<in> borel_measurable borel"
+lemma%unimportant borel_measurable_of_real [measurable]: "(of_real :: _ \<Rightarrow> (_::real_normed_algebra)) \<in> borel_measurable borel"
by (intro borel_measurable_continuous_on1 continuous_intros)
-lemma borel_measurable_sin [measurable]: "(sin :: _ \<Rightarrow> (_::{real_normed_field,banach})) \<in> borel_measurable borel"
+lemma%unimportant borel_measurable_sin [measurable]: "(sin :: _ \<Rightarrow> (_::{real_normed_field,banach})) \<in> borel_measurable borel"
by (intro borel_measurable_continuous_on1 continuous_intros)
-lemma borel_measurable_cos [measurable]: "(cos :: _ \<Rightarrow> (_::{real_normed_field,banach})) \<in> borel_measurable borel"
+lemma%unimportant borel_measurable_cos [measurable]: "(cos :: _ \<Rightarrow> (_::{real_normed_field,banach})) \<in> borel_measurable borel"
by (intro borel_measurable_continuous_on1 continuous_intros)
-lemma borel_measurable_arctan [measurable]: "arctan \<in> borel_measurable borel"
+lemma%unimportant borel_measurable_arctan [measurable]: "arctan \<in> borel_measurable borel"
by (intro borel_measurable_continuous_on1 continuous_intros)
-lemma borel_measurable_complex_iff:
+lemma%important borel_measurable_complex_iff:
"f \<in> borel_measurable M \<longleftrightarrow>
(\<lambda>x. Re (f x)) \<in> borel_measurable M \<and> (\<lambda>x. Im (f x)) \<in> borel_measurable M"
apply auto
@@ -1520,21 +1520,21 @@
apply auto
done
-lemma powr_real_measurable [measurable]:
+lemma%important powr_real_measurable [measurable]:
assumes "f \<in> measurable M borel" "g \<in> measurable M borel"
shows "(\<lambda>x. f x powr g x :: real) \<in> measurable M borel"
- using assms by (simp_all add: powr_def)
+ using%unimportant assms by (simp_all add: powr_def)
-lemma measurable_of_bool[measurable]: "of_bool \<in> count_space UNIV \<rightarrow>\<^sub>M borel"
+lemma%unimportant measurable_of_bool[measurable]: "of_bool \<in> count_space UNIV \<rightarrow>\<^sub>M borel"
by simp
-subsection "Borel space on the extended reals"
+subsection%important "Borel space on the extended reals"
-lemma borel_measurable_ereal[measurable (raw)]:
+lemma%unimportant borel_measurable_ereal[measurable (raw)]:
assumes f: "f \<in> borel_measurable M" shows "(\<lambda>x. ereal (f x)) \<in> borel_measurable M"
using continuous_on_ereal f by (rule borel_measurable_continuous_on) (rule continuous_on_id)
-lemma borel_measurable_real_of_ereal[measurable (raw)]:
+lemma%unimportant borel_measurable_real_of_ereal[measurable (raw)]:
fixes f :: "'a \<Rightarrow> ereal"
assumes f: "f \<in> borel_measurable M"
shows "(\<lambda>x. real_of_ereal (f x)) \<in> borel_measurable M"
@@ -1543,7 +1543,7 @@
apply (auto intro: continuous_on_real simp: Compl_eq_Diff_UNIV)
done
-lemma borel_measurable_ereal_cases:
+lemma%unimportant borel_measurable_ereal_cases:
fixes f :: "'a \<Rightarrow> ereal"
assumes f: "f \<in> borel_measurable M"
assumes H: "(\<lambda>x. H (ereal (real_of_ereal (f x)))) \<in> borel_measurable M"
@@ -1554,20 +1554,20 @@
with f H show ?thesis by simp
qed
-lemma
+lemma%unimportant (*FIX ME needs a name *)
fixes f :: "'a \<Rightarrow> ereal" assumes f[measurable]: "f \<in> borel_measurable M"
shows borel_measurable_ereal_abs[measurable(raw)]: "(\<lambda>x. \<bar>f x\<bar>) \<in> borel_measurable M"
and borel_measurable_ereal_inverse[measurable(raw)]: "(\<lambda>x. inverse (f x) :: ereal) \<in> borel_measurable M"
and borel_measurable_uminus_ereal[measurable(raw)]: "(\<lambda>x. - f x :: ereal) \<in> borel_measurable M"
by (auto simp del: abs_real_of_ereal simp: borel_measurable_ereal_cases[OF f] measurable_If)
-lemma borel_measurable_uminus_eq_ereal[simp]:
+lemma%unimportant borel_measurable_uminus_eq_ereal[simp]:
"(\<lambda>x. - f x :: ereal) \<in> borel_measurable M \<longleftrightarrow> f \<in> borel_measurable M" (is "?l = ?r")
proof
assume ?l from borel_measurable_uminus_ereal[OF this] show ?r by simp
qed auto
-lemma set_Collect_ereal2:
+lemma%important set_Collect_ereal2:
fixes f g :: "'a \<Rightarrow> ereal"
assumes f: "f \<in> borel_measurable M"
assumes g: "g \<in> borel_measurable M"
@@ -1577,7 +1577,7 @@
"{x \<in> space borel. H (ereal x) (-\<infinity>)} \<in> sets borel"
"{x \<in> space borel. H (ereal x) (\<infinity>)} \<in> sets borel"
shows "{x \<in> space M. H (f x) (g x)} \<in> sets M"
-proof -
+proof%unimportant -
let ?G = "\<lambda>y x. if g x = \<infinity> then H y \<infinity> else if g x = -\<infinity> then H y (-\<infinity>) else H y (ereal (real_of_ereal (g x)))"
let ?F = "\<lambda>x. if f x = \<infinity> then ?G \<infinity> x else if f x = -\<infinity> then ?G (-\<infinity>) x else ?G (ereal (real_of_ereal (f x))) x"
{ fix x have "H (f x) (g x) = ?F x" by (cases "f x" "g x" rule: ereal2_cases) auto }
@@ -1586,7 +1586,7 @@
by (subst *) (simp del: space_borel split del: if_split)
qed
-lemma borel_measurable_ereal_iff:
+lemma%unimportant borel_measurable_ereal_iff:
shows "(\<lambda>x. ereal (f x)) \<in> borel_measurable M \<longleftrightarrow> f \<in> borel_measurable M"
proof
assume "(\<lambda>x. ereal (f x)) \<in> borel_measurable M"
@@ -1594,15 +1594,15 @@
show "f \<in> borel_measurable M" by auto
qed auto
-lemma borel_measurable_erealD[measurable_dest]:
+lemma%unimportant borel_measurable_erealD[measurable_dest]:
"(\<lambda>x. ereal (f x)) \<in> borel_measurable M \<Longrightarrow> g \<in> measurable N M \<Longrightarrow> (\<lambda>x. f (g x)) \<in> borel_measurable N"
unfolding borel_measurable_ereal_iff by simp
-lemma borel_measurable_ereal_iff_real:
+lemma%important borel_measurable_ereal_iff_real:
fixes f :: "'a \<Rightarrow> ereal"
shows "f \<in> borel_measurable M \<longleftrightarrow>
((\<lambda>x. real_of_ereal (f x)) \<in> borel_measurable M \<and> f -` {\<infinity>} \<inter> space M \<in> sets M \<and> f -` {-\<infinity>} \<inter> space M \<in> sets M)"
-proof safe
+proof%unimportant safe
assume *: "(\<lambda>x. real_of_ereal (f x)) \<in> borel_measurable M" "f -` {\<infinity>} \<inter> space M \<in> sets M" "f -` {-\<infinity>} \<inter> space M \<in> sets M"
have "f -` {\<infinity>} \<inter> space M = {x\<in>space M. f x = \<infinity>}" "f -` {-\<infinity>} \<inter> space M = {x\<in>space M. f x = -\<infinity>}" by auto
with * have **: "{x\<in>space M. f x = \<infinity>} \<in> sets M" "{x\<in>space M. f x = -\<infinity>} \<in> sets M" by simp_all
@@ -1612,15 +1612,15 @@
finally show "f \<in> borel_measurable M" .
qed simp_all
-lemma borel_measurable_ereal_iff_Iio:
+lemma%unimportant borel_measurable_ereal_iff_Iio:
"(f::'a \<Rightarrow> ereal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. f -` {..< a} \<inter> space M \<in> sets M)"
by (auto simp: borel_Iio measurable_iff_measure_of)
-lemma borel_measurable_ereal_iff_Ioi:
+lemma%unimportant borel_measurable_ereal_iff_Ioi:
"(f::'a \<Rightarrow> ereal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. f -` {a <..} \<inter> space M \<in> sets M)"
by (auto simp: borel_Ioi measurable_iff_measure_of)
-lemma vimage_sets_compl_iff:
+lemma%unimportant vimage_sets_compl_iff:
"f -` A \<inter> space M \<in> sets M \<longleftrightarrow> f -` (- A) \<inter> space M \<in> sets M"
proof -
{ fix A assume "f -` A \<inter> space M \<in> sets M"
@@ -1630,15 +1630,15 @@
by (metis double_complement)
qed
-lemma borel_measurable_iff_Iic_ereal:
+lemma%unimportant borel_measurable_iff_Iic_ereal:
"(f::'a\<Rightarrow>ereal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. f -` {..a} \<inter> space M \<in> sets M)"
unfolding borel_measurable_ereal_iff_Ioi vimage_sets_compl_iff[where A="{a <..}" for a] by simp
-lemma borel_measurable_iff_Ici_ereal:
+lemma%unimportant borel_measurable_iff_Ici_ereal:
"(f::'a \<Rightarrow> ereal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. f -` {a..} \<inter> space M \<in> sets M)"
unfolding borel_measurable_ereal_iff_Iio vimage_sets_compl_iff[where A="{..< a}" for a] by simp
-lemma borel_measurable_ereal2:
+lemma%important borel_measurable_ereal2:
fixes f g :: "'a \<Rightarrow> ereal"
assumes f: "f \<in> borel_measurable M"
assumes g: "g \<in> borel_measurable M"
@@ -1648,7 +1648,7 @@
"(\<lambda>x. H (ereal (real_of_ereal (f x))) (-\<infinity>)) \<in> borel_measurable M"
"(\<lambda>x. H (ereal (real_of_ereal (f x))) (\<infinity>)) \<in> borel_measurable M"
shows "(\<lambda>x. H (f x) (g x)) \<in> borel_measurable M"
-proof -
+proof%unimportant -
let ?G = "\<lambda>y x. if g x = \<infinity> then H y \<infinity> else if g x = - \<infinity> then H y (-\<infinity>) else H y (ereal (real_of_ereal (g x)))"
let ?F = "\<lambda>x. if f x = \<infinity> then ?G \<infinity> x else if f x = - \<infinity> then ?G (-\<infinity>) x else ?G (ereal (real_of_ereal (f x))) x"
{ fix x have "H (f x) (g x) = ?F x" by (cases "f x" "g x" rule: ereal2_cases) auto }
@@ -1656,14 +1656,14 @@
from assms show ?thesis unfolding * by simp
qed
-lemma [measurable(raw)]:
+lemma%unimportant [measurable(raw)]:
fixes f :: "'a \<Rightarrow> ereal"
assumes [measurable]: "f \<in> borel_measurable M" "g \<in> borel_measurable M"
shows borel_measurable_ereal_add: "(\<lambda>x. f x + g x) \<in> borel_measurable M"
and borel_measurable_ereal_times: "(\<lambda>x. f x * g x) \<in> borel_measurable M"
by (simp_all add: borel_measurable_ereal2)
-lemma [measurable(raw)]:
+lemma%unimportant [measurable(raw)]:
fixes f g :: "'a \<Rightarrow> ereal"
assumes "f \<in> borel_measurable M"
assumes "g \<in> borel_measurable M"
@@ -1671,42 +1671,42 @@
and borel_measurable_ereal_divide: "(\<lambda>x. f x / g x) \<in> borel_measurable M"
using assms by (simp_all add: minus_ereal_def divide_ereal_def)
-lemma borel_measurable_ereal_sum[measurable (raw)]:
+lemma%unimportant borel_measurable_ereal_sum[measurable (raw)]:
fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> ereal"
assumes "\<And>i. i \<in> S \<Longrightarrow> f i \<in> borel_measurable M"
shows "(\<lambda>x. \<Sum>i\<in>S. f i x) \<in> borel_measurable M"
using assms by (induction S rule: infinite_finite_induct) auto
-lemma borel_measurable_ereal_prod[measurable (raw)]:
+lemma%unimportant borel_measurable_ereal_prod[measurable (raw)]:
fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> ereal"
assumes "\<And>i. i \<in> S \<Longrightarrow> f i \<in> borel_measurable M"
shows "(\<lambda>x. \<Prod>i\<in>S. f i x) \<in> borel_measurable M"
using assms by (induction S rule: infinite_finite_induct) auto
-lemma borel_measurable_extreal_suminf[measurable (raw)]:
+lemma%unimportant borel_measurable_extreal_suminf[measurable (raw)]:
fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> ereal"
assumes [measurable]: "\<And>i. f i \<in> borel_measurable M"
shows "(\<lambda>x. (\<Sum>i. f i x)) \<in> borel_measurable M"
unfolding suminf_def sums_def[abs_def] lim_def[symmetric] by simp
-subsection "Borel space on the extended non-negative reals"
+subsection%important "Borel space on the extended non-negative reals"
text \<open> @{type ennreal} is a topological monoid, so no rules for plus are required, also all order
statements are usually done on type classes. \<close>
-lemma measurable_enn2ereal[measurable]: "enn2ereal \<in> borel \<rightarrow>\<^sub>M borel"
+lemma%unimportant measurable_enn2ereal[measurable]: "enn2ereal \<in> borel \<rightarrow>\<^sub>M borel"
by (intro borel_measurable_continuous_on1 continuous_on_enn2ereal)
-lemma measurable_e2ennreal[measurable]: "e2ennreal \<in> borel \<rightarrow>\<^sub>M borel"
+lemma%unimportant measurable_e2ennreal[measurable]: "e2ennreal \<in> borel \<rightarrow>\<^sub>M borel"
by (intro borel_measurable_continuous_on1 continuous_on_e2ennreal)
-lemma borel_measurable_enn2real[measurable (raw)]:
+lemma%unimportant borel_measurable_enn2real[measurable (raw)]:
"f \<in> M \<rightarrow>\<^sub>M borel \<Longrightarrow> (\<lambda>x. enn2real (f x)) \<in> M \<rightarrow>\<^sub>M borel"
unfolding enn2real_def[abs_def] by measurable
-definition [simp]: "is_borel f M \<longleftrightarrow> f \<in> borel_measurable M"
+definition%important [simp]: "is_borel f M \<longleftrightarrow> f \<in> borel_measurable M"
-lemma is_borel_transfer[transfer_rule]: "rel_fun (rel_fun (=) pcr_ennreal) (=) is_borel is_borel"
+lemma%unimportant is_borel_transfer[transfer_rule]: "rel_fun (rel_fun (=) pcr_ennreal) (=) is_borel is_borel"
unfolding is_borel_def[abs_def]
proof (safe intro!: rel_funI ext dest!: rel_fun_eq_pcr_ennreal[THEN iffD1])
fix f and M :: "'a measure"
@@ -1718,14 +1718,14 @@
includes ennreal.lifting
begin
-lemma measurable_ennreal[measurable]: "ennreal \<in> borel \<rightarrow>\<^sub>M borel"
+lemma%unimportant measurable_ennreal[measurable]: "ennreal \<in> borel \<rightarrow>\<^sub>M borel"
unfolding is_borel_def[symmetric]
by transfer simp
-lemma borel_measurable_ennreal_iff[simp]:
+lemma%important borel_measurable_ennreal_iff[simp]:
assumes [simp]: "\<And>x. x \<in> space M \<Longrightarrow> 0 \<le> f x"
shows "(\<lambda>x. ennreal (f x)) \<in> M \<rightarrow>\<^sub>M borel \<longleftrightarrow> f \<in> M \<rightarrow>\<^sub>M borel"
-proof safe
+proof%unimportant safe
assume "(\<lambda>x. ennreal (f x)) \<in> M \<rightarrow>\<^sub>M borel"
then have "(\<lambda>x. enn2real (ennreal (f x))) \<in> M \<rightarrow>\<^sub>M borel"
by measurable
@@ -1733,44 +1733,44 @@
by (rule measurable_cong[THEN iffD1, rotated]) auto
qed measurable
-lemma borel_measurable_times_ennreal[measurable (raw)]:
+lemma%unimportant borel_measurable_times_ennreal[measurable (raw)]:
fixes f g :: "'a \<Rightarrow> ennreal"
shows "f \<in> M \<rightarrow>\<^sub>M borel \<Longrightarrow> g \<in> M \<rightarrow>\<^sub>M borel \<Longrightarrow> (\<lambda>x. f x * g x) \<in> M \<rightarrow>\<^sub>M borel"
unfolding is_borel_def[symmetric] by transfer simp
-lemma borel_measurable_inverse_ennreal[measurable (raw)]:
+lemma%unimportant borel_measurable_inverse_ennreal[measurable (raw)]:
fixes f :: "'a \<Rightarrow> ennreal"
shows "f \<in> M \<rightarrow>\<^sub>M borel \<Longrightarrow> (\<lambda>x. inverse (f x)) \<in> M \<rightarrow>\<^sub>M borel"
unfolding is_borel_def[symmetric] by transfer simp
-lemma borel_measurable_divide_ennreal[measurable (raw)]:
+lemma%unimportant borel_measurable_divide_ennreal[measurable (raw)]:
fixes f :: "'a \<Rightarrow> ennreal"
shows "f \<in> M \<rightarrow>\<^sub>M borel \<Longrightarrow> g \<in> M \<rightarrow>\<^sub>M borel \<Longrightarrow> (\<lambda>x. f x / g x) \<in> M \<rightarrow>\<^sub>M borel"
unfolding divide_ennreal_def by simp
-lemma borel_measurable_minus_ennreal[measurable (raw)]:
+lemma%unimportant borel_measurable_minus_ennreal[measurable (raw)]:
fixes f :: "'a \<Rightarrow> ennreal"
shows "f \<in> M \<rightarrow>\<^sub>M borel \<Longrightarrow> g \<in> M \<rightarrow>\<^sub>M borel \<Longrightarrow> (\<lambda>x. f x - g x) \<in> M \<rightarrow>\<^sub>M borel"
unfolding is_borel_def[symmetric] by transfer simp
-lemma borel_measurable_prod_ennreal[measurable (raw)]:
+lemma%important borel_measurable_prod_ennreal[measurable (raw)]:
fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> ennreal"
assumes "\<And>i. i \<in> S \<Longrightarrow> f i \<in> borel_measurable M"
shows "(\<lambda>x. \<Prod>i\<in>S. f i x) \<in> borel_measurable M"
- using assms by (induction S rule: infinite_finite_induct) auto
+ using%unimportant assms by (induction S rule: infinite_finite_induct) auto
end
hide_const (open) is_borel
-subsection \<open>LIMSEQ is borel measurable\<close>
+subsection%important \<open>LIMSEQ is borel measurable\<close>
-lemma borel_measurable_LIMSEQ_real:
+lemma%important borel_measurable_LIMSEQ_real:
fixes u :: "nat \<Rightarrow> 'a \<Rightarrow> real"
assumes u': "\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>i. u i x) \<longlonglongrightarrow> u' x"
and u: "\<And>i. u i \<in> borel_measurable M"
shows "u' \<in> borel_measurable M"
-proof -
+proof%unimportant -
have "\<And>x. x \<in> space M \<Longrightarrow> liminf (\<lambda>n. ereal (u n x)) = ereal (u' x)"
using u' by (simp add: lim_imp_Liminf)
moreover from u have "(\<lambda>x. liminf (\<lambda>n. ereal (u n x))) \<in> borel_measurable M"
@@ -1778,13 +1778,13 @@
ultimately show ?thesis by (simp cong: measurable_cong add: borel_measurable_ereal_iff)
qed
-lemma borel_measurable_LIMSEQ_metric:
+lemma%important borel_measurable_LIMSEQ_metric:
fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> 'b :: metric_space"
assumes [measurable]: "\<And>i. f i \<in> borel_measurable M"
assumes lim: "\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>i. f i x) \<longlonglongrightarrow> g x"
shows "g \<in> borel_measurable M"
unfolding borel_eq_closed
-proof (safe intro!: measurable_measure_of)
+proof%unimportant (safe intro!: measurable_measure_of)
fix A :: "'b set" assume "closed A"
have [measurable]: "(\<lambda>x. infdist (g x) A) \<in> borel_measurable M"
@@ -1809,13 +1809,13 @@
qed simp
qed auto
-lemma sets_Collect_Cauchy[measurable]:
+lemma%important sets_Collect_Cauchy[measurable]:
fixes f :: "nat \<Rightarrow> 'a => 'b::{metric_space, second_countable_topology}"
assumes f[measurable]: "\<And>i. f i \<in> borel_measurable M"
shows "{x\<in>space M. Cauchy (\<lambda>i. f i x)} \<in> sets M"
unfolding metric_Cauchy_iff2 using f by auto
-lemma borel_measurable_lim_metric[measurable (raw)]:
+lemma%unimportant borel_measurable_lim_metric[measurable (raw)]:
fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> 'b::{banach, second_countable_topology}"
assumes f[measurable]: "\<And>i. f i \<in> borel_measurable M"
shows "(\<lambda>x. lim (\<lambda>i. f i x)) \<in> borel_measurable M"
@@ -1837,17 +1837,17 @@
unfolding * by measurable
qed
-lemma borel_measurable_suminf[measurable (raw)]:
+lemma%unimportant borel_measurable_suminf[measurable (raw)]:
fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> 'b::{banach, second_countable_topology}"
assumes f[measurable]: "\<And>i. f i \<in> borel_measurable M"
shows "(\<lambda>x. suminf (\<lambda>i. f i x)) \<in> borel_measurable M"
unfolding suminf_def sums_def[abs_def] lim_def[symmetric] by simp
-lemma Collect_closed_imp_pred_borel: "closed {x. P x} \<Longrightarrow> Measurable.pred borel P"
+lemma%unimportant Collect_closed_imp_pred_borel: "closed {x. P x} \<Longrightarrow> Measurable.pred borel P"
by (simp add: pred_def)
(* Proof by Jeremy Avigad and Luke Serafin *)
-lemma isCont_borel_pred[measurable]:
+lemma%unimportant isCont_borel_pred[measurable]:
fixes f :: "'b::metric_space \<Rightarrow> 'a::metric_space"
shows "Measurable.pred borel (isCont f)"
proof (subst measurable_cong)
@@ -1898,21 +1898,21 @@
qed (intro pred_intros_countable closed_Collect_all closed_Collect_le open_Collect_less
Collect_closed_imp_pred_borel closed_Collect_imp open_Collect_conj continuous_intros)
-lemma isCont_borel:
+lemma%unimportant isCont_borel:
fixes f :: "'b::metric_space \<Rightarrow> 'a::metric_space"
shows "{x. isCont f x} \<in> sets borel"
by simp
-lemma is_real_interval:
+lemma%important is_real_interval:
assumes S: "is_interval S"
shows "\<exists>a b::real. S = {} \<or> S = UNIV \<or> S = {..<b} \<or> S = {..b} \<or> S = {a<..} \<or> S = {a..} \<or>
S = {a<..<b} \<or> S = {a<..b} \<or> S = {a..<b} \<or> S = {a..b}"
using S unfolding is_interval_1 by (blast intro: interval_cases)
-lemma real_interval_borel_measurable:
+lemma%important real_interval_borel_measurable:
assumes "is_interval (S::real set)"
shows "S \<in> sets borel"
-proof -
+proof%unimportant -
from assms is_real_interval have "\<exists>a b::real. S = {} \<or> S = UNIV \<or> S = {..<b} \<or> S = {..b} \<or>
S = {a<..} \<or> S = {a..} \<or> S = {a<..<b} \<or> S = {a<..b} \<or> S = {a..<b} \<or> S = {a..b}" by auto
then guess a ..
@@ -1924,7 +1924,7 @@
text \<open>The next lemmas hold in any second countable linorder (including ennreal or ereal for instance),
but in the current state they are restricted to reals.\<close>
-lemma borel_measurable_mono_on_fnc:
+lemma%important borel_measurable_mono_on_fnc:
fixes f :: "real \<Rightarrow> real" and A :: "real set"
assumes "mono_on f A"
shows "f \<in> borel_measurable (restrict_space borel A)"
@@ -1935,18 +1935,18 @@
intro!: borel_measurable_continuous_on_restrict intro: continuous_within_subset)
done
-lemma borel_measurable_piecewise_mono:
+lemma%unimportant borel_measurable_piecewise_mono:
fixes f::"real \<Rightarrow> real" and C::"real set set"
assumes "countable C" "\<And>c. c \<in> C \<Longrightarrow> c \<in> sets borel" "\<And>c. c \<in> C \<Longrightarrow> mono_on f c" "(\<Union>C) = UNIV"
shows "f \<in> borel_measurable borel"
-by (rule measurable_piecewise_restrict[of C], auto intro: borel_measurable_mono_on_fnc simp: assms)
+ by (rule measurable_piecewise_restrict[of C], auto intro: borel_measurable_mono_on_fnc simp: assms)
-lemma borel_measurable_mono:
+lemma%unimportant borel_measurable_mono:
fixes f :: "real \<Rightarrow> real"
shows "mono f \<Longrightarrow> f \<in> borel_measurable borel"
using borel_measurable_mono_on_fnc[of f UNIV] by (simp add: mono_def mono_on_def)
-lemma measurable_bdd_below_real[measurable (raw)]:
+lemma%unimportant measurable_bdd_below_real[measurable (raw)]:
fixes F :: "'a \<Rightarrow> 'i \<Rightarrow> real"
assumes [simp]: "countable I" and [measurable]: "\<And>i. i \<in> I \<Longrightarrow> F i \<in> M \<rightarrow>\<^sub>M borel"
shows "Measurable.pred M (\<lambda>x. bdd_below ((\<lambda>i. F i x)`I))"
@@ -1957,12 +1957,12 @@
using countable_int by measurable
qed
-lemma borel_measurable_cINF_real[measurable (raw)]:
+lemma%important borel_measurable_cINF_real[measurable (raw)]:
fixes F :: "_ \<Rightarrow> _ \<Rightarrow> real"
assumes [simp]: "countable I"
assumes F[measurable]: "\<And>i. i \<in> I \<Longrightarrow> F i \<in> borel_measurable M"
shows "(\<lambda>x. INF i:I. F i x) \<in> borel_measurable M"
-proof (rule measurable_piecewise_restrict)
+proof%unimportant (rule measurable_piecewise_restrict)
let ?\<Omega> = "{x\<in>space M. bdd_below ((\<lambda>i. F i x)`I)}"
show "countable {?\<Omega>, - ?\<Omega>}" "space M \<subseteq> \<Union>{?\<Omega>, - ?\<Omega>}" "\<And>X. X \<in> {?\<Omega>, - ?\<Omega>} \<Longrightarrow> X \<inter> space M \<in> sets M"
by auto
@@ -1982,7 +1982,7 @@
qed
qed
-lemma borel_Ici: "borel = sigma UNIV (range (\<lambda>x::real. {x ..}))"
+lemma%unimportant borel_Ici: "borel = sigma UNIV (range (\<lambda>x::real. {x ..}))"
proof (safe intro!: borel_eq_sigmaI1[OF borel_Iio])
fix x :: real
have eq: "{..<x} = space (sigma UNIV (range atLeast)) - {x ..}"
@@ -1991,7 +1991,7 @@
unfolding eq by (intro sets.compl_sets) auto
qed auto
-lemma borel_measurable_pred_less[measurable (raw)]:
+lemma%unimportant borel_measurable_pred_less[measurable (raw)]:
fixes f :: "'a \<Rightarrow> 'b::{second_countable_topology, linorder_topology}"
shows "f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> Measurable.pred M (\<lambda>w. f w < g w)"
unfolding Measurable.pred_def by (rule borel_measurable_less)
@@ -1999,19 +1999,19 @@
no_notation
eucl_less (infix "<e" 50)
-lemma borel_measurable_Max2[measurable (raw)]:
+lemma%important borel_measurable_Max2[measurable (raw)]:
fixes f::"_ \<Rightarrow> _ \<Rightarrow> 'a::{second_countable_topology, dense_linorder, linorder_topology}"
assumes "finite I"
and [measurable]: "\<And>i. f i \<in> borel_measurable M"
shows "(\<lambda>x. Max{f i x |i. i \<in> I}) \<in> borel_measurable M"
-by (simp add: borel_measurable_Max[OF assms(1), where ?f=f and ?M=M] Setcompr_eq_image)
+by%unimportant (simp add: borel_measurable_Max[OF assms(1), where ?f=f and ?M=M] Setcompr_eq_image)
-lemma measurable_compose_n [measurable (raw)]:
+lemma%unimportant measurable_compose_n [measurable (raw)]:
assumes "T \<in> measurable M M"
shows "(T^^n) \<in> measurable M M"
by (induction n, auto simp add: measurable_compose[OF _ assms])
-lemma measurable_real_imp_nat:
+lemma%unimportant measurable_real_imp_nat:
fixes f::"'a \<Rightarrow> nat"
assumes [measurable]: "(\<lambda>x. real(f x)) \<in> borel_measurable M"
shows "f \<in> measurable M (count_space UNIV)"
@@ -2023,7 +2023,7 @@
then show ?thesis using measurable_count_space_eq2_countable by blast
qed
-lemma measurable_equality_set [measurable]:
+lemma%unimportant measurable_equality_set [measurable]:
fixes f g::"_\<Rightarrow> 'a::{second_countable_topology, t2_space}"
assumes [measurable]: "f \<in> borel_measurable M" "g \<in> borel_measurable M"
shows "{x \<in> space M. f x = g x} \<in> sets M"
@@ -2038,7 +2038,7 @@
then show ?thesis unfolding A_def by simp
qed
-lemma measurable_inequality_set [measurable]:
+lemma%unimportant measurable_inequality_set [measurable]:
fixes f g::"_ \<Rightarrow> 'a::{second_countable_topology, linorder_topology}"
assumes [measurable]: "f \<in> borel_measurable M" "g \<in> borel_measurable M"
shows "{x \<in> space M. f x \<le> g x} \<in> sets M"
@@ -2066,7 +2066,7 @@
ultimately show "{x \<in> space M. f x > g x} \<in> sets M" using * by (metis (mono_tags, lifting) measurable_sets)
qed
-lemma measurable_limit [measurable]:
+lemma%unimportant measurable_limit [measurable]:
fixes f::"nat \<Rightarrow> 'a \<Rightarrow> 'b::first_countable_topology"
assumes [measurable]: "\<And>n::nat. f n \<in> borel_measurable M"
shows "Measurable.pred M (\<lambda>x. (\<lambda>n. f n x) \<longlonglongrightarrow> c)"
@@ -2106,11 +2106,11 @@
then show ?thesis by auto
qed
-lemma measurable_limit2 [measurable]:
+lemma%important measurable_limit2 [measurable]:
fixes u::"nat \<Rightarrow> 'a \<Rightarrow> real"
assumes [measurable]: "\<And>n. u n \<in> borel_measurable M" "v \<in> borel_measurable M"
shows "Measurable.pred M (\<lambda>x. (\<lambda>n. u n x) \<longlonglongrightarrow> v x)"
-proof -
+proof%unimportant -
define w where "w = (\<lambda>n x. u n x - v x)"
have [measurable]: "w n \<in> borel_measurable M" for n unfolding w_def by auto
have "((\<lambda>n. u n x) \<longlonglongrightarrow> v x) \<longleftrightarrow> ((\<lambda>n. w n x) \<longlonglongrightarrow> 0)" for x
@@ -2118,7 +2118,7 @@
then show ?thesis using measurable_limit by auto
qed
-lemma measurable_P_restriction [measurable (raw)]:
+lemma%unimportant measurable_P_restriction [measurable (raw)]:
assumes [measurable]: "Measurable.pred M P" "A \<in> sets M"
shows "{x \<in> A. P x} \<in> sets M"
proof -
@@ -2127,7 +2127,7 @@
then show ?thesis by auto
qed
-lemma measurable_sum_nat [measurable (raw)]:
+lemma%unimportant measurable_sum_nat [measurable (raw)]:
fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> nat"
assumes "\<And>i. i \<in> S \<Longrightarrow> f i \<in> measurable M (count_space UNIV)"
shows "(\<lambda>x. \<Sum>i\<in>S. f i x) \<in> measurable M (count_space UNIV)"
@@ -2137,7 +2137,7 @@
qed simp
-lemma measurable_abs_powr [measurable]:
+lemma%unimportant measurable_abs_powr [measurable]:
fixes p::real
assumes [measurable]: "f \<in> borel_measurable M"
shows "(\<lambda>x. \<bar>f x\<bar> powr p) \<in> borel_measurable M"
@@ -2145,7 +2145,7 @@
text \<open>The next one is a variation around \verb+measurable_restrict_space+.\<close>
-lemma measurable_restrict_space3:
+lemma%unimportant measurable_restrict_space3:
assumes "f \<in> measurable M N" and
"f \<in> A \<rightarrow> B"
shows "f \<in> measurable (restrict_space M A) (restrict_space N B)"
@@ -2157,12 +2157,12 @@
text \<open>The next one is a variation around \verb+measurable_piecewise_restrict+.\<close>
-lemma measurable_piecewise_restrict2:
+lemma%important measurable_piecewise_restrict2:
assumes [measurable]: "\<And>n. A n \<in> sets M"
and "space M = (\<Union>(n::nat). A n)"
"\<And>n. \<exists>h \<in> measurable M N. (\<forall>x \<in> A n. f x = h x)"
shows "f \<in> measurable M N"
-proof (rule measurableI)
+proof%unimportant (rule measurableI)
fix B assume [measurable]: "B \<in> sets N"
{
fix n::nat
--- a/src/HOL/Analysis/Caratheodory.thy Mon Aug 27 22:58:36 2018 +0200
+++ b/src/HOL/Analysis/Caratheodory.thy Tue Aug 28 13:28:39 2018 +0100
@@ -3,7 +3,7 @@
Author: Johannes Hölzl, TU München
*)
-section \<open>Caratheodory Extension Theorem\<close>
+section%important \<open>Caratheodory Extension Theorem\<close>
theory Caratheodory
imports Measure_Space
@@ -13,7 +13,7 @@
Originally from the Hurd/Coble measure theory development, translated by Lawrence Paulson.
\<close>
-lemma suminf_ennreal_2dimen:
+lemma%unimportant 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"
@@ -49,18 +49,18 @@
by (simp add: suminf_eq_SUP)
qed
-subsection \<open>Characterizations of Measures\<close>
+subsection%important \<open>Characterizations of Measures\<close>
-definition outer_measure_space where
+definition%important outer_measure_space where
"outer_measure_space M f \<longleftrightarrow> positive M f \<and> increasing M f \<and> countably_subadditive M f"
-subsubsection \<open>Lambda Systems\<close>
+subsubsection%important \<open>Lambda Systems\<close>
-definition lambda_system :: "'a set \<Rightarrow> 'a set set \<Rightarrow> ('a set \<Rightarrow> ennreal) \<Rightarrow> 'a set set"
+definition%important lambda_system :: "'a set \<Rightarrow> 'a set set \<Rightarrow> ('a set \<Rightarrow> ennreal) \<Rightarrow> 'a set set"
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 (in algebra) lambda_system_eq:
+lemma%unimportant (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 (in algebra) lambda_system_empty: "positive M f \<Longrightarrow> {} \<in> lambda_system \<Omega> M f"
+lemma%unimportant (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 lambda_system_sets: "x \<in> lambda_system \<Omega> M f \<Longrightarrow> x \<in> M"
+lemma%unimportant lambda_system_sets: "x \<in> lambda_system \<Omega> M f \<Longrightarrow> x \<in> M"
by (simp add: lambda_system_def)
-lemma (in algebra) lambda_system_Compl:
+lemma%unimportant (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 (in algebra) lambda_system_Int:
+lemma%unimportant (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 (in algebra) lambda_system_Un:
+lemma%unimportant (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 (in algebra) lambda_system_algebra:
+lemma%unimportant (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 (in algebra) lambda_system_strong_additive:
+lemma%unimportant (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 (in algebra) lambda_system_additive: "additive (lambda_system \<Omega> M f) f"
+lemma%unimportant (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 lambda_system_increasing: "increasing M f \<Longrightarrow> increasing (lambda_system \<Omega> M f) f"
+lemma%unimportant lambda_system_increasing: "increasing M f \<Longrightarrow> increasing (lambda_system \<Omega> M f) f"
by (simp add: increasing_def lambda_system_def)
-lemma lambda_system_positive: "positive M f \<Longrightarrow> positive (lambda_system \<Omega> M f) f"
+lemma%unimportant lambda_system_positive: "positive M f \<Longrightarrow> positive (lambda_system \<Omega> M f) f"
by (simp add: positive_def lambda_system_def)
-lemma (in algebra) lambda_system_strong_sum:
+lemma%unimportant (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 (in sigma_algebra) lambda_system_caratheodory:
+lemma%important (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 -
+proof%unimportant -
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 (in sigma_algebra) caratheodory_lemma:
+lemma%important (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 -
+proof%unimportant -
have pos: "positive M f"
by (metis oms outer_measure_space_def)
have alg: "algebra \<Omega> L"
@@ -297,11 +297,11 @@
using pos by (simp add: measure_space_def)
qed
-definition outer_measure :: "'a set set \<Rightarrow> ('a set \<Rightarrow> ennreal) \<Rightarrow> 'a set \<Rightarrow> ennreal" where
+definition%important outer_measure :: "'a set set \<Rightarrow> ('a set \<Rightarrow> ennreal) \<Rightarrow> 'a set \<Rightarrow> ennreal" where
"outer_measure M f X =
(INF A:{A. range A \<subseteq> M \<and> disjoint_family A \<and> X \<subseteq> (\<Union>i. A i)}. \<Sum>i. f (A i))"
-lemma (in ring_of_sets) outer_measure_agrees:
+lemma%unimportant (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 outer_measure_empty:
+lemma%unimportant 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 (in ring_of_sets) positive_outer_measure:
+lemma%unimportant (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 (in ring_of_sets) increasing_outer_measure: "increasing (Pow \<Omega>) (outer_measure M f)"
+lemma%unimportant (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 (in ring_of_sets) outer_measure_le:
+lemma%unimportant (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 (in ring_of_sets) outer_measure_close:
+lemma%unimportant (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 (in ring_of_sets) countably_subadditive_outer_measure:
+lemma%unimportant (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 (in ring_of_sets) outer_measure_space_outer_measure:
+lemma%unimportant (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 (in ring_of_sets) algebra_subset_lambda_system:
+lemma%unimportant (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 measure_down: "measure_space \<Omega> N \<mu> \<Longrightarrow> sigma_algebra \<Omega> M \<Longrightarrow> M \<subseteq> N \<Longrightarrow> measure_space \<Omega> M \<mu>"
+lemma%unimportant 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 \<open>Caratheodory's theorem\<close>
+subsection%important \<open>Caratheodory's theorem\<close>
-theorem (in ring_of_sets) caratheodory':
+theorem%important (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 -
+proof%unimportant -
have inc: "increasing M f"
by (metis additive_increasing ca countably_additive_additive posf)
let ?O = "outer_measure M f"
@@ -489,37 +489,37 @@
by (intro exI[of _ ?O]) auto
qed
-lemma (in ring_of_sets) caratheodory_empty_continuous:
+lemma%important (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 (intro caratheodory' empty_continuous_imp_countably_additive f)
+proof%unimportant (intro caratheodory' empty_continuous_imp_countably_additive f)
show "\<forall>A\<in>M. f A \<noteq> \<infinity>" using fin by auto
qed (rule cont)
-subsection \<open>Volumes\<close>
+subsection%important \<open>Volumes\<close>
-definition volume :: "'a set set \<Rightarrow> ('a set \<Rightarrow> ennreal) \<Rightarrow> bool" where
+definition%important volume :: "'a set set \<Rightarrow> ('a set \<Rightarrow> ennreal) \<Rightarrow> bool" where
"volume M f \<longleftrightarrow>
(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 volumeI:
+lemma%unimportant 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 volume_positive:
+lemma%unimportant volume_positive:
"volume M f \<Longrightarrow> a \<in> M \<Longrightarrow> 0 \<le> f a"
by (auto simp: volume_def)
-lemma volume_empty:
+lemma%unimportant volume_empty:
"volume M f \<Longrightarrow> f {} = 0"
by (auto simp: volume_def)
-lemma volume_finite_additive:
+lemma%unimportant 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 I A \<in> M"
shows "f (UNION I A) = (\<Sum>i\<in>I. f (A i))"
@@ -540,7 +540,7 @@
by simp
qed
-lemma (in ring_of_sets) volume_additiveI:
+lemma%unimportant (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 (in semiring_of_sets) extend_volume:
+lemma%important (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 -
+proof%unimportant -
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)
@@ -635,12 +635,12 @@
qed
qed
-subsubsection \<open>Caratheodory on semirings\<close>
+subsubsection%important \<open>Caratheodory on semirings\<close>
-theorem (in semiring_of_sets) caratheodory:
+theorem%important (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 -
+proof%unimportant -
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 extend_measure_caratheodory:
+lemma%important extend_measure_caratheodory:
fixes G :: "'i \<Rightarrow> 'a set"
assumes M: "M = extend_measure \<Omega> I G \<mu>"
assumes "i \<in> I"
@@ -827,7 +827,8 @@
assumes add: "\<And>A::nat \<Rightarrow> 'i. \<And>j. A \<in> UNIV \<rightarrow> I \<Longrightarrow> j \<in> I \<Longrightarrow> disjoint_family (G \<circ> A) \<Longrightarrow>
(\<Union>i. G (A i)) = G j \<Longrightarrow> (\<Sum>n. \<mu> (A n)) = \<mu> j"
shows "emeasure M (G i) = \<mu> i"
-proof -
+
+proof%unimportant -
interpret semiring_of_sets \<Omega> "G ` I"
by fact
have "\<forall>g\<in>G`I. \<exists>i\<in>I. g = G i"
@@ -860,7 +861,7 @@
qed fact
qed
-lemma extend_measure_caratheodory_pair:
+lemma%important 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"
@@ -872,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 -
+proof%unimportant -
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})"
--- a/src/HOL/Analysis/Cartesian_Euclidean_Space.thy Mon Aug 27 22:58:36 2018 +0200
+++ b/src/HOL/Analysis/Cartesian_Euclidean_Space.thy Tue Aug 28 13:28:39 2018 +0100
@@ -2,19 +2,19 @@
Some material by Jose Divasón, Tim Makarios and L C Paulson
*)
-section \<open>Instantiates the finite Cartesian product of Euclidean spaces as a Euclidean space\<close>
+section%important \<open>Instantiates the finite Cartesian product of Euclidean spaces as a Euclidean space\<close>
theory Cartesian_Euclidean_Space
imports Cartesian_Space Derivative
begin
-lemma subspace_special_hyperplane: "subspace {x. x $ k = 0}"
+lemma%unimportant subspace_special_hyperplane: "subspace {x. x $ k = 0}"
by (simp add: subspace_def)
-lemma sum_mult_product:
+lemma%important sum_mult_product:
"sum h {..<A * B :: nat} = (\<Sum>i\<in>{..<A}. \<Sum>j\<in>{..<B}. h (j + i * B))"
unfolding sum_nat_group[of h B A, unfolded atLeast0LessThan, symmetric]
-proof (rule sum.cong, simp, rule sum.reindex_cong)
+proof%unimportant (rule sum.cong, simp, rule sum.reindex_cong)
fix i
show "inj_on (\<lambda>j. j + i * B) {..<B}" by (auto intro!: inj_onI)
show "{i * B..<i * B + B} = (\<lambda>j. j + i * B) ` {..<B}"
@@ -25,32 +25,32 @@
qed simp
qed simp
-lemma interval_cbox_cart: "{a::real^'n..b} = cbox a b"
+lemma%unimportant interval_cbox_cart: "{a::real^'n..b} = cbox a b"
by (auto simp add: less_eq_vec_def mem_box Basis_vec_def inner_axis)
-lemma differentiable_vec:
+lemma%unimportant differentiable_vec:
fixes S :: "'a::euclidean_space set"
shows "vec differentiable_on S"
by (simp add: linear_linear bounded_linear_imp_differentiable_on)
-lemma continuous_vec [continuous_intros]:
+lemma%unimportant continuous_vec [continuous_intros]:
fixes x :: "'a::euclidean_space"
shows "isCont vec x"
apply (clarsimp simp add: continuous_def LIM_def dist_vec_def L2_set_def)
apply (rule_tac x="r / sqrt (real CARD('b))" in exI)
by (simp add: mult.commute pos_less_divide_eq real_sqrt_mult)
-lemma box_vec_eq_empty [simp]:
+lemma%unimportant box_vec_eq_empty [simp]:
shows "cbox (vec a) (vec b) = {} \<longleftrightarrow> cbox a b = {}"
"box (vec a) (vec b) = {} \<longleftrightarrow> box a b = {}"
by (auto simp: Basis_vec_def mem_box box_eq_empty inner_axis)
-subsection\<open>Closures and interiors of halfspaces\<close>
+subsection%important\<open>Closures and interiors of halfspaces\<close>
-lemma interior_halfspace_le [simp]:
+lemma%important interior_halfspace_le [simp]:
assumes "a \<noteq> 0"
shows "interior {x. a \<bullet> x \<le> b} = {x. a \<bullet> x < b}"
-proof -
+proof%unimportant -
have *: "a \<bullet> x < b" if x: "x \<in> S" and S: "S \<subseteq> {x. a \<bullet> x \<le> b}" and "open S" for S x
proof -
obtain e where "e>0" and e: "cball x e \<subseteq> S"
@@ -70,15 +70,15 @@
by (rule interior_unique) (auto simp: open_halfspace_lt *)
qed
-lemma interior_halfspace_ge [simp]:
+lemma%unimportant interior_halfspace_ge [simp]:
"a \<noteq> 0 \<Longrightarrow> interior {x. a \<bullet> x \<ge> b} = {x. a \<bullet> x > b}"
using interior_halfspace_le [of "-a" "-b"] by simp
-lemma interior_halfspace_component_le [simp]:
+lemma%important interior_halfspace_component_le [simp]:
"interior {x. x$k \<le> a} = {x :: (real^'n). x$k < a}" (is "?LE")
and interior_halfspace_component_ge [simp]:
"interior {x. x$k \<ge> a} = {x :: (real^'n). x$k > a}" (is "?GE")
-proof -
+proof%unimportant -
have "axis k (1::real) \<noteq> 0"
by (simp add: axis_def vec_eq_iff)
moreover have "axis k (1::real) \<bullet> x = x$k" for x
@@ -88,7 +88,7 @@
interior_halfspace_ge [of "axis k (1::real)" a] by auto
qed
-lemma closure_halfspace_lt [simp]:
+lemma%unimportant closure_halfspace_lt [simp]:
assumes "a \<noteq> 0"
shows "closure {x. a \<bullet> x < b} = {x. a \<bullet> x \<le> b}"
proof -
@@ -99,15 +99,15 @@
by (force simp: closure_interior)
qed
-lemma closure_halfspace_gt [simp]:
+lemma%unimportant closure_halfspace_gt [simp]:
"a \<noteq> 0 \<Longrightarrow> closure {x. a \<bullet> x > b} = {x. a \<bullet> x \<ge> b}"
using closure_halfspace_lt [of "-a" "-b"] by simp
-lemma closure_halfspace_component_lt [simp]:
+lemma%important closure_halfspace_component_lt [simp]:
"closure {x. x$k < a} = {x :: (real^'n). x$k \<le> a}" (is "?LE")
and closure_halfspace_component_gt [simp]:
"closure {x. x$k > a} = {x :: (real^'n). x$k \<ge> a}" (is "?GE")
-proof -
+proof%unimportant -
have "axis k (1::real) \<noteq> 0"
by (simp add: axis_def vec_eq_iff)
moreover have "axis k (1::real) \<bullet> x = x$k" for x
@@ -117,17 +117,17 @@
closure_halfspace_gt [of "axis k (1::real)" a] by auto
qed
-lemma interior_hyperplane [simp]:
+lemma%unimportant interior_hyperplane [simp]:
assumes "a \<noteq> 0"
shows "interior {x. a \<bullet> x = b} = {}"
-proof -
+proof%unimportant -
have [simp]: "{x. a \<bullet> x = b} = {x. a \<bullet> x \<le> b} \<inter> {x. a \<bullet> x \<ge> b}"
by (force simp:)
then show ?thesis
by (auto simp: assms)
qed
-lemma frontier_halfspace_le:
+lemma%unimportant frontier_halfspace_le:
assumes "a \<noteq> 0 \<or> b \<noteq> 0"
shows "frontier {x. a \<bullet> x \<le> b} = {x. a \<bullet> x = b}"
proof (cases "a = 0")
@@ -137,7 +137,7 @@
by (force simp: frontier_def closed_halfspace_le)
qed
-lemma frontier_halfspace_ge:
+lemma%unimportant frontier_halfspace_ge:
assumes "a \<noteq> 0 \<or> b \<noteq> 0"
shows "frontier {x. a \<bullet> x \<ge> b} = {x. a \<bullet> x = b}"
proof (cases "a = 0")
@@ -147,7 +147,7 @@
by (force simp: frontier_def closed_halfspace_ge)
qed
-lemma frontier_halfspace_lt:
+lemma%unimportant frontier_halfspace_lt:
assumes "a \<noteq> 0 \<or> b \<noteq> 0"
shows "frontier {x. a \<bullet> x < b} = {x. a \<bullet> x = b}"
proof (cases "a = 0")
@@ -157,19 +157,19 @@
by (force simp: frontier_def interior_open open_halfspace_lt)
qed
-lemma frontier_halfspace_gt:
+lemma%important frontier_halfspace_gt:
assumes "a \<noteq> 0 \<or> b \<noteq> 0"
shows "frontier {x. a \<bullet> x > b} = {x. a \<bullet> x = b}"
-proof (cases "a = 0")
+proof%unimportant (cases "a = 0")
case True with assms show ?thesis by simp
next
case False then show ?thesis
by (force simp: frontier_def interior_open open_halfspace_gt)
qed
-lemma interior_standard_hyperplane:
+lemma%important interior_standard_hyperplane:
"interior {x :: (real^'n). x$k = a} = {}"
-proof -
+proof%unimportant -
have "axis k (1::real) \<noteq> 0"
by (simp add: axis_def vec_eq_iff)
moreover have "axis k (1::real) \<bullet> x = x$k" for x
@@ -179,20 +179,20 @@
by force
qed
-lemma matrix_mult_transpose_dot_column:
+lemma%unimportant matrix_mult_transpose_dot_column:
shows "transpose A ** A = (\<chi> i j. inner (column i A) (column j A))"
by (simp add: matrix_matrix_mult_def vec_eq_iff transpose_def column_def inner_vec_def)
-lemma matrix_mult_transpose_dot_row:
+lemma%unimportant matrix_mult_transpose_dot_row:
shows "A ** transpose A = (\<chi> i j. inner (row i A) (row j A))"
by (simp add: matrix_matrix_mult_def vec_eq_iff transpose_def row_def inner_vec_def)
text\<open>Two sometimes fruitful ways of looking at matrix-vector multiplication.\<close>
-lemma matrix_mult_dot: "A *v x = (\<chi> i. inner (A$i) x)"
+lemma%important matrix_mult_dot: "A *v x = (\<chi> i. inner (A$i) x)"
by (simp add: matrix_vector_mult_def inner_vec_def)
-lemma adjoint_matrix: "adjoint(\<lambda>x. (A::real^'n^'m) *v x) = (\<lambda>x. transpose A *v x)"
+lemma%unimportant adjoint_matrix: "adjoint(\<lambda>x. (A::real^'n^'m) *v x) = (\<lambda>x. transpose A *v x)"
apply (rule adjoint_unique)
apply (simp add: transpose_def inner_vec_def matrix_vector_mult_def
sum_distrib_right sum_distrib_left)
@@ -200,9 +200,9 @@
apply (simp add: ac_simps)
done
-lemma matrix_adjoint: assumes lf: "linear (f :: real^'n \<Rightarrow> real ^'m)"
+lemma%important matrix_adjoint: assumes lf: "linear (f :: real^'n \<Rightarrow> real ^'m)"
shows "matrix(adjoint f) = transpose(matrix f)"
-proof -
+proof%unimportant -
have "matrix(adjoint f) = matrix(adjoint (( *v) (matrix f)))"
by (simp add: lf)
also have "\<dots> = transpose(matrix f)"
@@ -212,17 +212,17 @@
finally show ?thesis .
qed
-lemma matrix_vector_mul_bounded_linear[intro, simp]: "bounded_linear (( *v) A)" for A :: "'a::{euclidean_space,real_algebra_1}^'n^'m"
+lemma%unimportant matrix_vector_mul_bounded_linear[intro, simp]: "bounded_linear (( *v) A)" for A :: "'a::{euclidean_space,real_algebra_1}^'n^'m"
using matrix_vector_mul_linear[of A]
by (simp add: linear_conv_bounded_linear linear_matrix_vector_mul_eq)
-lemma
+lemma%unimportant (* FIX ME needs name*)
fixes A :: "'a::{euclidean_space,real_algebra_1}^'n^'m"
shows matrix_vector_mult_linear_continuous_at [continuous_intros]: "isCont (( *v) A) z"
and matrix_vector_mult_linear_continuous_on [continuous_intros]: "continuous_on S (( *v) A)"
by (simp_all add: linear_continuous_at linear_continuous_on)
-lemma scalar_invertible:
+lemma%unimportant scalar_invertible:
fixes A :: "('a::real_algebra_1)^'m^'n"
assumes "k \<noteq> 0" and "invertible A"
shows "invertible (k *\<^sub>R A)"
@@ -236,50 +236,50 @@
unfolding invertible_def by auto
qed
-lemma scalar_invertible_iff:
+lemma%unimportant scalar_invertible_iff:
fixes A :: "('a::real_algebra_1)^'m^'n"
assumes "k \<noteq> 0" and "invertible A"
shows "invertible (k *\<^sub>R A) \<longleftrightarrow> k \<noteq> 0 \<and> invertible A"
by (simp add: assms scalar_invertible)
-lemma vector_transpose_matrix [simp]: "x v* transpose A = A *v x"
+lemma%unimportant vector_transpose_matrix [simp]: "x v* transpose A = A *v x"
unfolding transpose_def vector_matrix_mult_def matrix_vector_mult_def
by simp
-lemma transpose_matrix_vector [simp]: "transpose A *v x = x v* A"
+lemma%unimportant transpose_matrix_vector [simp]: "transpose A *v x = x v* A"
unfolding transpose_def vector_matrix_mult_def matrix_vector_mult_def
by simp
-lemma vector_scalar_commute:
+lemma%unimportant vector_scalar_commute:
fixes A :: "'a::{field}^'m^'n"
shows "A *v (c *s x) = c *s (A *v x)"
by (simp add: vector_scalar_mult_def matrix_vector_mult_def mult_ac sum_distrib_left)
-lemma scalar_vector_matrix_assoc:
+lemma%unimportant scalar_vector_matrix_assoc:
fixes k :: "'a::{field}" and x :: "'a::{field}^'n" and A :: "'a^'m^'n"
shows "(k *s x) v* A = k *s (x v* A)"
by (metis transpose_matrix_vector vector_scalar_commute)
-lemma vector_matrix_mult_0 [simp]: "0 v* A = 0"
+lemma%unimportant vector_matrix_mult_0 [simp]: "0 v* A = 0"
unfolding vector_matrix_mult_def by (simp add: zero_vec_def)
-lemma vector_matrix_mult_0_right [simp]: "x v* 0 = 0"
+lemma%unimportant vector_matrix_mult_0_right [simp]: "x v* 0 = 0"
unfolding vector_matrix_mult_def by (simp add: zero_vec_def)
-lemma vector_matrix_mul_rid [simp]:
+lemma%unimportant vector_matrix_mul_rid [simp]:
fixes v :: "('a::semiring_1)^'n"
shows "v v* mat 1 = v"
by (metis matrix_vector_mul_lid transpose_mat vector_transpose_matrix)
-lemma scaleR_vector_matrix_assoc:
+lemma%unimportant scaleR_vector_matrix_assoc:
fixes k :: real and x :: "real^'n" and A :: "real^'m^'n"
shows "(k *\<^sub>R x) v* A = k *\<^sub>R (x v* A)"
by (metis matrix_vector_mult_scaleR transpose_matrix_vector)
-lemma vector_scaleR_matrix_ac:
+lemma%important vector_scaleR_matrix_ac:
fixes k :: real and x :: "real^'n" and A :: "real^'m^'n"
shows "x v* (k *\<^sub>R A) = k *\<^sub>R (x v* A)"
-proof -
+proof%unimportant -
have "x v* (k *\<^sub>R A) = (k *\<^sub>R x) v* A"
unfolding vector_matrix_mult_def
by (simp add: algebra_simps)
@@ -289,12 +289,12 @@
qed
-subsection\<open>Some bounds on components etc. relative to operator norm\<close>
+subsection%important\<open>Some bounds on components etc. relative to operator norm\<close>
-lemma norm_column_le_onorm:
+lemma%important norm_column_le_onorm:
fixes A :: "real^'n^'m"
shows "norm(column i A) \<le> onorm(( *v) A)"
-proof -
+proof%unimportant -
have "norm (\<chi> j. A $ j $ i) \<le> norm (A *v axis i 1)"
by (simp add: matrix_mult_dot cart_eq_inner_axis)
also have "\<dots> \<le> onorm (( *v) A)"
@@ -304,10 +304,10 @@
unfolding column_def .
qed
-lemma matrix_component_le_onorm:
+lemma%important matrix_component_le_onorm:
fixes A :: "real^'n^'m"
shows "\<bar>A $ i $ j\<bar> \<le> onorm(( *v) A)"
-proof -
+proof%unimportant -
have "\<bar>A $ i $ j\<bar> \<le> norm (\<chi> n. (A $ n $ j))"
by (metis (full_types, lifting) component_le_norm_cart vec_lambda_beta)
also have "\<dots> \<le> onorm (( *v) A)"
@@ -315,15 +315,15 @@
finally show ?thesis .
qed
-lemma component_le_onorm:
+lemma%unimportant component_le_onorm:
fixes f :: "real^'m \<Rightarrow> real^'n"
shows "linear f \<Longrightarrow> \<bar>matrix f $ i $ j\<bar> \<le> onorm f"
by (metis linear_matrix_vector_mul_eq matrix_component_le_onorm matrix_vector_mul)
-lemma onorm_le_matrix_component_sum:
+lemma%important onorm_le_matrix_component_sum:
fixes A :: "real^'n^'m"
shows "onorm(( *v) A) \<le> (\<Sum>i\<in>UNIV. \<Sum>j\<in>UNIV. \<bar>A $ i $ j\<bar>)"
-proof (rule onorm_le)
+proof%unimportant (rule onorm_le)
fix x
have "norm (A *v x) \<le> (\<Sum>i\<in>UNIV. \<bar>(A *v x) $ i\<bar>)"
by (rule norm_le_l1_cart)
@@ -342,11 +342,11 @@
by (simp add: sum_distrib_right)
qed
-lemma onorm_le_matrix_component:
+lemma%important onorm_le_matrix_component:
fixes A :: "real^'n^'m"
assumes "\<And>i j. abs(A$i$j) \<le> B"
shows "onorm(( *v) A) \<le> real (CARD('m)) * real (CARD('n)) * B"
-proof (rule onorm_le)
+proof%unimportant (rule onorm_le)
fix x :: "real^'n::_"
have "norm (A *v x) \<le> (\<Sum>i\<in>UNIV. \<bar>(A *v x) $ i\<bar>)"
by (rule norm_le_l1_cart)
@@ -366,11 +366,11 @@
finally show "norm (A *v x) \<le> CARD('m) * real (CARD('n)) * B * norm x" .
qed
-subsection \<open>lambda skolemization on cartesian products\<close>
+subsection%important \<open>lambda skolemization on cartesian products\<close>
-lemma lambda_skolem: "(\<forall>i. \<exists>x. P i x) \<longleftrightarrow>
+lemma%important lambda_skolem: "(\<forall>i. \<exists>x. P i x) \<longleftrightarrow>
(\<exists>x::'a ^ 'n. \<forall>i. P i (x $ i))" (is "?lhs \<longleftrightarrow> ?rhs")
-proof -
+proof%unimportant -
let ?S = "(UNIV :: 'n set)"
{ assume H: "?rhs"
then have ?lhs by auto }
@@ -387,16 +387,16 @@
ultimately show ?thesis by metis
qed
-lemma rational_approximation:
+lemma%unimportant rational_approximation:
assumes "e > 0"
obtains r::real where "r \<in> \<rat>" "\<bar>r - x\<bar> < e"
using Rats_dense_in_real [of "x - e/2" "x + e/2"] assms by auto
-lemma matrix_rational_approximation:
+lemma%important matrix_rational_approximation:
fixes A :: "real^'n^'m"
assumes "e > 0"
obtains B where "\<And>i j. B$i$j \<in> \<rat>" "onorm(\<lambda>x. (A - B) *v x) < e"
-proof -
+proof%unimportant -
have "\<forall>i j. \<exists>q \<in> \<rat>. \<bar>q - A $ i $ j\<bar> < e / (2 * CARD('m) * CARD('n))"
using assms by (force intro: rational_approximation [of "e / (2 * CARD('m) * CARD('n))"])
then obtain B where B: "\<And>i j. B$i$j \<in> \<rat>" and Bclo: "\<And>i j. \<bar>B$i$j - A $ i $ j\<bar> < e / (2 * CARD('m) * CARD('n))"
@@ -413,7 +413,7 @@
qed (use B in auto)
qed
-lemma vector_sub_project_orthogonal_cart: "(b::real^'n) \<bullet> (x - ((b \<bullet> x) / (b \<bullet> b)) *s b) = 0"
+lemma%unimportant vector_sub_project_orthogonal_cart: "(b::real^'n) \<bullet> (x - ((b \<bullet> x) / (b \<bullet> b)) *s b) = 0"
unfolding inner_simps scalar_mult_eq_scaleR by auto
@@ -422,51 +422,51 @@
text \<open>Considering an n-element vector as an n-by-1 or 1-by-n matrix.\<close>
-definition "rowvector v = (\<chi> i j. (v$j))"
+definition%unimportant "rowvector v = (\<chi> i j. (v$j))"
-definition "columnvector v = (\<chi> i j. (v$i))"
+definition%unimportant "columnvector v = (\<chi> i j. (v$i))"
-lemma transpose_columnvector: "transpose(columnvector v) = rowvector v"
+lemma%unimportant transpose_columnvector: "transpose(columnvector v) = rowvector v"
by (simp add: transpose_def rowvector_def columnvector_def vec_eq_iff)
-lemma transpose_rowvector: "transpose(rowvector v) = columnvector v"
+lemma%unimportant transpose_rowvector: "transpose(rowvector v) = columnvector v"
by (simp add: transpose_def columnvector_def rowvector_def vec_eq_iff)
-lemma dot_rowvector_columnvector: "columnvector (A *v v) = A ** columnvector v"
+lemma%unimportant dot_rowvector_columnvector: "columnvector (A *v v) = A ** columnvector v"
by (vector columnvector_def matrix_matrix_mult_def matrix_vector_mult_def)
-lemma dot_matrix_product:
+lemma%unimportant dot_matrix_product:
"(x::real^'n) \<bullet> y = (((rowvector x ::real^'n^1) ** (columnvector y :: real^1^'n))$1)$1"
by (vector matrix_matrix_mult_def rowvector_def columnvector_def inner_vec_def)
-lemma dot_matrix_vector_mul:
+lemma%unimportant dot_matrix_vector_mul:
fixes A B :: "real ^'n ^'n" and x y :: "real ^'n"
shows "(A *v x) \<bullet> (B *v y) =
(((rowvector x :: real^'n^1) ** ((transpose A ** B) ** (columnvector y :: real ^1^'n)))$1)$1"
unfolding dot_matrix_product transpose_columnvector[symmetric]
dot_rowvector_columnvector matrix_transpose_mul matrix_mul_assoc ..
-lemma infnorm_cart:"infnorm (x::real^'n) = Sup {\<bar>x$i\<bar> |i. i\<in>UNIV}"
+lemma%unimportant infnorm_cart:"infnorm (x::real^'n) = Sup {\<bar>x$i\<bar> |i. i\<in>UNIV}"
by (simp add: infnorm_def inner_axis Basis_vec_def) (metis (lifting) inner_axis real_inner_1_right)
-lemma component_le_infnorm_cart: "\<bar>x$i\<bar> \<le> infnorm (x::real^'n)"
+lemma%unimportant component_le_infnorm_cart: "\<bar>x$i\<bar> \<le> infnorm (x::real^'n)"
using Basis_le_infnorm[of "axis i 1" x]
by (simp add: Basis_vec_def axis_eq_axis inner_axis)
-lemma continuous_component[continuous_intros]: "continuous F f \<Longrightarrow> continuous F (\<lambda>x. f x $ i)"
+lemma%unimportant continuous_component[continuous_intros]: "continuous F f \<Longrightarrow> continuous F (\<lambda>x. f x $ i)"
unfolding continuous_def by (rule tendsto_vec_nth)
-lemma continuous_on_component[continuous_intros]: "continuous_on s f \<Longrightarrow> continuous_on s (\<lambda>x. f x $ i)"
+lemma%unimportant continuous_on_component[continuous_intros]: "continuous_on s f \<Longrightarrow> continuous_on s (\<lambda>x. f x $ i)"
unfolding continuous_on_def by (fast intro: tendsto_vec_nth)
-lemma continuous_on_vec_lambda[continuous_intros]:
+lemma%unimportant continuous_on_vec_lambda[continuous_intros]:
"(\<And>i. continuous_on S (f i)) \<Longrightarrow> continuous_on S (\<lambda>x. \<chi> i. f i x)"
unfolding continuous_on_def by (auto intro: tendsto_vec_lambda)
-lemma closed_positive_orthant: "closed {x::real^'n. \<forall>i. 0 \<le>x$i}"
+lemma%unimportant closed_positive_orthant: "closed {x::real^'n. \<forall>i. 0 \<le>x$i}"
by (simp add: Collect_all_eq closed_INT closed_Collect_le continuous_on_const continuous_on_id continuous_on_component)
-lemma bounded_component_cart: "bounded s \<Longrightarrow> bounded ((\<lambda>x. x $ i) ` s)"
+lemma%unimportant bounded_component_cart: "bounded s \<Longrightarrow> bounded ((\<lambda>x. x $ i) ` s)"
unfolding bounded_def
apply clarify
apply (rule_tac x="x $ i" in exI)
@@ -475,13 +475,13 @@
apply (rule order_trans [OF dist_vec_nth_le], simp)
done
-lemma compact_lemma_cart:
+lemma%important compact_lemma_cart:
fixes f :: "nat \<Rightarrow> 'a::heine_borel ^ 'n"
assumes f: "bounded (range f)"
shows "\<exists>l r. strict_mono r \<and>
(\<forall>e>0. eventually (\<lambda>n. \<forall>i\<in>d. dist (f (r n) $ i) (l $ i) < e) sequentially)"
(is "?th d")
-proof -
+proof%unimportant -
have "\<forall>d' \<subseteq> d. ?th d'"
by (rule compact_lemma_general[where unproj=vec_lambda])
(auto intro!: f bounded_component_cart simp: vec_lambda_eta)
@@ -517,19 +517,19 @@
with r show "\<exists>l r. strict_mono r \<and> ((f \<circ> r) \<longlongrightarrow> l) sequentially" by auto
qed
-lemma interval_cart:
+lemma%unimportant interval_cart:
fixes a :: "real^'n"
shows "box a b = {x::real^'n. \<forall>i. a$i < x$i \<and> x$i < b$i}"
and "cbox a b = {x::real^'n. \<forall>i. a$i \<le> x$i \<and> x$i \<le> b$i}"
by (auto simp add: set_eq_iff less_vec_def less_eq_vec_def mem_box Basis_vec_def inner_axis)
-lemma mem_box_cart:
+lemma%unimportant mem_box_cart:
fixes a :: "real^'n"
shows "x \<in> box a b \<longleftrightarrow> (\<forall>i. a$i < x$i \<and> x$i < b$i)"
and "x \<in> cbox a b \<longleftrightarrow> (\<forall>i. a$i \<le> x$i \<and> x$i \<le> b$i)"
using interval_cart[of a b] by (auto simp add: set_eq_iff less_vec_def less_eq_vec_def)
-lemma interval_eq_empty_cart:
+lemma%unimportant interval_eq_empty_cart:
fixes a :: "real^'n"
shows "(box a b = {} \<longleftrightarrow> (\<exists>i. b$i \<le> a$i))" (is ?th1)
and "(cbox a b = {} \<longleftrightarrow> (\<exists>i. b$i < a$i))" (is ?th2)
@@ -565,14 +565,14 @@
ultimately show ?th2 by blast
qed
-lemma interval_ne_empty_cart:
+lemma%unimportant interval_ne_empty_cart:
fixes a :: "real^'n"
shows "cbox a b \<noteq> {} \<longleftrightarrow> (\<forall>i. a$i \<le> b$i)"
and "box a b \<noteq> {} \<longleftrightarrow> (\<forall>i. a$i < b$i)"
unfolding interval_eq_empty_cart[of a b] by (auto simp add: not_less not_le)
(* BH: Why doesn't just "auto" work here? *)
-lemma subset_interval_imp_cart:
+lemma%unimportant subset_interval_imp_cart:
fixes a :: "real^'n"
shows "(\<forall>i. a$i \<le> c$i \<and> d$i \<le> b$i) \<Longrightarrow> cbox c d \<subseteq> cbox a b"
and "(\<forall>i. a$i < c$i \<and> d$i < b$i) \<Longrightarrow> cbox c d \<subseteq> box a b"
@@ -581,13 +581,13 @@
unfolding subset_eq[unfolded Ball_def] unfolding mem_box_cart
by (auto intro: order_trans less_le_trans le_less_trans less_imp_le) (* BH: Why doesn't just "auto" work here? *)
-lemma interval_sing:
+lemma%unimportant interval_sing:
fixes a :: "'a::linorder^'n"
shows "{a .. a} = {a} \<and> {a<..<a} = {}"
apply (auto simp add: set_eq_iff less_vec_def less_eq_vec_def vec_eq_iff)
done
-lemma subset_interval_cart:
+lemma%unimportant subset_interval_cart:
fixes a :: "real^'n"
shows "cbox c d \<subseteq> cbox a b \<longleftrightarrow> (\<forall>i. c$i \<le> d$i) --> (\<forall>i. a$i \<le> c$i \<and> d$i \<le> b$i)" (is ?th1)
and "cbox c d \<subseteq> box a b \<longleftrightarrow> (\<forall>i. c$i \<le> d$i) --> (\<forall>i. a$i < c$i \<and> d$i < b$i)" (is ?th2)
@@ -595,7 +595,7 @@
and "box c d \<subseteq> box a b \<longleftrightarrow> (\<forall>i. c$i < d$i) --> (\<forall>i. a$i \<le> c$i \<and> d$i \<le> b$i)" (is ?th4)
using subset_box[of c d a b] by (simp_all add: Basis_vec_def inner_axis)
-lemma disjoint_interval_cart:
+lemma%unimportant disjoint_interval_cart:
fixes a::"real^'n"
shows "cbox a b \<inter> cbox c d = {} \<longleftrightarrow> (\<exists>i. (b$i < a$i \<or> d$i < c$i \<or> b$i < c$i \<or> d$i < a$i))" (is ?th1)
and "cbox a b \<inter> box c d = {} \<longleftrightarrow> (\<exists>i. (b$i < a$i \<or> d$i \<le> c$i \<or> b$i \<le> c$i \<or> d$i \<le> a$i))" (is ?th2)
@@ -603,53 +603,53 @@
and "box a b \<inter> box c d = {} \<longleftrightarrow> (\<exists>i. (b$i \<le> a$i \<or> d$i \<le> c$i \<or> b$i \<le> c$i \<or> d$i \<le> a$i))" (is ?th4)
using disjoint_interval[of a b c d] by (simp_all add: Basis_vec_def inner_axis)
-lemma Int_interval_cart:
+lemma%unimportant Int_interval_cart:
fixes a :: "real^'n"
shows "cbox a b \<inter> cbox c d = {(\<chi> i. max (a$i) (c$i)) .. (\<chi> i. min (b$i) (d$i))}"
unfolding Int_interval
by (auto simp: mem_box less_eq_vec_def)
(auto simp: Basis_vec_def inner_axis)
-lemma closed_interval_left_cart:
+lemma%unimportant closed_interval_left_cart:
fixes b :: "real^'n"
shows "closed {x::real^'n. \<forall>i. x$i \<le> b$i}"
by (simp add: Collect_all_eq closed_INT closed_Collect_le continuous_on_const continuous_on_id continuous_on_component)
-lemma closed_interval_right_cart:
+lemma%unimportant closed_interval_right_cart:
fixes a::"real^'n"
shows "closed {x::real^'n. \<forall>i. a$i \<le> x$i}"
by (simp add: Collect_all_eq closed_INT closed_Collect_le continuous_on_const continuous_on_id continuous_on_component)
-lemma is_interval_cart:
+lemma%unimportant is_interval_cart:
"is_interval (s::(real^'n) set) \<longleftrightarrow>
(\<forall>a\<in>s. \<forall>b\<in>s. \<forall>x. (\<forall>i. ((a$i \<le> x$i \<and> x$i \<le> b$i) \<or> (b$i \<le> x$i \<and> x$i \<le> a$i))) \<longrightarrow> x \<in> s)"
by (simp add: is_interval_def Ball_def Basis_vec_def inner_axis imp_ex)
-lemma closed_halfspace_component_le_cart: "closed {x::real^'n. x$i \<le> a}"
+lemma%unimportant closed_halfspace_component_le_cart: "closed {x::real^'n. x$i \<le> a}"
by (simp add: closed_Collect_le continuous_on_const continuous_on_id continuous_on_component)
-lemma closed_halfspace_component_ge_cart: "closed {x::real^'n. x$i \<ge> a}"
+lemma%unimportant closed_halfspace_component_ge_cart: "closed {x::real^'n. x$i \<ge> a}"
by (simp add: closed_Collect_le continuous_on_const continuous_on_id continuous_on_component)
-lemma open_halfspace_component_lt_cart: "open {x::real^'n. x$i < a}"
+lemma%unimportant open_halfspace_component_lt_cart: "open {x::real^'n. x$i < a}"
by (simp add: open_Collect_less continuous_on_const continuous_on_id continuous_on_component)
-lemma open_halfspace_component_gt_cart: "open {x::real^'n. x$i > a}"
+lemma%unimportant open_halfspace_component_gt_cart: "open {x::real^'n. x$i > a}"
by (simp add: open_Collect_less continuous_on_const continuous_on_id continuous_on_component)
-lemma Lim_component_le_cart:
+lemma%unimportant Lim_component_le_cart:
fixes f :: "'a \<Rightarrow> real^'n"
assumes "(f \<longlongrightarrow> l) net" "\<not> (trivial_limit net)" "eventually (\<lambda>x. f x $i \<le> b) net"
shows "l$i \<le> b"
by (rule tendsto_le[OF assms(2) tendsto_const tendsto_vec_nth, OF assms(1, 3)])
-lemma Lim_component_ge_cart:
+lemma%unimportant Lim_component_ge_cart:
fixes f :: "'a \<Rightarrow> real^'n"
assumes "(f \<longlongrightarrow> l) net" "\<not> (trivial_limit net)" "eventually (\<lambda>x. b \<le> (f x)$i) net"
shows "b \<le> l$i"
by (rule tendsto_le[OF assms(2) tendsto_vec_nth tendsto_const, OF assms(1, 3)])
-lemma Lim_component_eq_cart:
+lemma%unimportant Lim_component_eq_cart:
fixes f :: "'a \<Rightarrow> real^'n"
assumes net: "(f \<longlongrightarrow> l) net" "~(trivial_limit net)" and ev:"eventually (\<lambda>x. f(x)$i = b) net"
shows "l$i = b"
@@ -657,18 +657,18 @@
Lim_component_ge_cart[OF net, of b i] and
Lim_component_le_cart[OF net, of i b] by auto
-lemma connected_ivt_component_cart:
+lemma%unimportant connected_ivt_component_cart:
fixes x :: "real^'n"
shows "connected s \<Longrightarrow> x \<in> s \<Longrightarrow> y \<in> s \<Longrightarrow> x$k \<le> a \<Longrightarrow> a \<le> y$k \<Longrightarrow> (\<exists>z\<in>s. z$k = a)"
using connected_ivt_hyperplane[of s x y "axis k 1" a]
by (auto simp add: inner_axis inner_commute)
-lemma subspace_substandard_cart: "vec.subspace {x. (\<forall>i. P i \<longrightarrow> x$i = 0)}"
+lemma%unimportant subspace_substandard_cart: "vec.subspace {x. (\<forall>i. P i \<longrightarrow> x$i = 0)}"
unfolding vec.subspace_def by auto
-lemma closed_substandard_cart:
+lemma%important closed_substandard_cart:
"closed {x::'a::real_normed_vector ^ 'n. \<forall>i. P i \<longrightarrow> x$i = 0}"
-proof -
+proof%unimportant -
{ fix i::'n
have "closed {x::'a ^ 'n. P i \<longrightarrow> x$i = 0}"
by (cases "P i") (simp_all add: closed_Collect_eq continuous_on_const continuous_on_id continuous_on_component) }
@@ -676,9 +676,9 @@
unfolding Collect_all_eq by (simp add: closed_INT)
qed
-lemma dim_substandard_cart: "vec.dim {x::'a::field^'n. \<forall>i. i \<notin> d \<longrightarrow> x$i = 0} = card d"
+lemma%important dim_substandard_cart: "vec.dim {x::'a::field^'n. \<forall>i. i \<notin> d \<longrightarrow> x$i = 0} = card d"
(is "vec.dim ?A = _")
-proof (rule vec.dim_unique)
+proof%unimportant (rule vec.dim_unique)
let ?B = "((\<lambda>x. axis x 1) ` d)"
have subset_basis: "?B \<subseteq> cart_basis"
by (auto simp: cart_basis_def)
@@ -703,27 +703,27 @@
then show "?A \<subseteq> vec.span ?B" by auto
qed (simp add: card_image inj_on_def axis_eq_axis)
-lemma dim_subset_UNIV_cart_gen:
+lemma%unimportant dim_subset_UNIV_cart_gen:
fixes S :: "('a::field^'n) set"
shows "vec.dim S \<le> CARD('n)"
by (metis vec.dim_eq_full vec.dim_subset_UNIV vec.span_UNIV vec_dim_card)
-lemma dim_subset_UNIV_cart:
+lemma%unimportant dim_subset_UNIV_cart:
fixes S :: "(real^'n) set"
shows "dim S \<le> CARD('n)"
using dim_subset_UNIV_cart_gen[of S] by (simp add: dim_vec_eq)
-lemma affinity_inverses:
+lemma%unimportant affinity_inverses:
assumes m0: "m \<noteq> (0::'a::field)"
shows "(\<lambda>x. m *s x + c) \<circ> (\<lambda>x. inverse(m) *s x + (-(inverse(m) *s c))) = id"
"(\<lambda>x. inverse(m) *s x + (-(inverse(m) *s c))) \<circ> (\<lambda>x. m *s x + c) = id"
using m0
by (auto simp add: fun_eq_iff vector_add_ldistrib diff_conv_add_uminus simp del: add_uminus_conv_diff)
-lemma vector_affinity_eq:
+lemma%important vector_affinity_eq:
assumes m0: "(m::'a::field) \<noteq> 0"
shows "m *s x + c = y \<longleftrightarrow> x = inverse m *s y + -(inverse m *s c)"
-proof
+proof%unimportant
assume h: "m *s x + c = y"
hence "m *s x = y - c" by (simp add: field_simps)
hence "inverse m *s (m *s x) = inverse m *s (y - c)" by simp
@@ -735,48 +735,48 @@
using m0 by (simp add: vector_smult_assoc vector_ssub_ldistrib)
qed
-lemma vector_eq_affinity:
+lemma%unimportant vector_eq_affinity:
"(m::'a::field) \<noteq> 0 ==> (y = m *s x + c \<longleftrightarrow> inverse(m) *s y + -(inverse(m) *s c) = x)"
using vector_affinity_eq[where m=m and x=x and y=y and c=c]
by metis
-lemma vector_cart:
+lemma%unimportant vector_cart:
fixes f :: "real^'n \<Rightarrow> real"
shows "(\<chi> i. f (axis i 1)) = (\<Sum>i\<in>Basis. f i *\<^sub>R i)"
unfolding euclidean_eq_iff[where 'a="real^'n"]
by simp (simp add: Basis_vec_def inner_axis)
-lemma const_vector_cart:"((\<chi> i. d)::real^'n) = (\<Sum>i\<in>Basis. d *\<^sub>R i)"
+lemma%unimportant const_vector_cart:"((\<chi> i. d)::real^'n) = (\<Sum>i\<in>Basis. d *\<^sub>R i)"
by (rule vector_cart)
-subsection "Convex Euclidean Space"
+subsection%important "Convex Euclidean Space"
-lemma Cart_1:"(1::real^'n) = \<Sum>Basis"
+lemma%unimportant Cart_1:"(1::real^'n) = \<Sum>Basis"
using const_vector_cart[of 1] by (simp add: one_vec_def)
declare vector_add_ldistrib[simp] vector_ssub_ldistrib[simp] vector_smult_assoc[simp] vector_smult_rneg[simp]
declare vector_sadd_rdistrib[simp] vector_sub_rdistrib[simp]
-lemmas vector_component_simps = vector_minus_component vector_smult_component vector_add_component less_eq_vec_def vec_lambda_beta vector_uminus_component
+lemmas%unimportant vector_component_simps = vector_minus_component vector_smult_component vector_add_component less_eq_vec_def vec_lambda_beta vector_uminus_component
-lemma convex_box_cart:
+lemma%unimportant convex_box_cart:
assumes "\<And>i. convex {x. P i x}"
shows "convex {x. \<forall>i. P i (x$i)}"
using assms unfolding convex_def by auto
-lemma convex_positive_orthant_cart: "convex {x::real^'n. (\<forall>i. 0 \<le> x$i)}"
+lemma%unimportant convex_positive_orthant_cart: "convex {x::real^'n. (\<forall>i. 0 \<le> x$i)}"
by (rule convex_box_cart) (simp add: atLeast_def[symmetric])
-lemma unit_interval_convex_hull_cart:
+lemma%unimportant unit_interval_convex_hull_cart:
"cbox (0::real^'n) 1 = convex hull {x. \<forall>i. (x$i = 0) \<or> (x$i = 1)}"
unfolding Cart_1 unit_interval_convex_hull[where 'a="real^'n"] box_real[symmetric]
by (rule arg_cong[where f="\<lambda>x. convex hull x"]) (simp add: Basis_vec_def inner_axis)
-lemma cube_convex_hull_cart:
+lemma%important cube_convex_hull_cart:
assumes "0 < d"
obtains s::"(real^'n) set"
where "finite s" "cbox (x - (\<chi> i. d)) (x + (\<chi> i. d)) = convex hull s"
-proof -
+proof%unimportant -
from assms obtain s where "finite s"
and "cbox (x - sum (( *\<^sub>R) d) Basis) (x + sum (( *\<^sub>R) d) Basis) = convex hull s"
by (rule cube_convex_hull)
@@ -785,14 +785,14 @@
qed
-subsection "Derivative"
+subsection%important "Derivative"
-definition "jacobian f net = matrix(frechet_derivative f net)"
+definition%important "jacobian f net = matrix(frechet_derivative f net)"
-lemma jacobian_works:
+lemma%important jacobian_works:
"(f::(real^'a) \<Rightarrow> (real^'b)) differentiable net \<longleftrightarrow>
(f has_derivative (\<lambda>h. (jacobian f net) *v h)) net" (is "?lhs = ?rhs")
-proof
+proof%unimportant
assume ?lhs then show ?rhs
by (simp add: frechet_derivative_works has_derivative_linear jacobian_def)
next
@@ -801,10 +801,10 @@
qed
-subsection \<open>Component of the differential must be zero if it exists at a local
+subsection%important \<open>Component of the differential must be zero if it exists at a local
maximum or minimum for that corresponding component\<close>
-lemma differential_zero_maxmin_cart:
+lemma%important differential_zero_maxmin_cart:
fixes f::"real^'a \<Rightarrow> real^'b"
assumes "0 < e" "((\<forall>y \<in> ball x e. (f y)$k \<le> (f x)$k) \<or> (\<forall>y\<in>ball x e. (f x)$k \<le> (f y)$k))"
"f differentiable (at x)"
@@ -813,7 +813,7 @@
vector_cart[of "\<lambda>j. frechet_derivative f (at x) j $ k"]
by (simp add: Basis_vec_def axis_eq_axis inner_axis jacobian_def matrix_def)
-subsection \<open>Lemmas for working on @{typ "real^1"}\<close>
+subsection%unimportant \<open>Lemmas for working on @{typ "real^1"}\<close>
lemma forall_1[simp]: "(\<forall>i::1. P i) \<longleftrightarrow> P 1"
by (metis (full_types) num1_eq_iff)
@@ -893,7 +893,7 @@
instance num1 :: wellorder
by intro_classes (auto simp: less_eq_num1_def less_num1_def)
-subsection\<open>The collapse of the general concepts to dimension one\<close>
+subsection%unimportant\<open>The collapse of the general concepts to dimension one\<close>
lemma vector_one: "(x::'a ^1) = (\<chi> i. (x$1))"
by (simp add: vec_eq_iff)
@@ -918,11 +918,11 @@
lemma dist_real: "dist(x::real ^ 1) y = \<bar>(x$1) - (y$1)\<bar>"
by (auto simp add: norm_real dist_norm)
-subsection\<open> Rank of a matrix\<close>
+subsection%important\<open> Rank of a matrix\<close>
text\<open>Equivalence of row and column rank is taken from George Mackiw's paper, Mathematics Magazine 1995, p. 285.\<close>
-lemma matrix_vector_mult_in_columnspace_gen:
+lemma%unimportant matrix_vector_mult_in_columnspace_gen:
fixes A :: "'a::field^'n^'m"
shows "(A *v x) \<in> vec.span(columns A)"
apply (simp add: matrix_vector_column columns_def transpose_def column_def)
@@ -930,17 +930,17 @@
apply (force intro: vec.span_base)
done
-lemma matrix_vector_mult_in_columnspace:
+lemma%unimportant matrix_vector_mult_in_columnspace:
fixes A :: "real^'n^'m"
shows "(A *v x) \<in> span(columns A)"
using matrix_vector_mult_in_columnspace_gen[of A x] by (simp add: span_vec_eq)
-lemma orthogonal_nullspace_rowspace:
+lemma%important orthogonal_nullspace_rowspace:
fixes A :: "real^'n^'m"
assumes 0: "A *v x = 0" and y: "y \<in> span(rows A)"
shows "orthogonal x y"
using y
-proof (induction rule: span_induct)
+proof%unimportant (induction rule: span_induct)
case base
then show ?case
by (simp add: subspace_orthogonal_to_vector)
@@ -953,28 +953,28 @@
by (simp add: mult.commute) (metis (no_types) vec_lambda_beta zero_index)
qed
-lemma nullspace_inter_rowspace:
+lemma%unimportant nullspace_inter_rowspace:
fixes A :: "real^'n^'m"
shows "A *v x = 0 \<and> x \<in> span(rows A) \<longleftrightarrow> x = 0"
using orthogonal_nullspace_rowspace orthogonal_self span_zero matrix_vector_mult_0_right
by blast
-lemma matrix_vector_mul_injective_on_rowspace:
+lemma%unimportant matrix_vector_mul_injective_on_rowspace:
fixes A :: "real^'n^'m"
shows "\<lbrakk>A *v x = A *v y; x \<in> span(rows A); y \<in> span(rows A)\<rbrakk> \<Longrightarrow> x = y"
using nullspace_inter_rowspace [of A "x-y"]
by (metis diff_eq_diff_eq diff_self matrix_vector_mult_diff_distrib span_diff)
-definition rank :: "'a::field^'n^'m=>nat"
+definition%important rank :: "'a::field^'n^'m=>nat"
where row_rank_def_gen: "rank A \<equiv> vec.dim(rows A)"
-lemma row_rank_def: "rank A = dim (rows A)" for A::"real^'n^'m"
- by (auto simp: row_rank_def_gen dim_vec_eq)
+lemma%important row_rank_def: "rank A = dim (rows A)" for A::"real^'n^'m"
+ by%unimportant (auto simp: row_rank_def_gen dim_vec_eq)
-lemma dim_rows_le_dim_columns:
+lemma%important dim_rows_le_dim_columns:
fixes A :: "real^'n^'m"
shows "dim(rows A) \<le> dim(columns A)"
-proof -
+proof%unimportant -
have "dim (span (rows A)) \<le> dim (span (columns A))"
proof -
obtain B where "independent B" "span(rows A) \<subseteq> span B"
@@ -999,32 +999,32 @@
by (simp add: dim_span)
qed
-lemma column_rank_def:
+lemma%unimportant column_rank_def:
fixes A :: "real^'n^'m"
shows "rank A = dim(columns A)"
unfolding row_rank_def
by (metis columns_transpose dim_rows_le_dim_columns le_antisym rows_transpose)
-lemma rank_transpose:
+lemma%unimportant rank_transpose:
fixes A :: "real^'n^'m"
shows "rank(transpose A) = rank A"
by (metis column_rank_def row_rank_def rows_transpose)
-lemma matrix_vector_mult_basis:
+lemma%unimportant matrix_vector_mult_basis:
fixes A :: "real^'n^'m"
shows "A *v (axis k 1) = column k A"
by (simp add: cart_eq_inner_axis column_def matrix_mult_dot)
-lemma columns_image_basis:
+lemma%unimportant columns_image_basis:
fixes A :: "real^'n^'m"
shows "columns A = ( *v) A ` (range (\<lambda>i. axis i 1))"
by (force simp: columns_def matrix_vector_mult_basis [symmetric])
-lemma rank_dim_range:
+lemma%important rank_dim_range:
fixes A :: "real^'n^'m"
shows "rank A = dim(range (\<lambda>x. A *v x))"
unfolding column_rank_def
-proof (rule span_eq_dim)
+proof%unimportant (rule span_eq_dim)
have "span (columns A) \<subseteq> span (range (( *v) A))" (is "?l \<subseteq> ?r")
by (simp add: columns_image_basis image_subsetI span_mono)
then show "?l = ?r"
@@ -1032,45 +1032,45 @@
span_eq span_span)
qed
-lemma rank_bound:
+lemma%unimportant rank_bound:
fixes A :: "real^'n^'m"
shows "rank A \<le> min CARD('m) (CARD('n))"
by (metis (mono_tags, lifting) dim_subset_UNIV_cart min.bounded_iff
column_rank_def row_rank_def)
-lemma full_rank_injective:
+lemma%unimportant full_rank_injective:
fixes A :: "real^'n^'m"
shows "rank A = CARD('n) \<longleftrightarrow> inj (( *v) A)"
by (simp add: matrix_left_invertible_injective [symmetric] matrix_left_invertible_span_rows row_rank_def
dim_eq_full [symmetric] card_cart_basis vec.dimension_def)
-lemma full_rank_surjective:
+lemma%unimportant full_rank_surjective:
fixes A :: "real^'n^'m"
shows "rank A = CARD('m) \<longleftrightarrow> surj (( *v) A)"
by (simp add: matrix_right_invertible_surjective [symmetric] left_invertible_transpose [symmetric]
matrix_left_invertible_injective full_rank_injective [symmetric] rank_transpose)
-lemma rank_I: "rank(mat 1::real^'n^'n) = CARD('n)"
+lemma%unimportant rank_I: "rank(mat 1::real^'n^'n) = CARD('n)"
by (simp add: full_rank_injective inj_on_def)
-lemma less_rank_noninjective:
+lemma%unimportant less_rank_noninjective:
fixes A :: "real^'n^'m"
shows "rank A < CARD('n) \<longleftrightarrow> \<not> inj (( *v) A)"
using less_le rank_bound by (auto simp: full_rank_injective [symmetric])
-lemma matrix_nonfull_linear_equations_eq:
+lemma%unimportant matrix_nonfull_linear_equations_eq:
fixes A :: "real^'n^'m"
shows "(\<exists>x. (x \<noteq> 0) \<and> A *v x = 0) \<longleftrightarrow> ~(rank A = CARD('n))"
by (meson matrix_left_invertible_injective full_rank_injective matrix_left_invertible_ker)
-lemma rank_eq_0: "rank A = 0 \<longleftrightarrow> A = 0" and rank_0 [simp]: "rank (0::real^'n^'m) = 0"
+lemma%unimportant rank_eq_0: "rank A = 0 \<longleftrightarrow> A = 0" and rank_0 [simp]: "rank (0::real^'n^'m) = 0"
for A :: "real^'n^'m"
by (auto simp: rank_dim_range matrix_eq)
-lemma rank_mul_le_right:
+lemma%important rank_mul_le_right:
fixes A :: "real^'n^'m" and B :: "real^'p^'n"
shows "rank(A ** B) \<le> rank B"
-proof -
+proof%unimportant -
have "rank(A ** B) \<le> dim (( *v) A ` range (( *v) B))"
by (auto simp: rank_dim_range image_comp o_def matrix_vector_mul_assoc)
also have "\<dots> \<le> rank B"
@@ -1078,12 +1078,12 @@
finally show ?thesis .
qed
-lemma rank_mul_le_left:
+lemma%unimportant rank_mul_le_left:
fixes A :: "real^'n^'m" and B :: "real^'p^'n"
shows "rank(A ** B) \<le> rank A"
by (metis matrix_transpose_mul rank_mul_le_right rank_transpose)
-subsection\<open>Routine results connecting the types @{typ "real^1"} and @{typ real}\<close>
+subsection%unimportant\<open>Routine results connecting the types @{typ "real^1"} and @{typ real}\<close>
lemma vector_one_nth [simp]:
fixes x :: "'a^1" shows "vec (x $ 1) = x"
@@ -1146,7 +1146,7 @@
done
-subsection\<open>Explicit vector construction from lists\<close>
+subsection%unimportant\<open>Explicit vector construction from lists\<close>
definition "vector l = (\<chi> i. foldr (\<lambda>x f n. fun_upd (f (n+1)) n x) l (\<lambda>n x. 0) 1 i)"
--- a/src/HOL/Analysis/Cartesian_Space.thy Mon Aug 27 22:58:36 2018 +0200
+++ b/src/HOL/Analysis/Cartesian_Space.thy Tue Aug 28 13:28:39 2018 +0100
@@ -11,19 +11,19 @@
Finite_Cartesian_Product Linear_Algebra
begin
-definition "cart_basis = {axis i 1 | i. i\<in>UNIV}"
+definition%unimportant "cart_basis = {axis i 1 | i. i\<in>UNIV}"
-lemma finite_cart_basis: "finite (cart_basis)" unfolding cart_basis_def
+lemma%unimportant finite_cart_basis: "finite (cart_basis)" unfolding cart_basis_def
using finite_Atleast_Atmost_nat by fastforce
-lemma card_cart_basis: "card (cart_basis::('a::zero_neq_one^'i) set) = CARD('i)"
+lemma%unimportant card_cart_basis: "card (cart_basis::('a::zero_neq_one^'i) set) = CARD('i)"
unfolding cart_basis_def Setcompr_eq_image
by (rule card_image) (auto simp: inj_on_def axis_eq_axis)
interpretation vec: vector_space "( *s) "
by unfold_locales (vector algebra_simps)+
-lemma independent_cart_basis:
+lemma%unimportant independent_cart_basis:
"vec.independent (cart_basis)"
proof (rule vec.independent_if_scalars_zero)
show "finite (cart_basis)" using finite_cart_basis .
@@ -48,7 +48,7 @@
finally show "f x = 0" ..
qed
-lemma span_cart_basis:
+lemma%unimportant span_cart_basis:
"vec.span (cart_basis) = UNIV"
proof (auto)
fix x::"('a, 'b) vec"
@@ -93,12 +93,12 @@
interpretation vec: finite_dimensional_vector_space "( *s)" "cart_basis"
by (unfold_locales, auto simp add: finite_cart_basis independent_cart_basis span_cart_basis)
-lemma matrix_vector_mul_linear_gen[intro, simp]:
+lemma%unimportant matrix_vector_mul_linear_gen[intro, simp]:
"Vector_Spaces.linear ( *s) ( *s) (( *v) A)"
by unfold_locales
(vector matrix_vector_mult_def sum.distrib algebra_simps)+
-lemma span_vec_eq: "vec.span X = span X"
+lemma%important span_vec_eq: "vec.span X = span X"
and dim_vec_eq: "vec.dim X = dim X"
and dependent_vec_eq: "vec.dependent X = dependent X"
and subspace_vec_eq: "vec.subspace X = subspace X"
@@ -106,11 +106,11 @@
unfolding span_raw_def dim_raw_def dependent_raw_def subspace_raw_def
by (auto simp: scalar_mult_eq_scaleR)
-lemma linear_componentwise:
+lemma%important linear_componentwise:
fixes f:: "'a::field ^'m \<Rightarrow> 'a ^ 'n"
assumes lf: "Vector_Spaces.linear ( *s) ( *s) f"
shows "(f x)$j = sum (\<lambda>i. (x$i) * (f (axis i 1)$j)) (UNIV :: 'm set)" (is "?lhs = ?rhs")
-proof -
+proof%unimportant -
interpret lf: Vector_Spaces.linear "( *s)" "( *s)" f
using lf .
let ?M = "(UNIV :: 'm set)"
@@ -128,7 +128,7 @@
interpretation vec: finite_dimensional_vector_space_pair "( *s)" cart_basis "( *s)" cart_basis ..
-lemma matrix_works:
+lemma%unimportant matrix_works:
assumes lf: "Vector_Spaces.linear ( *s) ( *s) f"
shows "matrix f *v x = f (x::'a::field ^ 'n)"
apply (simp add: matrix_def matrix_vector_mult_def vec_eq_iff mult.commute)
@@ -136,45 +136,45 @@
apply (rule linear_componentwise[OF lf, symmetric])
done
-lemma matrix_of_matrix_vector_mul[simp]: "matrix(\<lambda>x. A *v (x :: 'a::field ^ 'n)) = A"
+lemma%unimportant matrix_of_matrix_vector_mul[simp]: "matrix(\<lambda>x. A *v (x :: 'a::field ^ 'n)) = A"
by (simp add: matrix_eq matrix_works)
-lemma matrix_compose_gen:
+lemma%unimportant matrix_compose_gen:
assumes lf: "Vector_Spaces.linear ( *s) ( *s) (f::'a::{field}^'n \<Rightarrow> 'a^'m)"
and lg: "Vector_Spaces.linear ( *s) ( *s) (g::'a^'m \<Rightarrow> 'a^_)"
shows "matrix (g o f) = matrix g ** matrix f"
using lf lg Vector_Spaces.linear_compose[OF lf lg] matrix_works[OF Vector_Spaces.linear_compose[OF lf lg]]
by (simp add: matrix_eq matrix_works matrix_vector_mul_assoc[symmetric] o_def)
-lemma matrix_compose:
+lemma%unimportant matrix_compose:
assumes "linear (f::real^'n \<Rightarrow> real^'m)" "linear (g::real^'m \<Rightarrow> real^_)"
shows "matrix (g o f) = matrix g ** matrix f"
using matrix_compose_gen[of f g] assms
by (simp add: linear_def scalar_mult_eq_scaleR)
-lemma left_invertible_transpose:
+lemma%unimportant left_invertible_transpose:
"(\<exists>(B). B ** transpose (A) = mat (1::'a::comm_semiring_1)) \<longleftrightarrow> (\<exists>(B). A ** B = mat 1)"
by (metis matrix_transpose_mul transpose_mat transpose_transpose)
-lemma right_invertible_transpose:
+lemma%unimportant right_invertible_transpose:
"(\<exists>(B). transpose (A) ** B = mat (1::'a::comm_semiring_1)) \<longleftrightarrow> (\<exists>(B). B ** A = mat 1)"
by (metis matrix_transpose_mul transpose_mat transpose_transpose)
-lemma linear_matrix_vector_mul_eq:
+lemma%unimportant linear_matrix_vector_mul_eq:
"Vector_Spaces.linear ( *s) ( *s) f \<longleftrightarrow> linear (f :: real^'n \<Rightarrow> real ^'m)"
by (simp add: scalar_mult_eq_scaleR linear_def)
-lemma matrix_vector_mul[simp]:
+lemma%unimportant matrix_vector_mul[simp]:
"Vector_Spaces.linear ( *s) ( *s) g \<Longrightarrow> (\<lambda>y. matrix g *v y) = g"
"linear f \<Longrightarrow> (\<lambda>x. matrix f *v x) = f"
"bounded_linear f \<Longrightarrow> (\<lambda>x. matrix f *v x) = f"
for f :: "real^'n \<Rightarrow> real ^'m"
by (simp_all add: ext matrix_works linear_matrix_vector_mul_eq linear_linear)
-lemma matrix_left_invertible_injective:
+lemma%important matrix_left_invertible_injective:
fixes A :: "'a::field^'n^'m"
shows "(\<exists>B. B ** A = mat 1) \<longleftrightarrow> inj (( *v) A)"
-proof safe
+proof%unimportant safe
fix B
assume B: "B ** A = mat 1"
show "inj (( *v) A)"
@@ -192,15 +192,15 @@
by metis
qed
-lemma matrix_left_invertible_ker:
+lemma%unimportant matrix_left_invertible_ker:
"(\<exists>B. (B::'a::{field} ^'m^'n) ** (A::'a::{field}^'n^'m) = mat 1) \<longleftrightarrow> (\<forall>x. A *v x = 0 \<longrightarrow> x = 0)"
unfolding matrix_left_invertible_injective
using vec.inj_on_iff_eq_0[OF vec.subspace_UNIV, of A]
by (simp add: inj_on_def)
-lemma matrix_right_invertible_surjective:
+lemma%important matrix_right_invertible_surjective:
"(\<exists>B. (A::'a::field^'n^'m) ** (B::'a::field^'m^'n) = mat 1) \<longleftrightarrow> surj (\<lambda>x. A *v x)"
-proof -
+proof%unimportant -
{ fix B :: "'a ^'m^'n"
assume AB: "A ** B = mat 1"
{ fix x :: "'a ^ 'm"
@@ -223,12 +223,12 @@
ultimately show ?thesis unfolding surj_def by blast
qed
-lemma matrix_left_invertible_independent_columns:
+lemma%important matrix_left_invertible_independent_columns:
fixes A :: "'a::{field}^'n^'m"
shows "(\<exists>(B::'a ^'m^'n). B ** A = mat 1) \<longleftrightarrow>
(\<forall>c. sum (\<lambda>i. c i *s column i A) (UNIV :: 'n set) = 0 \<longrightarrow> (\<forall>i. c i = 0))"
(is "?lhs \<longleftrightarrow> ?rhs")
-proof -
+proof%unimportant -
let ?U = "UNIV :: 'n set"
{ assume k: "\<forall>x. A *v x = 0 \<longrightarrow> x = 0"
{ fix c i
@@ -250,7 +250,7 @@
ultimately show ?thesis unfolding matrix_left_invertible_ker by auto
qed
-lemma matrix_right_invertible_independent_rows:
+lemma%unimportant matrix_right_invertible_independent_rows:
fixes A :: "'a::{field}^'n^'m"
shows "(\<exists>(B::'a^'m^'n). A ** B = mat 1) \<longleftrightarrow>
(\<forall>c. sum (\<lambda>i. c i *s row i A) (UNIV :: 'm set) = 0 \<longrightarrow> (\<forall>i. c i = 0))"
@@ -258,10 +258,10 @@
matrix_left_invertible_independent_columns
by (simp add:)
-lemma matrix_right_invertible_span_columns:
+lemma%important matrix_right_invertible_span_columns:
"(\<exists>(B::'a::field ^'n^'m). (A::'a ^'m^'n) ** B = mat 1) \<longleftrightarrow>
vec.span (columns A) = UNIV" (is "?lhs = ?rhs")
-proof -
+proof%unimportant -
let ?U = "UNIV :: 'm set"
have fU: "finite ?U" by simp
have lhseq: "?lhs \<longleftrightarrow> (\<forall>y. \<exists>(x::'a^'m). sum (\<lambda>i. (x$i) *s column i A) ?U = y)"
@@ -322,21 +322,21 @@
ultimately show ?thesis by blast
qed
-lemma matrix_left_invertible_span_rows_gen:
+lemma%unimportant matrix_left_invertible_span_rows_gen:
"(\<exists>(B::'a^'m^'n). B ** (A::'a::field^'n^'m) = mat 1) \<longleftrightarrow> vec.span (rows A) = UNIV"
unfolding right_invertible_transpose[symmetric]
unfolding columns_transpose[symmetric]
unfolding matrix_right_invertible_span_columns
..
-lemma matrix_left_invertible_span_rows:
+lemma%unimportant matrix_left_invertible_span_rows:
"(\<exists>(B::real^'m^'n). B ** (A::real^'n^'m) = mat 1) \<longleftrightarrow> span (rows A) = UNIV"
using matrix_left_invertible_span_rows_gen[of A] by (simp add: span_vec_eq)
-lemma matrix_left_right_inverse:
+lemma%important matrix_left_right_inverse:
fixes A A' :: "'a::{field}^'n^'n"
shows "A ** A' = mat 1 \<longleftrightarrow> A' ** A = mat 1"
-proof -
+proof%unimportant -
{ fix A A' :: "'a ^'n^'n"
assume AA': "A ** A' = mat 1"
have sA: "surj (( *v) A)"
@@ -356,21 +356,21 @@
then show ?thesis by blast
qed
-lemma invertible_left_inverse:
+lemma%unimportant invertible_left_inverse:
fixes A :: "'a::{field}^'n^'n"
shows "invertible A \<longleftrightarrow> (\<exists>(B::'a^'n^'n). B ** A = mat 1)"
by (metis invertible_def matrix_left_right_inverse)
-lemma invertible_right_inverse:
+lemma%unimportant invertible_right_inverse:
fixes A :: "'a::{field}^'n^'n"
shows "invertible A \<longleftrightarrow> (\<exists>(B::'a^'n^'n). A** B = mat 1)"
by (metis invertible_def matrix_left_right_inverse)
-lemma invertible_mult:
+lemma%important invertible_mult:
assumes inv_A: "invertible A"
and inv_B: "invertible B"
shows "invertible (A**B)"
-proof -
+proof%unimportant -
obtain A' where AA': "A ** A' = mat 1" and A'A: "A' ** A = mat 1"
using inv_A unfolding invertible_def by blast
obtain B' where BB': "B ** B' = mat 1" and B'B: "B' ** B = mat 1"
@@ -393,28 +393,28 @@
qed
qed
-lemma transpose_invertible:
+lemma%unimportant transpose_invertible:
fixes A :: "real^'n^'n"
assumes "invertible A"
shows "invertible (transpose A)"
by (meson assms invertible_def matrix_left_right_inverse right_invertible_transpose)
-lemma vector_matrix_mul_assoc:
+lemma%important vector_matrix_mul_assoc:
fixes v :: "('a::comm_semiring_1)^'n"
shows "(v v* M) v* N = v v* (M ** N)"
-proof -
+proof%unimportant -
from matrix_vector_mul_assoc
have "transpose N *v (transpose M *v v) = (transpose N ** transpose M) *v v" by fast
thus "(v v* M) v* N = v v* (M ** N)"
by (simp add: matrix_transpose_mul [symmetric])
qed
-lemma matrix_scaleR_vector_ac:
+lemma%unimportant matrix_scaleR_vector_ac:
fixes A :: "real^('m::finite)^'n"
shows "A *v (k *\<^sub>R v) = k *\<^sub>R A *v v"
by (metis matrix_vector_mult_scaleR transpose_scalar vector_scaleR_matrix_ac vector_transpose_matrix)
-lemma scaleR_matrix_vector_assoc:
+lemma%unimportant scaleR_matrix_vector_assoc:
fixes A :: "real^('m::finite)^'n"
shows "k *\<^sub>R (A *v v) = k *\<^sub>R A *v v"
by (metis matrix_scaleR_vector_ac matrix_vector_mult_scaleR)
@@ -430,8 +430,8 @@
and BasisB :: "('b set)"
and f :: "('b=>'c)"
-lemma vec_dim_card: "vec.dim (UNIV::('a::{field}^'n) set) = CARD ('n)"
-proof -
+lemma%important vec_dim_card: "vec.dim (UNIV::('a::{field}^'n) set) = CARD ('n)"
+proof%unimportant -
let ?f="\<lambda>i::'n. axis i (1::'a)"
have "vec.dim (UNIV::('a::{field}^'n) set) = card (cart_basis::('a^'n) set)"
unfolding vec.dim_UNIV ..
--- a/src/HOL/Analysis/Change_Of_Vars.thy Mon Aug 27 22:58:36 2018 +0200
+++ b/src/HOL/Analysis/Change_Of_Vars.thy Tue Aug 28 13:28:39 2018 +0100
@@ -9,9 +9,9 @@
begin
-subsection\<open>Induction on matrix row operations\<close>
+subsection%important\<open>Induction on matrix row operations\<close>
-lemma induct_matrix_row_operations:
+lemma%unimportant induct_matrix_row_operations:
fixes P :: "real^'n^'n \<Rightarrow> bool"
assumes zero_row: "\<And>A i. row i A = 0 \<Longrightarrow> P A"
and diagonal: "\<And>A. (\<And>i j. i \<noteq> j \<Longrightarrow> A$i$j = 0) \<Longrightarrow> P A"
@@ -120,7 +120,7 @@
by blast
qed
-lemma induct_matrix_elementary:
+lemma%unimportant induct_matrix_elementary:
fixes P :: "real^'n^'n \<Rightarrow> bool"
assumes mult: "\<And>A B. \<lbrakk>P A; P B\<rbrakk> \<Longrightarrow> P(A ** B)"
and zero_row: "\<And>A i. row i A = 0 \<Longrightarrow> P A"
@@ -151,7 +151,7 @@
by (rule induct_matrix_row_operations [OF zero_row diagonal swap row])
qed
-lemma induct_matrix_elementary_alt:
+lemma%unimportant induct_matrix_elementary_alt:
fixes P :: "real^'n^'n \<Rightarrow> bool"
assumes mult: "\<And>A B. \<lbrakk>P A; P B\<rbrakk> \<Longrightarrow> P(A ** B)"
and zero_row: "\<And>A i. row i A = 0 \<Longrightarrow> P A"
@@ -185,11 +185,11 @@
by (rule induct_matrix_elementary) (auto intro: assms *)
qed
-lemma matrix_vector_mult_matrix_matrix_mult_compose:
+lemma%unimportant matrix_vector_mult_matrix_matrix_mult_compose:
"( *v) (A ** B) = ( *v) A \<circ> ( *v) B"
by (auto simp: matrix_vector_mul_assoc)
-lemma induct_linear_elementary:
+lemma%unimportant induct_linear_elementary:
fixes f :: "real^'n \<Rightarrow> real^'n"
assumes "linear f"
and comp: "\<And>f g. \<lbrakk>linear f; linear g; P f; P g\<rbrakk> \<Longrightarrow> P(f \<circ> g)"
@@ -250,14 +250,14 @@
qed
-proposition
+proposition%important
fixes a :: "real^'n"
assumes "m \<noteq> n" and ab_ne: "cbox a b \<noteq> {}" and an: "0 \<le> a$n"
shows measurable_shear_interval: "(\<lambda>x. \<chi> i. if i = m then x$m + x$n else x$i) ` (cbox a b) \<in> lmeasurable"
(is "?f ` _ \<in> _")
and measure_shear_interval: "measure lebesgue ((\<lambda>x. \<chi> i. if i = m then x$m + x$n else x$i) ` cbox a b)
= measure lebesgue (cbox a b)" (is "?Q")
-proof -
+proof%unimportant -
have lin: "linear ?f"
by (rule linearI) (auto simp: plus_vec_def scaleR_vec_def algebra_simps)
show fab: "?f ` cbox a b \<in> lmeasurable"
@@ -359,13 +359,13 @@
qed
-proposition
+proposition%important
fixes S :: "(real^'n) set"
assumes "S \<in> lmeasurable"
shows measurable_stretch: "((\<lambda>x. \<chi> k. m k * x$k) ` S) \<in> lmeasurable" (is "?f ` S \<in> _")
and measure_stretch: "measure lebesgue ((\<lambda>x. \<chi> k. m k * x$k) ` S) = \<bar>prod m UNIV\<bar> * measure lebesgue S"
(is "?MEQ")
-proof -
+proof%unimportant -
have "(?f ` S) \<in> lmeasurable \<and> ?MEQ"
proof (cases "\<forall>k. m k \<noteq> 0")
case True
@@ -463,12 +463,12 @@
qed
-proposition
+proposition%important
fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real^'n::_"
assumes "linear f" "S \<in> lmeasurable"
shows measurable_linear_image: "(f ` S) \<in> lmeasurable"
and measure_linear_image: "measure lebesgue (f ` S) = \<bar>det (matrix f)\<bar> * measure lebesgue S" (is "?Q f S")
-proof -
+proof%unimportant -
have "\<forall>S \<in> lmeasurable. (f ` S) \<in> lmeasurable \<and> ?Q f S"
proof (rule induct_linear_elementary [OF \<open>linear f\<close>]; intro ballI)
fix f g and S :: "(real,'n) vec set"
@@ -611,7 +611,7 @@
qed
-lemma
+lemma%unimportant
fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real^'n::_"
assumes f: "orthogonal_transformation f" and S: "S \<in> lmeasurable"
shows measurable_orthogonal_image: "f ` S \<in> lmeasurable"
@@ -625,19 +625,19 @@
by (simp add: measure_linear_image \<open>linear f\<close> S f)
qed
-subsection\<open>@{text F_sigma} and @{text G_delta} sets.\<close>
+subsection%important\<open>@{text F_sigma} and @{text G_delta} sets.\<close>(*FIX ME mv *)
(*https://en.wikipedia.org/wiki/F\<sigma>_set*)
-inductive fsigma :: "'a::topological_space set \<Rightarrow> bool" where
+inductive%important fsigma :: "'a::topological_space set \<Rightarrow> bool" where
"(\<And>n::nat. closed (F n)) \<Longrightarrow> fsigma (UNION UNIV F)"
-inductive gdelta :: "'a::topological_space set \<Rightarrow> bool" where
+inductive%important gdelta :: "'a::topological_space set \<Rightarrow> bool" where
"(\<And>n::nat. open (F n)) \<Longrightarrow> gdelta (INTER UNIV F)"
-lemma fsigma_Union_compact:
+lemma%important fsigma_Union_compact:
fixes S :: "'a::{real_normed_vector,heine_borel} set"
shows "fsigma S \<longleftrightarrow> (\<exists>F::nat \<Rightarrow> 'a set. range F \<subseteq> Collect compact \<and> S = UNION UNIV F)"
-proof safe
+proof%unimportant safe
assume "fsigma S"
then obtain F :: "nat \<Rightarrow> 'a set" where F: "range F \<subseteq> Collect closed" "S = UNION UNIV F"
by (meson fsigma.cases image_subsetI mem_Collect_eq)
@@ -668,7 +668,7 @@
by (simp add: compact_imp_closed fsigma.intros image_subset_iff)
qed
-lemma gdelta_imp_fsigma: "gdelta S \<Longrightarrow> fsigma (- S)"
+lemma%unimportant gdelta_imp_fsigma: "gdelta S \<Longrightarrow> fsigma (- S)"
proof (induction rule: gdelta.induct)
case (1 F)
have "- INTER UNIV F = (\<Union>i. -(F i))"
@@ -677,7 +677,7 @@
by (simp add: fsigma.intros closed_Compl 1)
qed
-lemma fsigma_imp_gdelta: "fsigma S \<Longrightarrow> gdelta (- S)"
+lemma%unimportant fsigma_imp_gdelta: "fsigma S \<Longrightarrow> gdelta (- S)"
proof (induction rule: fsigma.induct)
case (1 F)
have "- UNION UNIV F = (\<Inter>i. -(F i))"
@@ -686,11 +686,11 @@
by (simp add: 1 gdelta.intros open_closed)
qed
-lemma gdelta_complement: "gdelta(- S) \<longleftrightarrow> fsigma S"
+lemma%unimportant gdelta_complement: "gdelta(- S) \<longleftrightarrow> fsigma S"
using fsigma_imp_gdelta gdelta_imp_fsigma by force
text\<open>A Lebesgue set is almost an @{text F_sigma} or @{text G_delta}.\<close>
-lemma lebesgue_set_almost_fsigma:
+lemma%unimportant lebesgue_set_almost_fsigma:
assumes "S \<in> sets lebesgue"
obtains C T where "fsigma C" "negligible T" "C \<union> T = S" "disjnt C T"
proof -
@@ -725,7 +725,7 @@
qed
qed
-lemma lebesgue_set_almost_gdelta:
+lemma%unimportant lebesgue_set_almost_gdelta:
assumes "S \<in> sets lebesgue"
obtains C T where "gdelta C" "negligible T" "S \<union> T = C" "disjnt S T"
proof -
@@ -743,12 +743,12 @@
qed
-proposition measure_semicontinuous_with_hausdist_explicit:
+proposition%important measure_semicontinuous_with_hausdist_explicit:
assumes "bounded S" and neg: "negligible(frontier S)" and "e > 0"
obtains d where "d > 0"
"\<And>T. \<lbrakk>T \<in> lmeasurable; \<And>y. y \<in> T \<Longrightarrow> \<exists>x. x \<in> S \<and> dist x y < d\<rbrakk>
\<Longrightarrow> measure lebesgue T < measure lebesgue S + e"
-proof (cases "S = {}")
+proof%unimportant (cases "S = {}")
case True
with that \<open>e > 0\<close> show ?thesis by force
next
@@ -812,10 +812,10 @@
qed
qed
-proposition lebesgue_regular_inner:
+proposition%important lebesgue_regular_inner:
assumes "S \<in> sets lebesgue"
obtains K C where "negligible K" "\<And>n::nat. compact(C n)" "S = (\<Union>n. C n) \<union> K"
-proof -
+proof%unimportant -
have "\<exists>T. closed T \<and> T \<subseteq> S \<and> (S - T) \<in> lmeasurable \<and> measure lebesgue (S - T) < (1/2)^n" for n
using sets_lebesgue_inner_closed assms
by (metis sets_lebesgue_inner_closed zero_less_divide_1_iff zero_less_numeral zero_less_power)
@@ -865,7 +865,7 @@
qed
-lemma sets_lebesgue_continuous_image:
+lemma%unimportant sets_lebesgue_continuous_image:
assumes T: "T \<in> sets lebesgue" and contf: "continuous_on S f"
and negim: "\<And>T. \<lbrakk>negligible T; T \<subseteq> S\<rbrakk> \<Longrightarrow> negligible(f ` T)" and "T \<subseteq> S"
shows "f ` T \<in> sets lebesgue"
@@ -890,7 +890,7 @@
by (simp add: Teq image_Un image_Union)
qed
-lemma differentiable_image_in_sets_lebesgue:
+lemma%unimportant differentiable_image_in_sets_lebesgue:
fixes f :: "'m::euclidean_space \<Rightarrow> 'n::euclidean_space"
assumes S: "S \<in> sets lebesgue" and dim: "DIM('m) \<le> DIM('n)" and f: "f differentiable_on S"
shows "f`S \<in> sets lebesgue"
@@ -902,7 +902,7 @@
by (auto simp: intro!: negligible_differentiable_image_negligible [OF dim])
qed auto
-lemma sets_lebesgue_on_continuous_image:
+lemma%unimportant sets_lebesgue_on_continuous_image:
assumes S: "S \<in> sets lebesgue" and X: "X \<in> sets (lebesgue_on S)" and contf: "continuous_on S f"
and negim: "\<And>T. \<lbrakk>negligible T; T \<subseteq> S\<rbrakk> \<Longrightarrow> negligible(f ` T)"
shows "f ` X \<in> sets (lebesgue_on (f ` S))"
@@ -917,7 +917,7 @@
by (auto simp: sets_restrict_space_iff)
qed
-lemma differentiable_image_in_sets_lebesgue_on:
+lemma%unimportant differentiable_image_in_sets_lebesgue_on:
fixes f :: "'m::euclidean_space \<Rightarrow> 'n::euclidean_space"
assumes S: "S \<in> sets lebesgue" and X: "X \<in> sets (lebesgue_on S)" and dim: "DIM('m) \<le> DIM('n)"
and f: "f differentiable_on S"
@@ -931,7 +931,7 @@
qed
-proposition
+proposition%important
fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real^'n::_"
assumes S: "S \<in> lmeasurable"
and deriv: "\<And>x. x \<in> S \<Longrightarrow> (f has_derivative f' x) (at x within S)"
@@ -941,7 +941,7 @@
"f ` S \<in> lmeasurable"
and measure_bounded_differentiable_image:
"measure lebesgue (f ` S) \<le> B * measure lebesgue S" (is "?M")
-proof -
+proof%unimportant -
have "f ` S \<in> lmeasurable \<and> measure lebesgue (f ` S) \<le> B * measure lebesgue S"
proof (cases "B < 0")
case True
@@ -1166,13 +1166,13 @@
then show "f ` S \<in> lmeasurable" ?M by blast+
qed
-lemma
+lemma%important
fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real^'n::_"
assumes S: "S \<in> lmeasurable"
and deriv: "\<And>x. x \<in> S \<Longrightarrow> (f has_derivative f' x) (at x within S)"
and int: "(\<lambda>x. \<bar>det (matrix (f' x))\<bar>) integrable_on S"
shows m_diff_image_weak: "f ` S \<in> lmeasurable \<and> measure lebesgue (f ` S) \<le> integral S (\<lambda>x. \<bar>det (matrix (f' x))\<bar>)"
-proof -
+proof%unimportant -
let ?\<mu> = "measure lebesgue"
have aint_S: "(\<lambda>x. \<bar>det (matrix (f' x))\<bar>) absolutely_integrable_on S"
using int unfolding absolutely_integrable_on_def by auto
@@ -1364,7 +1364,7 @@
qed
-theorem
+theorem%important
fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real^'n::_"
assumes S: "S \<in> sets lebesgue"
and deriv: "\<And>x. x \<in> S \<Longrightarrow> (f has_derivative f' x) (at x within S)"
@@ -1372,7 +1372,7 @@
shows measurable_differentiable_image: "f ` S \<in> lmeasurable"
and measure_differentiable_image:
"measure lebesgue (f ` S) \<le> integral S (\<lambda>x. \<bar>det (matrix (f' x))\<bar>)" (is "?M")
-proof -
+proof%unimportant -
let ?I = "\<lambda>n::nat. cbox (vec (-n)) (vec n) \<inter> S"
let ?\<mu> = "measure lebesgue"
have "x \<in> cbox (vec (- real (nat \<lceil>norm x\<rceil>))) (vec (real (nat \<lceil>norm x\<rceil>)))" for x :: "real^'n::_"
@@ -1417,7 +1417,7 @@
qed
-lemma borel_measurable_simple_function_limit_increasing:
+lemma%unimportant borel_measurable_simple_function_limit_increasing:
fixes f :: "'a::euclidean_space \<Rightarrow> real"
shows "(f \<in> borel_measurable lebesgue \<and> (\<forall>x. 0 \<le> f x)) \<longleftrightarrow>
(\<exists>g. (\<forall>n x. 0 \<le> g n x \<and> g n x \<le> f x) \<and> (\<forall>n x. g n x \<le> (g(Suc n) x)) \<and>
@@ -1614,9 +1614,9 @@
by (blast intro: order_trans)
qed
-subsection\<open>Borel measurable Jacobian determinant\<close>
+subsection%important\<open>Borel measurable Jacobian determinant\<close>
-lemma lemma_partial_derivatives0:
+lemma%unimportant lemma_partial_derivatives0:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes "linear f" and lim0: "((\<lambda>x. f x /\<^sub>R norm x) \<longlongrightarrow> 0) (at 0 within S)"
and lb: "\<And>v. v \<noteq> 0 \<Longrightarrow> (\<exists>k>0. \<forall>e>0. \<exists>x. x \<in> S - {0} \<and> norm x < e \<and> k * norm x \<le> \<bar>v \<bullet> x\<bar>)"
@@ -1689,7 +1689,7 @@
mem_Collect_eq module_hom_zero span_base span_raw_def)
qed
-lemma lemma_partial_derivatives:
+lemma%unimportant lemma_partial_derivatives:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes "linear f" and lim: "((\<lambda>x. f (x - a) /\<^sub>R norm (x - a)) \<longlongrightarrow> 0) (at a within S)"
and lb: "\<And>v. v \<noteq> 0 \<Longrightarrow> (\<exists>k>0. \<forall>e>0. \<exists>x \<in> S - {a}. norm(a - x) < e \<and> k * norm(a - x) \<le> \<bar>v \<bullet> (x - a)\<bar>)"
@@ -1707,12 +1707,12 @@
qed
-proposition borel_measurable_partial_derivatives:
+proposition%important borel_measurable_partial_derivatives:
fixes f :: "real^'m::{finite,wellorder} \<Rightarrow> real^'n"
assumes S: "S \<in> sets lebesgue"
and f: "\<And>x. x \<in> S \<Longrightarrow> (f has_derivative f' x) (at x within S)"
shows "(\<lambda>x. (matrix(f' x)$m$n)) \<in> borel_measurable (lebesgue_on S)"
-proof -
+proof%unimportant -
have contf: "continuous_on S f"
using continuous_on_eq_continuous_within f has_derivative_continuous by blast
have "{x \<in> S. (matrix (f' x)$m$n) \<le> b} \<in> sets lebesgue" for b
@@ -2298,21 +2298,21 @@
qed
-theorem borel_measurable_det_Jacobian:
+theorem%important borel_measurable_det_Jacobian:
fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real^'n::_"
assumes S: "S \<in> sets lebesgue" and f: "\<And>x. x \<in> S \<Longrightarrow> (f has_derivative f' x) (at x within S)"
shows "(\<lambda>x. det(matrix(f' x))) \<in> borel_measurable (lebesgue_on S)"
unfolding det_def
- by (intro measurable) (auto intro: f borel_measurable_partial_derivatives [OF S])
+ by%unimportant (intro measurable) (auto intro: f borel_measurable_partial_derivatives [OF S])
text\<open>The localisation wrt S uses the same argument for many similar results.\<close>
(*See HOL Light's MEASURABLE_ON_LEBESGUE_MEASURABLE_PREIMAGE_BOREL, etc.*)
-lemma borel_measurable_lebesgue_on_preimage_borel:
+lemma%important borel_measurable_lebesgue_on_preimage_borel:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes "S \<in> sets lebesgue"
shows "f \<in> borel_measurable (lebesgue_on S) \<longleftrightarrow>
(\<forall>T. T \<in> sets borel \<longrightarrow> {x \<in> S. f x \<in> T} \<in> sets lebesgue)"
-proof -
+proof%unimportant -
have "{x. (if x \<in> S then f x else 0) \<in> T} \<in> sets lebesgue \<longleftrightarrow> {x \<in> S. f x \<in> T} \<in> sets lebesgue"
if "T \<in> sets borel" for T
proof (cases "0 \<in> T")
@@ -2334,7 +2334,7 @@
by blast
qed
-lemma sets_lebesgue_almost_borel:
+lemma%unimportant sets_lebesgue_almost_borel:
assumes "S \<in> sets lebesgue"
obtains B N where "B \<in> sets borel" "negligible N" "B \<union> N = S"
proof -
@@ -2344,7 +2344,7 @@
by (metis negligible_iff_null_sets negligible_subset null_sets_completionI that)
qed
-lemma double_lebesgue_sets:
+lemma%unimportant double_lebesgue_sets:
assumes S: "S \<in> sets lebesgue" and T: "T \<in> sets lebesgue" and fim: "f ` S \<subseteq> T"
shows "(\<forall>U. U \<in> sets lebesgue \<and> U \<subseteq> T \<longrightarrow> {x \<in> S. f x \<in> U} \<in> sets lebesgue) \<longleftrightarrow>
f \<in> borel_measurable (lebesgue_on S) \<and>
@@ -2387,9 +2387,9 @@
qed
-subsection\<open>Simplest case of Sard's theorem (we don't need continuity of derivative)\<close>
+subsection%important\<open>Simplest case of Sard's theorem (we don't need continuity of derivative)\<close>
-lemma Sard_lemma00:
+lemma%important Sard_lemma00:
fixes P :: "'b::euclidean_space set"
assumes "a \<ge> 0" and a: "a *\<^sub>R i \<noteq> 0" and i: "i \<in> Basis"
and P: "P \<subseteq> {x. a *\<^sub>R i \<bullet> x = 0}"
@@ -2397,7 +2397,7 @@
obtains S where "S \<in> lmeasurable"
and "{z. norm z \<le> m \<and> (\<exists>t \<in> P. norm(z - t) \<le> e)} \<subseteq> S"
and "measure lebesgue S \<le> (2 * e) * (2 * m) ^ (DIM('b) - 1)"
-proof -
+proof%unimportant -
have "a > 0"
using assms by simp
let ?v = "(\<Sum>j\<in>Basis. (if j = i then e else m) *\<^sub>R j)"
@@ -2425,7 +2425,7 @@
qed
text\<open>As above, but reorienting the vector (HOL Light's @text{GEOM\_BASIS\_MULTIPLE\_TAC})\<close>
-lemma Sard_lemma0:
+lemma%unimportant Sard_lemma0:
fixes P :: "(real^'n::{finite,wellorder}) set"
assumes "a \<noteq> 0"
and P: "P \<subseteq> {x. a \<bullet> x = 0}" and "0 \<le> m" "0 \<le> e"
@@ -2485,13 +2485,13 @@
qed
text\<open>As above, but translating the sets (HOL Light's @text{GEN\_GEOM\_ORIGIN\_TAC})\<close>
-lemma Sard_lemma1:
+lemma%important Sard_lemma1:
fixes P :: "(real^'n::{finite,wellorder}) set"
assumes P: "dim P < CARD('n)" and "0 \<le> m" "0 \<le> e"
obtains S where "S \<in> lmeasurable"
and "{z. norm(z - w) \<le> m \<and> (\<exists>t \<in> P. norm(z - w - t) \<le> e)} \<subseteq> S"
and "measure lebesgue S \<le> (2 * e) * (2 * m) ^ (CARD('n) - 1)"
-proof -
+proof%unimportant -
obtain a where "a \<noteq> 0" "P \<subseteq> {x. a \<bullet> x = 0}"
using lowdim_subset_hyperplane [of P] P span_base by auto
then obtain S where S: "S \<in> lmeasurable"
@@ -2509,7 +2509,7 @@
qed
qed
-lemma Sard_lemma2:
+lemma%important Sard_lemma2:
fixes f :: "real^'m::{finite,wellorder} \<Rightarrow> real^'n::{finite,wellorder}"
assumes mlen: "CARD('m) \<le> CARD('n)" (is "?m \<le> ?n")
and "B > 0" "bounded S"
@@ -2517,7 +2517,7 @@
and rank: "\<And>x. x \<in> S \<Longrightarrow> rank(matrix(f' x)) < CARD('n)"
and B: "\<And>x. x \<in> S \<Longrightarrow> onorm(f' x) \<le> B"
shows "negligible(f ` S)"
-proof -
+proof%unimportant -
have lin_f': "\<And>x. x \<in> S \<Longrightarrow> linear(f' x)"
using derS has_derivative_linear by blast
show ?thesis
@@ -2724,13 +2724,13 @@
qed
-theorem baby_Sard:
+theorem%important baby_Sard:
fixes f :: "real^'m::{finite,wellorder} \<Rightarrow> real^'n::{finite,wellorder}"
assumes mlen: "CARD('m) \<le> CARD('n)"
and der: "\<And>x. x \<in> S \<Longrightarrow> (f has_derivative f' x) (at x within S)"
and rank: "\<And>x. x \<in> S \<Longrightarrow> rank(matrix(f' x)) < CARD('n)"
shows "negligible(f ` S)"
-proof -
+proof%unimportant -
let ?U = "\<lambda>n. {x \<in> S. norm(x) \<le> n \<and> onorm(f' x) \<le> real n}"
have "\<And>x. x \<in> S \<Longrightarrow> \<exists>n. norm x \<le> real n \<and> onorm (f' x) \<le> real n"
by (meson linear order_trans real_arch_simple)
@@ -2750,9 +2750,9 @@
qed
-subsection\<open>A one-way version of change-of-variables not assuming injectivity. \<close>
+subsection%important\<open>A one-way version of change-of-variables not assuming injectivity. \<close>
-lemma integral_on_image_ubound_weak:
+lemma%important integral_on_image_ubound_weak:
fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real"
assumes S: "S \<in> sets lebesgue"
and f: "f \<in> borel_measurable (lebesgue_on (g ` S))"
@@ -2763,7 +2763,7 @@
shows "f integrable_on (g ` S) \<and>
integral (g ` S) f \<le> integral S (\<lambda>x. \<bar>det (matrix (g' x))\<bar> * f(g x))"
(is "_ \<and> _ \<le> ?b")
-proof -
+proof%unimportant -
let ?D = "\<lambda>x. \<bar>det (matrix (g' x))\<bar>"
have cont_g: "continuous_on S g"
using der_g has_derivative_continuous_on by blast
@@ -2941,14 +2941,14 @@
qed
-lemma integral_on_image_ubound_nonneg:
+lemma%important integral_on_image_ubound_nonneg:
fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real"
assumes nonneg_fg: "\<And>x. x \<in> S \<Longrightarrow> 0 \<le> f(g x)"
and der_g: "\<And>x. x \<in> S \<Longrightarrow> (g has_derivative g' x) (at x within S)"
and intS: "(\<lambda>x. \<bar>det (matrix (g' x))\<bar> * f(g x)) integrable_on S"
shows "f integrable_on (g ` S) \<and> integral (g ` S) f \<le> integral S (\<lambda>x. \<bar>det (matrix (g' x))\<bar> * f(g x))"
(is "_ \<and> _ \<le> ?b")
-proof -
+proof%unimportant -
let ?D = "\<lambda>x. det (matrix (g' x))"
define S' where "S' \<equiv> {x \<in> S. ?D x * f(g x) \<noteq> 0}"
then have der_gS': "\<And>x. x \<in> S' \<Longrightarrow> (g has_derivative g' x) (at x within S')"
@@ -3071,7 +3071,7 @@
qed
-lemma absolutely_integrable_on_image_real:
+lemma%unimportant absolutely_integrable_on_image_real:
fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real" and g :: "real^'n::_ \<Rightarrow> real^'n::_"
assumes der_g: "\<And>x. x \<in> S \<Longrightarrow> (g has_derivative g' x) (at x within S)"
and intS: "(\<lambda>x. \<bar>det (matrix (g' x))\<bar> * f(g x)) absolutely_integrable_on S"
@@ -3146,29 +3146,29 @@
qed
-proposition absolutely_integrable_on_image:
+proposition%important absolutely_integrable_on_image:
fixes f :: "real^'m::{finite,wellorder} \<Rightarrow> real^'n" and g :: "real^'m::_ \<Rightarrow> real^'m::_"
assumes der_g: "\<And>x. x \<in> S \<Longrightarrow> (g has_derivative g' x) (at x within S)"
and intS: "(\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R f(g x)) absolutely_integrable_on S"
shows "f absolutely_integrable_on (g ` S)"
apply (rule absolutely_integrable_componentwise [OF absolutely_integrable_on_image_real [OF der_g]])
- using absolutely_integrable_component [OF intS] by auto
+ using%unimportant absolutely_integrable_component [OF intS] by%unimportant auto
-proposition integral_on_image_ubound:
+proposition%important integral_on_image_ubound:
fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real" and g :: "real^'n::_ \<Rightarrow> real^'n::_"
assumes"\<And>x. x \<in> S \<Longrightarrow> 0 \<le> f(g x)"
and "\<And>x. x \<in> S \<Longrightarrow> (g has_derivative g' x) (at x within S)"
and "(\<lambda>x. \<bar>det (matrix (g' x))\<bar> * f(g x)) integrable_on S"
shows "integral (g ` S) f \<le> integral S (\<lambda>x. \<bar>det (matrix (g' x))\<bar> * f(g x))"
- using integral_on_image_ubound_nonneg [OF assms] by simp
+ using%unimportant integral_on_image_ubound_nonneg [OF assms] by%unimportant simp
-subsection\<open>Change-of-variables theorem\<close>
+subsection%important\<open>Change-of-variables theorem\<close>
text\<open>The classic change-of-variables theorem. We have two versions with quite general hypotheses,
the first that the transforming function has a continuous inverse, the second that the base set is
Lebesgue measurable.\<close>
-lemma cov_invertible_nonneg_le:
+lemma%unimportant cov_invertible_nonneg_le:
fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real" and g :: "real^'n::_ \<Rightarrow> real^'n::_"
assumes der_g: "\<And>x. x \<in> S \<Longrightarrow> (g has_derivative g' x) (at x within S)"
and der_h: "\<And>y. y \<in> T \<Longrightarrow> (h has_derivative h' y) (at y within T)"
@@ -3239,7 +3239,7 @@
qed
-lemma cov_invertible_nonneg_eq:
+lemma%unimportant cov_invertible_nonneg_eq:
fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real" and g :: "real^'n::_ \<Rightarrow> real^'n::_"
assumes "\<And>x. x \<in> S \<Longrightarrow> (g has_derivative g' x) (at x within S)"
and "\<And>y. y \<in> T \<Longrightarrow> (h has_derivative h' y) (at y within T)"
@@ -3252,7 +3252,7 @@
by (simp add: has_integral_iff) (meson eq_iff)
-lemma cov_invertible_real:
+lemma%unimportant cov_invertible_real:
fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real" and g :: "real^'n::_ \<Rightarrow> real^'n::_"
assumes der_g: "\<And>x. x \<in> S \<Longrightarrow> (g has_derivative g' x) (at x within S)"
and der_h: "\<And>y. y \<in> T \<Longrightarrow> (h has_derivative h' y) (at y within T)"
@@ -3418,7 +3418,7 @@
qed
-lemma cv_inv_version3:
+lemma%unimportant cv_inv_version3:
fixes f :: "real^'m::{finite,wellorder} \<Rightarrow> real^'n" and g :: "real^'m::_ \<Rightarrow> real^'m::_"
assumes der_g: "\<And>x. x \<in> S \<Longrightarrow> (g has_derivative g' x) (at x within S)"
and der_h: "\<And>y. y \<in> T \<Longrightarrow> (h has_derivative h' y) (at y within T)"
@@ -3444,7 +3444,7 @@
qed
-lemma cv_inv_version4:
+lemma%unimportant cv_inv_version4:
fixes f :: "real^'m::{finite,wellorder} \<Rightarrow> real^'n" and g :: "real^'m::_ \<Rightarrow> real^'m::_"
assumes der_g: "\<And>x. x \<in> S \<Longrightarrow> (g has_derivative g' x) (at x within S) \<and> invertible(matrix(g' x))"
and hg: "\<And>x. x \<in> S \<Longrightarrow> continuous_on (g ` S) h \<and> h(g x) = x"
@@ -3469,7 +3469,7 @@
qed
-proposition has_absolute_integral_change_of_variables_invertible:
+proposition%important has_absolute_integral_change_of_variables_invertible:
fixes f :: "real^'m::{finite,wellorder} \<Rightarrow> real^'n" and g :: "real^'m::_ \<Rightarrow> real^'m::_"
assumes der_g: "\<And>x. x \<in> S \<Longrightarrow> (g has_derivative g' x) (at x within S)"
and hg: "\<And>x. x \<in> S \<Longrightarrow> h(g x) = x"
@@ -3477,7 +3477,7 @@
shows "(\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R f(g x)) absolutely_integrable_on S \<and> integral S (\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R f(g x)) = b \<longleftrightarrow>
f absolutely_integrable_on (g ` S) \<and> integral (g ` S) f = b"
(is "?lhs = ?rhs")
-proof -
+proof%unimportant -
let ?S = "{x \<in> S. invertible (matrix (g' x))}" and ?D = "\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R f(g x)"
have *: "?D absolutely_integrable_on ?S \<and> integral ?S ?D = b
\<longleftrightarrow> f absolutely_integrable_on (g ` ?S) \<and> integral (g ` ?S) f = b"
@@ -3513,7 +3513,7 @@
-lemma has_absolute_integral_change_of_variables_compact:
+lemma%unimportant has_absolute_integral_change_of_variables_compact:
fixes f :: "real^'m::{finite,wellorder} \<Rightarrow> real^'n" and g :: "real^'m::_ \<Rightarrow> real^'m::_"
assumes "compact S"
and der_g: "\<And>x. x \<in> S \<Longrightarrow> (g has_derivative g' x) (at x within S)"
@@ -3531,7 +3531,7 @@
qed
-lemma has_absolute_integral_change_of_variables_compact_family:
+lemma%unimportant has_absolute_integral_change_of_variables_compact_family:
fixes f :: "real^'m::{finite,wellorder} \<Rightarrow> real^'n" and g :: "real^'m::_ \<Rightarrow> real^'m::_"
assumes compact: "\<And>n::nat. compact (F n)"
and der_g: "\<And>x. x \<in> (\<Union>n. F n) \<Longrightarrow> (g has_derivative g' x) (at x within (\<Union>n. F n))"
@@ -3710,7 +3710,7 @@
qed
-proposition has_absolute_integral_change_of_variables:
+proposition%important has_absolute_integral_change_of_variables:
fixes f :: "real^'m::{finite,wellorder} \<Rightarrow> real^'n" and g :: "real^'m::_ \<Rightarrow> real^'m::_"
assumes S: "S \<in> sets lebesgue"
and der_g: "\<And>x. x \<in> S \<Longrightarrow> (g has_derivative g' x) (at x within S)"
@@ -3718,7 +3718,7 @@
shows "(\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R f(g x)) absolutely_integrable_on S \<and>
integral S (\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R f(g x)) = b
\<longleftrightarrow> f absolutely_integrable_on (g ` S) \<and> integral (g ` S) f = b"
-proof -
+proof%unimportant -
obtain C N where "fsigma C" "negligible N" and CNS: "C \<union> N = S" and "disjnt C N"
using lebesgue_set_almost_fsigma [OF S] .
then obtain F :: "nat \<Rightarrow> (real^'m::_) set"
@@ -3772,16 +3772,16 @@
qed
-corollary absolutely_integrable_change_of_variables:
+corollary%important absolutely_integrable_change_of_variables:
fixes f :: "real^'m::{finite,wellorder} \<Rightarrow> real^'n" and g :: "real^'m::_ \<Rightarrow> real^'m::_"
assumes "S \<in> sets lebesgue"
and "\<And>x. x \<in> S \<Longrightarrow> (g has_derivative g' x) (at x within S)"
and "inj_on g S"
shows "f absolutely_integrable_on (g ` S)
\<longleftrightarrow> (\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R f(g x)) absolutely_integrable_on S"
- using assms has_absolute_integral_change_of_variables by blast
+ using%unimportant assms has_absolute_integral_change_of_variables by%unimportant blast
-corollary integral_change_of_variables:
+corollary%important integral_change_of_variables:
fixes f :: "real^'m::{finite,wellorder} \<Rightarrow> real^'n" and g :: "real^'m::_ \<Rightarrow> real^'m::_"
assumes S: "S \<in> sets lebesgue"
and der_g: "\<And>x. x \<in> S \<Longrightarrow> (g has_derivative g' x) (at x within S)"
@@ -3789,10 +3789,10 @@
and disj: "(f absolutely_integrable_on (g ` S) \<or>
(\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R f(g x)) absolutely_integrable_on S)"
shows "integral (g ` S) f = integral S (\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R f(g x))"
- using has_absolute_integral_change_of_variables [OF S der_g inj] disj
- by blast
+ using%unimportant has_absolute_integral_change_of_variables [OF S der_g inj] disj
+ by%unimportant blast
-lemma has_absolute_integral_change_of_variables_1:
+lemma%unimportant has_absolute_integral_change_of_variables_1:
fixes f :: "real \<Rightarrow> real^'n::{finite,wellorder}" and g :: "real \<Rightarrow> real"
assumes S: "S \<in> sets lebesgue"
and der_g: "\<And>x. x \<in> S \<Longrightarrow> (g has_vector_derivative g' x) (at x within S)"
@@ -3831,19 +3831,19 @@
qed
-corollary absolutely_integrable_change_of_variables_1:
+corollary%important absolutely_integrable_change_of_variables_1:
fixes f :: "real \<Rightarrow> real^'n::{finite,wellorder}" and g :: "real \<Rightarrow> real"
assumes S: "S \<in> sets lebesgue"
and der_g: "\<And>x. x \<in> S \<Longrightarrow> (g has_vector_derivative g' x) (at x within S)"
and inj: "inj_on g S"
shows "(f absolutely_integrable_on g ` S \<longleftrightarrow>
(\<lambda>x. \<bar>g' x\<bar> *\<^sub>R f(g x)) absolutely_integrable_on S)"
- using has_absolute_integral_change_of_variables_1 [OF assms] by auto
+ using%unimportant has_absolute_integral_change_of_variables_1 [OF assms] by%unimportant auto
-subsection\<open>Change of variables for integrals: special case of linear function\<close>
+subsection%important\<open>Change of variables for integrals: special case of linear function\<close>
-lemma has_absolute_integral_change_of_variables_linear:
+lemma%unimportant has_absolute_integral_change_of_variables_linear:
fixes f :: "real^'m::{finite,wellorder} \<Rightarrow> real^'n" and g :: "real^'m::_ \<Rightarrow> real^'m::_"
assumes "linear g"
shows "(\<lambda>x. \<bar>det (matrix g)\<bar> *\<^sub>R f(g x)) absolutely_integrable_on S \<and>
@@ -3868,14 +3868,14 @@
qed (use h in auto)
qed
-lemma absolutely_integrable_change_of_variables_linear:
+lemma%unimportant absolutely_integrable_change_of_variables_linear:
fixes f :: "real^'m::{finite,wellorder} \<Rightarrow> real^'n" and g :: "real^'m::_ \<Rightarrow> real^'m::_"
assumes "linear g"
shows "(\<lambda>x. \<bar>det (matrix g)\<bar> *\<^sub>R f(g x)) absolutely_integrable_on S
\<longleftrightarrow> f absolutely_integrable_on (g ` S)"
using assms has_absolute_integral_change_of_variables_linear by blast
-lemma absolutely_integrable_on_linear_image:
+lemma%unimportant absolutely_integrable_on_linear_image:
fixes f :: "real^'m::{finite,wellorder} \<Rightarrow> real^'n" and g :: "real^'m::_ \<Rightarrow> real^'m::_"
assumes "linear g"
shows "f absolutely_integrable_on (g ` S)
@@ -3883,12 +3883,12 @@
unfolding assms absolutely_integrable_change_of_variables_linear [OF assms, symmetric] absolutely_integrable_on_scaleR_iff
by (auto simp: set_integrable_def)
-lemma integral_change_of_variables_linear:
+lemma%important integral_change_of_variables_linear:
fixes f :: "real^'m::{finite,wellorder} \<Rightarrow> real^'n" and g :: "real^'m::_ \<Rightarrow> real^'m::_"
assumes "linear g"
and "f absolutely_integrable_on (g ` S) \<or> (f \<circ> g) absolutely_integrable_on S"
shows "integral (g ` S) f = \<bar>det (matrix g)\<bar> *\<^sub>R integral S (f \<circ> g)"
-proof -
+proof%unimportant -
have "((\<lambda>x. \<bar>det (matrix g)\<bar> *\<^sub>R f (g x)) absolutely_integrable_on S) \<or> (f absolutely_integrable_on g ` S)"
using absolutely_integrable_on_linear_image assms by blast
moreover
@@ -3900,20 +3900,20 @@
by blast
qed
-subsection\<open>Change of variable for measure\<close>
+subsection%important\<open>Change of variable for measure\<close>
-lemma has_measure_differentiable_image:
+lemma%unimportant has_measure_differentiable_image:
fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real^'n::_"
assumes "S \<in> sets lebesgue"
and "\<And>x. x \<in> S \<Longrightarrow> (f has_derivative f' x) (at x within S)"
and "inj_on f S"
shows "f ` S \<in> lmeasurable \<and> measure lebesgue (f ` S) = m
\<longleftrightarrow> ((\<lambda>x. \<bar>det (matrix (f' x))\<bar>) has_integral m) S"
- using has_absolute_integral_change_of_variables [OF assms, of "\<lambda>x. (1::real^1)" "vec m"]
- unfolding absolutely_integrable_on_1_iff integral_on_1_eq integrable_on_1_iff absolutely_integrable_on_def
- by (auto simp: has_integral_iff lmeasurable_iff_integrable_on lmeasure_integral)
+ using%unimportant has_absolute_integral_change_of_variables [OF assms, of "\<lambda>x. (1::real^1)" "vec m"]
+ unfolding%unimportant absolutely_integrable_on_1_iff integral_on_1_eq integrable_on_1_iff absolutely_integrable_on_def
+ by%unimportant (auto simp: has_integral_iff lmeasurable_iff_integrable_on lmeasure_integral)
-lemma measurable_differentiable_image_eq:
+lemma%unimportant measurable_differentiable_image_eq:
fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real^'n::_"
assumes "S \<in> sets lebesgue"
and "\<And>x. x \<in> S \<Longrightarrow> (f has_derivative f' x) (at x within S)"
@@ -3922,23 +3922,23 @@
using has_measure_differentiable_image [OF assms]
by blast
-lemma measurable_differentiable_image_alt:
+lemma%important measurable_differentiable_image_alt:
fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real^'n::_"
assumes "S \<in> sets lebesgue"
and "\<And>x. x \<in> S \<Longrightarrow> (f has_derivative f' x) (at x within S)"
and "inj_on f S"
shows "f ` S \<in> lmeasurable \<longleftrightarrow> (\<lambda>x. \<bar>det (matrix (f' x))\<bar>) absolutely_integrable_on S"
- using measurable_differentiable_image_eq [OF assms]
- by (simp only: absolutely_integrable_on_iff_nonneg)
+ using%unimportant measurable_differentiable_image_eq [OF assms]
+ by%unimportant (simp only: absolutely_integrable_on_iff_nonneg)
-lemma measure_differentiable_image_eq:
+lemma%important measure_differentiable_image_eq:
fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real^'n::_"
assumes S: "S \<in> sets lebesgue"
and der_f: "\<And>x. x \<in> S \<Longrightarrow> (f has_derivative f' x) (at x within S)"
and inj: "inj_on f S"
and intS: "(\<lambda>x. \<bar>det (matrix (f' x))\<bar>) integrable_on S"
shows "measure lebesgue (f ` S) = integral S (\<lambda>x. \<bar>det (matrix (f' x))\<bar>)"
- using measurable_differentiable_image_eq [OF S der_f inj]
- assms has_measure_differentiable_image by blast
+ using%unimportant measurable_differentiable_image_eq [OF S der_f inj]
+ assms has_measure_differentiable_image by%unimportant blast
end
--- a/src/HOL/Analysis/Cross3.thy Mon Aug 27 22:58:36 2018 +0200
+++ b/src/HOL/Analysis/Cross3.thy Tue Aug 28 13:28:39 2018 +0100
@@ -13,7 +13,7 @@
context includes no_Set_Product_syntax
begin \<comment>\<open>locally disable syntax for set product, to avoid warnings\<close>
-definition cross3 :: "[real^3, real^3] \<Rightarrow> real^3" (infixr "\<times>" 80)
+definition%important cross3 :: "[real^3, real^3] \<Rightarrow> real^3" (infixr "\<times>" 80)
where "a \<times> b \<equiv>
vector [a$2 * b$3 - a$3 * b$2,
a$3 * b$1 - a$1 * b$3,
@@ -33,76 +33,76 @@
unbundle cross3_syntax
-subsection\<open> Basic lemmas.\<close>
+subsection%important\<open> Basic lemmas.\<close>
lemmas cross3_simps = cross3_def inner_vec_def sum_3 det_3 vec_eq_iff vector_def algebra_simps
-lemma dot_cross_self: "x \<bullet> (x \<times> y) = 0" "x \<bullet> (y \<times> x) = 0" "(x \<times> y) \<bullet> y = 0" "(y \<times> x) \<bullet> y = 0"
+lemma%unimportant dot_cross_self: "x \<bullet> (x \<times> y) = 0" "x \<bullet> (y \<times> x) = 0" "(x \<times> y) \<bullet> y = 0" "(y \<times> x) \<bullet> y = 0"
by (simp_all add: orthogonal_def cross3_simps)
-lemma orthogonal_cross: "orthogonal (x \<times> y) x" "orthogonal (x \<times> y) y"
+lemma%unimportant orthogonal_cross: "orthogonal (x \<times> y) x" "orthogonal (x \<times> y) y"
"orthogonal y (x \<times> y)" "orthogonal (x \<times> y) x"
by (simp_all add: orthogonal_def dot_cross_self)
-lemma cross_zero_left [simp]: "0 \<times> x = 0" and cross_zero_right [simp]: "x \<times> 0 = 0" for x::"real^3"
+lemma%unimportant cross_zero_left [simp]: "0 \<times> x = 0" and cross_zero_right [simp]: "x \<times> 0 = 0" for x::"real^3"
by (simp_all add: cross3_simps)
-lemma cross_skew: "(x \<times> y) = -(y \<times> x)" for x::"real^3"
+lemma%unimportant cross_skew: "(x \<times> y) = -(y \<times> x)" for x::"real^3"
by (simp add: cross3_simps)
-lemma cross_refl [simp]: "x \<times> x = 0" for x::"real^3"
+lemma%unimportant cross_refl [simp]: "x \<times> x = 0" for x::"real^3"
by (simp add: cross3_simps)
-lemma cross_add_left: "(x + y) \<times> z = (x \<times> z) + (y \<times> z)" for x::"real^3"
+lemma%unimportant cross_add_left: "(x + y) \<times> z = (x \<times> z) + (y \<times> z)" for x::"real^3"
by (simp add: cross3_simps)
-lemma cross_add_right: "x \<times> (y + z) = (x \<times> y) + (x \<times> z)" for x::"real^3"
+lemma%unimportant cross_add_right: "x \<times> (y + z) = (x \<times> y) + (x \<times> z)" for x::"real^3"
by (simp add: cross3_simps)
-lemma cross_mult_left: "(c *\<^sub>R x) \<times> y = c *\<^sub>R (x \<times> y)" for x::"real^3"
+lemma%unimportant cross_mult_left: "(c *\<^sub>R x) \<times> y = c *\<^sub>R (x \<times> y)" for x::"real^3"
by (simp add: cross3_simps)
-lemma cross_mult_right: "x \<times> (c *\<^sub>R y) = c *\<^sub>R (x \<times> y)" for x::"real^3"
+lemma%unimportant cross_mult_right: "x \<times> (c *\<^sub>R y) = c *\<^sub>R (x \<times> y)" for x::"real^3"
by (simp add: cross3_simps)
-lemma cross_minus_left [simp]: "(-x) \<times> y = - (x \<times> y)" for x::"real^3"
+lemma%unimportant cross_minus_left [simp]: "(-x) \<times> y = - (x \<times> y)" for x::"real^3"
by (simp add: cross3_simps)
-lemma cross_minus_right [simp]: "x \<times> -y = - (x \<times> y)" for x::"real^3"
+lemma%unimportant cross_minus_right [simp]: "x \<times> -y = - (x \<times> y)" for x::"real^3"
by (simp add: cross3_simps)
-lemma left_diff_distrib: "(x - y) \<times> z = x \<times> z - y \<times> z" for x::"real^3"
+lemma%unimportant left_diff_distrib: "(x - y) \<times> z = x \<times> z - y \<times> z" for x::"real^3"
by (simp add: cross3_simps)
-lemma right_diff_distrib: "x \<times> (y - z) = x \<times> y - x \<times> z" for x::"real^3"
+lemma%unimportant right_diff_distrib: "x \<times> (y - z) = x \<times> y - x \<times> z" for x::"real^3"
by (simp add: cross3_simps)
hide_fact (open) left_diff_distrib right_diff_distrib
-lemma Jacobi: "x \<times> (y \<times> z) + y \<times> (z \<times> x) + z \<times> (x \<times> y) = 0" for x::"real^3"
- by (simp add: cross3_simps)
+lemma%important Jacobi: "x \<times> (y \<times> z) + y \<times> (z \<times> x) + z \<times> (x \<times> y) = 0" for x::"real^3"
+ by%unimportant (simp add: cross3_simps)
-lemma Lagrange: "x \<times> (y \<times> z) = (x \<bullet> z) *\<^sub>R y - (x \<bullet> y) *\<^sub>R z"
- by (simp add: cross3_simps) (metis (full_types) exhaust_3)
+lemma%important Lagrange: "x \<times> (y \<times> z) = (x \<bullet> z) *\<^sub>R y - (x \<bullet> y) *\<^sub>R z"
+ by%unimportant (simp add: cross3_simps) (metis (full_types) exhaust_3)
-lemma cross_triple: "(x \<times> y) \<bullet> z = (y \<times> z) \<bullet> x"
+lemma%unimportant cross_triple: "(x \<times> y) \<bullet> z = (y \<times> z) \<bullet> x"
by (simp add: cross3_def inner_vec_def sum_3 vec_eq_iff algebra_simps)
-lemma cross_components:
+lemma%unimportant cross_components:
"(x \<times> y)$1 = x$2 * y$3 - y$2 * x$3" "(x \<times> y)$2 = x$3 * y$1 - y$3 * x$1" "(x \<times> y)$3 = x$1 * y$2 - y$1 * x$2"
by (simp_all add: cross3_def inner_vec_def sum_3 vec_eq_iff algebra_simps)
-lemma cross_basis: "(axis 1 1) \<times> (axis 2 1) = axis 3 1" "(axis 2 1) \<times> (axis 1 1) = -(axis 3 1)"
+lemma%unimportant cross_basis: "(axis 1 1) \<times> (axis 2 1) = axis 3 1" "(axis 2 1) \<times> (axis 1 1) = -(axis 3 1)"
"(axis 2 1) \<times> (axis 3 1) = axis 1 1" "(axis 3 1) \<times> (axis 2 1) = -(axis 1 1)"
"(axis 3 1) \<times> (axis 1 1) = axis 2 1" "(axis 1 1) \<times> (axis 3 1) = -(axis 2 1)"
using exhaust_3
by (force simp add: axis_def cross3_simps)+
-lemma cross_basis_nonzero:
+lemma%unimportant cross_basis_nonzero:
"u \<noteq> 0 \<Longrightarrow> ~(u \<times> axis 1 1 = 0) \<or> ~(u \<times> axis 2 1 = 0) \<or> ~(u \<times> axis 3 1 = 0)"
by (clarsimp simp add: axis_def cross3_simps) (metis vector_3 exhaust_3)
-lemma cross_dot_cancel:
+lemma%unimportant cross_dot_cancel:
fixes x::"real^3"
assumes deq: "x \<bullet> y = x \<bullet> z" and veq: "x \<times> y = x \<times> z" and x: "x \<noteq> 0"
shows "y = z"
@@ -116,25 +116,25 @@
using eq_iff_diff_eq_0 by blast
qed
-lemma norm_cross_dot: "(norm (x \<times> y))\<^sup>2 + (x \<bullet> y)\<^sup>2 = (norm x * norm y)\<^sup>2"
+lemma%unimportant norm_cross_dot: "(norm (x \<times> y))\<^sup>2 + (x \<bullet> y)\<^sup>2 = (norm x * norm y)\<^sup>2"
unfolding power2_norm_eq_inner power_mult_distrib
by (simp add: cross3_simps power2_eq_square)
-lemma dot_cross_det: "x \<bullet> (y \<times> z) = det(vector[x,y,z])"
+lemma%unimportant dot_cross_det: "x \<bullet> (y \<times> z) = det(vector[x,y,z])"
by (simp add: cross3_simps)
-lemma cross_cross_det: "(w \<times> x) \<times> (y \<times> z) = det(vector[w,x,z]) *\<^sub>R y - det(vector[w,x,y]) *\<^sub>R z"
+lemma%unimportant cross_cross_det: "(w \<times> x) \<times> (y \<times> z) = det(vector[w,x,z]) *\<^sub>R y - det(vector[w,x,y]) *\<^sub>R z"
using exhaust_3 by (force simp add: cross3_simps)
-lemma dot_cross: "(w \<times> x) \<bullet> (y \<times> z) = (w \<bullet> y) * (x \<bullet> z) - (w \<bullet> z) * (x \<bullet> y)"
- by (force simp add: cross3_simps)
+lemma%important dot_cross: "(w \<times> x) \<bullet> (y \<times> z) = (w \<bullet> y) * (x \<bullet> z) - (w \<bullet> z) * (x \<bullet> y)"
+ by%unimportant (force simp add: cross3_simps)
-lemma norm_cross: "(norm (x \<times> y))\<^sup>2 = (norm x)\<^sup>2 * (norm y)\<^sup>2 - (x \<bullet> y)\<^sup>2"
+lemma%unimportant norm_cross: "(norm (x \<times> y))\<^sup>2 = (norm x)\<^sup>2 * (norm y)\<^sup>2 - (x \<bullet> y)\<^sup>2"
unfolding power2_norm_eq_inner power_mult_distrib
by (simp add: cross3_simps power2_eq_square)
-lemma cross_eq_0: "x \<times> y = 0 \<longleftrightarrow> collinear{0,x,y}"
-proof -
+lemma%important cross_eq_0: "x \<times> y = 0 \<longleftrightarrow> collinear{0,x,y}"
+proof%unimportant -
have "x \<times> y = 0 \<longleftrightarrow> norm (x \<times> y) = 0"
by simp
also have "... \<longleftrightarrow> (norm x * norm y)\<^sup>2 = (x \<bullet> y)\<^sup>2"
@@ -147,11 +147,11 @@
finally show ?thesis .
qed
-lemma cross_eq_self: "x \<times> y = x \<longleftrightarrow> x = 0" "x \<times> y = y \<longleftrightarrow> y = 0"
+lemma%unimportant cross_eq_self: "x \<times> y = x \<longleftrightarrow> x = 0" "x \<times> y = y \<longleftrightarrow> y = 0"
apply (metis cross_zero_left dot_cross_self(1) inner_eq_zero_iff)
by (metis cross_zero_right dot_cross_self(2) inner_eq_zero_iff)
-lemma norm_and_cross_eq_0:
+lemma%unimportant norm_and_cross_eq_0:
"x \<bullet> y = 0 \<and> x \<times> y = 0 \<longleftrightarrow> x = 0 \<or> y = 0" (is "?lhs = ?rhs")
proof
assume ?lhs
@@ -159,7 +159,7 @@
by (metis cross_dot_cancel cross_zero_right inner_zero_right)
qed auto
-lemma bilinear_cross: "bilinear(\<times>)"
+lemma%unimportant bilinear_cross: "bilinear(\<times>)"
apply (auto simp add: bilinear_def linear_def)
apply unfold_locales
apply (simp add: cross_add_right)
@@ -168,33 +168,33 @@
apply (simp add: cross_mult_left)
done
-subsection\<open>Preservation by rotation, or other orthogonal transformation up to sign.\<close>
+subsection%important \<open>Preservation by rotation, or other orthogonal transformation up to sign.\<close>
-lemma cross_matrix_mult: "transpose A *v ((A *v x) \<times> (A *v y)) = det A *\<^sub>R (x \<times> y)"
+lemma%unimportant cross_matrix_mult: "transpose A *v ((A *v x) \<times> (A *v y)) = det A *\<^sub>R (x \<times> y)"
apply (simp add: vec_eq_iff )
apply (simp add: vector_matrix_mult_def matrix_vector_mult_def forall_3 cross3_simps)
done
-lemma cross_orthogonal_matrix:
+lemma%important cross_orthogonal_matrix:
assumes "orthogonal_matrix A"
shows "(A *v x) \<times> (A *v y) = det A *\<^sub>R (A *v (x \<times> y))"
-proof -
+proof%unimportant -
have "mat 1 = transpose (A ** transpose A)"
by (metis (no_types) assms orthogonal_matrix_def transpose_mat)
then show ?thesis
by (metis (no_types) vector_matrix_mul_rid vector_transpose_matrix cross_matrix_mult matrix_vector_mul_assoc matrix_vector_mult_scaleR)
qed
-lemma cross_rotation_matrix: "rotation_matrix A \<Longrightarrow> (A *v x) \<times> (A *v y) = A *v (x \<times> y)"
+lemma%unimportant cross_rotation_matrix: "rotation_matrix A \<Longrightarrow> (A *v x) \<times> (A *v y) = A *v (x \<times> y)"
by (simp add: rotation_matrix_def cross_orthogonal_matrix)
-lemma cross_rotoinversion_matrix: "rotoinversion_matrix A \<Longrightarrow> (A *v x) \<times> (A *v y) = - A *v (x \<times> y)"
+lemma%unimportant cross_rotoinversion_matrix: "rotoinversion_matrix A \<Longrightarrow> (A *v x) \<times> (A *v y) = - A *v (x \<times> y)"
by (simp add: rotoinversion_matrix_def cross_orthogonal_matrix scaleR_matrix_vector_assoc)
-lemma cross_orthogonal_transformation:
+lemma%important cross_orthogonal_transformation:
assumes "orthogonal_transformation f"
shows "(f x) \<times> (f y) = det(matrix f) *\<^sub>R f(x \<times> y)"
-proof -
+proof%unimportant -
have orth: "orthogonal_matrix (matrix f)"
using assms orthogonal_transformation_matrix by blast
have "matrix f *v z = f z" for z
@@ -203,20 +203,20 @@
by simp
qed
-lemma cross_linear_image:
+lemma%unimportant cross_linear_image:
"\<lbrakk>linear f; \<And>x. norm(f x) = norm x; det(matrix f) = 1\<rbrakk>
\<Longrightarrow> (f x) \<times> (f y) = f(x \<times> y)"
by (simp add: cross_orthogonal_transformation orthogonal_transformation)
-subsection\<open>Continuity\<close>
+subsection%unimportant \<open>Continuity\<close>
-lemma continuous_cross: "\<lbrakk>continuous F f; continuous F g\<rbrakk> \<Longrightarrow> continuous F (\<lambda>x. (f x) \<times> (g x))"
+lemma%unimportant continuous_cross: "\<lbrakk>continuous F f; continuous F g\<rbrakk> \<Longrightarrow> continuous F (\<lambda>x. (f x) \<times> (g x))"
apply (subst continuous_componentwise)
apply (clarsimp simp add: cross3_simps)
apply (intro continuous_intros; simp)
done
-lemma continuous_on_cross:
+lemma%unimportant continuous_on_cross:
fixes f :: "'a::t2_space \<Rightarrow> real^3"
shows "\<lbrakk>continuous_on S f; continuous_on S g\<rbrakk> \<Longrightarrow> continuous_on S (\<lambda>x. (f x) \<times> (g x))"
by (simp add: continuous_on_eq_continuous_within continuous_cross)
--- a/src/HOL/Analysis/Determinants.thy Mon Aug 27 22:58:36 2018 +0200
+++ b/src/HOL/Analysis/Determinants.thy Tue Aug 28 13:28:39 2018 +0100
@@ -10,40 +10,40 @@
"HOL-Library.Permutations"
begin
-subsection \<open>Trace\<close>
+subsection%important \<open>Trace\<close>
-definition trace :: "'a::semiring_1^'n^'n \<Rightarrow> 'a"
+definition%important trace :: "'a::semiring_1^'n^'n \<Rightarrow> 'a"
where "trace A = sum (\<lambda>i. ((A$i)$i)) (UNIV::'n set)"
-lemma trace_0: "trace (mat 0) = 0"
+lemma%unimportant trace_0: "trace (mat 0) = 0"
by (simp add: trace_def mat_def)
-lemma trace_I: "trace (mat 1 :: 'a::semiring_1^'n^'n) = of_nat(CARD('n))"
+lemma%unimportant trace_I: "trace (mat 1 :: 'a::semiring_1^'n^'n) = of_nat(CARD('n))"
by (simp add: trace_def mat_def)
-lemma trace_add: "trace ((A::'a::comm_semiring_1^'n^'n) + B) = trace A + trace B"
+lemma%unimportant trace_add: "trace ((A::'a::comm_semiring_1^'n^'n) + B) = trace A + trace B"
by (simp add: trace_def sum.distrib)
-lemma trace_sub: "trace ((A::'a::comm_ring_1^'n^'n) - B) = trace A - trace B"
+lemma%unimportant trace_sub: "trace ((A::'a::comm_ring_1^'n^'n) - B) = trace A - trace B"
by (simp add: trace_def sum_subtractf)
-lemma trace_mul_sym: "trace ((A::'a::comm_semiring_1^'n^'m) ** B) = trace (B**A)"
+lemma%important trace_mul_sym: "trace ((A::'a::comm_semiring_1^'n^'m) ** B) = trace (B**A)"
apply (simp add: trace_def matrix_matrix_mult_def)
apply (subst sum.swap)
apply (simp add: mult.commute)
done
-subsubsection \<open>Definition of determinant\<close>
+subsubsection%important \<open>Definition of determinant\<close>
-definition det:: "'a::comm_ring_1^'n^'n \<Rightarrow> 'a" where
+definition%important det:: "'a::comm_ring_1^'n^'n \<Rightarrow> 'a" where
"det A =
sum (\<lambda>p. of_int (sign p) * prod (\<lambda>i. A$i$p i) (UNIV :: 'n set))
{p. p permutes (UNIV :: 'n set)}"
text \<open>Basic determinant properties\<close>
-lemma det_transpose [simp]: "det (transpose A) = det (A::'a::comm_ring_1 ^'n^'n)"
-proof -
+lemma%important det_transpose [simp]: "det (transpose A) = det (A::'a::comm_ring_1 ^'n^'n)"
+proof%unimportant -
let ?di = "\<lambda>A i j. A$i$j"
let ?U = "(UNIV :: 'n set)"
have fU: "finite ?U" by simp
@@ -80,7 +80,7 @@
by (subst sum_permutations_inverse) (blast intro: sum.cong)
qed
-lemma det_lowerdiagonal:
+lemma%unimportant det_lowerdiagonal:
fixes A :: "'a::comm_ring_1^('n::{finite,wellorder})^('n::{finite,wellorder})"
assumes ld: "\<And>i j. i < j \<Longrightarrow> A$i$j = 0"
shows "det A = prod (\<lambda>i. A$i$i) (UNIV:: 'n set)"
@@ -107,11 +107,11 @@
unfolding det_def by (simp add: sign_id)
qed
-lemma det_upperdiagonal:
+lemma%important det_upperdiagonal:
fixes A :: "'a::comm_ring_1^'n::{finite,wellorder}^'n::{finite,wellorder}"
assumes ld: "\<And>i j. i > j \<Longrightarrow> A$i$j = 0"
shows "det A = prod (\<lambda>i. A$i$i) (UNIV:: 'n set)"
-proof -
+proof%unimportant -
let ?U = "UNIV:: 'n set"
let ?PU = "{p. p permutes ?U}"
let ?pp = "(\<lambda>p. of_int (sign p) * prod (\<lambda>i. A$i$p i) (UNIV :: 'n set))"
@@ -134,11 +134,11 @@
unfolding det_def by (simp add: sign_id)
qed
-lemma det_diagonal:
+lemma%important det_diagonal:
fixes A :: "'a::comm_ring_1^'n^'n"
assumes ld: "\<And>i j. i \<noteq> j \<Longrightarrow> A$i$j = 0"
shows "det A = prod (\<lambda>i. A$i$i) (UNIV::'n set)"
-proof -
+proof%unimportant -
let ?U = "UNIV:: 'n set"
let ?PU = "{p. p permutes ?U}"
let ?pp = "\<lambda>p. of_int (sign p) * prod (\<lambda>i. A$i$p i) (UNIV :: 'n set)"
@@ -161,13 +161,13 @@
unfolding det_def by (simp add: sign_id)
qed
-lemma det_I [simp]: "det (mat 1 :: 'a::comm_ring_1^'n^'n) = 1"
+lemma%unimportant det_I [simp]: "det (mat 1 :: 'a::comm_ring_1^'n^'n) = 1"
by (simp add: det_diagonal mat_def)
-lemma det_0 [simp]: "det (mat 0 :: 'a::comm_ring_1^'n^'n) = 0"
+lemma%unimportant det_0 [simp]: "det (mat 0 :: 'a::comm_ring_1^'n^'n) = 0"
by (simp add: det_def prod_zero power_0_left)
-lemma det_permute_rows:
+lemma%unimportant det_permute_rows:
fixes A :: "'a::comm_ring_1^'n^'n"
assumes p: "p permutes (UNIV :: 'n::finite set)"
shows "det (\<chi> i. A$p i :: 'a^'n^'n) = of_int (sign p) * det A"
@@ -204,11 +204,11 @@
done
qed
-lemma det_permute_columns:
+lemma%important det_permute_columns:
fixes A :: "'a::comm_ring_1^'n^'n"
assumes p: "p permutes (UNIV :: 'n set)"
shows "det(\<chi> i j. A$i$ p j :: 'a^'n^'n) = of_int (sign p) * det A"
-proof -
+proof%unimportant -
let ?Ap = "\<chi> i j. A$i$ p j :: 'a^'n^'n"
let ?At = "transpose A"
have "of_int (sign p) * det A = det (transpose (\<chi> i. transpose A $ p i))"
@@ -220,7 +220,7 @@
by simp
qed
-lemma det_identical_columns:
+lemma%unimportant det_identical_columns:
fixes A :: "'a::comm_ring_1^'n^'n"
assumes jk: "j \<noteq> k"
and r: "column j A = column k A"
@@ -300,24 +300,24 @@
finally show "det A = 0" by simp
qed
-lemma det_identical_rows:
+lemma%unimportant det_identical_rows:
fixes A :: "'a::comm_ring_1^'n^'n"
assumes ij: "i \<noteq> j" and r: "row i A = row j A"
shows "det A = 0"
by (metis column_transpose det_identical_columns det_transpose ij r)
-lemma det_zero_row:
+lemma%unimportant det_zero_row:
fixes A :: "'a::{idom, ring_char_0}^'n^'n" and F :: "'b::{field}^'m^'m"
shows "row i A = 0 \<Longrightarrow> det A = 0" and "row j F = 0 \<Longrightarrow> det F = 0"
by (force simp: row_def det_def vec_eq_iff sign_nz intro!: sum.neutral)+
-lemma det_zero_column:
+lemma%unimportant det_zero_column:
fixes A :: "'a::{idom, ring_char_0}^'n^'n" and F :: "'b::{field}^'m^'m"
shows "column i A = 0 \<Longrightarrow> det A = 0" and "column j F = 0 \<Longrightarrow> det F = 0"
unfolding atomize_conj atomize_imp
by (metis det_transpose det_zero_row row_transpose)
-lemma det_row_add:
+lemma%unimportant det_row_add:
fixes a b c :: "'n::finite \<Rightarrow> _ ^ 'n"
shows "det((\<chi> i. if i = k then a i + b i else c i)::'a::comm_ring_1^'n^'n) =
det((\<chi> i. if i = k then a i else c i)::'a::comm_ring_1^'n^'n) +
@@ -358,7 +358,7 @@
by (simp add: field_simps)
qed auto
-lemma det_row_mul:
+lemma%unimportant det_row_mul:
fixes a b :: "'n::finite \<Rightarrow> _ ^ 'n"
shows "det((\<chi> i. if i = k then c *s a i else b i)::'a::comm_ring_1^'n^'n) =
c * det((\<chi> i. if i = k then a i else b i)::'a::comm_ring_1^'n^'n)"
@@ -395,7 +395,7 @@
by (simp add: field_simps)
qed auto
-lemma det_row_0:
+lemma%unimportant det_row_0:
fixes b :: "'n::finite \<Rightarrow> _ ^ 'n"
shows "det((\<chi> i. if i = k then 0 else b i)::'a::comm_ring_1^'n^'n) = 0"
using det_row_mul[of k 0 "\<lambda>i. 1" b]
@@ -403,11 +403,11 @@
apply (simp only: vector_smult_lzero)
done
-lemma det_row_operation:
+lemma%important det_row_operation:
fixes A :: "'a::{comm_ring_1}^'n^'n"
assumes ij: "i \<noteq> j"
shows "det (\<chi> k. if k = i then row i A + c *s row j A else row k A) = det A"
-proof -
+proof%unimportant -
let ?Z = "(\<chi> k. if k = i then row j A else row k A) :: 'a ^'n^'n"
have th: "row i ?Z = row j ?Z" by (vector row_def)
have th2: "((\<chi> k. if k = i then row i A else row k A) :: 'a^'n^'n) = A"
@@ -417,7 +417,7 @@
by simp
qed
-lemma det_row_span:
+lemma%unimportant det_row_span:
fixes A :: "'a::{field}^'n^'n"
assumes x: "x \<in> vec.span {row j A |j. j \<noteq> i}"
shows "det (\<chi> k. if k = i then row i A + x else row k A) = det A"
@@ -449,10 +449,10 @@
unfolding scalar_mult_eq_scaleR .
qed
-lemma matrix_id [simp]: "det (matrix id) = 1"
+lemma%unimportant matrix_id [simp]: "det (matrix id) = 1"
by (simp add: matrix_id_mat_1)
-lemma det_matrix_scaleR [simp]: "det (matrix ((( *\<^sub>R) r)) :: real^'n^'n) = r ^ CARD('n::finite)"
+lemma%important det_matrix_scaleR [simp]: "det (matrix ((( *\<^sub>R) r)) :: real^'n^'n) = r ^ CARD('n::finite)"
apply (subst det_diagonal)
apply (auto simp: matrix_def mat_def)
apply (simp add: cart_eq_inner_axis inner_axis_axis)
@@ -463,7 +463,7 @@
exact duplicates by considering the rows/columns as a set.
\<close>
-lemma det_dependent_rows:
+lemma%unimportant det_dependent_rows:
fixes A:: "'a::{field}^'n^'n"
assumes d: "vec.dependent (rows A)"
shows "det A = 0"
@@ -491,23 +491,23 @@
qed
qed
-lemma det_dependent_columns:
+lemma%unimportant det_dependent_columns:
assumes d: "vec.dependent (columns (A::real^'n^'n))"
shows "det A = 0"
by (metis d det_dependent_rows rows_transpose det_transpose)
text \<open>Multilinearity and the multiplication formula\<close>
-lemma Cart_lambda_cong: "(\<And>x. f x = g x) \<Longrightarrow> (vec_lambda f::'a^'n) = (vec_lambda g :: 'a^'n)"
+lemma%unimportant Cart_lambda_cong: "(\<And>x. f x = g x) \<Longrightarrow> (vec_lambda f::'a^'n) = (vec_lambda g :: 'a^'n)"
by auto
-lemma det_linear_row_sum:
+lemma%unimportant det_linear_row_sum:
assumes fS: "finite S"
shows "det ((\<chi> i. if i = k then sum (a i) S else c i)::'a::comm_ring_1^'n^'n) =
sum (\<lambda>j. det ((\<chi> i. if i = k then a i j else c i)::'a^'n^'n)) S"
using fS by (induct rule: finite_induct; simp add: det_row_0 det_row_add cong: if_cong)
-lemma finite_bounded_functions:
+lemma%unimportant finite_bounded_functions:
assumes fS: "finite S"
shows "finite {f. (\<forall>i \<in> {1.. (k::nat)}. f i \<in> S) \<and> (\<forall>i. i \<notin> {1 .. k} \<longrightarrow> f i = i)}"
proof (induct k)
@@ -532,14 +532,14 @@
qed
-lemma det_linear_rows_sum_lemma:
+lemma%important det_linear_rows_sum_lemma:
assumes fS: "finite S"
and fT: "finite T"
shows "det ((\<chi> i. if i \<in> T then sum (a i) S else c i):: 'a::comm_ring_1^'n^'n) =
sum (\<lambda>f. det((\<chi> i. if i \<in> T then a i (f i) else c i)::'a^'n^'n))
{f. (\<forall>i \<in> T. f i \<in> S) \<and> (\<forall>i. i \<notin> T \<longrightarrow> f i = i)}"
using fT
-proof (induct T arbitrary: a c set: finite)
+proof%unimportant (induct T arbitrary: a c set: finite)
case empty
have th0: "\<And>x y. (\<chi> i. if i \<in> {} then x i else y i) = (\<chi> i. y i)"
by vector
@@ -577,7 +577,7 @@
(auto intro!: cong[OF refl[of det]] simp: vec_eq_iff)
qed
-lemma det_linear_rows_sum:
+lemma%unimportant det_linear_rows_sum:
fixes S :: "'n::finite set"
assumes fS: "finite S"
shows "det (\<chi> i. sum (a i) S) =
@@ -589,12 +589,12 @@
show ?thesis by simp
qed
-lemma matrix_mul_sum_alt:
+lemma%unimportant matrix_mul_sum_alt:
fixes A B :: "'a::comm_ring_1^'n^'n"
shows "A ** B = (\<chi> i. sum (\<lambda>k. A$i$k *s B $ k) (UNIV :: 'n set))"
by (vector matrix_matrix_mult_def sum_component)
-lemma det_rows_mul:
+lemma%unimportant det_rows_mul:
"det((\<chi> i. c i *s a i)::'a::comm_ring_1^'n^'n) =
prod (\<lambda>i. c i) (UNIV:: 'n set) * det((\<chi> i. a i)::'a^'n^'n)"
proof (simp add: det_def sum_distrib_left cong add: prod.cong, rule sum.cong)
@@ -612,10 +612,10 @@
by (simp add: field_simps)
qed rule
-lemma det_mul:
+lemma%important det_mul:
fixes A B :: "'a::comm_ring_1^'n^'n"
shows "det (A ** B) = det A * det B"
-proof -
+proof%unimportant -
let ?U = "UNIV :: 'n set"
let ?F = "{f. (\<forall>i \<in> ?U. f i \<in> ?U) \<and> (\<forall>i. i \<notin> ?U \<longrightarrow> f i = i)}"
let ?PU = "{p. p permutes ?U}"
@@ -712,12 +712,12 @@
qed
-subsection \<open>Relation to invertibility\<close>
+subsection%important \<open>Relation to invertibility\<close>
-lemma invertible_det_nz:
+lemma%important invertible_det_nz:
fixes A::"'a::{field}^'n^'n"
shows "invertible A \<longleftrightarrow> det A \<noteq> 0"
-proof (cases "invertible A")
+proof%unimportant (cases "invertible A")
case True
then obtain B :: "'a^'n^'n" where B: "A ** B = mat 1"
unfolding invertible_right_inverse by blast
@@ -748,7 +748,7 @@
qed
-lemma det_nz_iff_inj_gen:
+lemma%unimportant det_nz_iff_inj_gen:
fixes f :: "'a::field^'n \<Rightarrow> 'a::field^'n"
assumes "Vector_Spaces.linear ( *s) ( *s) f"
shows "det (matrix f) \<noteq> 0 \<longleftrightarrow> inj f"
@@ -763,26 +763,26 @@
by (metis assms invertible_det_nz invertible_left_inverse matrix_compose_gen matrix_id_mat_1)
qed
-lemma det_nz_iff_inj:
+lemma%unimportant det_nz_iff_inj:
fixes f :: "real^'n \<Rightarrow> real^'n"
assumes "linear f"
shows "det (matrix f) \<noteq> 0 \<longleftrightarrow> inj f"
using det_nz_iff_inj_gen[of f] assms
unfolding linear_matrix_vector_mul_eq .
-lemma det_eq_0_rank:
+lemma%unimportant det_eq_0_rank:
fixes A :: "real^'n^'n"
shows "det A = 0 \<longleftrightarrow> rank A < CARD('n)"
using invertible_det_nz [of A]
by (auto simp: matrix_left_invertible_injective invertible_left_inverse less_rank_noninjective)
-subsubsection\<open>Invertibility of matrices and corresponding linear functions\<close>
+subsubsection%important \<open>Invertibility of matrices and corresponding linear functions\<close>
-lemma matrix_left_invertible_gen:
+lemma%important matrix_left_invertible_gen:
fixes f :: "'a::field^'m \<Rightarrow> 'a::field^'n"
assumes "Vector_Spaces.linear ( *s) ( *s) f"
shows "((\<exists>B. B ** matrix f = mat 1) \<longleftrightarrow> (\<exists>g. Vector_Spaces.linear ( *s) ( *s) g \<and> g \<circ> f = id))"
-proof safe
+proof%unimportant safe
fix B
assume 1: "B ** matrix f = mat 1"
show "\<exists>g. Vector_Spaces.linear ( *s) ( *s) g \<and> g \<circ> f = id"
@@ -801,12 +801,12 @@
then show "\<exists>B. B ** matrix f = mat 1" ..
qed
-lemma matrix_left_invertible:
+lemma%unimportant matrix_left_invertible:
"linear f \<Longrightarrow> ((\<exists>B. B ** matrix f = mat 1) \<longleftrightarrow> (\<exists>g. linear g \<and> g \<circ> f = id))" for f::"real^'m \<Rightarrow> real^'n"
using matrix_left_invertible_gen[of f]
by (auto simp: linear_matrix_vector_mul_eq)
-lemma matrix_right_invertible_gen:
+lemma%unimportant matrix_right_invertible_gen:
fixes f :: "'a::field^'m \<Rightarrow> 'a^'n"
assumes "Vector_Spaces.linear ( *s) ( *s) f"
shows "((\<exists>B. matrix f ** B = mat 1) \<longleftrightarrow> (\<exists>g. Vector_Spaces.linear ( *s) ( *s) g \<and> f \<circ> g = id))"
@@ -829,12 +829,12 @@
then show "\<exists>B. matrix f ** B = mat 1" ..
qed
-lemma matrix_right_invertible:
+lemma%unimportant matrix_right_invertible:
"linear f \<Longrightarrow> ((\<exists>B. matrix f ** B = mat 1) \<longleftrightarrow> (\<exists>g. linear g \<and> f \<circ> g = id))" for f::"real^'m \<Rightarrow> real^'n"
using matrix_right_invertible_gen[of f]
by (auto simp: linear_matrix_vector_mul_eq)
-lemma matrix_invertible_gen:
+lemma%unimportant matrix_invertible_gen:
fixes f :: "'a::field^'m \<Rightarrow> 'a::field^'n"
assumes "Vector_Spaces.linear ( *s) ( *s) f"
shows "invertible (matrix f) \<longleftrightarrow> (\<exists>g. Vector_Spaces.linear ( *s) ( *s) g \<and> f \<circ> g = id \<and> g \<circ> f = id)"
@@ -847,13 +847,13 @@
by (metis assms invertible_def matrix_compose_gen matrix_id_mat_1)
qed
-lemma matrix_invertible:
+lemma%unimportant matrix_invertible:
"linear f \<Longrightarrow> invertible (matrix f) \<longleftrightarrow> (\<exists>g. linear g \<and> f \<circ> g = id \<and> g \<circ> f = id)"
for f::"real^'m \<Rightarrow> real^'n"
using matrix_invertible_gen[of f]
by (auto simp: linear_matrix_vector_mul_eq)
-lemma invertible_eq_bij:
+lemma%unimportant invertible_eq_bij:
fixes m :: "'a::field^'m^'n"
shows "invertible m \<longleftrightarrow> bij (( *v) m)"
using matrix_invertible_gen[OF matrix_vector_mul_linear_gen, of m, simplified matrix_of_matrix_vector_mul]
@@ -861,15 +861,15 @@
vec.linear_injective_left_inverse vec.linear_surjective_right_inverse)
-subsection \<open>Cramer's rule\<close>
+subsection%important \<open>Cramer's rule\<close>
-lemma cramer_lemma_transpose:
+lemma%important cramer_lemma_transpose:
fixes A:: "'a::{field}^'n^'n"
and x :: "'a::{field}^'n"
shows "det ((\<chi> i. if i = k then sum (\<lambda>i. x$i *s row i A) (UNIV::'n set)
else row i A)::'a::{field}^'n^'n) = x$k * det A"
(is "?lhs = ?rhs")
-proof -
+proof%unimportant -
let ?U = "UNIV :: 'n set"
let ?Uk = "?U - {k}"
have U: "?U = insert k ?Uk"
@@ -899,10 +899,10 @@
done
qed
-lemma cramer_lemma:
+lemma%important cramer_lemma:
fixes A :: "'a::{field}^'n^'n"
shows "det((\<chi> i j. if j = k then (A *v x)$i else A$i$j):: 'a::{field}^'n^'n) = x$k * det A"
-proof -
+proof%unimportant -
let ?U = "UNIV :: 'n set"
have *: "\<And>c. sum (\<lambda>i. c i *s row i (transpose A)) ?U = sum (\<lambda>i. c i *s column i A) ?U"
by (auto intro: sum.cong)
@@ -916,11 +916,11 @@
done
qed
-lemma cramer:
+lemma%important cramer:
fixes A ::"'a::{field}^'n^'n"
assumes d0: "det A \<noteq> 0"
shows "A *v x = b \<longleftrightarrow> x = (\<chi> k. det(\<chi> i j. if j=k then b$i else A$i$j) / det A)"
-proof -
+proof%unimportant -
from d0 obtain B where B: "A ** B = mat 1" "B ** A = mat 1"
unfolding invertible_det_nz[symmetric] invertible_def
by blast
@@ -941,14 +941,14 @@
by auto
qed
-subsection \<open>Orthogonality of a transformation and matrix\<close>
+subsection%important \<open>Orthogonality of a transformation and matrix\<close>
-definition "orthogonal_transformation f \<longleftrightarrow> linear f \<and> (\<forall>v w. f v \<bullet> f w = v \<bullet> w)"
+definition%important "orthogonal_transformation f \<longleftrightarrow> linear f \<and> (\<forall>v w. f v \<bullet> f w = v \<bullet> w)"
-definition "orthogonal_matrix (Q::'a::semiring_1^'n^'n) \<longleftrightarrow>
+definition%important "orthogonal_matrix (Q::'a::semiring_1^'n^'n) \<longleftrightarrow>
transpose Q ** Q = mat 1 \<and> Q ** transpose Q = mat 1"
-lemma orthogonal_transformation:
+lemma%unimportant orthogonal_transformation:
"orthogonal_transformation f \<longleftrightarrow> linear f \<and> (\<forall>v. norm (f v) = norm v)"
unfolding orthogonal_transformation_def
apply auto
@@ -957,70 +957,70 @@
apply (simp add: dot_norm linear_add[symmetric])
done
-lemma orthogonal_transformation_id [simp]: "orthogonal_transformation (\<lambda>x. x)"
+lemma%unimportant orthogonal_transformation_id [simp]: "orthogonal_transformation (\<lambda>x. x)"
by (simp add: linear_iff orthogonal_transformation_def)
-lemma orthogonal_orthogonal_transformation:
+lemma%unimportant orthogonal_orthogonal_transformation:
"orthogonal_transformation f \<Longrightarrow> orthogonal (f x) (f y) \<longleftrightarrow> orthogonal x y"
by (simp add: orthogonal_def orthogonal_transformation_def)
-lemma orthogonal_transformation_compose:
+lemma%unimportant orthogonal_transformation_compose:
"\<lbrakk>orthogonal_transformation f; orthogonal_transformation g\<rbrakk> \<Longrightarrow> orthogonal_transformation(f \<circ> g)"
by (auto simp: orthogonal_transformation_def linear_compose)
-lemma orthogonal_transformation_neg:
+lemma%unimportant orthogonal_transformation_neg:
"orthogonal_transformation(\<lambda>x. -(f x)) \<longleftrightarrow> orthogonal_transformation f"
by (auto simp: orthogonal_transformation_def dest: linear_compose_neg)
-lemma orthogonal_transformation_scaleR: "orthogonal_transformation f \<Longrightarrow> f (c *\<^sub>R v) = c *\<^sub>R f v"
+lemma%unimportant orthogonal_transformation_scaleR: "orthogonal_transformation f \<Longrightarrow> f (c *\<^sub>R v) = c *\<^sub>R f v"
by (simp add: linear_iff orthogonal_transformation_def)
-lemma orthogonal_transformation_linear:
+lemma%unimportant orthogonal_transformation_linear:
"orthogonal_transformation f \<Longrightarrow> linear f"
by (simp add: orthogonal_transformation_def)
-lemma orthogonal_transformation_inj:
+lemma%unimportant orthogonal_transformation_inj:
"orthogonal_transformation f \<Longrightarrow> inj f"
unfolding orthogonal_transformation_def inj_on_def
by (metis vector_eq)
-lemma orthogonal_transformation_surj:
+lemma%unimportant orthogonal_transformation_surj:
"orthogonal_transformation f \<Longrightarrow> surj f"
for f :: "'a::euclidean_space \<Rightarrow> 'a::euclidean_space"
by (simp add: linear_injective_imp_surjective orthogonal_transformation_inj orthogonal_transformation_linear)
-lemma orthogonal_transformation_bij:
+lemma%unimportant orthogonal_transformation_bij:
"orthogonal_transformation f \<Longrightarrow> bij f"
for f :: "'a::euclidean_space \<Rightarrow> 'a::euclidean_space"
by (simp add: bij_def orthogonal_transformation_inj orthogonal_transformation_surj)
-lemma orthogonal_transformation_inv:
+lemma%unimportant orthogonal_transformation_inv:
"orthogonal_transformation f \<Longrightarrow> orthogonal_transformation (inv f)"
for f :: "'a::euclidean_space \<Rightarrow> 'a::euclidean_space"
by (metis (no_types, hide_lams) bijection.inv_right bijection_def inj_linear_imp_inv_linear orthogonal_transformation orthogonal_transformation_bij orthogonal_transformation_inj)
-lemma orthogonal_transformation_norm:
+lemma%unimportant orthogonal_transformation_norm:
"orthogonal_transformation f \<Longrightarrow> norm (f x) = norm x"
by (metis orthogonal_transformation)
-lemma orthogonal_matrix: "orthogonal_matrix (Q:: real ^'n^'n) \<longleftrightarrow> transpose Q ** Q = mat 1"
+lemma%unimportant orthogonal_matrix: "orthogonal_matrix (Q:: real ^'n^'n) \<longleftrightarrow> transpose Q ** Q = mat 1"
by (metis matrix_left_right_inverse orthogonal_matrix_def)
-lemma orthogonal_matrix_id: "orthogonal_matrix (mat 1 :: _^'n^'n)"
+lemma%unimportant orthogonal_matrix_id: "orthogonal_matrix (mat 1 :: _^'n^'n)"
by (simp add: orthogonal_matrix_def)
-lemma orthogonal_matrix_mul:
+lemma%unimportant orthogonal_matrix_mul:
fixes A :: "real ^'n^'n"
assumes "orthogonal_matrix A" "orthogonal_matrix B"
shows "orthogonal_matrix(A ** B)"
using assms
by (simp add: orthogonal_matrix matrix_transpose_mul matrix_left_right_inverse matrix_mul_assoc)
-lemma orthogonal_transformation_matrix:
+lemma%important orthogonal_transformation_matrix:
fixes f:: "real^'n \<Rightarrow> real^'n"
shows "orthogonal_transformation f \<longleftrightarrow> linear f \<and> orthogonal_matrix(matrix f)"
(is "?lhs \<longleftrightarrow> ?rhs")
-proof -
+proof%unimportant -
let ?mf = "matrix f"
let ?ot = "orthogonal_transformation f"
let ?U = "UNIV :: 'n set"
@@ -1063,11 +1063,11 @@
by (auto simp: linear_def scalar_mult_eq_scaleR)
qed
-lemma det_orthogonal_matrix:
+lemma%important det_orthogonal_matrix:
fixes Q:: "'a::linordered_idom^'n^'n"
assumes oQ: "orthogonal_matrix Q"
shows "det Q = 1 \<or> det Q = - 1"
-proof -
+proof%unimportant -
have "Q ** transpose Q = mat 1"
by (metis oQ orthogonal_matrix_def)
then have "det (Q ** transpose Q) = det (mat 1:: 'a^'n^'n)"
@@ -1078,20 +1078,20 @@
by (simp add: square_eq_1_iff)
qed
-lemma orthogonal_transformation_det [simp]:
+lemma%important orthogonal_transformation_det [simp]:
fixes f :: "real^'n \<Rightarrow> real^'n"
shows "orthogonal_transformation f \<Longrightarrow> \<bar>det (matrix f)\<bar> = 1"
- using det_orthogonal_matrix orthogonal_transformation_matrix by fastforce
+ using%unimportant det_orthogonal_matrix orthogonal_transformation_matrix by fastforce
-subsection \<open>Linearity of scaling, and hence isometry, that preserves origin\<close>
+subsection%important \<open>Linearity of scaling, and hence isometry, that preserves origin\<close>
-lemma scaling_linear:
+lemma%important scaling_linear:
fixes f :: "'a::real_inner \<Rightarrow> 'a::real_inner"
assumes f0: "f 0 = 0"
and fd: "\<forall>x y. dist (f x) (f y) = c * dist x y"
shows "linear f"
-proof -
+proof%unimportant -
{
fix v w
have "norm (f x) = c * norm x" for x
@@ -1105,13 +1105,13 @@
by (simp add: inner_add field_simps)
qed
-lemma isometry_linear:
+lemma%unimportant isometry_linear:
"f (0::'a::real_inner) = (0::'a) \<Longrightarrow> \<forall>x y. dist(f x) (f y) = dist x y \<Longrightarrow> linear f"
by (rule scaling_linear[where c=1]) simp_all
text \<open>Hence another formulation of orthogonal transformation\<close>
-lemma orthogonal_transformation_isometry:
+lemma%important orthogonal_transformation_isometry:
"orthogonal_transformation f \<longleftrightarrow> f(0::'a::real_inner) = (0::'a) \<and> (\<forall>x y. dist(f x) (f y) = dist x y)"
unfolding orthogonal_transformation
apply (auto simp: linear_0 isometry_linear)
@@ -1119,7 +1119,7 @@
by (metis dist_0_norm)
-lemma image_orthogonal_transformation_ball:
+lemma%unimportant image_orthogonal_transformation_ball:
fixes f :: "'a::euclidean_space \<Rightarrow> 'a"
assumes "orthogonal_transformation f"
shows "f ` ball x r = ball (f x) r"
@@ -1135,7 +1135,7 @@
by (auto simp: orthogonal_transformation_isometry)
qed
-lemma image_orthogonal_transformation_cball:
+lemma%unimportant image_orthogonal_transformation_cball:
fixes f :: "'a::euclidean_space \<Rightarrow> 'a"
assumes "orthogonal_transformation f"
shows "f ` cball x r = cball (f x) r"
@@ -1151,31 +1151,31 @@
by (auto simp: orthogonal_transformation_isometry)
qed
-subsection\<open> We can find an orthogonal matrix taking any unit vector to any other\<close>
+subsection%important \<open> We can find an orthogonal matrix taking any unit vector to any other\<close>
-lemma orthogonal_matrix_transpose [simp]:
+lemma%unimportant orthogonal_matrix_transpose [simp]:
"orthogonal_matrix(transpose A) \<longleftrightarrow> orthogonal_matrix A"
by (auto simp: orthogonal_matrix_def)
-lemma orthogonal_matrix_orthonormal_columns:
+lemma%unimportant orthogonal_matrix_orthonormal_columns:
fixes A :: "real^'n^'n"
shows "orthogonal_matrix A \<longleftrightarrow>
(\<forall>i. norm(column i A) = 1) \<and>
(\<forall>i j. i \<noteq> j \<longrightarrow> orthogonal (column i A) (column j A))"
by (auto simp: orthogonal_matrix matrix_mult_transpose_dot_column vec_eq_iff mat_def norm_eq_1 orthogonal_def)
-lemma orthogonal_matrix_orthonormal_rows:
+lemma%unimportant orthogonal_matrix_orthonormal_rows:
fixes A :: "real^'n^'n"
shows "orthogonal_matrix A \<longleftrightarrow>
(\<forall>i. norm(row i A) = 1) \<and>
(\<forall>i j. i \<noteq> j \<longrightarrow> orthogonal (row i A) (row j A))"
using orthogonal_matrix_orthonormal_columns [of "transpose A"] by simp
-lemma orthogonal_matrix_exists_basis:
+lemma%important orthogonal_matrix_exists_basis:
fixes a :: "real^'n"
assumes "norm a = 1"
obtains A where "orthogonal_matrix A" "A *v (axis k 1) = a"
-proof -
+proof%unimportant -
obtain S where "a \<in> S" "pairwise orthogonal S" and noS: "\<And>x. x \<in> S \<Longrightarrow> norm x = 1"
and "independent S" "card S = CARD('n)" "span S = UNIV"
using vector_in_orthonormal_basis assms by force
@@ -1198,11 +1198,11 @@
qed
qed
-lemma orthogonal_transformation_exists_1:
+lemma%unimportant orthogonal_transformation_exists_1:
fixes a b :: "real^'n"
assumes "norm a = 1" "norm b = 1"
obtains f where "orthogonal_transformation f" "f a = b"
-proof -
+proof%unimportant -
obtain k::'n where True
by simp
obtain A B where AB: "orthogonal_matrix A" "orthogonal_matrix B" and eq: "A *v (axis k 1) = a" "B *v (axis k 1) = b"
@@ -1220,11 +1220,11 @@
qed
qed
-lemma orthogonal_transformation_exists:
+lemma%important orthogonal_transformation_exists:
fixes a b :: "real^'n"
assumes "norm a = norm b"
obtains f where "orthogonal_transformation f" "f a = b"
-proof (cases "a = 0 \<or> b = 0")
+proof%unimportant (cases "a = 0 \<or> b = 0")
case True
with assms show ?thesis
using that by force
@@ -1246,14 +1246,14 @@
qed
-subsection \<open>Can extend an isometry from unit sphere\<close>
+subsection%important \<open>Can extend an isometry from unit sphere\<close>
-lemma isometry_sphere_extend:
+lemma%important isometry_sphere_extend:
fixes f:: "'a::real_inner \<Rightarrow> 'a"
assumes f1: "\<And>x. norm x = 1 \<Longrightarrow> norm (f x) = 1"
and fd1: "\<And>x y. \<lbrakk>norm x = 1; norm y = 1\<rbrakk> \<Longrightarrow> dist (f x) (f y) = dist x y"
shows "\<exists>g. orthogonal_transformation g \<and> (\<forall>x. norm x = 1 \<longrightarrow> g x = f x)"
-proof -
+proof%unimportant -
{
fix x y x' y' u v u' v' :: "'a"
assume H: "x = norm x *\<^sub>R u" "y = norm y *\<^sub>R v"
@@ -1286,31 +1286,31 @@
by (rule exI[where x= ?g]) (metis thfg thd)
qed
-subsection \<open>Rotation, reflection, rotoinversion\<close>
+subsection%important \<open>Rotation, reflection, rotoinversion\<close>
-definition "rotation_matrix Q \<longleftrightarrow> orthogonal_matrix Q \<and> det Q = 1"
-definition "rotoinversion_matrix Q \<longleftrightarrow> orthogonal_matrix Q \<and> det Q = - 1"
+definition%important "rotation_matrix Q \<longleftrightarrow> orthogonal_matrix Q \<and> det Q = 1"
+definition%important "rotoinversion_matrix Q \<longleftrightarrow> orthogonal_matrix Q \<and> det Q = - 1"
-lemma orthogonal_rotation_or_rotoinversion:
+lemma%important orthogonal_rotation_or_rotoinversion:
fixes Q :: "'a::linordered_idom^'n^'n"
shows " orthogonal_matrix Q \<longleftrightarrow> rotation_matrix Q \<or> rotoinversion_matrix Q"
- by (metis rotoinversion_matrix_def rotation_matrix_def det_orthogonal_matrix)
+ by%unimportant (metis rotoinversion_matrix_def rotation_matrix_def det_orthogonal_matrix)
text \<open>Explicit formulas for low dimensions\<close>
-lemma prod_neutral_const: "prod f {(1::nat)..1} = f 1"
+lemma%unimportant prod_neutral_const: "prod f {(1::nat)..1} = f 1"
by simp
-lemma prod_2: "prod f {(1::nat)..2} = f 1 * f 2"
+lemma%unimportant prod_2: "prod f {(1::nat)..2} = f 1 * f 2"
by (simp add: eval_nat_numeral atLeastAtMostSuc_conv mult.commute)
-lemma prod_3: "prod f {(1::nat)..3} = f 1 * f 2 * f 3"
+lemma%unimportant prod_3: "prod f {(1::nat)..3} = f 1 * f 2 * f 3"
by (simp add: eval_nat_numeral atLeastAtMostSuc_conv mult.commute)
-lemma det_1: "det (A::'a::comm_ring_1^1^1) = A$1$1"
+lemma%unimportant det_1: "det (A::'a::comm_ring_1^1^1) = A$1$1"
by (simp add: det_def sign_id)
-lemma det_2: "det (A::'a::comm_ring_1^2^2) = A$1$1 * A$2$2 - A$1$2 * A$2$1"
+lemma%unimportant det_2: "det (A::'a::comm_ring_1^2^2) = A$1$1 * A$2$2 - A$1$2 * A$2$1"
proof -
have f12: "finite {2::2}" "1 \<notin> {2::2}" by auto
show ?thesis
@@ -1320,7 +1320,7 @@
by (simp add: sign_swap_id sign_id swap_id_eq)
qed
-lemma det_3:
+lemma%unimportant det_3:
"det (A::'a::comm_ring_1^3^3) =
A$1$1 * A$2$2 * A$3$3 +
A$1$2 * A$2$3 * A$3$1 +
@@ -1344,7 +1344,7 @@
text\<open> Slightly stronger results giving rotation, but only in two or more dimensions\<close>
-lemma rotation_matrix_exists_basis:
+lemma%unimportant rotation_matrix_exists_basis:
fixes a :: "real^'n"
assumes 2: "2 \<le> CARD('n)" and "norm a = 1"
obtains A where "rotation_matrix A" "A *v (axis k 1) = a"
@@ -1383,7 +1383,7 @@
qed
qed
-lemma rotation_exists_1:
+lemma%unimportant rotation_exists_1:
fixes a :: "real^'n"
assumes "2 \<le> CARD('n)" "norm a = 1" "norm b = 1"
obtains f where "orthogonal_transformation f" "det(matrix f) = 1" "f a = b"
@@ -1406,7 +1406,7 @@
qed
qed
-lemma rotation_exists:
+lemma%unimportant rotation_exists:
fixes a :: "real^'n"
assumes 2: "2 \<le> CARD('n)" and eq: "norm a = norm b"
obtains f where "orthogonal_transformation f" "det(matrix f) = 1" "f a = b"
@@ -1428,7 +1428,7 @@
with f show thesis ..
qed
-lemma rotation_rightward_line:
+lemma%unimportant rotation_rightward_line:
fixes a :: "real^'n"
obtains f where "orthogonal_transformation f" "2 \<le> CARD('n) \<Longrightarrow> det(matrix f) = 1"
"f(norm a *\<^sub>R axis k 1) = a"
--- a/src/HOL/Analysis/Finite_Cartesian_Product.thy Mon Aug 27 22:58:36 2018 +0200
+++ b/src/HOL/Analysis/Finite_Cartesian_Product.thy Tue Aug 28 13:28:39 2018 +0100
@@ -2,7 +2,7 @@
Author: Amine Chaieb, University of Cambridge
*)
-section \<open>Definition of finite Cartesian product types\<close>
+section%important \<open>Definition of finite Cartesian product types\<close>
theory Finite_Cartesian_Product
imports
@@ -13,7 +13,7 @@
"HOL-Library.FuncSet"
begin
-subsection \<open>Finite Cartesian products, with indexing and lambdas\<close>
+subsection%unimportant \<open>Finite Cartesian products, with indexing and lambdas\<close>
typedef ('a, 'b) vec = "UNIV :: ('b::finite \<Rightarrow> 'a) set"
morphisms vec_nth vec_lambda ..
@@ -68,7 +68,7 @@
lemma vec_lambda_eta [simp]: "(\<chi> i. (g$i)) = g"
by (simp add: vec_eq_iff)
-subsection \<open>Cardinality of vectors\<close>
+subsection%important \<open>Cardinality of vectors\<close>
instance vec :: (finite, finite) finite
proof
@@ -104,10 +104,10 @@
by (auto elim!: countableE)
qed
-lemma infinite_UNIV_vec:
+lemma%important infinite_UNIV_vec:
assumes "infinite (UNIV :: 'a set)"
shows "infinite (UNIV :: ('a^'b) set)"
-proof (subst bij_betw_finite)
+proof%unimportant (subst bij_betw_finite)
show "bij_betw vec_nth UNIV (Pi (UNIV :: 'b set) (\<lambda>_. UNIV :: 'a set))"
by (intro bij_betwI[of _ _ _ vec_lambda]) (auto simp: vec_eq_iff)
have "infinite (PiE (UNIV :: 'b set) (\<lambda>_. UNIV :: 'a set))" (is "infinite ?A")
@@ -125,9 +125,9 @@
finally show "infinite (Pi (UNIV :: 'b set) (\<lambda>_. UNIV :: 'a set))" .
qed
-lemma CARD_vec [simp]:
+lemma%important CARD_vec [simp]:
"CARD('a^'b) = CARD('a) ^ CARD('b)"
-proof (cases "finite (UNIV :: 'a set)")
+proof%unimportant (cases "finite (UNIV :: 'a set)")
case True
show ?thesis
proof (subst bij_betw_same_card)
@@ -143,7 +143,7 @@
qed
qed (simp_all add: infinite_UNIV_vec)
-lemma countable_vector:
+lemma%unimportant countable_vector:
fixes B:: "'n::finite \<Rightarrow> 'a set"
assumes "\<And>i. countable (B i)"
shows "countable {V. \<forall>i::'n::finite. V $ i \<in> B i}"
@@ -165,7 +165,7 @@
by (simp add: image_comp o_def vec_nth_inverse)
qed
-subsection \<open>Group operations and class instances\<close>
+subsection%unimportant \<open>Group operations and class instances\<close>
instantiation vec :: (zero, finite) zero
begin
@@ -230,7 +230,7 @@
by standard (simp_all add: vec_eq_iff)
-subsection\<open>Basic componentwise operations on vectors\<close>
+subsection%unimportant\<open>Basic componentwise operations on vectors\<close>
instantiation vec :: (times, finite) times
begin
@@ -304,33 +304,33 @@
"scalar_product v w = (\<Sum> i \<in> UNIV. v $ i * w $ i)"
-subsection \<open>Real vector space\<close>
+subsection%important \<open>Real vector space\<close>
-instantiation vec :: (real_vector, finite) real_vector
+instantiation%unimportant vec :: (real_vector, finite) real_vector
begin
-definition "scaleR \<equiv> (\<lambda> r x. (\<chi> i. scaleR r (x$i)))"
+definition%important "scaleR \<equiv> (\<lambda> r x. (\<chi> i. scaleR r (x$i)))"
-lemma vector_scaleR_component [simp]: "(scaleR r x)$i = scaleR r (x$i)"
+lemma%important vector_scaleR_component [simp]: "(scaleR r x)$i = scaleR r (x$i)"
unfolding scaleR_vec_def by simp
-instance
+instance%unimportant
by standard (simp_all add: vec_eq_iff scaleR_left_distrib scaleR_right_distrib)
end
-subsection \<open>Topological space\<close>
+subsection%important \<open>Topological space\<close>
-instantiation vec :: (topological_space, finite) topological_space
+instantiation%unimportant vec :: (topological_space, finite) topological_space
begin
-definition [code del]:
+definition%important [code del]:
"open (S :: ('a ^ 'b) set) \<longleftrightarrow>
(\<forall>x\<in>S. \<exists>A. (\<forall>i. open (A i) \<and> x$i \<in> A i) \<and>
(\<forall>y. (\<forall>i. y$i \<in> A i) \<longrightarrow> y \<in> S))"
-instance proof
+instance proof%unimportant
show "open (UNIV :: ('a ^ 'b) set)"
unfolding open_vec_def by auto
next
@@ -358,30 +358,30 @@
end
-lemma open_vector_box: "\<forall>i. open (S i) \<Longrightarrow> open {x. \<forall>i. x $ i \<in> S i}"
+lemma%unimportant open_vector_box: "\<forall>i. open (S i) \<Longrightarrow> open {x. \<forall>i. x $ i \<in> S i}"
unfolding open_vec_def by auto
-lemma open_vimage_vec_nth: "open S \<Longrightarrow> open ((\<lambda>x. x $ i) -` S)"
+lemma%unimportant open_vimage_vec_nth: "open S \<Longrightarrow> open ((\<lambda>x. x $ i) -` S)"
unfolding open_vec_def
apply clarify
apply (rule_tac x="\<lambda>k. if k = i then S else UNIV" in exI, simp)
done
-lemma closed_vimage_vec_nth: "closed S \<Longrightarrow> closed ((\<lambda>x. x $ i) -` S)"
+lemma%unimportant closed_vimage_vec_nth: "closed S \<Longrightarrow> closed ((\<lambda>x. x $ i) -` S)"
unfolding closed_open vimage_Compl [symmetric]
by (rule open_vimage_vec_nth)
-lemma closed_vector_box: "\<forall>i. closed (S i) \<Longrightarrow> closed {x. \<forall>i. x $ i \<in> S i}"
+lemma%unimportant closed_vector_box: "\<forall>i. closed (S i) \<Longrightarrow> closed {x. \<forall>i. x $ i \<in> S i}"
proof -
have "{x. \<forall>i. x $ i \<in> S i} = (\<Inter>i. (\<lambda>x. x $ i) -` S i)" by auto
thus "\<forall>i. closed (S i) \<Longrightarrow> closed {x. \<forall>i. x $ i \<in> S i}"
by (simp add: closed_INT closed_vimage_vec_nth)
qed
-lemma tendsto_vec_nth [tendsto_intros]:
+lemma%important tendsto_vec_nth [tendsto_intros]:
assumes "((\<lambda>x. f x) \<longlongrightarrow> a) net"
shows "((\<lambda>x. f x $ i) \<longlongrightarrow> a $ i) net"
-proof (rule topological_tendstoI)
+proof%unimportant (rule topological_tendstoI)
fix S assume "open S" "a $ i \<in> S"
then have "open ((\<lambda>y. y $ i) -` S)" "a \<in> ((\<lambda>y. y $ i) -` S)"
by (simp_all add: open_vimage_vec_nth)
@@ -391,13 +391,13 @@
by simp
qed
-lemma isCont_vec_nth [simp]: "isCont f a \<Longrightarrow> isCont (\<lambda>x. f x $ i) a"
+lemma%unimportant isCont_vec_nth [simp]: "isCont f a \<Longrightarrow> isCont (\<lambda>x. f x $ i) a"
unfolding isCont_def by (rule tendsto_vec_nth)
-lemma vec_tendstoI:
+lemma%important vec_tendstoI:
assumes "\<And>i. ((\<lambda>x. f x $ i) \<longlongrightarrow> a $ i) net"
shows "((\<lambda>x. f x) \<longlongrightarrow> a) net"
-proof (rule topological_tendstoI)
+proof%unimportant (rule topological_tendstoI)
fix S assume "open S" and "a \<in> S"
then obtain A where A: "\<And>i. open (A i)" "\<And>i. a $ i \<in> A i"
and S: "\<And>y. \<forall>i. y $ i \<in> A i \<Longrightarrow> y \<in> S"
@@ -410,13 +410,13 @@
by (rule eventually_mono, simp add: S)
qed
-lemma tendsto_vec_lambda [tendsto_intros]:
+lemma%unimportant tendsto_vec_lambda [tendsto_intros]:
assumes "\<And>i. ((\<lambda>x. f x i) \<longlongrightarrow> a i) net"
shows "((\<lambda>x. \<chi> i. f x i) \<longlongrightarrow> (\<chi> i. a i)) net"
using assms by (simp add: vec_tendstoI)
-lemma open_image_vec_nth: assumes "open S" shows "open ((\<lambda>x. x $ i) ` S)"
-proof (rule openI)
+lemma%important open_image_vec_nth: assumes "open S" shows "open ((\<lambda>x. x $ i) ` S)"
+proof%unimportant (rule openI)
fix a assume "a \<in> (\<lambda>x. x $ i) ` S"
then obtain z where "a = z $ i" and "z \<in> S" ..
then obtain A where A: "\<forall>i. open (A i) \<and> z $ i \<in> A i"
@@ -442,38 +442,38 @@
qed
-subsection \<open>Metric space\<close>
+subsection%important \<open>Metric space\<close>
(* TODO: Product of uniform spaces and compatibility with metric_spaces! *)
-instantiation vec :: (metric_space, finite) dist
+instantiation%unimportant vec :: (metric_space, finite) dist
begin
-definition
+definition%important
"dist x y = L2_set (\<lambda>i. dist (x$i) (y$i)) UNIV"
instance ..
end
-instantiation vec :: (metric_space, finite) uniformity_dist
+instantiation%unimportant vec :: (metric_space, finite) uniformity_dist
begin
-definition [code del]:
+definition%important [code del]:
"(uniformity :: (('a^'b::_) \<times> ('a^'b::_)) filter) =
(INF e:{0 <..}. principal {(x, y). dist x y < e})"
-instance
+instance%unimportant
by standard (rule uniformity_vec_def)
end
declare uniformity_Abort[where 'a="'a :: metric_space ^ 'b :: finite", code]
-instantiation vec :: (metric_space, finite) metric_space
+instantiation%unimportant vec :: (metric_space, finite) metric_space
begin
-lemma dist_vec_nth_le: "dist (x $ i) (y $ i) \<le> dist x y"
+lemma%important dist_vec_nth_le: "dist (x $ i) (y $ i) \<le> dist x y"
unfolding dist_vec_def by (rule member_le_L2_set) simp_all
-instance proof
+instance proof%unimportant
fix x y :: "'a ^ 'b"
show "dist x y = 0 \<longleftrightarrow> x = y"
unfolding dist_vec_def
@@ -532,15 +532,15 @@
end
-lemma Cauchy_vec_nth:
+lemma%important Cauchy_vec_nth:
"Cauchy (\<lambda>n. X n) \<Longrightarrow> Cauchy (\<lambda>n. X n $ i)"
unfolding Cauchy_def by (fast intro: le_less_trans [OF dist_vec_nth_le])
-lemma vec_CauchyI:
+lemma%important vec_CauchyI:
fixes X :: "nat \<Rightarrow> 'a::metric_space ^ 'n"
assumes X: "\<And>i. Cauchy (\<lambda>n. X n $ i)"
shows "Cauchy (\<lambda>n. X n)"
-proof (rule metric_CauchyI)
+proof%unimportant (rule metric_CauchyI)
fix r :: real assume "0 < r"
hence "0 < r / of_nat CARD('n)" (is "0 < ?s") by simp
define N where "N i = (LEAST N. \<forall>m\<ge>N. \<forall>n\<ge>N. dist (X m $ i) (X n $ i) < ?s)" for i
@@ -582,16 +582,16 @@
qed
-subsection \<open>Normed vector space\<close>
+subsection%important \<open>Normed vector space\<close>
-instantiation vec :: (real_normed_vector, finite) real_normed_vector
+instantiation%unimportant vec :: (real_normed_vector, finite) real_normed_vector
begin
-definition "norm x = L2_set (\<lambda>i. norm (x$i)) UNIV"
+definition%important "norm x = L2_set (\<lambda>i. norm (x$i)) UNIV"
-definition "sgn (x::'a^'b) = scaleR (inverse (norm x)) x"
+definition%important "sgn (x::'a^'b) = scaleR (inverse (norm x)) x"
-instance proof
+instance proof%unimportant
fix a :: real and x y :: "'a ^ 'b"
show "norm x = 0 \<longleftrightarrow> x = 0"
unfolding norm_vec_def
@@ -613,32 +613,32 @@
end
-lemma norm_nth_le: "norm (x $ i) \<le> norm x"
+lemma%unimportant norm_nth_le: "norm (x $ i) \<le> norm x"
unfolding norm_vec_def
by (rule member_le_L2_set) simp_all
-lemma norm_le_componentwise_cart:
+lemma%important norm_le_componentwise_cart:
fixes x :: "'a::real_normed_vector^'n"
assumes "\<And>i. norm(x$i) \<le> norm(y$i)"
shows "norm x \<le> norm y"
- unfolding norm_vec_def
- by (rule L2_set_mono) (auto simp: assms)
+ unfolding%unimportant norm_vec_def
+ by%unimportant (rule L2_set_mono) (auto simp: assms)
-lemma component_le_norm_cart: "\<bar>x$i\<bar> \<le> norm x"
+lemma%unimportant component_le_norm_cart: "\<bar>x$i\<bar> \<le> norm x"
apply (simp add: norm_vec_def)
apply (rule member_le_L2_set, simp_all)
done
-lemma norm_bound_component_le_cart: "norm x \<le> e ==> \<bar>x$i\<bar> \<le> e"
+lemma%unimportant norm_bound_component_le_cart: "norm x \<le> e ==> \<bar>x$i\<bar> \<le> e"
by (metis component_le_norm_cart order_trans)
-lemma norm_bound_component_lt_cart: "norm x < e ==> \<bar>x$i\<bar> < e"
+lemma%unimportant norm_bound_component_lt_cart: "norm x < e ==> \<bar>x$i\<bar> < e"
by (metis component_le_norm_cart le_less_trans)
-lemma norm_le_l1_cart: "norm x \<le> sum(\<lambda>i. \<bar>x$i\<bar>) UNIV"
+lemma%unimportant norm_le_l1_cart: "norm x \<le> sum(\<lambda>i. \<bar>x$i\<bar>) UNIV"
by (simp add: norm_vec_def L2_set_le_sum)
-lemma bounded_linear_vec_nth: "bounded_linear (\<lambda>x. x $ i)"
+lemma%unimportant bounded_linear_vec_nth: "bounded_linear (\<lambda>x. x $ i)"
apply standard
apply (rule vector_add_component)
apply (rule vector_scaleR_component)
@@ -648,14 +648,14 @@
instance vec :: (banach, finite) banach ..
-subsection \<open>Inner product space\<close>
+subsection%important \<open>Inner product space\<close>
-instantiation vec :: (real_inner, finite) real_inner
+instantiation%unimportant vec :: (real_inner, finite) real_inner
begin
-definition "inner x y = sum (\<lambda>i. inner (x$i) (y$i)) UNIV"
+definition%important "inner x y = sum (\<lambda>i. inner (x$i) (y$i)) UNIV"
-instance proof
+instance proof%unimportant
fix r :: real and x y z :: "'a ^ 'b"
show "inner x y = inner y x"
unfolding inner_vec_def
@@ -680,19 +680,19 @@
end
-subsection \<open>Euclidean space\<close>
+subsection%important \<open>Euclidean space\<close>
text \<open>Vectors pointing along a single axis.\<close>
-definition "axis k x = (\<chi> i. if i = k then x else 0)"
+definition%important "axis k x = (\<chi> i. if i = k then x else 0)"
-lemma axis_nth [simp]: "axis i x $ i = x"
+lemma%unimportant axis_nth [simp]: "axis i x $ i = x"
unfolding axis_def by simp
-lemma axis_eq_axis: "axis i x = axis j y \<longleftrightarrow> x = y \<and> i = j \<or> x = 0 \<and> y = 0"
+lemma%unimportant axis_eq_axis: "axis i x = axis j y \<longleftrightarrow> x = y \<and> i = j \<or> x = 0 \<and> y = 0"
unfolding axis_def vec_eq_iff by auto
-lemma inner_axis_axis:
+lemma%unimportant inner_axis_axis:
"inner (axis i x) (axis j y) = (if i = j then inner x y else 0)"
unfolding inner_vec_def
apply (cases "i = j")
@@ -702,18 +702,18 @@
apply (rule sum.neutral, simp add: axis_def)
done
-lemma inner_axis: "inner x (axis i y) = inner (x $ i) y"
+lemma%unimportant inner_axis: "inner x (axis i y) = inner (x $ i) y"
by (simp add: inner_vec_def axis_def sum.remove [where x=i])
-lemma inner_axis': "inner(axis i y) x = inner y (x $ i)"
+lemma%unimportant inner_axis': "inner(axis i y) x = inner y (x $ i)"
by (simp add: inner_axis inner_commute)
-instantiation vec :: (euclidean_space, finite) euclidean_space
+instantiation%unimportant vec :: (euclidean_space, finite) euclidean_space
begin
-definition "Basis = (\<Union>i. \<Union>u\<in>Basis. {axis i u})"
+definition%important "Basis = (\<Union>i. \<Union>u\<in>Basis. {axis i u})"
-instance proof
+instance proof%unimportant
show "(Basis :: ('a ^ 'b) set) \<noteq> {}"
unfolding Basis_vec_def by simp
next
@@ -732,8 +732,8 @@
by (simp add: inner_axis euclidean_all_zero_iff vec_eq_iff)
qed
-lemma DIM_cart [simp]: "DIM('a^'b) = CARD('b) * DIM('a)"
-proof -
+lemma%important DIM_cart [simp]: "DIM('a^'b) = CARD('b) * DIM('a)"
+proof%unimportant -
have "card (\<Union>i::'b. \<Union>u::'a\<in>Basis. {axis i u}) = (\<Sum>i::'b\<in>UNIV. card (\<Union>u::'a\<in>Basis. {axis i u}))"
by (rule card_UN_disjoint) (auto simp: axis_eq_axis)
also have "... = CARD('b) * DIM('a)"
@@ -744,30 +744,30 @@
end
-lemma norm_axis_1 [simp]: "norm (axis m (1::real)) = 1"
+lemma%unimportant norm_axis_1 [simp]: "norm (axis m (1::real)) = 1"
by (simp add: inner_axis' norm_eq_1)
-lemma sum_norm_allsubsets_bound_cart:
+lemma%important sum_norm_allsubsets_bound_cart:
fixes f:: "'a \<Rightarrow> real ^'n"
assumes fP: "finite P" and fPs: "\<And>Q. Q \<subseteq> P \<Longrightarrow> norm (sum f Q) \<le> e"
shows "sum (\<lambda>x. norm (f x)) P \<le> 2 * real CARD('n) * e"
- using sum_norm_allsubsets_bound[OF assms]
- by simp
+ using%unimportant sum_norm_allsubsets_bound[OF assms]
+ by%unimportant simp
-lemma cart_eq_inner_axis: "a $ i = inner a (axis i 1)"
+lemma%unimportant cart_eq_inner_axis: "a $ i = inner a (axis i 1)"
by (simp add: inner_axis)
-lemma axis_eq_0_iff [simp]:
+lemma%unimportant axis_eq_0_iff [simp]:
shows "axis m x = 0 \<longleftrightarrow> x = 0"
by (simp add: axis_def vec_eq_iff)
-lemma axis_in_Basis_iff [simp]: "axis i a \<in> Basis \<longleftrightarrow> a \<in> Basis"
+lemma%unimportant axis_in_Basis_iff [simp]: "axis i a \<in> Basis \<longleftrightarrow> a \<in> Basis"
by (auto simp: Basis_vec_def axis_eq_axis)
text\<open>Mapping each basis element to the corresponding finite index\<close>
definition axis_index :: "('a::comm_ring_1)^'n \<Rightarrow> 'n" where "axis_index v \<equiv> SOME i. v = axis i 1"
-lemma axis_inverse:
+lemma%unimportant axis_inverse:
fixes v :: "real^'n"
assumes "v \<in> Basis"
shows "\<exists>i. v = axis i 1"
@@ -778,20 +778,20 @@
by (force simp add: vec_eq_iff)
qed
-lemma axis_index:
+lemma%unimportant axis_index:
fixes v :: "real^'n"
assumes "v \<in> Basis"
shows "v = axis (axis_index v) 1"
by (metis (mono_tags) assms axis_inverse axis_index_def someI_ex)
-lemma axis_index_axis [simp]:
+lemma%unimportant axis_index_axis [simp]:
fixes UU :: "real^'n"
shows "(axis_index (axis u 1 :: real^'n)) = (u::'n)"
by (simp add: axis_eq_axis axis_index_def)
-subsection \<open>A naive proof procedure to lift really trivial arithmetic stuff from the basis of the vector space\<close>
+subsection%unimportant \<open>A naive proof procedure to lift really trivial arithmetic stuff from the basis of the vector space\<close>
-lemma sum_cong_aux:
+lemma%unimportant sum_cong_aux:
"(\<And>x. x \<in> A \<Longrightarrow> f x = g x) \<Longrightarrow> sum f A = sum g A"
by (auto intro: sum.cong)
@@ -823,22 +823,22 @@
end
\<close> "lift trivial vector statements to real arith statements"
-lemma vec_0[simp]: "vec 0 = 0" by vector
-lemma vec_1[simp]: "vec 1 = 1" by vector
+lemma%unimportant vec_0[simp]: "vec 0 = 0" by vector
+lemma%unimportant vec_1[simp]: "vec 1 = 1" by vector
-lemma vec_inj[simp]: "vec x = vec y \<longleftrightarrow> x = y" by vector
+lemma%unimportant vec_inj[simp]: "vec x = vec y \<longleftrightarrow> x = y" by vector
-lemma vec_in_image_vec: "vec x \<in> (vec ` S) \<longleftrightarrow> x \<in> S" by auto
+lemma%unimportant vec_in_image_vec: "vec x \<in> (vec ` S) \<longleftrightarrow> x \<in> S" by auto
-lemma vec_add: "vec(x + y) = vec x + vec y" by vector
-lemma vec_sub: "vec(x - y) = vec x - vec y" by vector
-lemma vec_cmul: "vec(c * x) = c *s vec x " by vector
-lemma vec_neg: "vec(- x) = - vec x " by vector
+lemma%unimportant vec_add: "vec(x + y) = vec x + vec y" by vector
+lemma%unimportant vec_sub: "vec(x - y) = vec x - vec y" by vector
+lemma%unimportant vec_cmul: "vec(c * x) = c *s vec x " by vector
+lemma%unimportant vec_neg: "vec(- x) = - vec x " by vector
-lemma vec_scaleR: "vec(c * x) = c *\<^sub>R vec x"
+lemma%unimportant vec_scaleR: "vec(c * x) = c *\<^sub>R vec x"
by vector
-lemma vec_sum:
+lemma%unimportant vec_sum:
assumes "finite S"
shows "vec(sum f S) = sum (vec \<circ> f) S"
using assms
@@ -852,24 +852,24 @@
text\<open>Obvious "component-pushing".\<close>
-lemma vec_component [simp]: "vec x $ i = x"
+lemma%unimportant vec_component [simp]: "vec x $ i = x"
by vector
-lemma vector_mult_component [simp]: "(x * y)$i = x$i * y$i"
+lemma%unimportant vector_mult_component [simp]: "(x * y)$i = x$i * y$i"
by vector
-lemma vector_smult_component [simp]: "(c *s y)$i = c * (y$i)"
+lemma%unimportant vector_smult_component [simp]: "(c *s y)$i = c * (y$i)"
by vector
-lemma cond_component: "(if b then x else y)$i = (if b then x$i else y$i)" by vector
+lemma%unimportant cond_component: "(if b then x else y)$i = (if b then x$i else y$i)" by vector
-lemmas vector_component =
+lemmas%unimportant vector_component =
vec_component vector_add_component vector_mult_component
vector_smult_component vector_minus_component vector_uminus_component
vector_scaleR_component cond_component
-subsection \<open>Some frequently useful arithmetic lemmas over vectors\<close>
+subsection%unimportant \<open>Some frequently useful arithmetic lemmas over vectors\<close>
instance vec :: (semigroup_mult, finite) semigroup_mult
by standard (vector mult.assoc)
@@ -1004,47 +1004,47 @@
apply (auto )
by unfold_locales (vector algebra_simps)+
-subsection \<open>Matrix operations\<close>
+subsection%important \<open>Matrix operations\<close>
text\<open>Matrix notation. NB: an MxN matrix is of type @{typ "'a^'n^'m"}, not @{typ "'a^'m^'n"}\<close>
-definition map_matrix::"('a \<Rightarrow> 'b) \<Rightarrow> (('a, 'i::finite)vec, 'j::finite) vec \<Rightarrow> (('b, 'i)vec, 'j) vec" where
+definition%important map_matrix::"('a \<Rightarrow> 'b) \<Rightarrow> (('a, 'i::finite)vec, 'j::finite) vec \<Rightarrow> (('b, 'i)vec, 'j) vec" where
"map_matrix f x = (\<chi> i j. f (x $ i $ j))"
-lemma nth_map_matrix[simp]: "map_matrix f x $ i $ j = f (x $ i $ j)"
+lemma%unimportant nth_map_matrix[simp]: "map_matrix f x $ i $ j = f (x $ i $ j)"
by (simp add: map_matrix_def)
-definition matrix_matrix_mult :: "('a::semiring_1) ^'n^'m \<Rightarrow> 'a ^'p^'n \<Rightarrow> 'a ^ 'p ^'m"
+definition%important matrix_matrix_mult :: "('a::semiring_1) ^'n^'m \<Rightarrow> 'a ^'p^'n \<Rightarrow> 'a ^ 'p ^'m"
(infixl "**" 70)
where "m ** m' == (\<chi> i j. sum (\<lambda>k. ((m$i)$k) * ((m'$k)$j)) (UNIV :: 'n set)) ::'a ^ 'p ^'m"
-definition matrix_vector_mult :: "('a::semiring_1) ^'n^'m \<Rightarrow> 'a ^'n \<Rightarrow> 'a ^ 'm"
+definition%important matrix_vector_mult :: "('a::semiring_1) ^'n^'m \<Rightarrow> 'a ^'n \<Rightarrow> 'a ^ 'm"
(infixl "*v" 70)
where "m *v x \<equiv> (\<chi> i. sum (\<lambda>j. ((m$i)$j) * (x$j)) (UNIV ::'n set)) :: 'a^'m"
-definition vector_matrix_mult :: "'a ^ 'm \<Rightarrow> ('a::semiring_1) ^'n^'m \<Rightarrow> 'a ^'n "
+definition%important vector_matrix_mult :: "'a ^ 'm \<Rightarrow> ('a::semiring_1) ^'n^'m \<Rightarrow> 'a ^'n "
(infixl "v*" 70)
where "v v* m == (\<chi> j. sum (\<lambda>i. ((m$i)$j) * (v$i)) (UNIV :: 'm set)) :: 'a^'n"
-definition "(mat::'a::zero => 'a ^'n^'n) k = (\<chi> i j. if i = j then k else 0)"
-definition transpose where
+definition%unimportant "(mat::'a::zero => 'a ^'n^'n) k = (\<chi> i j. if i = j then k else 0)"
+definition%unimportant transpose where
"(transpose::'a^'n^'m \<Rightarrow> 'a^'m^'n) A = (\<chi> i j. ((A$j)$i))"
-definition "(row::'m => 'a ^'n^'m \<Rightarrow> 'a ^'n) i A = (\<chi> j. ((A$i)$j))"
-definition "(column::'n =>'a^'n^'m =>'a^'m) j A = (\<chi> i. ((A$i)$j))"
-definition "rows(A::'a^'n^'m) = { row i A | i. i \<in> (UNIV :: 'm set)}"
-definition "columns(A::'a^'n^'m) = { column i A | i. i \<in> (UNIV :: 'n set)}"
+definition%unimportant "(row::'m => 'a ^'n^'m \<Rightarrow> 'a ^'n) i A = (\<chi> j. ((A$i)$j))"
+definition%unimportant "(column::'n =>'a^'n^'m =>'a^'m) j A = (\<chi> i. ((A$i)$j))"
+definition%unimportant "rows(A::'a^'n^'m) = { row i A | i. i \<in> (UNIV :: 'm set)}"
+definition%unimportant "columns(A::'a^'n^'m) = { column i A | i. i \<in> (UNIV :: 'n set)}"
-lemma times0_left [simp]: "(0::'a::semiring_1^'n^'m) ** (A::'a ^'p^'n) = 0"
+lemma%unimportant times0_left [simp]: "(0::'a::semiring_1^'n^'m) ** (A::'a ^'p^'n) = 0"
by (simp add: matrix_matrix_mult_def zero_vec_def)
-lemma times0_right [simp]: "(A::'a::semiring_1^'n^'m) ** (0::'a ^'p^'n) = 0"
+lemma%unimportant times0_right [simp]: "(A::'a::semiring_1^'n^'m) ** (0::'a ^'p^'n) = 0"
by (simp add: matrix_matrix_mult_def zero_vec_def)
-lemma mat_0[simp]: "mat 0 = 0" by (vector mat_def)
-lemma matrix_add_ldistrib: "(A ** (B + C)) = (A ** B) + (A ** C)"
+lemma%unimportant mat_0[simp]: "mat 0 = 0" by (vector mat_def)
+lemma%unimportant matrix_add_ldistrib: "(A ** (B + C)) = (A ** B) + (A ** C)"
by (vector matrix_matrix_mult_def sum.distrib[symmetric] field_simps)
-lemma matrix_mul_lid [simp]:
+lemma%unimportant matrix_mul_lid [simp]:
fixes A :: "'a::semiring_1 ^ 'm ^ 'n"
shows "mat 1 ** A = A"
apply (simp add: matrix_matrix_mult_def mat_def)
@@ -1053,7 +1053,7 @@
mult_1_left mult_zero_left if_True UNIV_I)
done
-lemma matrix_mul_rid [simp]:
+lemma%unimportant matrix_mul_rid [simp]:
fixes A :: "'a::semiring_1 ^ 'm ^ 'n"
shows "A ** mat 1 = A"
apply (simp add: matrix_matrix_mult_def mat_def)
@@ -1062,39 +1062,39 @@
mult_1_right mult_zero_right if_True UNIV_I cong: if_cong)
done
-lemma matrix_mul_assoc: "A ** (B ** C) = (A ** B) ** C"
+lemma%unimportant matrix_mul_assoc: "A ** (B ** C) = (A ** B) ** C"
apply (vector matrix_matrix_mult_def sum_distrib_left sum_distrib_right mult.assoc)
apply (subst sum.swap)
apply simp
done
-lemma matrix_vector_mul_assoc: "A *v (B *v x) = (A ** B) *v x"
+lemma%unimportant matrix_vector_mul_assoc: "A *v (B *v x) = (A ** B) *v x"
apply (vector matrix_matrix_mult_def matrix_vector_mult_def
sum_distrib_left sum_distrib_right mult.assoc)
apply (subst sum.swap)
apply simp
done
-lemma scalar_matrix_assoc:
+lemma%unimportant scalar_matrix_assoc:
fixes A :: "('a::real_algebra_1)^'m^'n"
shows "k *\<^sub>R (A ** B) = (k *\<^sub>R A) ** B"
by (simp add: matrix_matrix_mult_def sum_distrib_left mult_ac vec_eq_iff scaleR_sum_right)
-lemma matrix_scalar_ac:
+lemma%unimportant matrix_scalar_ac:
fixes A :: "('a::real_algebra_1)^'m^'n"
shows "A ** (k *\<^sub>R B) = k *\<^sub>R A ** B"
by (simp add: matrix_matrix_mult_def sum_distrib_left mult_ac vec_eq_iff)
-lemma matrix_vector_mul_lid [simp]: "mat 1 *v x = (x::'a::semiring_1 ^ 'n)"
+lemma%unimportant matrix_vector_mul_lid [simp]: "mat 1 *v x = (x::'a::semiring_1 ^ 'n)"
apply (vector matrix_vector_mult_def mat_def)
apply (simp add: if_distrib if_distribR sum.delta' cong del: if_weak_cong)
done
-lemma matrix_transpose_mul:
+lemma%unimportant matrix_transpose_mul:
"transpose(A ** B) = transpose B ** transpose (A::'a::comm_semiring_1^_^_)"
by (simp add: matrix_matrix_mult_def transpose_def vec_eq_iff mult.commute)
-lemma matrix_eq:
+lemma%unimportant matrix_eq:
fixes A B :: "'a::semiring_1 ^ 'n ^ 'm"
shows "A = B \<longleftrightarrow> (\<forall>x. A *v x = B *v x)" (is "?lhs \<longleftrightarrow> ?rhs")
apply auto
@@ -1107,125 +1107,125 @@
sum.delta[OF finite] cong del: if_weak_cong)
done
-lemma matrix_vector_mul_component: "(A *v x)$k = inner (A$k) x"
+lemma%unimportant matrix_vector_mul_component: "(A *v x)$k = inner (A$k) x"
by (simp add: matrix_vector_mult_def inner_vec_def)
-lemma dot_lmul_matrix: "inner ((x::real ^_) v* A) y = inner x (A *v y)"
+lemma%unimportant dot_lmul_matrix: "inner ((x::real ^_) v* A) y = inner x (A *v y)"
apply (simp add: inner_vec_def matrix_vector_mult_def vector_matrix_mult_def sum_distrib_right sum_distrib_left ac_simps)
apply (subst sum.swap)
apply simp
done
-lemma transpose_mat [simp]: "transpose (mat n) = mat n"
+lemma%unimportant transpose_mat [simp]: "transpose (mat n) = mat n"
by (vector transpose_def mat_def)
-lemma transpose_transpose [simp]: "transpose(transpose A) = A"
+lemma%unimportant transpose_transpose [simp]: "transpose(transpose A) = A"
by (vector transpose_def)
-lemma row_transpose [simp]: "row i (transpose A) = column i A"
+lemma%unimportant row_transpose [simp]: "row i (transpose A) = column i A"
by (simp add: row_def column_def transpose_def vec_eq_iff)
-lemma column_transpose [simp]: "column i (transpose A) = row i A"
+lemma%unimportant column_transpose [simp]: "column i (transpose A) = row i A"
by (simp add: row_def column_def transpose_def vec_eq_iff)
-lemma rows_transpose [simp]: "rows(transpose A) = columns A"
+lemma%unimportant rows_transpose [simp]: "rows(transpose A) = columns A"
by (auto simp add: rows_def columns_def intro: set_eqI)
-lemma columns_transpose [simp]: "columns(transpose A) = rows A"
+lemma%unimportant columns_transpose [simp]: "columns(transpose A) = rows A"
by (metis transpose_transpose rows_transpose)
-lemma transpose_scalar: "transpose (k *\<^sub>R A) = k *\<^sub>R transpose A"
+lemma%unimportant transpose_scalar: "transpose (k *\<^sub>R A) = k *\<^sub>R transpose A"
unfolding transpose_def
by (simp add: vec_eq_iff)
-lemma transpose_iff [iff]: "transpose A = transpose B \<longleftrightarrow> A = B"
+lemma%unimportant transpose_iff [iff]: "transpose A = transpose B \<longleftrightarrow> A = B"
by (metis transpose_transpose)
-lemma matrix_mult_sum:
+lemma%unimportant matrix_mult_sum:
"(A::'a::comm_semiring_1^'n^'m) *v x = sum (\<lambda>i. (x$i) *s column i A) (UNIV:: 'n set)"
by (simp add: matrix_vector_mult_def vec_eq_iff column_def mult.commute)
-lemma vector_componentwise:
+lemma%unimportant vector_componentwise:
"(x::'a::ring_1^'n) = (\<chi> j. \<Sum>i\<in>UNIV. (x$i) * (axis i 1 :: 'a^'n) $ j)"
by (simp add: axis_def if_distrib sum.If_cases vec_eq_iff)
-lemma basis_expansion: "sum (\<lambda>i. (x$i) *s axis i 1) UNIV = (x::('a::ring_1) ^'n)"
+lemma%unimportant basis_expansion: "sum (\<lambda>i. (x$i) *s axis i 1) UNIV = (x::('a::ring_1) ^'n)"
by (auto simp add: axis_def vec_eq_iff if_distrib sum.If_cases cong del: if_weak_cong)
text\<open>Correspondence between matrices and linear operators.\<close>
-definition matrix :: "('a::{plus,times, one, zero}^'m \<Rightarrow> 'a ^ 'n) \<Rightarrow> 'a^'m^'n"
+definition%important matrix :: "('a::{plus,times, one, zero}^'m \<Rightarrow> 'a ^ 'n) \<Rightarrow> 'a^'m^'n"
where "matrix f = (\<chi> i j. (f(axis j 1))$i)"
-lemma matrix_id_mat_1: "matrix id = mat 1"
+lemma%unimportant matrix_id_mat_1: "matrix id = mat 1"
by (simp add: mat_def matrix_def axis_def)
-lemma matrix_scaleR: "(matrix (( *\<^sub>R) r)) = mat r"
+lemma%unimportant matrix_scaleR: "(matrix (( *\<^sub>R) r)) = mat r"
by (simp add: mat_def matrix_def axis_def if_distrib cong: if_cong)
-lemma matrix_vector_mul_linear[intro, simp]: "linear (\<lambda>x. A *v (x::'a::real_algebra_1 ^ _))"
+lemma%unimportant matrix_vector_mul_linear[intro, simp]: "linear (\<lambda>x. A *v (x::'a::real_algebra_1 ^ _))"
by (simp add: linear_iff matrix_vector_mult_def vec_eq_iff field_simps sum_distrib_left
sum.distrib scaleR_right.sum)
-lemma vector_matrix_left_distrib [algebra_simps]:
+lemma%unimportant vector_matrix_left_distrib [algebra_simps]:
shows "(x + y) v* A = x v* A + y v* A"
unfolding vector_matrix_mult_def
by (simp add: algebra_simps sum.distrib vec_eq_iff)
-lemma matrix_vector_right_distrib [algebra_simps]:
+lemma%unimportant matrix_vector_right_distrib [algebra_simps]:
"A *v (x + y) = A *v x + A *v y"
by (vector matrix_vector_mult_def sum.distrib distrib_left)
-lemma matrix_vector_mult_diff_distrib [algebra_simps]:
+lemma%unimportant matrix_vector_mult_diff_distrib [algebra_simps]:
fixes A :: "'a::ring_1^'n^'m"
shows "A *v (x - y) = A *v x - A *v y"
by (vector matrix_vector_mult_def sum_subtractf right_diff_distrib)
-lemma matrix_vector_mult_scaleR[algebra_simps]:
+lemma%unimportant matrix_vector_mult_scaleR[algebra_simps]:
fixes A :: "real^'n^'m"
shows "A *v (c *\<^sub>R x) = c *\<^sub>R (A *v x)"
using linear_iff matrix_vector_mul_linear by blast
-lemma matrix_vector_mult_0_right [simp]: "A *v 0 = 0"
+lemma%unimportant matrix_vector_mult_0_right [simp]: "A *v 0 = 0"
by (simp add: matrix_vector_mult_def vec_eq_iff)
-lemma matrix_vector_mult_0 [simp]: "0 *v w = 0"
+lemma%unimportant matrix_vector_mult_0 [simp]: "0 *v w = 0"
by (simp add: matrix_vector_mult_def vec_eq_iff)
-lemma matrix_vector_mult_add_rdistrib [algebra_simps]:
+lemma%unimportant matrix_vector_mult_add_rdistrib [algebra_simps]:
"(A + B) *v x = (A *v x) + (B *v x)"
by (vector matrix_vector_mult_def sum.distrib distrib_right)
-lemma matrix_vector_mult_diff_rdistrib [algebra_simps]:
+lemma%unimportant matrix_vector_mult_diff_rdistrib [algebra_simps]:
fixes A :: "'a :: ring_1^'n^'m"
shows "(A - B) *v x = (A *v x) - (B *v x)"
by (vector matrix_vector_mult_def sum_subtractf left_diff_distrib)
-lemma matrix_vector_column:
+lemma%unimportant matrix_vector_column:
"(A::'a::comm_semiring_1^'n^_) *v x = sum (\<lambda>i. (x$i) *s ((transpose A)$i)) (UNIV:: 'n set)"
by (simp add: matrix_vector_mult_def transpose_def vec_eq_iff mult.commute)
-subsection\<open>Inverse matrices (not necessarily square)\<close>
+subsection%important\<open>Inverse matrices (not necessarily square)\<close>
-definition
+definition%important
"invertible(A::'a::semiring_1^'n^'m) \<longleftrightarrow> (\<exists>A'::'a^'m^'n. A ** A' = mat 1 \<and> A' ** A = mat 1)"
-definition
+definition%important
"matrix_inv(A:: 'a::semiring_1^'n^'m) =
(SOME A'::'a^'m^'n. A ** A' = mat 1 \<and> A' ** A = mat 1)"
-lemma inj_matrix_vector_mult:
+lemma%unimportant inj_matrix_vector_mult:
fixes A::"'a::field^'n^'m"
assumes "invertible A"
shows "inj (( *v) A)"
by (metis assms inj_on_inverseI invertible_def matrix_vector_mul_assoc matrix_vector_mul_lid)
-lemma scalar_invertible:
+lemma%important scalar_invertible:
fixes A :: "('a::real_algebra_1)^'m^'n"
assumes "k \<noteq> 0" and "invertible A"
shows "invertible (k *\<^sub>R A)"
-proof -
+proof%unimportant -
obtain A' where "A ** A' = mat 1" and "A' ** A = mat 1"
using assms unfolding invertible_def by auto
with `k \<noteq> 0`
@@ -1235,50 +1235,50 @@
unfolding invertible_def by auto
qed
-lemma scalar_invertible_iff:
+lemma%unimportant scalar_invertible_iff:
fixes A :: "('a::real_algebra_1)^'m^'n"
assumes "k \<noteq> 0" and "invertible A"
shows "invertible (k *\<^sub>R A) \<longleftrightarrow> k \<noteq> 0 \<and> invertible A"
by (simp add: assms scalar_invertible)
-lemma vector_transpose_matrix [simp]: "x v* transpose A = A *v x"
+lemma%unimportant vector_transpose_matrix [simp]: "x v* transpose A = A *v x"
unfolding transpose_def vector_matrix_mult_def matrix_vector_mult_def
by simp
-lemma transpose_matrix_vector [simp]: "transpose A *v x = x v* A"
+lemma%unimportant transpose_matrix_vector [simp]: "transpose A *v x = x v* A"
unfolding transpose_def vector_matrix_mult_def matrix_vector_mult_def
by simp
-lemma vector_scalar_commute:
+lemma%unimportant vector_scalar_commute:
fixes A :: "'a::{field}^'m^'n"
shows "A *v (c *s x) = c *s (A *v x)"
by (simp add: vector_scalar_mult_def matrix_vector_mult_def mult_ac sum_distrib_left)
-lemma scalar_vector_matrix_assoc:
+lemma%unimportant scalar_vector_matrix_assoc:
fixes k :: "'a::{field}" and x :: "'a::{field}^'n" and A :: "'a^'m^'n"
shows "(k *s x) v* A = k *s (x v* A)"
by (metis transpose_matrix_vector vector_scalar_commute)
-lemma vector_matrix_mult_0 [simp]: "0 v* A = 0"
+lemma%unimportant vector_matrix_mult_0 [simp]: "0 v* A = 0"
unfolding vector_matrix_mult_def by (simp add: zero_vec_def)
-lemma vector_matrix_mult_0_right [simp]: "x v* 0 = 0"
+lemma%unimportant vector_matrix_mult_0_right [simp]: "x v* 0 = 0"
unfolding vector_matrix_mult_def by (simp add: zero_vec_def)
-lemma vector_matrix_mul_rid [simp]:
+lemma%unimportant vector_matrix_mul_rid [simp]:
fixes v :: "('a::semiring_1)^'n"
shows "v v* mat 1 = v"
by (metis matrix_vector_mul_lid transpose_mat vector_transpose_matrix)
-lemma scaleR_vector_matrix_assoc:
+lemma%unimportant scaleR_vector_matrix_assoc:
fixes k :: real and x :: "real^'n" and A :: "real^'m^'n"
shows "(k *\<^sub>R x) v* A = k *\<^sub>R (x v* A)"
by (metis matrix_vector_mult_scaleR transpose_matrix_vector)
-lemma vector_scaleR_matrix_ac:
+lemma%important vector_scaleR_matrix_ac:
fixes k :: real and x :: "real^'n" and A :: "real^'m^'n"
shows "x v* (k *\<^sub>R A) = k *\<^sub>R (x v* A)"
-proof -
+proof%unimportant -
have "x v* (k *\<^sub>R A) = (k *\<^sub>R x) v* A"
unfolding vector_matrix_mult_def
by (simp add: algebra_simps)
--- a/src/HOL/Analysis/Finite_Product_Measure.thy Mon Aug 27 22:58:36 2018 +0200
+++ b/src/HOL/Analysis/Finite_Product_Measure.thy Tue Aug 28 13:28:39 2018 +0100
@@ -2,20 +2,20 @@
Author: Johannes Hölzl, TU München
*)
-section \<open>Finite product measures\<close>
+section%important \<open>Finite product measures\<close>
theory Finite_Product_Measure
imports Binary_Product_Measure
begin
-lemma PiE_choice: "(\<exists>f\<in>Pi\<^sub>E I F. \<forall>i\<in>I. P i (f i)) \<longleftrightarrow> (\<forall>i\<in>I. \<exists>x\<in>F i. P i x)"
+lemma%unimportant PiE_choice: "(\<exists>f\<in>Pi\<^sub>E I F. \<forall>i\<in>I. P i (f i)) \<longleftrightarrow> (\<forall>i\<in>I. \<exists>x\<in>F i. P i x)"
by (auto simp: Bex_def PiE_iff Ball_def dest!: choice_iff'[THEN iffD1])
(force intro: exI[of _ "restrict f I" for f])
-lemma case_prod_const: "(\<lambda>(i, j). c) = (\<lambda>_. c)"
+lemma%unimportant case_prod_const: "(\<lambda>(i, j). c) = (\<lambda>_. c)"
by auto
-subsubsection \<open>More about Function restricted by @{const extensional}\<close>
+subsubsection%unimportant \<open>More about Function restricted by @{const extensional}\<close>
definition
"merge I J = (\<lambda>(x, y) i. if i \<in> I then x i else if i \<in> J then y i else undefined)"
@@ -109,59 +109,59 @@
"I \<inter> J = {} \<Longrightarrow> merge I J -` Pi\<^sub>E (I \<union> J) E = Pi I E \<times> Pi J E"
by (auto simp: restrict_Pi_cancel PiE_def)
-subsection \<open>Finite product spaces\<close>
+subsection%important \<open>Finite product spaces\<close>
-subsubsection \<open>Products\<close>
+subsubsection%important \<open>Products\<close>
-definition prod_emb where
+definition%important prod_emb where
"prod_emb I M K X = (\<lambda>x. restrict x K) -` X \<inter> (\<Pi>\<^sub>E i\<in>I. space (M i))"
-lemma prod_emb_iff:
+lemma%important prod_emb_iff:
"f \<in> prod_emb I M K X \<longleftrightarrow> f \<in> extensional I \<and> (restrict f K \<in> X) \<and> (\<forall>i\<in>I. f i \<in> space (M i))"
- unfolding prod_emb_def PiE_def by auto
+ unfolding%unimportant prod_emb_def PiE_def by auto
-lemma
+lemma%unimportant (*FIX ME needs a name *)
shows prod_emb_empty[simp]: "prod_emb M L K {} = {}"
and prod_emb_Un[simp]: "prod_emb M L K (A \<union> B) = prod_emb M L K A \<union> prod_emb M L K B"
and prod_emb_Int: "prod_emb M L K (A \<inter> B) = prod_emb M L K A \<inter> prod_emb M L K B"
and prod_emb_UN[simp]: "prod_emb M L K (\<Union>i\<in>I. F i) = (\<Union>i\<in>I. prod_emb M L K (F i))"
and prod_emb_INT[simp]: "I \<noteq> {} \<Longrightarrow> prod_emb M L K (\<Inter>i\<in>I. F i) = (\<Inter>i\<in>I. prod_emb M L K (F i))"
and prod_emb_Diff[simp]: "prod_emb M L K (A - B) = prod_emb M L K A - prod_emb M L K B"
- by (auto simp: prod_emb_def)
+ by%unimportant (auto simp: prod_emb_def)
-lemma prod_emb_PiE: "J \<subseteq> I \<Longrightarrow> (\<And>i. i \<in> J \<Longrightarrow> E i \<subseteq> space (M i)) \<Longrightarrow>
+lemma%unimportant prod_emb_PiE: "J \<subseteq> I \<Longrightarrow> (\<And>i. i \<in> J \<Longrightarrow> E i \<subseteq> space (M i)) \<Longrightarrow>
prod_emb I M J (\<Pi>\<^sub>E i\<in>J. E i) = (\<Pi>\<^sub>E i\<in>I. if i \<in> J then E i else space (M i))"
by (force simp: prod_emb_def PiE_iff if_split_mem2)
-lemma prod_emb_PiE_same_index[simp]:
+lemma%unimportant prod_emb_PiE_same_index[simp]:
"(\<And>i. i \<in> I \<Longrightarrow> E i \<subseteq> space (M i)) \<Longrightarrow> prod_emb I M I (Pi\<^sub>E I E) = Pi\<^sub>E I E"
by (auto simp: prod_emb_def PiE_iff)
-lemma prod_emb_trans[simp]:
+lemma%unimportant prod_emb_trans[simp]:
"J \<subseteq> K \<Longrightarrow> K \<subseteq> L \<Longrightarrow> prod_emb L M K (prod_emb K M J X) = prod_emb L M J X"
by (auto simp add: Int_absorb1 prod_emb_def PiE_def)
-lemma prod_emb_Pi:
+lemma%unimportant prod_emb_Pi:
assumes "X \<in> (\<Pi> j\<in>J. sets (M j))" "J \<subseteq> K"
shows "prod_emb K M J (Pi\<^sub>E J X) = (\<Pi>\<^sub>E i\<in>K. if i \<in> J then X i else space (M i))"
using assms sets.space_closed
by (auto simp: prod_emb_def PiE_iff split: if_split_asm) blast+
-lemma prod_emb_id:
+lemma%unimportant prod_emb_id:
"B \<subseteq> (\<Pi>\<^sub>E i\<in>L. space (M i)) \<Longrightarrow> prod_emb L M L B = B"
by (auto simp: prod_emb_def subset_eq extensional_restrict)
-lemma prod_emb_mono:
+lemma%unimportant prod_emb_mono:
"F \<subseteq> G \<Longrightarrow> prod_emb A M B F \<subseteq> prod_emb A M B G"
by (auto simp: prod_emb_def)
-definition PiM :: "'i set \<Rightarrow> ('i \<Rightarrow> 'a measure) \<Rightarrow> ('i \<Rightarrow> 'a) measure" where
+definition%important PiM :: "'i set \<Rightarrow> ('i \<Rightarrow> 'a measure) \<Rightarrow> ('i \<Rightarrow> 'a) measure" where
"PiM I M = extend_measure (\<Pi>\<^sub>E i\<in>I. space (M i))
{(J, X). (J \<noteq> {} \<or> I = {}) \<and> finite J \<and> J \<subseteq> I \<and> X \<in> (\<Pi> j\<in>J. sets (M j))}
(\<lambda>(J, X). prod_emb I M J (\<Pi>\<^sub>E j\<in>J. X j))
(\<lambda>(J, X). \<Prod>j\<in>J \<union> {i\<in>I. emeasure (M i) (space (M i)) \<noteq> 1}. if j \<in> J then emeasure (M j) (X j) else emeasure (M j) (space (M j)))"
-definition prod_algebra :: "'i set \<Rightarrow> ('i \<Rightarrow> 'a measure) \<Rightarrow> ('i \<Rightarrow> 'a) set set" where
+definition%important prod_algebra :: "'i set \<Rightarrow> ('i \<Rightarrow> 'a measure) \<Rightarrow> ('i \<Rightarrow> 'a) set set" where
"prod_algebra I M = (\<lambda>(J, X). prod_emb I M J (\<Pi>\<^sub>E j\<in>J. X j)) `
{(J, X). (J \<noteq> {} \<or> I = {}) \<and> finite J \<and> J \<subseteq> I \<and> X \<in> (\<Pi> j\<in>J. sets (M j))}"
@@ -173,20 +173,20 @@
translations
"\<Pi>\<^sub>M x\<in>I. M" == "CONST PiM I (%x. M)"
-lemma extend_measure_cong:
+lemma%important extend_measure_cong:
assumes "\<Omega> = \<Omega>'" "I = I'" "G = G'" "\<And>i. i \<in> I' \<Longrightarrow> \<mu> i = \<mu>' i"
shows "extend_measure \<Omega> I G \<mu> = extend_measure \<Omega>' I' G' \<mu>'"
- unfolding extend_measure_def by (auto simp add: assms)
+ unfolding%unimportant extend_measure_def by (auto simp add: assms)
-lemma Pi_cong_sets:
+lemma%unimportant Pi_cong_sets:
"\<lbrakk>I = J; \<And>x. x \<in> I \<Longrightarrow> M x = N x\<rbrakk> \<Longrightarrow> Pi I M = Pi J N"
unfolding Pi_def by auto
-lemma PiM_cong:
+lemma%important PiM_cong:
assumes "I = J" "\<And>x. x \<in> I \<Longrightarrow> M x = N x"
shows "PiM I M = PiM J N"
unfolding PiM_def
-proof (rule extend_measure_cong, goal_cases)
+proof%unimportant (rule extend_measure_cong, goal_cases)
case 1
show ?case using assms
by (subst assms(1), intro PiE_cong[of J "\<lambda>i. space (M i)" "\<lambda>i. space (N i)"]) simp_all
@@ -206,14 +206,14 @@
qed
-lemma prod_algebra_sets_into_space:
+lemma%unimportant prod_algebra_sets_into_space:
"prod_algebra I M \<subseteq> Pow (\<Pi>\<^sub>E i\<in>I. space (M i))"
by (auto simp: prod_emb_def prod_algebra_def)
-lemma prod_algebra_eq_finite:
+lemma%important prod_algebra_eq_finite:
assumes I: "finite I"
shows "prod_algebra I M = {(\<Pi>\<^sub>E i\<in>I. X i) |X. X \<in> (\<Pi> j\<in>I. sets (M j))}" (is "?L = ?R")
-proof (intro iffI set_eqI)
+proof%unimportant (intro iffI set_eqI)
fix A assume "A \<in> ?L"
then obtain J E where J: "J \<noteq> {} \<or> I = {}" "finite J" "J \<subseteq> I" "\<forall>i\<in>J. E i \<in> sets (M i)"
and A: "A = prod_emb I M J (\<Pi>\<^sub>E j\<in>J. E j)"
@@ -232,32 +232,32 @@
by (auto simp: prod_algebra_def)
qed
-lemma prod_algebraI:
+lemma%unimportant prod_algebraI:
"finite J \<Longrightarrow> (J \<noteq> {} \<or> I = {}) \<Longrightarrow> J \<subseteq> I \<Longrightarrow> (\<And>i. i \<in> J \<Longrightarrow> E i \<in> sets (M i))
\<Longrightarrow> prod_emb I M J (\<Pi>\<^sub>E j\<in>J. E j) \<in> prod_algebra I M"
by (auto simp: prod_algebra_def)
-lemma prod_algebraI_finite:
+lemma%unimportant prod_algebraI_finite:
"finite I \<Longrightarrow> (\<forall>i\<in>I. E i \<in> sets (M i)) \<Longrightarrow> (Pi\<^sub>E I E) \<in> prod_algebra I M"
using prod_algebraI[of I I E M] prod_emb_PiE_same_index[of I E M, OF sets.sets_into_space] by simp
-lemma Int_stable_PiE: "Int_stable {Pi\<^sub>E J E | E. \<forall>i\<in>I. E i \<in> sets (M i)}"
+lemma%unimportant Int_stable_PiE: "Int_stable {Pi\<^sub>E J E | E. \<forall>i\<in>I. E i \<in> sets (M i)}"
proof (safe intro!: Int_stableI)
fix E F assume "\<forall>i\<in>I. E i \<in> sets (M i)" "\<forall>i\<in>I. F i \<in> sets (M i)"
then show "\<exists>G. Pi\<^sub>E J E \<inter> Pi\<^sub>E J F = Pi\<^sub>E J G \<and> (\<forall>i\<in>I. G i \<in> sets (M i))"
by (auto intro!: exI[of _ "\<lambda>i. E i \<inter> F i"] simp: PiE_Int)
qed
-lemma prod_algebraE:
+lemma%unimportant prod_algebraE:
assumes A: "A \<in> prod_algebra I M"
obtains J E where "A = prod_emb I M J (\<Pi>\<^sub>E j\<in>J. E j)"
"finite J" "J \<noteq> {} \<or> I = {}" "J \<subseteq> I" "\<And>i. i \<in> J \<Longrightarrow> E i \<in> sets (M i)"
using A by (auto simp: prod_algebra_def)
-lemma prod_algebraE_all:
+lemma%important prod_algebraE_all:
assumes A: "A \<in> prod_algebra I M"
obtains E where "A = Pi\<^sub>E I E" "E \<in> (\<Pi> i\<in>I. sets (M i))"
-proof -
+proof%unimportant -
from A obtain E J where A: "A = prod_emb I M J (Pi\<^sub>E J E)"
and J: "J \<subseteq> I" and E: "E \<in> (\<Pi> i\<in>J. sets (M i))"
by (auto simp: prod_algebra_def)
@@ -270,7 +270,7 @@
ultimately show ?thesis using that by auto
qed
-lemma Int_stable_prod_algebra: "Int_stable (prod_algebra I M)"
+lemma%unimportant Int_stable_prod_algebra: "Int_stable (prod_algebra I M)"
proof (unfold Int_stable_def, safe)
fix A assume "A \<in> prod_algebra I M"
from prod_algebraE[OF this] guess J E . note A = this
@@ -291,11 +291,11 @@
finally show "A \<inter> B \<in> prod_algebra I M" .
qed
-lemma prod_algebra_mono:
+lemma%unimportant prod_algebra_mono:
assumes space: "\<And>i. i \<in> I \<Longrightarrow> space (E i) = space (F i)"
assumes sets: "\<And>i. i \<in> I \<Longrightarrow> sets (E i) \<subseteq> sets (F i)"
shows "prod_algebra I E \<subseteq> prod_algebra I F"
-proof
+proof%unimportant
fix A assume "A \<in> prod_algebra I E"
then obtain J G where J: "J \<noteq> {} \<or> I = {}" "finite J" "J \<subseteq> I"
and A: "A = prod_emb I E J (\<Pi>\<^sub>E i\<in>J. G i)"
@@ -316,17 +316,17 @@
done
qed
-lemma prod_algebra_cong:
+lemma%unimportant prod_algebra_cong:
assumes "I = J" and sets: "(\<And>i. i \<in> I \<Longrightarrow> sets (M i) = sets (N i))"
shows "prod_algebra I M = prod_algebra J N"
-proof -
+proof%unimportant -
have space: "\<And>i. i \<in> I \<Longrightarrow> space (M i) = space (N i)"
using sets_eq_imp_space_eq[OF sets] by auto
with sets show ?thesis unfolding \<open>I = J\<close>