wenzelm@41983: (* Title: HOL/Probability/Sigma_Algebra.thy hoelzl@42067: Author: Stefan Richter, Markus Wenzel, TU München hoelzl@42067: Author: Johannes Hölzl, TU München hoelzl@41981: Plus material from the Hurd/Coble measure theory development, hoelzl@41981: translated by Lawrence Paulson. paulson@33271: *) paulson@33271: hoelzl@56994: header {* Describing measurable sets *} paulson@33271: wenzelm@41413: theory Sigma_Algebra wenzelm@41413: imports hoelzl@42145: Complex_Main immler@50245: "~~/src/HOL/Library/Countable_Set" wenzelm@41413: "~~/src/HOL/Library/FuncSet" wenzelm@41413: "~~/src/HOL/Library/Indicator_Function" hoelzl@47694: "~~/src/HOL/Library/Extended_Real" wenzelm@41413: begin paulson@33271: paulson@33271: text {* Sigma algebras are an elementary concept in measure paulson@33271: theory. To measure --- that is to integrate --- functions, we first have paulson@33271: to measure sets. Unfortunately, when dealing with a large universe, paulson@33271: it is often not possible to consistently assign a measure to every paulson@33271: subset. Therefore it is necessary to define the set of measurable paulson@33271: subsets of the universe. A sigma algebra is such a set that has paulson@33271: three very natural and desirable properties. *} paulson@33271: hoelzl@47762: subsection {* Families of sets *} paulson@33271: hoelzl@47694: locale subset_class = hoelzl@47694: fixes \ :: "'a set" and M :: "'a set set" hoelzl@47694: assumes space_closed: "M \ Pow \" paulson@33271: hoelzl@47694: lemma (in subset_class) sets_into_space: "x \ M \ x \ \" paulson@33271: by (metis PowD contra_subsetD space_closed) paulson@33271: hoelzl@56994: subsubsection {* Semiring of sets *} hoelzl@47762: hoelzl@47762: definition "disjoint A \ (\a\A. \b\A. a \ b \ a \ b = {})" hoelzl@47762: hoelzl@47762: lemma disjointI: hoelzl@47762: "(\a b. a \ A \ b \ A \ a \ b \ a \ b = {}) \ disjoint A" hoelzl@47762: unfolding disjoint_def by auto hoelzl@47762: hoelzl@47762: lemma disjointD: hoelzl@47762: "disjoint A \ a \ A \ b \ A \ a \ b \ a \ b = {}" hoelzl@47762: unfolding disjoint_def by auto hoelzl@47762: hoelzl@47762: lemma disjoint_empty[iff]: "disjoint {}" hoelzl@47762: by (auto simp: disjoint_def) hoelzl@42065: hoelzl@47762: lemma disjoint_union: hoelzl@47762: assumes C: "disjoint C" and B: "disjoint B" and disj: "\C \ \B = {}" hoelzl@47762: shows "disjoint (C \ B)" hoelzl@47762: proof (rule disjointI) hoelzl@47762: fix c d assume sets: "c \ C \ B" "d \ C \ B" and "c \ d" hoelzl@47762: show "c \ d = {}" hoelzl@47762: proof cases hoelzl@47762: assume "(c \ C \ d \ C) \ (c \ B \ d \ B)" hoelzl@47762: then show ?thesis hoelzl@47762: proof hoelzl@47762: assume "c \ C \ d \ C" with `c \ d` C show "c \ d = {}" hoelzl@47762: by (auto simp: disjoint_def) hoelzl@47762: next hoelzl@47762: assume "c \ B \ d \ B" with `c \ d` B show "c \ d = {}" hoelzl@47762: by (auto simp: disjoint_def) hoelzl@47762: qed hoelzl@47762: next hoelzl@47762: assume "\ ((c \ C \ d \ C) \ (c \ B \ d \ B))" hoelzl@47762: with sets have "(c \ \C \ d \ \B) \ (c \ \B \ d \ \C)" hoelzl@47762: by auto hoelzl@47762: with disj show "c \ d = {}" by auto hoelzl@47762: qed hoelzl@47762: qed hoelzl@47762: Andreas@53816: lemma disjoint_singleton [simp]: "disjoint {A}" Andreas@53816: by(simp add: disjoint_def) Andreas@53816: hoelzl@47762: locale semiring_of_sets = subset_class + hoelzl@47762: assumes empty_sets[iff]: "{} \ M" hoelzl@47762: assumes Int[intro]: "\a b. a \ M \ b \ M \ a \ b \ M" hoelzl@47762: assumes Diff_cover: hoelzl@47762: "\a b. a \ M \ b \ M \ \C\M. finite C \ disjoint C \ a - b = \C" hoelzl@47762: hoelzl@47762: lemma (in semiring_of_sets) finite_INT[intro]: hoelzl@47762: assumes "finite I" "I \ {}" "\i. i \ I \ A i \ M" hoelzl@47762: shows "(\i\I. A i) \ M" hoelzl@47762: using assms by (induct rule: finite_ne_induct) auto hoelzl@47762: hoelzl@47762: lemma (in semiring_of_sets) Int_space_eq1 [simp]: "x \ M \ \ \ x = x" hoelzl@47762: by (metis Int_absorb1 sets_into_space) hoelzl@47762: hoelzl@47762: lemma (in semiring_of_sets) Int_space_eq2 [simp]: "x \ M \ x \ \ = x" hoelzl@47762: by (metis Int_absorb2 sets_into_space) hoelzl@47762: hoelzl@47762: lemma (in semiring_of_sets) sets_Collect_conj: hoelzl@47762: assumes "{x\\. P x} \ M" "{x\\. Q x} \ M" hoelzl@47762: shows "{x\\. Q x \ P x} \ M" paulson@33271: proof - hoelzl@47762: have "{x\\. Q x \ P x} = {x\\. Q x} \ {x\\. P x}" hoelzl@42065: by auto hoelzl@47762: with assms show ?thesis by auto paulson@33271: qed paulson@33271: hoelzl@50002: lemma (in semiring_of_sets) sets_Collect_finite_All': hoelzl@47762: assumes "\i. i \ S \ {x\\. P i x} \ M" "finite S" "S \ {}" hoelzl@47762: shows "{x\\. \i\S. P i x} \ M" hoelzl@47762: proof - hoelzl@47762: have "{x\\. \i\S. P i x} = (\i\S. {x\\. P i x})" hoelzl@47762: using `S \ {}` by auto hoelzl@47762: with assms show ?thesis by auto hoelzl@47762: qed hoelzl@47762: hoelzl@47762: locale ring_of_sets = semiring_of_sets + hoelzl@47762: assumes Un [intro]: "\a b. a \ M \ b \ M \ a \ b \ M" hoelzl@47762: hoelzl@42065: lemma (in ring_of_sets) finite_Union [intro]: hoelzl@47694: "finite X \ X \ M \ Union X \ M" hoelzl@38656: by (induct set: finite) (auto simp add: Un) paulson@33271: hoelzl@42065: lemma (in ring_of_sets) finite_UN[intro]: hoelzl@47694: assumes "finite I" and "\i. i \ I \ A i \ M" hoelzl@47694: shows "(\i\I. A i) \ M" hoelzl@41981: using assms by induct auto hoelzl@41981: hoelzl@47762: lemma (in ring_of_sets) Diff [intro]: hoelzl@47762: assumes "a \ M" "b \ M" shows "a - b \ M" hoelzl@47762: using Diff_cover[OF assms] by auto hoelzl@47762: hoelzl@47762: lemma ring_of_setsI: hoelzl@47762: assumes space_closed: "M \ Pow \" hoelzl@47762: assumes empty_sets[iff]: "{} \ M" hoelzl@47762: assumes Un[intro]: "\a b. a \ M \ b \ M \ a \ b \ M" hoelzl@47762: assumes Diff[intro]: "\a b. a \ M \ b \ M \ a - b \ M" hoelzl@47762: shows "ring_of_sets \ M" hoelzl@47762: proof hoelzl@47762: fix a b assume ab: "a \ M" "b \ M" hoelzl@47762: from ab show "\C\M. finite C \ disjoint C \ a - b = \C" hoelzl@47762: by (intro exI[of _ "{a - b}"]) (auto simp: disjoint_def) hoelzl@47762: have "a \ b = a - (a - b)" by auto hoelzl@47762: also have "\ \ M" using ab by auto hoelzl@47762: finally show "a \ b \ M" . hoelzl@47762: qed fact+ hoelzl@47762: hoelzl@47762: lemma ring_of_sets_iff: "ring_of_sets \ M \ M \ Pow \ \ {} \ M \ (\a\M. \b\M. a \ b \ M) \ (\a\M. \b\M. a - b \ M)" hoelzl@47762: proof hoelzl@47762: assume "ring_of_sets \ M" hoelzl@47762: then interpret ring_of_sets \ M . hoelzl@47762: show "M \ Pow \ \ {} \ M \ (\a\M. \b\M. a \ b \ M) \ (\a\M. \b\M. a - b \ M)" hoelzl@47762: using space_closed by auto hoelzl@47762: qed (auto intro!: ring_of_setsI) hoelzl@41981: hoelzl@42065: lemma (in ring_of_sets) insert_in_sets: hoelzl@47694: assumes "{x} \ M" "A \ M" shows "insert x A \ M" hoelzl@38656: proof - hoelzl@47694: have "{x} \ A \ M" using assms by (rule Un) hoelzl@38656: thus ?thesis by auto hoelzl@38656: qed hoelzl@38656: hoelzl@42867: lemma (in ring_of_sets) sets_Collect_disj: hoelzl@47694: assumes "{x\\. P x} \ M" "{x\\. Q x} \ M" hoelzl@47694: shows "{x\\. Q x \ P x} \ M" hoelzl@42867: proof - hoelzl@47694: have "{x\\. Q x \ P x} = {x\\. Q x} \ {x\\. P x}" hoelzl@42867: by auto hoelzl@42867: with assms show ?thesis by auto hoelzl@42867: qed hoelzl@42867: hoelzl@42867: lemma (in ring_of_sets) sets_Collect_finite_Ex: hoelzl@47694: assumes "\i. i \ S \ {x\\. P i x} \ M" "finite S" hoelzl@47694: shows "{x\\. \i\S. P i x} \ M" hoelzl@42867: proof - hoelzl@47694: have "{x\\. \i\S. P i x} = (\i\S. {x\\. P i x})" hoelzl@42867: by auto hoelzl@42867: with assms show ?thesis by auto hoelzl@42867: qed hoelzl@42867: hoelzl@42065: locale algebra = ring_of_sets + hoelzl@47694: assumes top [iff]: "\ \ M" hoelzl@42065: hoelzl@42065: lemma (in algebra) compl_sets [intro]: hoelzl@47694: "a \ M \ \ - a \ M" hoelzl@42065: by auto hoelzl@42065: hoelzl@42065: lemma algebra_iff_Un: hoelzl@47694: "algebra \ M \ hoelzl@47694: M \ Pow \ \ hoelzl@47694: {} \ M \ hoelzl@47694: (\a \ M. \ - a \ M) \ hoelzl@47694: (\a \ M. \ b \ M. a \ b \ M)" (is "_ \ ?Un") hoelzl@42065: proof hoelzl@47694: assume "algebra \ M" hoelzl@47694: then interpret algebra \ M . hoelzl@42065: show ?Un using sets_into_space by auto hoelzl@42065: next hoelzl@42065: assume ?Un hoelzl@47762: then have "\ \ M" by auto hoelzl@47762: interpret ring_of_sets \ M hoelzl@47762: proof (rule ring_of_setsI) hoelzl@47762: show \: "M \ Pow \" "{} \ M" hoelzl@42065: using `?Un` by auto hoelzl@47694: fix a b assume a: "a \ M" and b: "b \ M" hoelzl@47694: then show "a \ b \ M" using `?Un` by auto hoelzl@47694: have "a - b = \ - ((\ - a) \ b)" hoelzl@47694: using \ a b by auto hoelzl@47694: then show "a - b \ M" hoelzl@42065: using a b `?Un` by auto hoelzl@42065: qed hoelzl@47762: show "algebra \ M" proof qed fact hoelzl@42065: qed hoelzl@42065: hoelzl@42065: lemma algebra_iff_Int: hoelzl@47694: "algebra \ M \ hoelzl@47694: M \ Pow \ & {} \ M & hoelzl@47694: (\a \ M. \ - a \ M) & hoelzl@47694: (\a \ M. \ b \ M. a \ b \ M)" (is "_ \ ?Int") hoelzl@42065: proof hoelzl@47694: assume "algebra \ M" hoelzl@47694: then interpret algebra \ M . hoelzl@42065: show ?Int using sets_into_space by auto hoelzl@42065: next hoelzl@42065: assume ?Int hoelzl@47694: show "algebra \ M" hoelzl@42065: proof (unfold algebra_iff_Un, intro conjI ballI) hoelzl@47694: show \: "M \ Pow \" "{} \ M" hoelzl@42065: using `?Int` by auto hoelzl@47694: from `?Int` show "\a. a \ M \ \ - a \ M" by auto hoelzl@47694: fix a b assume M: "a \ M" "b \ M" hoelzl@47694: hence "a \ b = \ - ((\ - a) \ (\ - b))" hoelzl@47694: using \ by blast hoelzl@47694: also have "... \ M" hoelzl@47694: using M `?Int` by auto hoelzl@47694: finally show "a \ b \ M" . hoelzl@42065: qed hoelzl@42065: qed hoelzl@42065: hoelzl@42867: lemma (in algebra) sets_Collect_neg: hoelzl@47694: assumes "{x\\. P x} \ M" hoelzl@47694: shows "{x\\. \ P x} \ M" hoelzl@42867: proof - hoelzl@47694: have "{x\\. \ P x} = \ - {x\\. P x}" by auto hoelzl@42867: with assms show ?thesis by auto hoelzl@42867: qed hoelzl@42867: hoelzl@42867: lemma (in algebra) sets_Collect_imp: hoelzl@47694: "{x\\. P x} \ M \ {x\\. Q x} \ M \ {x\\. Q x \ P x} \ M" hoelzl@42867: unfolding imp_conv_disj by (intro sets_Collect_disj sets_Collect_neg) hoelzl@42867: hoelzl@42867: lemma (in algebra) sets_Collect_const: hoelzl@47694: "{x\\. P} \ M" hoelzl@42867: by (cases P) auto hoelzl@42867: hoelzl@42984: lemma algebra_single_set: hoelzl@47762: "X \ S \ algebra S { {}, X, S - X, S }" hoelzl@47762: by (auto simp: algebra_iff_Int) hoelzl@42984: hoelzl@56994: subsubsection {* Restricted algebras *} hoelzl@39092: hoelzl@39092: abbreviation (in algebra) hoelzl@47694: "restricted_space A \ (op \ A) ` M" hoelzl@39092: hoelzl@38656: lemma (in algebra) restricted_algebra: hoelzl@47694: assumes "A \ M" shows "algebra A (restricted_space A)" hoelzl@47762: using assms by (auto simp: algebra_iff_Int) paulson@33271: hoelzl@56994: subsubsection {* Sigma Algebras *} paulson@33271: paulson@33271: locale sigma_algebra = algebra + hoelzl@47694: assumes countable_nat_UN [intro]: "\A. range A \ M \ (\i::nat. A i) \ M" paulson@33271: hoelzl@42984: lemma (in algebra) is_sigma_algebra: hoelzl@47694: assumes "finite M" hoelzl@47694: shows "sigma_algebra \ M" hoelzl@42984: proof hoelzl@47694: fix A :: "nat \ 'a set" assume "range A \ M" hoelzl@47694: then have "(\i. A i) = (\s\M \ range A. s)" hoelzl@42984: by auto hoelzl@47694: also have "(\s\M \ range A. s) \ M" hoelzl@47694: using `finite M` by auto hoelzl@47694: finally show "(\i. A i) \ M" . hoelzl@42984: qed hoelzl@42984: hoelzl@38656: lemma countable_UN_eq: hoelzl@38656: fixes A :: "'i::countable \ 'a set" hoelzl@47694: shows "(range A \ M \ (\i. A i) \ M) \ hoelzl@47694: (range (A \ from_nat) \ M \ (\i. (A \ from_nat) i) \ M)" hoelzl@38656: proof - hoelzl@38656: let ?A' = "A \ from_nat" hoelzl@38656: have *: "(\i. ?A' i) = (\i. A i)" (is "?l = ?r") hoelzl@38656: proof safe hoelzl@38656: fix x i assume "x \ A i" thus "x \ ?l" hoelzl@38656: by (auto intro!: exI[of _ "to_nat i"]) hoelzl@38656: next hoelzl@38656: fix x i assume "x \ ?A' i" thus "x \ ?r" hoelzl@38656: by (auto intro!: exI[of _ "from_nat i"]) hoelzl@38656: qed hoelzl@38656: have **: "range ?A' = range A" hoelzl@40702: using surj_from_nat haftmann@56154: by (auto simp: image_comp [symmetric] intro!: imageI) hoelzl@38656: show ?thesis unfolding * ** .. hoelzl@38656: qed hoelzl@38656: immler@50245: lemma (in sigma_algebra) countable_Union [intro]: immler@50245: assumes "countable X" "X \ M" shows "Union X \ M" immler@50245: proof cases immler@50245: assume "X \ {}" immler@50245: hence "\X = (\n. from_nat_into X n)" immler@50245: using assms by (auto intro: from_nat_into) (metis from_nat_into_surj) immler@50245: also have "\ \ M" using assms immler@50245: by (auto intro!: countable_nat_UN) (metis `X \ {}` from_nat_into set_mp) immler@50245: finally show ?thesis . immler@50245: qed simp immler@50245: hoelzl@38656: lemma (in sigma_algebra) countable_UN[intro]: hoelzl@38656: fixes A :: "'i::countable \ 'a set" hoelzl@47694: assumes "A`X \ M" hoelzl@47694: shows "(\x\X. A x) \ M" hoelzl@38656: proof - wenzelm@46731: let ?A = "\i. if i \ X then A i else {}" hoelzl@47694: from assms have "range ?A \ M" by auto hoelzl@38656: with countable_nat_UN[of "?A \ from_nat"] countable_UN_eq[of ?A M] hoelzl@47694: have "(\x. ?A x) \ M" by auto hoelzl@38656: moreover have "(\x. ?A x) = (\x\X. A x)" by (auto split: split_if_asm) hoelzl@38656: ultimately show ?thesis by simp hoelzl@38656: qed hoelzl@38656: hoelzl@50526: lemma (in sigma_algebra) countable_UN': hoelzl@50526: fixes A :: "'i \ 'a set" hoelzl@50526: assumes X: "countable X" hoelzl@50526: assumes A: "A`X \ M" hoelzl@50526: shows "(\x\X. A x) \ M" hoelzl@50526: proof - hoelzl@50526: have "(\x\X. A x) = (\i\to_nat_on X ` X. A (from_nat_into X i))" hoelzl@50526: using X by auto hoelzl@50526: also have "\ \ M" hoelzl@50526: using A X hoelzl@50526: by (intro countable_UN) auto hoelzl@50526: finally show ?thesis . hoelzl@50526: qed hoelzl@50526: paulson@33533: lemma (in sigma_algebra) countable_INT [intro]: hoelzl@38656: fixes A :: "'i::countable \ 'a set" hoelzl@47694: assumes A: "A`X \ M" "X \ {}" hoelzl@47694: shows "(\i\X. A i) \ M" paulson@33271: proof - hoelzl@47694: from A have "\i\X. A i \ M" by fast hoelzl@47694: hence "\ - (\i\X. \ - A i) \ M" by blast paulson@33271: moreover hoelzl@47694: have "(\i\X. A i) = \ - (\i\X. \ - A i)" using space_closed A paulson@33271: by blast paulson@33271: ultimately show ?thesis by metis paulson@33271: qed paulson@33271: hoelzl@50526: lemma (in sigma_algebra) countable_INT': hoelzl@50526: fixes A :: "'i \ 'a set" hoelzl@50526: assumes X: "countable X" "X \ {}" hoelzl@50526: assumes A: "A`X \ M" hoelzl@50526: shows "(\x\X. A x) \ M" hoelzl@50526: proof - hoelzl@50526: have "(\x\X. A x) = (\i\to_nat_on X ` X. A (from_nat_into X i))" hoelzl@50526: using X by auto hoelzl@50526: also have "\ \ M" hoelzl@50526: using A X hoelzl@50526: by (intro countable_INT) auto hoelzl@50526: finally show ?thesis . hoelzl@50526: qed hoelzl@50526: hoelzl@57275: hoelzl@57275: lemma (in sigma_algebra) countable: hoelzl@57275: assumes "\a. a \ A \ {a} \ M" "countable A" hoelzl@57275: shows "A \ M" hoelzl@57275: proof - hoelzl@57275: have "(\a\A. {a}) \ M" hoelzl@57275: using assms by (intro countable_UN') auto hoelzl@57275: also have "(\a\A. {a}) = A" by auto hoelzl@57275: finally show ?thesis by auto hoelzl@57275: qed hoelzl@57275: hoelzl@47694: lemma ring_of_sets_Pow: "ring_of_sets sp (Pow sp)" hoelzl@47762: by (auto simp: ring_of_sets_iff) hoelzl@42145: hoelzl@47694: lemma algebra_Pow: "algebra sp (Pow sp)" hoelzl@47762: by (auto simp: algebra_iff_Un) hoelzl@38656: hoelzl@38656: lemma sigma_algebra_iff: hoelzl@47694: "sigma_algebra \ M \ hoelzl@47694: algebra \ M \ (\A. range A \ M \ (\i::nat. A i) \ M)" hoelzl@38656: by (simp add: sigma_algebra_def sigma_algebra_axioms_def) paulson@33271: hoelzl@47762: lemma sigma_algebra_Pow: "sigma_algebra sp (Pow sp)" hoelzl@47762: by (auto simp: sigma_algebra_iff algebra_iff_Int) hoelzl@47762: hoelzl@42867: lemma (in sigma_algebra) sets_Collect_countable_All: hoelzl@47694: assumes "\i. {x\\. P i x} \ M" hoelzl@47694: shows "{x\\. \i::'i::countable. P i x} \ M" hoelzl@42867: proof - hoelzl@47694: have "{x\\. \i::'i::countable. P i x} = (\i. {x\\. P i x})" by auto hoelzl@42867: with assms show ?thesis by auto hoelzl@42867: qed hoelzl@42867: hoelzl@42867: lemma (in sigma_algebra) sets_Collect_countable_Ex: hoelzl@47694: assumes "\i. {x\\. P i x} \ M" hoelzl@47694: shows "{x\\. \i::'i::countable. P i x} \ M" hoelzl@42867: proof - hoelzl@47694: have "{x\\. \i::'i::countable. P i x} = (\i. {x\\. P i x})" by auto hoelzl@42867: with assms show ?thesis by auto hoelzl@42867: qed hoelzl@42867: hoelzl@50526: lemma (in sigma_algebra) sets_Collect_countable_Ex': hoelzl@54418: assumes "\i. i \ I \ {x\\. P i x} \ M" hoelzl@50526: assumes "countable I" hoelzl@50526: shows "{x\\. \i\I. P i x} \ M" hoelzl@50526: proof - hoelzl@50526: have "{x\\. \i\I. P i x} = (\i\I. {x\\. P i x})" by auto hoelzl@50526: with assms show ?thesis hoelzl@50526: by (auto intro!: countable_UN') hoelzl@50526: qed hoelzl@50526: hoelzl@54418: lemma (in sigma_algebra) sets_Collect_countable_All': hoelzl@54418: assumes "\i. i \ I \ {x\\. P i x} \ M" hoelzl@54418: assumes "countable I" hoelzl@54418: shows "{x\\. \i\I. P i x} \ M" hoelzl@54418: proof - hoelzl@54418: have "{x\\. \i\I. P i x} = (\i\I. {x\\. P i x}) \ \" by auto hoelzl@54418: with assms show ?thesis hoelzl@54418: by (cases "I = {}") (auto intro!: countable_INT') hoelzl@54418: qed hoelzl@54418: hoelzl@54418: lemma (in sigma_algebra) sets_Collect_countable_Ex1': hoelzl@54418: assumes "\i. i \ I \ {x\\. P i x} \ M" hoelzl@54418: assumes "countable I" hoelzl@54418: shows "{x\\. \!i\I. P i x} \ M" hoelzl@54418: proof - hoelzl@54418: have "{x\\. \!i\I. P i x} = {x\\. \i\I. P i x \ (\j\I. P j x \ i = j)}" hoelzl@54418: by auto hoelzl@54418: with assms show ?thesis hoelzl@54418: by (auto intro!: sets_Collect_countable_All' sets_Collect_countable_Ex' sets_Collect_conj sets_Collect_imp sets_Collect_const) hoelzl@54418: qed hoelzl@54418: hoelzl@42867: lemmas (in sigma_algebra) sets_Collect = hoelzl@42867: sets_Collect_imp sets_Collect_disj sets_Collect_conj sets_Collect_neg sets_Collect_const hoelzl@42867: sets_Collect_countable_All sets_Collect_countable_Ex sets_Collect_countable_All hoelzl@42867: hoelzl@47694: lemma (in sigma_algebra) sets_Collect_countable_Ball: hoelzl@47694: assumes "\i. {x\\. P i x} \ M" hoelzl@47694: shows "{x\\. \i::'i::countable\X. P i x} \ M" hoelzl@47694: unfolding Ball_def by (intro sets_Collect assms) hoelzl@47694: hoelzl@47694: lemma (in sigma_algebra) sets_Collect_countable_Bex: hoelzl@47694: assumes "\i. {x\\. P i x} \ M" hoelzl@47694: shows "{x\\. \i::'i::countable\X. P i x} \ M" hoelzl@47694: unfolding Bex_def by (intro sets_Collect assms) hoelzl@47694: hoelzl@42984: lemma sigma_algebra_single_set: hoelzl@42984: assumes "X \ S" hoelzl@47694: shows "sigma_algebra S { {}, X, S - X, S }" hoelzl@42984: using algebra.is_sigma_algebra[OF algebra_single_set[OF `X \ S`]] by simp hoelzl@42984: hoelzl@56994: subsubsection {* Binary Unions *} paulson@33271: paulson@33271: definition binary :: "'a \ 'a \ nat \ 'a" wenzelm@50252: where "binary a b = (\x. b)(0 := a)" paulson@33271: hoelzl@38656: lemma range_binary_eq: "range(binary a b) = {a,b}" hoelzl@38656: by (auto simp add: binary_def) paulson@33271: hoelzl@38656: lemma Un_range_binary: "a \ b = (\i::nat. binary a b i)" haftmann@44106: by (simp add: SUP_def range_binary_eq) paulson@33271: hoelzl@38656: lemma Int_range_binary: "a \ b = (\i::nat. binary a b i)" haftmann@44106: by (simp add: INF_def range_binary_eq) paulson@33271: paulson@33271: lemma sigma_algebra_iff2: hoelzl@47694: "sigma_algebra \ M \ hoelzl@47694: M \ Pow \ \ hoelzl@47694: {} \ M \ (\s \ M. \ - s \ M) \ hoelzl@47694: (\A. range A \ M \ (\i::nat. A i) \ M)" hoelzl@38656: by (auto simp add: range_binary_eq sigma_algebra_def sigma_algebra_axioms_def hoelzl@42065: algebra_iff_Un Un_range_binary) paulson@33271: hoelzl@56994: subsubsection {* Initial Sigma Algebra *} paulson@33271: paulson@33271: text {*Sigma algebras can naturally be created as the closure of any set of hoelzl@47694: M with regard to the properties just postulated. *} paulson@33271: hoelzl@51683: inductive_set sigma_sets :: "'a set \ 'a set set \ 'a set set" paulson@33271: for sp :: "'a set" and A :: "'a set set" paulson@33271: where hoelzl@47694: Basic[intro, simp]: "a \ A \ a \ sigma_sets sp A" paulson@33271: | Empty: "{} \ sigma_sets sp A" paulson@33271: | Compl: "a \ sigma_sets sp A \ sp - a \ sigma_sets sp A" paulson@33271: | Union: "(\i::nat. a i \ sigma_sets sp A) \ (\i. a i) \ sigma_sets sp A" paulson@33271: hoelzl@41543: lemma (in sigma_algebra) sigma_sets_subset: hoelzl@47694: assumes a: "a \ M" hoelzl@47694: shows "sigma_sets \ a \ M" hoelzl@41543: proof hoelzl@41543: fix x hoelzl@47694: assume "x \ sigma_sets \ a" hoelzl@47694: from this show "x \ M" hoelzl@41543: by (induct rule: sigma_sets.induct, auto) (metis a subsetD) hoelzl@41543: qed hoelzl@41543: hoelzl@41543: lemma sigma_sets_into_sp: "A \ Pow sp \ x \ sigma_sets sp A \ x \ sp" hoelzl@41543: by (erule sigma_sets.induct, auto) hoelzl@41543: hoelzl@41543: lemma sigma_algebra_sigma_sets: hoelzl@47694: "a \ Pow \ \ sigma_algebra \ (sigma_sets \ a)" hoelzl@41543: by (auto simp add: sigma_algebra_iff2 dest: sigma_sets_into_sp hoelzl@41543: intro!: sigma_sets.Union sigma_sets.Empty sigma_sets.Compl) hoelzl@41543: hoelzl@41543: lemma sigma_sets_least_sigma_algebra: hoelzl@41543: assumes "A \ Pow S" hoelzl@47694: shows "sigma_sets S A = \{B. A \ B \ sigma_algebra S B}" hoelzl@41543: proof safe hoelzl@47694: fix B X assume "A \ B" and sa: "sigma_algebra S B" hoelzl@41543: and X: "X \ sigma_sets S A" hoelzl@41543: from sigma_algebra.sigma_sets_subset[OF sa, simplified, OF `A \ B`] X hoelzl@41543: show "X \ B" by auto hoelzl@41543: next hoelzl@47694: fix X assume "X \ \{B. A \ B \ sigma_algebra S B}" hoelzl@47694: then have [intro!]: "\B. A \ B \ sigma_algebra S B \ X \ B" hoelzl@41543: by simp hoelzl@47694: have "A \ sigma_sets S A" using assms by auto hoelzl@47694: moreover have "sigma_algebra S (sigma_sets S A)" hoelzl@41543: using assms by (intro sigma_algebra_sigma_sets[of A]) auto hoelzl@41543: ultimately show "X \ sigma_sets S A" by auto hoelzl@41543: qed paulson@33271: paulson@33271: lemma sigma_sets_top: "sp \ sigma_sets sp A" paulson@33271: by (metis Diff_empty sigma_sets.Compl sigma_sets.Empty) paulson@33271: hoelzl@38656: lemma sigma_sets_Un: paulson@33271: "a \ sigma_sets sp A \ b \ sigma_sets sp A \ a \ b \ sigma_sets sp A" hoelzl@38656: apply (simp add: Un_range_binary range_binary_eq) hoelzl@40859: apply (rule Union, simp add: binary_def) paulson@33271: done paulson@33271: paulson@33271: lemma sigma_sets_Inter: paulson@33271: assumes Asb: "A \ Pow sp" paulson@33271: shows "(\i::nat. a i \ sigma_sets sp A) \ (\i. a i) \ sigma_sets sp A" paulson@33271: proof - paulson@33271: assume ai: "\i::nat. a i \ sigma_sets sp A" hoelzl@38656: hence "\i::nat. sp-(a i) \ sigma_sets sp A" paulson@33271: by (rule sigma_sets.Compl) hoelzl@38656: hence "(\i. sp-(a i)) \ sigma_sets sp A" paulson@33271: by (rule sigma_sets.Union) hoelzl@38656: hence "sp-(\i. sp-(a i)) \ sigma_sets sp A" paulson@33271: by (rule sigma_sets.Compl) hoelzl@38656: also have "sp-(\i. sp-(a i)) = sp Int (\i. a i)" paulson@33271: by auto paulson@33271: also have "... = (\i. a i)" using ai hoelzl@38656: by (blast dest: sigma_sets_into_sp [OF Asb]) hoelzl@38656: finally show ?thesis . paulson@33271: qed paulson@33271: paulson@33271: lemma sigma_sets_INTER: hoelzl@38656: assumes Asb: "A \ Pow sp" paulson@33271: and ai: "\i::nat. i \ S \ a i \ sigma_sets sp A" and non: "S \ {}" paulson@33271: shows "(\i\S. a i) \ sigma_sets sp A" paulson@33271: proof - paulson@33271: from ai have "\i. (if i\S then a i else sp) \ sigma_sets sp A" hoelzl@47694: by (simp add: sigma_sets.intros(2-) sigma_sets_top) paulson@33271: hence "(\i. (if i\S then a i else sp)) \ sigma_sets sp A" paulson@33271: by (rule sigma_sets_Inter [OF Asb]) paulson@33271: also have "(\i. (if i\S then a i else sp)) = (\i\S. a i)" paulson@33271: by auto (metis ai non sigma_sets_into_sp subset_empty subset_iff Asb)+ paulson@33271: finally show ?thesis . paulson@33271: qed paulson@33271: hoelzl@51683: lemma sigma_sets_UNION: "countable B \ (\b. b \ B \ b \ sigma_sets X A) \ (\B) \ sigma_sets X A" hoelzl@51683: using from_nat_into[of B] range_from_nat_into[of B] sigma_sets.Union[of "from_nat_into B" X A] hoelzl@51683: apply (cases "B = {}") hoelzl@51683: apply (simp add: sigma_sets.Empty) hoelzl@51683: apply (simp del: Union_image_eq add: Union_image_eq[symmetric]) hoelzl@51683: done hoelzl@51683: paulson@33271: lemma (in sigma_algebra) sigma_sets_eq: hoelzl@47694: "sigma_sets \ M = M" paulson@33271: proof hoelzl@47694: show "M \ sigma_sets \ M" huffman@37032: by (metis Set.subsetI sigma_sets.Basic) paulson@33271: next hoelzl@47694: show "sigma_sets \ M \ M" paulson@33271: by (metis sigma_sets_subset subset_refl) paulson@33271: qed paulson@33271: hoelzl@42981: lemma sigma_sets_eqI: hoelzl@42981: assumes A: "\a. a \ A \ a \ sigma_sets M B" hoelzl@42981: assumes B: "\b. b \ B \ b \ sigma_sets M A" hoelzl@42981: shows "sigma_sets M A = sigma_sets M B" hoelzl@42981: proof (intro set_eqI iffI) hoelzl@42981: fix a assume "a \ sigma_sets M A" hoelzl@42981: from this A show "a \ sigma_sets M B" hoelzl@47694: by induct (auto intro!: sigma_sets.intros(2-) del: sigma_sets.Basic) hoelzl@42981: next hoelzl@42981: fix b assume "b \ sigma_sets M B" hoelzl@42981: from this B show "b \ sigma_sets M A" hoelzl@47694: by induct (auto intro!: sigma_sets.intros(2-) del: sigma_sets.Basic) hoelzl@42981: qed hoelzl@42981: hoelzl@42984: lemma sigma_sets_subseteq: assumes "A \ B" shows "sigma_sets X A \ sigma_sets X B" hoelzl@42984: proof hoelzl@42984: fix x assume "x \ sigma_sets X A" then show "x \ sigma_sets X B" hoelzl@47694: by induct (insert `A \ B`, auto intro: sigma_sets.intros(2-)) hoelzl@42984: qed hoelzl@42984: hoelzl@47762: lemma sigma_sets_mono: assumes "A \ sigma_sets X B" shows "sigma_sets X A \ sigma_sets X B" hoelzl@47762: proof hoelzl@47762: fix x assume "x \ sigma_sets X A" then show "x \ sigma_sets X B" hoelzl@47762: by induct (insert `A \ sigma_sets X B`, auto intro: sigma_sets.intros(2-)) hoelzl@47762: qed hoelzl@47762: hoelzl@47762: lemma sigma_sets_mono': assumes "A \ B" shows "sigma_sets X A \ sigma_sets X B" hoelzl@47762: proof hoelzl@47762: fix x assume "x \ sigma_sets X A" then show "x \ sigma_sets X B" hoelzl@47762: by induct (insert `A \ B`, auto intro: sigma_sets.intros(2-)) hoelzl@47762: qed hoelzl@47762: hoelzl@47762: lemma sigma_sets_superset_generator: "A \ sigma_sets X A" hoelzl@47762: by (auto intro: sigma_sets.Basic) hoelzl@47762: hoelzl@38656: lemma (in sigma_algebra) restriction_in_sets: hoelzl@38656: fixes A :: "nat \ 'a set" hoelzl@47694: assumes "S \ M" hoelzl@47694: and *: "range A \ (\A. S \ A) ` M" (is "_ \ ?r") hoelzl@47694: shows "range A \ M" "(\i. A i) \ (\A. S \ A) ` M" hoelzl@38656: proof - hoelzl@38656: { fix i have "A i \ ?r" using * by auto hoelzl@47694: hence "\B. A i = B \ S \ B \ M" by auto hoelzl@47694: hence "A i \ S" "A i \ M" using `S \ M` by auto } hoelzl@47694: thus "range A \ M" "(\i. A i) \ (\A. S \ A) ` M" hoelzl@38656: by (auto intro!: image_eqI[of _ _ "(\i. A i)"]) hoelzl@38656: qed hoelzl@38656: hoelzl@38656: lemma (in sigma_algebra) restricted_sigma_algebra: hoelzl@47694: assumes "S \ M" hoelzl@47694: shows "sigma_algebra S (restricted_space S)" hoelzl@38656: unfolding sigma_algebra_def sigma_algebra_axioms_def hoelzl@38656: proof safe hoelzl@47694: show "algebra S (restricted_space S)" using restricted_algebra[OF assms] . hoelzl@38656: next hoelzl@47694: fix A :: "nat \ 'a set" assume "range A \ restricted_space S" hoelzl@38656: from restriction_in_sets[OF assms this[simplified]] hoelzl@47694: show "(\i. A i) \ restricted_space S" by simp hoelzl@38656: qed hoelzl@38656: hoelzl@40859: lemma sigma_sets_Int: hoelzl@41689: assumes "A \ sigma_sets sp st" "A \ sp" hoelzl@41689: shows "op \ A ` sigma_sets sp st = sigma_sets A (op \ A ` st)" hoelzl@40859: proof (intro equalityI subsetI) hoelzl@40859: fix x assume "x \ op \ A ` sigma_sets sp st" hoelzl@40859: then obtain y where "y \ sigma_sets sp st" "x = y \ A" by auto hoelzl@41689: then have "x \ sigma_sets (A \ sp) (op \ A ` st)" hoelzl@40859: proof (induct arbitrary: x) hoelzl@40859: case (Compl a) hoelzl@40859: then show ?case hoelzl@40859: by (force intro!: sigma_sets.Compl simp: Diff_Int_distrib ac_simps) hoelzl@40859: next hoelzl@40859: case (Union a) hoelzl@40859: then show ?case hoelzl@40859: by (auto intro!: sigma_sets.Union hoelzl@40859: simp add: UN_extend_simps simp del: UN_simps) hoelzl@47694: qed (auto intro!: sigma_sets.intros(2-)) hoelzl@41689: then show "x \ sigma_sets A (op \ A ` st)" hoelzl@41689: using `A \ sp` by (simp add: Int_absorb2) hoelzl@40859: next hoelzl@41689: fix x assume "x \ sigma_sets A (op \ A ` st)" hoelzl@40859: then show "x \ op \ A ` sigma_sets sp st" hoelzl@40859: proof induct hoelzl@40859: case (Compl a) hoelzl@40859: then obtain x where "a = A \ x" "x \ sigma_sets sp st" by auto hoelzl@41689: then show ?case using `A \ sp` hoelzl@40859: by (force simp add: image_iff intro!: bexI[of _ "sp - x"] sigma_sets.Compl) hoelzl@40859: next hoelzl@40859: case (Union a) hoelzl@40859: then have "\i. \x. x \ sigma_sets sp st \ a i = A \ x" hoelzl@40859: by (auto simp: image_iff Bex_def) hoelzl@40859: from choice[OF this] guess f .. hoelzl@40859: then show ?case hoelzl@40859: by (auto intro!: bexI[of _ "(\x. f x)"] sigma_sets.Union hoelzl@40859: simp add: image_iff) hoelzl@47694: qed (auto intro!: sigma_sets.intros(2-)) hoelzl@40859: qed hoelzl@40859: hoelzl@47694: lemma sigma_sets_empty_eq: "sigma_sets A {} = {{}, A}" hoelzl@40859: proof (intro set_eqI iffI) hoelzl@47694: fix a assume "a \ sigma_sets A {}" then show "a \ {{}, A}" hoelzl@47694: by induct blast+ hoelzl@47694: qed (auto intro: sigma_sets.Empty sigma_sets_top) hoelzl@47694: hoelzl@47694: lemma sigma_sets_single[simp]: "sigma_sets A {A} = {{}, A}" hoelzl@47694: proof (intro set_eqI iffI) hoelzl@47694: fix x assume "x \ sigma_sets A {A}" hoelzl@47694: then show "x \ {{}, A}" hoelzl@47694: by induct blast+ hoelzl@40859: next hoelzl@47694: fix x assume "x \ {{}, A}" hoelzl@47694: then show "x \ sigma_sets A {A}" hoelzl@40859: by (auto intro: sigma_sets.Empty sigma_sets_top) hoelzl@40859: qed hoelzl@40859: hoelzl@42987: lemma sigma_sets_sigma_sets_eq: hoelzl@42987: "M \ Pow S \ sigma_sets S (sigma_sets S M) = sigma_sets S M" hoelzl@47694: by (rule sigma_algebra.sigma_sets_eq[OF sigma_algebra_sigma_sets, of M S]) auto hoelzl@42987: hoelzl@42984: lemma sigma_sets_singleton: hoelzl@42984: assumes "X \ S" hoelzl@42984: shows "sigma_sets S { X } = { {}, X, S - X, S }" hoelzl@42984: proof - hoelzl@47694: interpret sigma_algebra S "{ {}, X, S - X, S }" hoelzl@42984: by (rule sigma_algebra_single_set) fact hoelzl@42984: have "sigma_sets S { X } \ sigma_sets S { {}, X, S - X, S }" hoelzl@42984: by (rule sigma_sets_subseteq) simp hoelzl@42984: moreover have "\ = { {}, X, S - X, S }" hoelzl@47694: using sigma_sets_eq by simp hoelzl@42984: moreover hoelzl@42984: { fix A assume "A \ { {}, X, S - X, S }" hoelzl@42984: then have "A \ sigma_sets S { X }" hoelzl@47694: by (auto intro: sigma_sets.intros(2-) sigma_sets_top) } hoelzl@42984: ultimately have "sigma_sets S { X } = sigma_sets S { {}, X, S - X, S }" hoelzl@42984: by (intro antisym) auto hoelzl@47694: with sigma_sets_eq show ?thesis by simp hoelzl@42984: qed hoelzl@42984: hoelzl@42863: lemma restricted_sigma: hoelzl@47694: assumes S: "S \ sigma_sets \ M" and M: "M \ Pow \" hoelzl@47694: shows "algebra.restricted_space (sigma_sets \ M) S = hoelzl@47694: sigma_sets S (algebra.restricted_space M S)" hoelzl@42863: proof - hoelzl@42863: from S sigma_sets_into_sp[OF M] hoelzl@47694: have "S \ sigma_sets \ M" "S \ \" by auto hoelzl@42863: from sigma_sets_Int[OF this] hoelzl@47694: show ?thesis by simp hoelzl@42863: qed hoelzl@42863: hoelzl@42987: lemma sigma_sets_vimage_commute: hoelzl@47694: assumes X: "X \ \ \ \'" hoelzl@47694: shows "{X -` A \ \ |A. A \ sigma_sets \' M'} hoelzl@47694: = sigma_sets \ {X -` A \ \ |A. A \ M'}" (is "?L = ?R") hoelzl@42987: proof hoelzl@42987: show "?L \ ?R" hoelzl@42987: proof clarify hoelzl@47694: fix A assume "A \ sigma_sets \' M'" hoelzl@47694: then show "X -` A \ \ \ ?R" hoelzl@42987: proof induct hoelzl@42987: case Empty then show ?case hoelzl@42987: by (auto intro!: sigma_sets.Empty) hoelzl@42987: next hoelzl@42987: case (Compl B) hoelzl@47694: have [simp]: "X -` (\' - B) \ \ = \ - (X -` B \ \)" hoelzl@42987: by (auto simp add: funcset_mem [OF X]) hoelzl@42987: with Compl show ?case hoelzl@42987: by (auto intro!: sigma_sets.Compl) hoelzl@42987: next hoelzl@42987: case (Union F) hoelzl@42987: then show ?case hoelzl@42987: by (auto simp add: vimage_UN UN_extend_simps(4) simp del: UN_simps hoelzl@42987: intro!: sigma_sets.Union) hoelzl@47694: qed auto hoelzl@42987: qed hoelzl@42987: show "?R \ ?L" hoelzl@42987: proof clarify hoelzl@42987: fix A assume "A \ ?R" hoelzl@47694: then show "\B. A = X -` B \ \ \ B \ sigma_sets \' M'" hoelzl@42987: proof induct hoelzl@42987: case (Basic B) then show ?case by auto hoelzl@42987: next hoelzl@42987: case Empty then show ?case hoelzl@47694: by (auto intro!: sigma_sets.Empty exI[of _ "{}"]) hoelzl@42987: next hoelzl@42987: case (Compl B) hoelzl@47694: then obtain A where A: "B = X -` A \ \" "A \ sigma_sets \' M'" by auto hoelzl@47694: then have [simp]: "\ - B = X -` (\' - A) \ \" hoelzl@42987: by (auto simp add: funcset_mem [OF X]) hoelzl@42987: with A(2) show ?case hoelzl@47694: by (auto intro: sigma_sets.Compl) hoelzl@42987: next hoelzl@42987: case (Union F) hoelzl@47694: then have "\i. \B. F i = X -` B \ \ \ B \ sigma_sets \' M'" by auto hoelzl@42987: from choice[OF this] guess A .. note A = this hoelzl@42987: with A show ?case hoelzl@47694: by (auto simp: vimage_UN[symmetric] intro: sigma_sets.Union) hoelzl@42987: qed hoelzl@42987: qed hoelzl@42987: qed hoelzl@42987: hoelzl@56994: subsubsection "Disjoint families" hoelzl@38656: hoelzl@38656: definition hoelzl@38656: disjoint_family_on where hoelzl@38656: "disjoint_family_on A S \ (\m\S. \n\S. m \ n \ A m \ A n = {})" hoelzl@38656: hoelzl@38656: abbreviation hoelzl@38656: "disjoint_family A \ disjoint_family_on A UNIV" hoelzl@38656: hoelzl@38656: lemma range_subsetD: "range f \ B \ f i \ B" hoelzl@38656: by blast hoelzl@38656: hoelzl@57447: lemma disjoint_family_onD: "disjoint_family_on A I \ i \ I \ j \ I \ i \ j \ A i \ A j = {}" hoelzl@57447: by (auto simp: disjoint_family_on_def) hoelzl@57447: hoelzl@38656: lemma Int_Diff_disjoint: "A \ B \ (A - B) = {}" hoelzl@38656: by blast hoelzl@38656: hoelzl@38656: lemma Int_Diff_Un: "A \ B \ (A - B) = A" hoelzl@38656: by blast hoelzl@38656: hoelzl@38656: lemma disjoint_family_subset: hoelzl@38656: "disjoint_family A \ (!!x. B x \ A x) \ disjoint_family B" hoelzl@38656: by (force simp add: disjoint_family_on_def) hoelzl@38656: hoelzl@40859: lemma disjoint_family_on_bisimulation: hoelzl@40859: assumes "disjoint_family_on f S" hoelzl@40859: and "\n m. n \ S \ m \ S \ n \ m \ f n \ f m = {} \ g n \ g m = {}" hoelzl@40859: shows "disjoint_family_on g S" hoelzl@40859: using assms unfolding disjoint_family_on_def by auto hoelzl@40859: hoelzl@38656: lemma disjoint_family_on_mono: hoelzl@38656: "A \ B \ disjoint_family_on f B \ disjoint_family_on f A" hoelzl@38656: unfolding disjoint_family_on_def by auto hoelzl@38656: hoelzl@38656: lemma disjoint_family_Suc: hoelzl@38656: assumes Suc: "!!n. A n \ A (Suc n)" hoelzl@38656: shows "disjoint_family (\i. A (Suc i) - A i)" hoelzl@38656: proof - hoelzl@38656: { hoelzl@38656: fix m hoelzl@38656: have "!!n. A n \ A (m+n)" hoelzl@38656: proof (induct m) hoelzl@38656: case 0 show ?case by simp hoelzl@38656: next hoelzl@38656: case (Suc m) thus ?case haftmann@57512: by (metis Suc_eq_plus1 assms add.commute add.left_commute subset_trans) hoelzl@38656: qed hoelzl@38656: } hoelzl@38656: hence "!!m n. m < n \ A m \ A n" haftmann@57512: by (metis add.commute le_add_diff_inverse nat_less_le) hoelzl@38656: thus ?thesis hoelzl@38656: by (auto simp add: disjoint_family_on_def) hoelzl@38656: (metis insert_absorb insert_subset le_SucE le_antisym not_leE) hoelzl@38656: qed hoelzl@38656: hoelzl@39092: lemma setsum_indicator_disjoint_family: hoelzl@39092: fixes f :: "'d \ 'e::semiring_1" hoelzl@39092: assumes d: "disjoint_family_on A P" and "x \ A j" and "finite P" and "j \ P" hoelzl@39092: shows "(\i\P. f i * indicator (A i) x) = f j" hoelzl@39092: proof - hoelzl@39092: have "P \ {i. x \ A i} = {j}" hoelzl@39092: using d `x \ A j` `j \ P` unfolding disjoint_family_on_def hoelzl@39092: by auto hoelzl@39092: thus ?thesis hoelzl@39092: unfolding indicator_def haftmann@57418: by (simp add: if_distrib setsum.If_cases[OF `finite P`]) hoelzl@39092: qed hoelzl@39092: hoelzl@38656: definition disjointed :: "(nat \ 'a set) \ nat \ 'a set " hoelzl@38656: where "disjointed A n = A n - (\i\{0..i\{0..i\{0..i. disjointed A i) = (\i. A i)" hoelzl@38656: apply (rule UN_finite2_eq [where k=0]) hoelzl@38656: apply (simp add: finite_UN_disjointed_eq) hoelzl@38656: done hoelzl@38656: hoelzl@38656: lemma less_disjoint_disjointed: "m disjointed A m \ disjointed A n = {}" hoelzl@38656: by (auto simp add: disjointed_def) hoelzl@38656: hoelzl@38656: lemma disjoint_family_disjointed: "disjoint_family (disjointed A)" hoelzl@38656: by (simp add: disjoint_family_on_def) hoelzl@38656: (metis neq_iff Int_commute less_disjoint_disjointed) hoelzl@38656: hoelzl@38656: lemma disjointed_subset: "disjointed A n \ A n" hoelzl@38656: by (auto simp add: disjointed_def) hoelzl@38656: hoelzl@42065: lemma (in ring_of_sets) UNION_in_sets: hoelzl@38656: fixes A:: "nat \ 'a set" hoelzl@47694: assumes A: "range A \ M" hoelzl@47694: shows "(\i\{0.. M" hoelzl@38656: proof (induct n) hoelzl@38656: case 0 show ?case by simp hoelzl@38656: next hoelzl@38656: case (Suc n) hoelzl@38656: thus ?case hoelzl@38656: by (simp add: atLeastLessThanSuc) (metis A Un UNIV_I image_subset_iff) hoelzl@38656: qed hoelzl@38656: hoelzl@42065: lemma (in ring_of_sets) range_disjointed_sets: hoelzl@47694: assumes A: "range A \ M" hoelzl@47694: shows "range (disjointed A) \ M" hoelzl@38656: proof (auto simp add: disjointed_def) hoelzl@38656: fix n hoelzl@47694: show "A n - (\i\{0.. M" using UNION_in_sets hoelzl@38656: by (metis A Diff UNIV_I image_subset_iff) hoelzl@38656: qed hoelzl@38656: hoelzl@42065: lemma (in algebra) range_disjointed_sets': hoelzl@47694: "range A \ M \ range (disjointed A) \ M" hoelzl@42065: using range_disjointed_sets . hoelzl@42065: hoelzl@42145: lemma disjointed_0[simp]: "disjointed A 0 = A 0" hoelzl@42145: by (simp add: disjointed_def) hoelzl@42145: hoelzl@42145: lemma incseq_Un: hoelzl@42145: "incseq A \ (\i\n. A i) = A n" hoelzl@42145: unfolding incseq_def by auto hoelzl@42145: hoelzl@42145: lemma disjointed_incseq: hoelzl@42145: "incseq A \ disjointed A (Suc n) = A (Suc n) - A n" hoelzl@42145: using incseq_Un[of A] hoelzl@42145: by (simp add: disjointed_def atLeastLessThanSuc_atLeastAtMost atLeast0AtMost) hoelzl@42145: hoelzl@38656: lemma sigma_algebra_disjoint_iff: hoelzl@47694: "sigma_algebra \ M \ algebra \ M \ hoelzl@47694: (\A. range A \ M \ disjoint_family A \ (\i::nat. A i) \ M)" hoelzl@38656: proof (auto simp add: sigma_algebra_iff) hoelzl@38656: fix A :: "nat \ 'a set" hoelzl@47694: assume M: "algebra \ M" hoelzl@47694: and A: "range A \ M" hoelzl@47694: and UnA: "\A. range A \ M \ disjoint_family A \ (\i::nat. A i) \ M" hoelzl@47694: hence "range (disjointed A) \ M \ hoelzl@38656: disjoint_family (disjointed A) \ hoelzl@47694: (\i. disjointed A i) \ M" by blast hoelzl@47694: hence "(\i. disjointed A i) \ M" hoelzl@47694: by (simp add: algebra.range_disjointed_sets'[of \] M A disjoint_family_disjointed) hoelzl@47694: thus "(\i::nat. A i) \ M" by (simp add: UN_disjointed_eq) hoelzl@47694: qed hoelzl@47694: hoelzl@47762: lemma disjoint_family_on_disjoint_image: hoelzl@47762: "disjoint_family_on A I \ disjoint (A ` I)" hoelzl@47762: unfolding disjoint_family_on_def disjoint_def by force hoelzl@47762: hoelzl@47762: lemma disjoint_image_disjoint_family_on: hoelzl@47762: assumes d: "disjoint (A ` I)" and i: "inj_on A I" hoelzl@47762: shows "disjoint_family_on A I" hoelzl@47762: unfolding disjoint_family_on_def hoelzl@47762: proof (intro ballI impI) hoelzl@47762: fix n m assume nm: "m \ I" "n \ I" and "n \ m" hoelzl@47762: with i[THEN inj_onD, of n m] show "A n \ A m = {}" hoelzl@47762: by (intro disjointD[OF d]) auto hoelzl@47762: qed hoelzl@47762: hoelzl@56994: subsubsection {* Ring generated by a semiring *} hoelzl@47762: hoelzl@47762: definition (in semiring_of_sets) hoelzl@47762: "generated_ring = { \C | C. C \ M \ finite C \ disjoint C }" hoelzl@47762: hoelzl@47762: lemma (in semiring_of_sets) generated_ringE[elim?]: hoelzl@47762: assumes "a \ generated_ring" hoelzl@47762: obtains C where "finite C" "disjoint C" "C \ M" "a = \C" hoelzl@47762: using assms unfolding generated_ring_def by auto hoelzl@47762: hoelzl@47762: lemma (in semiring_of_sets) generated_ringI[intro?]: hoelzl@47762: assumes "finite C" "disjoint C" "C \ M" "a = \C" hoelzl@47762: shows "a \ generated_ring" hoelzl@47762: using assms unfolding generated_ring_def by auto hoelzl@47762: hoelzl@47762: lemma (in semiring_of_sets) generated_ringI_Basic: hoelzl@47762: "A \ M \ A \ generated_ring" hoelzl@47762: by (rule generated_ringI[of "{A}"]) (auto simp: disjoint_def) hoelzl@47762: hoelzl@47762: lemma (in semiring_of_sets) generated_ring_disjoint_Un[intro]: hoelzl@47762: assumes a: "a \ generated_ring" and b: "b \ generated_ring" hoelzl@47762: and "a \ b = {}" hoelzl@47762: shows "a \ b \ generated_ring" hoelzl@47762: proof - hoelzl@47762: from a guess Ca .. note Ca = this hoelzl@47762: from b guess Cb .. note Cb = this hoelzl@47762: show ?thesis hoelzl@47762: proof hoelzl@47762: show "disjoint (Ca \ Cb)" hoelzl@47762: using `a \ b = {}` Ca Cb by (auto intro!: disjoint_union) hoelzl@47762: qed (insert Ca Cb, auto) hoelzl@47762: qed hoelzl@47762: hoelzl@47762: lemma (in semiring_of_sets) generated_ring_empty: "{} \ generated_ring" hoelzl@47762: by (auto simp: generated_ring_def disjoint_def) hoelzl@47762: hoelzl@47762: lemma (in semiring_of_sets) generated_ring_disjoint_Union: hoelzl@47762: assumes "finite A" shows "A \ generated_ring \ disjoint A \ \A \ generated_ring" hoelzl@47762: using assms by (induct A) (auto simp: disjoint_def intro!: generated_ring_disjoint_Un generated_ring_empty) hoelzl@47762: hoelzl@47762: lemma (in semiring_of_sets) generated_ring_disjoint_UNION: hoelzl@47762: "finite I \ disjoint (A ` I) \ (\i. i \ I \ A i \ generated_ring) \ UNION I A \ generated_ring" hoelzl@47762: unfolding SUP_def by (intro generated_ring_disjoint_Union) auto hoelzl@47762: hoelzl@47762: lemma (in semiring_of_sets) generated_ring_Int: hoelzl@47762: assumes a: "a \ generated_ring" and b: "b \ generated_ring" hoelzl@47762: shows "a \ b \ generated_ring" hoelzl@47762: proof - hoelzl@47762: from a guess Ca .. note Ca = this hoelzl@47762: from b guess Cb .. note Cb = this hoelzl@47762: def C \ "(\(a,b). a \ b)` (Ca\Cb)" hoelzl@47762: show ?thesis hoelzl@47762: proof hoelzl@47762: show "disjoint C" hoelzl@47762: proof (simp add: disjoint_def C_def, intro ballI impI) hoelzl@47762: fix a1 b1 a2 b2 assume sets: "a1 \ Ca" "b1 \ Cb" "a2 \ Ca" "b2 \ Cb" hoelzl@47762: assume "a1 \ b1 \ a2 \ b2" hoelzl@47762: then have "a1 \ a2 \ b1 \ b2" by auto hoelzl@47762: then show "(a1 \ b1) \ (a2 \ b2) = {}" hoelzl@47762: proof hoelzl@47762: assume "a1 \ a2" hoelzl@47762: with sets Ca have "a1 \ a2 = {}" hoelzl@47762: by (auto simp: disjoint_def) hoelzl@47762: then show ?thesis by auto hoelzl@47762: next hoelzl@47762: assume "b1 \ b2" hoelzl@47762: with sets Cb have "b1 \ b2 = {}" hoelzl@47762: by (auto simp: disjoint_def) hoelzl@47762: then show ?thesis by auto hoelzl@47762: qed hoelzl@47762: qed hoelzl@47762: qed (insert Ca Cb, auto simp: C_def) hoelzl@47762: qed hoelzl@47762: hoelzl@47762: lemma (in semiring_of_sets) generated_ring_Inter: hoelzl@47762: assumes "finite A" "A \ {}" shows "A \ generated_ring \ \A \ generated_ring" hoelzl@47762: using assms by (induct A rule: finite_ne_induct) (auto intro: generated_ring_Int) hoelzl@47762: hoelzl@47762: lemma (in semiring_of_sets) generated_ring_INTER: hoelzl@47762: "finite I \ I \ {} \ (\i. i \ I \ A i \ generated_ring) \ INTER I A \ generated_ring" hoelzl@47762: unfolding INF_def by (intro generated_ring_Inter) auto hoelzl@47762: hoelzl@47762: lemma (in semiring_of_sets) generating_ring: hoelzl@47762: "ring_of_sets \ generated_ring" hoelzl@47762: proof (rule ring_of_setsI) hoelzl@47762: let ?R = generated_ring hoelzl@47762: show "?R \ Pow \" hoelzl@47762: using sets_into_space by (auto simp: generated_ring_def generated_ring_empty) hoelzl@47762: show "{} \ ?R" by (rule generated_ring_empty) hoelzl@47762: hoelzl@47762: { fix a assume a: "a \ ?R" then guess Ca .. note Ca = this hoelzl@47762: fix b assume b: "b \ ?R" then guess Cb .. note Cb = this hoelzl@47762: hoelzl@47762: show "a - b \ ?R" hoelzl@47762: proof cases hoelzl@47762: assume "Cb = {}" with Cb `a \ ?R` show ?thesis hoelzl@47762: by simp hoelzl@47762: next hoelzl@47762: assume "Cb \ {}" hoelzl@47762: with Ca Cb have "a - b = (\a'\Ca. \b'\Cb. a' - b')" by auto hoelzl@47762: also have "\ \ ?R" hoelzl@47762: proof (intro generated_ring_INTER generated_ring_disjoint_UNION) hoelzl@47762: fix a b assume "a \ Ca" "b \ Cb" hoelzl@47762: with Ca Cb Diff_cover[of a b] show "a - b \ ?R" hoelzl@47762: by (auto simp add: generated_ring_def) hoelzl@47762: next hoelzl@47762: show "disjoint ((\a'. \b'\Cb. a' - b')`Ca)" hoelzl@47762: using Ca by (auto simp add: disjoint_def `Cb \ {}`) hoelzl@47762: next hoelzl@47762: show "finite Ca" "finite Cb" "Cb \ {}" by fact+ hoelzl@47762: qed hoelzl@47762: finally show "a - b \ ?R" . hoelzl@47762: qed } hoelzl@47762: note Diff = this hoelzl@47762: hoelzl@47762: fix a b assume sets: "a \ ?R" "b \ ?R" hoelzl@47762: have "a \ b = (a - b) \ (a \ b) \ (b - a)" by auto hoelzl@47762: also have "\ \ ?R" hoelzl@47762: by (intro sets generated_ring_disjoint_Un generated_ring_Int Diff) auto hoelzl@47762: finally show "a \ b \ ?R" . hoelzl@47762: qed hoelzl@47762: hoelzl@47762: lemma (in semiring_of_sets) sigma_sets_generated_ring_eq: "sigma_sets \ generated_ring = sigma_sets \ M" hoelzl@47762: proof hoelzl@47762: interpret M: sigma_algebra \ "sigma_sets \ M" hoelzl@47762: using space_closed by (rule sigma_algebra_sigma_sets) hoelzl@47762: show "sigma_sets \ generated_ring \ sigma_sets \ M" hoelzl@47762: by (blast intro!: sigma_sets_mono elim: generated_ringE) hoelzl@47762: qed (auto intro!: generated_ringI_Basic sigma_sets_mono) hoelzl@47762: hoelzl@56994: subsubsection {* A Two-Element Series *} hoelzl@38656: hoelzl@38656: definition binaryset :: "'a set \ 'a set \ nat \ 'a set " wenzelm@50252: where "binaryset A B = (\x. {})(0 := A, Suc 0 := B)" hoelzl@38656: hoelzl@38656: lemma range_binaryset_eq: "range(binaryset A B) = {A,B,{}}" hoelzl@38656: apply (simp add: binaryset_def) nipkow@39302: apply (rule set_eqI) hoelzl@38656: apply (auto simp add: image_iff) hoelzl@38656: done hoelzl@38656: hoelzl@38656: lemma UN_binaryset_eq: "(\i. binaryset A B i) = A \ B" haftmann@44106: by (simp add: SUP_def range_binaryset_eq) hoelzl@38656: hoelzl@56994: subsubsection {* Closed CDI *} hoelzl@38656: hoelzl@47694: definition closed_cdi where hoelzl@47694: "closed_cdi \ M \ hoelzl@47694: M \ Pow \ & hoelzl@47694: (\s \ M. \ - s \ M) & hoelzl@47694: (\A. (range A \ M) & (A 0 = {}) & (\n. A n \ A (Suc n)) \ hoelzl@47694: (\i. A i) \ M) & hoelzl@47694: (\A. (range A \ M) & disjoint_family A \ (\i::nat. A i) \ M)" hoelzl@38656: hoelzl@38656: inductive_set hoelzl@47694: smallest_ccdi_sets :: "'a set \ 'a set set \ 'a set set" hoelzl@47694: for \ M hoelzl@38656: where hoelzl@38656: Basic [intro]: hoelzl@47694: "a \ M \ a \ smallest_ccdi_sets \ M" hoelzl@38656: | Compl [intro]: hoelzl@47694: "a \ smallest_ccdi_sets \ M \ \ - a \ smallest_ccdi_sets \ M" hoelzl@38656: | Inc: hoelzl@47694: "range A \ Pow(smallest_ccdi_sets \ M) \ A 0 = {} \ (\n. A n \ A (Suc n)) hoelzl@47694: \ (\i. A i) \ smallest_ccdi_sets \ M" hoelzl@38656: | Disj: hoelzl@47694: "range A \ Pow(smallest_ccdi_sets \ M) \ disjoint_family A hoelzl@47694: \ (\i::nat. A i) \ smallest_ccdi_sets \ M" hoelzl@38656: hoelzl@47694: lemma (in subset_class) smallest_closed_cdi1: "M \ smallest_ccdi_sets \ M" hoelzl@47694: by auto hoelzl@38656: hoelzl@47694: lemma (in subset_class) smallest_ccdi_sets: "smallest_ccdi_sets \ M \ Pow \" hoelzl@38656: apply (rule subsetI) hoelzl@38656: apply (erule smallest_ccdi_sets.induct) hoelzl@38656: apply (auto intro: range_subsetD dest: sets_into_space) hoelzl@38656: done hoelzl@38656: hoelzl@47694: lemma (in subset_class) smallest_closed_cdi2: "closed_cdi \ (smallest_ccdi_sets \ M)" hoelzl@47694: apply (auto simp add: closed_cdi_def smallest_ccdi_sets) hoelzl@38656: apply (blast intro: smallest_ccdi_sets.Inc smallest_ccdi_sets.Disj) + hoelzl@38656: done hoelzl@38656: hoelzl@47694: lemma closed_cdi_subset: "closed_cdi \ M \ M \ Pow \" hoelzl@38656: by (simp add: closed_cdi_def) hoelzl@38656: hoelzl@47694: lemma closed_cdi_Compl: "closed_cdi \ M \ s \ M \ \ - s \ M" hoelzl@38656: by (simp add: closed_cdi_def) hoelzl@38656: hoelzl@38656: lemma closed_cdi_Inc: hoelzl@47694: "closed_cdi \ M \ range A \ M \ A 0 = {} \ (!!n. A n \ A (Suc n)) \ (\i. A i) \ M" hoelzl@38656: by (simp add: closed_cdi_def) hoelzl@38656: hoelzl@38656: lemma closed_cdi_Disj: hoelzl@47694: "closed_cdi \ M \ range A \ M \ disjoint_family A \ (\i::nat. A i) \ M" hoelzl@38656: by (simp add: closed_cdi_def) hoelzl@38656: hoelzl@38656: lemma closed_cdi_Un: hoelzl@47694: assumes cdi: "closed_cdi \ M" and empty: "{} \ M" hoelzl@47694: and A: "A \ M" and B: "B \ M" hoelzl@38656: and disj: "A \ B = {}" hoelzl@47694: shows "A \ B \ M" hoelzl@38656: proof - hoelzl@47694: have ra: "range (binaryset A B) \ M" hoelzl@38656: by (simp add: range_binaryset_eq empty A B) hoelzl@38656: have di: "disjoint_family (binaryset A B)" using disj hoelzl@38656: by (simp add: disjoint_family_on_def binaryset_def Int_commute) hoelzl@38656: from closed_cdi_Disj [OF cdi ra di] hoelzl@38656: show ?thesis hoelzl@38656: by (simp add: UN_binaryset_eq) hoelzl@38656: qed hoelzl@38656: hoelzl@38656: lemma (in algebra) smallest_ccdi_sets_Un: hoelzl@47694: assumes A: "A \ smallest_ccdi_sets \ M" and B: "B \ smallest_ccdi_sets \ M" hoelzl@38656: and disj: "A \ B = {}" hoelzl@47694: shows "A \ B \ smallest_ccdi_sets \ M" hoelzl@38656: proof - hoelzl@47694: have ra: "range (binaryset A B) \ Pow (smallest_ccdi_sets \ M)" hoelzl@38656: by (simp add: range_binaryset_eq A B smallest_ccdi_sets.Basic) hoelzl@38656: have di: "disjoint_family (binaryset A B)" using disj hoelzl@38656: by (simp add: disjoint_family_on_def binaryset_def Int_commute) hoelzl@38656: from Disj [OF ra di] hoelzl@38656: show ?thesis hoelzl@38656: by (simp add: UN_binaryset_eq) hoelzl@38656: qed hoelzl@38656: hoelzl@38656: lemma (in algebra) smallest_ccdi_sets_Int1: hoelzl@47694: assumes a: "a \ M" hoelzl@47694: shows "b \ smallest_ccdi_sets \ M \ a \ b \ smallest_ccdi_sets \ M" hoelzl@38656: proof (induct rule: smallest_ccdi_sets.induct) hoelzl@38656: case (Basic x) hoelzl@38656: thus ?case hoelzl@38656: by (metis a Int smallest_ccdi_sets.Basic) hoelzl@38656: next hoelzl@38656: case (Compl x) hoelzl@47694: have "a \ (\ - x) = \ - ((\ - a) \ (a \ x))" hoelzl@38656: by blast hoelzl@47694: also have "... \ smallest_ccdi_sets \ M" hoelzl@38656: by (metis smallest_ccdi_sets.Compl a Compl(2) Diff_Int2 Diff_Int_distrib2 hoelzl@47694: Diff_disjoint Int_Diff Int_empty_right smallest_ccdi_sets_Un hoelzl@47694: smallest_ccdi_sets.Basic smallest_ccdi_sets.Compl) hoelzl@38656: finally show ?case . hoelzl@38656: next hoelzl@38656: case (Inc A) hoelzl@38656: have 1: "(\i. (\i. a \ A i) i) = a \ (\i. A i)" hoelzl@38656: by blast hoelzl@47694: have "range (\i. a \ A i) \ Pow(smallest_ccdi_sets \ M)" using Inc hoelzl@38656: by blast hoelzl@38656: moreover have "(\i. a \ A i) 0 = {}" hoelzl@38656: by (simp add: Inc) hoelzl@38656: moreover have "!!n. (\i. a \ A i) n \ (\i. a \ A i) (Suc n)" using Inc hoelzl@38656: by blast hoelzl@47694: ultimately have 2: "(\i. (\i. a \ A i) i) \ smallest_ccdi_sets \ M" hoelzl@38656: by (rule smallest_ccdi_sets.Inc) hoelzl@38656: show ?case hoelzl@38656: by (metis 1 2) hoelzl@38656: next hoelzl@38656: case (Disj A) hoelzl@38656: have 1: "(\i. (\i. a \ A i) i) = a \ (\i. A i)" hoelzl@38656: by blast hoelzl@47694: have "range (\i. a \ A i) \ Pow(smallest_ccdi_sets \ M)" using Disj hoelzl@38656: by blast hoelzl@38656: moreover have "disjoint_family (\i. a \ A i)" using Disj hoelzl@38656: by (auto simp add: disjoint_family_on_def) hoelzl@47694: ultimately have 2: "(\i. (\i. a \ A i) i) \ smallest_ccdi_sets \ M" hoelzl@38656: by (rule smallest_ccdi_sets.Disj) hoelzl@38656: show ?case hoelzl@38656: by (metis 1 2) hoelzl@38656: qed hoelzl@38656: hoelzl@38656: hoelzl@38656: lemma (in algebra) smallest_ccdi_sets_Int: hoelzl@47694: assumes b: "b \ smallest_ccdi_sets \ M" hoelzl@47694: shows "a \ smallest_ccdi_sets \ M \ a \ b \ smallest_ccdi_sets \ M" hoelzl@38656: proof (induct rule: smallest_ccdi_sets.induct) hoelzl@38656: case (Basic x) hoelzl@38656: thus ?case hoelzl@38656: by (metis b smallest_ccdi_sets_Int1) hoelzl@38656: next hoelzl@38656: case (Compl x) hoelzl@47694: have "(\ - x) \ b = \ - (x \ b \ (\ - b))" hoelzl@38656: by blast hoelzl@47694: also have "... \ smallest_ccdi_sets \ M" hoelzl@38656: by (metis Compl(2) Diff_disjoint Int_Diff Int_commute Int_empty_right b hoelzl@38656: smallest_ccdi_sets.Compl smallest_ccdi_sets_Un) hoelzl@38656: finally show ?case . hoelzl@38656: next hoelzl@38656: case (Inc A) hoelzl@38656: have 1: "(\i. (\i. A i \ b) i) = (\i. A i) \ b" hoelzl@38656: by blast hoelzl@47694: have "range (\i. A i \ b) \ Pow(smallest_ccdi_sets \ M)" using Inc hoelzl@38656: by blast hoelzl@38656: moreover have "(\i. A i \ b) 0 = {}" hoelzl@38656: by (simp add: Inc) hoelzl@38656: moreover have "!!n. (\i. A i \ b) n \ (\i. A i \ b) (Suc n)" using Inc hoelzl@38656: by blast hoelzl@47694: ultimately have 2: "(\i. (\i. A i \ b) i) \ smallest_ccdi_sets \ M" hoelzl@38656: by (rule smallest_ccdi_sets.Inc) hoelzl@38656: show ?case hoelzl@38656: by (metis 1 2) hoelzl@38656: next hoelzl@38656: case (Disj A) hoelzl@38656: have 1: "(\i. (\i. A i \ b) i) = (\i. A i) \ b" hoelzl@38656: by blast hoelzl@47694: have "range (\i. A i \ b) \ Pow(smallest_ccdi_sets \ M)" using Disj hoelzl@38656: by blast hoelzl@38656: moreover have "disjoint_family (\i. A i \ b)" using Disj hoelzl@38656: by (auto simp add: disjoint_family_on_def) hoelzl@47694: ultimately have 2: "(\i. (\i. A i \ b) i) \ smallest_ccdi_sets \ M" hoelzl@38656: by (rule smallest_ccdi_sets.Disj) hoelzl@38656: show ?case hoelzl@38656: by (metis 1 2) hoelzl@38656: qed hoelzl@38656: hoelzl@38656: lemma (in algebra) sigma_property_disjoint_lemma: hoelzl@47694: assumes sbC: "M \ C" hoelzl@47694: and ccdi: "closed_cdi \ C" hoelzl@47694: shows "sigma_sets \ M \ C" hoelzl@38656: proof - hoelzl@47694: have "smallest_ccdi_sets \ M \ {B . M \ B \ sigma_algebra \ B}" hoelzl@38656: apply (auto simp add: sigma_algebra_disjoint_iff algebra_iff_Int hoelzl@38656: smallest_ccdi_sets_Int) hoelzl@38656: apply (metis Union_Pow_eq Union_upper subsetD smallest_ccdi_sets) hoelzl@38656: apply (blast intro: smallest_ccdi_sets.Disj) hoelzl@38656: done hoelzl@47694: hence "sigma_sets (\) (M) \ smallest_ccdi_sets \ M" hoelzl@38656: by clarsimp hoelzl@47694: (drule sigma_algebra.sigma_sets_subset [where a="M"], auto) hoelzl@38656: also have "... \ C" hoelzl@38656: proof hoelzl@38656: fix x hoelzl@47694: assume x: "x \ smallest_ccdi_sets \ M" hoelzl@38656: thus "x \ C" hoelzl@38656: proof (induct rule: smallest_ccdi_sets.induct) hoelzl@38656: case (Basic x) hoelzl@38656: thus ?case hoelzl@38656: by (metis Basic subsetD sbC) hoelzl@38656: next hoelzl@38656: case (Compl x) hoelzl@38656: thus ?case hoelzl@38656: by (blast intro: closed_cdi_Compl [OF ccdi, simplified]) hoelzl@38656: next hoelzl@38656: case (Inc A) hoelzl@38656: thus ?case hoelzl@38656: by (auto intro: closed_cdi_Inc [OF ccdi, simplified]) hoelzl@38656: next hoelzl@38656: case (Disj A) hoelzl@38656: thus ?case hoelzl@38656: by (auto intro: closed_cdi_Disj [OF ccdi, simplified]) hoelzl@38656: qed hoelzl@38656: qed hoelzl@38656: finally show ?thesis . hoelzl@38656: qed hoelzl@38656: hoelzl@38656: lemma (in algebra) sigma_property_disjoint: hoelzl@47694: assumes sbC: "M \ C" hoelzl@47694: and compl: "!!s. s \ C \ sigma_sets (\) (M) \ \ - s \ C" hoelzl@47694: and inc: "!!A. range A \ C \ sigma_sets (\) (M) hoelzl@38656: \ A 0 = {} \ (!!n. A n \ A (Suc n)) hoelzl@38656: \ (\i. A i) \ C" hoelzl@47694: and disj: "!!A. range A \ C \ sigma_sets (\) (M) hoelzl@38656: \ disjoint_family A \ (\i::nat. A i) \ C" hoelzl@47694: shows "sigma_sets (\) (M) \ C" hoelzl@38656: proof - hoelzl@47694: have "sigma_sets (\) (M) \ C \ sigma_sets (\) (M)" hoelzl@38656: proof (rule sigma_property_disjoint_lemma) hoelzl@47694: show "M \ C \ sigma_sets (\) (M)" hoelzl@38656: by (metis Int_greatest Set.subsetI sbC sigma_sets.Basic) hoelzl@38656: next hoelzl@47694: show "closed_cdi \ (C \ sigma_sets (\) (M))" hoelzl@38656: by (simp add: closed_cdi_def compl inc disj) hoelzl@38656: (metis PowI Set.subsetI le_infI2 sigma_sets_into_sp space_closed hoelzl@38656: IntE sigma_sets.Compl range_subsetD sigma_sets.Union) hoelzl@38656: qed hoelzl@38656: thus ?thesis hoelzl@38656: by blast hoelzl@38656: qed hoelzl@38656: hoelzl@56994: subsubsection {* Dynkin systems *} hoelzl@40859: hoelzl@42065: locale dynkin_system = subset_class + hoelzl@47694: assumes space: "\ \ M" hoelzl@47694: and compl[intro!]: "\A. A \ M \ \ - A \ M" hoelzl@47694: and UN[intro!]: "\A. disjoint_family A \ range A \ M hoelzl@47694: \ (\i::nat. A i) \ M" hoelzl@40859: hoelzl@47694: lemma (in dynkin_system) empty[intro, simp]: "{} \ M" hoelzl@47694: using space compl[of "\"] by simp hoelzl@40859: hoelzl@40859: lemma (in dynkin_system) diff: hoelzl@47694: assumes sets: "D \ M" "E \ M" and "D \ E" hoelzl@47694: shows "E - D \ M" hoelzl@40859: proof - hoelzl@47694: let ?f = "\x. if x = 0 then D else if x = Suc 0 then \ - E else {}" hoelzl@47694: have "range ?f = {D, \ - E, {}}" hoelzl@40859: by (auto simp: image_iff) hoelzl@47694: moreover have "D \ (\ - E) = (\i. ?f i)" hoelzl@40859: by (auto simp: image_iff split: split_if_asm) hoelzl@40859: moreover wenzelm@53374: have "disjoint_family ?f" unfolding disjoint_family_on_def hoelzl@47694: using `D \ M`[THEN sets_into_space] `D \ E` by auto hoelzl@47694: ultimately have "\ - (D \ (\ - E)) \ M" hoelzl@40859: using sets by auto hoelzl@47694: also have "\ - (D \ (\ - E)) = E - D" hoelzl@40859: using assms sets_into_space by auto hoelzl@40859: finally show ?thesis . hoelzl@40859: qed hoelzl@40859: hoelzl@40859: lemma dynkin_systemI: hoelzl@47694: assumes "\ A. A \ M \ A \ \" "\ \ M" hoelzl@47694: assumes "\ A. A \ M \ \ - A \ M" hoelzl@47694: assumes "\ A. disjoint_family A \ range A \ M hoelzl@47694: \ (\i::nat. A i) \ M" hoelzl@47694: shows "dynkin_system \ M" hoelzl@42065: using assms by (auto simp: dynkin_system_def dynkin_system_axioms_def subset_class_def) hoelzl@40859: hoelzl@42988: lemma dynkin_systemI': hoelzl@47694: assumes 1: "\ A. A \ M \ A \ \" hoelzl@47694: assumes empty: "{} \ M" hoelzl@47694: assumes Diff: "\ A. A \ M \ \ - A \ M" hoelzl@47694: assumes 2: "\ A. disjoint_family A \ range A \ M hoelzl@47694: \ (\i::nat. A i) \ M" hoelzl@47694: shows "dynkin_system \ M" hoelzl@42988: proof - hoelzl@47694: from Diff[OF empty] have "\ \ M" by auto hoelzl@42988: from 1 this Diff 2 show ?thesis hoelzl@42988: by (intro dynkin_systemI) auto hoelzl@42988: qed hoelzl@42988: hoelzl@40859: lemma dynkin_system_trivial: hoelzl@47694: shows "dynkin_system A (Pow A)" hoelzl@40859: by (rule dynkin_systemI) auto hoelzl@40859: hoelzl@40859: lemma sigma_algebra_imp_dynkin_system: hoelzl@47694: assumes "sigma_algebra \ M" shows "dynkin_system \ M" hoelzl@40859: proof - hoelzl@47694: interpret sigma_algebra \ M by fact nipkow@44890: show ?thesis using sets_into_space by (fastforce intro!: dynkin_systemI) hoelzl@40859: qed hoelzl@40859: hoelzl@56994: subsubsection "Intersection sets systems" hoelzl@40859: hoelzl@47694: definition "Int_stable M \ (\ a \ M. \ b \ M. a \ b \ M)" hoelzl@40859: hoelzl@40859: lemma (in algebra) Int_stable: "Int_stable M" hoelzl@40859: unfolding Int_stable_def by auto hoelzl@40859: hoelzl@42981: lemma Int_stableI: hoelzl@47694: "(\a b. a \ A \ b \ A \ a \ b \ A) \ Int_stable A" hoelzl@42981: unfolding Int_stable_def by auto hoelzl@42981: hoelzl@42981: lemma Int_stableD: hoelzl@47694: "Int_stable M \ a \ M \ b \ M \ a \ b \ M" hoelzl@42981: unfolding Int_stable_def by auto hoelzl@42981: hoelzl@40859: lemma (in dynkin_system) sigma_algebra_eq_Int_stable: hoelzl@47694: "sigma_algebra \ M \ Int_stable M" hoelzl@40859: proof hoelzl@47694: assume "sigma_algebra \ M" then show "Int_stable M" hoelzl@40859: unfolding sigma_algebra_def using algebra.Int_stable by auto hoelzl@40859: next hoelzl@40859: assume "Int_stable M" hoelzl@47694: show "sigma_algebra \ M" hoelzl@42065: unfolding sigma_algebra_disjoint_iff algebra_iff_Un hoelzl@40859: proof (intro conjI ballI allI impI) hoelzl@47694: show "M \ Pow (\)" using sets_into_space by auto hoelzl@40859: next hoelzl@47694: fix A B assume "A \ M" "B \ M" hoelzl@47694: then have "A \ B = \ - ((\ - A) \ (\ - B))" hoelzl@47694: "\ - A \ M" "\ - B \ M" hoelzl@40859: using sets_into_space by auto hoelzl@47694: then show "A \ B \ M" hoelzl@40859: using `Int_stable M` unfolding Int_stable_def by auto hoelzl@40859: qed auto hoelzl@40859: qed hoelzl@40859: hoelzl@56994: subsubsection "Smallest Dynkin systems" hoelzl@40859: hoelzl@41689: definition dynkin where hoelzl@47694: "dynkin \ M = (\{D. dynkin_system \ D \ M \ D})" hoelzl@40859: hoelzl@40859: lemma dynkin_system_dynkin: hoelzl@47694: assumes "M \ Pow (\)" hoelzl@47694: shows "dynkin_system \ (dynkin \ M)" hoelzl@40859: proof (rule dynkin_systemI) hoelzl@47694: fix A assume "A \ dynkin \ M" hoelzl@40859: moreover hoelzl@47694: { fix D assume "A \ D" and d: "dynkin_system \ D" hoelzl@47694: then have "A \ \" by (auto simp: dynkin_system_def subset_class_def) } hoelzl@47694: moreover have "{D. dynkin_system \ D \ M \ D} \ {}" nipkow@44890: using assms dynkin_system_trivial by fastforce hoelzl@47694: ultimately show "A \ \" hoelzl@40859: unfolding dynkin_def using assms hoelzl@47694: by auto hoelzl@40859: next hoelzl@47694: show "\ \ dynkin \ M" nipkow@44890: unfolding dynkin_def using dynkin_system.space by fastforce hoelzl@40859: next hoelzl@47694: fix A assume "A \ dynkin \ M" hoelzl@47694: then show "\ - A \ dynkin \ M" hoelzl@40859: unfolding dynkin_def using dynkin_system.compl by force hoelzl@40859: next hoelzl@40859: fix A :: "nat \ 'a set" hoelzl@47694: assume A: "disjoint_family A" "range A \ dynkin \ M" hoelzl@47694: show "(\i. A i) \ dynkin \ M" unfolding dynkin_def hoelzl@40859: proof (simp, safe) hoelzl@47694: fix D assume "dynkin_system \ D" "M \ D" hoelzl@47694: with A have "(\i. A i) \ D" hoelzl@40859: by (intro dynkin_system.UN) (auto simp: dynkin_def) hoelzl@40859: then show "(\i. A i) \ D" by auto hoelzl@40859: qed hoelzl@40859: qed hoelzl@40859: hoelzl@47694: lemma dynkin_Basic[intro]: "A \ M \ A \ dynkin \ M" hoelzl@40859: unfolding dynkin_def by auto hoelzl@40859: hoelzl@40859: lemma (in dynkin_system) restricted_dynkin_system: hoelzl@47694: assumes "D \ M" hoelzl@47694: shows "dynkin_system \ {Q. Q \ \ \ Q \ D \ M}" hoelzl@40859: proof (rule dynkin_systemI, simp_all) hoelzl@47694: have "\ \ D = D" hoelzl@47694: using `D \ M` sets_into_space by auto hoelzl@47694: then show "\ \ D \ M" hoelzl@47694: using `D \ M` by auto hoelzl@40859: next hoelzl@47694: fix A assume "A \ \ \ A \ D \ M" hoelzl@47694: moreover have "(\ - A) \ D = (\ - (A \ D)) - (\ - D)" hoelzl@40859: by auto hoelzl@47694: ultimately show "\ - A \ \ \ (\ - A) \ D \ M" hoelzl@47694: using `D \ M` by (auto intro: diff) hoelzl@40859: next hoelzl@40859: fix A :: "nat \ 'a set" hoelzl@47694: assume "disjoint_family A" "range A \ {Q. Q \ \ \ Q \ D \ M}" hoelzl@47694: then have "\i. A i \ \" "disjoint_family (\i. A i \ D)" hoelzl@47694: "range (\i. A i \ D) \ M" "(\x. A x) \ D = (\x. A x \ D)" nipkow@44890: by ((fastforce simp: disjoint_family_on_def)+) hoelzl@47694: then show "(\x. A x) \ \ \ (\x. A x) \ D \ M" hoelzl@40859: by (auto simp del: UN_simps) hoelzl@40859: qed hoelzl@40859: hoelzl@40859: lemma (in dynkin_system) dynkin_subset: hoelzl@47694: assumes "N \ M" hoelzl@47694: shows "dynkin \ N \ M" hoelzl@40859: proof - hoelzl@47694: have "dynkin_system \ M" by default hoelzl@47694: then have "dynkin_system \ M" hoelzl@42065: using assms unfolding dynkin_system_def dynkin_system_axioms_def subset_class_def by simp hoelzl@47694: with `N \ M` show ?thesis by (auto simp add: dynkin_def) hoelzl@40859: qed hoelzl@40859: hoelzl@40859: lemma sigma_eq_dynkin: hoelzl@47694: assumes sets: "M \ Pow \" hoelzl@40859: assumes "Int_stable M" hoelzl@47694: shows "sigma_sets \ M = dynkin \ M" hoelzl@40859: proof - hoelzl@47694: have "dynkin \ M \ sigma_sets (\) (M)" hoelzl@40859: using sigma_algebra_imp_dynkin_system hoelzl@47694: unfolding dynkin_def sigma_sets_least_sigma_algebra[OF sets] by auto hoelzl@40859: moreover hoelzl@47694: interpret dynkin_system \ "dynkin \ M" hoelzl@40859: using dynkin_system_dynkin[OF sets] . hoelzl@47694: have "sigma_algebra \ (dynkin \ M)" hoelzl@40859: unfolding sigma_algebra_eq_Int_stable Int_stable_def hoelzl@40859: proof (intro ballI) hoelzl@47694: fix A B assume "A \ dynkin \ M" "B \ dynkin \ M" hoelzl@47694: let ?D = "\E. {Q. Q \ \ \ Q \ E \ dynkin \ M}" hoelzl@47694: have "M \ ?D B" hoelzl@40859: proof hoelzl@47694: fix E assume "E \ M" hoelzl@47694: then have "M \ ?D E" "E \ dynkin \ M" hoelzl@40859: using sets_into_space `Int_stable M` by (auto simp: Int_stable_def) hoelzl@47694: then have "dynkin \ M \ ?D E" hoelzl@47694: using restricted_dynkin_system `E \ dynkin \ M` hoelzl@40859: by (intro dynkin_system.dynkin_subset) simp_all hoelzl@47694: then have "B \ ?D E" hoelzl@47694: using `B \ dynkin \ M` by auto hoelzl@47694: then have "E \ B \ dynkin \ M" hoelzl@40859: by (subst Int_commute) simp hoelzl@47694: then show "E \ ?D B" hoelzl@47694: using sets `E \ M` by auto hoelzl@40859: qed hoelzl@47694: then have "dynkin \ M \ ?D B" hoelzl@47694: using restricted_dynkin_system `B \ dynkin \ M` hoelzl@40859: by (intro dynkin_system.dynkin_subset) simp_all hoelzl@47694: then show "A \ B \ dynkin \ M" hoelzl@47694: using `A \ dynkin \ M` sets_into_space by auto hoelzl@40859: qed hoelzl@47694: from sigma_algebra.sigma_sets_subset[OF this, of "M"] hoelzl@47694: have "sigma_sets (\) (M) \ dynkin \ M" by auto hoelzl@47694: ultimately have "sigma_sets (\) (M) = dynkin \ M" by auto hoelzl@40859: then show ?thesis hoelzl@47694: by (auto simp: dynkin_def) hoelzl@40859: qed hoelzl@40859: hoelzl@40859: lemma (in dynkin_system) dynkin_idem: hoelzl@47694: "dynkin \ M = M" hoelzl@40859: proof - hoelzl@47694: have "dynkin \ M = M" hoelzl@40859: proof hoelzl@47694: show "M \ dynkin \ M" hoelzl@40859: using dynkin_Basic by auto hoelzl@47694: show "dynkin \ M \ M" hoelzl@40859: by (intro dynkin_subset) auto hoelzl@40859: qed hoelzl@40859: then show ?thesis hoelzl@47694: by (auto simp: dynkin_def) hoelzl@40859: qed hoelzl@40859: hoelzl@40859: lemma (in dynkin_system) dynkin_lemma: hoelzl@41689: assumes "Int_stable E" hoelzl@47694: and E: "E \ M" "M \ sigma_sets \ E" hoelzl@47694: shows "sigma_sets \ E = M" hoelzl@40859: proof - hoelzl@47694: have "E \ Pow \" hoelzl@41689: using E sets_into_space by force wenzelm@53374: then have *: "sigma_sets \ E = dynkin \ E" hoelzl@40859: using `Int_stable E` by (rule sigma_eq_dynkin) wenzelm@53374: then have "dynkin \ E = M" hoelzl@47694: using assms dynkin_subset[OF E(1)] by simp wenzelm@53374: with * show ?thesis hoelzl@47694: using assms by (auto simp: dynkin_def) hoelzl@42864: qed hoelzl@42864: hoelzl@56994: subsubsection {* Induction rule for intersection-stable generators *} hoelzl@56994: hoelzl@56994: text {* The reason to introduce Dynkin-systems is the following induction rules for $\sigma$-algebras hoelzl@56994: generated by a generator closed under intersection. *} hoelzl@56994: hoelzl@49789: lemma sigma_sets_induct_disjoint[consumes 3, case_names basic empty compl union]: hoelzl@49789: assumes "Int_stable G" hoelzl@49789: and closed: "G \ Pow \" hoelzl@49789: and A: "A \ sigma_sets \ G" hoelzl@49789: assumes basic: "\A. A \ G \ P A" hoelzl@49789: and empty: "P {}" hoelzl@49789: and compl: "\A. A \ sigma_sets \ G \ P A \ P (\ - A)" hoelzl@49789: and union: "\A. disjoint_family A \ range A \ sigma_sets \ G \ (\i. P (A i)) \ P (\i::nat. A i)" hoelzl@49789: shows "P A" hoelzl@49789: proof - hoelzl@49789: let ?D = "{ A \ sigma_sets \ G. P A }" hoelzl@49789: interpret sigma_algebra \ "sigma_sets \ G" hoelzl@49789: using closed by (rule sigma_algebra_sigma_sets) hoelzl@49789: from compl[OF _ empty] closed have space: "P \" by simp hoelzl@49789: interpret dynkin_system \ ?D hoelzl@49789: by default (auto dest: sets_into_space intro!: space compl union) hoelzl@49789: have "sigma_sets \ G = ?D" hoelzl@49789: by (rule dynkin_lemma) (auto simp: basic `Int_stable G`) hoelzl@49789: with A show ?thesis by auto hoelzl@49789: qed hoelzl@49789: hoelzl@56994: subsection {* Measure type *} hoelzl@56994: hoelzl@56994: definition positive :: "'a set set \ ('a set \ ereal) \ bool" where hoelzl@56994: "positive M \ \ \ {} = 0 \ (\A\M. 0 \ \ A)" hoelzl@56994: hoelzl@56994: definition countably_additive :: "'a set set \ ('a set \ ereal) \ bool" where hoelzl@56994: "countably_additive M f \ (\A. range A \ M \ disjoint_family A \ (\i. A i) \ M \ hoelzl@56994: (\i. f (A i)) = f (\i. A i))" hoelzl@56994: hoelzl@56994: definition measure_space :: "'a set \ 'a set set \ ('a set \ ereal) \ bool" where hoelzl@56994: "measure_space \ A \ \ sigma_algebra \ A \ positive A \ \ countably_additive A \" hoelzl@56994: hoelzl@56994: typedef 'a measure = "{(\::'a set, A, \). (\a\-A. \ a = 0) \ measure_space \ A \ }" hoelzl@56994: proof hoelzl@56994: have "sigma_algebra UNIV {{}, UNIV}" hoelzl@56994: by (auto simp: sigma_algebra_iff2) hoelzl@56994: then show "(UNIV, {{}, UNIV}, \A. 0) \ {(\, A, \). (\a\-A. \ a = 0) \ measure_space \ A \} " hoelzl@56994: by (auto simp: measure_space_def positive_def countably_additive_def) hoelzl@56994: qed hoelzl@56994: hoelzl@56994: definition space :: "'a measure \ 'a set" where hoelzl@56994: "space M = fst (Rep_measure M)" hoelzl@56994: hoelzl@56994: definition sets :: "'a measure \ 'a set set" where hoelzl@56994: "sets M = fst (snd (Rep_measure M))" hoelzl@56994: hoelzl@56994: definition emeasure :: "'a measure \ 'a set \ ereal" where hoelzl@56994: "emeasure M = snd (snd (Rep_measure M))" hoelzl@56994: hoelzl@56994: definition measure :: "'a measure \ 'a set \ real" where hoelzl@56994: "measure M A = real (emeasure M A)" hoelzl@56994: hoelzl@56994: declare [[coercion sets]] hoelzl@56994: hoelzl@56994: declare [[coercion measure]] hoelzl@56994: hoelzl@56994: declare [[coercion emeasure]] hoelzl@56994: hoelzl@56994: lemma measure_space: "measure_space (space M) (sets M) (emeasure M)" hoelzl@56994: by (cases M) (auto simp: space_def sets_def emeasure_def Abs_measure_inverse) hoelzl@56994: hoelzl@56994: interpretation sets!: sigma_algebra "space M" "sets M" for M :: "'a measure" hoelzl@56994: using measure_space[of M] by (auto simp: measure_space_def) hoelzl@56994: hoelzl@56994: definition measure_of :: "'a set \ 'a set set \ ('a set \ ereal) \ 'a measure" where hoelzl@56994: "measure_of \ A \ = Abs_measure (\, if A \ Pow \ then sigma_sets \ A else {{}, \}, hoelzl@56994: \a. if a \ sigma_sets \ A \ measure_space \ (sigma_sets \ A) \ then \ a else 0)" hoelzl@56994: hoelzl@56994: abbreviation "sigma \ A \ measure_of \ A (\x. 0)" hoelzl@56994: hoelzl@56994: lemma measure_space_0: "A \ Pow \ \ measure_space \ (sigma_sets \ A) (\x. 0)" hoelzl@56994: unfolding measure_space_def hoelzl@56994: by (auto intro!: sigma_algebra_sigma_sets simp: positive_def countably_additive_def) hoelzl@56994: hoelzl@56994: lemma sigma_algebra_trivial: "sigma_algebra \ {{}, \}" hoelzl@56994: by unfold_locales(fastforce intro: exI[where x="{{}}"] exI[where x="{\}"])+ hoelzl@56994: hoelzl@56994: lemma measure_space_0': "measure_space \ {{}, \} (\x. 0)" hoelzl@56994: by(simp add: measure_space_def positive_def countably_additive_def sigma_algebra_trivial) hoelzl@56994: hoelzl@56994: lemma measure_space_closed: hoelzl@56994: assumes "measure_space \ M \" hoelzl@56994: shows "M \ Pow \" hoelzl@56994: proof - hoelzl@56994: interpret sigma_algebra \ M using assms by(simp add: measure_space_def) hoelzl@56994: show ?thesis by(rule space_closed) hoelzl@56994: qed hoelzl@56994: hoelzl@56994: lemma (in ring_of_sets) positive_cong_eq: hoelzl@56994: "(\a. a \ M \ \' a = \ a) \ positive M \' = positive M \" hoelzl@56994: by (auto simp add: positive_def) hoelzl@56994: hoelzl@56994: lemma (in sigma_algebra) countably_additive_eq: hoelzl@56994: "(\a. a \ M \ \' a = \ a) \ countably_additive M \' = countably_additive M \" hoelzl@56994: unfolding countably_additive_def hoelzl@56994: by (intro arg_cong[where f=All] ext) (auto simp add: countably_additive_def subset_eq) hoelzl@56994: hoelzl@56994: lemma measure_space_eq: hoelzl@56994: assumes closed: "A \ Pow \" and eq: "\a. a \ sigma_sets \ A \ \ a = \' a" hoelzl@56994: shows "measure_space \ (sigma_sets \ A) \ = measure_space \ (sigma_sets \ A) \'" hoelzl@56994: proof - hoelzl@56994: interpret sigma_algebra \ "sigma_sets \ A" using closed by (rule sigma_algebra_sigma_sets) hoelzl@56994: from positive_cong_eq[OF eq, of "\i. i"] countably_additive_eq[OF eq, of "\i. i"] show ?thesis hoelzl@56994: by (auto simp: measure_space_def) hoelzl@56994: qed hoelzl@56994: hoelzl@56994: lemma measure_of_eq: hoelzl@56994: assumes closed: "A \ Pow \" and eq: "(\a. a \ sigma_sets \ A \ \ a = \' a)" hoelzl@56994: shows "measure_of \ A \ = measure_of \ A \'" hoelzl@56994: proof - hoelzl@56994: have "measure_space \ (sigma_sets \ A) \ = measure_space \ (sigma_sets \ A) \'" hoelzl@56994: using assms by (rule measure_space_eq) hoelzl@56994: with eq show ?thesis hoelzl@56994: by (auto simp add: measure_of_def intro!: arg_cong[where f=Abs_measure]) hoelzl@56994: qed hoelzl@56994: hoelzl@56994: lemma hoelzl@56994: shows space_measure_of_conv: "space (measure_of \ A \) = \" (is ?space) hoelzl@56994: and sets_measure_of_conv: hoelzl@56994: "sets (measure_of \ A \) = (if A \ Pow \ then sigma_sets \ A else {{}, \})" (is ?sets) hoelzl@56994: and emeasure_measure_of_conv: hoelzl@56994: "emeasure (measure_of \ A \) = hoelzl@56994: (\B. if B \ sigma_sets \ A \ measure_space \ (sigma_sets \ A) \ then \ B else 0)" (is ?emeasure) hoelzl@56994: proof - hoelzl@56994: have "?space \ ?sets \ ?emeasure" hoelzl@56994: proof(cases "measure_space \ (sigma_sets \ A) \") hoelzl@56994: case True hoelzl@56994: from measure_space_closed[OF this] sigma_sets_superset_generator[of A \] hoelzl@56994: have "A \ Pow \" by simp hoelzl@56994: hence "measure_space \ (sigma_sets \ A) \ = measure_space \ (sigma_sets \ A) hoelzl@56994: (\a. if a \ sigma_sets \ A then \ a else 0)" hoelzl@56994: by(rule measure_space_eq) auto hoelzl@56994: with True `A \ Pow \` show ?thesis hoelzl@56994: by(simp add: measure_of_def space_def sets_def emeasure_def Abs_measure_inverse) hoelzl@56994: next hoelzl@56994: case False thus ?thesis hoelzl@56994: by(cases "A \ Pow \")(simp_all add: Abs_measure_inverse measure_of_def sets_def space_def emeasure_def measure_space_0 measure_space_0') hoelzl@56994: qed hoelzl@56994: thus ?space ?sets ?emeasure by simp_all hoelzl@56994: qed hoelzl@56994: hoelzl@56994: lemma [simp]: hoelzl@56994: assumes A: "A \ Pow \" hoelzl@56994: shows sets_measure_of: "sets (measure_of \ A \) = sigma_sets \ A" hoelzl@56994: and space_measure_of: "space (measure_of \ A \) = \" hoelzl@56994: using assms hoelzl@56994: by(simp_all add: sets_measure_of_conv space_measure_of_conv) hoelzl@56994: hoelzl@56994: lemma (in sigma_algebra) sets_measure_of_eq[simp]: "sets (measure_of \ M \) = M" hoelzl@56994: using space_closed by (auto intro!: sigma_sets_eq) hoelzl@56994: hoelzl@56994: lemma (in sigma_algebra) space_measure_of_eq[simp]: "space (measure_of \ M \) = \" hoelzl@56994: by (rule space_measure_of_conv) hoelzl@56994: hoelzl@56994: lemma measure_of_subset: "M \ Pow \ \ M' \ M \ sets (measure_of \ M' \) \ sets (measure_of \ M \')" hoelzl@56994: by (auto intro!: sigma_sets_subseteq) hoelzl@56994: hoelzl@56994: lemma sigma_sets_mono'': hoelzl@56994: assumes "A \ sigma_sets C D" hoelzl@56994: assumes "B \ D" hoelzl@56994: assumes "D \ Pow C" hoelzl@56994: shows "sigma_sets A B \ sigma_sets C D" hoelzl@56994: proof hoelzl@56994: fix x assume "x \ sigma_sets A B" hoelzl@56994: thus "x \ sigma_sets C D" hoelzl@56994: proof induct hoelzl@56994: case (Basic a) with assms have "a \ D" by auto hoelzl@56994: thus ?case .. hoelzl@56994: next hoelzl@56994: case Empty show ?case by (rule sigma_sets.Empty) hoelzl@56994: next hoelzl@56994: from assms have "A \ sets (sigma C D)" by (subst sets_measure_of[OF `D \ Pow C`]) hoelzl@56994: moreover case (Compl a) hence "a \ sets (sigma C D)" by (subst sets_measure_of[OF `D \ Pow C`]) hoelzl@56994: ultimately have "A - a \ sets (sigma C D)" .. hoelzl@56994: thus ?case by (subst (asm) sets_measure_of[OF `D \ Pow C`]) hoelzl@56994: next hoelzl@56994: case (Union a) hoelzl@56994: thus ?case by (intro sigma_sets.Union) hoelzl@56994: qed hoelzl@56994: qed hoelzl@56994: hoelzl@56994: lemma in_measure_of[intro, simp]: "M \ Pow \ \ A \ M \ A \ sets (measure_of \ M \)" hoelzl@56994: by auto hoelzl@56994: hoelzl@56994: subsubsection {* Constructing simple @{typ "'a measure"} *} hoelzl@56994: hoelzl@56994: lemma emeasure_measure_of: hoelzl@56994: assumes M: "M = measure_of \ A \" hoelzl@56994: assumes ms: "A \ Pow \" "positive (sets M) \" "countably_additive (sets M) \" hoelzl@56994: assumes X: "X \ sets M" hoelzl@56994: shows "emeasure M X = \ X" hoelzl@56994: proof - hoelzl@56994: interpret sigma_algebra \ "sigma_sets \ A" by (rule sigma_algebra_sigma_sets) fact hoelzl@56994: have "measure_space \ (sigma_sets \ A) \" hoelzl@56994: using ms M by (simp add: measure_space_def sigma_algebra_sigma_sets) hoelzl@56994: thus ?thesis using X ms hoelzl@56994: by(simp add: M emeasure_measure_of_conv sets_measure_of_conv) hoelzl@56994: qed hoelzl@56994: hoelzl@56994: lemma emeasure_measure_of_sigma: hoelzl@56994: assumes ms: "sigma_algebra \ M" "positive M \" "countably_additive M \" hoelzl@56994: assumes A: "A \ M" hoelzl@56994: shows "emeasure (measure_of \ M \) A = \ A" hoelzl@56994: proof - hoelzl@56994: interpret sigma_algebra \ M by fact hoelzl@56994: have "measure_space \ (sigma_sets \ M) \" hoelzl@56994: using ms sigma_sets_eq by (simp add: measure_space_def) hoelzl@56994: thus ?thesis by(simp add: emeasure_measure_of_conv A) hoelzl@56994: qed hoelzl@56994: hoelzl@56994: lemma measure_cases[cases type: measure]: hoelzl@56994: obtains (measure) \ A \ where "x = Abs_measure (\, A, \)" "\a\-A. \ a = 0" "measure_space \ A \" hoelzl@56994: by atomize_elim (cases x, auto) hoelzl@56994: hoelzl@56994: lemma sets_eq_imp_space_eq: hoelzl@56994: "sets M = sets M' \ space M = space M'" hoelzl@56994: using sets.top[of M] sets.top[of M'] sets.space_closed[of M] sets.space_closed[of M'] hoelzl@56994: by blast hoelzl@56994: hoelzl@56994: lemma emeasure_notin_sets: "A \ sets M \ emeasure M A = 0" hoelzl@56994: by (cases M) (auto simp: sets_def emeasure_def Abs_measure_inverse measure_space_def) hoelzl@56994: hoelzl@56994: lemma emeasure_neq_0_sets: "emeasure M A \ 0 \ A \ sets M" hoelzl@56994: using emeasure_notin_sets[of A M] by blast hoelzl@56994: hoelzl@56994: lemma measure_notin_sets: "A \ sets M \ measure M A = 0" hoelzl@56994: by (simp add: measure_def emeasure_notin_sets) hoelzl@56994: hoelzl@56994: lemma measure_eqI: hoelzl@56994: fixes M N :: "'a measure" hoelzl@56994: assumes "sets M = sets N" and eq: "\A. A \ sets M \ emeasure M A = emeasure N A" hoelzl@56994: shows "M = N" hoelzl@56994: proof (cases M N rule: measure_cases[case_product measure_cases]) hoelzl@56994: case (measure_measure \ A \ \' A' \') hoelzl@56994: interpret M: sigma_algebra \ A using measure_measure by (auto simp: measure_space_def) hoelzl@56994: interpret N: sigma_algebra \' A' using measure_measure by (auto simp: measure_space_def) hoelzl@56994: have "A = sets M" "A' = sets N" hoelzl@56994: using measure_measure by (simp_all add: sets_def Abs_measure_inverse) hoelzl@56994: with `sets M = sets N` have AA': "A = A'" by simp hoelzl@56994: moreover from M.top N.top M.space_closed N.space_closed AA' have "\ = \'" by auto hoelzl@56994: moreover { fix B have "\ B = \' B" hoelzl@56994: proof cases hoelzl@56994: assume "B \ A" hoelzl@56994: with eq `A = sets M` have "emeasure M B = emeasure N B" by simp hoelzl@56994: with measure_measure show "\ B = \' B" hoelzl@56994: by (simp add: emeasure_def Abs_measure_inverse) hoelzl@56994: next hoelzl@56994: assume "B \ A" hoelzl@56994: with `A = sets M` `A' = sets N` `A = A'` have "B \ sets M" "B \ sets N" hoelzl@56994: by auto hoelzl@56994: then have "emeasure M B = 0" "emeasure N B = 0" hoelzl@56994: by (simp_all add: emeasure_notin_sets) hoelzl@56994: with measure_measure show "\ B = \' B" hoelzl@56994: by (simp add: emeasure_def Abs_measure_inverse) hoelzl@56994: qed } hoelzl@56994: then have "\ = \'" by auto hoelzl@56994: ultimately show "M = N" hoelzl@56994: by (simp add: measure_measure) hoelzl@56994: qed hoelzl@56994: hoelzl@56994: lemma emeasure_sigma: "A \ Pow \ \ emeasure (sigma \ A) = (\_. 0)" hoelzl@56994: using measure_space_0[of A \] hoelzl@56994: by (simp add: measure_of_def emeasure_def Abs_measure_inverse) hoelzl@56994: hoelzl@56994: lemma sigma_eqI: hoelzl@56994: assumes [simp]: "M \ Pow \" "N \ Pow \" "sigma_sets \ M = sigma_sets \ N" hoelzl@56994: shows "sigma \ M = sigma \ N" hoelzl@56994: by (rule measure_eqI) (simp_all add: emeasure_sigma) hoelzl@56994: hoelzl@56994: subsubsection {* Measurable functions *} hoelzl@56994: hoelzl@56994: definition measurable :: "'a measure \ 'b measure \ ('a \ 'b) set" where hoelzl@56994: "measurable A B = {f \ space A -> space B. \y \ sets B. f -` y \ space A \ sets A}" hoelzl@56994: hoelzl@56994: lemma measurable_space: hoelzl@56994: "f \ measurable M A \ x \ space M \ f x \ space A" hoelzl@56994: unfolding measurable_def by auto hoelzl@56994: hoelzl@56994: lemma measurable_sets: hoelzl@56994: "f \ measurable M A \ S \ sets A \ f -` S \ space M \ sets M" hoelzl@56994: unfolding measurable_def by auto hoelzl@56994: hoelzl@56994: lemma measurable_sets_Collect: hoelzl@56994: assumes f: "f \ measurable M N" and P: "{x\space N. P x} \ sets N" shows "{x\space M. P (f x)} \ sets M" hoelzl@56994: proof - hoelzl@56994: have "f -` {x \ space N. P x} \ space M = {x\space M. P (f x)}" hoelzl@56994: using measurable_space[OF f] by auto hoelzl@56994: with measurable_sets[OF f P] show ?thesis hoelzl@56994: by simp hoelzl@56994: qed hoelzl@56994: hoelzl@56994: lemma measurable_sigma_sets: hoelzl@56994: assumes B: "sets N = sigma_sets \ A" "A \ Pow \" hoelzl@56994: and f: "f \ space M \ \" hoelzl@56994: and ba: "\y. y \ A \ (f -` y) \ space M \ sets M" hoelzl@56994: shows "f \ measurable M N" hoelzl@56994: proof - hoelzl@56994: interpret A: sigma_algebra \ "sigma_sets \ A" using B(2) by (rule sigma_algebra_sigma_sets) hoelzl@56994: from B sets.top[of N] A.top sets.space_closed[of N] A.space_closed have \: "\ = space N" by force hoelzl@56994: hoelzl@56994: { fix X assume "X \ sigma_sets \ A" hoelzl@56994: then have "f -` X \ space M \ sets M \ X \ \" hoelzl@56994: proof induct hoelzl@56994: case (Basic a) then show ?case hoelzl@56994: by (auto simp add: ba) (metis B(2) subsetD PowD) hoelzl@56994: next hoelzl@56994: case (Compl a) hoelzl@56994: have [simp]: "f -` \ \ space M = space M" hoelzl@56994: by (auto simp add: funcset_mem [OF f]) hoelzl@56994: then show ?case hoelzl@56994: by (auto simp add: vimage_Diff Diff_Int_distrib2 sets.compl_sets Compl) hoelzl@56994: next hoelzl@56994: case (Union a) hoelzl@56994: then show ?case hoelzl@56994: by (simp add: vimage_UN, simp only: UN_extend_simps(4)) blast hoelzl@56994: qed auto } hoelzl@56994: with f show ?thesis hoelzl@56994: by (auto simp add: measurable_def B \) hoelzl@56994: qed hoelzl@56994: hoelzl@56994: lemma measurable_measure_of: hoelzl@56994: assumes B: "N \ Pow \" hoelzl@56994: and f: "f \ space M \ \" hoelzl@56994: and ba: "\y. y \ N \ (f -` y) \ space M \ sets M" hoelzl@56994: shows "f \ measurable M (measure_of \ N \)" hoelzl@56994: proof - hoelzl@56994: have "sets (measure_of \ N \) = sigma_sets \ N" hoelzl@56994: using B by (rule sets_measure_of) hoelzl@56994: from this assms show ?thesis by (rule measurable_sigma_sets) hoelzl@56994: qed hoelzl@56994: hoelzl@56994: lemma measurable_iff_measure_of: hoelzl@56994: assumes "N \ Pow \" "f \ space M \ \" hoelzl@56994: shows "f \ measurable M (measure_of \ N \) \ (\A\N. f -` A \ space M \ sets M)" hoelzl@56994: by (metis assms in_measure_of measurable_measure_of assms measurable_sets) hoelzl@56994: hoelzl@56994: lemma measurable_cong_sets: hoelzl@56994: assumes sets: "sets M = sets M'" "sets N = sets N'" hoelzl@56994: shows "measurable M N = measurable M' N'" hoelzl@56994: using sets[THEN sets_eq_imp_space_eq] sets by (simp add: measurable_def) hoelzl@56994: hoelzl@56994: lemma measurable_cong: hoelzl@56994: assumes "\ w. w \ space M \ f w = g w" hoelzl@56994: shows "f \ measurable M M' \ g \ measurable M M'" hoelzl@56994: unfolding measurable_def using assms hoelzl@56994: by (simp cong: vimage_inter_cong Pi_cong) hoelzl@56994: hoelzl@56994: lemma measurable_cong_strong: hoelzl@56994: "M = N \ M' = N' \ (\w. w \ space M \ f w = g w) \ hoelzl@56994: f \ measurable M M' \ g \ measurable N N'" hoelzl@56994: by (metis measurable_cong) hoelzl@56994: hoelzl@56994: lemma measurable_eqI: hoelzl@56994: "\ space m1 = space m1' ; space m2 = space m2' ; hoelzl@56994: sets m1 = sets m1' ; sets m2 = sets m2' \ hoelzl@56994: \ measurable m1 m2 = measurable m1' m2'" hoelzl@56994: by (simp add: measurable_def sigma_algebra_iff2) hoelzl@56994: hoelzl@56994: lemma measurable_compose: hoelzl@56994: assumes f: "f \ measurable M N" and g: "g \ measurable N L" hoelzl@56994: shows "(\x. g (f x)) \ measurable M L" hoelzl@56994: proof - hoelzl@56994: have "\A. (\x. g (f x)) -` A \ space M = f -` (g -` A \ space N) \ space M" hoelzl@56994: using measurable_space[OF f] by auto hoelzl@56994: with measurable_space[OF f] measurable_space[OF g] show ?thesis hoelzl@56994: by (auto intro: measurable_sets[OF f] measurable_sets[OF g] hoelzl@56994: simp del: vimage_Int simp add: measurable_def) hoelzl@56994: qed hoelzl@56994: hoelzl@56994: lemma measurable_comp: hoelzl@56994: "f \ measurable M N \ g \ measurable N L \ g \ f \ measurable M L" hoelzl@56994: using measurable_compose[of f M N g L] by (simp add: comp_def) hoelzl@56994: hoelzl@56994: lemma measurable_const: hoelzl@56994: "c \ space M' \ (\x. c) \ measurable M M'" hoelzl@56994: by (auto simp add: measurable_def) hoelzl@56994: hoelzl@56994: lemma measurable_If: hoelzl@56994: assumes measure: "f \ measurable M M'" "g \ measurable M M'" hoelzl@56994: assumes P: "{x\space M. P x} \ sets M" hoelzl@56994: shows "(\x. if P x then f x else g x) \ measurable M M'" hoelzl@56994: unfolding measurable_def hoelzl@56994: proof safe hoelzl@56994: fix x assume "x \ space M" hoelzl@56994: thus "(if P x then f x else g x) \ space M'" hoelzl@56994: using measure unfolding measurable_def by auto hoelzl@56994: next hoelzl@56994: fix A assume "A \ sets M'" hoelzl@56994: hence *: "(\x. if P x then f x else g x) -` A \ space M = hoelzl@56994: ((f -` A \ space M) \ {x\space M. P x}) \ hoelzl@56994: ((g -` A \ space M) \ (space M - {x\space M. P x}))" hoelzl@56994: using measure unfolding measurable_def by (auto split: split_if_asm) hoelzl@56994: show "(\x. if P x then f x else g x) -` A \ space M \ sets M" hoelzl@56994: using `A \ sets M'` measure P unfolding * measurable_def hoelzl@56994: by (auto intro!: sets.Un) hoelzl@56994: qed hoelzl@56994: hoelzl@56994: lemma measurable_If_set: hoelzl@56994: assumes measure: "f \ measurable M M'" "g \ measurable M M'" hoelzl@56994: assumes P: "A \ space M \ sets M" hoelzl@56994: shows "(\x. if x \ A then f x else g x) \ measurable M M'" hoelzl@56994: proof (rule measurable_If[OF measure]) hoelzl@56994: have "{x \ space M. x \ A} = A \ space M" by auto hoelzl@56994: thus "{x \ space M. x \ A} \ sets M" using `A \ space M \ sets M` by auto hoelzl@56994: qed hoelzl@56994: hoelzl@56994: lemma measurable_ident: "id \ measurable M M" hoelzl@56994: by (auto simp add: measurable_def) hoelzl@56994: hoelzl@56994: lemma measurable_ident_sets: hoelzl@56994: assumes eq: "sets M = sets M'" shows "(\x. x) \ measurable M M'" hoelzl@56994: using measurable_ident[of M] hoelzl@56994: unfolding id_def measurable_def eq sets_eq_imp_space_eq[OF eq] . hoelzl@56994: hoelzl@56994: lemma sets_Least: hoelzl@56994: assumes meas: "\i::nat. {x\space M. P i x} \ M" hoelzl@56994: shows "(\x. LEAST j. P j x) -` A \ space M \ sets M" hoelzl@56994: proof - hoelzl@56994: { fix i have "(\x. LEAST j. P j x) -` {i} \ space M \ sets M" hoelzl@56994: proof cases hoelzl@56994: assume i: "(LEAST j. False) = i" hoelzl@56994: have "(\x. LEAST j. P j x) -` {i} \ space M = hoelzl@56994: {x\space M. P i x} \ (space M - (\jspace M. P j x})) \ (space M - (\i. {x\space M. P i x}))" hoelzl@56994: by (simp add: set_eq_iff, safe) hoelzl@56994: (insert i, auto dest: Least_le intro: LeastI intro!: Least_equality) hoelzl@56994: with meas show ?thesis hoelzl@56994: by (auto intro!: sets.Int) hoelzl@56994: next hoelzl@56994: assume i: "(LEAST j. False) \ i" hoelzl@56994: then have "(\x. LEAST j. P j x) -` {i} \ space M = hoelzl@56994: {x\space M. P i x} \ (space M - (\jspace M. P j x}))" hoelzl@56994: proof (simp add: set_eq_iff, safe) hoelzl@56994: fix x assume neq: "(LEAST j. False) \ (LEAST j. P j x)" hoelzl@56994: have "\j. P j x" hoelzl@56994: by (rule ccontr) (insert neq, auto) hoelzl@56994: then show "P (LEAST j. P j x) x" by (rule LeastI_ex) hoelzl@56994: qed (auto dest: Least_le intro!: Least_equality) hoelzl@56994: with meas show ?thesis hoelzl@56994: by auto hoelzl@56994: qed } hoelzl@56994: then have "(\i\A. (\x. LEAST j. P j x) -` {i} \ space M) \ sets M" hoelzl@56994: by (intro sets.countable_UN) auto hoelzl@56994: moreover have "(\i\A. (\x. LEAST j. P j x) -` {i} \ space M) = hoelzl@56994: (\x. LEAST j. P j x) -` A \ space M" by auto hoelzl@56994: ultimately show ?thesis by auto hoelzl@56994: qed hoelzl@56994: hoelzl@56994: lemma measurable_strong: hoelzl@56994: fixes f :: "'a \ 'b" and g :: "'b \ 'c" hoelzl@56994: assumes f: "f \ measurable a b" and g: "g \ space b \ space c" hoelzl@56994: and t: "f ` (space a) \ t" hoelzl@56994: and cb: "\s. s \ sets c \ (g -` s) \ t \ sets b" hoelzl@56994: shows "(g o f) \ measurable a c" hoelzl@56994: proof - hoelzl@56994: have fab: "f \ (space a -> space b)" hoelzl@56994: and ba: "\y. y \ sets b \ (f -` y) \ (space a) \ sets a" using f hoelzl@56994: by (auto simp add: measurable_def) hoelzl@56994: have eq: "\y. (g \ f) -` y \ space a = f -` (g -` y \ t) \ space a" using t hoelzl@56994: by force hoelzl@56994: show ?thesis hoelzl@56994: apply (auto simp add: measurable_def vimage_comp) hoelzl@56994: apply (metis funcset_mem fab g) hoelzl@56994: apply (subst eq) hoelzl@56994: apply (metis ba cb) hoelzl@56994: done hoelzl@56994: qed hoelzl@56994: hoelzl@57275: lemma measurable_discrete_difference: hoelzl@57275: assumes f: "f \ measurable M N" hoelzl@57275: assumes X: "countable X" hoelzl@57275: assumes sets: "\x. x \ X \ {x} \ sets M" hoelzl@57275: assumes space: "\x. x \ X \ g x \ space N" hoelzl@57275: assumes eq: "\x. x \ space M \ x \ X \ f x = g x" hoelzl@57275: shows "g \ measurable M N" hoelzl@57275: unfolding measurable_def hoelzl@57275: proof safe hoelzl@57275: fix x assume "x \ space M" then show "g x \ space N" hoelzl@57275: using measurable_space[OF f, of x] eq[of x] space[of x] by (cases "x \ X") auto hoelzl@57275: next hoelzl@57275: fix S assume S: "S \ sets N" hoelzl@57275: have "g -` S \ space M = (f -` S \ space M) - (\x\X. {x}) \ (\x\{x\X. g x \ S}. {x})" hoelzl@57275: using sets.sets_into_space[OF sets] eq by auto hoelzl@57275: also have "\ \ sets M" hoelzl@57275: by (safe intro!: sets.Diff sets.Un measurable_sets[OF f] S sets.countable_UN' X countable_Collect sets) hoelzl@57275: finally show "g -` S \ space M \ sets M" . hoelzl@57275: qed hoelzl@57275: hoelzl@56994: lemma measurable_mono1: hoelzl@56994: "M' \ Pow \ \ M \ M' \ hoelzl@56994: measurable (measure_of \ M \) N \ measurable (measure_of \ M' \') N" hoelzl@56994: using measure_of_subset[of M' \ M] by (auto simp add: measurable_def) hoelzl@56994: hoelzl@56994: subsubsection {* Counting space *} hoelzl@56994: hoelzl@56994: definition count_space :: "'a set \ 'a measure" where hoelzl@56994: "count_space \ = measure_of \ (Pow \) (\A. if finite A then ereal (card A) else \)" hoelzl@56994: hoelzl@56994: lemma hoelzl@56994: shows space_count_space[simp]: "space (count_space \) = \" hoelzl@56994: and sets_count_space[simp]: "sets (count_space \) = Pow \" hoelzl@56994: using sigma_sets_into_sp[of "Pow \" \] hoelzl@56994: by (auto simp: count_space_def) hoelzl@56994: hoelzl@56994: lemma measurable_count_space_eq1[simp]: hoelzl@56994: "f \ measurable (count_space A) M \ f \ A \ space M" hoelzl@56994: unfolding measurable_def by simp hoelzl@56994: hoelzl@56994: lemma measurable_count_space_eq2: hoelzl@56994: assumes "finite A" hoelzl@56994: shows "f \ measurable M (count_space A) \ (f \ space M \ A \ (\a\A. f -` {a} \ space M \ sets M))" hoelzl@56994: proof - hoelzl@56994: { fix X assume "X \ A" "f \ space M \ A" hoelzl@56994: with `finite A` have "f -` X \ space M = (\a\X. f -` {a} \ space M)" "finite X" hoelzl@56994: by (auto dest: finite_subset) hoelzl@56994: moreover assume "\a\A. f -` {a} \ space M \ sets M" hoelzl@56994: ultimately have "f -` X \ space M \ sets M" hoelzl@56994: using `X \ A` by (auto intro!: sets.finite_UN simp del: UN_simps) } hoelzl@56994: then show ?thesis hoelzl@56994: unfolding measurable_def by auto hoelzl@56994: qed hoelzl@56994: hoelzl@56994: lemma measurable_count_space_eq2_countable: hoelzl@56994: fixes f :: "'a => 'c::countable" hoelzl@56994: shows "f \ measurable M (count_space A) \ (f \ space M \ A \ (\a\A. f -` {a} \ space M \ sets M))" hoelzl@56994: proof - hoelzl@56994: { fix X assume "X \ A" "f \ space M \ A" hoelzl@56994: assume *: "\a. a\A \ f -` {a} \ space M \ sets M" hoelzl@56994: have "f -` X \ space M = (\a\X. f -` {a} \ space M)" hoelzl@56994: by auto hoelzl@56994: also have "\ \ sets M" hoelzl@56994: using * `X \ A` by (intro sets.countable_UN) auto hoelzl@56994: finally have "f -` X \ space M \ sets M" . } hoelzl@56994: then show ?thesis hoelzl@56994: unfolding measurable_def by auto hoelzl@56994: qed hoelzl@56994: hoelzl@56994: lemma measurable_compose_countable: hoelzl@56994: assumes f: "\i::'i::countable. (\x. f i x) \ measurable M N" and g: "g \ measurable M (count_space UNIV)" hoelzl@56994: shows "(\x. f (g x) x) \ measurable M N" hoelzl@56994: unfolding measurable_def hoelzl@56994: proof safe hoelzl@56994: fix x assume "x \ space M" then show "f (g x) x \ space N" hoelzl@56994: using f[THEN measurable_space] g[THEN measurable_space] by auto hoelzl@56994: next hoelzl@56994: fix A assume A: "A \ sets N" hoelzl@56994: have "(\x. f (g x) x) -` A \ space M = (\i. (g -` {i} \ space M) \ (f i -` A \ space M))" hoelzl@56994: by auto hoelzl@56994: also have "\ \ sets M" using f[THEN measurable_sets, OF A] g[THEN measurable_sets] hoelzl@56994: by (auto intro!: sets.countable_UN measurable_sets) hoelzl@56994: finally show "(\x. f (g x) x) -` A \ space M \ sets M" . hoelzl@56994: qed hoelzl@56994: hoelzl@56994: lemma measurable_count_space_const: hoelzl@56994: "(\x. c) \ measurable M (count_space UNIV)" hoelzl@56994: by (simp add: measurable_const) hoelzl@56994: hoelzl@56994: lemma measurable_count_space: hoelzl@56994: "f \ measurable (count_space A) (count_space UNIV)" hoelzl@56994: by simp hoelzl@56994: hoelzl@56994: lemma measurable_compose_rev: hoelzl@56994: assumes f: "f \ measurable L N" and g: "g \ measurable M L" hoelzl@56994: shows "(\x. f (g x)) \ measurable M N" hoelzl@56994: using measurable_compose[OF g f] . hoelzl@56994: hoelzl@56994: lemma measurable_count_space_eq_countable: hoelzl@56994: assumes "countable A" hoelzl@56994: shows "f \ measurable M (count_space A) \ (f \ space M \ A \ (\a\A. f -` {a} \ space M \ sets M))" hoelzl@56994: proof - hoelzl@56994: { fix X assume "X \ A" "f \ space M \ A" hoelzl@56994: with `countable A` have "f -` X \ space M = (\a\X. f -` {a} \ space M)" "countable X" hoelzl@56994: by (auto dest: countable_subset) hoelzl@56994: moreover assume "\a\A. f -` {a} \ space M \ sets M" hoelzl@56994: ultimately have "f -` X \ space M \ sets M" hoelzl@56994: using `X \ A` by (auto intro!: sets.countable_UN' simp del: UN_simps) } hoelzl@56994: then show ?thesis hoelzl@56994: unfolding measurable_def by auto hoelzl@56994: qed hoelzl@56994: hoelzl@56994: subsubsection {* Extend measure *} hoelzl@56994: hoelzl@56994: definition "extend_measure \ I G \ = hoelzl@56994: (if (\\'. (\i\I. \' (G i) = \ i) \ measure_space \ (sigma_sets \ (G`I)) \') \ \ (\i\I. \ i = 0) hoelzl@56994: then measure_of \ (G`I) (SOME \'. (\i\I. \' (G i) = \ i) \ measure_space \ (sigma_sets \ (G`I)) \') hoelzl@56994: else measure_of \ (G`I) (\_. 0))" hoelzl@56994: hoelzl@56994: lemma space_extend_measure: "G ` I \ Pow \ \ space (extend_measure \ I G \) = \" hoelzl@56994: unfolding extend_measure_def by simp hoelzl@56994: hoelzl@56994: lemma sets_extend_measure: "G ` I \ Pow \ \ sets (extend_measure \ I G \) = sigma_sets \ (G`I)" hoelzl@56994: unfolding extend_measure_def by simp hoelzl@56994: hoelzl@56994: lemma emeasure_extend_measure: hoelzl@56994: assumes M: "M = extend_measure \ I G \" hoelzl@56994: and eq: "\i. i \ I \ \' (G i) = \ i" hoelzl@56994: and ms: "G ` I \ Pow \" "positive (sets M) \'" "countably_additive (sets M) \'" hoelzl@56994: and "i \ I" hoelzl@56994: shows "emeasure M (G i) = \ i" hoelzl@56994: proof cases hoelzl@56994: assume *: "(\i\I. \ i = 0)" hoelzl@56994: with M have M_eq: "M = measure_of \ (G`I) (\_. 0)" hoelzl@56994: by (simp add: extend_measure_def) hoelzl@56994: from measure_space_0[OF ms(1)] ms `i\I` hoelzl@56994: have "emeasure M (G i) = 0" hoelzl@56994: by (intro emeasure_measure_of[OF M_eq]) (auto simp add: M measure_space_def sets_extend_measure) hoelzl@56994: with `i\I` * show ?thesis hoelzl@56994: by simp hoelzl@56994: next hoelzl@56994: def P \ "\\'. (\i\I. \' (G i) = \ i) \ measure_space \ (sigma_sets \ (G`I)) \'" hoelzl@56994: assume "\ (\i\I. \ i = 0)" hoelzl@56994: moreover hoelzl@56994: have "measure_space (space M) (sets M) \'" hoelzl@56994: using ms unfolding measure_space_def by auto default hoelzl@56994: with ms eq have "\\'. P \'" hoelzl@56994: unfolding P_def hoelzl@56994: by (intro exI[of _ \']) (auto simp add: M space_extend_measure sets_extend_measure) hoelzl@56994: ultimately have M_eq: "M = measure_of \ (G`I) (Eps P)" hoelzl@56994: by (simp add: M extend_measure_def P_def[symmetric]) hoelzl@56994: hoelzl@56994: from `\\'. P \'` have P: "P (Eps P)" by (rule someI_ex) hoelzl@56994: show "emeasure M (G i) = \ i" hoelzl@56994: proof (subst emeasure_measure_of[OF M_eq]) hoelzl@56994: have sets_M: "sets M = sigma_sets \ (G`I)" hoelzl@56994: using M_eq ms by (auto simp: sets_extend_measure) hoelzl@56994: then show "G i \ sets M" using `i \ I` by auto hoelzl@56994: show "positive (sets M) (Eps P)" "countably_additive (sets M) (Eps P)" "Eps P (G i) = \ i" hoelzl@56994: using P `i\I` by (auto simp add: sets_M measure_space_def P_def) hoelzl@56994: qed fact hoelzl@56994: qed hoelzl@56994: hoelzl@56994: lemma emeasure_extend_measure_Pair: hoelzl@56994: assumes M: "M = extend_measure \ {(i, j). I i j} (\(i, j). G i j) (\(i, j). \ i j)" hoelzl@56994: and eq: "\i j. I i j \ \' (G i j) = \ i j" hoelzl@56994: and ms: "\i j. I i j \ G i j \ Pow \" "positive (sets M) \'" "countably_additive (sets M) \'" hoelzl@56994: and "I i j" hoelzl@56994: shows "emeasure M (G i j) = \ i j" hoelzl@56994: using emeasure_extend_measure[OF M _ _ ms(2,3), of "(i,j)"] eq ms(1) `I i j` hoelzl@56994: by (auto simp: subset_eq) hoelzl@56994: hoelzl@56994: subsubsection {* Sigma algebra generated by function preimages *} hoelzl@56994: hoelzl@56994: definition hoelzl@56994: "vimage_algebra M S X = sigma S ((\A. X -` A \ S) ` sets M)" hoelzl@56994: hoelzl@56994: lemma sigma_algebra_preimages: hoelzl@56994: fixes f :: "'x \ 'a" hoelzl@56994: assumes "f \ S \ space M" hoelzl@56994: shows "sigma_algebra S ((\A. f -` A \ S) ` sets M)" hoelzl@56994: (is "sigma_algebra _ (?F ` sets M)") hoelzl@56994: proof (simp add: sigma_algebra_iff2, safe) hoelzl@56994: show "{} \ ?F ` sets M" by blast hoelzl@56994: next hoelzl@56994: fix A assume "A \ sets M" hoelzl@56994: moreover have "S - ?F A = ?F (space M - A)" hoelzl@56994: using assms by auto hoelzl@56994: ultimately show "S - ?F A \ ?F ` sets M" hoelzl@56994: by blast hoelzl@56994: next hoelzl@56994: fix A :: "nat \ 'x set" assume *: "range A \ ?F ` M" hoelzl@56994: have "\i. \b. b \ M \ A i = ?F b" hoelzl@56994: proof safe hoelzl@56994: fix i hoelzl@56994: have "A i \ ?F ` M" using * by auto hoelzl@56994: then show "\b. b \ M \ A i = ?F b" by auto hoelzl@56994: qed hoelzl@56994: from choice[OF this] obtain b where b: "range b \ M" "\i. A i = ?F (b i)" hoelzl@56994: by auto hoelzl@56994: then have "(\i. A i) = ?F (\i. b i)" by auto hoelzl@56994: then show "(\i. A i) \ ?F ` M" using b(1) by blast hoelzl@56994: qed hoelzl@56994: hoelzl@56994: lemma sets_vimage_algebra[simp]: hoelzl@56994: "f \ S \ space M \ sets (vimage_algebra M S f) = (\A. f -` A \ S) ` sets M" hoelzl@56994: using sigma_algebra.sets_measure_of_eq[OF sigma_algebra_preimages, of f S M] hoelzl@56994: by (simp add: vimage_algebra_def) hoelzl@56994: hoelzl@56994: lemma space_vimage_algebra[simp]: hoelzl@56994: "f \ S \ space M \ space (vimage_algebra M S f) = S" hoelzl@56994: using sigma_algebra.space_measure_of_eq[OF sigma_algebra_preimages, of f S M] hoelzl@56994: by (simp add: vimage_algebra_def) hoelzl@56994: hoelzl@56994: lemma in_vimage_algebra[simp]: hoelzl@56994: "f \ S \ space M \ A \ sets (vimage_algebra M S f) \ (\B\sets M. A = f -` B \ S)" hoelzl@56994: by (simp add: image_iff) hoelzl@56994: hoelzl@56994: lemma measurable_vimage_algebra: hoelzl@56994: fixes S :: "'c set" assumes "f \ S \ space M" hoelzl@56994: shows "f \ measurable (vimage_algebra M S f) M" hoelzl@56994: unfolding measurable_def using assms by force hoelzl@56994: hoelzl@56994: lemma measurable_vimage: hoelzl@56994: fixes g :: "'a \ 'c" and f :: "'d \ 'a" hoelzl@56994: assumes "g \ measurable M M2" "f \ S \ space M" hoelzl@56994: shows "(\x. g (f x)) \ measurable (vimage_algebra M S f) M2" hoelzl@56994: proof - hoelzl@56994: note measurable_vimage_algebra[OF assms(2)] hoelzl@56994: from measurable_comp[OF this assms(1)] hoelzl@56994: show ?thesis by (simp add: comp_def) hoelzl@56994: qed hoelzl@56994: hoelzl@56994: lemma sigma_sets_vimage: hoelzl@56994: assumes "f \ S' \ S" and "A \ Pow S" hoelzl@56994: shows "sigma_sets S' ((\X. f -` X \ S') ` A) = (\X. f -` X \ S') ` sigma_sets S A" hoelzl@56994: proof (intro set_eqI iffI) hoelzl@56994: let ?F = "\X. f -` X \ S'" hoelzl@56994: fix X assume "X \ sigma_sets S' (?F ` A)" hoelzl@56994: then show "X \ ?F ` sigma_sets S A" hoelzl@56994: proof induct hoelzl@56994: case (Basic X) then obtain X' where "X = ?F X'" "X' \ A" hoelzl@56994: by auto hoelzl@56994: then show ?case by auto hoelzl@56994: next hoelzl@56994: case Empty then show ?case hoelzl@56994: by (auto intro!: image_eqI[of _ _ "{}"] sigma_sets.Empty) hoelzl@56994: next hoelzl@56994: case (Compl X) then obtain X' where X: "X = ?F X'" and "X' \ sigma_sets S A" hoelzl@56994: by auto hoelzl@56994: then have "S - X' \ sigma_sets S A" hoelzl@56994: by (auto intro!: sigma_sets.Compl) hoelzl@56994: then show ?case hoelzl@56994: using X assms by (auto intro!: image_eqI[where x="S - X'"]) hoelzl@56994: next hoelzl@56994: case (Union F) hoelzl@56994: then have "\i. \F'. F' \ sigma_sets S A \ F i = f -` F' \ S'" hoelzl@56994: by (auto simp: image_iff Bex_def) hoelzl@56994: from choice[OF this] obtain F' where hoelzl@56994: "\i. F' i \ sigma_sets S A" and "\i. F i = f -` F' i \ S'" hoelzl@56994: by auto hoelzl@56994: then show ?case hoelzl@56994: by (auto intro!: sigma_sets.Union image_eqI[where x="\i. F' i"]) hoelzl@56994: qed hoelzl@56994: next hoelzl@56994: let ?F = "\X. f -` X \ S'" hoelzl@56994: fix X assume "X \ ?F ` sigma_sets S A" hoelzl@56994: then obtain X' where "X' \ sigma_sets S A" "X = ?F X'" by auto hoelzl@56994: then show "X \ sigma_sets S' (?F ` A)" hoelzl@56994: proof (induct arbitrary: X) hoelzl@56994: case Empty then show ?case by (auto intro: sigma_sets.Empty) hoelzl@56994: next hoelzl@56994: case (Compl X') hoelzl@56994: have "S' - (S' - X) \ sigma_sets S' (?F ` A)" hoelzl@56994: apply (rule sigma_sets.Compl) hoelzl@56994: using assms by (auto intro!: Compl.hyps simp: Compl.prems) hoelzl@56994: also have "S' - (S' - X) = X" hoelzl@56994: using assms Compl by auto hoelzl@56994: finally show ?case . hoelzl@56994: next hoelzl@56994: case (Union F) hoelzl@56994: have "(\i. f -` F i \ S') \ sigma_sets S' (?F ` A)" hoelzl@56994: by (intro sigma_sets.Union Union.hyps) simp hoelzl@56994: also have "(\i. f -` F i \ S') = X" hoelzl@56994: using assms Union by auto hoelzl@56994: finally show ?case . hoelzl@56994: qed auto hoelzl@56994: qed hoelzl@56994: hoelzl@56994: subsubsection {* Restricted Space Sigma Algebra *} hoelzl@56994: hoelzl@57025: definition restrict_space where hoelzl@57025: "restrict_space M \ = measure_of (\ \ space M) ((op \ \) ` sets M) (emeasure M)" hoelzl@56994: hoelzl@57025: lemma space_restrict_space: "space (restrict_space M \) = \ \ space M" hoelzl@57025: using sets.sets_into_space unfolding restrict_space_def by (subst space_measure_of) auto hoelzl@57025: hoelzl@57025: lemma space_restrict_space2: "\ \ sets M \ space (restrict_space M \) = \" hoelzl@57025: by (simp add: space_restrict_space sets.sets_into_space) hoelzl@56994: hoelzl@57025: lemma sets_restrict_space: "sets (restrict_space M \) = (op \ \) ` sets M" hoelzl@57025: proof - hoelzl@57025: have "sigma_sets (\ \ space M) ((\X. X \ (\ \ space M)) ` sets M) = hoelzl@57025: (\X. X \ (\ \ space M)) ` sets M" hoelzl@57025: using sigma_sets_vimage[of "\x. x" "\ \ space M" "space M" "sets M"] sets.space_closed[of M] hoelzl@57025: by (simp add: sets.sigma_sets_eq) hoelzl@57025: moreover have "(\X. X \ (\ \ space M)) ` sets M = (op \ \) ` sets M" hoelzl@57025: using sets.sets_into_space by (intro image_cong) auto hoelzl@57025: ultimately show ?thesis hoelzl@57025: using sets.sets_into_space[of _ M] unfolding restrict_space_def hoelzl@57025: by (subst sets_measure_of) fastforce+ hoelzl@57025: qed hoelzl@56994: hoelzl@56994: lemma sets_restrict_space_iff: hoelzl@57025: "\ \ space M \ sets M \ A \ sets (restrict_space M \) \ (A \ \ \ A \ sets M)" hoelzl@57025: proof (subst sets_restrict_space, safe) hoelzl@57025: fix A assume "\ \ space M \ sets M" and A: "A \ sets M" hoelzl@57025: then have "(\ \ space M) \ A \ sets M" hoelzl@57025: by rule hoelzl@57025: also have "(\ \ space M) \ A = \ \ A" hoelzl@57025: using sets.sets_into_space[OF A] by auto hoelzl@57025: finally show "\ \ A \ sets M" hoelzl@57025: by auto hoelzl@57025: qed auto hoelzl@56994: hoelzl@56994: lemma measurable_restrict_space1: hoelzl@57025: assumes \: "\ \ space M \ sets M" and f: "f \ measurable M N" hoelzl@57025: shows "f \ measurable (restrict_space M \) N" hoelzl@56994: unfolding measurable_def hoelzl@56994: proof (intro CollectI conjI ballI) hoelzl@56994: show sp: "f \ space (restrict_space M \) \ space N" hoelzl@56994: using measurable_space[OF f] sets.sets_into_space[OF \] by (auto simp: space_restrict_space) hoelzl@56994: hoelzl@56994: fix A assume "A \ sets N" hoelzl@57025: have "f -` A \ space (restrict_space M \) = (f -` A \ space M) \ (\ \ space M)" hoelzl@56994: using sets.sets_into_space[OF \] by (auto simp: space_restrict_space) hoelzl@56994: also have "\ \ sets (restrict_space M \)" hoelzl@56994: unfolding sets_restrict_space_iff[OF \] hoelzl@56994: using measurable_sets[OF f `A \ sets N`] \ by blast hoelzl@56994: finally show "f -` A \ space (restrict_space M \) \ sets (restrict_space M \)" . hoelzl@56994: qed hoelzl@56994: hoelzl@56994: lemma measurable_restrict_space2: hoelzl@57137: "\ \ space N \ sets N \ f \ space M \ \ \ f \ measurable M N \ hoelzl@57025: f \ measurable M (restrict_space N \)" hoelzl@57137: by (simp add: measurable_def space_restrict_space sets_restrict_space_iff Pi_Int[symmetric]) hoelzl@56994: hoelzl@57138: lemma measurable_restrict_space_iff: hoelzl@57138: assumes \[simp, intro]: "\ \ space M \ sets M" "c \ space N" hoelzl@57138: shows "f \ measurable (restrict_space M \) N \ hoelzl@57138: (\x. if x \ \ then f x else c) \ measurable M N" (is "f \ measurable ?R N \ ?f \ measurable M N") hoelzl@57138: unfolding measurable_def hoelzl@57138: proof safe hoelzl@57138: fix x assume "f \ space ?R \ space N" "x \ space M" then show "?f x \ space N" hoelzl@57138: using `c\space N` by (auto simp: space_restrict_space) hoelzl@57138: next hoelzl@57138: fix x assume "?f \ space M \ space N" "x \ space ?R" then show "f x \ space N" hoelzl@57138: using `c\space N` by (auto simp: space_restrict_space Pi_iff) hoelzl@57138: next hoelzl@57138: fix X assume X: "X \ sets N" hoelzl@57138: assume *[THEN bspec]: "\y\sets N. f -` y \ space ?R \ sets ?R" hoelzl@57138: have "?f -` X \ space M = (f -` X \ (\ \ space M)) \ (if c \ X then (space M - (\ \ space M)) else {})" hoelzl@57138: by (auto split: split_if_asm) hoelzl@57138: also have "\ \ sets M" hoelzl@57138: using *[OF X] by (auto simp add: space_restrict_space sets_restrict_space_iff) hoelzl@57138: finally show "?f -` X \ space M \ sets M" . hoelzl@57138: next hoelzl@57138: assume *[THEN bspec]: "\y\sets N. ?f -` y \ space M \ sets M" hoelzl@57138: fix X :: "'b set" assume X: "X \ sets N" hoelzl@57138: have "f -` X \ (\ \ space M) = (?f -` X \ space M) \ (\ \ space M)" hoelzl@57138: by (auto simp: space_restrict_space) hoelzl@57138: also have "\ \ sets M" hoelzl@57138: using *[OF X] by auto hoelzl@57138: finally show "f -` X \ space ?R \ sets ?R" hoelzl@57138: by (auto simp add: sets_restrict_space_iff space_restrict_space) hoelzl@57138: qed hoelzl@57138: paulson@33271: end hoelzl@57025: