wenzelm@41983: (* Title: HOL/Probability/Complete_Measure.thy hoelzl@40859: Author: Robert Himmelmann, Johannes Hoelzl, TU Muenchen hoelzl@40859: *) wenzelm@41983: hoelzl@40859: theory Complete_Measure hoelzl@56993: imports Bochner_Integration hoelzl@40859: begin hoelzl@40859: hoelzl@47694: definition hoelzl@47694: "split_completion M A p = (if A \ sets M then p = (A, {}) else hoelzl@47694: \N'. A = fst p \ snd p \ fst p \ snd p = {} \ fst p \ sets M \ snd p \ N' \ N' \ null_sets M)" hoelzl@41689: hoelzl@47694: definition hoelzl@47694: "main_part M A = fst (Eps (split_completion M A))" hoelzl@41689: hoelzl@47694: definition hoelzl@47694: "null_part M A = snd (Eps (split_completion M A))" hoelzl@40859: hoelzl@47694: definition completion :: "'a measure \ 'a measure" where hoelzl@47694: "completion M = measure_of (space M) { S \ N |S N N'. S \ sets M \ N' \ null_sets M \ N \ N' } hoelzl@47694: (emeasure M \ main_part M)" hoelzl@40859: hoelzl@47694: lemma completion_into_space: hoelzl@47694: "{ S \ N |S N N'. S \ sets M \ N' \ null_sets M \ N \ N' } \ Pow (space M)" immler@50244: using sets.sets_into_space by auto hoelzl@40859: hoelzl@47694: lemma space_completion[simp]: "space (completion M) = space M" hoelzl@47694: unfolding completion_def using space_measure_of[OF completion_into_space] by simp hoelzl@40859: hoelzl@47694: lemma completionI: hoelzl@47694: assumes "A = S \ N" "N \ N'" "N' \ null_sets M" "S \ sets M" hoelzl@47694: shows "A \ { S \ N |S N N'. S \ sets M \ N' \ null_sets M \ N \ N' }" hoelzl@40859: using assms by auto hoelzl@40859: hoelzl@47694: lemma completionE: hoelzl@47694: assumes "A \ { S \ N |S N N'. S \ sets M \ N' \ null_sets M \ N \ N' }" hoelzl@47694: obtains S N N' where "A = S \ N" "N \ N'" "N' \ null_sets M" "S \ sets M" hoelzl@47694: using assms by auto hoelzl@47694: hoelzl@47694: lemma sigma_algebra_completion: hoelzl@47694: "sigma_algebra (space M) { S \ N |S N N'. S \ sets M \ N' \ null_sets M \ N \ N' }" hoelzl@47694: (is "sigma_algebra _ ?A") hoelzl@47694: unfolding sigma_algebra_iff2 hoelzl@47694: proof (intro conjI ballI allI impI) hoelzl@47694: show "?A \ Pow (space M)" immler@50244: using sets.sets_into_space by auto hoelzl@40859: next hoelzl@47694: show "{} \ ?A" by auto hoelzl@40859: next hoelzl@47694: let ?C = "space M" hoelzl@47694: fix A assume "A \ ?A" from completionE[OF this] guess S N N' . hoelzl@47694: then show "space M - A \ ?A" hoelzl@47694: by (intro completionI[of _ "(?C - S) \ (?C - N')" "(?C - S) \ N' \ (?C - N)"]) auto hoelzl@47694: next hoelzl@47694: fix A :: "nat \ 'a set" assume A: "range A \ ?A" hoelzl@47694: then have "\n. \S N N'. A n = S \ N \ S \ sets M \ N' \ null_sets M \ N \ N'" hoelzl@47694: by (auto simp: image_subset_iff) hoelzl@40859: from choice[OF this] guess S .. hoelzl@40859: from choice[OF this] guess N .. hoelzl@40859: from choice[OF this] guess N' .. hoelzl@47694: then show "UNION UNIV A \ ?A" hoelzl@40859: using null_sets_UN[of N'] hoelzl@47694: by (intro completionI[of _ "UNION UNIV S" "UNION UNIV N" "UNION UNIV N'"]) auto hoelzl@47694: qed hoelzl@47694: hoelzl@47694: lemma sets_completion: hoelzl@47694: "sets (completion M) = { S \ N |S N N'. S \ sets M \ N' \ null_sets M \ N \ N' }" hoelzl@47694: using sigma_algebra.sets_measure_of_eq[OF sigma_algebra_completion] by (simp add: completion_def) hoelzl@47694: hoelzl@47694: lemma sets_completionE: hoelzl@47694: assumes "A \ sets (completion M)" hoelzl@47694: obtains S N N' where "A = S \ N" "N \ N'" "N' \ null_sets M" "S \ sets M" hoelzl@47694: using assms unfolding sets_completion by auto hoelzl@47694: hoelzl@47694: lemma sets_completionI: hoelzl@47694: assumes "A = S \ N" "N \ N'" "N' \ null_sets M" "S \ sets M" hoelzl@47694: shows "A \ sets (completion M)" hoelzl@47694: using assms unfolding sets_completion by auto hoelzl@47694: hoelzl@47694: lemma sets_completionI_sets[intro, simp]: hoelzl@47694: "A \ sets M \ A \ sets (completion M)" hoelzl@47694: unfolding sets_completion by force hoelzl@40859: hoelzl@47694: lemma null_sets_completion: hoelzl@47694: assumes "N' \ null_sets M" "N \ N'" shows "N \ sets (completion M)" hoelzl@47694: using assms by (intro sets_completionI[of N "{}" N N']) auto hoelzl@47694: hoelzl@47694: lemma split_completion: hoelzl@47694: assumes "A \ sets (completion M)" hoelzl@47694: shows "split_completion M A (main_part M A, null_part M A)" hoelzl@47694: proof cases hoelzl@47694: assume "A \ sets M" then show ?thesis hoelzl@47694: by (simp add: split_completion_def[abs_def] main_part_def null_part_def) hoelzl@47694: next hoelzl@47694: assume nA: "A \ sets M" hoelzl@47694: show ?thesis hoelzl@47694: unfolding main_part_def null_part_def if_not_P[OF nA] hoelzl@47694: proof (rule someI2_ex) hoelzl@47694: from assms[THEN sets_completionE] guess S N N' . note A = this hoelzl@47694: let ?P = "(S, N - S)" hoelzl@47694: show "\p. split_completion M A p" hoelzl@47694: unfolding split_completion_def if_not_P[OF nA] using A hoelzl@47694: proof (intro exI conjI) hoelzl@47694: show "A = fst ?P \ snd ?P" using A by auto hoelzl@47694: show "snd ?P \ N'" using A by auto hoelzl@47694: qed auto hoelzl@40859: qed auto hoelzl@47694: qed hoelzl@40859: hoelzl@47694: lemma hoelzl@47694: assumes "S \ sets (completion M)" hoelzl@47694: shows main_part_sets[intro, simp]: "main_part M S \ sets M" hoelzl@47694: and main_part_null_part_Un[simp]: "main_part M S \ null_part M S = S" hoelzl@47694: and main_part_null_part_Int[simp]: "main_part M S \ null_part M S = {}" hoelzl@47694: using split_completion[OF assms] hoelzl@47694: by (auto simp: split_completion_def split: split_if_asm) hoelzl@40859: hoelzl@47694: lemma main_part[simp]: "S \ sets M \ main_part M S = S" hoelzl@47694: using split_completion[of S M] hoelzl@47694: by (auto simp: split_completion_def split: split_if_asm) hoelzl@40859: hoelzl@47694: lemma null_part: hoelzl@47694: assumes "S \ sets (completion M)" shows "\N. N\null_sets M \ null_part M S \ N" hoelzl@47694: using split_completion[OF assms] by (auto simp: split_completion_def split: split_if_asm) hoelzl@47694: hoelzl@47694: lemma null_part_sets[intro, simp]: hoelzl@47694: assumes "S \ sets M" shows "null_part M S \ sets M" "emeasure M (null_part M S) = 0" hoelzl@40859: proof - hoelzl@47694: have S: "S \ sets (completion M)" using assms by auto hoelzl@47694: have "S - main_part M S \ sets M" using assms by auto hoelzl@40859: moreover hoelzl@40859: from main_part_null_part_Un[OF S] main_part_null_part_Int[OF S] hoelzl@47694: have "S - main_part M S = null_part M S" by auto hoelzl@47694: ultimately show sets: "null_part M S \ sets M" by auto hoelzl@40859: from null_part[OF S] guess N .. hoelzl@47694: with emeasure_eq_0[of N _ "null_part M S"] sets hoelzl@47694: show "emeasure M (null_part M S) = 0" by auto hoelzl@40859: qed hoelzl@40859: hoelzl@47694: lemma emeasure_main_part_UN: hoelzl@40859: fixes S :: "nat \ 'a set" hoelzl@47694: assumes "range S \ sets (completion M)" hoelzl@47694: shows "emeasure M (main_part M (\i. (S i))) = emeasure M (\i. main_part M (S i))" hoelzl@40859: proof - hoelzl@47694: have S: "\i. S i \ sets (completion M)" using assms by auto hoelzl@47694: then have UN: "(\i. S i) \ sets (completion M)" by auto hoelzl@47694: have "\i. \N. N \ null_sets M \ null_part M (S i) \ N" hoelzl@40859: using null_part[OF S] by auto hoelzl@40859: from choice[OF this] guess N .. note N = this hoelzl@47694: then have UN_N: "(\i. N i) \ null_sets M" by (intro null_sets_UN) auto hoelzl@47694: have "(\i. S i) \ sets (completion M)" using S by auto hoelzl@40859: from null_part[OF this] guess N' .. note N' = this hoelzl@40859: let ?N = "(\i. N i) \ N'" hoelzl@47694: have null_set: "?N \ null_sets M" using N' UN_N by (intro null_sets.Un) auto hoelzl@47694: have "main_part M (\i. S i) \ ?N = (main_part M (\i. S i) \ null_part M (\i. S i)) \ ?N" hoelzl@40859: using N' by auto hoelzl@47694: also have "\ = (\i. main_part M (S i) \ null_part M (S i)) \ ?N" hoelzl@40859: unfolding main_part_null_part_Un[OF S] main_part_null_part_Un[OF UN] by auto hoelzl@47694: also have "\ = (\i. main_part M (S i)) \ ?N" hoelzl@40859: using N by auto hoelzl@47694: finally have *: "main_part M (\i. S i) \ ?N = (\i. main_part M (S i)) \ ?N" . hoelzl@47694: have "emeasure M (main_part M (\i. S i)) = emeasure M (main_part M (\i. S i) \ ?N)" hoelzl@47694: using null_set UN by (intro emeasure_Un_null_set[symmetric]) auto hoelzl@47694: also have "\ = emeasure M ((\i. main_part M (S i)) \ ?N)" hoelzl@40859: unfolding * .. hoelzl@47694: also have "\ = emeasure M (\i. main_part M (S i))" hoelzl@47694: using null_set S by (intro emeasure_Un_null_set) auto hoelzl@41689: finally show ?thesis . hoelzl@40859: qed hoelzl@40859: hoelzl@47694: lemma emeasure_completion[simp]: hoelzl@47694: assumes S: "S \ sets (completion M)" shows "emeasure (completion M) S = emeasure M (main_part M S)" hoelzl@47694: proof (subst emeasure_measure_of[OF completion_def completion_into_space]) hoelzl@47694: let ?\ = "emeasure M \ main_part M" hoelzl@47694: show "S \ sets (completion M)" "?\ S = emeasure M (main_part M S) " using S by simp_all hoelzl@47694: show "positive (sets (completion M)) ?\" hoelzl@47694: by (simp add: positive_def emeasure_nonneg) hoelzl@47694: show "countably_additive (sets (completion M)) ?\" hoelzl@47694: proof (intro countably_additiveI) hoelzl@47694: fix A :: "nat \ 'a set" assume A: "range A \ sets (completion M)" "disjoint_family A" hoelzl@47694: have "disjoint_family (\i. main_part M (A i))" hoelzl@47694: proof (intro disjoint_family_on_bisimulation[OF A(2)]) hoelzl@47694: fix n m assume "A n \ A m = {}" hoelzl@47694: then have "(main_part M (A n) \ null_part M (A n)) \ (main_part M (A m) \ null_part M (A m)) = {}" hoelzl@47694: using A by (subst (1 2) main_part_null_part_Un) auto hoelzl@47694: then show "main_part M (A n) \ main_part M (A m) = {}" by auto hoelzl@47694: qed hoelzl@47694: then have "(\n. emeasure M (main_part M (A n))) = emeasure M (\i. main_part M (A i))" hoelzl@47694: using A by (auto intro!: suminf_emeasure) hoelzl@47694: then show "(\n. ?\ (A n)) = ?\ (UNION UNIV A)" hoelzl@47694: by (simp add: completion_def emeasure_main_part_UN[OF A(1)]) hoelzl@47694: qed hoelzl@40859: qed hoelzl@40859: hoelzl@47694: lemma emeasure_completion_UN: hoelzl@47694: "range S \ sets (completion M) \ hoelzl@47694: emeasure (completion M) (\i::nat. (S i)) = emeasure M (\i. main_part M (S i))" hoelzl@47694: by (subst emeasure_completion) (auto simp add: emeasure_main_part_UN) hoelzl@40859: hoelzl@47694: lemma emeasure_completion_Un: hoelzl@47694: assumes S: "S \ sets (completion M)" and T: "T \ sets (completion M)" hoelzl@47694: shows "emeasure (completion M) (S \ T) = emeasure M (main_part M S \ main_part M T)" hoelzl@47694: proof (subst emeasure_completion) hoelzl@47694: have UN: "(\i. binary (main_part M S) (main_part M T) i) = (\i. main_part M (binary S T i))" hoelzl@47694: unfolding binary_def by (auto split: split_if_asm) hoelzl@47694: show "emeasure M (main_part M (S \ T)) = emeasure M (main_part M S \ main_part M T)" hoelzl@47694: using emeasure_main_part_UN[of "binary S T" M] assms hoelzl@47694: unfolding range_binary_eq Un_range_binary UN by auto hoelzl@47694: qed (auto intro: S T) hoelzl@47694: hoelzl@47694: lemma sets_completionI_sub: hoelzl@47694: assumes N: "N' \ null_sets M" "N \ N'" hoelzl@47694: shows "N \ sets (completion M)" hoelzl@47694: using assms by (intro sets_completionI[of _ "{}" N N']) auto hoelzl@47694: hoelzl@47694: lemma completion_ex_simple_function: hoelzl@47694: assumes f: "simple_function (completion M) f" hoelzl@47694: shows "\f'. simple_function M f' \ (AE x in M. f x = f' x)" hoelzl@40859: proof - wenzelm@46731: let ?F = "\x. f -` {x} \ space M" hoelzl@47694: have F: "\x. ?F x \ sets (completion M)" and fin: "finite (f`space M)" hoelzl@47694: using simple_functionD[OF f] simple_functionD[OF f] by simp_all hoelzl@47694: have "\x. \N. N \ null_sets M \ null_part M (?F x) \ N" hoelzl@40859: using F null_part by auto hoelzl@40859: from choice[OF this] obtain N where hoelzl@47694: N: "\x. null_part M (?F x) \ N x" "\x. N x \ null_sets M" by auto wenzelm@46731: let ?N = "\x\f`space M. N x" wenzelm@46731: let ?f' = "\x. if x \ ?N then undefined else f x" hoelzl@47694: have sets: "?N \ null_sets M" using N fin by (intro null_sets.finite_UN) auto hoelzl@40859: show ?thesis unfolding simple_function_def hoelzl@40859: proof (safe intro!: exI[of _ ?f']) hoelzl@40859: have "?f' ` space M \ f`space M \ {undefined}" by auto hoelzl@47694: from finite_subset[OF this] simple_functionD(1)[OF f] hoelzl@40859: show "finite (?f' ` space M)" by auto hoelzl@40859: next hoelzl@40859: fix x assume "x \ space M" hoelzl@40859: have "?f' -` {?f' x} \ space M = hoelzl@40859: (if x \ ?N then ?F undefined \ ?N hoelzl@40859: else if f x = undefined then ?F (f x) \ ?N hoelzl@40859: else ?F (f x) - ?N)" immler@50244: using N(2) sets.sets_into_space by (auto split: split_if_asm simp: null_sets_def) hoelzl@40859: moreover { fix y have "?F y \ ?N \ sets M" hoelzl@40859: proof cases hoelzl@40859: assume y: "y \ f`space M" hoelzl@47694: have "?F y \ ?N = (main_part M (?F y) \ null_part M (?F y)) \ ?N" hoelzl@40859: using main_part_null_part_Un[OF F] by auto hoelzl@47694: also have "\ = main_part M (?F y) \ ?N" hoelzl@40859: using y N by auto hoelzl@40859: finally show ?thesis hoelzl@40859: using F sets by auto hoelzl@40859: next hoelzl@40859: assume "y \ f`space M" then have "?F y = {}" by auto hoelzl@40859: then show ?thesis using sets by auto hoelzl@40859: qed } hoelzl@40859: moreover { hoelzl@47694: have "?F (f x) - ?N = main_part M (?F (f x)) \ null_part M (?F (f x)) - ?N" hoelzl@40859: using main_part_null_part_Un[OF F] by auto hoelzl@47694: also have "\ = main_part M (?F (f x)) - ?N" hoelzl@40859: using N `x \ space M` by auto hoelzl@40859: finally have "?F (f x) - ?N \ sets M" hoelzl@40859: using F sets by auto } hoelzl@40859: ultimately show "?f' -` {?f' x} \ space M \ sets M" by auto hoelzl@40859: next hoelzl@47694: show "AE x in M. f x = ?f' x" hoelzl@40859: by (rule AE_I', rule sets) auto hoelzl@40859: qed hoelzl@40859: qed hoelzl@40859: hoelzl@47694: lemma completion_ex_borel_measurable_pos: hoelzl@43920: fixes g :: "'a \ ereal" hoelzl@47694: assumes g: "g \ borel_measurable (completion M)" and "\x. 0 \ g x" hoelzl@47694: shows "\g'\borel_measurable M. (AE x in M. g x = g' x)" hoelzl@40859: proof - hoelzl@47694: from g[THEN borel_measurable_implies_simple_function_sequence'] guess f . note f = this hoelzl@41981: from this(1)[THEN completion_ex_simple_function] hoelzl@47694: have "\i. \f'. simple_function M f' \ (AE x in M. f i x = f' x)" .. hoelzl@40859: from this[THEN choice] obtain f' where hoelzl@41689: sf: "\i. simple_function M (f' i)" and hoelzl@47694: AE: "\i. AE x in M. f i x = f' i x" by auto hoelzl@40859: show ?thesis hoelzl@40859: proof (intro bexI) hoelzl@41981: from AE[unfolded AE_all_countable[symmetric]] hoelzl@47694: show "AE x in M. g x = (SUP i. f' i x)" (is "AE x in M. g x = ?f x") hoelzl@41705: proof (elim AE_mp, safe intro!: AE_I2) hoelzl@40859: fix x assume eq: "\i. f i x = f' i x" hoelzl@41981: moreover have "g x = (SUP i. f i x)" hoelzl@41981: unfolding f using `0 \ g x` by (auto split: split_max) hoelzl@41981: ultimately show "g x = ?f x" by auto hoelzl@40859: qed hoelzl@40859: show "?f \ borel_measurable M" hoelzl@56949: using sf[THEN borel_measurable_simple_function] by auto hoelzl@40859: qed hoelzl@40859: qed hoelzl@40859: hoelzl@47694: lemma completion_ex_borel_measurable: hoelzl@43920: fixes g :: "'a \ ereal" hoelzl@47694: assumes g: "g \ borel_measurable (completion M)" hoelzl@47694: shows "\g'\borel_measurable M. (AE x in M. g x = g' x)" hoelzl@41981: proof - hoelzl@47694: have "(\x. max 0 (g x)) \ borel_measurable (completion M)" "\x. 0 \ max 0 (g x)" using g by auto hoelzl@41981: from completion_ex_borel_measurable_pos[OF this] guess g_pos .. hoelzl@41981: moreover hoelzl@47694: have "(\x. max 0 (- g x)) \ borel_measurable (completion M)" "\x. 0 \ max 0 (- g x)" using g by auto hoelzl@41981: from completion_ex_borel_measurable_pos[OF this] guess g_neg .. hoelzl@41981: ultimately hoelzl@41981: show ?thesis hoelzl@41981: proof (safe intro!: bexI[of _ "\x. g_pos x - g_neg x"]) hoelzl@47694: show "AE x in M. max 0 (- g x) = g_neg x \ max 0 (g x) = g_pos x \ g x = g_pos x - g_neg x" hoelzl@41981: proof (intro AE_I2 impI) hoelzl@41981: fix x assume g: "max 0 (- g x) = g_neg x" "max 0 (g x) = g_pos x" hoelzl@41981: show "g x = g_pos x - g_neg x" unfolding g[symmetric] hoelzl@41981: by (cases "g x") (auto split: split_max) hoelzl@41981: qed hoelzl@41981: qed auto hoelzl@41981: qed hoelzl@41981: hoelzl@40859: end