diff -r 5001ed24e129 -r d5d342611edb src/HOL/Probability/Sigma_Algebra.thy --- a/src/HOL/Probability/Sigma_Algebra.thy Mon Aug 23 17:46:13 2010 +0200 +++ b/src/HOL/Probability/Sigma_Algebra.thy Mon Aug 23 19:35:57 2010 +0200 @@ -1,12 +1,12 @@ (* Title: Sigma_Algebra.thy Author: Stefan Richter, Markus Wenzel, TU Muenchen - Plus material from the Hurd/Coble measure theory development, + Plus material from the Hurd/Coble measure theory development, translated by Lawrence Paulson. *) header {* Sigma Algebras *} -theory Sigma_Algebra imports Complex_Main begin +theory Sigma_Algebra imports Main Countable FuncSet begin text {* Sigma algebras are an elementary concept in measure theory. To measure --- that is to integrate --- functions, we first have @@ -18,8 +18,8 @@ subsection {* Algebras *} -record 'a algebra = - space :: "'a set" +record 'a algebra = + space :: "'a set" sets :: "'a set set" locale algebra = @@ -35,20 +35,20 @@ lemma (in algebra) sets_into_space: "x \ sets M \ x \ space M" by (metis PowD contra_subsetD space_closed) -lemma (in algebra) Int [intro]: +lemma (in algebra) Int [intro]: assumes a: "a \ sets M" and b: "b \ sets M" shows "a \ b \ sets M" proof - - have "((space M - a) \ (space M - b)) \ sets M" + have "((space M - a) \ (space M - b)) \ sets M" by (metis a b compl_sets Un) moreover - have "a \ b = space M - ((space M - a) \ (space M - b))" + have "a \ b = space M - ((space M - a) \ (space M - b))" using space_closed a b by blast ultimately show ?thesis by (metis compl_sets) qed -lemma (in algebra) Diff [intro]: +lemma (in algebra) Diff [intro]: assumes a: "a \ sets M" and b: "b \ sets M" shows "a - b \ sets M" proof - have "(a \ (space M - b)) \ sets M" @@ -60,74 +60,143 @@ by metis qed -lemma (in algebra) finite_union [intro]: +lemma (in algebra) finite_union [intro]: "finite X \ X \ sets M \ Union X \ sets M" - by (induct set: finite) (auto simp add: Un) + by (induct set: finite) (auto simp add: Un) +lemma algebra_iff_Int: + "algebra M \ + sets M \ Pow (space M) & {} \ sets M & + (\a \ sets M. space M - a \ sets M) & + (\a \ sets M. \ b \ sets M. a \ b \ sets M)" +proof (auto simp add: algebra.Int, auto simp add: algebra_def) + fix a b + assume ab: "sets M \ Pow (space M)" + "\a\sets M. space M - a \ sets M" + "\a\sets M. \b\sets M. a \ b \ sets M" + "a \ sets M" "b \ sets M" + hence "a \ b = space M - ((space M - a) \ (space M - b))" + by blast + also have "... \ sets M" + by (metis ab) + finally show "a \ b \ sets M" . +qed + +lemma (in algebra) insert_in_sets: + assumes "{x} \ sets M" "A \ sets M" shows "insert x A \ sets M" +proof - + have "{x} \ A \ sets M" using assms by (rule Un) + thus ?thesis by auto +qed + +lemma (in algebra) Int_space_eq1 [simp]: "x \ sets M \ space M \ x = x" + by (metis Int_absorb1 sets_into_space) + +lemma (in algebra) Int_space_eq2 [simp]: "x \ sets M \ x \ space M = x" + by (metis Int_absorb2 sets_into_space) + +lemma (in algebra) restricted_algebra: + assumes "S \ sets M" + shows "algebra (M\ space := S, sets := (\A. S \ A) ` sets M \)" + (is "algebra ?r") + using assms by unfold_locales auto subsection {* Sigma Algebras *} locale sigma_algebra = algebra + - assumes countable_UN [intro]: + assumes countable_nat_UN [intro]: "!!A. range A \ sets M \ (\i::nat. A i) \ sets M" +lemma countable_UN_eq: + fixes A :: "'i::countable \ 'a set" + shows "(range A \ sets M \ (\i. A i) \ sets M) \ + (range (A \ from_nat) \ sets M \ (\i. (A \ from_nat) i) \ sets M)" +proof - + let ?A' = "A \ from_nat" + have *: "(\i. ?A' i) = (\i. A i)" (is "?l = ?r") + proof safe + fix x i assume "x \ A i" thus "x \ ?l" + by (auto intro!: exI[of _ "to_nat i"]) + next + fix x i assume "x \ ?A' i" thus "x \ ?r" + by (auto intro!: exI[of _ "from_nat i"]) + qed + have **: "range ?A' = range A" + using surj_range[OF surj_from_nat] + by (auto simp: image_compose intro!: imageI) + show ?thesis unfolding * ** .. +qed + +lemma (in sigma_algebra) countable_UN[intro]: + fixes A :: "'i::countable \ 'a set" + assumes "A`X \ sets M" + shows "(\x\X. A x) \ sets M" +proof - + let "?A i" = "if i \ X then A i else {}" + from assms have "range ?A \ sets M" by auto + with countable_nat_UN[of "?A \ from_nat"] countable_UN_eq[of ?A M] + have "(\x. ?A x) \ sets M" by auto + moreover have "(\x. ?A x) = (\x\X. A x)" by (auto split: split_if_asm) + ultimately show ?thesis by simp +qed + +lemma (in sigma_algebra) finite_UN: + assumes "finite I" and "\i. i \ I \ A i \ sets M" + shows "(\i\I. A i) \ sets M" + using assms by induct auto + lemma (in sigma_algebra) countable_INT [intro]: - assumes a: "range a \ sets M" - shows "(\i::nat. a i) \ sets M" + fixes A :: "'i::countable \ 'a set" + assumes A: "A`X \ sets M" "X \ {}" + shows "(\i\X. A i) \ sets M" proof - - from a have "\i. a i \ sets M" by fast - hence "space M - (\i. space M - a i) \ sets M" by blast + from A have "\i\X. A i \ sets M" by fast + hence "space M - (\i\X. space M - A i) \ sets M" by blast moreover - have "(\i. a i) = space M - (\i. space M - a i)" using space_closed a + have "(\i\X. A i) = space M - (\i\X. space M - A i)" using space_closed A by blast ultimately show ?thesis by metis qed -lemma (in sigma_algebra) gen_countable_UN: - fixes f :: "nat \ 'c" - shows "I = range f \ range A \ sets M \ (\x\I. A x) \ sets M" -by auto - -lemma (in sigma_algebra) gen_countable_INT: - fixes f :: "nat \ 'c" - shows "I = range f \ range A \ sets M \ (\x\I. A x) \ sets M" -by auto +lemma (in sigma_algebra) finite_INT: + assumes "finite I" "I \ {}" "\i. i \ I \ A i \ sets M" + shows "(\i\I. A i) \ sets M" + using assms by (induct rule: finite_ne_induct) auto lemma algebra_Pow: - "algebra (| space = sp, sets = Pow sp |)" - by (auto simp add: algebra_def) + "algebra \ space = sp, sets = Pow sp, \ = X \" + by (auto simp add: algebra_def) lemma sigma_algebra_Pow: - "sigma_algebra (| space = sp, sets = Pow sp |)" - by (auto simp add: sigma_algebra_def sigma_algebra_axioms_def algebra_Pow) + "sigma_algebra \ space = sp, sets = Pow sp, \ = X \" + by (auto simp add: sigma_algebra_def sigma_algebra_axioms_def algebra_Pow) + +lemma sigma_algebra_iff: + "sigma_algebra M \ + algebra M \ (\A. range A \ sets M \ (\i::nat. A i) \ sets M)" + by (simp add: sigma_algebra_def sigma_algebra_axioms_def) subsection {* Binary Unions *} definition binary :: "'a \ 'a \ nat \ 'a" where "binary a b = (\\<^isup>x. b)(0 := a)" -lemma range_binary_eq: "range(binary a b) = {a,b}" - by (auto simp add: binary_def) - -lemma Un_range_binary: "a \ b = (\i::nat. binary a b i)" - by (simp add: UNION_eq_Union_image range_binary_eq) +lemma range_binary_eq: "range(binary a b) = {a,b}" + by (auto simp add: binary_def) -lemma Int_range_binary: "a \ b = (\i::nat. binary a b i)" - by (simp add: INTER_eq_Inter_image range_binary_eq) +lemma Un_range_binary: "a \ b = (\i::nat. binary a b i)" + by (simp add: UNION_eq_Union_image range_binary_eq) -lemma sigma_algebra_iff: - "sigma_algebra M \ - algebra M & (\A. range A \ sets M \ (\i::nat. A i) \ sets M)" - by (simp add: sigma_algebra_def sigma_algebra_axioms_def) +lemma Int_range_binary: "a \ b = (\i::nat. binary a b i)" + by (simp add: INTER_eq_Inter_image range_binary_eq) lemma sigma_algebra_iff2: "sigma_algebra M \ - sets M \ Pow (space M) & - {} \ sets M & (\s \ sets M. space M - s \ sets M) & + sets M \ Pow (space M) \ + {} \ sets M \ (\s \ sets M. space M - s \ sets M) \ (\A. range A \ sets M \ (\i::nat. A i) \ sets M)" - by (force simp add: range_binary_eq sigma_algebra_def sigma_algebra_axioms_def - algebra_def Un_range_binary) - + by (auto simp add: range_binary_eq sigma_algebra_def sigma_algebra_axioms_def + algebra_def Un_range_binary) subsection {* Initial Sigma Algebra *} @@ -148,19 +217,21 @@ sigma where "sigma sp A = (| space = sp, sets = sigma_sets sp A |)" +lemma sets_sigma: "sets (sigma A B) = sigma_sets A B" + unfolding sigma_def by simp lemma space_sigma [simp]: "space (sigma X B) = X" - by (simp add: sigma_def) + by (simp add: sigma_def) lemma sigma_sets_top: "sp \ sigma_sets sp A" by (metis Diff_empty sigma_sets.Compl sigma_sets.Empty) lemma sigma_sets_into_sp: "A \ Pow sp \ x \ sigma_sets sp A \ x \ sp" - by (erule sigma_sets.induct, auto) + by (erule sigma_sets.induct, auto) -lemma sigma_sets_Un: +lemma sigma_sets_Un: "a \ sigma_sets sp A \ b \ sigma_sets sp A \ a \ b \ sigma_sets sp A" -apply (simp add: Un_range_binary range_binary_eq) +apply (simp add: Un_range_binary range_binary_eq) apply (rule Union, simp add: binary_def COMBK_def fun_upd_apply) done @@ -169,21 +240,21 @@ shows "(\i::nat. a i \ sigma_sets sp A) \ (\i. a i) \ sigma_sets sp A" proof - assume ai: "\i::nat. a i \ sigma_sets sp A" - hence "\i::nat. sp-(a i) \ sigma_sets sp A" + hence "\i::nat. sp-(a i) \ sigma_sets sp A" by (rule sigma_sets.Compl) - hence "(\i. sp-(a i)) \ sigma_sets sp A" + hence "(\i. sp-(a i)) \ sigma_sets sp A" by (rule sigma_sets.Union) - hence "sp-(\i. sp-(a i)) \ sigma_sets sp A" + hence "sp-(\i. sp-(a i)) \ sigma_sets sp A" by (rule sigma_sets.Compl) - also have "sp-(\i. sp-(a i)) = sp Int (\i. a i)" + also have "sp-(\i. sp-(a i)) = sp Int (\i. a i)" by auto also have "... = (\i. a i)" using ai - by (blast dest: sigma_sets_into_sp [OF Asb]) - finally show ?thesis . + by (blast dest: sigma_sets_into_sp [OF Asb]) + finally show ?thesis . qed lemma sigma_sets_INTER: - assumes Asb: "A \ Pow sp" + assumes Asb: "A \ Pow sp" and ai: "\i::nat. i \ S \ a i \ sigma_sets sp A" and non: "S \ {}" shows "(\i\S. a i) \ sigma_sets sp A" proof - @@ -197,13 +268,13 @@ qed lemma (in sigma_algebra) sigma_sets_subset: - assumes a: "a \ sets M" + assumes a: "a \ sets M" shows "sigma_sets (space M) a \ sets M" proof fix x assume "x \ sigma_sets (space M) a" from this show "x \ sets M" - by (induct rule: sigma_sets.induct, auto) (metis a subsetD) + by (induct rule: sigma_sets.induct, auto) (metis a subsetD) qed lemma (in sigma_algebra) sigma_sets_eq: @@ -219,19 +290,612 @@ lemma sigma_algebra_sigma_sets: "a \ Pow (space M) \ sets M = sigma_sets (space M) a \ sigma_algebra M" apply (auto simp add: sigma_algebra_def sigma_algebra_axioms_def - algebra_def sigma_sets.Empty sigma_sets.Compl sigma_sets_Un) + algebra_def sigma_sets.Empty sigma_sets.Compl sigma_sets_Un) apply (blast dest: sigma_sets_into_sp) apply (rule sigma_sets.Union, fast) done lemma sigma_algebra_sigma: "a \ Pow X \ sigma_algebra (sigma X a)" - apply (rule sigma_algebra_sigma_sets) - apply (auto simp add: sigma_def) + apply (rule sigma_algebra_sigma_sets) + apply (auto simp add: sigma_def) done lemma (in sigma_algebra) sigma_subset: "a \ sets M ==> sets (sigma (space M) a) \ (sets M)" by (simp add: sigma_def sigma_sets_subset) +lemma (in sigma_algebra) restriction_in_sets: + fixes A :: "nat \ 'a set" + assumes "S \ sets M" + and *: "range A \ (\A. S \ A) ` sets M" (is "_ \ ?r") + shows "range A \ sets M" "(\i. A i) \ (\A. S \ A) ` sets M" +proof - + { fix i have "A i \ ?r" using * by auto + hence "\B. A i = B \ S \ B \ sets M" by auto + hence "A i \ S" "A i \ sets M" using `S \ sets M` by auto } + thus "range A \ sets M" "(\i. A i) \ (\A. S \ A) ` sets M" + by (auto intro!: image_eqI[of _ _ "(\i. A i)"]) +qed + +lemma (in sigma_algebra) restricted_sigma_algebra: + assumes "S \ sets M" + shows "sigma_algebra (M\ space := S, sets := (\A. S \ A) ` sets M \)" + (is "sigma_algebra ?r") + unfolding sigma_algebra_def sigma_algebra_axioms_def +proof safe + show "algebra ?r" using restricted_algebra[OF assms] . +next + fix A :: "nat \ 'a set" assume "range A \ sets ?r" + from restriction_in_sets[OF assms this[simplified]] + show "(\i. A i) \ sets ?r" by simp +qed + +section {* Measurable functions *} + +definition + "measurable A B = {f \ space A -> space B. \y \ sets B. f -` y \ space A \ sets A}" + +lemma (in sigma_algebra) measurable_sigma: + assumes B: "B \ Pow X" + and f: "f \ space M -> X" + and ba: "\y. y \ B \ (f -` y) \ space M \ sets M" + shows "f \ measurable M (sigma X B)" +proof - + have "sigma_sets X B \ {y . (f -` y) \ space M \ sets M & y \ X}" + proof clarify + fix x + assume "x \ sigma_sets X B" + thus "f -` x \ space M \ sets M \ x \ X" + proof induct + case (Basic a) + thus ?case + by (auto simp add: ba) (metis B subsetD PowD) + next + case Empty + thus ?case + by auto + next + case (Compl a) + have [simp]: "f -` X \ space M = space M" + by (auto simp add: funcset_mem [OF f]) + thus ?case + by (auto simp add: vimage_Diff Diff_Int_distrib2 compl_sets Compl) + next + case (Union a) + thus ?case + by (simp add: vimage_UN, simp only: UN_extend_simps(4)) + (blast intro: countable_UN) + qed + qed + thus ?thesis + by (simp add: measurable_def sigma_algebra_axioms sigma_algebra_sigma B f) + (auto simp add: sigma_def) +qed + +lemma measurable_cong: + assumes "\ w. w \ space M \ f w = g w" + shows "f \ measurable M M' \ g \ measurable M M'" + unfolding measurable_def using assms + by (simp cong: vimage_inter_cong Pi_cong) + +lemma measurable_space: + "f \ measurable M A \ x \ space M \ f x \ space A" + unfolding measurable_def by auto + +lemma measurable_sets: + "f \ measurable M A \ S \ sets A \ f -` S \ space M \ sets M" + unfolding measurable_def by auto + +lemma (in sigma_algebra) measurable_subset: + "(\S. S \ sets A \ S \ space A) \ measurable M A \ measurable M (sigma (space A) (sets A))" + by (auto intro: measurable_sigma measurable_sets measurable_space) + +lemma measurable_eqI: + "\ space m1 = space m1' ; space m2 = space m2' ; + sets m1 = sets m1' ; sets m2 = sets m2' \ + \ measurable m1 m2 = measurable m1' m2'" + by (simp add: measurable_def sigma_algebra_iff2) + +lemma (in sigma_algebra) measurable_const[intro, simp]: + assumes "c \ space M'" + shows "(\x. c) \ measurable M M'" + using assms by (auto simp add: measurable_def) + +lemma (in sigma_algebra) measurable_If: + assumes measure: "f \ measurable M M'" "g \ measurable M M'" + assumes P: "{x\space M. P x} \ sets M" + shows "(\x. if P x then f x else g x) \ measurable M M'" + unfolding measurable_def +proof safe + fix x assume "x \ space M" + thus "(if P x then f x else g x) \ space M'" + using measure unfolding measurable_def by auto +next + fix A assume "A \ sets M'" + hence *: "(\x. if P x then f x else g x) -` A \ space M = + ((f -` A \ space M) \ {x\space M. P x}) \ + ((g -` A \ space M) \ (space M - {x\space M. P x}))" + using measure unfolding measurable_def by (auto split: split_if_asm) + show "(\x. if P x then f x else g x) -` A \ space M \ sets M" + using `A \ sets M'` measure P unfolding * measurable_def + by (auto intro!: Un) +qed + +lemma (in sigma_algebra) measurable_If_set: + assumes measure: "f \ measurable M M'" "g \ measurable M M'" + assumes P: "A \ sets M" + shows "(\x. if x \ A then f x else g x) \ measurable M M'" +proof (rule measurable_If[OF measure]) + have "{x \ space M. x \ A} = A" using `A \ sets M` sets_into_space by auto + thus "{x \ space M. x \ A} \ sets M" using `A \ sets M` by auto +qed + +lemma (in algebra) measurable_ident[intro, simp]: "id \ measurable M M" + by (auto simp add: measurable_def) + +lemma measurable_comp[intro]: + fixes f :: "'a \ 'b" and g :: "'b \ 'c" + shows "f \ measurable a b \ g \ measurable b c \ (g o f) \ measurable a c" + apply (auto simp add: measurable_def vimage_compose) + apply (subgoal_tac "f -` g -` y \ space a = f -` (g -` y \ space b) \ space a") + apply force+ + done + +lemma measurable_strong: + fixes f :: "'a \ 'b" and g :: "'b \ 'c" + assumes f: "f \ measurable a b" and g: "g \ (space b -> space c)" + and a: "sigma_algebra a" and b: "sigma_algebra b" and c: "sigma_algebra c" + and t: "f ` (space a) \ t" + and cb: "\s. s \ sets c \ (g -` s) \ t \ sets b" + shows "(g o f) \ measurable a c" +proof - + have fab: "f \ (space a -> space b)" + and ba: "\y. y \ sets b \ (f -` y) \ (space a) \ sets a" using f + by (auto simp add: measurable_def) + have eq: "f -` g -` y \ space a = f -` (g -` y \ t) \ space a" using t + by force + show ?thesis + apply (auto simp add: measurable_def vimage_compose a c) + apply (metis funcset_mem fab g) + apply (subst eq, metis ba cb) + done +qed + +lemma measurable_mono1: + "a \ b \ sigma_algebra \space = X, sets = b\ + \ measurable \space = X, sets = a\ c \ measurable \space = X, sets = b\ c" + by (auto simp add: measurable_def) + +lemma measurable_up_sigma: + "measurable A M \ measurable (sigma (space A) (sets A)) M" + unfolding measurable_def + by (auto simp: sigma_def intro: sigma_sets.Basic) + +lemma (in sigma_algebra) measurable_range_reduce: + "\ f \ measurable M \space = s, sets = Pow s\ ; s \ {} \ + \ f \ measurable M \space = s \ (f ` space M), sets = Pow (s \ (f ` space M))\" + by (simp add: measurable_def sigma_algebra_Pow del: Pow_Int_eq) blast + +lemma (in sigma_algebra) measurable_Pow_to_Pow: + "(sets M = Pow (space M)) \ f \ measurable M \space = UNIV, sets = Pow UNIV\" + by (auto simp add: measurable_def sigma_algebra_def sigma_algebra_axioms_def algebra_def) + +lemma (in sigma_algebra) measurable_Pow_to_Pow_image: + "sets M = Pow (space M) + \ f \ measurable M \space = f ` space M, sets = Pow (f ` space M)\" + by (simp add: measurable_def sigma_algebra_Pow) intro_locales + +lemma (in sigma_algebra) sigma_algebra_preimages: + fixes f :: "'x \ 'a" + assumes "f \ A \ space M" + shows "sigma_algebra \ space = A, sets = (\M. f -` M \ A) ` sets M \" + (is "sigma_algebra \ space = _, sets = ?F ` sets M \") +proof (simp add: sigma_algebra_iff2, safe) + show "{} \ ?F ` sets M" by blast +next + fix S assume "S \ sets M" + moreover have "A - ?F S = ?F (space M - S)" + using assms by auto + ultimately show "A - ?F S \ ?F ` sets M" + by blast +next + fix S :: "nat \ 'x set" assume *: "range S \ ?F ` sets M" + have "\i. \b. b \ sets M \ S i = ?F b" + proof safe + fix i + have "S i \ ?F ` sets M" using * by auto + then show "\b. b \ sets M \ S i = ?F b" by auto + qed + from choice[OF this] obtain b where b: "range b \ sets M" "\i. S i = ?F (b i)" + by auto + then have "(\i. S i) = ?F (\i. b i)" by auto + then show "(\i. S i) \ ?F ` sets M" using b(1) by blast +qed + +section "Disjoint families" + +definition + disjoint_family_on where + "disjoint_family_on A S \ (\m\S. \n\S. m \ n \ A m \ A n = {})" + +abbreviation + "disjoint_family A \ disjoint_family_on A UNIV" + +lemma range_subsetD: "range f \ B \ f i \ B" + by blast + +lemma Int_Diff_disjoint: "A \ B \ (A - B) = {}" + by blast + +lemma Int_Diff_Un: "A \ B \ (A - B) = A" + by blast + +lemma disjoint_family_subset: + "disjoint_family A \ (!!x. B x \ A x) \ disjoint_family B" + by (force simp add: disjoint_family_on_def) + +lemma disjoint_family_on_mono: + "A \ B \ disjoint_family_on f B \ disjoint_family_on f A" + unfolding disjoint_family_on_def by auto + +lemma disjoint_family_Suc: + assumes Suc: "!!n. A n \ A (Suc n)" + shows "disjoint_family (\i. A (Suc i) - A i)" +proof - + { + fix m + have "!!n. A n \ A (m+n)" + proof (induct m) + case 0 show ?case by simp + next + case (Suc m) thus ?case + by (metis Suc_eq_plus1 assms nat_add_commute nat_add_left_commute subset_trans) + qed + } + hence "!!m n. m < n \ A m \ A n" + by (metis add_commute le_add_diff_inverse nat_less_le) + thus ?thesis + by (auto simp add: disjoint_family_on_def) + (metis insert_absorb insert_subset le_SucE le_antisym not_leE) +qed + +definition disjointed :: "(nat \ 'a set) \ nat \ 'a set " + where "disjointed A n = A n - (\i\{0..i\{0..i\{0..i. disjointed A i) = (\i. A i)" + apply (rule UN_finite2_eq [where k=0]) + apply (simp add: finite_UN_disjointed_eq) + done + +lemma less_disjoint_disjointed: "m disjointed A m \ disjointed A n = {}" + by (auto simp add: disjointed_def) + +lemma disjoint_family_disjointed: "disjoint_family (disjointed A)" + by (simp add: disjoint_family_on_def) + (metis neq_iff Int_commute less_disjoint_disjointed) + +lemma disjointed_subset: "disjointed A n \ A n" + by (auto simp add: disjointed_def) + +lemma (in algebra) UNION_in_sets: + fixes A:: "nat \ 'a set" + assumes A: "range A \ sets M " + shows "(\i\{0.. sets M" +proof (induct n) + case 0 show ?case by simp +next + case (Suc n) + thus ?case + by (simp add: atLeastLessThanSuc) (metis A Un UNIV_I image_subset_iff) +qed + +lemma (in algebra) range_disjointed_sets: + assumes A: "range A \ sets M " + shows "range (disjointed A) \ sets M" +proof (auto simp add: disjointed_def) + fix n + show "A n - (\i\{0.. sets M" using UNION_in_sets + by (metis A Diff UNIV_I image_subset_iff) +qed + +lemma sigma_algebra_disjoint_iff: + "sigma_algebra M \ + algebra M & + (\A. range A \ sets M \ disjoint_family A \ + (\i::nat. A i) \ sets M)" +proof (auto simp add: sigma_algebra_iff) + fix A :: "nat \ 'a set" + assume M: "algebra M" + and A: "range A \ sets M" + and UnA: "\A. range A \ sets M \ + disjoint_family A \ (\i::nat. A i) \ sets M" + hence "range (disjointed A) \ sets M \ + disjoint_family (disjointed A) \ + (\i. disjointed A i) \ sets M" by blast + hence "(\i. disjointed A i) \ sets M" + by (simp add: algebra.range_disjointed_sets M A disjoint_family_disjointed) + thus "(\i::nat. A i) \ sets M" by (simp add: UN_disjointed_eq) +qed + +subsection {* A Two-Element Series *} + +definition binaryset :: "'a set \ 'a set \ nat \ 'a set " + where "binaryset A B = (\\<^isup>x. {})(0 := A, Suc 0 := B)" + +lemma range_binaryset_eq: "range(binaryset A B) = {A,B,{}}" + apply (simp add: binaryset_def) + apply (rule set_ext) + apply (auto simp add: image_iff) + done + +lemma UN_binaryset_eq: "(\i. binaryset A B i) = A \ B" + by (simp add: UNION_eq_Union_image range_binaryset_eq) + +section {* Closed CDI *} + +definition + closed_cdi where + "closed_cdi M \ + sets M \ Pow (space M) & + (\s \ sets M. space M - s \ sets M) & + (\A. (range A \ sets M) & (A 0 = {}) & (\n. A n \ A (Suc n)) \ + (\i. A i) \ sets M) & + (\A. (range A \ sets M) & disjoint_family A \ (\i::nat. A i) \ sets M)" + + +inductive_set + smallest_ccdi_sets :: "('a, 'b) algebra_scheme \ 'a set set" + for M + where + Basic [intro]: + "a \ sets M \ a \ smallest_ccdi_sets M" + | Compl [intro]: + "a \ smallest_ccdi_sets M \ space M - a \ smallest_ccdi_sets M" + | Inc: + "range A \ Pow(smallest_ccdi_sets M) \ A 0 = {} \ (\n. A n \ A (Suc n)) + \ (\i. A i) \ smallest_ccdi_sets M" + | Disj: + "range A \ Pow(smallest_ccdi_sets M) \ disjoint_family A + \ (\i::nat. A i) \ smallest_ccdi_sets M" + monos Pow_mono + + +definition + smallest_closed_cdi where + "smallest_closed_cdi M = (|space = space M, sets = smallest_ccdi_sets M|)" + +lemma space_smallest_closed_cdi [simp]: + "space (smallest_closed_cdi M) = space M" + by (simp add: smallest_closed_cdi_def) + +lemma (in algebra) smallest_closed_cdi1: "sets M \ sets (smallest_closed_cdi M)" + by (auto simp add: smallest_closed_cdi_def) + +lemma (in algebra) smallest_ccdi_sets: + "smallest_ccdi_sets M \ Pow (space M)" + apply (rule subsetI) + apply (erule smallest_ccdi_sets.induct) + apply (auto intro: range_subsetD dest: sets_into_space) + done + +lemma (in algebra) smallest_closed_cdi2: "closed_cdi (smallest_closed_cdi M)" + apply (auto simp add: closed_cdi_def smallest_closed_cdi_def smallest_ccdi_sets) + apply (blast intro: smallest_ccdi_sets.Inc smallest_ccdi_sets.Disj) + + done + +lemma (in algebra) smallest_closed_cdi3: + "sets (smallest_closed_cdi M) \ Pow (space M)" + by (simp add: smallest_closed_cdi_def smallest_ccdi_sets) + +lemma closed_cdi_subset: "closed_cdi M \ sets M \ Pow (space M)" + by (simp add: closed_cdi_def) + +lemma closed_cdi_Compl: "closed_cdi M \ s \ sets M \ space M - s \ sets M" + by (simp add: closed_cdi_def) + +lemma closed_cdi_Inc: + "closed_cdi M \ range A \ sets M \ A 0 = {} \ (!!n. A n \ A (Suc n)) \ + (\i. A i) \ sets M" + by (simp add: closed_cdi_def) + +lemma closed_cdi_Disj: + "closed_cdi M \ range A \ sets M \ disjoint_family A \ (\i::nat. A i) \ sets M" + by (simp add: closed_cdi_def) + +lemma closed_cdi_Un: + assumes cdi: "closed_cdi M" and empty: "{} \ sets M" + and A: "A \ sets M" and B: "B \ sets M" + and disj: "A \ B = {}" + shows "A \ B \ sets M" +proof - + have ra: "range (binaryset A B) \ sets M" + by (simp add: range_binaryset_eq empty A B) + have di: "disjoint_family (binaryset A B)" using disj + by (simp add: disjoint_family_on_def binaryset_def Int_commute) + from closed_cdi_Disj [OF cdi ra di] + show ?thesis + by (simp add: UN_binaryset_eq) +qed + +lemma (in algebra) smallest_ccdi_sets_Un: + assumes A: "A \ smallest_ccdi_sets M" and B: "B \ smallest_ccdi_sets M" + and disj: "A \ B = {}" + shows "A \ B \ smallest_ccdi_sets M" +proof - + have ra: "range (binaryset A B) \ Pow (smallest_ccdi_sets M)" + by (simp add: range_binaryset_eq A B smallest_ccdi_sets.Basic) + have di: "disjoint_family (binaryset A B)" using disj + by (simp add: disjoint_family_on_def binaryset_def Int_commute) + from Disj [OF ra di] + show ?thesis + by (simp add: UN_binaryset_eq) +qed + +lemma (in algebra) smallest_ccdi_sets_Int1: + assumes a: "a \ sets M" + shows "b \ smallest_ccdi_sets M \ a \ b \ smallest_ccdi_sets M" +proof (induct rule: smallest_ccdi_sets.induct) + case (Basic x) + thus ?case + by (metis a Int smallest_ccdi_sets.Basic) +next + case (Compl x) + have "a \ (space M - x) = space M - ((space M - a) \ (a \ x))" + by blast + also have "... \ smallest_ccdi_sets M" + by (metis smallest_ccdi_sets.Compl a Compl(2) Diff_Int2 Diff_Int_distrib2 + Diff_disjoint Int_Diff Int_empty_right Un_commute + smallest_ccdi_sets.Basic smallest_ccdi_sets.Compl + smallest_ccdi_sets_Un) + finally show ?case . +next + case (Inc A) + have 1: "(\i. (\i. a \ A i) i) = a \ (\i. A i)" + by blast + have "range (\i. a \ A i) \ Pow(smallest_ccdi_sets M)" using Inc + by blast + moreover have "(\i. a \ A i) 0 = {}" + by (simp add: Inc) + moreover have "!!n. (\i. a \ A i) n \ (\i. a \ A i) (Suc n)" using Inc + by blast + ultimately have 2: "(\i. (\i. a \ A i) i) \ smallest_ccdi_sets M" + by (rule smallest_ccdi_sets.Inc) + show ?case + by (metis 1 2) +next + case (Disj A) + have 1: "(\i. (\i. a \ A i) i) = a \ (\i. A i)" + by blast + have "range (\i. a \ A i) \ Pow(smallest_ccdi_sets M)" using Disj + by blast + moreover have "disjoint_family (\i. a \ A i)" using Disj + by (auto simp add: disjoint_family_on_def) + ultimately have 2: "(\i. (\i. a \ A i) i) \ smallest_ccdi_sets M" + by (rule smallest_ccdi_sets.Disj) + show ?case + by (metis 1 2) +qed + + +lemma (in algebra) smallest_ccdi_sets_Int: + assumes b: "b \ smallest_ccdi_sets M" + shows "a \ smallest_ccdi_sets M \ a \ b \ smallest_ccdi_sets M" +proof (induct rule: smallest_ccdi_sets.induct) + case (Basic x) + thus ?case + by (metis b smallest_ccdi_sets_Int1) +next + case (Compl x) + have "(space M - x) \ b = space M - (x \ b \ (space M - b))" + by blast + also have "... \ smallest_ccdi_sets M" + by (metis Compl(2) Diff_disjoint Int_Diff Int_commute Int_empty_right b + smallest_ccdi_sets.Compl smallest_ccdi_sets_Un) + finally show ?case . +next + case (Inc A) + have 1: "(\i. (\i. A i \ b) i) = (\i. A i) \ b" + by blast + have "range (\i. A i \ b) \ Pow(smallest_ccdi_sets M)" using Inc + by blast + moreover have "(\i. A i \ b) 0 = {}" + by (simp add: Inc) + moreover have "!!n. (\i. A i \ b) n \ (\i. A i \ b) (Suc n)" using Inc + by blast + ultimately have 2: "(\i. (\i. A i \ b) i) \ smallest_ccdi_sets M" + by (rule smallest_ccdi_sets.Inc) + show ?case + by (metis 1 2) +next + case (Disj A) + have 1: "(\i. (\i. A i \ b) i) = (\i. A i) \ b" + by blast + have "range (\i. A i \ b) \ Pow(smallest_ccdi_sets M)" using Disj + by blast + moreover have "disjoint_family (\i. A i \ b)" using Disj + by (auto simp add: disjoint_family_on_def) + ultimately have 2: "(\i. (\i. A i \ b) i) \ smallest_ccdi_sets M" + by (rule smallest_ccdi_sets.Disj) + show ?case + by (metis 1 2) +qed + +lemma (in algebra) sets_smallest_closed_cdi_Int: + "a \ sets (smallest_closed_cdi M) \ b \ sets (smallest_closed_cdi M) + \ a \ b \ sets (smallest_closed_cdi M)" + by (simp add: smallest_ccdi_sets_Int smallest_closed_cdi_def) + +lemma (in algebra) sigma_property_disjoint_lemma: + assumes sbC: "sets M \ C" + and ccdi: "closed_cdi (|space = space M, sets = C|)" + shows "sigma_sets (space M) (sets M) \ C" +proof - + have "smallest_ccdi_sets M \ {B . sets M \ B \ sigma_algebra (|space = space M, sets = B|)}" + apply (auto simp add: sigma_algebra_disjoint_iff algebra_iff_Int + smallest_ccdi_sets_Int) + apply (metis Union_Pow_eq Union_upper subsetD smallest_ccdi_sets) + apply (blast intro: smallest_ccdi_sets.Disj) + done + hence "sigma_sets (space M) (sets M) \ smallest_ccdi_sets M" + by clarsimp + (drule sigma_algebra.sigma_sets_subset [where a="sets M"], auto) + also have "... \ C" + proof + fix x + assume x: "x \ smallest_ccdi_sets M" + thus "x \ C" + proof (induct rule: smallest_ccdi_sets.induct) + case (Basic x) + thus ?case + by (metis Basic subsetD sbC) + next + case (Compl x) + thus ?case + by (blast intro: closed_cdi_Compl [OF ccdi, simplified]) + next + case (Inc A) + thus ?case + by (auto intro: closed_cdi_Inc [OF ccdi, simplified]) + next + case (Disj A) + thus ?case + by (auto intro: closed_cdi_Disj [OF ccdi, simplified]) + qed + qed + finally show ?thesis . +qed + +lemma (in algebra) sigma_property_disjoint: + assumes sbC: "sets M \ C" + and compl: "!!s. s \ C \ sigma_sets (space M) (sets M) \ space M - s \ C" + and inc: "!!A. range A \ C \ sigma_sets (space M) (sets M) + \ A 0 = {} \ (!!n. A n \ A (Suc n)) + \ (\i. A i) \ C" + and disj: "!!A. range A \ C \ sigma_sets (space M) (sets M) + \ disjoint_family A \ (\i::nat. A i) \ C" + shows "sigma_sets (space M) (sets M) \ C" +proof - + have "sigma_sets (space M) (sets M) \ C \ sigma_sets (space M) (sets M)" + proof (rule sigma_property_disjoint_lemma) + show "sets M \ C \ sigma_sets (space M) (sets M)" + by (metis Int_greatest Set.subsetI sbC sigma_sets.Basic) + next + show "closed_cdi \space = space M, sets = C \ sigma_sets (space M) (sets M)\" + by (simp add: closed_cdi_def compl inc disj) + (metis PowI Set.subsetI le_infI2 sigma_sets_into_sp space_closed + IntE sigma_sets.Compl range_subsetD sigma_sets.Union) + qed + thus ?thesis + by blast +qed + end