# HG changeset patch # User hoelzl # Date 1301401659 -7200 # Node ID 5b52c6a9c62741cf0a5cd8bdb72ff29b6fdf0314 # Parent 8448713d48b72dc8bebd9be5686e14478a23fd2a split Product_Measure into Binary_Product_Measure and Finite_Product_Measure diff -r 8448713d48b7 -r 5b52c6a9c627 src/HOL/IsaMakefile --- a/src/HOL/IsaMakefile Tue Mar 29 14:27:31 2011 +0200 +++ b/src/HOL/IsaMakefile Tue Mar 29 14:27:39 2011 +0200 @@ -1186,17 +1186,16 @@ HOL-Probability: HOL-Multivariate_Analysis $(OUT)/HOL-Probability $(OUT)/HOL-Probability: $(OUT)/HOL-Multivariate_Analysis \ - Probability/Borel_Space.thy Probability/Caratheodory.thy \ - Probability/Complete_Measure.thy \ + Probability/Binary_Product_Measure.thy Probability/Borel_Space.thy \ + Probability/Caratheodory.thy Probability/Complete_Measure.thy \ Probability/ex/Dining_Cryptographers.thy \ Probability/ex/Koepf_Duermuth_Countermeasure.thy \ - Probability/Information.thy Probability/Lebesgue_Integration.thy \ - Probability/Lebesgue_Measure.thy Probability/Measure.thy \ - Probability/Probability_Space.thy Probability/Probability.thy \ - Probability/Product_Measure.thy Probability/Radon_Nikodym.thy \ + Probability/Finite_Product_Measure.thy Probability/Information.thy \ + Probability/Lebesgue_Integration.thy Probability/Lebesgue_Measure.thy \ + Probability/Measure.thy Probability/Probability_Space.thy \ + Probability/Probability.thy Probability/Radon_Nikodym.thy \ Probability/ROOT.ML Probability/Sigma_Algebra.thy \ - Library/Countable.thy Library/FuncSet.thy \ - Library/Nat_Bijection.thy + Library/Countable.thy Library/FuncSet.thy Library/Nat_Bijection.thy @cd Probability; $(ISABELLE_TOOL) usedir -b -g true $(OUT)/HOL-Multivariate_Analysis HOL-Probability diff -r 8448713d48b7 -r 5b52c6a9c627 src/HOL/Probability/Binary_Product_Measure.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Probability/Binary_Product_Measure.thy Tue Mar 29 14:27:39 2011 +0200 @@ -0,0 +1,975 @@ +(* Title: HOL/Probability/Binary_Product_Measure.thy + Author: Johannes Hölzl, TU München +*) + +header {*Binary product measures*} + +theory Binary_Product_Measure +imports Lebesgue_Integration +begin + +lemma times_Int_times: "A \ B \ C \ D = (A \ C) \ (B \ D)" + by auto + +lemma Pair_vimage_times[simp]: "\A B x. Pair x -` (A \ B) = (if x \ A then B else {})" + by auto + +lemma rev_Pair_vimage_times[simp]: "\A B y. (\x. (x, y)) -` (A \ B) = (if y \ B then A else {})" + by auto + +lemma case_prod_distrib: "f (case x of (x, y) \ g x y) = (case x of (x, y) \ f (g x y))" + by (cases x) simp + +lemma split_beta': "(\(x,y). f x y) = (\x. f (fst x) (snd x))" + by (auto simp: fun_eq_iff) + +section "Binary products" + +definition + "pair_measure_generator A B = + \ space = space A \ space B, + sets = {a \ b | a b. a \ sets A \ b \ sets B}, + measure = \X. \\<^isup>+x. (\\<^isup>+y. indicator X (x,y) \B) \A \" + +definition pair_measure (infixr "\\<^isub>M" 80) where + "A \\<^isub>M B = sigma (pair_measure_generator A B)" + +locale pair_sigma_algebra = M1: sigma_algebra M1 + M2: sigma_algebra M2 + for M1 :: "('a, 'c) measure_space_scheme" and M2 :: "('b, 'd) measure_space_scheme" + +abbreviation (in pair_sigma_algebra) + "E \ pair_measure_generator M1 M2" + +abbreviation (in pair_sigma_algebra) + "P \ M1 \\<^isub>M M2" + +lemma sigma_algebra_pair_measure: + "sets M1 \ Pow (space M1) \ sets M2 \ Pow (space M2) \ sigma_algebra (pair_measure M1 M2)" + by (force simp: pair_measure_def pair_measure_generator_def intro!: sigma_algebra_sigma) + +sublocale pair_sigma_algebra \ sigma_algebra P + using M1.space_closed M2.space_closed + by (rule sigma_algebra_pair_measure) + +lemma pair_measure_generatorI[intro, simp]: + "x \ sets A \ y \ sets B \ x \ y \ sets (pair_measure_generator A B)" + by (auto simp add: pair_measure_generator_def) + +lemma pair_measureI[intro, simp]: + "x \ sets A \ y \ sets B \ x \ y \ sets (A \\<^isub>M B)" + by (auto simp add: pair_measure_def) + +lemma space_pair_measure: + "space (A \\<^isub>M B) = space A \ space B" + by (simp add: pair_measure_def pair_measure_generator_def) + +lemma sets_pair_measure_generator: + "sets (pair_measure_generator N M) = (\(x, y). x \ y) ` (sets N \ sets M)" + unfolding pair_measure_generator_def by auto + +lemma pair_measure_generator_sets_into_space: + assumes "sets M \ Pow (space M)" "sets N \ Pow (space N)" + shows "sets (pair_measure_generator M N) \ Pow (space (pair_measure_generator M N))" + using assms by (auto simp: pair_measure_generator_def) + +lemma pair_measure_generator_Int_snd: + assumes "sets S1 \ Pow (space S1)" + shows "sets (pair_measure_generator S1 (algebra.restricted_space S2 A)) = + sets (algebra.restricted_space (pair_measure_generator S1 S2) (space S1 \ A))" + (is "?L = ?R") + apply (auto simp: pair_measure_generator_def image_iff) + using assms + apply (rule_tac x="a \ xa" in exI) + apply force + using assms + apply (rule_tac x="a" in exI) + apply (rule_tac x="b \ A" in exI) + apply auto + done + +lemma (in pair_sigma_algebra) + shows measurable_fst[intro!, simp]: + "fst \ measurable P M1" (is ?fst) + and measurable_snd[intro!, simp]: + "snd \ measurable P M2" (is ?snd) +proof - + { fix X assume "X \ sets M1" + then have "\X1\sets M1. \X2\sets M2. fst -` X \ space M1 \ space M2 = X1 \ X2" + apply - apply (rule bexI[of _ X]) apply (rule bexI[of _ "space M2"]) + using M1.sets_into_space by force+ } + moreover + { fix X assume "X \ sets M2" + then have "\X1\sets M1. \X2\sets M2. snd -` X \ space M1 \ space M2 = X1 \ X2" + apply - apply (rule bexI[of _ "space M1"]) apply (rule bexI[of _ X]) + using M2.sets_into_space by force+ } + ultimately have "?fst \ ?snd" + by (fastsimp simp: measurable_def sets_sigma space_pair_measure + intro!: sigma_sets.Basic) + then show ?fst ?snd by auto +qed + +lemma (in pair_sigma_algebra) measurable_pair_iff: + assumes "sigma_algebra M" + shows "f \ measurable M P \ + (fst \ f) \ measurable M M1 \ (snd \ f) \ measurable M M2" +proof - + interpret M: sigma_algebra M by fact + from assms show ?thesis + proof (safe intro!: measurable_comp[where b=P]) + assume f: "(fst \ f) \ measurable M M1" and s: "(snd \ f) \ measurable M M2" + show "f \ measurable M P" unfolding pair_measure_def + proof (rule M.measurable_sigma) + show "sets (pair_measure_generator M1 M2) \ Pow (space E)" + unfolding pair_measure_generator_def using M1.sets_into_space M2.sets_into_space by auto + show "f \ space M \ space E" + using f s by (auto simp: mem_Times_iff measurable_def comp_def space_sigma pair_measure_generator_def) + fix A assume "A \ sets E" + then obtain B C where "B \ sets M1" "C \ sets M2" "A = B \ C" + unfolding pair_measure_generator_def by auto + moreover have "(fst \ f) -` B \ space M \ sets M" + using f `B \ sets M1` unfolding measurable_def by auto + moreover have "(snd \ f) -` C \ space M \ sets M" + using s `C \ sets M2` unfolding measurable_def by auto + moreover have "f -` A \ space M = ((fst \ f) -` B \ space M) \ ((snd \ f) -` C \ space M)" + unfolding `A = B \ C` by (auto simp: vimage_Times) + ultimately show "f -` A \ space M \ sets M" by auto + qed + qed +qed + +lemma (in pair_sigma_algebra) measurable_pair: + assumes "sigma_algebra M" + assumes "(fst \ f) \ measurable M M1" "(snd \ f) \ measurable M M2" + shows "f \ measurable M P" + unfolding measurable_pair_iff[OF assms(1)] using assms(2,3) by simp + +lemma pair_measure_generatorE: + assumes "X \ sets (pair_measure_generator M1 M2)" + obtains A B where "X = A \ B" "A \ sets M1" "B \ sets M2" + using assms unfolding pair_measure_generator_def by auto + +lemma (in pair_sigma_algebra) pair_measure_generator_swap: + "(\X. (\(x,y). (y,x)) -` X \ space M2 \ space M1) ` sets E = sets (pair_measure_generator M2 M1)" +proof (safe elim!: pair_measure_generatorE) + fix A B assume "A \ sets M1" "B \ sets M2" + moreover then have "(\(x, y). (y, x)) -` (A \ B) \ space M2 \ space M1 = B \ A" + using M1.sets_into_space M2.sets_into_space by auto + ultimately show "(\(x, y). (y, x)) -` (A \ B) \ space M2 \ space M1 \ sets (pair_measure_generator M2 M1)" + by (auto intro: pair_measure_generatorI) +next + fix A B assume "A \ sets M1" "B \ sets M2" + then show "B \ A \ (\X. (\(x, y). (y, x)) -` X \ space M2 \ space M1) ` sets E" + using M1.sets_into_space M2.sets_into_space + by (auto intro!: image_eqI[where x="A \ B"] pair_measure_generatorI) +qed + +lemma (in pair_sigma_algebra) sets_pair_sigma_algebra_swap: + assumes Q: "Q \ sets P" + shows "(\(x,y). (y, x)) -` Q \ sets (M2 \\<^isub>M M1)" (is "_ \ sets ?Q") +proof - + let "?f Q" = "(\(x,y). (y, x)) -` Q \ space M2 \ space M1" + have *: "(\(x,y). (y, x)) -` Q = ?f Q" + using sets_into_space[OF Q] by (auto simp: space_pair_measure) + have "sets (M2 \\<^isub>M M1) = sets (sigma (pair_measure_generator M2 M1))" + unfolding pair_measure_def .. + also have "\ = sigma_sets (space M2 \ space M1) (?f ` sets E)" + unfolding sigma_def pair_measure_generator_swap[symmetric] + by (simp add: pair_measure_generator_def) + also have "\ = ?f ` sigma_sets (space M1 \ space M2) (sets E)" + using M1.sets_into_space M2.sets_into_space + by (intro sigma_sets_vimage) (auto simp: pair_measure_generator_def) + also have "\ = ?f ` sets P" + unfolding pair_measure_def pair_measure_generator_def sigma_def by simp + finally show ?thesis + using Q by (subst *) auto +qed + +lemma (in pair_sigma_algebra) pair_sigma_algebra_swap_measurable: + shows "(\(x,y). (y, x)) \ measurable P (M2 \\<^isub>M M1)" + (is "?f \ measurable ?P ?Q") + unfolding measurable_def +proof (intro CollectI conjI Pi_I ballI) + fix x assume "x \ space ?P" then show "(case x of (x, y) \ (y, x)) \ space ?Q" + unfolding pair_measure_generator_def pair_measure_def by auto +next + fix A assume "A \ sets (M2 \\<^isub>M M1)" + interpret Q: pair_sigma_algebra M2 M1 by default + with Q.sets_pair_sigma_algebra_swap[OF `A \ sets (M2 \\<^isub>M M1)`] + show "?f -` A \ space ?P \ sets ?P" by simp +qed + +lemma (in pair_sigma_algebra) measurable_cut_fst[simp,intro]: + assumes "Q \ sets P" shows "Pair x -` Q \ sets M2" +proof - + let ?Q' = "{Q. Q \ space P \ Pair x -` Q \ sets M2}" + let ?Q = "\ space = space P, sets = ?Q' \" + interpret Q: sigma_algebra ?Q + proof qed (auto simp: vimage_UN vimage_Diff space_pair_measure) + have "sets E \ sets ?Q" + using M1.sets_into_space M2.sets_into_space + by (auto simp: pair_measure_generator_def space_pair_measure) + then have "sets P \ sets ?Q" + apply (subst pair_measure_def, intro Q.sets_sigma_subset) + by (simp add: pair_measure_def) + with assms show ?thesis by auto +qed + +lemma (in pair_sigma_algebra) measurable_cut_snd: + assumes Q: "Q \ sets P" shows "(\x. (x, y)) -` Q \ sets M1" (is "?cut Q \ sets M1") +proof - + interpret Q: pair_sigma_algebra M2 M1 by default + with Q.measurable_cut_fst[OF sets_pair_sigma_algebra_swap[OF Q], of y] + show ?thesis by (simp add: vimage_compose[symmetric] comp_def) +qed + +lemma (in pair_sigma_algebra) measurable_pair_image_snd: + assumes m: "f \ measurable P M" and "x \ space M1" + shows "(\y. f (x, y)) \ measurable M2 M" + unfolding measurable_def +proof (intro CollectI conjI Pi_I ballI) + fix y assume "y \ space M2" with `f \ measurable P M` `x \ space M1` + show "f (x, y) \ space M" + unfolding measurable_def pair_measure_generator_def pair_measure_def by auto +next + fix A assume "A \ sets M" + then have "Pair x -` (f -` A \ space P) \ sets M2" (is "?C \ _") + using `f \ measurable P M` + by (intro measurable_cut_fst) (auto simp: measurable_def) + also have "?C = (\y. f (x, y)) -` A \ space M2" + using `x \ space M1` by (auto simp: pair_measure_generator_def pair_measure_def) + finally show "(\y. f (x, y)) -` A \ space M2 \ sets M2" . +qed + +lemma (in pair_sigma_algebra) measurable_pair_image_fst: + assumes m: "f \ measurable P M" and "y \ space M2" + shows "(\x. f (x, y)) \ measurable M1 M" +proof - + interpret Q: pair_sigma_algebra M2 M1 by default + from Q.measurable_pair_image_snd[OF measurable_comp `y \ space M2`, + OF Q.pair_sigma_algebra_swap_measurable m] + show ?thesis by simp +qed + +lemma (in pair_sigma_algebra) Int_stable_pair_measure_generator: "Int_stable E" + unfolding Int_stable_def +proof (intro ballI) + fix A B assume "A \ sets E" "B \ sets E" + then obtain A1 A2 B1 B2 where "A = A1 \ A2" "B = B1 \ B2" + "A1 \ sets M1" "A2 \ sets M2" "B1 \ sets M1" "B2 \ sets M2" + unfolding pair_measure_generator_def by auto + then show "A \ B \ sets E" + by (auto simp add: times_Int_times pair_measure_generator_def) +qed + +lemma finite_measure_cut_measurable: + fixes M1 :: "('a, 'c) measure_space_scheme" and M2 :: "('b, 'd) measure_space_scheme" + assumes "sigma_finite_measure M1" "finite_measure M2" + assumes "Q \ sets (M1 \\<^isub>M M2)" + shows "(\x. measure M2 (Pair x -` Q)) \ borel_measurable M1" + (is "?s Q \ _") +proof - + interpret M1: sigma_finite_measure M1 by fact + interpret M2: finite_measure M2 by fact + interpret pair_sigma_algebra M1 M2 by default + have [intro]: "sigma_algebra M1" by fact + have [intro]: "sigma_algebra M2" by fact + let ?D = "\ space = space P, sets = {A\sets P. ?s A \ borel_measurable M1} \" + note space_pair_measure[simp] + interpret dynkin_system ?D + proof (intro dynkin_systemI) + fix A assume "A \ sets ?D" then show "A \ space ?D" + using sets_into_space by simp + next + from top show "space ?D \ sets ?D" + by (auto simp add: if_distrib intro!: M1.measurable_If) + next + fix A assume "A \ sets ?D" + with sets_into_space have "\x. measure M2 (Pair x -` (space M1 \ space M2 - A)) = + (if x \ space M1 then measure M2 (space M2) - ?s A x else 0)" + by (auto intro!: M2.measure_compl simp: vimage_Diff) + with `A \ sets ?D` top show "space ?D - A \ sets ?D" + by (auto intro!: Diff M1.measurable_If M1.borel_measurable_extreal_diff) + next + fix F :: "nat \ ('a\'b) set" assume "disjoint_family F" "range F \ sets ?D" + moreover then have "\x. measure M2 (\i. Pair x -` F i) = (\i. ?s (F i) x)" + by (intro M2.measure_countably_additive[symmetric]) + (auto simp: disjoint_family_on_def) + ultimately show "(\i. F i) \ sets ?D" + by (auto simp: vimage_UN intro!: M1.borel_measurable_psuminf) + qed + have "sets P = sets ?D" apply (subst pair_measure_def) + proof (intro dynkin_lemma) + show "Int_stable E" by (rule Int_stable_pair_measure_generator) + from M1.sets_into_space have "\A. A \ sets M1 \ {x \ space M1. x \ A} = A" + by auto + then show "sets E \ sets ?D" + by (auto simp: pair_measure_generator_def sets_sigma if_distrib + intro: sigma_sets.Basic intro!: M1.measurable_If) + qed (auto simp: pair_measure_def) + with `Q \ sets P` have "Q \ sets ?D" by simp + then show "?s Q \ borel_measurable M1" by simp +qed + +subsection {* Binary products of $\sigma$-finite measure spaces *} + +locale pair_sigma_finite = M1: sigma_finite_measure M1 + M2: sigma_finite_measure M2 + for M1 :: "('a, 'c) measure_space_scheme" and M2 :: "('b, 'd) measure_space_scheme" + +sublocale pair_sigma_finite \ pair_sigma_algebra M1 M2 + by default + +lemma times_eq_iff: "A \ B = C \ D \ A = C \ B = D \ ((A = {} \ B = {}) \ (C = {} \ D = {}))" + by auto + +lemma sigma_sets_subseteq: assumes "A \ B" shows "sigma_sets X A \ sigma_sets X B" +proof + fix x assume "x \ sigma_sets X A" then show "x \ sigma_sets X B" + by induct (insert `A \ B`, auto intro: sigma_sets.intros) +qed + +lemma (in pair_sigma_finite) measure_cut_measurable_fst: + assumes "Q \ sets P" shows "(\x. measure M2 (Pair x -` Q)) \ borel_measurable M1" (is "?s Q \ _") +proof - + have [intro]: "sigma_algebra M1" and [intro]: "sigma_algebra M2" by default+ + have M1: "sigma_finite_measure M1" by default + from M2.disjoint_sigma_finite guess F .. note F = this + then have F_sets: "\i. F i \ sets M2" by auto + let "?C x i" = "F i \ Pair x -` Q" + { fix i + let ?R = "M2.restricted_space (F i)" + have [simp]: "space M1 \ F i \ space M1 \ space M2 = space M1 \ F i" + using F M2.sets_into_space by auto + let ?R2 = "M2.restricted_space (F i)" + have "(\x. measure ?R2 (Pair x -` (space M1 \ space ?R2 \ Q))) \ borel_measurable M1" + proof (intro finite_measure_cut_measurable[OF M1]) + show "finite_measure ?R2" + using F by (intro M2.restricted_to_finite_measure) auto + have "(space M1 \ space ?R2) \ Q \ (op \ (space M1 \ F i)) ` sets P" + using `Q \ sets P` by (auto simp: image_iff) + also have "\ = sigma_sets (space M1 \ F i) ((op \ (space M1 \ F i)) ` sets E)" + unfolding pair_measure_def pair_measure_generator_def sigma_def + using `F i \ sets M2` M2.sets_into_space + by (auto intro!: sigma_sets_Int sigma_sets.Basic) + also have "\ \ sets (M1 \\<^isub>M ?R2)" + using M1.sets_into_space + apply (auto simp: times_Int_times pair_measure_def pair_measure_generator_def sigma_def + intro!: sigma_sets_subseteq) + apply (rule_tac x="a" in exI) + apply (rule_tac x="b \ F i" in exI) + by auto + finally show "(space M1 \ space ?R2) \ Q \ sets (M1 \\<^isub>M ?R2)" . + qed + moreover have "\x. Pair x -` (space M1 \ F i \ Q) = ?C x i" + using `Q \ sets P` sets_into_space by (auto simp: space_pair_measure) + ultimately have "(\x. measure M2 (?C x i)) \ borel_measurable M1" + by simp } + moreover + { fix x + have "(\i. measure M2 (?C x i)) = measure M2 (\i. ?C x i)" + proof (intro M2.measure_countably_additive) + show "range (?C x) \ sets M2" + using F `Q \ sets P` by (auto intro!: M2.Int) + have "disjoint_family F" using F by auto + show "disjoint_family (?C x)" + by (rule disjoint_family_on_bisimulation[OF `disjoint_family F`]) auto + qed + also have "(\i. ?C x i) = Pair x -` Q" + using F sets_into_space `Q \ sets P` + by (auto simp: space_pair_measure) + finally have "measure M2 (Pair x -` Q) = (\i. measure M2 (?C x i))" + by simp } + ultimately show ?thesis using `Q \ sets P` F_sets + by (auto intro!: M1.borel_measurable_psuminf M2.Int) +qed + +lemma (in pair_sigma_finite) measure_cut_measurable_snd: + assumes "Q \ sets P" shows "(\y. M1.\ ((\x. (x, y)) -` Q)) \ borel_measurable M2" +proof - + interpret Q: pair_sigma_finite M2 M1 by default + note sets_pair_sigma_algebra_swap[OF assms] + from Q.measure_cut_measurable_fst[OF this] + show ?thesis by (simp add: vimage_compose[symmetric] comp_def) +qed + +lemma (in pair_sigma_algebra) pair_sigma_algebra_measurable: + assumes "f \ measurable P M" shows "(\(x,y). f (y, x)) \ measurable (M2 \\<^isub>M M1) M" +proof - + interpret Q: pair_sigma_algebra M2 M1 by default + have *: "(\(x,y). f (y, x)) = f \ (\(x,y). (y, x))" by (simp add: fun_eq_iff) + show ?thesis + using Q.pair_sigma_algebra_swap_measurable assms + unfolding * by (rule measurable_comp) +qed + +lemma (in pair_sigma_finite) pair_measure_alt: + assumes "A \ sets P" + shows "measure (M1 \\<^isub>M M2) A = (\\<^isup>+ x. measure M2 (Pair x -` A) \M1)" + apply (simp add: pair_measure_def pair_measure_generator_def) +proof (rule M1.positive_integral_cong) + fix x assume "x \ space M1" + have *: "\y. indicator A (x, y) = (indicator (Pair x -` A) y :: extreal)" + unfolding indicator_def by auto + show "(\\<^isup>+ y. indicator A (x, y) \M2) = measure M2 (Pair x -` A)" + unfolding * + apply (subst M2.positive_integral_indicator) + apply (rule measurable_cut_fst[OF assms]) + by simp +qed + +lemma (in pair_sigma_finite) pair_measure_times: + assumes A: "A \ sets M1" and "B \ sets M2" + shows "measure (M1 \\<^isub>M M2) (A \ B) = M1.\ A * measure M2 B" +proof - + have "measure (M1 \\<^isub>M M2) (A \ B) = (\\<^isup>+ x. measure M2 B * indicator A x \M1)" + using assms by (auto intro!: M1.positive_integral_cong simp: pair_measure_alt) + with assms show ?thesis + by (simp add: M1.positive_integral_cmult_indicator ac_simps) +qed + +lemma (in measure_space) measure_not_negative[simp,intro]: + assumes A: "A \ sets M" shows "\ A \ - \" + using positive_measure[OF A] by auto + +lemma (in pair_sigma_finite) sigma_finite_up_in_pair_measure_generator: + "\F::nat \ ('a \ 'b) set. range F \ sets E \ incseq F \ (\i. F i) = space E \ + (\i. measure (M1 \\<^isub>M M2) (F i) \ \)" +proof - + obtain F1 :: "nat \ 'a set" and F2 :: "nat \ 'b set" where + F1: "range F1 \ sets M1" "incseq F1" "(\i. F1 i) = space M1" "\i. M1.\ (F1 i) \ \" and + F2: "range F2 \ sets M2" "incseq F2" "(\i. F2 i) = space M2" "\i. M2.\ (F2 i) \ \" + using M1.sigma_finite_up M2.sigma_finite_up by auto + then have space: "space M1 = (\i. F1 i)" "space M2 = (\i. F2 i)" by auto + let ?F = "\i. F1 i \ F2 i" + show ?thesis unfolding space_pair_measure + proof (intro exI[of _ ?F] conjI allI) + show "range ?F \ sets E" using F1 F2 + by (fastsimp intro!: pair_measure_generatorI) + next + have "space M1 \ space M2 \ (\i. ?F i)" + proof (intro subsetI) + fix x assume "x \ space M1 \ space M2" + then obtain i j where "fst x \ F1 i" "snd x \ F2 j" + by (auto simp: space) + then have "fst x \ F1 (max i j)" "snd x \ F2 (max j i)" + using `incseq F1` `incseq F2` unfolding incseq_def + by (force split: split_max)+ + then have "(fst x, snd x) \ F1 (max i j) \ F2 (max i j)" + by (intro SigmaI) (auto simp add: min_max.sup_commute) + then show "x \ (\i. ?F i)" by auto + qed + then show "(\i. ?F i) = space E" + using space by (auto simp: space pair_measure_generator_def) + next + fix i show "incseq (\i. F1 i \ F2 i)" + using `incseq F1` `incseq F2` unfolding incseq_Suc_iff by auto + next + fix i + from F1 F2 have "F1 i \ sets M1" "F2 i \ sets M2" by auto + with F1 F2 M1.positive_measure[OF this(1)] M2.positive_measure[OF this(2)] + show "measure P (F1 i \ F2 i) \ \" + by (simp add: pair_measure_times) + qed +qed + +sublocale pair_sigma_finite \ sigma_finite_measure P +proof + show "positive P (measure P)" + unfolding pair_measure_def pair_measure_generator_def sigma_def positive_def + by (auto intro: M1.positive_integral_positive M2.positive_integral_positive) + + show "countably_additive P (measure P)" + unfolding countably_additive_def + proof (intro allI impI) + fix F :: "nat \ ('a \ 'b) set" + assume F: "range F \ sets P" "disjoint_family F" + from F have *: "\i. F i \ sets P" "(\i. F i) \ sets P" by auto + moreover from F have "\i. (\x. measure M2 (Pair x -` F i)) \ borel_measurable M1" + by (intro measure_cut_measurable_fst) auto + moreover have "\x. disjoint_family (\i. Pair x -` F i)" + by (intro disjoint_family_on_bisimulation[OF F(2)]) auto + moreover have "\x. x \ space M1 \ range (\i. Pair x -` F i) \ sets M2" + using F by auto + ultimately show "(\n. measure P (F n)) = measure P (\i. F i)" + by (simp add: pair_measure_alt vimage_UN M1.positive_integral_suminf[symmetric] + M2.measure_countably_additive + cong: M1.positive_integral_cong) + qed + + from sigma_finite_up_in_pair_measure_generator guess F :: "nat \ ('a \ 'b) set" .. note F = this + show "\F::nat \ ('a \ 'b) set. range F \ sets P \ (\i. F i) = space P \ (\i. measure P (F i) \ \)" + proof (rule exI[of _ F], intro conjI) + show "range F \ sets P" using F by (auto simp: pair_measure_def) + show "(\i. F i) = space P" + using F by (auto simp: pair_measure_def pair_measure_generator_def) + show "\i. measure P (F i) \ \" using F by auto + qed +qed + +lemma (in pair_sigma_algebra) sets_swap: + assumes "A \ sets P" + shows "(\(x, y). (y, x)) -` A \ space (M2 \\<^isub>M M1) \ sets (M2 \\<^isub>M M1)" + (is "_ -` A \ space ?Q \ sets ?Q") +proof - + have *: "(\(x, y). (y, x)) -` A \ space ?Q = (\(x, y). (y, x)) -` A" + using `A \ sets P` sets_into_space by (auto simp: space_pair_measure) + show ?thesis + unfolding * using assms by (rule sets_pair_sigma_algebra_swap) +qed + +lemma (in pair_sigma_finite) pair_measure_alt2: + assumes A: "A \ sets P" + shows "\ A = (\\<^isup>+y. M1.\ ((\x. (x, y)) -` A) \M2)" + (is "_ = ?\ A") +proof - + interpret Q: pair_sigma_finite M2 M1 by default + from sigma_finite_up_in_pair_measure_generator guess F :: "nat \ ('a \ 'b) set" .. note F = this + have [simp]: "\m. \ space = space E, sets = sets (sigma E), measure = m \ = P\ measure := m \" + unfolding pair_measure_def by simp + + have "\ A = Q.\ ((\(y, x). (x, y)) -` A \ space Q.P)" + proof (rule measure_unique_Int_stable_vimage[OF Int_stable_pair_measure_generator]) + show "measure_space P" "measure_space Q.P" by default + show "(\(y, x). (x, y)) \ measurable Q.P P" by (rule Q.pair_sigma_algebra_swap_measurable) + show "sets (sigma E) = sets P" "space E = space P" "A \ sets (sigma E)" + using assms unfolding pair_measure_def by auto + show "range F \ sets E" "incseq F" "(\i. F i) = space E" "\i. \ (F i) \ \" + using F `A \ sets P` by (auto simp: pair_measure_def) + fix X assume "X \ sets E" + then obtain A B where X[simp]: "X = A \ B" and AB: "A \ sets M1" "B \ sets M2" + unfolding pair_measure_def pair_measure_generator_def by auto + then have "(\(y, x). (x, y)) -` X \ space Q.P = B \ A" + using M1.sets_into_space M2.sets_into_space by (auto simp: space_pair_measure) + then show "\ X = Q.\ ((\(y, x). (x, y)) -` X \ space Q.P)" + using AB by (simp add: pair_measure_times Q.pair_measure_times ac_simps) + qed + then show ?thesis + using sets_into_space[OF A] Q.pair_measure_alt[OF sets_swap[OF A]] + by (auto simp add: Q.pair_measure_alt space_pair_measure + intro!: M2.positive_integral_cong arg_cong[where f="M1.\"]) +qed + +lemma pair_sigma_algebra_sigma: + assumes 1: "incseq S1" "(\i. S1 i) = space E1" "range S1 \ sets E1" and E1: "sets E1 \ Pow (space E1)" + assumes 2: "decseq S2" "(\i. S2 i) = space E2" "range S2 \ sets E2" and E2: "sets E2 \ Pow (space E2)" + shows "sets (sigma (pair_measure_generator (sigma E1) (sigma E2))) = sets (sigma (pair_measure_generator E1 E2))" + (is "sets ?S = sets ?E") +proof - + interpret M1: sigma_algebra "sigma E1" using E1 by (rule sigma_algebra_sigma) + interpret M2: sigma_algebra "sigma E2" using E2 by (rule sigma_algebra_sigma) + have P: "sets (pair_measure_generator E1 E2) \ Pow (space E1 \ space E2)" + using E1 E2 by (auto simp add: pair_measure_generator_def) + interpret E: sigma_algebra ?E unfolding pair_measure_generator_def + using E1 E2 by (intro sigma_algebra_sigma) auto + { fix A assume "A \ sets E1" + then have "fst -` A \ space ?E = A \ (\i. S2 i)" + using E1 2 unfolding pair_measure_generator_def by auto + also have "\ = (\i. A \ S2 i)" by auto + also have "\ \ sets ?E" unfolding pair_measure_generator_def sets_sigma + using 2 `A \ sets E1` + by (intro sigma_sets.Union) + (force simp: image_subset_iff intro!: sigma_sets.Basic) + finally have "fst -` A \ space ?E \ sets ?E" . } + moreover + { fix B assume "B \ sets E2" + then have "snd -` B \ space ?E = (\i. S1 i) \ B" + using E2 1 unfolding pair_measure_generator_def by auto + also have "\ = (\i. S1 i \ B)" by auto + also have "\ \ sets ?E" + using 1 `B \ sets E2` unfolding pair_measure_generator_def sets_sigma + by (intro sigma_sets.Union) + (force simp: image_subset_iff intro!: sigma_sets.Basic) + finally have "snd -` B \ space ?E \ sets ?E" . } + ultimately have proj: + "fst \ measurable ?E (sigma E1) \ snd \ measurable ?E (sigma E2)" + using E1 E2 by (subst (1 2) E.measurable_iff_sigma) + (auto simp: pair_measure_generator_def sets_sigma) + { fix A B assume A: "A \ sets (sigma E1)" and B: "B \ sets (sigma E2)" + with proj have "fst -` A \ space ?E \ sets ?E" "snd -` B \ space ?E \ sets ?E" + unfolding measurable_def by simp_all + moreover have "A \ B = (fst -` A \ space ?E) \ (snd -` B \ space ?E)" + using A B M1.sets_into_space M2.sets_into_space + by (auto simp: pair_measure_generator_def) + ultimately have "A \ B \ sets ?E" by auto } + then have "sigma_sets (space ?E) (sets (pair_measure_generator (sigma E1) (sigma E2))) \ sets ?E" + by (intro E.sigma_sets_subset) (auto simp add: pair_measure_generator_def sets_sigma) + then have subset: "sets ?S \ sets ?E" + by (simp add: sets_sigma pair_measure_generator_def) + show "sets ?S = sets ?E" + proof (intro set_eqI iffI) + fix A assume "A \ sets ?E" then show "A \ sets ?S" + unfolding sets_sigma + proof induct + case (Basic A) then show ?case + by (auto simp: pair_measure_generator_def sets_sigma intro: sigma_sets.Basic) + qed (auto intro: sigma_sets.intros simp: pair_measure_generator_def) + next + fix A assume "A \ sets ?S" then show "A \ sets ?E" using subset by auto + qed +qed + +section "Fubinis theorem" + +lemma (in pair_sigma_finite) simple_function_cut: + assumes f: "simple_function P f" "\x. 0 \ f x" + shows "(\x. \\<^isup>+y. f (x, y) \M2) \ borel_measurable M1" + and "(\\<^isup>+ x. (\\<^isup>+ y. f (x, y) \M2) \M1) = integral\<^isup>P P f" +proof - + have f_borel: "f \ borel_measurable P" + using f(1) by (rule borel_measurable_simple_function) + let "?F z" = "f -` {z} \ space P" + let "?F' x z" = "Pair x -` ?F z" + { fix x assume "x \ space M1" + have [simp]: "\z y. indicator (?F z) (x, y) = indicator (?F' x z) y" + by (auto simp: indicator_def) + have "\y. y \ space M2 \ (x, y) \ space P" using `x \ space M1` + by (simp add: space_pair_measure) + moreover have "\x z. ?F' x z \ sets M2" using f_borel + by (intro borel_measurable_vimage measurable_cut_fst) + ultimately have "simple_function M2 (\ y. f (x, y))" + apply (rule_tac M2.simple_function_cong[THEN iffD2, OF _]) + apply (rule simple_function_indicator_representation[OF f(1)]) + using `x \ space M1` by (auto simp del: space_sigma) } + note M2_sf = this + { fix x assume x: "x \ space M1" + then have "(\\<^isup>+y. f (x, y) \M2) = (\z\f ` space P. z * M2.\ (?F' x z))" + unfolding M2.positive_integral_eq_simple_integral[OF M2_sf[OF x] f(2)] + unfolding simple_integral_def + proof (safe intro!: setsum_mono_zero_cong_left) + from f(1) show "finite (f ` space P)" by (rule simple_functionD) + next + fix y assume "y \ space M2" then show "f (x, y) \ f ` space P" + using `x \ space M1` by (auto simp: space_pair_measure) + next + fix x' y assume "(x', y) \ space P" + "f (x', y) \ (\y. f (x, y)) ` space M2" + then have *: "?F' x (f (x', y)) = {}" + by (force simp: space_pair_measure) + show "f (x', y) * M2.\ (?F' x (f (x', y))) = 0" + unfolding * by simp + qed (simp add: vimage_compose[symmetric] comp_def + space_pair_measure) } + note eq = this + moreover have "\z. ?F z \ sets P" + by (auto intro!: f_borel borel_measurable_vimage simp del: space_sigma) + moreover then have "\z. (\x. M2.\ (?F' x z)) \ borel_measurable M1" + by (auto intro!: measure_cut_measurable_fst simp del: vimage_Int) + moreover have *: "\i x. 0 \ M2.\ (Pair x -` (f -` {i} \ space P))" + using f(1)[THEN simple_functionD(2)] f(2) by (intro M2.positive_measure measurable_cut_fst) + moreover { fix i assume "i \ f`space P" + with * have "\x. 0 \ i * M2.\ (Pair x -` (f -` {i} \ space P))" + using f(2) by auto } + ultimately + show "(\x. \\<^isup>+y. f (x, y) \M2) \ borel_measurable M1" + and "(\\<^isup>+ x. (\\<^isup>+ y. f (x, y) \M2) \M1) = integral\<^isup>P P f" using f(2) + by (auto simp del: vimage_Int cong: measurable_cong + intro!: M1.borel_measurable_extreal_setsum setsum_cong + simp add: M1.positive_integral_setsum simple_integral_def + M1.positive_integral_cmult + M1.positive_integral_cong[OF eq] + positive_integral_eq_simple_integral[OF f] + pair_measure_alt[symmetric]) +qed + +lemma (in pair_sigma_finite) positive_integral_fst_measurable: + assumes f: "f \ borel_measurable P" + shows "(\x. \\<^isup>+ y. f (x, y) \M2) \ borel_measurable M1" + (is "?C f \ borel_measurable M1") + and "(\\<^isup>+ x. (\\<^isup>+ y. f (x, y) \M2) \M1) = integral\<^isup>P P f" +proof - + from borel_measurable_implies_simple_function_sequence'[OF f] guess F . note F = this + then have F_borel: "\i. F i \ borel_measurable P" + by (auto intro: borel_measurable_simple_function) + note sf = simple_function_cut[OF F(1,5)] + then have "(\x. SUP i. ?C (F i) x) \ borel_measurable M1" + using F(1) by auto + moreover + { fix x assume "x \ space M1" + from F measurable_pair_image_snd[OF F_borel`x \ space M1`] + have "(\\<^isup>+y. (SUP i. F i (x, y)) \M2) = (SUP i. ?C (F i) x)" + by (intro M2.positive_integral_monotone_convergence_SUP) + (auto simp: incseq_Suc_iff le_fun_def) + then have "(SUP i. ?C (F i) x) = ?C f x" + unfolding F(4) positive_integral_max_0 by simp } + note SUPR_C = this + ultimately show "?C f \ borel_measurable M1" + by (simp cong: measurable_cong) + have "(\\<^isup>+x. (SUP i. F i x) \P) = (SUP i. integral\<^isup>P P (F i))" + using F_borel F + by (intro positive_integral_monotone_convergence_SUP) auto + also have "(SUP i. integral\<^isup>P P (F i)) = (SUP i. \\<^isup>+ x. (\\<^isup>+ y. F i (x, y) \M2) \M1)" + unfolding sf(2) by simp + also have "\ = \\<^isup>+ x. (SUP i. \\<^isup>+ y. F i (x, y) \M2) \M1" using F sf(1) + by (intro M1.positive_integral_monotone_convergence_SUP[symmetric]) + (auto intro!: M2.positive_integral_mono M2.positive_integral_positive + simp: incseq_Suc_iff le_fun_def) + also have "\ = \\<^isup>+ x. (\\<^isup>+ y. (SUP i. F i (x, y)) \M2) \M1" + using F_borel F(2,5) + by (auto intro!: M1.positive_integral_cong M2.positive_integral_monotone_convergence_SUP[symmetric] + simp: incseq_Suc_iff le_fun_def measurable_pair_image_snd) + finally show "(\\<^isup>+ x. (\\<^isup>+ y. f (x, y) \M2) \M1) = integral\<^isup>P P f" + using F by (simp add: positive_integral_max_0) +qed + +lemma (in pair_sigma_finite) measure_preserving_swap: + "(\(x,y). (y, x)) \ measure_preserving (M1 \\<^isub>M M2) (M2 \\<^isub>M M1)" +proof + interpret Q: pair_sigma_finite M2 M1 by default + show *: "(\(x,y). (y, x)) \ measurable (M1 \\<^isub>M M2) (M2 \\<^isub>M M1)" + using pair_sigma_algebra_swap_measurable . + fix X assume "X \ sets (M2 \\<^isub>M M1)" + from measurable_sets[OF * this] this Q.sets_into_space[OF this] + show "measure (M1 \\<^isub>M M2) ((\(x, y). (y, x)) -` X \ space P) = measure (M2 \\<^isub>M M1) X" + by (auto intro!: M1.positive_integral_cong arg_cong[where f="M2.\"] + simp: pair_measure_alt Q.pair_measure_alt2 space_pair_measure) +qed + +lemma (in pair_sigma_finite) positive_integral_product_swap: + assumes f: "f \ borel_measurable P" + shows "(\\<^isup>+x. f (case x of (x,y)\(y,x)) \(M2 \\<^isub>M M1)) = integral\<^isup>P P f" +proof - + interpret Q: pair_sigma_finite M2 M1 by default + have "sigma_algebra P" by default + with f show ?thesis + by (subst Q.positive_integral_vimage[OF _ Q.measure_preserving_swap]) auto +qed + +lemma (in pair_sigma_finite) positive_integral_snd_measurable: + assumes f: "f \ borel_measurable P" + shows "(\\<^isup>+ y. (\\<^isup>+ x. f (x, y) \M1) \M2) = integral\<^isup>P P f" +proof - + interpret Q: pair_sigma_finite M2 M1 by default + note pair_sigma_algebra_measurable[OF f] + from Q.positive_integral_fst_measurable[OF this] + have "(\\<^isup>+ y. (\\<^isup>+ x. f (x, y) \M1) \M2) = (\\<^isup>+ (x, y). f (y, x) \Q.P)" + by simp + also have "(\\<^isup>+ (x, y). f (y, x) \Q.P) = integral\<^isup>P P f" + unfolding positive_integral_product_swap[OF f, symmetric] + by (auto intro!: Q.positive_integral_cong) + finally show ?thesis . +qed + +lemma (in pair_sigma_finite) Fubini: + assumes f: "f \ borel_measurable P" + shows "(\\<^isup>+ y. (\\<^isup>+ x. f (x, y) \M1) \M2) = (\\<^isup>+ x. (\\<^isup>+ y. f (x, y) \M2) \M1)" + unfolding positive_integral_snd_measurable[OF assms] + unfolding positive_integral_fst_measurable[OF assms] .. + +lemma (in pair_sigma_finite) AE_pair: + assumes "AE x in P. Q x" + shows "AE x in M1. (AE y in M2. Q (x, y))" +proof - + obtain N where N: "N \ sets P" "\ N = 0" "{x\space P. \ Q x} \ N" + using assms unfolding almost_everywhere_def by auto + show ?thesis + proof (rule M1.AE_I) + from N measure_cut_measurable_fst[OF `N \ sets P`] + show "M1.\ {x\space M1. M2.\ (Pair x -` N) \ 0} = 0" + by (auto simp: pair_measure_alt M1.positive_integral_0_iff) + show "{x \ space M1. M2.\ (Pair x -` N) \ 0} \ sets M1" + by (intro M1.borel_measurable_extreal_neq_const measure_cut_measurable_fst N) + { fix x assume "x \ space M1" "M2.\ (Pair x -` N) = 0" + have "M2.almost_everywhere (\y. Q (x, y))" + proof (rule M2.AE_I) + show "M2.\ (Pair x -` N) = 0" by fact + show "Pair x -` N \ sets M2" by (intro measurable_cut_fst N) + show "{y \ space M2. \ Q (x, y)} \ Pair x -` N" + using N `x \ space M1` unfolding space_sigma space_pair_measure by auto + qed } + then show "{x \ space M1. \ M2.almost_everywhere (\y. Q (x, y))} \ {x \ space M1. M2.\ (Pair x -` N) \ 0}" + by auto + qed +qed + +lemma (in pair_sigma_algebra) measurable_product_swap: + "f \ measurable (M2 \\<^isub>M M1) M \ (\(x,y). f (y,x)) \ measurable P M" +proof - + interpret Q: pair_sigma_algebra M2 M1 by default + show ?thesis + using pair_sigma_algebra_measurable[of "\(x,y). f (y, x)"] + by (auto intro!: pair_sigma_algebra_measurable Q.pair_sigma_algebra_measurable iffI) +qed + +lemma (in pair_sigma_finite) integrable_product_swap: + assumes "integrable P f" + shows "integrable (M2 \\<^isub>M M1) (\(x,y). f (y,x))" +proof - + interpret Q: pair_sigma_finite M2 M1 by default + have *: "(\(x,y). f (y,x)) = (\x. f (case x of (x,y)\(y,x)))" by (auto simp: fun_eq_iff) + show ?thesis unfolding * + using assms unfolding integrable_def + apply (subst (1 2) positive_integral_product_swap) + using `integrable P f` unfolding integrable_def + by (auto simp: *[symmetric] Q.measurable_product_swap[symmetric]) +qed + +lemma (in pair_sigma_finite) integrable_product_swap_iff: + "integrable (M2 \\<^isub>M M1) (\(x,y). f (y,x)) \ integrable P f" +proof - + interpret Q: pair_sigma_finite M2 M1 by default + from Q.integrable_product_swap[of "\(x,y). f (y,x)"] integrable_product_swap[of f] + show ?thesis by auto +qed + +lemma (in pair_sigma_finite) integral_product_swap: + assumes "integrable P f" + shows "(\(x,y). f (y,x) \(M2 \\<^isub>M M1)) = integral\<^isup>L P f" +proof - + interpret Q: pair_sigma_finite M2 M1 by default + have *: "(\(x,y). f (y,x)) = (\x. f (case x of (x,y)\(y,x)))" by (auto simp: fun_eq_iff) + show ?thesis + unfolding lebesgue_integral_def * + apply (subst (1 2) positive_integral_product_swap) + using `integrable P f` unfolding integrable_def + by (auto simp: *[symmetric] Q.measurable_product_swap[symmetric]) +qed + +lemma (in pair_sigma_finite) integrable_fst_measurable: + assumes f: "integrable P f" + shows "M1.almost_everywhere (\x. integrable M2 (\ y. f (x, y)))" (is "?AE") + and "(\x. (\y. f (x, y) \M2) \M1) = integral\<^isup>L P f" (is "?INT") +proof - + let "?pf x" = "extreal (f x)" and "?nf x" = "extreal (- f x)" + have + borel: "?nf \ borel_measurable P""?pf \ borel_measurable P" and + int: "integral\<^isup>P P ?nf \ \" "integral\<^isup>P P ?pf \ \" + using assms by auto + have "(\\<^isup>+x. (\\<^isup>+y. extreal (f (x, y)) \M2) \M1) \ \" + "(\\<^isup>+x. (\\<^isup>+y. extreal (- f (x, y)) \M2) \M1) \ \" + using borel[THEN positive_integral_fst_measurable(1)] int + unfolding borel[THEN positive_integral_fst_measurable(2)] by simp_all + with borel[THEN positive_integral_fst_measurable(1)] + have AE_pos: "AE x in M1. (\\<^isup>+y. extreal (f (x, y)) \M2) \ \" + "AE x in M1. (\\<^isup>+y. extreal (- f (x, y)) \M2) \ \" + by (auto intro!: M1.positive_integral_PInf_AE ) + then have AE: "AE x in M1. \\\<^isup>+y. extreal (f (x, y)) \M2\ \ \" + "AE x in M1. \\\<^isup>+y. extreal (- f (x, y)) \M2\ \ \" + by (auto simp: M2.positive_integral_positive) + from AE_pos show ?AE using assms + by (simp add: measurable_pair_image_snd integrable_def) + { fix f have "(\\<^isup>+ x. - \\<^isup>+ y. extreal (f x y) \M2 \M1) = (\\<^isup>+x. 0 \M1)" + using M2.positive_integral_positive + by (intro M1.positive_integral_cong_pos) (auto simp: extreal_uminus_le_reorder) + then have "(\\<^isup>+ x. - \\<^isup>+ y. extreal (f x y) \M2 \M1) = 0" by simp } + note this[simp] + { fix f assume borel: "(\x. extreal (f x)) \ borel_measurable P" + and int: "integral\<^isup>P P (\x. extreal (f x)) \ \" + and AE: "M1.almost_everywhere (\x. (\\<^isup>+y. extreal (f (x, y)) \M2) \ \)" + have "integrable M1 (\x. real (\\<^isup>+y. extreal (f (x, y)) \M2))" (is "integrable M1 ?f") + proof (intro integrable_def[THEN iffD2] conjI) + show "?f \ borel_measurable M1" + using borel by (auto intro!: M1.borel_measurable_real_of_extreal positive_integral_fst_measurable) + have "(\\<^isup>+x. extreal (?f x) \M1) = (\\<^isup>+x. (\\<^isup>+y. extreal (f (x, y)) \M2) \M1)" + using AE M2.positive_integral_positive + by (auto intro!: M1.positive_integral_cong_AE simp: extreal_real) + then show "(\\<^isup>+x. extreal (?f x) \M1) \ \" + using positive_integral_fst_measurable[OF borel] int by simp + have "(\\<^isup>+x. extreal (- ?f x) \M1) = (\\<^isup>+x. 0 \M1)" + by (intro M1.positive_integral_cong_pos) + (simp add: M2.positive_integral_positive real_of_extreal_pos) + then show "(\\<^isup>+x. extreal (- ?f x) \M1) \ \" by simp + qed } + with this[OF borel(1) int(1) AE_pos(2)] this[OF borel(2) int(2) AE_pos(1)] + show ?INT + unfolding lebesgue_integral_def[of P] lebesgue_integral_def[of M2] + borel[THEN positive_integral_fst_measurable(2), symmetric] + using AE[THEN M1.integral_real] + by simp +qed + +lemma (in pair_sigma_finite) integrable_snd_measurable: + assumes f: "integrable P f" + shows "M2.almost_everywhere (\y. integrable M1 (\x. f (x, y)))" (is "?AE") + and "(\y. (\x. f (x, y) \M1) \M2) = integral\<^isup>L P f" (is "?INT") +proof - + interpret Q: pair_sigma_finite M2 M1 by default + have Q_int: "integrable Q.P (\(x, y). f (y, x))" + using f unfolding integrable_product_swap_iff . + show ?INT + using Q.integrable_fst_measurable(2)[OF Q_int] + using integral_product_swap[OF f] by simp + show ?AE + using Q.integrable_fst_measurable(1)[OF Q_int] + by simp +qed + +lemma (in pair_sigma_finite) Fubini_integral: + assumes f: "integrable P f" + shows "(\y. (\x. f (x, y) \M1) \M2) = (\x. (\y. f (x, y) \M2) \M1)" + unfolding integrable_snd_measurable[OF assms] + unfolding integrable_fst_measurable[OF assms] .. + +section "Products on finite spaces" + +lemma sigma_sets_pair_measure_generator_finite: + assumes "finite A" and "finite B" + shows "sigma_sets (A \ B) { a \ b | a b. a \ Pow A \ b \ Pow B} = Pow (A \ B)" + (is "sigma_sets ?prod ?sets = _") +proof safe + have fin: "finite (A \ B)" using assms by (rule finite_cartesian_product) + fix x assume subset: "x \ A \ B" + hence "finite x" using fin by (rule finite_subset) + from this subset show "x \ sigma_sets ?prod ?sets" + proof (induct x) + case empty show ?case by (rule sigma_sets.Empty) + next + case (insert a x) + hence "{a} \ sigma_sets ?prod ?sets" + by (auto simp: pair_measure_generator_def intro!: sigma_sets.Basic) + moreover have "x \ sigma_sets ?prod ?sets" using insert by auto + ultimately show ?case unfolding insert_is_Un[of a x] by (rule sigma_sets_Un) + qed +next + fix x a b + assume "x \ sigma_sets ?prod ?sets" and "(a, b) \ x" + from sigma_sets_into_sp[OF _ this(1)] this(2) + show "a \ A" and "b \ B" by auto +qed + +locale pair_finite_sigma_algebra = M1: finite_sigma_algebra M1 + M2: finite_sigma_algebra M2 + for M1 :: "('a, 'c) measure_space_scheme" and M2 :: "('b, 'd) measure_space_scheme" + +sublocale pair_finite_sigma_algebra \ pair_sigma_algebra by default + +lemma (in pair_finite_sigma_algebra) finite_pair_sigma_algebra: + shows "P = \ space = space M1 \ space M2, sets = Pow (space M1 \ space M2), \ = algebra.more P \" +proof - + show ?thesis + using sigma_sets_pair_measure_generator_finite[OF M1.finite_space M2.finite_space] + by (intro algebra.equality) (simp_all add: pair_measure_def pair_measure_generator_def sigma_def) +qed + +sublocale pair_finite_sigma_algebra \ finite_sigma_algebra P + apply default + using M1.finite_space M2.finite_space + apply (subst finite_pair_sigma_algebra) apply simp + apply (subst (1 2) finite_pair_sigma_algebra) apply simp + done + +locale pair_finite_space = M1: finite_measure_space M1 + M2: finite_measure_space M2 + for M1 M2 + +sublocale pair_finite_space \ pair_finite_sigma_algebra + by default + +sublocale pair_finite_space \ pair_sigma_finite + by default + +lemma (in pair_finite_space) pair_measure_Pair[simp]: + assumes "a \ space M1" "b \ space M2" + shows "\ {(a, b)} = M1.\ {a} * M2.\ {b}" +proof - + have "\ ({a}\{b}) = M1.\ {a} * M2.\ {b}" + using M1.sets_eq_Pow M2.sets_eq_Pow assms + by (subst pair_measure_times) auto + then show ?thesis by simp +qed + +lemma (in pair_finite_space) pair_measure_singleton[simp]: + assumes "x \ space M1 \ space M2" + shows "\ {x} = M1.\ {fst x} * M2.\ {snd x}" + using pair_measure_Pair assms by (cases x) auto + +sublocale pair_finite_space \ finite_measure_space P + by default (auto simp: space_pair_measure) + +end \ No newline at end of file diff -r 8448713d48b7 -r 5b52c6a9c627 src/HOL/Probability/Complete_Measure.thy --- a/src/HOL/Probability/Complete_Measure.thy Tue Mar 29 14:27:31 2011 +0200 +++ b/src/HOL/Probability/Complete_Measure.thy Tue Mar 29 14:27:39 2011 +0200 @@ -3,7 +3,7 @@ *) theory Complete_Measure -imports Product_Measure +imports Lebesgue_Integration begin locale completeable_measure_space = measure_space diff -r 8448713d48b7 -r 5b52c6a9c627 src/HOL/Probability/Finite_Product_Measure.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Probability/Finite_Product_Measure.thy Tue Mar 29 14:27:39 2011 +0200 @@ -0,0 +1,1014 @@ +(* Title: HOL/Probability/Finite_Product_Measure.thy + Author: Johannes Hölzl, TU München +*) + +header {*Finite product measures*} + +theory Finite_Product_Measure +imports Binary_Product_Measure +begin + +lemma Pi_iff: "f \ Pi I X \ (\i\I. f i \ X i)" + unfolding Pi_def by auto + +abbreviation + "Pi\<^isub>E A B \ Pi A B \ extensional A" + +syntax + "_PiE" :: "[pttrn, 'a set, 'b set] => ('a => 'b) set" ("(3PIE _:_./ _)" 10) + +syntax (xsymbols) + "_PiE" :: "[pttrn, 'a set, 'b set] => ('a => 'b) set" ("(3\\<^isub>E _\_./ _)" 10) + +syntax (HTML output) + "_PiE" :: "[pttrn, 'a set, 'b set] => ('a => 'b) set" ("(3\\<^isub>E _\_./ _)" 10) + +translations + "PIE x:A. B" == "CONST Pi\<^isub>E A (%x. B)" + +abbreviation + funcset_extensional :: "['a set, 'b set] => ('a => 'b) set" + (infixr "->\<^isub>E" 60) where + "A ->\<^isub>E B \ Pi\<^isub>E A (%_. B)" + +notation (xsymbols) + funcset_extensional (infixr "\\<^isub>E" 60) + +lemma extensional_empty[simp]: "extensional {} = {\x. undefined}" + by safe (auto simp add: extensional_def fun_eq_iff) + +lemma extensional_insert[intro, simp]: + assumes "a \ extensional (insert i I)" + shows "a(i := b) \ extensional (insert i I)" + using assms unfolding extensional_def by auto + +lemma extensional_Int[simp]: + "extensional I \ extensional I' = extensional (I \ I')" + unfolding extensional_def by auto + +definition + "merge I x J y = (\i. if i \ I then x i else if i \ J then y i else undefined)" + +lemma merge_apply[simp]: + "I \ J = {} \ i \ I \ merge I x J y i = x i" + "I \ J = {} \ i \ J \ merge I x J y i = y i" + "J \ I = {} \ i \ I \ merge I x J y i = x i" + "J \ I = {} \ i \ J \ merge I x J y i = y i" + "i \ I \ i \ J \ merge I x J y i = undefined" + unfolding merge_def by auto + +lemma merge_commute: + "I \ J = {} \ merge I x J y = merge J y I x" + by (auto simp: merge_def intro!: ext) + +lemma Pi_cancel_merge_range[simp]: + "I \ J = {} \ x \ Pi I (merge I A J B) \ x \ Pi I A" + "I \ J = {} \ x \ Pi I (merge J B I A) \ x \ Pi I A" + "J \ I = {} \ x \ Pi I (merge I A J B) \ x \ Pi I A" + "J \ I = {} \ x \ Pi I (merge J B I A) \ x \ Pi I A" + by (auto simp: Pi_def) + +lemma Pi_cancel_merge[simp]: + "I \ J = {} \ merge I x J y \ Pi I B \ x \ Pi I B" + "J \ I = {} \ merge I x J y \ Pi I B \ x \ Pi I B" + "I \ J = {} \ merge I x J y \ Pi J B \ y \ Pi J B" + "J \ I = {} \ merge I x J y \ Pi J B \ y \ Pi J B" + by (auto simp: Pi_def) + +lemma extensional_merge[simp]: "merge I x J y \ extensional (I \ J)" + by (auto simp: extensional_def) + +lemma restrict_Pi_cancel: "restrict x I \ Pi I A \ x \ Pi I A" + by (auto simp: restrict_def Pi_def) + +lemma restrict_merge[simp]: + "I \ J = {} \ restrict (merge I x J y) I = restrict x I" + "I \ J = {} \ restrict (merge I x J y) J = restrict y J" + "J \ I = {} \ restrict (merge I x J y) I = restrict x I" + "J \ I = {} \ restrict (merge I x J y) J = restrict y J" + by (auto simp: restrict_def intro!: ext) + +lemma extensional_insert_undefined[intro, simp]: + assumes "a \ extensional (insert i I)" + shows "a(i := undefined) \ extensional I" + using assms unfolding extensional_def by auto + +lemma extensional_insert_cancel[intro, simp]: + assumes "a \ extensional I" + shows "a \ extensional (insert i I)" + using assms unfolding extensional_def by auto + +lemma merge_singleton[simp]: "i \ I \ merge I x {i} y = restrict (x(i := y i)) (insert i I)" + unfolding merge_def by (auto simp: fun_eq_iff) + +lemma Pi_Int: "Pi I E \ Pi I F = (\ i\I. E i \ F i)" + by auto + +lemma PiE_Int: "(Pi\<^isub>E I A) \ (Pi\<^isub>E I B) = Pi\<^isub>E I (\x. A x \ B x)" + by auto + +lemma Pi_cancel_fupd_range[simp]: "i \ I \ x \ Pi I (B(i := b)) \ x \ Pi I B" + by (auto simp: Pi_def) + +lemma Pi_split_insert_domain[simp]: "x \ Pi (insert i I) X \ x \ Pi I X \ x i \ X i" + by (auto simp: Pi_def) + +lemma Pi_split_domain[simp]: "x \ Pi (I \ J) X \ x \ Pi I X \ x \ Pi J X" + by (auto simp: Pi_def) + +lemma Pi_cancel_fupd[simp]: "i \ I \ x(i := a) \ Pi I B \ x \ Pi I B" + by (auto simp: Pi_def) + +lemma restrict_vimage: + assumes "I \ J = {}" + shows "(\x. (restrict x I, restrict x J)) -` (Pi\<^isub>E I E \ Pi\<^isub>E J F) = Pi (I \ J) (merge I E J F)" + using assms by (auto simp: restrict_Pi_cancel) + +lemma merge_vimage: + assumes "I \ J = {}" + shows "(\(x,y). merge I x J y) -` Pi\<^isub>E (I \ J) E = Pi I E \ Pi J E" + using assms by (auto simp: restrict_Pi_cancel) + +lemma restrict_fupd[simp]: "i \ I \ restrict (f (i := x)) I = restrict f I" + by (auto simp: restrict_def intro!: ext) + +lemma merge_restrict[simp]: + "merge I (restrict x I) J y = merge I x J y" + "merge I x J (restrict y J) = merge I x J y" + unfolding merge_def by (auto intro!: ext) + +lemma merge_x_x_eq_restrict[simp]: + "merge I x J x = restrict x (I \ J)" + unfolding merge_def by (auto intro!: ext) + +lemma Pi_fupd_iff: "i \ I \ f \ Pi I (B(i := A)) \ f \ Pi (I - {i}) B \ f i \ A" + apply auto + apply (drule_tac x=x in Pi_mem) + apply (simp_all split: split_if_asm) + apply (drule_tac x=i in Pi_mem) + apply (auto dest!: Pi_mem) + done + +lemma Pi_UN: + fixes A :: "nat \ 'i \ 'a set" + assumes "finite I" and mono: "\i n m. i \ I \ n \ m \ A n i \ A m i" + shows "(\n. Pi I (A n)) = (\ i\I. \n. A n i)" +proof (intro set_eqI iffI) + fix f assume "f \ (\ i\I. \n. A n i)" + then have "\i\I. \n. f i \ A n i" by auto + from bchoice[OF this] obtain n where n: "\i. i \ I \ f i \ (A (n i) i)" by auto + obtain k where k: "\i. i \ I \ n i \ k" + using `finite I` finite_nat_set_iff_bounded_le[of "n`I"] by auto + have "f \ Pi I (A k)" + proof (intro Pi_I) + fix i assume "i \ I" + from mono[OF this, of "n i" k] k[OF this] n[OF this] + show "f i \ A k i" by auto + qed + then show "f \ (\n. Pi I (A n))" by auto +qed auto + +lemma PiE_cong: + assumes "\i. i\I \ A i = B i" + shows "Pi\<^isub>E I A = Pi\<^isub>E I B" + using assms by (auto intro!: Pi_cong) + +lemma restrict_upd[simp]: + "i \ I \ (restrict f I)(i := y) = restrict (f(i := y)) (insert i I)" + by (auto simp: fun_eq_iff) + +lemma Pi_eq_subset: + assumes ne: "\i. i \ I \ F i \ {}" "\i. i \ I \ F' i \ {}" + assumes eq: "Pi\<^isub>E I F = Pi\<^isub>E I F'" and "i \ I" + shows "F i \ F' i" +proof + fix x assume "x \ F i" + with ne have "\j. \y. ((j \ I \ y \ F j \ (i = j \ x = y)) \ (j \ I \ y = undefined))" by auto + from choice[OF this] guess f .. note f = this + then have "f \ Pi\<^isub>E I F" by (auto simp: extensional_def) + then have "f \ Pi\<^isub>E I F'" using assms by simp + then show "x \ F' i" using f `i \ I` by auto +qed + +lemma Pi_eq_iff_not_empty: + assumes ne: "\i. i \ I \ F i \ {}" "\i. i \ I \ F' i \ {}" + shows "Pi\<^isub>E I F = Pi\<^isub>E I F' \ (\i\I. F i = F' i)" +proof (intro iffI ballI) + fix i assume eq: "Pi\<^isub>E I F = Pi\<^isub>E I F'" and i: "i \ I" + show "F i = F' i" + using Pi_eq_subset[of I F F', OF ne eq i] + using Pi_eq_subset[of I F' F, OF ne(2,1) eq[symmetric] i] + by auto +qed auto + +lemma Pi_eq_empty_iff: + "Pi\<^isub>E I F = {} \ (\i\I. F i = {})" +proof + assume "Pi\<^isub>E I F = {}" + show "\i\I. F i = {}" + proof (rule ccontr) + assume "\ ?thesis" + then have "\i. \y. (i \ I \ y \ F i) \ (i \ I \ y = undefined)" by auto + from choice[OF this] guess f .. + then have "f \ Pi\<^isub>E I F" by (auto simp: extensional_def) + with `Pi\<^isub>E I F = {}` show False by auto + qed +qed auto + +lemma Pi_eq_iff: + "Pi\<^isub>E I F = Pi\<^isub>E I F' \ (\i\I. F i = F' i) \ ((\i\I. F i = {}) \ (\i\I. F' i = {}))" +proof (intro iffI disjCI) + assume eq[simp]: "Pi\<^isub>E I F = Pi\<^isub>E I F'" + assume "\ ((\i\I. F i = {}) \ (\i\I. F' i = {}))" + then have "(\i\I. F i \ {}) \ (\i\I. F' i \ {})" + using Pi_eq_empty_iff[of I F] Pi_eq_empty_iff[of I F'] by auto + with Pi_eq_iff_not_empty[of I F F'] show "\i\I. F i = F' i" by auto +next + assume "(\i\I. F i = F' i) \ (\i\I. F i = {}) \ (\i\I. F' i = {})" + then show "Pi\<^isub>E I F = Pi\<^isub>E I F'" + using Pi_eq_empty_iff[of I F] Pi_eq_empty_iff[of I F'] by auto +qed + +section "Finite product spaces" + +section "Products" + +locale product_sigma_algebra = + fixes M :: "'i \ ('a, 'b) measure_space_scheme" + assumes sigma_algebras: "\i. sigma_algebra (M i)" + +locale finite_product_sigma_algebra = product_sigma_algebra M + for M :: "'i \ ('a, 'b) measure_space_scheme" + + fixes I :: "'i set" + assumes finite_index: "finite I" + +definition + "product_algebra_generator I M = \ space = (\\<^isub>E i \ I. space (M i)), + sets = Pi\<^isub>E I ` (\ i \ I. sets (M i)), + measure = \A. (\i\I. measure (M i) ((SOME F. A = Pi\<^isub>E I F) i)) \" + +definition product_algebra_def: + "Pi\<^isub>M I M = sigma (product_algebra_generator I M) + \ measure := (SOME \. sigma_finite_measure (sigma (product_algebra_generator I M) \ measure := \ \) \ + (\A\\ i\I. sets (M i). \ (Pi\<^isub>E I A) = (\i\I. measure (M i) (A i))))\" + +syntax + "_PiM" :: "[pttrn, 'i set, ('a, 'b) measure_space_scheme] => + ('i => 'a, 'b) measure_space_scheme" ("(3PIM _:_./ _)" 10) + +syntax (xsymbols) + "_PiM" :: "[pttrn, 'i set, ('a, 'b) measure_space_scheme] => + ('i => 'a, 'b) measure_space_scheme" ("(3\\<^isub>M _\_./ _)" 10) + +syntax (HTML output) + "_PiM" :: "[pttrn, 'i set, ('a, 'b) measure_space_scheme] => + ('i => 'a, 'b) measure_space_scheme" ("(3\\<^isub>M _\_./ _)" 10) + +translations + "PIM x:I. M" == "CONST Pi\<^isub>M I (%x. M)" + +abbreviation (in finite_product_sigma_algebra) "G \ product_algebra_generator I M" +abbreviation (in finite_product_sigma_algebra) "P \ Pi\<^isub>M I M" + +sublocale product_sigma_algebra \ M: sigma_algebra "M i" for i by (rule sigma_algebras) + +lemma sigma_into_space: + assumes "sets M \ Pow (space M)" + shows "sets (sigma M) \ Pow (space M)" + using sigma_sets_into_sp[OF assms] unfolding sigma_def by auto + +lemma (in product_sigma_algebra) product_algebra_generator_into_space: + "sets (product_algebra_generator I M) \ Pow (space (product_algebra_generator I M))" + using M.sets_into_space unfolding product_algebra_generator_def + by auto blast + +lemma (in product_sigma_algebra) product_algebra_into_space: + "sets (Pi\<^isub>M I M) \ Pow (space (Pi\<^isub>M I M))" + using product_algebra_generator_into_space + by (auto intro!: sigma_into_space simp add: product_algebra_def) + +lemma (in product_sigma_algebra) sigma_algebra_product_algebra: "sigma_algebra (Pi\<^isub>M I M)" + using product_algebra_generator_into_space unfolding product_algebra_def + by (rule sigma_algebra.sigma_algebra_cong[OF sigma_algebra_sigma]) simp_all + +sublocale finite_product_sigma_algebra \ sigma_algebra P + using sigma_algebra_product_algebra . + +lemma product_algebraE: + assumes "A \ sets (product_algebra_generator I M)" + obtains E where "A = Pi\<^isub>E I E" "E \ (\ i\I. sets (M i))" + using assms unfolding product_algebra_generator_def by auto + +lemma product_algebra_generatorI[intro]: + assumes "E \ (\ i\I. sets (M i))" + shows "Pi\<^isub>E I E \ sets (product_algebra_generator I M)" + using assms unfolding product_algebra_generator_def by auto + +lemma space_product_algebra_generator[simp]: + "space (product_algebra_generator I M) = Pi\<^isub>E I (\i. space (M i))" + unfolding product_algebra_generator_def by simp + +lemma space_product_algebra[simp]: + "space (Pi\<^isub>M I M) = (\\<^isub>E i\I. space (M i))" + unfolding product_algebra_def product_algebra_generator_def by simp + +lemma sets_product_algebra: + "sets (Pi\<^isub>M I M) = sets (sigma (product_algebra_generator I M))" + unfolding product_algebra_def sigma_def by simp + +lemma product_algebra_generator_sets_into_space: + assumes "\i. i\I \ sets (M i) \ Pow (space (M i))" + shows "sets (product_algebra_generator I M) \ Pow (space (product_algebra_generator I M))" + using assms by (auto simp: product_algebra_generator_def) blast + +lemma (in finite_product_sigma_algebra) in_P[simp, intro]: + "\ \i. i \ I \ A i \ sets (M i) \ \ Pi\<^isub>E I A \ sets P" + by (auto simp: sets_product_algebra) + +section "Generating set generates also product algebra" + +lemma sigma_product_algebra_sigma_eq: + assumes "finite I" + assumes mono: "\i. i \ I \ incseq (S i)" + assumes union: "\i. i \ I \ (\j. S i j) = space (E i)" + assumes sets_into: "\i. i \ I \ range (S i) \ sets (E i)" + and E: "\i. sets (E i) \ Pow (space (E i))" + shows "sets (\\<^isub>M i\I. sigma (E i)) = sets (\\<^isub>M i\I. E i)" + (is "sets ?S = sets ?E") +proof cases + assume "I = {}" then show ?thesis + by (simp add: product_algebra_def product_algebra_generator_def) +next + assume "I \ {}" + interpret E: sigma_algebra "sigma (E i)" for i + using E by (rule sigma_algebra_sigma) + have into_space[intro]: "\i x A. A \ sets (E i) \ x i \ A \ x i \ space (E i)" + using E by auto + interpret G: sigma_algebra ?E + unfolding product_algebra_def product_algebra_generator_def using E + by (intro sigma_algebra.sigma_algebra_cong[OF sigma_algebra_sigma]) (auto dest: Pi_mem) + { fix A i assume "i \ I" and A: "A \ sets (E i)" + then have "(\x. x i) -` A \ space ?E = (\\<^isub>E j\I. if j = i then A else \n. S j n) \ space ?E" + using mono union unfolding incseq_Suc_iff space_product_algebra + by (auto dest: Pi_mem) + also have "\ = (\n. (\\<^isub>E j\I. if j = i then A else S j n))" + unfolding space_product_algebra + apply simp + apply (subst Pi_UN[OF `finite I`]) + using mono[THEN incseqD] apply simp + apply (simp add: PiE_Int) + apply (intro PiE_cong) + using A sets_into by (auto intro!: into_space) + also have "\ \ sets ?E" + using sets_into `A \ sets (E i)` + unfolding sets_product_algebra sets_sigma + by (intro sigma_sets.Union) + (auto simp: image_subset_iff intro!: sigma_sets.Basic) + finally have "(\x. x i) -` A \ space ?E \ sets ?E" . } + then have proj: + "\i. i\I \ (\x. x i) \ measurable ?E (sigma (E i))" + using E by (subst G.measurable_iff_sigma) + (auto simp: sets_product_algebra sets_sigma) + { fix A assume A: "\i. i \ I \ A i \ sets (sigma (E i))" + with proj have basic: "\i. i \ I \ (\x. x i) -` (A i) \ space ?E \ sets ?E" + unfolding measurable_def by simp + have "Pi\<^isub>E I A = (\i\I. (\x. x i) -` (A i) \ space ?E)" + using A E.sets_into_space `I \ {}` unfolding product_algebra_def by auto blast + then have "Pi\<^isub>E I A \ sets ?E" + using G.finite_INT[OF `finite I` `I \ {}` basic, of "\i. i"] by simp } + then have "sigma_sets (space ?E) (sets (product_algebra_generator I (\i. sigma (E i)))) \ sets ?E" + by (intro G.sigma_sets_subset) (auto simp add: product_algebra_generator_def) + then have subset: "sets ?S \ sets ?E" + by (simp add: sets_sigma sets_product_algebra) + show "sets ?S = sets ?E" + proof (intro set_eqI iffI) + fix A assume "A \ sets ?E" then show "A \ sets ?S" + unfolding sets_sigma sets_product_algebra + proof induct + case (Basic A) then show ?case + by (auto simp: sets_sigma product_algebra_generator_def intro: sigma_sets.Basic) + qed (auto intro: sigma_sets.intros simp: product_algebra_generator_def) + next + fix A assume "A \ sets ?S" then show "A \ sets ?E" using subset by auto + qed +qed + +lemma product_algebraI[intro]: + "E \ (\ i\I. sets (M i)) \ Pi\<^isub>E I E \ sets (Pi\<^isub>M I M)" + using assms unfolding product_algebra_def by (auto intro: product_algebra_generatorI) + +lemma (in product_sigma_algebra) measurable_component_update: + assumes "x \ space (Pi\<^isub>M I M)" and "i \ I" + shows "(\v. x(i := v)) \ measurable (M i) (Pi\<^isub>M (insert i I) M)" (is "?f \ _") + unfolding product_algebra_def apply simp +proof (intro measurable_sigma) + let ?G = "product_algebra_generator (insert i I) M" + show "sets ?G \ Pow (space ?G)" using product_algebra_generator_into_space . + show "?f \ space (M i) \ space ?G" + using M.sets_into_space assms by auto + fix A assume "A \ sets ?G" + from product_algebraE[OF this] guess E . note E = this + then have "?f -` A \ space (M i) = E i \ ?f -` A \ space (M i) = {}" + using M.sets_into_space assms by auto + then show "?f -` A \ space (M i) \ sets (M i)" + using E by (auto intro!: product_algebraI) +qed + +lemma (in product_sigma_algebra) measurable_add_dim: + assumes "i \ I" + shows "(\(f, y). f(i := y)) \ measurable (Pi\<^isub>M I M \\<^isub>M M i) (Pi\<^isub>M (insert i I) M)" +proof - + let ?f = "(\(f, y). f(i := y))" and ?G = "product_algebra_generator (insert i I) M" + interpret Ii: pair_sigma_algebra "Pi\<^isub>M I M" "M i" + unfolding pair_sigma_algebra_def + by (intro sigma_algebra_product_algebra sigma_algebras conjI) + have "?f \ measurable Ii.P (sigma ?G)" + proof (rule Ii.measurable_sigma) + show "sets ?G \ Pow (space ?G)" + using product_algebra_generator_into_space . + show "?f \ space (Pi\<^isub>M I M \\<^isub>M M i) \ space ?G" + by (auto simp: space_pair_measure) + next + fix A assume "A \ sets ?G" + then obtain F where "A = Pi\<^isub>E (insert i I) F" + and F: "\j. j \ I \ F j \ sets (M j)" "F i \ sets (M i)" + by (auto elim!: product_algebraE) + then have "?f -` A \ space (Pi\<^isub>M I M \\<^isub>M M i) = Pi\<^isub>E I F \ (F i)" + using sets_into_space `i \ I` + by (auto simp add: space_pair_measure) blast+ + then show "?f -` A \ space (Pi\<^isub>M I M \\<^isub>M M i) \ sets (Pi\<^isub>M I M \\<^isub>M M i)" + using F by (auto intro!: pair_measureI) + qed + then show ?thesis + by (simp add: product_algebra_def) +qed + +lemma (in product_sigma_algebra) measurable_merge: + assumes [simp]: "I \ J = {}" + shows "(\(x, y). merge I x J y) \ measurable (Pi\<^isub>M I M \\<^isub>M Pi\<^isub>M J M) (Pi\<^isub>M (I \ J) M)" +proof - + let ?I = "Pi\<^isub>M I M" and ?J = "Pi\<^isub>M J M" + interpret P: sigma_algebra "?I \\<^isub>M ?J" + by (intro sigma_algebra_pair_measure product_algebra_into_space) + let ?f = "\(x, y). merge I x J y" + let ?G = "product_algebra_generator (I \ J) M" + have "?f \ measurable (?I \\<^isub>M ?J) (sigma ?G)" + proof (rule P.measurable_sigma) + fix A assume "A \ sets ?G" + from product_algebraE[OF this] + obtain E where E: "A = Pi\<^isub>E (I \ J) E" "E \ (\ i\I \ J. sets (M i))" . + then have *: "?f -` A \ space (?I \\<^isub>M ?J) = Pi\<^isub>E I E \ Pi\<^isub>E J E" + using sets_into_space `I \ J = {}` + by (auto simp add: space_pair_measure) fast+ + then show "?f -` A \ space (?I \\<^isub>M ?J) \ sets (?I \\<^isub>M ?J)" + using E unfolding * by (auto intro!: pair_measureI in_sigma product_algebra_generatorI + simp: product_algebra_def) + qed (insert product_algebra_generator_into_space, auto simp: space_pair_measure) + then show "?f \ measurable (?I \\<^isub>M ?J) (Pi\<^isub>M (I \ J) M)" + unfolding product_algebra_def[of "I \ J"] by simp +qed + +lemma (in product_sigma_algebra) measurable_component_singleton: + assumes "i \ I" shows "(\x. x i) \ measurable (Pi\<^isub>M I M) (M i)" +proof (unfold measurable_def, intro CollectI conjI ballI) + fix A assume "A \ sets (M i)" + then have "(\x. x i) -` A \ space (Pi\<^isub>M I M) = (\\<^isub>E j\I. if i = j then A else space (M j))" + using M.sets_into_space `i \ I` by (fastsimp dest: Pi_mem split: split_if_asm) + then show "(\x. x i) -` A \ space (Pi\<^isub>M I M) \ sets (Pi\<^isub>M I M)" + using `A \ sets (M i)` by (auto intro!: product_algebraI) +qed (insert `i \ I`, auto) + +locale product_sigma_finite = + fixes M :: "'i \ ('a,'b) measure_space_scheme" + assumes sigma_finite_measures: "\i. sigma_finite_measure (M i)" + +locale finite_product_sigma_finite = product_sigma_finite M + for M :: "'i \ ('a,'b) measure_space_scheme" + + fixes I :: "'i set" assumes finite_index'[intro]: "finite I" + +sublocale product_sigma_finite \ M: sigma_finite_measure "M i" for i + by (rule sigma_finite_measures) + +sublocale product_sigma_finite \ product_sigma_algebra + by default + +sublocale finite_product_sigma_finite \ finite_product_sigma_algebra + by default (fact finite_index') + +lemma setprod_extreal_0: + fixes f :: "'a \ extreal" + shows "(\i\A. f i) = 0 \ (finite A \ (\i\A. f i = 0))" +proof cases + assume "finite A" + then show ?thesis by (induct A) auto +qed auto + +lemma setprod_extreal_pos: + fixes f :: "'a \ extreal" assumes pos: "\i. i \ I \ 0 \ f i" shows "0 \ (\i\I. f i)" +proof cases + assume "finite I" from this pos show ?thesis by induct auto +qed simp + +lemma setprod_PInf: + assumes "\i. i \ I \ 0 \ f i" + shows "(\i\I. f i) = \ \ finite I \ (\i\I. f i = \) \ (\i\I. f i \ 0)" +proof cases + assume "finite I" from this assms show ?thesis + proof (induct I) + case (insert i I) + then have pos: "0 \ f i" "0 \ setprod f I" by (auto intro!: setprod_extreal_pos) + from insert have "(\j\insert i I. f j) = \ \ setprod f I * f i = \" by auto + also have "\ \ (setprod f I = \ \ f i = \) \ f i \ 0 \ setprod f I \ 0" + using setprod_extreal_pos[of I f] pos + by (cases rule: extreal2_cases[of "f i" "setprod f I"]) auto + also have "\ \ finite (insert i I) \ (\j\insert i I. f j = \) \ (\j\insert i I. f j \ 0)" + using insert by (auto simp: setprod_extreal_0) + finally show ?case . + qed simp +qed simp + +lemma setprod_extreal: "(\i\A. extreal (f i)) = extreal (setprod f A)" +proof cases + assume "finite A" then show ?thesis + by induct (auto simp: one_extreal_def) +qed (simp add: one_extreal_def) + +lemma (in finite_product_sigma_finite) product_algebra_generator_measure: + assumes "Pi\<^isub>E I F \ sets G" + shows "measure G (Pi\<^isub>E I F) = (\i\I. M.\ i (F i))" +proof cases + assume ne: "\i\I. F i \ {}" + have "\i\I. (SOME F'. Pi\<^isub>E I F = Pi\<^isub>E I F') i = F i" + by (rule someI2[where P="\F'. Pi\<^isub>E I F = Pi\<^isub>E I F'"]) + (insert ne, auto simp: Pi_eq_iff) + then show ?thesis + unfolding product_algebra_generator_def by simp +next + assume empty: "\ (\j\I. F j \ {})" + then have "(\j\I. M.\ j (F j)) = 0" + by (auto simp: setprod_extreal_0 intro!: bexI) + moreover + have "\j\I. (SOME F'. Pi\<^isub>E I F = Pi\<^isub>E I F') j = {}" + by (rule someI2[where P="\F'. Pi\<^isub>E I F = Pi\<^isub>E I F'"]) + (insert empty, auto simp: Pi_eq_empty_iff[symmetric]) + then have "(\j\I. M.\ j ((SOME F'. Pi\<^isub>E I F = Pi\<^isub>E I F') j)) = 0" + by (auto simp: setprod_extreal_0 intro!: bexI) + ultimately show ?thesis + unfolding product_algebra_generator_def by simp +qed + +lemma (in finite_product_sigma_finite) sigma_finite_pairs: + "\F::'i \ nat \ 'a set. + (\i\I. range (F i) \ sets (M i)) \ + (\k. \i\I. \ i (F i k) \ \) \ incseq (\k. \\<^isub>E i\I. F i k) \ + (\k. \\<^isub>E i\I. F i k) = space G" +proof - + have "\i::'i. \F::nat \ 'a set. range F \ sets (M i) \ incseq F \ (\i. F i) = space (M i) \ (\k. \ i (F k) \ \)" + using M.sigma_finite_up by simp + from choice[OF this] guess F :: "'i \ nat \ 'a set" .. + then have F: "\i. range (F i) \ sets (M i)" "\i. incseq (F i)" "\i. (\j. F i j) = space (M i)" "\i k. \ i (F i k) \ \" + by auto + let ?F = "\k. \\<^isub>E i\I. F i k" + note space_product_algebra[simp] + show ?thesis + proof (intro exI[of _ F] conjI allI incseq_SucI set_eqI iffI ballI) + fix i show "range (F i) \ sets (M i)" by fact + next + fix i k show "\ i (F i k) \ \" by fact + next + fix A assume "A \ (\i. ?F i)" then show "A \ space G" + using `\i. range (F i) \ sets (M i)` M.sets_into_space + by (force simp: image_subset_iff) + next + fix f assume "f \ space G" + with Pi_UN[OF finite_index, of "\k i. F i k"] F + show "f \ (\i. ?F i)" by (auto simp: incseq_def) + next + fix i show "?F i \ ?F (Suc i)" + using `\i. incseq (F i)`[THEN incseq_SucD] by auto + qed +qed + +lemma sets_pair_cancel_measure[simp]: + "sets (M1\measure := m1\ \\<^isub>M M2) = sets (M1 \\<^isub>M M2)" + "sets (M1 \\<^isub>M M2\measure := m2\) = sets (M1 \\<^isub>M M2)" + unfolding pair_measure_def pair_measure_generator_def sets_sigma + by simp_all + +lemma measurable_pair_cancel_measure[simp]: + "measurable (M1\measure := m1\ \\<^isub>M M2) M = measurable (M1 \\<^isub>M M2) M" + "measurable (M1 \\<^isub>M M2\measure := m2\) M = measurable (M1 \\<^isub>M M2) M" + "measurable M (M1\measure := m3\ \\<^isub>M M2) = measurable M (M1 \\<^isub>M M2)" + "measurable M (M1 \\<^isub>M M2\measure := m4\) = measurable M (M1 \\<^isub>M M2)" + unfolding measurable_def by (auto simp add: space_pair_measure) + +lemma (in product_sigma_finite) product_measure_exists: + assumes "finite I" + shows "\\. sigma_finite_measure (sigma (product_algebra_generator I M) \ measure := \ \) \ + (\A\\ i\I. sets (M i). \ (Pi\<^isub>E I A) = (\i\I. M.\ i (A i)))" +using `finite I` proof induct + case empty + interpret finite_product_sigma_finite M "{}" by default simp + let ?\ = "(\A. if A = {} then 0 else 1) :: 'd set \ extreal" + show ?case + proof (intro exI conjI ballI) + have "sigma_algebra (sigma G \measure := ?\\)" + by (rule sigma_algebra_cong) (simp_all add: product_algebra_def) + then have "measure_space (sigma G\measure := ?\\)" + by (rule finite_additivity_sufficient) + (simp_all add: positive_def additive_def sets_sigma + product_algebra_generator_def image_constant) + then show "sigma_finite_measure (sigma G\measure := ?\\)" + by (auto intro!: exI[of _ "\i. {\_. undefined}"] + simp: sigma_finite_measure_def sigma_finite_measure_axioms_def + product_algebra_generator_def) + qed auto +next + case (insert i I) + interpret finite_product_sigma_finite M I by default fact + have "finite (insert i I)" using `finite I` by auto + interpret I': finite_product_sigma_finite M "insert i I" by default fact + from insert obtain \ where + prod: "\A. A \ (\ i\I. sets (M i)) \ \ (Pi\<^isub>E I A) = (\i\I. M.\ i (A i))" and + "sigma_finite_measure (sigma G\ measure := \ \)" by auto + then interpret I: sigma_finite_measure "P\ measure := \\" unfolding product_algebra_def by simp + interpret P: pair_sigma_finite "P\ measure := \\" "M i" .. + let ?h = "(\(f, y). f(i := y))" + let ?\ = "\A. P.\ (?h -` A \ space P.P)" + have I': "sigma_algebra (I'.P\ measure := ?\ \)" + by (rule I'.sigma_algebra_cong) simp_all + interpret I'': measure_space "I'.P\ measure := ?\ \" + using measurable_add_dim[OF `i \ I`] + by (intro P.measure_space_vimage[OF I']) (auto simp add: measure_preserving_def) + show ?case + proof (intro exI[of _ ?\] conjI ballI) + let "?m A" = "measure (Pi\<^isub>M I M\measure := \\ \\<^isub>M M i) (?h -` A \ space P.P)" + { fix A assume A: "A \ (\ i\insert i I. sets (M i))" + then have *: "?h -` Pi\<^isub>E (insert i I) A \ space P.P = Pi\<^isub>E I A \ A i" + using `i \ I` M.sets_into_space by (auto simp: space_pair_measure space_product_algebra) blast + show "?m (Pi\<^isub>E (insert i I) A) = (\i\insert i I. M.\ i (A i))" + unfolding * using A + apply (subst P.pair_measure_times) + using A apply fastsimp + using A apply fastsimp + using `i \ I` `finite I` prod[of A] A by (auto simp: ac_simps) } + note product = this + have *: "sigma I'.G\ measure := ?\ \ = I'.P\ measure := ?\ \" + by (simp add: product_algebra_def) + show "sigma_finite_measure (sigma I'.G\ measure := ?\ \)" + proof (unfold *, default, simp) + from I'.sigma_finite_pairs guess F :: "'i \ nat \ 'a set" .. + then have F: "\j. j \ insert i I \ range (F j) \ sets (M j)" + "incseq (\k. \\<^isub>E j \ insert i I. F j k)" + "(\k. \\<^isub>E j \ insert i I. F j k) = space I'.G" + "\k. \j. j \ insert i I \ \ j (F j k) \ \" + by blast+ + let "?F k" = "\\<^isub>E j \ insert i I. F j k" + show "\F::nat \ ('i \ 'a) set. range F \ sets I'.P \ + (\i. F i) = (\\<^isub>E i\insert i I. space (M i)) \ (\i. ?m (F i) \ \)" + proof (intro exI[of _ ?F] conjI allI) + show "range ?F \ sets I'.P" using F(1) by auto + next + from F(3) show "(\i. ?F i) = (\\<^isub>E i\insert i I. space (M i))" by simp + next + fix j + have "\k. k \ insert i I \ 0 \ measure (M k) (F k j)" + using F(1) by auto + with F `finite I` setprod_PInf[of "insert i I", OF this] show "?\ (?F j) \ \" + by (subst product) auto + qed + qed + qed +qed + +sublocale finite_product_sigma_finite \ sigma_finite_measure P + unfolding product_algebra_def + using product_measure_exists[OF finite_index] + by (rule someI2_ex) auto + +lemma (in finite_product_sigma_finite) measure_times: + assumes "\i. i \ I \ A i \ sets (M i)" + shows "\ (Pi\<^isub>E I A) = (\i\I. M.\ i (A i))" + unfolding product_algebra_def + using product_measure_exists[OF finite_index] + proof (rule someI2_ex, elim conjE) + fix \ assume *: "\A\\ i\I. sets (M i). \ (Pi\<^isub>E I A) = (\i\I. M.\ i (A i))" + have "Pi\<^isub>E I A = Pi\<^isub>E I (\i\I. A i)" by (auto dest: Pi_mem) + then have "\ (Pi\<^isub>E I A) = \ (Pi\<^isub>E I (\i\I. A i))" by simp + also have "\ = (\i\I. M.\ i ((\i\I. A i) i))" using assms * by auto + finally show "measure (sigma G\measure := \\) (Pi\<^isub>E I A) = (\i\I. M.\ i (A i))" + by (simp add: sigma_def) + qed + +lemma (in product_sigma_finite) product_measure_empty[simp]: + "measure (Pi\<^isub>M {} M) {\x. undefined} = 1" +proof - + interpret finite_product_sigma_finite M "{}" by default auto + from measure_times[of "\x. {}"] show ?thesis by simp +qed + +lemma (in finite_product_sigma_algebra) P_empty: + assumes "I = {}" + shows "space P = {\k. undefined}" "sets P = { {}, {\k. undefined} }" + unfolding product_algebra_def product_algebra_generator_def `I = {}` + by (simp_all add: sigma_def image_constant) + +lemma (in product_sigma_finite) positive_integral_empty: + assumes pos: "0 \ f (\k. undefined)" + shows "integral\<^isup>P (Pi\<^isub>M {} M) f = f (\k. undefined)" +proof - + interpret finite_product_sigma_finite M "{}" by default (fact finite.emptyI) + have "\A. measure (Pi\<^isub>M {} M) (Pi\<^isub>E {} A) = 1" + using assms by (subst measure_times) auto + then show ?thesis + unfolding positive_integral_def simple_function_def simple_integral_def_raw + proof (simp add: P_empty, intro antisym) + show "f (\k. undefined) \ (SUP f:{g. g \ max 0 \ f}. f (\k. undefined))" + by (intro le_SUPI) (auto simp: le_fun_def split: split_max) + show "(SUP f:{g. g \ max 0 \ f}. f (\k. undefined)) \ f (\k. undefined)" using pos + by (intro SUP_leI) (auto simp: le_fun_def simp: max_def split: split_if_asm) + qed +qed + +lemma (in product_sigma_finite) measure_fold: + assumes IJ[simp]: "I \ J = {}" and fin: "finite I" "finite J" + assumes A: "A \ sets (Pi\<^isub>M (I \ J) M)" + shows "measure (Pi\<^isub>M (I \ J) M) A = + measure (Pi\<^isub>M I M \\<^isub>M Pi\<^isub>M J M) ((\(x,y). merge I x J y) -` A \ space (Pi\<^isub>M I M \\<^isub>M Pi\<^isub>M J M))" +proof - + interpret I: finite_product_sigma_finite M I by default fact + interpret J: finite_product_sigma_finite M J by default fact + have "finite (I \ J)" using fin by auto + interpret IJ: finite_product_sigma_finite M "I \ J" by default fact + interpret P: pair_sigma_finite I.P J.P by default + let ?g = "\(x,y). merge I x J y" + let "?X S" = "?g -` S \ space P.P" + from IJ.sigma_finite_pairs obtain F where + F: "\i. i\ I \ J \ range (F i) \ sets (M i)" + "incseq (\k. \\<^isub>E i\I \ J. F i k)" + "(\k. \\<^isub>E i\I \ J. F i k) = space IJ.G" + "\k. \i\I\J. \ i (F i k) \ \" + by auto + let ?F = "\k. \\<^isub>E i\I \ J. F i k" + show "IJ.\ A = P.\ (?X A)" + proof (rule measure_unique_Int_stable_vimage) + show "measure_space IJ.P" "measure_space P.P" by default + show "sets (sigma IJ.G) = sets IJ.P" "space IJ.G = space IJ.P" "A \ sets (sigma IJ.G)" + using A unfolding product_algebra_def by auto + next + show "Int_stable IJ.G" + by (simp add: PiE_Int Int_stable_def product_algebra_def + product_algebra_generator_def) + auto + show "range ?F \ sets IJ.G" using F + by (simp add: image_subset_iff product_algebra_def + product_algebra_generator_def) + show "incseq ?F" "(\i. ?F i) = space IJ.G " by fact+ + next + fix k + have "\j. j \ I \ J \ 0 \ measure (M j) (F j k)" + using F(1) by auto + with F `finite I` setprod_PInf[of "I \ J", OF this] + show "IJ.\ (?F k) \ \" by (subst IJ.measure_times) auto + next + fix A assume "A \ sets IJ.G" + then obtain F where A: "A = Pi\<^isub>E (I \ J) F" + and F: "\i. i \ I \ J \ F i \ sets (M i)" + by (auto simp: product_algebra_generator_def) + then have X: "?X A = (Pi\<^isub>E I F \ Pi\<^isub>E J F)" + using sets_into_space by (auto simp: space_pair_measure) blast+ + then have "P.\ (?X A) = (\i\I. \ i (F i)) * (\i\J. \ i (F i))" + using `finite J` `finite I` F + by (simp add: P.pair_measure_times I.measure_times J.measure_times) + also have "\ = (\i\I \ J. \ i (F i))" + using `finite J` `finite I` `I \ J = {}` by (simp add: setprod_Un_one) + also have "\ = IJ.\ A" + using `finite J` `finite I` F unfolding A + by (intro IJ.measure_times[symmetric]) auto + finally show "IJ.\ A = P.\ (?X A)" by (rule sym) + qed (rule measurable_merge[OF IJ]) +qed + +lemma (in product_sigma_finite) measure_preserving_merge: + assumes IJ: "I \ J = {}" and "finite I" "finite J" + shows "(\(x,y). merge I x J y) \ measure_preserving (Pi\<^isub>M I M \\<^isub>M Pi\<^isub>M J M) (Pi\<^isub>M (I \ J) M)" + by (intro measure_preservingI measurable_merge[OF IJ] measure_fold[symmetric] assms) + +lemma (in product_sigma_finite) product_positive_integral_fold: + assumes IJ[simp]: "I \ J = {}" "finite I" "finite J" + and f: "f \ borel_measurable (Pi\<^isub>M (I \ J) M)" + shows "integral\<^isup>P (Pi\<^isub>M (I \ J) M) f = + (\\<^isup>+ x. (\\<^isup>+ y. f (merge I x J y) \(Pi\<^isub>M J M)) \(Pi\<^isub>M I M))" +proof - + interpret I: finite_product_sigma_finite M I by default fact + interpret J: finite_product_sigma_finite M J by default fact + interpret P: pair_sigma_finite "Pi\<^isub>M I M" "Pi\<^isub>M J M" by default + interpret IJ: finite_product_sigma_finite M "I \ J" by default simp + have P_borel: "(\x. f (case x of (x, y) \ merge I x J y)) \ borel_measurable P.P" + using measurable_comp[OF measurable_merge[OF IJ(1)] f] by (simp add: comp_def) + show ?thesis + unfolding P.positive_integral_fst_measurable[OF P_borel, simplified] + proof (rule P.positive_integral_vimage) + show "sigma_algebra IJ.P" by default + show "(\(x, y). merge I x J y) \ measure_preserving P.P IJ.P" + using IJ by (rule measure_preserving_merge) + show "f \ borel_measurable IJ.P" using f by simp + qed +qed + +lemma (in product_sigma_finite) measure_preserving_component_singelton: + "(\x. x i) \ measure_preserving (Pi\<^isub>M {i} M) (M i)" +proof (intro measure_preservingI measurable_component_singleton) + interpret I: finite_product_sigma_finite M "{i}" by default simp + fix A let ?P = "(\x. x i) -` A \ space I.P" + assume A: "A \ sets (M i)" + then have *: "?P = {i} \\<^isub>E A" using sets_into_space by auto + show "I.\ ?P = M.\ i A" unfolding * + using A I.measure_times[of "\_. A"] by auto +qed auto + +lemma (in product_sigma_finite) product_positive_integral_singleton: + assumes f: "f \ borel_measurable (M i)" + shows "integral\<^isup>P (Pi\<^isub>M {i} M) (\x. f (x i)) = integral\<^isup>P (M i) f" +proof - + interpret I: finite_product_sigma_finite M "{i}" by default simp + show ?thesis + proof (rule I.positive_integral_vimage[symmetric]) + show "sigma_algebra (M i)" by (rule sigma_algebras) + show "(\x. x i) \ measure_preserving (Pi\<^isub>M {i} M) (M i)" + by (rule measure_preserving_component_singelton) + show "f \ borel_measurable (M i)" by fact + qed +qed + +lemma (in product_sigma_finite) product_positive_integral_insert: + assumes [simp]: "finite I" "i \ I" + and f: "f \ borel_measurable (Pi\<^isub>M (insert i I) M)" + shows "integral\<^isup>P (Pi\<^isub>M (insert i I) M) f = (\\<^isup>+ x. (\\<^isup>+ y. f (x(i := y)) \(M i)) \(Pi\<^isub>M I M))" +proof - + interpret I: finite_product_sigma_finite M I by default auto + interpret i: finite_product_sigma_finite M "{i}" by default auto + interpret P: pair_sigma_algebra I.P i.P .. + have IJ: "I \ {i} = {}" and insert: "I \ {i} = insert i I" + using f by auto + show ?thesis + unfolding product_positive_integral_fold[OF IJ, unfolded insert, simplified, OF f] + proof (rule I.positive_integral_cong, subst product_positive_integral_singleton) + fix x assume x: "x \ space I.P" + let "?f y" = "f (restrict (x(i := y)) (insert i I))" + have f'_eq: "\y. ?f y = f (x(i := y))" + using x by (auto intro!: arg_cong[where f=f] simp: fun_eq_iff extensional_def) + show "?f \ borel_measurable (M i)" unfolding f'_eq + using measurable_comp[OF measurable_component_update f] x `i \ I` + by (simp add: comp_def) + show "integral\<^isup>P (M i) ?f = \\<^isup>+ y. f (x(i:=y)) \M i" + unfolding f'_eq by simp + qed +qed + +lemma (in product_sigma_finite) product_positive_integral_setprod: + fixes f :: "'i \ 'a \ extreal" + assumes "finite I" and borel: "\i. i \ I \ f i \ borel_measurable (M i)" + and pos: "\i x. i \ I \ 0 \ f i x" + shows "(\\<^isup>+ x. (\i\I. f i (x i)) \Pi\<^isub>M I M) = (\i\I. integral\<^isup>P (M i) (f i))" +using assms proof induct + case empty + interpret finite_product_sigma_finite M "{}" by default auto + then show ?case by simp +next + case (insert i I) + note `finite I`[intro, simp] + interpret I: finite_product_sigma_finite M I by default auto + have *: "\x y. (\j\I. f j (if j = i then y else x j)) = (\j\I. f j (x j))" + using insert by (auto intro!: setprod_cong) + have prod: "\J. J \ insert i I \ (\x. (\i\J. f i (x i))) \ borel_measurable (Pi\<^isub>M J M)" + using sets_into_space insert + by (intro sigma_algebra.borel_measurable_extreal_setprod sigma_algebra_product_algebra + measurable_comp[OF measurable_component_singleton, unfolded comp_def]) + auto + then show ?case + apply (simp add: product_positive_integral_insert[OF insert(1,2) prod]) + apply (simp add: insert * pos borel setprod_extreal_pos M.positive_integral_multc) + apply (subst I.positive_integral_cmult) + apply (auto simp add: pos borel insert setprod_extreal_pos positive_integral_positive) + done +qed + +lemma (in product_sigma_finite) product_integral_singleton: + assumes f: "f \ borel_measurable (M i)" + shows "(\x. f (x i) \Pi\<^isub>M {i} M) = integral\<^isup>L (M i) f" +proof - + interpret I: finite_product_sigma_finite M "{i}" by default simp + have *: "(\x. extreal (f x)) \ borel_measurable (M i)" + "(\x. extreal (- f x)) \ borel_measurable (M i)" + using assms by auto + show ?thesis + unfolding lebesgue_integral_def *[THEN product_positive_integral_singleton] .. +qed + +lemma (in product_sigma_finite) product_integral_fold: + assumes IJ[simp]: "I \ J = {}" and fin: "finite I" "finite J" + and f: "integrable (Pi\<^isub>M (I \ J) M) f" + shows "integral\<^isup>L (Pi\<^isub>M (I \ J) M) f = (\x. (\y. f (merge I x J y) \Pi\<^isub>M J M) \Pi\<^isub>M I M)" +proof - + interpret I: finite_product_sigma_finite M I by default fact + interpret J: finite_product_sigma_finite M J by default fact + have "finite (I \ J)" using fin by auto + interpret IJ: finite_product_sigma_finite M "I \ J" by default fact + interpret P: pair_sigma_finite I.P J.P by default + let ?M = "\(x, y). merge I x J y" + let ?f = "\x. f (?M x)" + show ?thesis + proof (subst P.integrable_fst_measurable(2)[of ?f, simplified]) + have 1: "sigma_algebra IJ.P" by default + have 2: "?M \ measure_preserving P.P IJ.P" using measure_preserving_merge[OF assms(1,2,3)] . + have 3: "integrable (Pi\<^isub>M (I \ J) M) f" by fact + then have 4: "f \ borel_measurable (Pi\<^isub>M (I \ J) M)" + by (simp add: integrable_def) + show "integrable P.P ?f" + by (rule P.integrable_vimage[where f=f, OF 1 2 3]) + show "integral\<^isup>L IJ.P f = integral\<^isup>L P.P ?f" + by (rule P.integral_vimage[where f=f, OF 1 2 4]) + qed +qed + +lemma (in product_sigma_finite) product_integral_insert: + assumes [simp]: "finite I" "i \ I" + and f: "integrable (Pi\<^isub>M (insert i I) M) f" + shows "integral\<^isup>L (Pi\<^isub>M (insert i I) M) f = (\x. (\y. f (x(i:=y)) \M i) \Pi\<^isub>M I M)" +proof - + interpret I: finite_product_sigma_finite M I by default auto + interpret I': finite_product_sigma_finite M "insert i I" by default auto + interpret i: finite_product_sigma_finite M "{i}" by default auto + interpret P: pair_sigma_finite I.P i.P .. + have IJ: "I \ {i} = {}" by auto + show ?thesis + unfolding product_integral_fold[OF IJ, simplified, OF f] + proof (rule I.integral_cong, subst product_integral_singleton) + fix x assume x: "x \ space I.P" + let "?f y" = "f (restrict (x(i := y)) (insert i I))" + have f'_eq: "\y. ?f y = f (x(i := y))" + using x by (auto intro!: arg_cong[where f=f] simp: fun_eq_iff extensional_def) + have f: "f \ borel_measurable I'.P" using f unfolding integrable_def by auto + show "?f \ borel_measurable (M i)" + unfolding measurable_cong[OF f'_eq] + using measurable_comp[OF measurable_component_update f] x `i \ I` + by (simp add: comp_def) + show "integral\<^isup>L (M i) ?f = integral\<^isup>L (M i) (\y. f (x(i := y)))" + unfolding f'_eq by simp + qed +qed + +lemma (in product_sigma_finite) product_integrable_setprod: + fixes f :: "'i \ 'a \ real" + assumes [simp]: "finite I" and integrable: "\i. i \ I \ integrable (M i) (f i)" + shows "integrable (Pi\<^isub>M I M) (\x. (\i\I. f i (x i)))" (is "integrable _ ?f") +proof - + interpret finite_product_sigma_finite M I by default fact + have f: "\i. i \ I \ f i \ borel_measurable (M i)" + using integrable unfolding integrable_def by auto + then have borel: "?f \ borel_measurable P" + using measurable_comp[OF measurable_component_singleton f] + by (auto intro!: borel_measurable_setprod simp: comp_def) + moreover have "integrable (Pi\<^isub>M I M) (\x. \\i\I. f i (x i)\)" + proof (unfold integrable_def, intro conjI) + show "(\x. abs (?f x)) \ borel_measurable P" + using borel by auto + have "(\\<^isup>+x. extreal (abs (?f x)) \P) = (\\<^isup>+x. (\i\I. extreal (abs (f i (x i)))) \P)" + by (simp add: setprod_extreal abs_setprod) + also have "\ = (\i\I. (\\<^isup>+x. extreal (abs (f i x)) \M i))" + using f by (subst product_positive_integral_setprod) auto + also have "\ < \" + using integrable[THEN M.integrable_abs] + by (simp add: setprod_PInf integrable_def M.positive_integral_positive) + finally show "(\\<^isup>+x. extreal (abs (?f x)) \P) \ \" by auto + have "(\\<^isup>+x. extreal (- abs (?f x)) \P) = (\\<^isup>+x. 0 \P)" + by (intro positive_integral_cong_pos) auto + then show "(\\<^isup>+x. extreal (- abs (?f x)) \P) \ \" by simp + qed + ultimately show ?thesis + by (rule integrable_abs_iff[THEN iffD1]) +qed + +lemma (in product_sigma_finite) product_integral_setprod: + fixes f :: "'i \ 'a \ real" + assumes "finite I" "I \ {}" and integrable: "\i. i \ I \ integrable (M i) (f i)" + shows "(\x. (\i\I. f i (x i)) \Pi\<^isub>M I M) = (\i\I. integral\<^isup>L (M i) (f i))" +using assms proof (induct rule: finite_ne_induct) + case (singleton i) + then show ?case by (simp add: product_integral_singleton integrable_def) +next + case (insert i I) + then have iI: "finite (insert i I)" by auto + then have prod: "\J. J \ insert i I \ + integrable (Pi\<^isub>M J M) (\x. (\i\J. f i (x i)))" + by (intro product_integrable_setprod insert(5)) (auto intro: finite_subset) + interpret I: finite_product_sigma_finite M I by default fact + have *: "\x y. (\j\I. f j (if j = i then y else x j)) = (\j\I. f j (x j))" + using `i \ I` by (auto intro!: setprod_cong) + show ?case + unfolding product_integral_insert[OF insert(1,3) prod[OF subset_refl]] + by (simp add: * insert integral_multc I.integral_cmult[OF prod] subset_insertI) +qed + +end \ No newline at end of file diff -r 8448713d48b7 -r 5b52c6a9c627 src/HOL/Probability/Lebesgue_Measure.thy --- a/src/HOL/Probability/Lebesgue_Measure.thy Tue Mar 29 14:27:31 2011 +0200 +++ b/src/HOL/Probability/Lebesgue_Measure.thy Tue Mar 29 14:27:39 2011 +0200 @@ -6,7 +6,7 @@ header {* Lebsegue measure *} theory Lebesgue_Measure - imports Product_Measure + imports Finite_Product_Measure begin subsection {* Standard Cubes *} @@ -50,9 +50,6 @@ lemma cube_subset_Suc[intro]: "cube n \ cube (Suc n)" unfolding cube_def_raw subset_eq apply safe unfolding mem_interval by auto -lemma Pi_iff: "f \ Pi I X \ (\i\I. f i \ X i)" - unfolding Pi_def by auto - subsection {* Lebesgue measure *} definition lebesgue :: "'a::ordered_euclidean_space measure_space" where diff -r 8448713d48b7 -r 5b52c6a9c627 src/HOL/Probability/Probability_Space.thy --- a/src/HOL/Probability/Probability_Space.thy Tue Mar 29 14:27:31 2011 +0200 +++ b/src/HOL/Probability/Probability_Space.thy Tue Mar 29 14:27:39 2011 +0200 @@ -6,7 +6,7 @@ header {*Probability spaces*} theory Probability_Space -imports Lebesgue_Integration Radon_Nikodym Product_Measure +imports Lebesgue_Integration Radon_Nikodym Finite_Product_Measure begin lemma real_of_extreal_inverse[simp]: diff -r 8448713d48b7 -r 5b52c6a9c627 src/HOL/Probability/Product_Measure.thy --- a/src/HOL/Probability/Product_Measure.thy Tue Mar 29 14:27:31 2011 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,1975 +0,0 @@ -(* Title: HOL/Probability/Product_Measure.thy - Author: Johannes Hölzl, TU München -*) - -header {*Product measure spaces*} - -theory Product_Measure -imports Lebesgue_Integration -begin - -lemma sigma_sets_subseteq: assumes "A \ B" shows "sigma_sets X A \ sigma_sets X B" -proof - fix x assume "x \ sigma_sets X A" then show "x \ sigma_sets X B" - by induct (insert `A \ B`, auto intro: sigma_sets.intros) -qed - -lemma times_Int_times: "A \ B \ C \ D = (A \ C) \ (B \ D)" - by auto - -lemma Pair_vimage_times[simp]: "\A B x. Pair x -` (A \ B) = (if x \ A then B else {})" - by auto - -lemma rev_Pair_vimage_times[simp]: "\A B y. (\x. (x, y)) -` (A \ B) = (if y \ B then A else {})" - by auto - -lemma case_prod_distrib: "f (case x of (x, y) \ g x y) = (case x of (x, y) \ f (g x y))" - by (cases x) simp - -lemma split_beta': "(\(x,y). f x y) = (\x. f (fst x) (snd x))" - by (auto simp: fun_eq_iff) - -abbreviation - "Pi\<^isub>E A B \ Pi A B \ extensional A" - -syntax - "_PiE" :: "[pttrn, 'a set, 'b set] => ('a => 'b) set" ("(3PIE _:_./ _)" 10) - -syntax (xsymbols) - "_PiE" :: "[pttrn, 'a set, 'b set] => ('a => 'b) set" ("(3\\<^isub>E _\_./ _)" 10) - -syntax (HTML output) - "_PiE" :: "[pttrn, 'a set, 'b set] => ('a => 'b) set" ("(3\\<^isub>E _\_./ _)" 10) - -translations - "PIE x:A. B" == "CONST Pi\<^isub>E A (%x. B)" - -abbreviation - funcset_extensional :: "['a set, 'b set] => ('a => 'b) set" - (infixr "->\<^isub>E" 60) where - "A ->\<^isub>E B \ Pi\<^isub>E A (%_. B)" - -notation (xsymbols) - funcset_extensional (infixr "\\<^isub>E" 60) - -lemma extensional_empty[simp]: "extensional {} = {\x. undefined}" - by safe (auto simp add: extensional_def fun_eq_iff) - -lemma extensional_insert[intro, simp]: - assumes "a \ extensional (insert i I)" - shows "a(i := b) \ extensional (insert i I)" - using assms unfolding extensional_def by auto - -lemma extensional_Int[simp]: - "extensional I \ extensional I' = extensional (I \ I')" - unfolding extensional_def by auto - -definition - "merge I x J y = (\i. if i \ I then x i else if i \ J then y i else undefined)" - -lemma merge_apply[simp]: - "I \ J = {} \ i \ I \ merge I x J y i = x i" - "I \ J = {} \ i \ J \ merge I x J y i = y i" - "J \ I = {} \ i \ I \ merge I x J y i = x i" - "J \ I = {} \ i \ J \ merge I x J y i = y i" - "i \ I \ i \ J \ merge I x J y i = undefined" - unfolding merge_def by auto - -lemma merge_commute: - "I \ J = {} \ merge I x J y = merge J y I x" - by (auto simp: merge_def intro!: ext) - -lemma Pi_cancel_merge_range[simp]: - "I \ J = {} \ x \ Pi I (merge I A J B) \ x \ Pi I A" - "I \ J = {} \ x \ Pi I (merge J B I A) \ x \ Pi I A" - "J \ I = {} \ x \ Pi I (merge I A J B) \ x \ Pi I A" - "J \ I = {} \ x \ Pi I (merge J B I A) \ x \ Pi I A" - by (auto simp: Pi_def) - -lemma Pi_cancel_merge[simp]: - "I \ J = {} \ merge I x J y \ Pi I B \ x \ Pi I B" - "J \ I = {} \ merge I x J y \ Pi I B \ x \ Pi I B" - "I \ J = {} \ merge I x J y \ Pi J B \ y \ Pi J B" - "J \ I = {} \ merge I x J y \ Pi J B \ y \ Pi J B" - by (auto simp: Pi_def) - -lemma extensional_merge[simp]: "merge I x J y \ extensional (I \ J)" - by (auto simp: extensional_def) - -lemma restrict_Pi_cancel: "restrict x I \ Pi I A \ x \ Pi I A" - by (auto simp: restrict_def Pi_def) - -lemma restrict_merge[simp]: - "I \ J = {} \ restrict (merge I x J y) I = restrict x I" - "I \ J = {} \ restrict (merge I x J y) J = restrict y J" - "J \ I = {} \ restrict (merge I x J y) I = restrict x I" - "J \ I = {} \ restrict (merge I x J y) J = restrict y J" - by (auto simp: restrict_def intro!: ext) - -lemma extensional_insert_undefined[intro, simp]: - assumes "a \ extensional (insert i I)" - shows "a(i := undefined) \ extensional I" - using assms unfolding extensional_def by auto - -lemma extensional_insert_cancel[intro, simp]: - assumes "a \ extensional I" - shows "a \ extensional (insert i I)" - using assms unfolding extensional_def by auto - -lemma merge_singleton[simp]: "i \ I \ merge I x {i} y = restrict (x(i := y i)) (insert i I)" - unfolding merge_def by (auto simp: fun_eq_iff) - -lemma Pi_Int: "Pi I E \ Pi I F = (\ i\I. E i \ F i)" - by auto - -lemma PiE_Int: "(Pi\<^isub>E I A) \ (Pi\<^isub>E I B) = Pi\<^isub>E I (\x. A x \ B x)" - by auto - -lemma Pi_cancel_fupd_range[simp]: "i \ I \ x \ Pi I (B(i := b)) \ x \ Pi I B" - by (auto simp: Pi_def) - -lemma Pi_split_insert_domain[simp]: "x \ Pi (insert i I) X \ x \ Pi I X \ x i \ X i" - by (auto simp: Pi_def) - -lemma Pi_split_domain[simp]: "x \ Pi (I \ J) X \ x \ Pi I X \ x \ Pi J X" - by (auto simp: Pi_def) - -lemma Pi_cancel_fupd[simp]: "i \ I \ x(i := a) \ Pi I B \ x \ Pi I B" - by (auto simp: Pi_def) - -lemma restrict_vimage: - assumes "I \ J = {}" - shows "(\x. (restrict x I, restrict x J)) -` (Pi\<^isub>E I E \ Pi\<^isub>E J F) = Pi (I \ J) (merge I E J F)" - using assms by (auto simp: restrict_Pi_cancel) - -lemma merge_vimage: - assumes "I \ J = {}" - shows "(\(x,y). merge I x J y) -` Pi\<^isub>E (I \ J) E = Pi I E \ Pi J E" - using assms by (auto simp: restrict_Pi_cancel) - -lemma restrict_fupd[simp]: "i \ I \ restrict (f (i := x)) I = restrict f I" - by (auto simp: restrict_def intro!: ext) - -lemma merge_restrict[simp]: - "merge I (restrict x I) J y = merge I x J y" - "merge I x J (restrict y J) = merge I x J y" - unfolding merge_def by (auto intro!: ext) - -lemma merge_x_x_eq_restrict[simp]: - "merge I x J x = restrict x (I \ J)" - unfolding merge_def by (auto intro!: ext) - -lemma Pi_fupd_iff: "i \ I \ f \ Pi I (B(i := A)) \ f \ Pi (I - {i}) B \ f i \ A" - apply auto - apply (drule_tac x=x in Pi_mem) - apply (simp_all split: split_if_asm) - apply (drule_tac x=i in Pi_mem) - apply (auto dest!: Pi_mem) - done - -lemma Pi_UN: - fixes A :: "nat \ 'i \ 'a set" - assumes "finite I" and mono: "\i n m. i \ I \ n \ m \ A n i \ A m i" - shows "(\n. Pi I (A n)) = (\ i\I. \n. A n i)" -proof (intro set_eqI iffI) - fix f assume "f \ (\ i\I. \n. A n i)" - then have "\i\I. \n. f i \ A n i" by auto - from bchoice[OF this] obtain n where n: "\i. i \ I \ f i \ (A (n i) i)" by auto - obtain k where k: "\i. i \ I \ n i \ k" - using `finite I` finite_nat_set_iff_bounded_le[of "n`I"] by auto - have "f \ Pi I (A k)" - proof (intro Pi_I) - fix i assume "i \ I" - from mono[OF this, of "n i" k] k[OF this] n[OF this] - show "f i \ A k i" by auto - qed - then show "f \ (\n. Pi I (A n))" by auto -qed auto - -lemma PiE_cong: - assumes "\i. i\I \ A i = B i" - shows "Pi\<^isub>E I A = Pi\<^isub>E I B" - using assms by (auto intro!: Pi_cong) - -lemma restrict_upd[simp]: - "i \ I \ (restrict f I)(i := y) = restrict (f(i := y)) (insert i I)" - by (auto simp: fun_eq_iff) - -lemma Pi_eq_subset: - assumes ne: "\i. i \ I \ F i \ {}" "\i. i \ I \ F' i \ {}" - assumes eq: "Pi\<^isub>E I F = Pi\<^isub>E I F'" and "i \ I" - shows "F i \ F' i" -proof - fix x assume "x \ F i" - with ne have "\j. \y. ((j \ I \ y \ F j \ (i = j \ x = y)) \ (j \ I \ y = undefined))" by auto - from choice[OF this] guess f .. note f = this - then have "f \ Pi\<^isub>E I F" by (auto simp: extensional_def) - then have "f \ Pi\<^isub>E I F'" using assms by simp - then show "x \ F' i" using f `i \ I` by auto -qed - -lemma Pi_eq_iff_not_empty: - assumes ne: "\i. i \ I \ F i \ {}" "\i. i \ I \ F' i \ {}" - shows "Pi\<^isub>E I F = Pi\<^isub>E I F' \ (\i\I. F i = F' i)" -proof (intro iffI ballI) - fix i assume eq: "Pi\<^isub>E I F = Pi\<^isub>E I F'" and i: "i \ I" - show "F i = F' i" - using Pi_eq_subset[of I F F', OF ne eq i] - using Pi_eq_subset[of I F' F, OF ne(2,1) eq[symmetric] i] - by auto -qed auto - -lemma Pi_eq_empty_iff: - "Pi\<^isub>E I F = {} \ (\i\I. F i = {})" -proof - assume "Pi\<^isub>E I F = {}" - show "\i\I. F i = {}" - proof (rule ccontr) - assume "\ ?thesis" - then have "\i. \y. (i \ I \ y \ F i) \ (i \ I \ y = undefined)" by auto - from choice[OF this] guess f .. - then have "f \ Pi\<^isub>E I F" by (auto simp: extensional_def) - with `Pi\<^isub>E I F = {}` show False by auto - qed -qed auto - -lemma Pi_eq_iff: - "Pi\<^isub>E I F = Pi\<^isub>E I F' \ (\i\I. F i = F' i) \ ((\i\I. F i = {}) \ (\i\I. F' i = {}))" -proof (intro iffI disjCI) - assume eq[simp]: "Pi\<^isub>E I F = Pi\<^isub>E I F'" - assume "\ ((\i\I. F i = {}) \ (\i\I. F' i = {}))" - then have "(\i\I. F i \ {}) \ (\i\I. F' i \ {})" - using Pi_eq_empty_iff[of I F] Pi_eq_empty_iff[of I F'] by auto - with Pi_eq_iff_not_empty[of I F F'] show "\i\I. F i = F' i" by auto -next - assume "(\i\I. F i = F' i) \ (\i\I. F i = {}) \ (\i\I. F' i = {})" - then show "Pi\<^isub>E I F = Pi\<^isub>E I F'" - using Pi_eq_empty_iff[of I F] Pi_eq_empty_iff[of I F'] by auto -qed - -section "Binary products" - -definition - "pair_measure_generator A B = - \ space = space A \ space B, - sets = {a \ b | a b. a \ sets A \ b \ sets B}, - measure = \X. \\<^isup>+x. (\\<^isup>+y. indicator X (x,y) \B) \A \" - -definition pair_measure (infixr "\\<^isub>M" 80) where - "A \\<^isub>M B = sigma (pair_measure_generator A B)" - -locale pair_sigma_algebra = M1: sigma_algebra M1 + M2: sigma_algebra M2 - for M1 :: "('a, 'c) measure_space_scheme" and M2 :: "('b, 'd) measure_space_scheme" - -abbreviation (in pair_sigma_algebra) - "E \ pair_measure_generator M1 M2" - -abbreviation (in pair_sigma_algebra) - "P \ M1 \\<^isub>M M2" - -lemma sigma_algebra_pair_measure: - "sets M1 \ Pow (space M1) \ sets M2 \ Pow (space M2) \ sigma_algebra (pair_measure M1 M2)" - by (force simp: pair_measure_def pair_measure_generator_def intro!: sigma_algebra_sigma) - -sublocale pair_sigma_algebra \ sigma_algebra P - using M1.space_closed M2.space_closed - by (rule sigma_algebra_pair_measure) - -lemma pair_measure_generatorI[intro, simp]: - "x \ sets A \ y \ sets B \ x \ y \ sets (pair_measure_generator A B)" - by (auto simp add: pair_measure_generator_def) - -lemma pair_measureI[intro, simp]: - "x \ sets A \ y \ sets B \ x \ y \ sets (A \\<^isub>M B)" - by (auto simp add: pair_measure_def) - -lemma space_pair_measure: - "space (A \\<^isub>M B) = space A \ space B" - by (simp add: pair_measure_def pair_measure_generator_def) - -lemma sets_pair_measure_generator: - "sets (pair_measure_generator N M) = (\(x, y). x \ y) ` (sets N \ sets M)" - unfolding pair_measure_generator_def by auto - -lemma pair_measure_generator_sets_into_space: - assumes "sets M \ Pow (space M)" "sets N \ Pow (space N)" - shows "sets (pair_measure_generator M N) \ Pow (space (pair_measure_generator M N))" - using assms by (auto simp: pair_measure_generator_def) - -lemma pair_measure_generator_Int_snd: - assumes "sets S1 \ Pow (space S1)" - shows "sets (pair_measure_generator S1 (algebra.restricted_space S2 A)) = - sets (algebra.restricted_space (pair_measure_generator S1 S2) (space S1 \ A))" - (is "?L = ?R") - apply (auto simp: pair_measure_generator_def image_iff) - using assms - apply (rule_tac x="a \ xa" in exI) - apply force - using assms - apply (rule_tac x="a" in exI) - apply (rule_tac x="b \ A" in exI) - apply auto - done - -lemma (in pair_sigma_algebra) - shows measurable_fst[intro!, simp]: - "fst \ measurable P M1" (is ?fst) - and measurable_snd[intro!, simp]: - "snd \ measurable P M2" (is ?snd) -proof - - { fix X assume "X \ sets M1" - then have "\X1\sets M1. \X2\sets M2. fst -` X \ space M1 \ space M2 = X1 \ X2" - apply - apply (rule bexI[of _ X]) apply (rule bexI[of _ "space M2"]) - using M1.sets_into_space by force+ } - moreover - { fix X assume "X \ sets M2" - then have "\X1\sets M1. \X2\sets M2. snd -` X \ space M1 \ space M2 = X1 \ X2" - apply - apply (rule bexI[of _ "space M1"]) apply (rule bexI[of _ X]) - using M2.sets_into_space by force+ } - ultimately have "?fst \ ?snd" - by (fastsimp simp: measurable_def sets_sigma space_pair_measure - intro!: sigma_sets.Basic) - then show ?fst ?snd by auto -qed - -lemma (in pair_sigma_algebra) measurable_pair_iff: - assumes "sigma_algebra M" - shows "f \ measurable M P \ - (fst \ f) \ measurable M M1 \ (snd \ f) \ measurable M M2" -proof - - interpret M: sigma_algebra M by fact - from assms show ?thesis - proof (safe intro!: measurable_comp[where b=P]) - assume f: "(fst \ f) \ measurable M M1" and s: "(snd \ f) \ measurable M M2" - show "f \ measurable M P" unfolding pair_measure_def - proof (rule M.measurable_sigma) - show "sets (pair_measure_generator M1 M2) \ Pow (space E)" - unfolding pair_measure_generator_def using M1.sets_into_space M2.sets_into_space by auto - show "f \ space M \ space E" - using f s by (auto simp: mem_Times_iff measurable_def comp_def space_sigma pair_measure_generator_def) - fix A assume "A \ sets E" - then obtain B C where "B \ sets M1" "C \ sets M2" "A = B \ C" - unfolding pair_measure_generator_def by auto - moreover have "(fst \ f) -` B \ space M \ sets M" - using f `B \ sets M1` unfolding measurable_def by auto - moreover have "(snd \ f) -` C \ space M \ sets M" - using s `C \ sets M2` unfolding measurable_def by auto - moreover have "f -` A \ space M = ((fst \ f) -` B \ space M) \ ((snd \ f) -` C \ space M)" - unfolding `A = B \ C` by (auto simp: vimage_Times) - ultimately show "f -` A \ space M \ sets M" by auto - qed - qed -qed - -lemma (in pair_sigma_algebra) measurable_pair: - assumes "sigma_algebra M" - assumes "(fst \ f) \ measurable M M1" "(snd \ f) \ measurable M M2" - shows "f \ measurable M P" - unfolding measurable_pair_iff[OF assms(1)] using assms(2,3) by simp - -lemma pair_measure_generatorE: - assumes "X \ sets (pair_measure_generator M1 M2)" - obtains A B where "X = A \ B" "A \ sets M1" "B \ sets M2" - using assms unfolding pair_measure_generator_def by auto - -lemma (in pair_sigma_algebra) pair_measure_generator_swap: - "(\X. (\(x,y). (y,x)) -` X \ space M2 \ space M1) ` sets E = sets (pair_measure_generator M2 M1)" -proof (safe elim!: pair_measure_generatorE) - fix A B assume "A \ sets M1" "B \ sets M2" - moreover then have "(\(x, y). (y, x)) -` (A \ B) \ space M2 \ space M1 = B \ A" - using M1.sets_into_space M2.sets_into_space by auto - ultimately show "(\(x, y). (y, x)) -` (A \ B) \ space M2 \ space M1 \ sets (pair_measure_generator M2 M1)" - by (auto intro: pair_measure_generatorI) -next - fix A B assume "A \ sets M1" "B \ sets M2" - then show "B \ A \ (\X. (\(x, y). (y, x)) -` X \ space M2 \ space M1) ` sets E" - using M1.sets_into_space M2.sets_into_space - by (auto intro!: image_eqI[where x="A \ B"] pair_measure_generatorI) -qed - -lemma (in pair_sigma_algebra) sets_pair_sigma_algebra_swap: - assumes Q: "Q \ sets P" - shows "(\(x,y). (y, x)) -` Q \ sets (M2 \\<^isub>M M1)" (is "_ \ sets ?Q") -proof - - let "?f Q" = "(\(x,y). (y, x)) -` Q \ space M2 \ space M1" - have *: "(\(x,y). (y, x)) -` Q = ?f Q" - using sets_into_space[OF Q] by (auto simp: space_pair_measure) - have "sets (M2 \\<^isub>M M1) = sets (sigma (pair_measure_generator M2 M1))" - unfolding pair_measure_def .. - also have "\ = sigma_sets (space M2 \ space M1) (?f ` sets E)" - unfolding sigma_def pair_measure_generator_swap[symmetric] - by (simp add: pair_measure_generator_def) - also have "\ = ?f ` sigma_sets (space M1 \ space M2) (sets E)" - using M1.sets_into_space M2.sets_into_space - by (intro sigma_sets_vimage) (auto simp: pair_measure_generator_def) - also have "\ = ?f ` sets P" - unfolding pair_measure_def pair_measure_generator_def sigma_def by simp - finally show ?thesis - using Q by (subst *) auto -qed - -lemma (in pair_sigma_algebra) pair_sigma_algebra_swap_measurable: - shows "(\(x,y). (y, x)) \ measurable P (M2 \\<^isub>M M1)" - (is "?f \ measurable ?P ?Q") - unfolding measurable_def -proof (intro CollectI conjI Pi_I ballI) - fix x assume "x \ space ?P" then show "(case x of (x, y) \ (y, x)) \ space ?Q" - unfolding pair_measure_generator_def pair_measure_def by auto -next - fix A assume "A \ sets (M2 \\<^isub>M M1)" - interpret Q: pair_sigma_algebra M2 M1 by default - with Q.sets_pair_sigma_algebra_swap[OF `A \ sets (M2 \\<^isub>M M1)`] - show "?f -` A \ space ?P \ sets ?P" by simp -qed - -lemma (in pair_sigma_algebra) measurable_cut_fst[simp,intro]: - assumes "Q \ sets P" shows "Pair x -` Q \ sets M2" -proof - - let ?Q' = "{Q. Q \ space P \ Pair x -` Q \ sets M2}" - let ?Q = "\ space = space P, sets = ?Q' \" - interpret Q: sigma_algebra ?Q - proof qed (auto simp: vimage_UN vimage_Diff space_pair_measure) - have "sets E \ sets ?Q" - using M1.sets_into_space M2.sets_into_space - by (auto simp: pair_measure_generator_def space_pair_measure) - then have "sets P \ sets ?Q" - apply (subst pair_measure_def, intro Q.sets_sigma_subset) - by (simp add: pair_measure_def) - with assms show ?thesis by auto -qed - -lemma (in pair_sigma_algebra) measurable_cut_snd: - assumes Q: "Q \ sets P" shows "(\x. (x, y)) -` Q \ sets M1" (is "?cut Q \ sets M1") -proof - - interpret Q: pair_sigma_algebra M2 M1 by default - with Q.measurable_cut_fst[OF sets_pair_sigma_algebra_swap[OF Q], of y] - show ?thesis by (simp add: vimage_compose[symmetric] comp_def) -qed - -lemma (in pair_sigma_algebra) measurable_pair_image_snd: - assumes m: "f \ measurable P M" and "x \ space M1" - shows "(\y. f (x, y)) \ measurable M2 M" - unfolding measurable_def -proof (intro CollectI conjI Pi_I ballI) - fix y assume "y \ space M2" with `f \ measurable P M` `x \ space M1` - show "f (x, y) \ space M" - unfolding measurable_def pair_measure_generator_def pair_measure_def by auto -next - fix A assume "A \ sets M" - then have "Pair x -` (f -` A \ space P) \ sets M2" (is "?C \ _") - using `f \ measurable P M` - by (intro measurable_cut_fst) (auto simp: measurable_def) - also have "?C = (\y. f (x, y)) -` A \ space M2" - using `x \ space M1` by (auto simp: pair_measure_generator_def pair_measure_def) - finally show "(\y. f (x, y)) -` A \ space M2 \ sets M2" . -qed - -lemma (in pair_sigma_algebra) measurable_pair_image_fst: - assumes m: "f \ measurable P M" and "y \ space M2" - shows "(\x. f (x, y)) \ measurable M1 M" -proof - - interpret Q: pair_sigma_algebra M2 M1 by default - from Q.measurable_pair_image_snd[OF measurable_comp `y \ space M2`, - OF Q.pair_sigma_algebra_swap_measurable m] - show ?thesis by simp -qed - -lemma (in pair_sigma_algebra) Int_stable_pair_measure_generator: "Int_stable E" - unfolding Int_stable_def -proof (intro ballI) - fix A B assume "A \ sets E" "B \ sets E" - then obtain A1 A2 B1 B2 where "A = A1 \ A2" "B = B1 \ B2" - "A1 \ sets M1" "A2 \ sets M2" "B1 \ sets M1" "B2 \ sets M2" - unfolding pair_measure_generator_def by auto - then show "A \ B \ sets E" - by (auto simp add: times_Int_times pair_measure_generator_def) -qed - -lemma finite_measure_cut_measurable: - fixes M1 :: "('a, 'c) measure_space_scheme" and M2 :: "('b, 'd) measure_space_scheme" - assumes "sigma_finite_measure M1" "finite_measure M2" - assumes "Q \ sets (M1 \\<^isub>M M2)" - shows "(\x. measure M2 (Pair x -` Q)) \ borel_measurable M1" - (is "?s Q \ _") -proof - - interpret M1: sigma_finite_measure M1 by fact - interpret M2: finite_measure M2 by fact - interpret pair_sigma_algebra M1 M2 by default - have [intro]: "sigma_algebra M1" by fact - have [intro]: "sigma_algebra M2" by fact - let ?D = "\ space = space P, sets = {A\sets P. ?s A \ borel_measurable M1} \" - note space_pair_measure[simp] - interpret dynkin_system ?D - proof (intro dynkin_systemI) - fix A assume "A \ sets ?D" then show "A \ space ?D" - using sets_into_space by simp - next - from top show "space ?D \ sets ?D" - by (auto simp add: if_distrib intro!: M1.measurable_If) - next - fix A assume "A \ sets ?D" - with sets_into_space have "\x. measure M2 (Pair x -` (space M1 \ space M2 - A)) = - (if x \ space M1 then measure M2 (space M2) - ?s A x else 0)" - by (auto intro!: M2.measure_compl simp: vimage_Diff) - with `A \ sets ?D` top show "space ?D - A \ sets ?D" - by (auto intro!: Diff M1.measurable_If M1.borel_measurable_extreal_diff) - next - fix F :: "nat \ ('a\'b) set" assume "disjoint_family F" "range F \ sets ?D" - moreover then have "\x. measure M2 (\i. Pair x -` F i) = (\i. ?s (F i) x)" - by (intro M2.measure_countably_additive[symmetric]) - (auto simp: disjoint_family_on_def) - ultimately show "(\i. F i) \ sets ?D" - by (auto simp: vimage_UN intro!: M1.borel_measurable_psuminf) - qed - have "sets P = sets ?D" apply (subst pair_measure_def) - proof (intro dynkin_lemma) - show "Int_stable E" by (rule Int_stable_pair_measure_generator) - from M1.sets_into_space have "\A. A \ sets M1 \ {x \ space M1. x \ A} = A" - by auto - then show "sets E \ sets ?D" - by (auto simp: pair_measure_generator_def sets_sigma if_distrib - intro: sigma_sets.Basic intro!: M1.measurable_If) - qed (auto simp: pair_measure_def) - with `Q \ sets P` have "Q \ sets ?D" by simp - then show "?s Q \ borel_measurable M1" by simp -qed - -subsection {* Binary products of $\sigma$-finite measure spaces *} - -locale pair_sigma_finite = M1: sigma_finite_measure M1 + M2: sigma_finite_measure M2 - for M1 :: "('a, 'c) measure_space_scheme" and M2 :: "('b, 'd) measure_space_scheme" - -sublocale pair_sigma_finite \ pair_sigma_algebra M1 M2 - by default - -lemma times_eq_iff: "A \ B = C \ D \ A = C \ B = D \ ((A = {} \ B = {}) \ (C = {} \ D = {}))" - by auto - -lemma (in pair_sigma_finite) measure_cut_measurable_fst: - assumes "Q \ sets P" shows "(\x. measure M2 (Pair x -` Q)) \ borel_measurable M1" (is "?s Q \ _") -proof - - have [intro]: "sigma_algebra M1" and [intro]: "sigma_algebra M2" by default+ - have M1: "sigma_finite_measure M1" by default - from M2.disjoint_sigma_finite guess F .. note F = this - then have F_sets: "\i. F i \ sets M2" by auto - let "?C x i" = "F i \ Pair x -` Q" - { fix i - let ?R = "M2.restricted_space (F i)" - have [simp]: "space M1 \ F i \ space M1 \ space M2 = space M1 \ F i" - using F M2.sets_into_space by auto - let ?R2 = "M2.restricted_space (F i)" - have "(\x. measure ?R2 (Pair x -` (space M1 \ space ?R2 \ Q))) \ borel_measurable M1" - proof (intro finite_measure_cut_measurable[OF M1]) - show "finite_measure ?R2" - using F by (intro M2.restricted_to_finite_measure) auto - have "(space M1 \ space ?R2) \ Q \ (op \ (space M1 \ F i)) ` sets P" - using `Q \ sets P` by (auto simp: image_iff) - also have "\ = sigma_sets (space M1 \ F i) ((op \ (space M1 \ F i)) ` sets E)" - unfolding pair_measure_def pair_measure_generator_def sigma_def - using `F i \ sets M2` M2.sets_into_space - by (auto intro!: sigma_sets_Int sigma_sets.Basic) - also have "\ \ sets (M1 \\<^isub>M ?R2)" - using M1.sets_into_space - apply (auto simp: times_Int_times pair_measure_def pair_measure_generator_def sigma_def - intro!: sigma_sets_subseteq) - apply (rule_tac x="a" in exI) - apply (rule_tac x="b \ F i" in exI) - by auto - finally show "(space M1 \ space ?R2) \ Q \ sets (M1 \\<^isub>M ?R2)" . - qed - moreover have "\x. Pair x -` (space M1 \ F i \ Q) = ?C x i" - using `Q \ sets P` sets_into_space by (auto simp: space_pair_measure) - ultimately have "(\x. measure M2 (?C x i)) \ borel_measurable M1" - by simp } - moreover - { fix x - have "(\i. measure M2 (?C x i)) = measure M2 (\i. ?C x i)" - proof (intro M2.measure_countably_additive) - show "range (?C x) \ sets M2" - using F `Q \ sets P` by (auto intro!: M2.Int) - have "disjoint_family F" using F by auto - show "disjoint_family (?C x)" - by (rule disjoint_family_on_bisimulation[OF `disjoint_family F`]) auto - qed - also have "(\i. ?C x i) = Pair x -` Q" - using F sets_into_space `Q \ sets P` - by (auto simp: space_pair_measure) - finally have "measure M2 (Pair x -` Q) = (\i. measure M2 (?C x i))" - by simp } - ultimately show ?thesis using `Q \ sets P` F_sets - by (auto intro!: M1.borel_measurable_psuminf M2.Int) -qed - -lemma (in pair_sigma_finite) measure_cut_measurable_snd: - assumes "Q \ sets P" shows "(\y. M1.\ ((\x. (x, y)) -` Q)) \ borel_measurable M2" -proof - - interpret Q: pair_sigma_finite M2 M1 by default - note sets_pair_sigma_algebra_swap[OF assms] - from Q.measure_cut_measurable_fst[OF this] - show ?thesis by (simp add: vimage_compose[symmetric] comp_def) -qed - -lemma (in pair_sigma_algebra) pair_sigma_algebra_measurable: - assumes "f \ measurable P M" shows "(\(x,y). f (y, x)) \ measurable (M2 \\<^isub>M M1) M" -proof - - interpret Q: pair_sigma_algebra M2 M1 by default - have *: "(\(x,y). f (y, x)) = f \ (\(x,y). (y, x))" by (simp add: fun_eq_iff) - show ?thesis - using Q.pair_sigma_algebra_swap_measurable assms - unfolding * by (rule measurable_comp) -qed - -lemma (in pair_sigma_finite) pair_measure_alt: - assumes "A \ sets P" - shows "measure (M1 \\<^isub>M M2) A = (\\<^isup>+ x. measure M2 (Pair x -` A) \M1)" - apply (simp add: pair_measure_def pair_measure_generator_def) -proof (rule M1.positive_integral_cong) - fix x assume "x \ space M1" - have *: "\y. indicator A (x, y) = (indicator (Pair x -` A) y :: extreal)" - unfolding indicator_def by auto - show "(\\<^isup>+ y. indicator A (x, y) \M2) = measure M2 (Pair x -` A)" - unfolding * - apply (subst M2.positive_integral_indicator) - apply (rule measurable_cut_fst[OF assms]) - by simp -qed - -lemma (in pair_sigma_finite) pair_measure_times: - assumes A: "A \ sets M1" and "B \ sets M2" - shows "measure (M1 \\<^isub>M M2) (A \ B) = M1.\ A * measure M2 B" -proof - - have "measure (M1 \\<^isub>M M2) (A \ B) = (\\<^isup>+ x. measure M2 B * indicator A x \M1)" - using assms by (auto intro!: M1.positive_integral_cong simp: pair_measure_alt) - with assms show ?thesis - by (simp add: M1.positive_integral_cmult_indicator ac_simps) -qed - -lemma (in measure_space) measure_not_negative[simp,intro]: - assumes A: "A \ sets M" shows "\ A \ - \" - using positive_measure[OF A] by auto - -lemma (in pair_sigma_finite) sigma_finite_up_in_pair_measure_generator: - "\F::nat \ ('a \ 'b) set. range F \ sets E \ incseq F \ (\i. F i) = space E \ - (\i. measure (M1 \\<^isub>M M2) (F i) \ \)" -proof - - obtain F1 :: "nat \ 'a set" and F2 :: "nat \ 'b set" where - F1: "range F1 \ sets M1" "incseq F1" "(\i. F1 i) = space M1" "\i. M1.\ (F1 i) \ \" and - F2: "range F2 \ sets M2" "incseq F2" "(\i. F2 i) = space M2" "\i. M2.\ (F2 i) \ \" - using M1.sigma_finite_up M2.sigma_finite_up by auto - then have space: "space M1 = (\i. F1 i)" "space M2 = (\i. F2 i)" by auto - let ?F = "\i. F1 i \ F2 i" - show ?thesis unfolding space_pair_measure - proof (intro exI[of _ ?F] conjI allI) - show "range ?F \ sets E" using F1 F2 - by (fastsimp intro!: pair_measure_generatorI) - next - have "space M1 \ space M2 \ (\i. ?F i)" - proof (intro subsetI) - fix x assume "x \ space M1 \ space M2" - then obtain i j where "fst x \ F1 i" "snd x \ F2 j" - by (auto simp: space) - then have "fst x \ F1 (max i j)" "snd x \ F2 (max j i)" - using `incseq F1` `incseq F2` unfolding incseq_def - by (force split: split_max)+ - then have "(fst x, snd x) \ F1 (max i j) \ F2 (max i j)" - by (intro SigmaI) (auto simp add: min_max.sup_commute) - then show "x \ (\i. ?F i)" by auto - qed - then show "(\i. ?F i) = space E" - using space by (auto simp: space pair_measure_generator_def) - next - fix i show "incseq (\i. F1 i \ F2 i)" - using `incseq F1` `incseq F2` unfolding incseq_Suc_iff by auto - next - fix i - from F1 F2 have "F1 i \ sets M1" "F2 i \ sets M2" by auto - with F1 F2 M1.positive_measure[OF this(1)] M2.positive_measure[OF this(2)] - show "measure P (F1 i \ F2 i) \ \" - by (simp add: pair_measure_times) - qed -qed - -sublocale pair_sigma_finite \ sigma_finite_measure P -proof - show "positive P (measure P)" - unfolding pair_measure_def pair_measure_generator_def sigma_def positive_def - by (auto intro: M1.positive_integral_positive M2.positive_integral_positive) - - show "countably_additive P (measure P)" - unfolding countably_additive_def - proof (intro allI impI) - fix F :: "nat \ ('a \ 'b) set" - assume F: "range F \ sets P" "disjoint_family F" - from F have *: "\i. F i \ sets P" "(\i. F i) \ sets P" by auto - moreover from F have "\i. (\x. measure M2 (Pair x -` F i)) \ borel_measurable M1" - by (intro measure_cut_measurable_fst) auto - moreover have "\x. disjoint_family (\i. Pair x -` F i)" - by (intro disjoint_family_on_bisimulation[OF F(2)]) auto - moreover have "\x. x \ space M1 \ range (\i. Pair x -` F i) \ sets M2" - using F by auto - ultimately show "(\n. measure P (F n)) = measure P (\i. F i)" - by (simp add: pair_measure_alt vimage_UN M1.positive_integral_suminf[symmetric] - M2.measure_countably_additive - cong: M1.positive_integral_cong) - qed - - from sigma_finite_up_in_pair_measure_generator guess F :: "nat \ ('a \ 'b) set" .. note F = this - show "\F::nat \ ('a \ 'b) set. range F \ sets P \ (\i. F i) = space P \ (\i. measure P (F i) \ \)" - proof (rule exI[of _ F], intro conjI) - show "range F \ sets P" using F by (auto simp: pair_measure_def) - show "(\i. F i) = space P" - using F by (auto simp: pair_measure_def pair_measure_generator_def) - show "\i. measure P (F i) \ \" using F by auto - qed -qed - -lemma (in pair_sigma_algebra) sets_swap: - assumes "A \ sets P" - shows "(\(x, y). (y, x)) -` A \ space (M2 \\<^isub>M M1) \ sets (M2 \\<^isub>M M1)" - (is "_ -` A \ space ?Q \ sets ?Q") -proof - - have *: "(\(x, y). (y, x)) -` A \ space ?Q = (\(x, y). (y, x)) -` A" - using `A \ sets P` sets_into_space by (auto simp: space_pair_measure) - show ?thesis - unfolding * using assms by (rule sets_pair_sigma_algebra_swap) -qed - -lemma (in pair_sigma_finite) pair_measure_alt2: - assumes A: "A \ sets P" - shows "\ A = (\\<^isup>+y. M1.\ ((\x. (x, y)) -` A) \M2)" - (is "_ = ?\ A") -proof - - interpret Q: pair_sigma_finite M2 M1 by default - from sigma_finite_up_in_pair_measure_generator guess F :: "nat \ ('a \ 'b) set" .. note F = this - have [simp]: "\m. \ space = space E, sets = sets (sigma E), measure = m \ = P\ measure := m \" - unfolding pair_measure_def by simp - - have "\ A = Q.\ ((\(y, x). (x, y)) -` A \ space Q.P)" - proof (rule measure_unique_Int_stable_vimage[OF Int_stable_pair_measure_generator]) - show "measure_space P" "measure_space Q.P" by default - show "(\(y, x). (x, y)) \ measurable Q.P P" by (rule Q.pair_sigma_algebra_swap_measurable) - show "sets (sigma E) = sets P" "space E = space P" "A \ sets (sigma E)" - using assms unfolding pair_measure_def by auto - show "range F \ sets E" "incseq F" "(\i. F i) = space E" "\i. \ (F i) \ \" - using F `A \ sets P` by (auto simp: pair_measure_def) - fix X assume "X \ sets E" - then obtain A B where X[simp]: "X = A \ B" and AB: "A \ sets M1" "B \ sets M2" - unfolding pair_measure_def pair_measure_generator_def by auto - then have "(\(y, x). (x, y)) -` X \ space Q.P = B \ A" - using M1.sets_into_space M2.sets_into_space by (auto simp: space_pair_measure) - then show "\ X = Q.\ ((\(y, x). (x, y)) -` X \ space Q.P)" - using AB by (simp add: pair_measure_times Q.pair_measure_times ac_simps) - qed - then show ?thesis - using sets_into_space[OF A] Q.pair_measure_alt[OF sets_swap[OF A]] - by (auto simp add: Q.pair_measure_alt space_pair_measure - intro!: M2.positive_integral_cong arg_cong[where f="M1.\"]) -qed - -lemma pair_sigma_algebra_sigma: - assumes 1: "incseq S1" "(\i. S1 i) = space E1" "range S1 \ sets E1" and E1: "sets E1 \ Pow (space E1)" - assumes 2: "decseq S2" "(\i. S2 i) = space E2" "range S2 \ sets E2" and E2: "sets E2 \ Pow (space E2)" - shows "sets (sigma (pair_measure_generator (sigma E1) (sigma E2))) = sets (sigma (pair_measure_generator E1 E2))" - (is "sets ?S = sets ?E") -proof - - interpret M1: sigma_algebra "sigma E1" using E1 by (rule sigma_algebra_sigma) - interpret M2: sigma_algebra "sigma E2" using E2 by (rule sigma_algebra_sigma) - have P: "sets (pair_measure_generator E1 E2) \ Pow (space E1 \ space E2)" - using E1 E2 by (auto simp add: pair_measure_generator_def) - interpret E: sigma_algebra ?E unfolding pair_measure_generator_def - using E1 E2 by (intro sigma_algebra_sigma) auto - { fix A assume "A \ sets E1" - then have "fst -` A \ space ?E = A \ (\i. S2 i)" - using E1 2 unfolding pair_measure_generator_def by auto - also have "\ = (\i. A \ S2 i)" by auto - also have "\ \ sets ?E" unfolding pair_measure_generator_def sets_sigma - using 2 `A \ sets E1` - by (intro sigma_sets.Union) - (force simp: image_subset_iff intro!: sigma_sets.Basic) - finally have "fst -` A \ space ?E \ sets ?E" . } - moreover - { fix B assume "B \ sets E2" - then have "snd -` B \ space ?E = (\i. S1 i) \ B" - using E2 1 unfolding pair_measure_generator_def by auto - also have "\ = (\i. S1 i \ B)" by auto - also have "\ \ sets ?E" - using 1 `B \ sets E2` unfolding pair_measure_generator_def sets_sigma - by (intro sigma_sets.Union) - (force simp: image_subset_iff intro!: sigma_sets.Basic) - finally have "snd -` B \ space ?E \ sets ?E" . } - ultimately have proj: - "fst \ measurable ?E (sigma E1) \ snd \ measurable ?E (sigma E2)" - using E1 E2 by (subst (1 2) E.measurable_iff_sigma) - (auto simp: pair_measure_generator_def sets_sigma) - { fix A B assume A: "A \ sets (sigma E1)" and B: "B \ sets (sigma E2)" - with proj have "fst -` A \ space ?E \ sets ?E" "snd -` B \ space ?E \ sets ?E" - unfolding measurable_def by simp_all - moreover have "A \ B = (fst -` A \ space ?E) \ (snd -` B \ space ?E)" - using A B M1.sets_into_space M2.sets_into_space - by (auto simp: pair_measure_generator_def) - ultimately have "A \ B \ sets ?E" by auto } - then have "sigma_sets (space ?E) (sets (pair_measure_generator (sigma E1) (sigma E2))) \ sets ?E" - by (intro E.sigma_sets_subset) (auto simp add: pair_measure_generator_def sets_sigma) - then have subset: "sets ?S \ sets ?E" - by (simp add: sets_sigma pair_measure_generator_def) - show "sets ?S = sets ?E" - proof (intro set_eqI iffI) - fix A assume "A \ sets ?E" then show "A \ sets ?S" - unfolding sets_sigma - proof induct - case (Basic A) then show ?case - by (auto simp: pair_measure_generator_def sets_sigma intro: sigma_sets.Basic) - qed (auto intro: sigma_sets.intros simp: pair_measure_generator_def) - next - fix A assume "A \ sets ?S" then show "A \ sets ?E" using subset by auto - qed -qed - -section "Fubinis theorem" - -lemma (in pair_sigma_finite) simple_function_cut: - assumes f: "simple_function P f" "\x. 0 \ f x" - shows "(\x. \\<^isup>+y. f (x, y) \M2) \ borel_measurable M1" - and "(\\<^isup>+ x. (\\<^isup>+ y. f (x, y) \M2) \M1) = integral\<^isup>P P f" -proof - - have f_borel: "f \ borel_measurable P" - using f(1) by (rule borel_measurable_simple_function) - let "?F z" = "f -` {z} \ space P" - let "?F' x z" = "Pair x -` ?F z" - { fix x assume "x \ space M1" - have [simp]: "\z y. indicator (?F z) (x, y) = indicator (?F' x z) y" - by (auto simp: indicator_def) - have "\y. y \ space M2 \ (x, y) \ space P" using `x \ space M1` - by (simp add: space_pair_measure) - moreover have "\x z. ?F' x z \ sets M2" using f_borel - by (intro borel_measurable_vimage measurable_cut_fst) - ultimately have "simple_function M2 (\ y. f (x, y))" - apply (rule_tac M2.simple_function_cong[THEN iffD2, OF _]) - apply (rule simple_function_indicator_representation[OF f(1)]) - using `x \ space M1` by (auto simp del: space_sigma) } - note M2_sf = this - { fix x assume x: "x \ space M1" - then have "(\\<^isup>+y. f (x, y) \M2) = (\z\f ` space P. z * M2.\ (?F' x z))" - unfolding M2.positive_integral_eq_simple_integral[OF M2_sf[OF x] f(2)] - unfolding simple_integral_def - proof (safe intro!: setsum_mono_zero_cong_left) - from f(1) show "finite (f ` space P)" by (rule simple_functionD) - next - fix y assume "y \ space M2" then show "f (x, y) \ f ` space P" - using `x \ space M1` by (auto simp: space_pair_measure) - next - fix x' y assume "(x', y) \ space P" - "f (x', y) \ (\y. f (x, y)) ` space M2" - then have *: "?F' x (f (x', y)) = {}" - by (force simp: space_pair_measure) - show "f (x', y) * M2.\ (?F' x (f (x', y))) = 0" - unfolding * by simp - qed (simp add: vimage_compose[symmetric] comp_def - space_pair_measure) } - note eq = this - moreover have "\z. ?F z \ sets P" - by (auto intro!: f_borel borel_measurable_vimage simp del: space_sigma) - moreover then have "\z. (\x. M2.\ (?F' x z)) \ borel_measurable M1" - by (auto intro!: measure_cut_measurable_fst simp del: vimage_Int) - moreover have *: "\i x. 0 \ M2.\ (Pair x -` (f -` {i} \ space P))" - using f(1)[THEN simple_functionD(2)] f(2) by (intro M2.positive_measure measurable_cut_fst) - moreover { fix i assume "i \ f`space P" - with * have "\x. 0 \ i * M2.\ (Pair x -` (f -` {i} \ space P))" - using f(2) by auto } - ultimately - show "(\x. \\<^isup>+y. f (x, y) \M2) \ borel_measurable M1" - and "(\\<^isup>+ x. (\\<^isup>+ y. f (x, y) \M2) \M1) = integral\<^isup>P P f" using f(2) - by (auto simp del: vimage_Int cong: measurable_cong - intro!: M1.borel_measurable_extreal_setsum setsum_cong - simp add: M1.positive_integral_setsum simple_integral_def - M1.positive_integral_cmult - M1.positive_integral_cong[OF eq] - positive_integral_eq_simple_integral[OF f] - pair_measure_alt[symmetric]) -qed - -lemma (in pair_sigma_finite) positive_integral_fst_measurable: - assumes f: "f \ borel_measurable P" - shows "(\x. \\<^isup>+ y. f (x, y) \M2) \ borel_measurable M1" - (is "?C f \ borel_measurable M1") - and "(\\<^isup>+ x. (\\<^isup>+ y. f (x, y) \M2) \M1) = integral\<^isup>P P f" -proof - - from borel_measurable_implies_simple_function_sequence'[OF f] guess F . note F = this - then have F_borel: "\i. F i \ borel_measurable P" - by (auto intro: borel_measurable_simple_function) - note sf = simple_function_cut[OF F(1,5)] - then have "(\x. SUP i. ?C (F i) x) \ borel_measurable M1" - using F(1) by auto - moreover - { fix x assume "x \ space M1" - from F measurable_pair_image_snd[OF F_borel`x \ space M1`] - have "(\\<^isup>+y. (SUP i. F i (x, y)) \M2) = (SUP i. ?C (F i) x)" - by (intro M2.positive_integral_monotone_convergence_SUP) - (auto simp: incseq_Suc_iff le_fun_def) - then have "(SUP i. ?C (F i) x) = ?C f x" - unfolding F(4) positive_integral_max_0 by simp } - note SUPR_C = this - ultimately show "?C f \ borel_measurable M1" - by (simp cong: measurable_cong) - have "(\\<^isup>+x. (SUP i. F i x) \P) = (SUP i. integral\<^isup>P P (F i))" - using F_borel F - by (intro positive_integral_monotone_convergence_SUP) auto - also have "(SUP i. integral\<^isup>P P (F i)) = (SUP i. \\<^isup>+ x. (\\<^isup>+ y. F i (x, y) \M2) \M1)" - unfolding sf(2) by simp - also have "\ = \\<^isup>+ x. (SUP i. \\<^isup>+ y. F i (x, y) \M2) \M1" using F sf(1) - by (intro M1.positive_integral_monotone_convergence_SUP[symmetric]) - (auto intro!: M2.positive_integral_mono M2.positive_integral_positive - simp: incseq_Suc_iff le_fun_def) - also have "\ = \\<^isup>+ x. (\\<^isup>+ y. (SUP i. F i (x, y)) \M2) \M1" - using F_borel F(2,5) - by (auto intro!: M1.positive_integral_cong M2.positive_integral_monotone_convergence_SUP[symmetric] - simp: incseq_Suc_iff le_fun_def measurable_pair_image_snd) - finally show "(\\<^isup>+ x. (\\<^isup>+ y. f (x, y) \M2) \M1) = integral\<^isup>P P f" - using F by (simp add: positive_integral_max_0) -qed - -lemma (in pair_sigma_finite) measure_preserving_swap: - "(\(x,y). (y, x)) \ measure_preserving (M1 \\<^isub>M M2) (M2 \\<^isub>M M1)" -proof - interpret Q: pair_sigma_finite M2 M1 by default - show *: "(\(x,y). (y, x)) \ measurable (M1 \\<^isub>M M2) (M2 \\<^isub>M M1)" - using pair_sigma_algebra_swap_measurable . - fix X assume "X \ sets (M2 \\<^isub>M M1)" - from measurable_sets[OF * this] this Q.sets_into_space[OF this] - show "measure (M1 \\<^isub>M M2) ((\(x, y). (y, x)) -` X \ space P) = measure (M2 \\<^isub>M M1) X" - by (auto intro!: M1.positive_integral_cong arg_cong[where f="M2.\"] - simp: pair_measure_alt Q.pair_measure_alt2 space_pair_measure) -qed - -lemma (in pair_sigma_finite) positive_integral_product_swap: - assumes f: "f \ borel_measurable P" - shows "(\\<^isup>+x. f (case x of (x,y)\(y,x)) \(M2 \\<^isub>M M1)) = integral\<^isup>P P f" -proof - - interpret Q: pair_sigma_finite M2 M1 by default - have "sigma_algebra P" by default - with f show ?thesis - by (subst Q.positive_integral_vimage[OF _ Q.measure_preserving_swap]) auto -qed - -lemma (in pair_sigma_finite) positive_integral_snd_measurable: - assumes f: "f \ borel_measurable P" - shows "(\\<^isup>+ y. (\\<^isup>+ x. f (x, y) \M1) \M2) = integral\<^isup>P P f" -proof - - interpret Q: pair_sigma_finite M2 M1 by default - note pair_sigma_algebra_measurable[OF f] - from Q.positive_integral_fst_measurable[OF this] - have "(\\<^isup>+ y. (\\<^isup>+ x. f (x, y) \M1) \M2) = (\\<^isup>+ (x, y). f (y, x) \Q.P)" - by simp - also have "(\\<^isup>+ (x, y). f (y, x) \Q.P) = integral\<^isup>P P f" - unfolding positive_integral_product_swap[OF f, symmetric] - by (auto intro!: Q.positive_integral_cong) - finally show ?thesis . -qed - -lemma (in pair_sigma_finite) Fubini: - assumes f: "f \ borel_measurable P" - shows "(\\<^isup>+ y. (\\<^isup>+ x. f (x, y) \M1) \M2) = (\\<^isup>+ x. (\\<^isup>+ y. f (x, y) \M2) \M1)" - unfolding positive_integral_snd_measurable[OF assms] - unfolding positive_integral_fst_measurable[OF assms] .. - -lemma (in pair_sigma_finite) AE_pair: - assumes "AE x in P. Q x" - shows "AE x in M1. (AE y in M2. Q (x, y))" -proof - - obtain N where N: "N \ sets P" "\ N = 0" "{x\space P. \ Q x} \ N" - using assms unfolding almost_everywhere_def by auto - show ?thesis - proof (rule M1.AE_I) - from N measure_cut_measurable_fst[OF `N \ sets P`] - show "M1.\ {x\space M1. M2.\ (Pair x -` N) \ 0} = 0" - by (auto simp: pair_measure_alt M1.positive_integral_0_iff) - show "{x \ space M1. M2.\ (Pair x -` N) \ 0} \ sets M1" - by (intro M1.borel_measurable_extreal_neq_const measure_cut_measurable_fst N) - { fix x assume "x \ space M1" "M2.\ (Pair x -` N) = 0" - have "M2.almost_everywhere (\y. Q (x, y))" - proof (rule M2.AE_I) - show "M2.\ (Pair x -` N) = 0" by fact - show "Pair x -` N \ sets M2" by (intro measurable_cut_fst N) - show "{y \ space M2. \ Q (x, y)} \ Pair x -` N" - using N `x \ space M1` unfolding space_sigma space_pair_measure by auto - qed } - then show "{x \ space M1. \ M2.almost_everywhere (\y. Q (x, y))} \ {x \ space M1. M2.\ (Pair x -` N) \ 0}" - by auto - qed -qed - -lemma (in pair_sigma_algebra) measurable_product_swap: - "f \ measurable (M2 \\<^isub>M M1) M \ (\(x,y). f (y,x)) \ measurable P M" -proof - - interpret Q: pair_sigma_algebra M2 M1 by default - show ?thesis - using pair_sigma_algebra_measurable[of "\(x,y). f (y, x)"] - by (auto intro!: pair_sigma_algebra_measurable Q.pair_sigma_algebra_measurable iffI) -qed - -lemma (in pair_sigma_finite) integrable_product_swap: - assumes "integrable P f" - shows "integrable (M2 \\<^isub>M M1) (\(x,y). f (y,x))" -proof - - interpret Q: pair_sigma_finite M2 M1 by default - have *: "(\(x,y). f (y,x)) = (\x. f (case x of (x,y)\(y,x)))" by (auto simp: fun_eq_iff) - show ?thesis unfolding * - using assms unfolding integrable_def - apply (subst (1 2) positive_integral_product_swap) - using `integrable P f` unfolding integrable_def - by (auto simp: *[symmetric] Q.measurable_product_swap[symmetric]) -qed - -lemma (in pair_sigma_finite) integrable_product_swap_iff: - "integrable (M2 \\<^isub>M M1) (\(x,y). f (y,x)) \ integrable P f" -proof - - interpret Q: pair_sigma_finite M2 M1 by default - from Q.integrable_product_swap[of "\(x,y). f (y,x)"] integrable_product_swap[of f] - show ?thesis by auto -qed - -lemma (in pair_sigma_finite) integral_product_swap: - assumes "integrable P f" - shows "(\(x,y). f (y,x) \(M2 \\<^isub>M M1)) = integral\<^isup>L P f" -proof - - interpret Q: pair_sigma_finite M2 M1 by default - have *: "(\(x,y). f (y,x)) = (\x. f (case x of (x,y)\(y,x)))" by (auto simp: fun_eq_iff) - show ?thesis - unfolding lebesgue_integral_def * - apply (subst (1 2) positive_integral_product_swap) - using `integrable P f` unfolding integrable_def - by (auto simp: *[symmetric] Q.measurable_product_swap[symmetric]) -qed - -lemma (in pair_sigma_finite) integrable_fst_measurable: - assumes f: "integrable P f" - shows "M1.almost_everywhere (\x. integrable M2 (\ y. f (x, y)))" (is "?AE") - and "(\x. (\y. f (x, y) \M2) \M1) = integral\<^isup>L P f" (is "?INT") -proof - - let "?pf x" = "extreal (f x)" and "?nf x" = "extreal (- f x)" - have - borel: "?nf \ borel_measurable P""?pf \ borel_measurable P" and - int: "integral\<^isup>P P ?nf \ \" "integral\<^isup>P P ?pf \ \" - using assms by auto - have "(\\<^isup>+x. (\\<^isup>+y. extreal (f (x, y)) \M2) \M1) \ \" - "(\\<^isup>+x. (\\<^isup>+y. extreal (- f (x, y)) \M2) \M1) \ \" - using borel[THEN positive_integral_fst_measurable(1)] int - unfolding borel[THEN positive_integral_fst_measurable(2)] by simp_all - with borel[THEN positive_integral_fst_measurable(1)] - have AE_pos: "AE x in M1. (\\<^isup>+y. extreal (f (x, y)) \M2) \ \" - "AE x in M1. (\\<^isup>+y. extreal (- f (x, y)) \M2) \ \" - by (auto intro!: M1.positive_integral_PInf_AE ) - then have AE: "AE x in M1. \\\<^isup>+y. extreal (f (x, y)) \M2\ \ \" - "AE x in M1. \\\<^isup>+y. extreal (- f (x, y)) \M2\ \ \" - by (auto simp: M2.positive_integral_positive) - from AE_pos show ?AE using assms - by (simp add: measurable_pair_image_snd integrable_def) - { fix f have "(\\<^isup>+ x. - \\<^isup>+ y. extreal (f x y) \M2 \M1) = (\\<^isup>+x. 0 \M1)" - using M2.positive_integral_positive - by (intro M1.positive_integral_cong_pos) (auto simp: extreal_uminus_le_reorder) - then have "(\\<^isup>+ x. - \\<^isup>+ y. extreal (f x y) \M2 \M1) = 0" by simp } - note this[simp] - { fix f assume borel: "(\x. extreal (f x)) \ borel_measurable P" - and int: "integral\<^isup>P P (\x. extreal (f x)) \ \" - and AE: "M1.almost_everywhere (\x. (\\<^isup>+y. extreal (f (x, y)) \M2) \ \)" - have "integrable M1 (\x. real (\\<^isup>+y. extreal (f (x, y)) \M2))" (is "integrable M1 ?f") - proof (intro integrable_def[THEN iffD2] conjI) - show "?f \ borel_measurable M1" - using borel by (auto intro!: M1.borel_measurable_real_of_extreal positive_integral_fst_measurable) - have "(\\<^isup>+x. extreal (?f x) \M1) = (\\<^isup>+x. (\\<^isup>+y. extreal (f (x, y)) \M2) \M1)" - using AE M2.positive_integral_positive - by (auto intro!: M1.positive_integral_cong_AE simp: extreal_real) - then show "(\\<^isup>+x. extreal (?f x) \M1) \ \" - using positive_integral_fst_measurable[OF borel] int by simp - have "(\\<^isup>+x. extreal (- ?f x) \M1) = (\\<^isup>+x. 0 \M1)" - by (intro M1.positive_integral_cong_pos) - (simp add: M2.positive_integral_positive real_of_extreal_pos) - then show "(\\<^isup>+x. extreal (- ?f x) \M1) \ \" by simp - qed } - with this[OF borel(1) int(1) AE_pos(2)] this[OF borel(2) int(2) AE_pos(1)] - show ?INT - unfolding lebesgue_integral_def[of P] lebesgue_integral_def[of M2] - borel[THEN positive_integral_fst_measurable(2), symmetric] - using AE[THEN M1.integral_real] - by simp -qed - -lemma (in pair_sigma_finite) integrable_snd_measurable: - assumes f: "integrable P f" - shows "M2.almost_everywhere (\y. integrable M1 (\x. f (x, y)))" (is "?AE") - and "(\y. (\x. f (x, y) \M1) \M2) = integral\<^isup>L P f" (is "?INT") -proof - - interpret Q: pair_sigma_finite M2 M1 by default - have Q_int: "integrable Q.P (\(x, y). f (y, x))" - using f unfolding integrable_product_swap_iff . - show ?INT - using Q.integrable_fst_measurable(2)[OF Q_int] - using integral_product_swap[OF f] by simp - show ?AE - using Q.integrable_fst_measurable(1)[OF Q_int] - by simp -qed - -lemma (in pair_sigma_finite) Fubini_integral: - assumes f: "integrable P f" - shows "(\y. (\x. f (x, y) \M1) \M2) = (\x. (\y. f (x, y) \M2) \M1)" - unfolding integrable_snd_measurable[OF assms] - unfolding integrable_fst_measurable[OF assms] .. - -section "Finite product spaces" - -section "Products" - -locale product_sigma_algebra = - fixes M :: "'i \ ('a, 'b) measure_space_scheme" - assumes sigma_algebras: "\i. sigma_algebra (M i)" - -locale finite_product_sigma_algebra = product_sigma_algebra M - for M :: "'i \ ('a, 'b) measure_space_scheme" + - fixes I :: "'i set" - assumes finite_index: "finite I" - -definition - "product_algebra_generator I M = \ space = (\\<^isub>E i \ I. space (M i)), - sets = Pi\<^isub>E I ` (\ i \ I. sets (M i)), - measure = \A. (\i\I. measure (M i) ((SOME F. A = Pi\<^isub>E I F) i)) \" - -definition product_algebra_def: - "Pi\<^isub>M I M = sigma (product_algebra_generator I M) - \ measure := (SOME \. sigma_finite_measure (sigma (product_algebra_generator I M) \ measure := \ \) \ - (\A\\ i\I. sets (M i). \ (Pi\<^isub>E I A) = (\i\I. measure (M i) (A i))))\" - -syntax - "_PiM" :: "[pttrn, 'i set, ('a, 'b) measure_space_scheme] => - ('i => 'a, 'b) measure_space_scheme" ("(3PIM _:_./ _)" 10) - -syntax (xsymbols) - "_PiM" :: "[pttrn, 'i set, ('a, 'b) measure_space_scheme] => - ('i => 'a, 'b) measure_space_scheme" ("(3\\<^isub>M _\_./ _)" 10) - -syntax (HTML output) - "_PiM" :: "[pttrn, 'i set, ('a, 'b) measure_space_scheme] => - ('i => 'a, 'b) measure_space_scheme" ("(3\\<^isub>M _\_./ _)" 10) - -translations - "PIM x:I. M" == "CONST Pi\<^isub>M I (%x. M)" - -abbreviation (in finite_product_sigma_algebra) "G \ product_algebra_generator I M" -abbreviation (in finite_product_sigma_algebra) "P \ Pi\<^isub>M I M" - -sublocale product_sigma_algebra \ M: sigma_algebra "M i" for i by (rule sigma_algebras) - -lemma sigma_into_space: - assumes "sets M \ Pow (space M)" - shows "sets (sigma M) \ Pow (space M)" - using sigma_sets_into_sp[OF assms] unfolding sigma_def by auto - -lemma (in product_sigma_algebra) product_algebra_generator_into_space: - "sets (product_algebra_generator I M) \ Pow (space (product_algebra_generator I M))" - using M.sets_into_space unfolding product_algebra_generator_def - by auto blast - -lemma (in product_sigma_algebra) product_algebra_into_space: - "sets (Pi\<^isub>M I M) \ Pow (space (Pi\<^isub>M I M))" - using product_algebra_generator_into_space - by (auto intro!: sigma_into_space simp add: product_algebra_def) - -lemma (in product_sigma_algebra) sigma_algebra_product_algebra: "sigma_algebra (Pi\<^isub>M I M)" - using product_algebra_generator_into_space unfolding product_algebra_def - by (rule sigma_algebra.sigma_algebra_cong[OF sigma_algebra_sigma]) simp_all - -sublocale finite_product_sigma_algebra \ sigma_algebra P - using sigma_algebra_product_algebra . - -lemma product_algebraE: - assumes "A \ sets (product_algebra_generator I M)" - obtains E where "A = Pi\<^isub>E I E" "E \ (\ i\I. sets (M i))" - using assms unfolding product_algebra_generator_def by auto - -lemma product_algebra_generatorI[intro]: - assumes "E \ (\ i\I. sets (M i))" - shows "Pi\<^isub>E I E \ sets (product_algebra_generator I M)" - using assms unfolding product_algebra_generator_def by auto - -lemma space_product_algebra_generator[simp]: - "space (product_algebra_generator I M) = Pi\<^isub>E I (\i. space (M i))" - unfolding product_algebra_generator_def by simp - -lemma space_product_algebra[simp]: - "space (Pi\<^isub>M I M) = (\\<^isub>E i\I. space (M i))" - unfolding product_algebra_def product_algebra_generator_def by simp - -lemma sets_product_algebra: - "sets (Pi\<^isub>M I M) = sets (sigma (product_algebra_generator I M))" - unfolding product_algebra_def sigma_def by simp - -lemma product_algebra_generator_sets_into_space: - assumes "\i. i\I \ sets (M i) \ Pow (space (M i))" - shows "sets (product_algebra_generator I M) \ Pow (space (product_algebra_generator I M))" - using assms by (auto simp: product_algebra_generator_def) blast - -lemma (in finite_product_sigma_algebra) in_P[simp, intro]: - "\ \i. i \ I \ A i \ sets (M i) \ \ Pi\<^isub>E I A \ sets P" - by (auto simp: sets_product_algebra) - -section "Generating set generates also product algebra" - -lemma sigma_product_algebra_sigma_eq: - assumes "finite I" - assumes mono: "\i. i \ I \ incseq (S i)" - assumes union: "\i. i \ I \ (\j. S i j) = space (E i)" - assumes sets_into: "\i. i \ I \ range (S i) \ sets (E i)" - and E: "\i. sets (E i) \ Pow (space (E i))" - shows "sets (\\<^isub>M i\I. sigma (E i)) = sets (\\<^isub>M i\I. E i)" - (is "sets ?S = sets ?E") -proof cases - assume "I = {}" then show ?thesis - by (simp add: product_algebra_def product_algebra_generator_def) -next - assume "I \ {}" - interpret E: sigma_algebra "sigma (E i)" for i - using E by (rule sigma_algebra_sigma) - have into_space[intro]: "\i x A. A \ sets (E i) \ x i \ A \ x i \ space (E i)" - using E by auto - interpret G: sigma_algebra ?E - unfolding product_algebra_def product_algebra_generator_def using E - by (intro sigma_algebra.sigma_algebra_cong[OF sigma_algebra_sigma]) (auto dest: Pi_mem) - { fix A i assume "i \ I" and A: "A \ sets (E i)" - then have "(\x. x i) -` A \ space ?E = (\\<^isub>E j\I. if j = i then A else \n. S j n) \ space ?E" - using mono union unfolding incseq_Suc_iff space_product_algebra - by (auto dest: Pi_mem) - also have "\ = (\n. (\\<^isub>E j\I. if j = i then A else S j n))" - unfolding space_product_algebra - apply simp - apply (subst Pi_UN[OF `finite I`]) - using mono[THEN incseqD] apply simp - apply (simp add: PiE_Int) - apply (intro PiE_cong) - using A sets_into by (auto intro!: into_space) - also have "\ \ sets ?E" - using sets_into `A \ sets (E i)` - unfolding sets_product_algebra sets_sigma - by (intro sigma_sets.Union) - (auto simp: image_subset_iff intro!: sigma_sets.Basic) - finally have "(\x. x i) -` A \ space ?E \ sets ?E" . } - then have proj: - "\i. i\I \ (\x. x i) \ measurable ?E (sigma (E i))" - using E by (subst G.measurable_iff_sigma) - (auto simp: sets_product_algebra sets_sigma) - { fix A assume A: "\i. i \ I \ A i \ sets (sigma (E i))" - with proj have basic: "\i. i \ I \ (\x. x i) -` (A i) \ space ?E \ sets ?E" - unfolding measurable_def by simp - have "Pi\<^isub>E I A = (\i\I. (\x. x i) -` (A i) \ space ?E)" - using A E.sets_into_space `I \ {}` unfolding product_algebra_def by auto blast - then have "Pi\<^isub>E I A \ sets ?E" - using G.finite_INT[OF `finite I` `I \ {}` basic, of "\i. i"] by simp } - then have "sigma_sets (space ?E) (sets (product_algebra_generator I (\i. sigma (E i)))) \ sets ?E" - by (intro G.sigma_sets_subset) (auto simp add: product_algebra_generator_def) - then have subset: "sets ?S \ sets ?E" - by (simp add: sets_sigma sets_product_algebra) - show "sets ?S = sets ?E" - proof (intro set_eqI iffI) - fix A assume "A \ sets ?E" then show "A \ sets ?S" - unfolding sets_sigma sets_product_algebra - proof induct - case (Basic A) then show ?case - by (auto simp: sets_sigma product_algebra_generator_def intro: sigma_sets.Basic) - qed (auto intro: sigma_sets.intros simp: product_algebra_generator_def) - next - fix A assume "A \ sets ?S" then show "A \ sets ?E" using subset by auto - qed -qed - -lemma product_algebraI[intro]: - "E \ (\ i\I. sets (M i)) \ Pi\<^isub>E I E \ sets (Pi\<^isub>M I M)" - using assms unfolding product_algebra_def by (auto intro: product_algebra_generatorI) - -lemma (in product_sigma_algebra) measurable_component_update: - assumes "x \ space (Pi\<^isub>M I M)" and "i \ I" - shows "(\v. x(i := v)) \ measurable (M i) (Pi\<^isub>M (insert i I) M)" (is "?f \ _") - unfolding product_algebra_def apply simp -proof (intro measurable_sigma) - let ?G = "product_algebra_generator (insert i I) M" - show "sets ?G \ Pow (space ?G)" using product_algebra_generator_into_space . - show "?f \ space (M i) \ space ?G" - using M.sets_into_space assms by auto - fix A assume "A \ sets ?G" - from product_algebraE[OF this] guess E . note E = this - then have "?f -` A \ space (M i) = E i \ ?f -` A \ space (M i) = {}" - using M.sets_into_space assms by auto - then show "?f -` A \ space (M i) \ sets (M i)" - using E by (auto intro!: product_algebraI) -qed - -lemma (in product_sigma_algebra) measurable_add_dim: - assumes "i \ I" - shows "(\(f, y). f(i := y)) \ measurable (Pi\<^isub>M I M \\<^isub>M M i) (Pi\<^isub>M (insert i I) M)" -proof - - let ?f = "(\(f, y). f(i := y))" and ?G = "product_algebra_generator (insert i I) M" - interpret Ii: pair_sigma_algebra "Pi\<^isub>M I M" "M i" - unfolding pair_sigma_algebra_def - by (intro sigma_algebra_product_algebra sigma_algebras conjI) - have "?f \ measurable Ii.P (sigma ?G)" - proof (rule Ii.measurable_sigma) - show "sets ?G \ Pow (space ?G)" - using product_algebra_generator_into_space . - show "?f \ space (Pi\<^isub>M I M \\<^isub>M M i) \ space ?G" - by (auto simp: space_pair_measure) - next - fix A assume "A \ sets ?G" - then obtain F where "A = Pi\<^isub>E (insert i I) F" - and F: "\j. j \ I \ F j \ sets (M j)" "F i \ sets (M i)" - by (auto elim!: product_algebraE) - then have "?f -` A \ space (Pi\<^isub>M I M \\<^isub>M M i) = Pi\<^isub>E I F \ (F i)" - using sets_into_space `i \ I` - by (auto simp add: space_pair_measure) blast+ - then show "?f -` A \ space (Pi\<^isub>M I M \\<^isub>M M i) \ sets (Pi\<^isub>M I M \\<^isub>M M i)" - using F by (auto intro!: pair_measureI) - qed - then show ?thesis - by (simp add: product_algebra_def) -qed - -lemma (in product_sigma_algebra) measurable_merge: - assumes [simp]: "I \ J = {}" - shows "(\(x, y). merge I x J y) \ measurable (Pi\<^isub>M I M \\<^isub>M Pi\<^isub>M J M) (Pi\<^isub>M (I \ J) M)" -proof - - let ?I = "Pi\<^isub>M I M" and ?J = "Pi\<^isub>M J M" - interpret P: sigma_algebra "?I \\<^isub>M ?J" - by (intro sigma_algebra_pair_measure product_algebra_into_space) - let ?f = "\(x, y). merge I x J y" - let ?G = "product_algebra_generator (I \ J) M" - have "?f \ measurable (?I \\<^isub>M ?J) (sigma ?G)" - proof (rule P.measurable_sigma) - fix A assume "A \ sets ?G" - from product_algebraE[OF this] - obtain E where E: "A = Pi\<^isub>E (I \ J) E" "E \ (\ i\I \ J. sets (M i))" . - then have *: "?f -` A \ space (?I \\<^isub>M ?J) = Pi\<^isub>E I E \ Pi\<^isub>E J E" - using sets_into_space `I \ J = {}` - by (auto simp add: space_pair_measure) fast+ - then show "?f -` A \ space (?I \\<^isub>M ?J) \ sets (?I \\<^isub>M ?J)" - using E unfolding * by (auto intro!: pair_measureI in_sigma product_algebra_generatorI - simp: product_algebra_def) - qed (insert product_algebra_generator_into_space, auto simp: space_pair_measure) - then show "?f \ measurable (?I \\<^isub>M ?J) (Pi\<^isub>M (I \ J) M)" - unfolding product_algebra_def[of "I \ J"] by simp -qed - -lemma (in product_sigma_algebra) measurable_component_singleton: - assumes "i \ I" shows "(\x. x i) \ measurable (Pi\<^isub>M I M) (M i)" -proof (unfold measurable_def, intro CollectI conjI ballI) - fix A assume "A \ sets (M i)" - then have "(\x. x i) -` A \ space (Pi\<^isub>M I M) = (\\<^isub>E j\I. if i = j then A else space (M j))" - using M.sets_into_space `i \ I` by (fastsimp dest: Pi_mem split: split_if_asm) - then show "(\x. x i) -` A \ space (Pi\<^isub>M I M) \ sets (Pi\<^isub>M I M)" - using `A \ sets (M i)` by (auto intro!: product_algebraI) -qed (insert `i \ I`, auto) - -locale product_sigma_finite = - fixes M :: "'i \ ('a,'b) measure_space_scheme" - assumes sigma_finite_measures: "\i. sigma_finite_measure (M i)" - -locale finite_product_sigma_finite = product_sigma_finite M - for M :: "'i \ ('a,'b) measure_space_scheme" + - fixes I :: "'i set" assumes finite_index'[intro]: "finite I" - -sublocale product_sigma_finite \ M: sigma_finite_measure "M i" for i - by (rule sigma_finite_measures) - -sublocale product_sigma_finite \ product_sigma_algebra - by default - -sublocale finite_product_sigma_finite \ finite_product_sigma_algebra - by default (fact finite_index') - -lemma setprod_extreal_0: - fixes f :: "'a \ extreal" - shows "(\i\A. f i) = 0 \ (finite A \ (\i\A. f i = 0))" -proof cases - assume "finite A" - then show ?thesis by (induct A) auto -qed auto - -lemma setprod_extreal_pos: - fixes f :: "'a \ extreal" assumes pos: "\i. i \ I \ 0 \ f i" shows "0 \ (\i\I. f i)" -proof cases - assume "finite I" from this pos show ?thesis by induct auto -qed simp - -lemma setprod_PInf: - assumes "\i. i \ I \ 0 \ f i" - shows "(\i\I. f i) = \ \ finite I \ (\i\I. f i = \) \ (\i\I. f i \ 0)" -proof cases - assume "finite I" from this assms show ?thesis - proof (induct I) - case (insert i I) - then have pos: "0 \ f i" "0 \ setprod f I" by (auto intro!: setprod_extreal_pos) - from insert have "(\j\insert i I. f j) = \ \ setprod f I * f i = \" by auto - also have "\ \ (setprod f I = \ \ f i = \) \ f i \ 0 \ setprod f I \ 0" - using setprod_extreal_pos[of I f] pos - by (cases rule: extreal2_cases[of "f i" "setprod f I"]) auto - also have "\ \ finite (insert i I) \ (\j\insert i I. f j = \) \ (\j\insert i I. f j \ 0)" - using insert by (auto simp: setprod_extreal_0) - finally show ?case . - qed simp -qed simp - -lemma setprod_extreal: "(\i\A. extreal (f i)) = extreal (setprod f A)" -proof cases - assume "finite A" then show ?thesis - by induct (auto simp: one_extreal_def) -qed (simp add: one_extreal_def) - -lemma (in finite_product_sigma_finite) product_algebra_generator_measure: - assumes "Pi\<^isub>E I F \ sets G" - shows "measure G (Pi\<^isub>E I F) = (\i\I. M.\ i (F i))" -proof cases - assume ne: "\i\I. F i \ {}" - have "\i\I. (SOME F'. Pi\<^isub>E I F = Pi\<^isub>E I F') i = F i" - by (rule someI2[where P="\F'. Pi\<^isub>E I F = Pi\<^isub>E I F'"]) - (insert ne, auto simp: Pi_eq_iff) - then show ?thesis - unfolding product_algebra_generator_def by simp -next - assume empty: "\ (\j\I. F j \ {})" - then have "(\j\I. M.\ j (F j)) = 0" - by (auto simp: setprod_extreal_0 intro!: bexI) - moreover - have "\j\I. (SOME F'. Pi\<^isub>E I F = Pi\<^isub>E I F') j = {}" - by (rule someI2[where P="\F'. Pi\<^isub>E I F = Pi\<^isub>E I F'"]) - (insert empty, auto simp: Pi_eq_empty_iff[symmetric]) - then have "(\j\I. M.\ j ((SOME F'. Pi\<^isub>E I F = Pi\<^isub>E I F') j)) = 0" - by (auto simp: setprod_extreal_0 intro!: bexI) - ultimately show ?thesis - unfolding product_algebra_generator_def by simp -qed - -lemma (in finite_product_sigma_finite) sigma_finite_pairs: - "\F::'i \ nat \ 'a set. - (\i\I. range (F i) \ sets (M i)) \ - (\k. \i\I. \ i (F i k) \ \) \ incseq (\k. \\<^isub>E i\I. F i k) \ - (\k. \\<^isub>E i\I. F i k) = space G" -proof - - have "\i::'i. \F::nat \ 'a set. range F \ sets (M i) \ incseq F \ (\i. F i) = space (M i) \ (\k. \ i (F k) \ \)" - using M.sigma_finite_up by simp - from choice[OF this] guess F :: "'i \ nat \ 'a set" .. - then have F: "\i. range (F i) \ sets (M i)" "\i. incseq (F i)" "\i. (\j. F i j) = space (M i)" "\i k. \ i (F i k) \ \" - by auto - let ?F = "\k. \\<^isub>E i\I. F i k" - note space_product_algebra[simp] - show ?thesis - proof (intro exI[of _ F] conjI allI incseq_SucI set_eqI iffI ballI) - fix i show "range (F i) \ sets (M i)" by fact - next - fix i k show "\ i (F i k) \ \" by fact - next - fix A assume "A \ (\i. ?F i)" then show "A \ space G" - using `\i. range (F i) \ sets (M i)` M.sets_into_space - by (force simp: image_subset_iff) - next - fix f assume "f \ space G" - with Pi_UN[OF finite_index, of "\k i. F i k"] F - show "f \ (\i. ?F i)" by (auto simp: incseq_def) - next - fix i show "?F i \ ?F (Suc i)" - using `\i. incseq (F i)`[THEN incseq_SucD] by auto - qed -qed - -lemma sets_pair_cancel_measure[simp]: - "sets (M1\measure := m1\ \\<^isub>M M2) = sets (M1 \\<^isub>M M2)" - "sets (M1 \\<^isub>M M2\measure := m2\) = sets (M1 \\<^isub>M M2)" - unfolding pair_measure_def pair_measure_generator_def sets_sigma - by simp_all - -lemma measurable_pair_cancel_measure[simp]: - "measurable (M1\measure := m1\ \\<^isub>M M2) M = measurable (M1 \\<^isub>M M2) M" - "measurable (M1 \\<^isub>M M2\measure := m2\) M = measurable (M1 \\<^isub>M M2) M" - "measurable M (M1\measure := m3\ \\<^isub>M M2) = measurable M (M1 \\<^isub>M M2)" - "measurable M (M1 \\<^isub>M M2\measure := m4\) = measurable M (M1 \\<^isub>M M2)" - unfolding measurable_def by (auto simp add: space_pair_measure) - -lemma (in product_sigma_finite) product_measure_exists: - assumes "finite I" - shows "\\. sigma_finite_measure (sigma (product_algebra_generator I M) \ measure := \ \) \ - (\A\\ i\I. sets (M i). \ (Pi\<^isub>E I A) = (\i\I. M.\ i (A i)))" -using `finite I` proof induct - case empty - interpret finite_product_sigma_finite M "{}" by default simp - let ?\ = "(\A. if A = {} then 0 else 1) :: 'd set \ extreal" - show ?case - proof (intro exI conjI ballI) - have "sigma_algebra (sigma G \measure := ?\\)" - by (rule sigma_algebra_cong) (simp_all add: product_algebra_def) - then have "measure_space (sigma G\measure := ?\\)" - by (rule finite_additivity_sufficient) - (simp_all add: positive_def additive_def sets_sigma - product_algebra_generator_def image_constant) - then show "sigma_finite_measure (sigma G\measure := ?\\)" - by (auto intro!: exI[of _ "\i. {\_. undefined}"] - simp: sigma_finite_measure_def sigma_finite_measure_axioms_def - product_algebra_generator_def) - qed auto -next - case (insert i I) - interpret finite_product_sigma_finite M I by default fact - have "finite (insert i I)" using `finite I` by auto - interpret I': finite_product_sigma_finite M "insert i I" by default fact - from insert obtain \ where - prod: "\A. A \ (\ i\I. sets (M i)) \ \ (Pi\<^isub>E I A) = (\i\I. M.\ i (A i))" and - "sigma_finite_measure (sigma G\ measure := \ \)" by auto - then interpret I: sigma_finite_measure "P\ measure := \\" unfolding product_algebra_def by simp - interpret P: pair_sigma_finite "P\ measure := \\" "M i" .. - let ?h = "(\(f, y). f(i := y))" - let ?\ = "\A. P.\ (?h -` A \ space P.P)" - have I': "sigma_algebra (I'.P\ measure := ?\ \)" - by (rule I'.sigma_algebra_cong) simp_all - interpret I'': measure_space "I'.P\ measure := ?\ \" - using measurable_add_dim[OF `i \ I`] - by (intro P.measure_space_vimage[OF I']) (auto simp add: measure_preserving_def) - show ?case - proof (intro exI[of _ ?\] conjI ballI) - let "?m A" = "measure (Pi\<^isub>M I M\measure := \\ \\<^isub>M M i) (?h -` A \ space P.P)" - { fix A assume A: "A \ (\ i\insert i I. sets (M i))" - then have *: "?h -` Pi\<^isub>E (insert i I) A \ space P.P = Pi\<^isub>E I A \ A i" - using `i \ I` M.sets_into_space by (auto simp: space_pair_measure space_product_algebra) blast - show "?m (Pi\<^isub>E (insert i I) A) = (\i\insert i I. M.\ i (A i))" - unfolding * using A - apply (subst P.pair_measure_times) - using A apply fastsimp - using A apply fastsimp - using `i \ I` `finite I` prod[of A] A by (auto simp: ac_simps) } - note product = this - have *: "sigma I'.G\ measure := ?\ \ = I'.P\ measure := ?\ \" - by (simp add: product_algebra_def) - show "sigma_finite_measure (sigma I'.G\ measure := ?\ \)" - proof (unfold *, default, simp) - from I'.sigma_finite_pairs guess F :: "'i \ nat \ 'a set" .. - then have F: "\j. j \ insert i I \ range (F j) \ sets (M j)" - "incseq (\k. \\<^isub>E j \ insert i I. F j k)" - "(\k. \\<^isub>E j \ insert i I. F j k) = space I'.G" - "\k. \j. j \ insert i I \ \ j (F j k) \ \" - by blast+ - let "?F k" = "\\<^isub>E j \ insert i I. F j k" - show "\F::nat \ ('i \ 'a) set. range F \ sets I'.P \ - (\i. F i) = (\\<^isub>E i\insert i I. space (M i)) \ (\i. ?m (F i) \ \)" - proof (intro exI[of _ ?F] conjI allI) - show "range ?F \ sets I'.P" using F(1) by auto - next - from F(3) show "(\i. ?F i) = (\\<^isub>E i\insert i I. space (M i))" by simp - next - fix j - have "\k. k \ insert i I \ 0 \ measure (M k) (F k j)" - using F(1) by auto - with F `finite I` setprod_PInf[of "insert i I", OF this] show "?\ (?F j) \ \" - by (subst product) auto - qed - qed - qed -qed - -sublocale finite_product_sigma_finite \ sigma_finite_measure P - unfolding product_algebra_def - using product_measure_exists[OF finite_index] - by (rule someI2_ex) auto - -lemma (in finite_product_sigma_finite) measure_times: - assumes "\i. i \ I \ A i \ sets (M i)" - shows "\ (Pi\<^isub>E I A) = (\i\I. M.\ i (A i))" - unfolding product_algebra_def - using product_measure_exists[OF finite_index] - proof (rule someI2_ex, elim conjE) - fix \ assume *: "\A\\ i\I. sets (M i). \ (Pi\<^isub>E I A) = (\i\I. M.\ i (A i))" - have "Pi\<^isub>E I A = Pi\<^isub>E I (\i\I. A i)" by (auto dest: Pi_mem) - then have "\ (Pi\<^isub>E I A) = \ (Pi\<^isub>E I (\i\I. A i))" by simp - also have "\ = (\i\I. M.\ i ((\i\I. A i) i))" using assms * by auto - finally show "measure (sigma G\measure := \\) (Pi\<^isub>E I A) = (\i\I. M.\ i (A i))" - by (simp add: sigma_def) - qed - -lemma (in product_sigma_finite) product_measure_empty[simp]: - "measure (Pi\<^isub>M {} M) {\x. undefined} = 1" -proof - - interpret finite_product_sigma_finite M "{}" by default auto - from measure_times[of "\x. {}"] show ?thesis by simp -qed - -lemma (in finite_product_sigma_algebra) P_empty: - assumes "I = {}" - shows "space P = {\k. undefined}" "sets P = { {}, {\k. undefined} }" - unfolding product_algebra_def product_algebra_generator_def `I = {}` - by (simp_all add: sigma_def image_constant) - -lemma (in product_sigma_finite) positive_integral_empty: - assumes pos: "0 \ f (\k. undefined)" - shows "integral\<^isup>P (Pi\<^isub>M {} M) f = f (\k. undefined)" -proof - - interpret finite_product_sigma_finite M "{}" by default (fact finite.emptyI) - have "\A. measure (Pi\<^isub>M {} M) (Pi\<^isub>E {} A) = 1" - using assms by (subst measure_times) auto - then show ?thesis - unfolding positive_integral_def simple_function_def simple_integral_def_raw - proof (simp add: P_empty, intro antisym) - show "f (\k. undefined) \ (SUP f:{g. g \ max 0 \ f}. f (\k. undefined))" - by (intro le_SUPI) (auto simp: le_fun_def split: split_max) - show "(SUP f:{g. g \ max 0 \ f}. f (\k. undefined)) \ f (\k. undefined)" using pos - by (intro SUP_leI) (auto simp: le_fun_def simp: max_def split: split_if_asm) - qed -qed - -lemma (in product_sigma_finite) measure_fold: - assumes IJ[simp]: "I \ J = {}" and fin: "finite I" "finite J" - assumes A: "A \ sets (Pi\<^isub>M (I \ J) M)" - shows "measure (Pi\<^isub>M (I \ J) M) A = - measure (Pi\<^isub>M I M \\<^isub>M Pi\<^isub>M J M) ((\(x,y). merge I x J y) -` A \ space (Pi\<^isub>M I M \\<^isub>M Pi\<^isub>M J M))" -proof - - interpret I: finite_product_sigma_finite M I by default fact - interpret J: finite_product_sigma_finite M J by default fact - have "finite (I \ J)" using fin by auto - interpret IJ: finite_product_sigma_finite M "I \ J" by default fact - interpret P: pair_sigma_finite I.P J.P by default - let ?g = "\(x,y). merge I x J y" - let "?X S" = "?g -` S \ space P.P" - from IJ.sigma_finite_pairs obtain F where - F: "\i. i\ I \ J \ range (F i) \ sets (M i)" - "incseq (\k. \\<^isub>E i\I \ J. F i k)" - "(\k. \\<^isub>E i\I \ J. F i k) = space IJ.G" - "\k. \i\I\J. \ i (F i k) \ \" - by auto - let ?F = "\k. \\<^isub>E i\I \ J. F i k" - show "IJ.\ A = P.\ (?X A)" - proof (rule measure_unique_Int_stable_vimage) - show "measure_space IJ.P" "measure_space P.P" by default - show "sets (sigma IJ.G) = sets IJ.P" "space IJ.G = space IJ.P" "A \ sets (sigma IJ.G)" - using A unfolding product_algebra_def by auto - next - show "Int_stable IJ.G" - by (simp add: PiE_Int Int_stable_def product_algebra_def - product_algebra_generator_def) - auto - show "range ?F \ sets IJ.G" using F - by (simp add: image_subset_iff product_algebra_def - product_algebra_generator_def) - show "incseq ?F" "(\i. ?F i) = space IJ.G " by fact+ - next - fix k - have "\j. j \ I \ J \ 0 \ measure (M j) (F j k)" - using F(1) by auto - with F `finite I` setprod_PInf[of "I \ J", OF this] - show "IJ.\ (?F k) \ \" by (subst IJ.measure_times) auto - next - fix A assume "A \ sets IJ.G" - then obtain F where A: "A = Pi\<^isub>E (I \ J) F" - and F: "\i. i \ I \ J \ F i \ sets (M i)" - by (auto simp: product_algebra_generator_def) - then have X: "?X A = (Pi\<^isub>E I F \ Pi\<^isub>E J F)" - using sets_into_space by (auto simp: space_pair_measure) blast+ - then have "P.\ (?X A) = (\i\I. \ i (F i)) * (\i\J. \ i (F i))" - using `finite J` `finite I` F - by (simp add: P.pair_measure_times I.measure_times J.measure_times) - also have "\ = (\i\I \ J. \ i (F i))" - using `finite J` `finite I` `I \ J = {}` by (simp add: setprod_Un_one) - also have "\ = IJ.\ A" - using `finite J` `finite I` F unfolding A - by (intro IJ.measure_times[symmetric]) auto - finally show "IJ.\ A = P.\ (?X A)" by (rule sym) - qed (rule measurable_merge[OF IJ]) -qed - -lemma (in product_sigma_finite) measure_preserving_merge: - assumes IJ: "I \ J = {}" and "finite I" "finite J" - shows "(\(x,y). merge I x J y) \ measure_preserving (Pi\<^isub>M I M \\<^isub>M Pi\<^isub>M J M) (Pi\<^isub>M (I \ J) M)" - by (intro measure_preservingI measurable_merge[OF IJ] measure_fold[symmetric] assms) - -lemma (in product_sigma_finite) product_positive_integral_fold: - assumes IJ[simp]: "I \ J = {}" "finite I" "finite J" - and f: "f \ borel_measurable (Pi\<^isub>M (I \ J) M)" - shows "integral\<^isup>P (Pi\<^isub>M (I \ J) M) f = - (\\<^isup>+ x. (\\<^isup>+ y. f (merge I x J y) \(Pi\<^isub>M J M)) \(Pi\<^isub>M I M))" -proof - - interpret I: finite_product_sigma_finite M I by default fact - interpret J: finite_product_sigma_finite M J by default fact - interpret P: pair_sigma_finite "Pi\<^isub>M I M" "Pi\<^isub>M J M" by default - interpret IJ: finite_product_sigma_finite M "I \ J" by default simp - have P_borel: "(\x. f (case x of (x, y) \ merge I x J y)) \ borel_measurable P.P" - using measurable_comp[OF measurable_merge[OF IJ(1)] f] by (simp add: comp_def) - show ?thesis - unfolding P.positive_integral_fst_measurable[OF P_borel, simplified] - proof (rule P.positive_integral_vimage) - show "sigma_algebra IJ.P" by default - show "(\(x, y). merge I x J y) \ measure_preserving P.P IJ.P" - using IJ by (rule measure_preserving_merge) - show "f \ borel_measurable IJ.P" using f by simp - qed -qed - -lemma (in product_sigma_finite) measure_preserving_component_singelton: - "(\x. x i) \ measure_preserving (Pi\<^isub>M {i} M) (M i)" -proof (intro measure_preservingI measurable_component_singleton) - interpret I: finite_product_sigma_finite M "{i}" by default simp - fix A let ?P = "(\x. x i) -` A \ space I.P" - assume A: "A \ sets (M i)" - then have *: "?P = {i} \\<^isub>E A" using sets_into_space by auto - show "I.\ ?P = M.\ i A" unfolding * - using A I.measure_times[of "\_. A"] by auto -qed auto - -lemma (in product_sigma_finite) product_positive_integral_singleton: - assumes f: "f \ borel_measurable (M i)" - shows "integral\<^isup>P (Pi\<^isub>M {i} M) (\x. f (x i)) = integral\<^isup>P (M i) f" -proof - - interpret I: finite_product_sigma_finite M "{i}" by default simp - show ?thesis - proof (rule I.positive_integral_vimage[symmetric]) - show "sigma_algebra (M i)" by (rule sigma_algebras) - show "(\x. x i) \ measure_preserving (Pi\<^isub>M {i} M) (M i)" - by (rule measure_preserving_component_singelton) - show "f \ borel_measurable (M i)" by fact - qed -qed - -lemma (in product_sigma_finite) product_positive_integral_insert: - assumes [simp]: "finite I" "i \ I" - and f: "f \ borel_measurable (Pi\<^isub>M (insert i I) M)" - shows "integral\<^isup>P (Pi\<^isub>M (insert i I) M) f = (\\<^isup>+ x. (\\<^isup>+ y. f (x(i := y)) \(M i)) \(Pi\<^isub>M I M))" -proof - - interpret I: finite_product_sigma_finite M I by default auto - interpret i: finite_product_sigma_finite M "{i}" by default auto - interpret P: pair_sigma_algebra I.P i.P .. - have IJ: "I \ {i} = {}" and insert: "I \ {i} = insert i I" - using f by auto - show ?thesis - unfolding product_positive_integral_fold[OF IJ, unfolded insert, simplified, OF f] - proof (rule I.positive_integral_cong, subst product_positive_integral_singleton) - fix x assume x: "x \ space I.P" - let "?f y" = "f (restrict (x(i := y)) (insert i I))" - have f'_eq: "\y. ?f y = f (x(i := y))" - using x by (auto intro!: arg_cong[where f=f] simp: fun_eq_iff extensional_def) - show "?f \ borel_measurable (M i)" unfolding f'_eq - using measurable_comp[OF measurable_component_update f] x `i \ I` - by (simp add: comp_def) - show "integral\<^isup>P (M i) ?f = \\<^isup>+ y. f (x(i:=y)) \M i" - unfolding f'_eq by simp - qed -qed - -lemma (in product_sigma_finite) product_positive_integral_setprod: - fixes f :: "'i \ 'a \ extreal" - assumes "finite I" and borel: "\i. i \ I \ f i \ borel_measurable (M i)" - and pos: "\i x. i \ I \ 0 \ f i x" - shows "(\\<^isup>+ x. (\i\I. f i (x i)) \Pi\<^isub>M I M) = (\i\I. integral\<^isup>P (M i) (f i))" -using assms proof induct - case empty - interpret finite_product_sigma_finite M "{}" by default auto - then show ?case by simp -next - case (insert i I) - note `finite I`[intro, simp] - interpret I: finite_product_sigma_finite M I by default auto - have *: "\x y. (\j\I. f j (if j = i then y else x j)) = (\j\I. f j (x j))" - using insert by (auto intro!: setprod_cong) - have prod: "\J. J \ insert i I \ (\x. (\i\J. f i (x i))) \ borel_measurable (Pi\<^isub>M J M)" - using sets_into_space insert - by (intro sigma_algebra.borel_measurable_extreal_setprod sigma_algebra_product_algebra - measurable_comp[OF measurable_component_singleton, unfolded comp_def]) - auto - then show ?case - apply (simp add: product_positive_integral_insert[OF insert(1,2) prod]) - apply (simp add: insert * pos borel setprod_extreal_pos M.positive_integral_multc) - apply (subst I.positive_integral_cmult) - apply (auto simp add: pos borel insert setprod_extreal_pos positive_integral_positive) - done -qed - -lemma (in product_sigma_finite) product_integral_singleton: - assumes f: "f \ borel_measurable (M i)" - shows "(\x. f (x i) \Pi\<^isub>M {i} M) = integral\<^isup>L (M i) f" -proof - - interpret I: finite_product_sigma_finite M "{i}" by default simp - have *: "(\x. extreal (f x)) \ borel_measurable (M i)" - "(\x. extreal (- f x)) \ borel_measurable (M i)" - using assms by auto - show ?thesis - unfolding lebesgue_integral_def *[THEN product_positive_integral_singleton] .. -qed - -lemma (in product_sigma_finite) product_integral_fold: - assumes IJ[simp]: "I \ J = {}" and fin: "finite I" "finite J" - and f: "integrable (Pi\<^isub>M (I \ J) M) f" - shows "integral\<^isup>L (Pi\<^isub>M (I \ J) M) f = (\x. (\y. f (merge I x J y) \Pi\<^isub>M J M) \Pi\<^isub>M I M)" -proof - - interpret I: finite_product_sigma_finite M I by default fact - interpret J: finite_product_sigma_finite M J by default fact - have "finite (I \ J)" using fin by auto - interpret IJ: finite_product_sigma_finite M "I \ J" by default fact - interpret P: pair_sigma_finite I.P J.P by default - let ?M = "\(x, y). merge I x J y" - let ?f = "\x. f (?M x)" - show ?thesis - proof (subst P.integrable_fst_measurable(2)[of ?f, simplified]) - have 1: "sigma_algebra IJ.P" by default - have 2: "?M \ measure_preserving P.P IJ.P" using measure_preserving_merge[OF assms(1,2,3)] . - have 3: "integrable (Pi\<^isub>M (I \ J) M) f" by fact - then have 4: "f \ borel_measurable (Pi\<^isub>M (I \ J) M)" - by (simp add: integrable_def) - show "integrable P.P ?f" - by (rule P.integrable_vimage[where f=f, OF 1 2 3]) - show "integral\<^isup>L IJ.P f = integral\<^isup>L P.P ?f" - by (rule P.integral_vimage[where f=f, OF 1 2 4]) - qed -qed - -lemma (in product_sigma_finite) product_integral_insert: - assumes [simp]: "finite I" "i \ I" - and f: "integrable (Pi\<^isub>M (insert i I) M) f" - shows "integral\<^isup>L (Pi\<^isub>M (insert i I) M) f = (\x. (\y. f (x(i:=y)) \M i) \Pi\<^isub>M I M)" -proof - - interpret I: finite_product_sigma_finite M I by default auto - interpret I': finite_product_sigma_finite M "insert i I" by default auto - interpret i: finite_product_sigma_finite M "{i}" by default auto - interpret P: pair_sigma_finite I.P i.P .. - have IJ: "I \ {i} = {}" by auto - show ?thesis - unfolding product_integral_fold[OF IJ, simplified, OF f] - proof (rule I.integral_cong, subst product_integral_singleton) - fix x assume x: "x \ space I.P" - let "?f y" = "f (restrict (x(i := y)) (insert i I))" - have f'_eq: "\y. ?f y = f (x(i := y))" - using x by (auto intro!: arg_cong[where f=f] simp: fun_eq_iff extensional_def) - have f: "f \ borel_measurable I'.P" using f unfolding integrable_def by auto - show "?f \ borel_measurable (M i)" - unfolding measurable_cong[OF f'_eq] - using measurable_comp[OF measurable_component_update f] x `i \ I` - by (simp add: comp_def) - show "integral\<^isup>L (M i) ?f = integral\<^isup>L (M i) (\y. f (x(i := y)))" - unfolding f'_eq by simp - qed -qed - -lemma (in product_sigma_finite) product_integrable_setprod: - fixes f :: "'i \ 'a \ real" - assumes [simp]: "finite I" and integrable: "\i. i \ I \ integrable (M i) (f i)" - shows "integrable (Pi\<^isub>M I M) (\x. (\i\I. f i (x i)))" (is "integrable _ ?f") -proof - - interpret finite_product_sigma_finite M I by default fact - have f: "\i. i \ I \ f i \ borel_measurable (M i)" - using integrable unfolding integrable_def by auto - then have borel: "?f \ borel_measurable P" - using measurable_comp[OF measurable_component_singleton f] - by (auto intro!: borel_measurable_setprod simp: comp_def) - moreover have "integrable (Pi\<^isub>M I M) (\x. \\i\I. f i (x i)\)" - proof (unfold integrable_def, intro conjI) - show "(\x. abs (?f x)) \ borel_measurable P" - using borel by auto - have "(\\<^isup>+x. extreal (abs (?f x)) \P) = (\\<^isup>+x. (\i\I. extreal (abs (f i (x i)))) \P)" - by (simp add: setprod_extreal abs_setprod) - also have "\ = (\i\I. (\\<^isup>+x. extreal (abs (f i x)) \M i))" - using f by (subst product_positive_integral_setprod) auto - also have "\ < \" - using integrable[THEN M.integrable_abs] - by (simp add: setprod_PInf integrable_def M.positive_integral_positive) - finally show "(\\<^isup>+x. extreal (abs (?f x)) \P) \ \" by auto - have "(\\<^isup>+x. extreal (- abs (?f x)) \P) = (\\<^isup>+x. 0 \P)" - by (intro positive_integral_cong_pos) auto - then show "(\\<^isup>+x. extreal (- abs (?f x)) \P) \ \" by simp - qed - ultimately show ?thesis - by (rule integrable_abs_iff[THEN iffD1]) -qed - -lemma (in product_sigma_finite) product_integral_setprod: - fixes f :: "'i \ 'a \ real" - assumes "finite I" "I \ {}" and integrable: "\i. i \ I \ integrable (M i) (f i)" - shows "(\x. (\i\I. f i (x i)) \Pi\<^isub>M I M) = (\i\I. integral\<^isup>L (M i) (f i))" -using assms proof (induct rule: finite_ne_induct) - case (singleton i) - then show ?case by (simp add: product_integral_singleton integrable_def) -next - case (insert i I) - then have iI: "finite (insert i I)" by auto - then have prod: "\J. J \ insert i I \ - integrable (Pi\<^isub>M J M) (\x. (\i\J. f i (x i)))" - by (intro product_integrable_setprod insert(5)) (auto intro: finite_subset) - interpret I: finite_product_sigma_finite M I by default fact - have *: "\x y. (\j\I. f j (if j = i then y else x j)) = (\j\I. f j (x j))" - using `i \ I` by (auto intro!: setprod_cong) - show ?case - unfolding product_integral_insert[OF insert(1,3) prod[OF subset_refl]] - by (simp add: * insert integral_multc I.integral_cmult[OF prod] subset_insertI) -qed - -section "Products on finite spaces" - -lemma sigma_sets_pair_measure_generator_finite: - assumes "finite A" and "finite B" - shows "sigma_sets (A \ B) { a \ b | a b. a \ Pow A \ b \ Pow B} = Pow (A \ B)" - (is "sigma_sets ?prod ?sets = _") -proof safe - have fin: "finite (A \ B)" using assms by (rule finite_cartesian_product) - fix x assume subset: "x \ A \ B" - hence "finite x" using fin by (rule finite_subset) - from this subset show "x \ sigma_sets ?prod ?sets" - proof (induct x) - case empty show ?case by (rule sigma_sets.Empty) - next - case (insert a x) - hence "{a} \ sigma_sets ?prod ?sets" - by (auto simp: pair_measure_generator_def intro!: sigma_sets.Basic) - moreover have "x \ sigma_sets ?prod ?sets" using insert by auto - ultimately show ?case unfolding insert_is_Un[of a x] by (rule sigma_sets_Un) - qed -next - fix x a b - assume "x \ sigma_sets ?prod ?sets" and "(a, b) \ x" - from sigma_sets_into_sp[OF _ this(1)] this(2) - show "a \ A" and "b \ B" by auto -qed - -locale pair_finite_sigma_algebra = M1: finite_sigma_algebra M1 + M2: finite_sigma_algebra M2 - for M1 :: "('a, 'c) measure_space_scheme" and M2 :: "('b, 'd) measure_space_scheme" - -sublocale pair_finite_sigma_algebra \ pair_sigma_algebra by default - -lemma (in pair_finite_sigma_algebra) finite_pair_sigma_algebra: - shows "P = \ space = space M1 \ space M2, sets = Pow (space M1 \ space M2), \ = algebra.more P \" -proof - - show ?thesis - using sigma_sets_pair_measure_generator_finite[OF M1.finite_space M2.finite_space] - by (intro algebra.equality) (simp_all add: pair_measure_def pair_measure_generator_def sigma_def) -qed - -sublocale pair_finite_sigma_algebra \ finite_sigma_algebra P - apply default - using M1.finite_space M2.finite_space - apply (subst finite_pair_sigma_algebra) apply simp - apply (subst (1 2) finite_pair_sigma_algebra) apply simp - done - -locale pair_finite_space = M1: finite_measure_space M1 + M2: finite_measure_space M2 - for M1 M2 - -sublocale pair_finite_space \ pair_finite_sigma_algebra - by default - -sublocale pair_finite_space \ pair_sigma_finite - by default - -lemma (in pair_finite_space) pair_measure_Pair[simp]: - assumes "a \ space M1" "b \ space M2" - shows "\ {(a, b)} = M1.\ {a} * M2.\ {b}" -proof - - have "\ ({a}\{b}) = M1.\ {a} * M2.\ {b}" - using M1.sets_eq_Pow M2.sets_eq_Pow assms - by (subst pair_measure_times) auto - then show ?thesis by simp -qed - -lemma (in pair_finite_space) pair_measure_singleton[simp]: - assumes "x \ space M1 \ space M2" - shows "\ {x} = M1.\ {fst x} * M2.\ {snd x}" - using pair_measure_Pair assms by (cases x) auto - -sublocale pair_finite_space \ finite_measure_space P - by default (auto simp: space_pair_measure) - -end \ No newline at end of file