src/HOL/Probability/Binary_Product_Measure.thy
author wenzelm
Tue, 13 Mar 2012 13:31:26 +0100
changeset 47770 1570b30ee040
parent 47605 5302e932d1e5
child 48565 05663f75964c
permissions -rw-r--r--
tuned proofs -- eliminated pointless chaining of facts after 'interpret';
     1 (*  Title:      HOL/Probability/Binary_Product_Measure.thy
     2     Author:     Johannes Hölzl, TU München
     3 *)
     4 
     5 header {*Binary product measures*}
     6 
     7 theory Binary_Product_Measure
     8 imports Lebesgue_Integration
     9 begin
    10 
    11 lemma times_eq_iff: "A \<times> B = C \<times> D \<longleftrightarrow> A = C \<and> B = D \<or> ((A = {} \<or> B = {}) \<and> (C = {} \<or> D = {}))"
    12   by auto
    13 
    14 lemma times_Int_times: "A \<times> B \<inter> C \<times> D = (A \<inter> C) \<times> (B \<inter> D)"
    15   by auto
    16 
    17 lemma Pair_vimage_times[simp]: "\<And>A B x. Pair x -` (A \<times> B) = (if x \<in> A then B else {})"
    18   by auto
    19 
    20 lemma rev_Pair_vimage_times[simp]: "\<And>A B y. (\<lambda>x. (x, y)) -` (A \<times> B) = (if y \<in> B then A else {})"
    21   by auto
    22 
    23 lemma case_prod_distrib: "f (case x of (x, y) \<Rightarrow> g x y) = (case x of (x, y) \<Rightarrow> f (g x y))"
    24   by (cases x) simp
    25 
    26 lemma split_beta': "(\<lambda>(x,y). f x y) = (\<lambda>x. f (fst x) (snd x))"
    27   by (auto simp: fun_eq_iff)
    28 
    29 section "Binary products"
    30 
    31 definition
    32   "pair_measure_generator A B =
    33     \<lparr> space = space A \<times> space B,
    34       sets = {a \<times> b | a b. a \<in> sets A \<and> b \<in> sets B},
    35       measure = \<lambda>X. \<integral>\<^isup>+x. (\<integral>\<^isup>+y. indicator X (x,y) \<partial>B) \<partial>A \<rparr>"
    36 
    37 definition pair_measure (infixr "\<Otimes>\<^isub>M" 80) where
    38   "A \<Otimes>\<^isub>M B = sigma (pair_measure_generator A B)"
    39 
    40 locale pair_sigma_algebra = M1: sigma_algebra M1 + M2: sigma_algebra M2
    41   for M1 :: "('a, 'c) measure_space_scheme" and M2 :: "('b, 'd) measure_space_scheme"
    42 
    43 abbreviation (in pair_sigma_algebra)
    44   "E \<equiv> pair_measure_generator M1 M2"
    45 
    46 abbreviation (in pair_sigma_algebra)
    47   "P \<equiv> M1 \<Otimes>\<^isub>M M2"
    48 
    49 lemma sigma_algebra_pair_measure:
    50   "sets M1 \<subseteq> Pow (space M1) \<Longrightarrow> sets M2 \<subseteq> Pow (space M2) \<Longrightarrow> sigma_algebra (pair_measure M1 M2)"
    51   by (force simp: pair_measure_def pair_measure_generator_def intro!: sigma_algebra_sigma)
    52 
    53 sublocale pair_sigma_algebra \<subseteq> sigma_algebra P
    54   using M1.space_closed M2.space_closed
    55   by (rule sigma_algebra_pair_measure)
    56 
    57 lemma pair_measure_generatorI[intro, simp]:
    58   "x \<in> sets A \<Longrightarrow> y \<in> sets B \<Longrightarrow> x \<times> y \<in> sets (pair_measure_generator A B)"
    59   by (auto simp add: pair_measure_generator_def)
    60 
    61 lemma pair_measureI[intro, simp]:
    62   "x \<in> sets A \<Longrightarrow> y \<in> sets B \<Longrightarrow> x \<times> y \<in> sets (A \<Otimes>\<^isub>M B)"
    63   by (auto simp add: pair_measure_def)
    64 
    65 lemma space_pair_measure:
    66   "space (A \<Otimes>\<^isub>M B) = space A \<times> space B"
    67   by (simp add: pair_measure_def pair_measure_generator_def)
    68 
    69 lemma sets_pair_measure_generator:
    70   "sets (pair_measure_generator N M) = (\<lambda>(x, y). x \<times> y) ` (sets N \<times> sets M)"
    71   unfolding pair_measure_generator_def by auto
    72 
    73 lemma pair_measure_generator_sets_into_space:
    74   assumes "sets M \<subseteq> Pow (space M)" "sets N \<subseteq> Pow (space N)"
    75   shows "sets (pair_measure_generator M N) \<subseteq> Pow (space (pair_measure_generator M N))"
    76   using assms by (auto simp: pair_measure_generator_def)
    77 
    78 lemma pair_measure_generator_Int_snd:
    79   assumes "sets S1 \<subseteq> Pow (space S1)"
    80   shows "sets (pair_measure_generator S1 (algebra.restricted_space S2 A)) =
    81          sets (algebra.restricted_space (pair_measure_generator S1 S2) (space S1 \<times> A))"
    82   (is "?L = ?R")
    83   apply (auto simp: pair_measure_generator_def image_iff)
    84   using assms
    85   apply (rule_tac x="a \<times> xa" in exI)
    86   apply force
    87   using assms
    88   apply (rule_tac x="a" in exI)
    89   apply (rule_tac x="b \<inter> A" in exI)
    90   apply auto
    91   done
    92 
    93 lemma (in pair_sigma_algebra)
    94   shows measurable_fst[intro!, simp]:
    95     "fst \<in> measurable P M1" (is ?fst)
    96   and measurable_snd[intro!, simp]:
    97     "snd \<in> measurable P M2" (is ?snd)
    98 proof -
    99   { fix X assume "X \<in> sets M1"
   100     then have "\<exists>X1\<in>sets M1. \<exists>X2\<in>sets M2. fst -` X \<inter> space M1 \<times> space M2 = X1 \<times> X2"
   101       apply - apply (rule bexI[of _ X]) apply (rule bexI[of _ "space M2"])
   102       using M1.sets_into_space by force+ }
   103   moreover
   104   { fix X assume "X \<in> sets M2"
   105     then have "\<exists>X1\<in>sets M1. \<exists>X2\<in>sets M2. snd -` X \<inter> space M1 \<times> space M2 = X1 \<times> X2"
   106       apply - apply (rule bexI[of _ "space M1"]) apply (rule bexI[of _ X])
   107       using M2.sets_into_space by force+ }
   108   ultimately have "?fst \<and> ?snd"
   109     by (fastforce simp: measurable_def sets_sigma space_pair_measure
   110                  intro!: sigma_sets.Basic)
   111   then show ?fst ?snd by auto
   112 qed
   113 
   114 lemma (in pair_sigma_algebra) measurable_pair_iff:
   115   assumes "sigma_algebra M"
   116   shows "f \<in> measurable M P \<longleftrightarrow>
   117     (fst \<circ> f) \<in> measurable M M1 \<and> (snd \<circ> f) \<in> measurable M M2"
   118 proof -
   119   interpret M: sigma_algebra M by fact
   120   from assms show ?thesis
   121   proof (safe intro!: measurable_comp[where b=P])
   122     assume f: "(fst \<circ> f) \<in> measurable M M1" and s: "(snd \<circ> f) \<in> measurable M M2"
   123     show "f \<in> measurable M P" unfolding pair_measure_def
   124     proof (rule M.measurable_sigma)
   125       show "sets (pair_measure_generator M1 M2) \<subseteq> Pow (space E)"
   126         unfolding pair_measure_generator_def using M1.sets_into_space M2.sets_into_space by auto
   127       show "f \<in> space M \<rightarrow> space E"
   128         using f s by (auto simp: mem_Times_iff measurable_def comp_def space_sigma pair_measure_generator_def)
   129       fix A assume "A \<in> sets E"
   130       then obtain B C where "B \<in> sets M1" "C \<in> sets M2" "A = B \<times> C"
   131         unfolding pair_measure_generator_def by auto
   132       moreover have "(fst \<circ> f) -` B \<inter> space M \<in> sets M"
   133         using f `B \<in> sets M1` unfolding measurable_def by auto
   134       moreover have "(snd \<circ> f) -` C \<inter> space M \<in> sets M"
   135         using s `C \<in> sets M2` unfolding measurable_def by auto
   136       moreover have "f -` A \<inter> space M = ((fst \<circ> f) -` B \<inter> space M) \<inter> ((snd \<circ> f) -` C \<inter> space M)"
   137         unfolding `A = B \<times> C` by (auto simp: vimage_Times)
   138       ultimately show "f -` A \<inter> space M \<in> sets M" by auto
   139     qed
   140   qed
   141 qed
   142 
   143 lemma (in pair_sigma_algebra) measurable_pair:
   144   assumes "sigma_algebra M"
   145   assumes "(fst \<circ> f) \<in> measurable M M1" "(snd \<circ> f) \<in> measurable M M2"
   146   shows "f \<in> measurable M P"
   147   unfolding measurable_pair_iff[OF assms(1)] using assms(2,3) by simp
   148 
   149 lemma pair_measure_generatorE:
   150   assumes "X \<in> sets (pair_measure_generator M1 M2)"
   151   obtains A B where "X = A \<times> B" "A \<in> sets M1" "B \<in> sets M2"
   152   using assms unfolding pair_measure_generator_def by auto
   153 
   154 lemma (in pair_sigma_algebra) pair_measure_generator_swap:
   155   "(\<lambda>X. (\<lambda>(x,y). (y,x)) -` X \<inter> space M2 \<times> space M1) ` sets E = sets (pair_measure_generator M2 M1)"
   156 proof (safe elim!: pair_measure_generatorE)
   157   fix A B assume "A \<in> sets M1" "B \<in> sets M2"
   158   moreover then have "(\<lambda>(x, y). (y, x)) -` (A \<times> B) \<inter> space M2 \<times> space M1 = B \<times> A"
   159     using M1.sets_into_space M2.sets_into_space by auto
   160   ultimately show "(\<lambda>(x, y). (y, x)) -` (A \<times> B) \<inter> space M2 \<times> space M1 \<in> sets (pair_measure_generator M2 M1)"
   161     by (auto intro: pair_measure_generatorI)
   162 next
   163   fix A B assume "A \<in> sets M1" "B \<in> sets M2"
   164   then show "B \<times> A \<in> (\<lambda>X. (\<lambda>(x, y). (y, x)) -` X \<inter> space M2 \<times> space M1) ` sets E"
   165     using M1.sets_into_space M2.sets_into_space
   166     by (auto intro!: image_eqI[where x="A \<times> B"] pair_measure_generatorI)
   167 qed
   168 
   169 lemma (in pair_sigma_algebra) sets_pair_sigma_algebra_swap:
   170   assumes Q: "Q \<in> sets P"
   171   shows "(\<lambda>(x,y). (y, x)) -` Q \<in> sets (M2 \<Otimes>\<^isub>M M1)" (is "_ \<in> sets ?Q")
   172 proof -
   173   let ?f = "\<lambda>Q. (\<lambda>(x,y). (y, x)) -` Q \<inter> space M2 \<times> space M1"
   174   have *: "(\<lambda>(x,y). (y, x)) -` Q = ?f Q"
   175     using sets_into_space[OF Q] by (auto simp: space_pair_measure)
   176   have "sets (M2 \<Otimes>\<^isub>M M1) = sets (sigma (pair_measure_generator M2 M1))"
   177     unfolding pair_measure_def ..
   178   also have "\<dots> = sigma_sets (space M2 \<times> space M1) (?f ` sets E)"
   179     unfolding sigma_def pair_measure_generator_swap[symmetric]
   180     by (simp add: pair_measure_generator_def)
   181   also have "\<dots> = ?f ` sigma_sets (space M1 \<times> space M2) (sets E)"
   182     using M1.sets_into_space M2.sets_into_space
   183     by (intro sigma_sets_vimage) (auto simp: pair_measure_generator_def)
   184   also have "\<dots> = ?f ` sets P"
   185     unfolding pair_measure_def pair_measure_generator_def sigma_def by simp
   186   finally show ?thesis
   187     using Q by (subst *) auto
   188 qed
   189 
   190 lemma (in pair_sigma_algebra) pair_sigma_algebra_swap_measurable:
   191   shows "(\<lambda>(x,y). (y, x)) \<in> measurable P (M2 \<Otimes>\<^isub>M M1)"
   192     (is "?f \<in> measurable ?P ?Q")
   193   unfolding measurable_def
   194 proof (intro CollectI conjI Pi_I ballI)
   195   fix x assume "x \<in> space ?P" then show "(case x of (x, y) \<Rightarrow> (y, x)) \<in> space ?Q"
   196     unfolding pair_measure_generator_def pair_measure_def by auto
   197 next
   198   fix A assume "A \<in> sets (M2 \<Otimes>\<^isub>M M1)"
   199   interpret Q: pair_sigma_algebra M2 M1 by default
   200   from Q.sets_pair_sigma_algebra_swap[OF `A \<in> sets (M2 \<Otimes>\<^isub>M M1)`]
   201   show "?f -` A \<inter> space ?P \<in> sets ?P" by simp
   202 qed
   203 
   204 lemma (in pair_sigma_algebra) measurable_cut_fst[simp,intro]:
   205   assumes "Q \<in> sets P" shows "Pair x -` Q \<in> sets M2"
   206 proof -
   207   let ?Q' = "{Q. Q \<subseteq> space P \<and> Pair x -` Q \<in> sets M2}"
   208   let ?Q = "\<lparr> space = space P, sets = ?Q' \<rparr>"
   209   interpret Q: sigma_algebra ?Q
   210     proof qed (auto simp: vimage_UN vimage_Diff space_pair_measure)
   211   have "sets E \<subseteq> sets ?Q"
   212     using M1.sets_into_space M2.sets_into_space
   213     by (auto simp: pair_measure_generator_def space_pair_measure)
   214   then have "sets P \<subseteq> sets ?Q"
   215     apply (subst pair_measure_def, intro Q.sets_sigma_subset)
   216     by (simp add: pair_measure_def)
   217   with assms show ?thesis by auto
   218 qed
   219 
   220 lemma (in pair_sigma_algebra) measurable_cut_snd:
   221   assumes Q: "Q \<in> sets P" shows "(\<lambda>x. (x, y)) -` Q \<in> sets M1" (is "?cut Q \<in> sets M1")
   222 proof -
   223   interpret Q: pair_sigma_algebra M2 M1 by default
   224   from Q.measurable_cut_fst[OF sets_pair_sigma_algebra_swap[OF Q], of y]
   225   show ?thesis by (simp add: vimage_compose[symmetric] comp_def)
   226 qed
   227 
   228 lemma (in pair_sigma_algebra) measurable_pair_image_snd:
   229   assumes m: "f \<in> measurable P M" and "x \<in> space M1"
   230   shows "(\<lambda>y. f (x, y)) \<in> measurable M2 M"
   231   unfolding measurable_def
   232 proof (intro CollectI conjI Pi_I ballI)
   233   fix y assume "y \<in> space M2" with `f \<in> measurable P M` `x \<in> space M1`
   234   show "f (x, y) \<in> space M"
   235     unfolding measurable_def pair_measure_generator_def pair_measure_def by auto
   236 next
   237   fix A assume "A \<in> sets M"
   238   then have "Pair x -` (f -` A \<inter> space P) \<in> sets M2" (is "?C \<in> _")
   239     using `f \<in> measurable P M`
   240     by (intro measurable_cut_fst) (auto simp: measurable_def)
   241   also have "?C = (\<lambda>y. f (x, y)) -` A \<inter> space M2"
   242     using `x \<in> space M1` by (auto simp: pair_measure_generator_def pair_measure_def)
   243   finally show "(\<lambda>y. f (x, y)) -` A \<inter> space M2 \<in> sets M2" .
   244 qed
   245 
   246 lemma (in pair_sigma_algebra) measurable_pair_image_fst:
   247   assumes m: "f \<in> measurable P M" and "y \<in> space M2"
   248   shows "(\<lambda>x. f (x, y)) \<in> measurable M1 M"
   249 proof -
   250   interpret Q: pair_sigma_algebra M2 M1 by default
   251   from Q.measurable_pair_image_snd[OF measurable_comp `y \<in> space M2`,
   252                                       OF Q.pair_sigma_algebra_swap_measurable m]
   253   show ?thesis by simp
   254 qed
   255 
   256 lemma (in pair_sigma_algebra) Int_stable_pair_measure_generator: "Int_stable E"
   257   unfolding Int_stable_def
   258 proof (intro ballI)
   259   fix A B assume "A \<in> sets E" "B \<in> sets E"
   260   then obtain A1 A2 B1 B2 where "A = A1 \<times> A2" "B = B1 \<times> B2"
   261     "A1 \<in> sets M1" "A2 \<in> sets M2" "B1 \<in> sets M1" "B2 \<in> sets M2"
   262     unfolding pair_measure_generator_def by auto
   263   then show "A \<inter> B \<in> sets E"
   264     by (auto simp add: times_Int_times pair_measure_generator_def)
   265 qed
   266 
   267 lemma finite_measure_cut_measurable:
   268   fixes M1 :: "('a, 'c) measure_space_scheme" and M2 :: "('b, 'd) measure_space_scheme"
   269   assumes "sigma_finite_measure M1" "finite_measure M2"
   270   assumes "Q \<in> sets (M1 \<Otimes>\<^isub>M M2)"
   271   shows "(\<lambda>x. measure M2 (Pair x -` Q)) \<in> borel_measurable M1"
   272     (is "?s Q \<in> _")
   273 proof -
   274   interpret M1: sigma_finite_measure M1 by fact
   275   interpret M2: finite_measure M2 by fact
   276   interpret pair_sigma_algebra M1 M2 by default
   277   have [intro]: "sigma_algebra M1" by fact
   278   have [intro]: "sigma_algebra M2" by fact
   279   let ?D = "\<lparr> space = space P, sets = {A\<in>sets P. ?s A \<in> borel_measurable M1}  \<rparr>"
   280   note space_pair_measure[simp]
   281   interpret dynkin_system ?D
   282   proof (intro dynkin_systemI)
   283     fix A assume "A \<in> sets ?D" then show "A \<subseteq> space ?D"
   284       using sets_into_space by simp
   285   next
   286     from top show "space ?D \<in> sets ?D"
   287       by (auto simp add: if_distrib intro!: M1.measurable_If)
   288   next
   289     fix A assume "A \<in> sets ?D"
   290     with sets_into_space have "\<And>x. measure M2 (Pair x -` (space M1 \<times> space M2 - A)) =
   291         (if x \<in> space M1 then measure M2 (space M2) - ?s A x else 0)"
   292       by (auto intro!: M2.measure_compl simp: vimage_Diff)
   293     with `A \<in> sets ?D` top show "space ?D - A \<in> sets ?D"
   294       by (auto intro!: Diff M1.measurable_If M1.borel_measurable_ereal_diff)
   295   next
   296     fix F :: "nat \<Rightarrow> ('a\<times>'b) set" assume "disjoint_family F" "range F \<subseteq> sets ?D"
   297     moreover then have "\<And>x. measure M2 (\<Union>i. Pair x -` F i) = (\<Sum>i. ?s (F i) x)"
   298       by (intro M2.measure_countably_additive[symmetric])
   299          (auto simp: disjoint_family_on_def)
   300     ultimately show "(\<Union>i. F i) \<in> sets ?D"
   301       by (auto simp: vimage_UN intro!: M1.borel_measurable_psuminf)
   302   qed
   303   have "sets P = sets ?D" apply (subst pair_measure_def)
   304   proof (intro dynkin_lemma)
   305     show "Int_stable E" by (rule Int_stable_pair_measure_generator)
   306     from M1.sets_into_space have "\<And>A. A \<in> sets M1 \<Longrightarrow> {x \<in> space M1. x \<in> A} = A"
   307       by auto
   308     then show "sets E \<subseteq> sets ?D"
   309       by (auto simp: pair_measure_generator_def sets_sigma if_distrib
   310                intro: sigma_sets.Basic intro!: M1.measurable_If)
   311   qed (auto simp: pair_measure_def)
   312   with `Q \<in> sets P` have "Q \<in> sets ?D" by simp
   313   then show "?s Q \<in> borel_measurable M1" by simp
   314 qed
   315 
   316 subsection {* Binary products of $\sigma$-finite measure spaces *}
   317 
   318 locale pair_sigma_finite = pair_sigma_algebra M1 M2 + M1: sigma_finite_measure M1 + M2: sigma_finite_measure M2
   319   for M1 :: "('a, 'c) measure_space_scheme" and M2 :: "('b, 'd) measure_space_scheme"
   320 
   321 lemma (in pair_sigma_finite) measure_cut_measurable_fst:
   322   assumes "Q \<in> sets P" shows "(\<lambda>x. measure M2 (Pair x -` Q)) \<in> borel_measurable M1" (is "?s Q \<in> _")
   323 proof -
   324   have [intro]: "sigma_algebra M1" and [intro]: "sigma_algebra M2" by default+
   325   have M1: "sigma_finite_measure M1" by default
   326   from M2.disjoint_sigma_finite guess F .. note F = this
   327   then have F_sets: "\<And>i. F i \<in> sets M2" by auto
   328   let ?C = "\<lambda>x i. F i \<inter> Pair x -` Q"
   329   { fix i
   330     let ?R = "M2.restricted_space (F i)"
   331     have [simp]: "space M1 \<times> F i \<inter> space M1 \<times> space M2 = space M1 \<times> F i"
   332       using F M2.sets_into_space by auto
   333     let ?R2 = "M2.restricted_space (F i)"
   334     have "(\<lambda>x. measure ?R2 (Pair x -` (space M1 \<times> space ?R2 \<inter> Q))) \<in> borel_measurable M1"
   335     proof (intro finite_measure_cut_measurable[OF M1])
   336       show "finite_measure ?R2"
   337         using F by (intro M2.restricted_to_finite_measure) auto
   338       have "(space M1 \<times> space ?R2) \<inter> Q \<in> (op \<inter> (space M1 \<times> F i)) ` sets P"
   339         using `Q \<in> sets P` by (auto simp: image_iff)
   340       also have "\<dots> = sigma_sets (space M1 \<times> F i) ((op \<inter> (space M1 \<times> F i)) ` sets E)"
   341         unfolding pair_measure_def pair_measure_generator_def sigma_def
   342         using `F i \<in> sets M2` M2.sets_into_space
   343         by (auto intro!: sigma_sets_Int sigma_sets.Basic)
   344       also have "\<dots> \<subseteq> sets (M1 \<Otimes>\<^isub>M ?R2)"
   345         using M1.sets_into_space
   346         apply (auto simp: times_Int_times pair_measure_def pair_measure_generator_def sigma_def
   347                     intro!: sigma_sets_subseteq)
   348         apply (rule_tac x="a" in exI)
   349         apply (rule_tac x="b \<inter> F i" in exI)
   350         by auto
   351       finally show "(space M1 \<times> space ?R2) \<inter> Q \<in> sets (M1 \<Otimes>\<^isub>M ?R2)" .
   352     qed
   353     moreover have "\<And>x. Pair x -` (space M1 \<times> F i \<inter> Q) = ?C x i"
   354       using `Q \<in> sets P` sets_into_space by (auto simp: space_pair_measure)
   355     ultimately have "(\<lambda>x. measure M2 (?C x i)) \<in> borel_measurable M1"
   356       by simp }
   357   moreover
   358   { fix x
   359     have "(\<Sum>i. measure M2 (?C x i)) = measure M2 (\<Union>i. ?C x i)"
   360     proof (intro M2.measure_countably_additive)
   361       show "range (?C x) \<subseteq> sets M2"
   362         using F `Q \<in> sets P` by (auto intro!: M2.Int)
   363       have "disjoint_family F" using F by auto
   364       show "disjoint_family (?C x)"
   365         by (rule disjoint_family_on_bisimulation[OF `disjoint_family F`]) auto
   366     qed
   367     also have "(\<Union>i. ?C x i) = Pair x -` Q"
   368       using F sets_into_space `Q \<in> sets P`
   369       by (auto simp: space_pair_measure)
   370     finally have "measure M2 (Pair x -` Q) = (\<Sum>i. measure M2 (?C x i))"
   371       by simp }
   372   ultimately show ?thesis using `Q \<in> sets P` F_sets
   373     by (auto intro!: M1.borel_measurable_psuminf M2.Int)
   374 qed
   375 
   376 lemma (in pair_sigma_finite) measure_cut_measurable_snd:
   377   assumes "Q \<in> sets P" shows "(\<lambda>y. M1.\<mu> ((\<lambda>x. (x, y)) -` Q)) \<in> borel_measurable M2"
   378 proof -
   379   interpret Q: pair_sigma_finite M2 M1 by default
   380   note sets_pair_sigma_algebra_swap[OF assms]
   381   from Q.measure_cut_measurable_fst[OF this]
   382   show ?thesis by (simp add: vimage_compose[symmetric] comp_def)
   383 qed
   384 
   385 lemma (in pair_sigma_algebra) pair_sigma_algebra_measurable:
   386   assumes "f \<in> measurable P M" shows "(\<lambda>(x,y). f (y, x)) \<in> measurable (M2 \<Otimes>\<^isub>M M1) M"
   387 proof -
   388   interpret Q: pair_sigma_algebra M2 M1 by default
   389   have *: "(\<lambda>(x,y). f (y, x)) = f \<circ> (\<lambda>(x,y). (y, x))" by (simp add: fun_eq_iff)
   390   show ?thesis
   391     using Q.pair_sigma_algebra_swap_measurable assms
   392     unfolding * by (rule measurable_comp)
   393 qed
   394 
   395 lemma (in pair_sigma_finite) pair_measure_alt:
   396   assumes "A \<in> sets P"
   397   shows "measure (M1 \<Otimes>\<^isub>M M2) A = (\<integral>\<^isup>+ x. measure M2 (Pair x -` A) \<partial>M1)"
   398   apply (simp add: pair_measure_def pair_measure_generator_def)
   399 proof (rule M1.positive_integral_cong)
   400   fix x assume "x \<in> space M1"
   401   have *: "\<And>y. indicator A (x, y) = (indicator (Pair x -` A) y :: ereal)"
   402     unfolding indicator_def by auto
   403   show "(\<integral>\<^isup>+ y. indicator A (x, y) \<partial>M2) = measure M2 (Pair x -` A)"
   404     unfolding *
   405     apply (subst M2.positive_integral_indicator)
   406     apply (rule measurable_cut_fst[OF assms])
   407     by simp
   408 qed
   409 
   410 lemma (in pair_sigma_finite) pair_measure_times:
   411   assumes A: "A \<in> sets M1" and "B \<in> sets M2"
   412   shows "measure (M1 \<Otimes>\<^isub>M M2) (A \<times> B) = M1.\<mu> A * measure M2 B"
   413 proof -
   414   have "measure (M1 \<Otimes>\<^isub>M M2) (A \<times> B) = (\<integral>\<^isup>+ x. measure M2 B * indicator A x \<partial>M1)"
   415     using assms by (auto intro!: M1.positive_integral_cong simp: pair_measure_alt)
   416   with assms show ?thesis
   417     by (simp add: M1.positive_integral_cmult_indicator ac_simps)
   418 qed
   419 
   420 lemma (in measure_space) measure_not_negative[simp,intro]:
   421   assumes A: "A \<in> sets M" shows "\<mu> A \<noteq> - \<infinity>"
   422   using positive_measure[OF A] by auto
   423 
   424 lemma (in pair_sigma_finite) sigma_finite_up_in_pair_measure_generator:
   425   "\<exists>F::nat \<Rightarrow> ('a \<times> 'b) set. range F \<subseteq> sets E \<and> incseq F \<and> (\<Union>i. F i) = space E \<and>
   426     (\<forall>i. measure (M1 \<Otimes>\<^isub>M M2) (F i) \<noteq> \<infinity>)"
   427 proof -
   428   obtain F1 :: "nat \<Rightarrow> 'a set" and F2 :: "nat \<Rightarrow> 'b set" where
   429     F1: "range F1 \<subseteq> sets M1" "incseq F1" "(\<Union>i. F1 i) = space M1" "\<And>i. M1.\<mu> (F1 i) \<noteq> \<infinity>" and
   430     F2: "range F2 \<subseteq> sets M2" "incseq F2" "(\<Union>i. F2 i) = space M2" "\<And>i. M2.\<mu> (F2 i) \<noteq> \<infinity>"
   431     using M1.sigma_finite_up M2.sigma_finite_up by auto
   432   then have space: "space M1 = (\<Union>i. F1 i)" "space M2 = (\<Union>i. F2 i)" by auto
   433   let ?F = "\<lambda>i. F1 i \<times> F2 i"
   434   show ?thesis unfolding space_pair_measure
   435   proof (intro exI[of _ ?F] conjI allI)
   436     show "range ?F \<subseteq> sets E" using F1 F2
   437       by (fastforce intro!: pair_measure_generatorI)
   438   next
   439     have "space M1 \<times> space M2 \<subseteq> (\<Union>i. ?F i)"
   440     proof (intro subsetI)
   441       fix x assume "x \<in> space M1 \<times> space M2"
   442       then obtain i j where "fst x \<in> F1 i" "snd x \<in> F2 j"
   443         by (auto simp: space)
   444       then have "fst x \<in> F1 (max i j)" "snd x \<in> F2 (max j i)"
   445         using `incseq F1` `incseq F2` unfolding incseq_def
   446         by (force split: split_max)+
   447       then have "(fst x, snd x) \<in> F1 (max i j) \<times> F2 (max i j)"
   448         by (intro SigmaI) (auto simp add: min_max.sup_commute)
   449       then show "x \<in> (\<Union>i. ?F i)" by auto
   450     qed
   451     then show "(\<Union>i. ?F i) = space E"
   452       using space by (auto simp: space pair_measure_generator_def)
   453   next
   454     fix i show "incseq (\<lambda>i. F1 i \<times> F2 i)"
   455       using `incseq F1` `incseq F2` unfolding incseq_Suc_iff by auto
   456   next
   457     fix i
   458     from F1 F2 have "F1 i \<in> sets M1" "F2 i \<in> sets M2" by auto
   459     with F1 F2 M1.positive_measure[OF this(1)] M2.positive_measure[OF this(2)]
   460     show "measure P (F1 i \<times> F2 i) \<noteq> \<infinity>"
   461       by (simp add: pair_measure_times)
   462   qed
   463 qed
   464 
   465 sublocale pair_sigma_finite \<subseteq> sigma_finite_measure P
   466 proof
   467   show "positive P (measure P)"
   468     unfolding pair_measure_def pair_measure_generator_def sigma_def positive_def
   469     by (auto intro: M1.positive_integral_positive M2.positive_integral_positive)
   470 
   471   show "countably_additive P (measure P)"
   472     unfolding countably_additive_def
   473   proof (intro allI impI)
   474     fix F :: "nat \<Rightarrow> ('a \<times> 'b) set"
   475     assume F: "range F \<subseteq> sets P" "disjoint_family F"
   476     from F have *: "\<And>i. F i \<in> sets P" "(\<Union>i. F i) \<in> sets P" by auto
   477     moreover from F have "\<And>i. (\<lambda>x. measure M2 (Pair x -` F i)) \<in> borel_measurable M1"
   478       by (intro measure_cut_measurable_fst) auto
   479     moreover have "\<And>x. disjoint_family (\<lambda>i. Pair x -` F i)"
   480       by (intro disjoint_family_on_bisimulation[OF F(2)]) auto
   481     moreover have "\<And>x. x \<in> space M1 \<Longrightarrow> range (\<lambda>i. Pair x -` F i) \<subseteq> sets M2"
   482       using F by auto
   483     ultimately show "(\<Sum>n. measure P (F n)) = measure P (\<Union>i. F i)"
   484       by (simp add: pair_measure_alt vimage_UN M1.positive_integral_suminf[symmetric]
   485                     M2.measure_countably_additive
   486                cong: M1.positive_integral_cong)
   487   qed
   488 
   489   from sigma_finite_up_in_pair_measure_generator guess F :: "nat \<Rightarrow> ('a \<times> 'b) set" .. note F = this
   490   show "\<exists>F::nat \<Rightarrow> ('a \<times> 'b) set. range F \<subseteq> sets P \<and> (\<Union>i. F i) = space P \<and> (\<forall>i. measure P (F i) \<noteq> \<infinity>)"
   491   proof (rule exI[of _ F], intro conjI)
   492     show "range F \<subseteq> sets P" using F by (auto simp: pair_measure_def)
   493     show "(\<Union>i. F i) = space P"
   494       using F by (auto simp: pair_measure_def pair_measure_generator_def)
   495     show "\<forall>i. measure P (F i) \<noteq> \<infinity>" using F by auto
   496   qed
   497 qed
   498 
   499 lemma (in pair_sigma_algebra) sets_swap:
   500   assumes "A \<in> sets P"
   501   shows "(\<lambda>(x, y). (y, x)) -` A \<inter> space (M2 \<Otimes>\<^isub>M M1) \<in> sets (M2 \<Otimes>\<^isub>M M1)"
   502     (is "_ -` A \<inter> space ?Q \<in> sets ?Q")
   503 proof -
   504   have *: "(\<lambda>(x, y). (y, x)) -` A \<inter> space ?Q = (\<lambda>(x, y). (y, x)) -` A"
   505     using `A \<in> sets P` sets_into_space by (auto simp: space_pair_measure)
   506   show ?thesis
   507     unfolding * using assms by (rule sets_pair_sigma_algebra_swap)
   508 qed
   509 
   510 lemma (in pair_sigma_finite) pair_measure_alt2:
   511   assumes A: "A \<in> sets P"
   512   shows "\<mu> A = (\<integral>\<^isup>+y. M1.\<mu> ((\<lambda>x. (x, y)) -` A) \<partial>M2)"
   513     (is "_ = ?\<nu> A")
   514 proof -
   515   interpret Q: pair_sigma_finite M2 M1 by default
   516   from sigma_finite_up_in_pair_measure_generator guess F :: "nat \<Rightarrow> ('a \<times> 'b) set" .. note F = this
   517   have [simp]: "\<And>m. \<lparr> space = space E, sets = sets (sigma E), measure = m \<rparr> = P\<lparr> measure := m \<rparr>"
   518     unfolding pair_measure_def by simp
   519 
   520   have "\<mu> A = Q.\<mu> ((\<lambda>(y, x). (x, y)) -` A \<inter> space Q.P)"
   521   proof (rule measure_unique_Int_stable_vimage[OF Int_stable_pair_measure_generator])
   522     show "measure_space P" "measure_space Q.P" by default
   523     show "(\<lambda>(y, x). (x, y)) \<in> measurable Q.P P" by (rule Q.pair_sigma_algebra_swap_measurable)
   524     show "sets (sigma E) = sets P" "space E = space P" "A \<in> sets (sigma E)"
   525       using assms unfolding pair_measure_def by auto
   526     show "range F \<subseteq> sets E" "incseq F" "(\<Union>i. F i) = space E" "\<And>i. \<mu> (F i) \<noteq> \<infinity>"
   527       using F `A \<in> sets P` by (auto simp: pair_measure_def)
   528     fix X assume "X \<in> sets E"
   529     then obtain A B where X[simp]: "X = A \<times> B" and AB: "A \<in> sets M1" "B \<in> sets M2"
   530       unfolding pair_measure_def pair_measure_generator_def by auto
   531     then have "(\<lambda>(y, x). (x, y)) -` X \<inter> space Q.P = B \<times> A"
   532       using M1.sets_into_space M2.sets_into_space by (auto simp: space_pair_measure)
   533     then show "\<mu> X = Q.\<mu> ((\<lambda>(y, x). (x, y)) -` X \<inter> space Q.P)"
   534       using AB by (simp add: pair_measure_times Q.pair_measure_times ac_simps)
   535   qed
   536   then show ?thesis
   537     using sets_into_space[OF A] Q.pair_measure_alt[OF sets_swap[OF A]]
   538     by (auto simp add: Q.pair_measure_alt space_pair_measure
   539              intro!: M2.positive_integral_cong arg_cong[where f="M1.\<mu>"])
   540 qed
   541 
   542 lemma pair_sigma_algebra_sigma:
   543   assumes 1: "incseq S1" "(\<Union>i. S1 i) = space E1" "range S1 \<subseteq> sets E1" and E1: "sets E1 \<subseteq> Pow (space E1)"
   544   assumes 2: "decseq S2" "(\<Union>i. S2 i) = space E2" "range S2 \<subseteq> sets E2" and E2: "sets E2 \<subseteq> Pow (space E2)"
   545   shows "sets (sigma (pair_measure_generator (sigma E1) (sigma E2))) = sets (sigma (pair_measure_generator E1 E2))"
   546     (is "sets ?S = sets ?E")
   547 proof -
   548   interpret M1: sigma_algebra "sigma E1" using E1 by (rule sigma_algebra_sigma)
   549   interpret M2: sigma_algebra "sigma E2" using E2 by (rule sigma_algebra_sigma)
   550   have P: "sets (pair_measure_generator E1 E2) \<subseteq> Pow (space E1 \<times> space E2)"
   551     using E1 E2 by (auto simp add: pair_measure_generator_def)
   552   interpret E: sigma_algebra ?E unfolding pair_measure_generator_def
   553     using E1 E2 by (intro sigma_algebra_sigma) auto
   554   { fix A assume "A \<in> sets E1"
   555     then have "fst -` A \<inter> space ?E = A \<times> (\<Union>i. S2 i)"
   556       using E1 2 unfolding pair_measure_generator_def by auto
   557     also have "\<dots> = (\<Union>i. A \<times> S2 i)" by auto
   558     also have "\<dots> \<in> sets ?E" unfolding pair_measure_generator_def sets_sigma
   559       using 2 `A \<in> sets E1`
   560       by (intro sigma_sets.Union)
   561          (force simp: image_subset_iff intro!: sigma_sets.Basic)
   562     finally have "fst -` A \<inter> space ?E \<in> sets ?E" . }
   563   moreover
   564   { fix B assume "B \<in> sets E2"
   565     then have "snd -` B \<inter> space ?E = (\<Union>i. S1 i) \<times> B"
   566       using E2 1 unfolding pair_measure_generator_def by auto
   567     also have "\<dots> = (\<Union>i. S1 i \<times> B)" by auto
   568     also have "\<dots> \<in> sets ?E"
   569       using 1 `B \<in> sets E2` unfolding pair_measure_generator_def sets_sigma
   570       by (intro sigma_sets.Union)
   571          (force simp: image_subset_iff intro!: sigma_sets.Basic)
   572     finally have "snd -` B \<inter> space ?E \<in> sets ?E" . }
   573   ultimately have proj:
   574     "fst \<in> measurable ?E (sigma E1) \<and> snd \<in> measurable ?E (sigma E2)"
   575     using E1 E2 by (subst (1 2) E.measurable_iff_sigma)
   576                    (auto simp: pair_measure_generator_def sets_sigma)
   577   { fix A B assume A: "A \<in> sets (sigma E1)" and B: "B \<in> sets (sigma E2)"
   578     with proj have "fst -` A \<inter> space ?E \<in> sets ?E" "snd -` B \<inter> space ?E \<in> sets ?E"
   579       unfolding measurable_def by simp_all
   580     moreover have "A \<times> B = (fst -` A \<inter> space ?E) \<inter> (snd -` B \<inter> space ?E)"
   581       using A B M1.sets_into_space M2.sets_into_space
   582       by (auto simp: pair_measure_generator_def)
   583     ultimately have "A \<times> B \<in> sets ?E" by auto }
   584   then have "sigma_sets (space ?E) (sets (pair_measure_generator (sigma E1) (sigma E2))) \<subseteq> sets ?E"
   585     by (intro E.sigma_sets_subset) (auto simp add: pair_measure_generator_def sets_sigma)
   586   then have subset: "sets ?S \<subseteq> sets ?E"
   587     by (simp add: sets_sigma pair_measure_generator_def)
   588   show "sets ?S = sets ?E"
   589   proof (intro set_eqI iffI)
   590     fix A assume "A \<in> sets ?E" then show "A \<in> sets ?S"
   591       unfolding sets_sigma
   592     proof induct
   593       case (Basic A) then show ?case
   594         by (auto simp: pair_measure_generator_def sets_sigma intro: sigma_sets.Basic)
   595     qed (auto intro: sigma_sets.intros simp: pair_measure_generator_def)
   596   next
   597     fix A assume "A \<in> sets ?S" then show "A \<in> sets ?E" using subset by auto
   598   qed
   599 qed
   600 
   601 section "Fubinis theorem"
   602 
   603 lemma (in pair_sigma_finite) simple_function_cut:
   604   assumes f: "simple_function P f" "\<And>x. 0 \<le> f x"
   605   shows "(\<lambda>x. \<integral>\<^isup>+y. f (x, y) \<partial>M2) \<in> borel_measurable M1"
   606     and "(\<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. f (x, y) \<partial>M2) \<partial>M1) = integral\<^isup>P P f"
   607 proof -
   608   have f_borel: "f \<in> borel_measurable P"
   609     using f(1) by (rule borel_measurable_simple_function)
   610   let ?F = "\<lambda>z. f -` {z} \<inter> space P"
   611   let ?F' = "\<lambda>x z. Pair x -` ?F z"
   612   { fix x assume "x \<in> space M1"
   613     have [simp]: "\<And>z y. indicator (?F z) (x, y) = indicator (?F' x z) y"
   614       by (auto simp: indicator_def)
   615     have "\<And>y. y \<in> space M2 \<Longrightarrow> (x, y) \<in> space P" using `x \<in> space M1`
   616       by (simp add: space_pair_measure)
   617     moreover have "\<And>x z. ?F' x z \<in> sets M2" using f_borel
   618       by (intro borel_measurable_vimage measurable_cut_fst)
   619     ultimately have "simple_function M2 (\<lambda> y. f (x, y))"
   620       apply (rule_tac M2.simple_function_cong[THEN iffD2, OF _])
   621       apply (rule simple_function_indicator_representation[OF f(1)])
   622       using `x \<in> space M1` by (auto simp del: space_sigma) }
   623   note M2_sf = this
   624   { fix x assume x: "x \<in> space M1"
   625     then have "(\<integral>\<^isup>+y. f (x, y) \<partial>M2) = (\<Sum>z\<in>f ` space P. z * M2.\<mu> (?F' x z))"
   626       unfolding M2.positive_integral_eq_simple_integral[OF M2_sf[OF x] f(2)]
   627       unfolding simple_integral_def
   628     proof (safe intro!: setsum_mono_zero_cong_left)
   629       from f(1) show "finite (f ` space P)" by (rule simple_functionD)
   630     next
   631       fix y assume "y \<in> space M2" then show "f (x, y) \<in> f ` space P"
   632         using `x \<in> space M1` by (auto simp: space_pair_measure)
   633     next
   634       fix x' y assume "(x', y) \<in> space P"
   635         "f (x', y) \<notin> (\<lambda>y. f (x, y)) ` space M2"
   636       then have *: "?F' x (f (x', y)) = {}"
   637         by (force simp: space_pair_measure)
   638       show  "f (x', y) * M2.\<mu> (?F' x (f (x', y))) = 0"
   639         unfolding * by simp
   640     qed (simp add: vimage_compose[symmetric] comp_def
   641                    space_pair_measure) }
   642   note eq = this
   643   moreover have "\<And>z. ?F z \<in> sets P"
   644     by (auto intro!: f_borel borel_measurable_vimage simp del: space_sigma)
   645   moreover then have "\<And>z. (\<lambda>x. M2.\<mu> (?F' x z)) \<in> borel_measurable M1"
   646     by (auto intro!: measure_cut_measurable_fst simp del: vimage_Int)
   647   moreover have *: "\<And>i x. 0 \<le> M2.\<mu> (Pair x -` (f -` {i} \<inter> space P))"
   648     using f(1)[THEN simple_functionD(2)] f(2) by (intro M2.positive_measure measurable_cut_fst)
   649   moreover { fix i assume "i \<in> f`space P"
   650     with * have "\<And>x. 0 \<le> i * M2.\<mu> (Pair x -` (f -` {i} \<inter> space P))"
   651       using f(2) by auto }
   652   ultimately
   653   show "(\<lambda>x. \<integral>\<^isup>+y. f (x, y) \<partial>M2) \<in> borel_measurable M1"
   654     and "(\<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. f (x, y) \<partial>M2) \<partial>M1) = integral\<^isup>P P f" using f(2)
   655     by (auto simp del: vimage_Int cong: measurable_cong
   656              intro!: M1.borel_measurable_ereal_setsum setsum_cong
   657              simp add: M1.positive_integral_setsum simple_integral_def
   658                        M1.positive_integral_cmult
   659                        M1.positive_integral_cong[OF eq]
   660                        positive_integral_eq_simple_integral[OF f]
   661                        pair_measure_alt[symmetric])
   662 qed
   663 
   664 lemma (in pair_sigma_finite) positive_integral_fst_measurable:
   665   assumes f: "f \<in> borel_measurable P"
   666   shows "(\<lambda>x. \<integral>\<^isup>+ y. f (x, y) \<partial>M2) \<in> borel_measurable M1"
   667       (is "?C f \<in> borel_measurable M1")
   668     and "(\<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. f (x, y) \<partial>M2) \<partial>M1) = integral\<^isup>P P f"
   669 proof -
   670   from borel_measurable_implies_simple_function_sequence'[OF f] guess F . note F = this
   671   then have F_borel: "\<And>i. F i \<in> borel_measurable P"
   672     by (auto intro: borel_measurable_simple_function)
   673   note sf = simple_function_cut[OF F(1,5)]
   674   then have "(\<lambda>x. SUP i. ?C (F i) x) \<in> borel_measurable M1"
   675     using F(1) by auto
   676   moreover
   677   { fix x assume "x \<in> space M1"
   678     from F measurable_pair_image_snd[OF F_borel`x \<in> space M1`]
   679     have "(\<integral>\<^isup>+y. (SUP i. F i (x, y)) \<partial>M2) = (SUP i. ?C (F i) x)"
   680       by (intro M2.positive_integral_monotone_convergence_SUP)
   681          (auto simp: incseq_Suc_iff le_fun_def)
   682     then have "(SUP i. ?C (F i) x) = ?C f x"
   683       unfolding F(4) positive_integral_max_0 by simp }
   684   note SUPR_C = this
   685   ultimately show "?C f \<in> borel_measurable M1"
   686     by (simp cong: measurable_cong)
   687   have "(\<integral>\<^isup>+x. (SUP i. F i x) \<partial>P) = (SUP i. integral\<^isup>P P (F i))"
   688     using F_borel F
   689     by (intro positive_integral_monotone_convergence_SUP) auto
   690   also have "(SUP i. integral\<^isup>P P (F i)) = (SUP i. \<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. F i (x, y) \<partial>M2) \<partial>M1)"
   691     unfolding sf(2) by simp
   692   also have "\<dots> = \<integral>\<^isup>+ x. (SUP i. \<integral>\<^isup>+ y. F i (x, y) \<partial>M2) \<partial>M1" using F sf(1)
   693     by (intro M1.positive_integral_monotone_convergence_SUP[symmetric])
   694        (auto intro!: M2.positive_integral_mono M2.positive_integral_positive
   695                 simp: incseq_Suc_iff le_fun_def)
   696   also have "\<dots> = \<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. (SUP i. F i (x, y)) \<partial>M2) \<partial>M1"
   697     using F_borel F(2,5)
   698     by (auto intro!: M1.positive_integral_cong M2.positive_integral_monotone_convergence_SUP[symmetric]
   699              simp: incseq_Suc_iff le_fun_def measurable_pair_image_snd)
   700   finally show "(\<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. f (x, y) \<partial>M2) \<partial>M1) = integral\<^isup>P P f"
   701     using F by (simp add: positive_integral_max_0)
   702 qed
   703 
   704 lemma (in pair_sigma_finite) measure_preserving_swap:
   705   "(\<lambda>(x,y). (y, x)) \<in> measure_preserving (M1 \<Otimes>\<^isub>M M2) (M2 \<Otimes>\<^isub>M M1)"
   706 proof
   707   interpret Q: pair_sigma_finite M2 M1 by default
   708   show *: "(\<lambda>(x,y). (y, x)) \<in> measurable (M1 \<Otimes>\<^isub>M M2) (M2 \<Otimes>\<^isub>M M1)"
   709     using pair_sigma_algebra_swap_measurable .
   710   fix X assume "X \<in> sets (M2 \<Otimes>\<^isub>M M1)"
   711   from measurable_sets[OF * this] this Q.sets_into_space[OF this]
   712   show "measure (M1 \<Otimes>\<^isub>M M2) ((\<lambda>(x, y). (y, x)) -` X \<inter> space P) = measure (M2 \<Otimes>\<^isub>M M1) X"
   713     by (auto intro!: M1.positive_integral_cong arg_cong[where f="M2.\<mu>"]
   714       simp: pair_measure_alt Q.pair_measure_alt2 space_pair_measure)
   715 qed
   716 
   717 lemma (in pair_sigma_finite) positive_integral_product_swap:
   718   assumes f: "f \<in> borel_measurable P"
   719   shows "(\<integral>\<^isup>+x. f (case x of (x,y)\<Rightarrow>(y,x)) \<partial>(M2 \<Otimes>\<^isub>M M1)) = integral\<^isup>P P f"
   720 proof -
   721   interpret Q: pair_sigma_finite M2 M1 by default
   722   have "sigma_algebra P" by default
   723   with f show ?thesis
   724     by (subst Q.positive_integral_vimage[OF _ Q.measure_preserving_swap]) auto
   725 qed
   726 
   727 lemma (in pair_sigma_finite) positive_integral_snd_measurable:
   728   assumes f: "f \<in> borel_measurable P"
   729   shows "(\<integral>\<^isup>+ y. (\<integral>\<^isup>+ x. f (x, y) \<partial>M1) \<partial>M2) = integral\<^isup>P P f"
   730 proof -
   731   interpret Q: pair_sigma_finite M2 M1 by default
   732   note pair_sigma_algebra_measurable[OF f]
   733   from Q.positive_integral_fst_measurable[OF this]
   734   have "(\<integral>\<^isup>+ y. (\<integral>\<^isup>+ x. f (x, y) \<partial>M1) \<partial>M2) = (\<integral>\<^isup>+ (x, y). f (y, x) \<partial>Q.P)"
   735     by simp
   736   also have "(\<integral>\<^isup>+ (x, y). f (y, x) \<partial>Q.P) = integral\<^isup>P P f"
   737     unfolding positive_integral_product_swap[OF f, symmetric]
   738     by (auto intro!: Q.positive_integral_cong)
   739   finally show ?thesis .
   740 qed
   741 
   742 lemma (in pair_sigma_finite) Fubini:
   743   assumes f: "f \<in> borel_measurable P"
   744   shows "(\<integral>\<^isup>+ y. (\<integral>\<^isup>+ x. f (x, y) \<partial>M1) \<partial>M2) = (\<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. f (x, y) \<partial>M2) \<partial>M1)"
   745   unfolding positive_integral_snd_measurable[OF assms]
   746   unfolding positive_integral_fst_measurable[OF assms] ..
   747 
   748 lemma (in pair_sigma_finite) AE_pair:
   749   assumes "AE x in P. Q x"
   750   shows "AE x in M1. (AE y in M2. Q (x, y))"
   751 proof -
   752   obtain N where N: "N \<in> sets P" "\<mu> N = 0" "{x\<in>space P. \<not> Q x} \<subseteq> N"
   753     using assms unfolding almost_everywhere_def by auto
   754   show ?thesis
   755   proof (rule M1.AE_I)
   756     from N measure_cut_measurable_fst[OF `N \<in> sets P`]
   757     show "M1.\<mu> {x\<in>space M1. M2.\<mu> (Pair x -` N) \<noteq> 0} = 0"
   758       by (auto simp: pair_measure_alt M1.positive_integral_0_iff)
   759     show "{x \<in> space M1. M2.\<mu> (Pair x -` N) \<noteq> 0} \<in> sets M1"
   760       by (intro M1.borel_measurable_ereal_neq_const measure_cut_measurable_fst N)
   761     { fix x assume "x \<in> space M1" "M2.\<mu> (Pair x -` N) = 0"
   762       have "M2.almost_everywhere (\<lambda>y. Q (x, y))"
   763       proof (rule M2.AE_I)
   764         show "M2.\<mu> (Pair x -` N) = 0" by fact
   765         show "Pair x -` N \<in> sets M2" by (intro measurable_cut_fst N)
   766         show "{y \<in> space M2. \<not> Q (x, y)} \<subseteq> Pair x -` N"
   767           using N `x \<in> space M1` unfolding space_sigma space_pair_measure by auto
   768       qed }
   769     then show "{x \<in> space M1. \<not> M2.almost_everywhere (\<lambda>y. Q (x, y))} \<subseteq> {x \<in> space M1. M2.\<mu> (Pair x -` N) \<noteq> 0}"
   770       by auto
   771   qed
   772 qed
   773 
   774 lemma (in pair_sigma_algebra) measurable_product_swap:
   775   "f \<in> measurable (M2 \<Otimes>\<^isub>M M1) M \<longleftrightarrow> (\<lambda>(x,y). f (y,x)) \<in> measurable P M"
   776 proof -
   777   interpret Q: pair_sigma_algebra M2 M1 by default
   778   show ?thesis
   779     using pair_sigma_algebra_measurable[of "\<lambda>(x,y). f (y, x)"]
   780     by (auto intro!: pair_sigma_algebra_measurable Q.pair_sigma_algebra_measurable iffI)
   781 qed
   782 
   783 lemma (in pair_sigma_finite) integrable_product_swap:
   784   assumes "integrable P f"
   785   shows "integrable (M2 \<Otimes>\<^isub>M M1) (\<lambda>(x,y). f (y,x))"
   786 proof -
   787   interpret Q: pair_sigma_finite M2 M1 by default
   788   have *: "(\<lambda>(x,y). f (y,x)) = (\<lambda>x. f (case x of (x,y)\<Rightarrow>(y,x)))" by (auto simp: fun_eq_iff)
   789   show ?thesis unfolding *
   790     using assms unfolding integrable_def
   791     apply (subst (1 2) positive_integral_product_swap)
   792     using `integrable P f` unfolding integrable_def
   793     by (auto simp: *[symmetric] Q.measurable_product_swap[symmetric])
   794 qed
   795 
   796 lemma (in pair_sigma_finite) integrable_product_swap_iff:
   797   "integrable (M2 \<Otimes>\<^isub>M M1) (\<lambda>(x,y). f (y,x)) \<longleftrightarrow> integrable P f"
   798 proof -
   799   interpret Q: pair_sigma_finite M2 M1 by default
   800   from Q.integrable_product_swap[of "\<lambda>(x,y). f (y,x)"] integrable_product_swap[of f]
   801   show ?thesis by auto
   802 qed
   803 
   804 lemma (in pair_sigma_finite) integral_product_swap:
   805   assumes "integrable P f"
   806   shows "(\<integral>(x,y). f (y,x) \<partial>(M2 \<Otimes>\<^isub>M M1)) = integral\<^isup>L P f"
   807 proof -
   808   interpret Q: pair_sigma_finite M2 M1 by default
   809   have *: "(\<lambda>(x,y). f (y,x)) = (\<lambda>x. f (case x of (x,y)\<Rightarrow>(y,x)))" by (auto simp: fun_eq_iff)
   810   show ?thesis
   811     unfolding lebesgue_integral_def *
   812     apply (subst (1 2) positive_integral_product_swap)
   813     using `integrable P f` unfolding integrable_def
   814     by (auto simp: *[symmetric] Q.measurable_product_swap[symmetric])
   815 qed
   816 
   817 lemma (in pair_sigma_finite) integrable_fst_measurable:
   818   assumes f: "integrable P f"
   819   shows "M1.almost_everywhere (\<lambda>x. integrable M2 (\<lambda> y. f (x, y)))" (is "?AE")
   820     and "(\<integral>x. (\<integral>y. f (x, y) \<partial>M2) \<partial>M1) = integral\<^isup>L P f" (is "?INT")
   821 proof -
   822   let ?pf = "\<lambda>x. ereal (f x)" and ?nf = "\<lambda>x. ereal (- f x)"
   823   have
   824     borel: "?nf \<in> borel_measurable P""?pf \<in> borel_measurable P" and
   825     int: "integral\<^isup>P P ?nf \<noteq> \<infinity>" "integral\<^isup>P P ?pf \<noteq> \<infinity>"
   826     using assms by auto
   827   have "(\<integral>\<^isup>+x. (\<integral>\<^isup>+y. ereal (f (x, y)) \<partial>M2) \<partial>M1) \<noteq> \<infinity>"
   828      "(\<integral>\<^isup>+x. (\<integral>\<^isup>+y. ereal (- f (x, y)) \<partial>M2) \<partial>M1) \<noteq> \<infinity>"
   829     using borel[THEN positive_integral_fst_measurable(1)] int
   830     unfolding borel[THEN positive_integral_fst_measurable(2)] by simp_all
   831   with borel[THEN positive_integral_fst_measurable(1)]
   832   have AE_pos: "AE x in M1. (\<integral>\<^isup>+y. ereal (f (x, y)) \<partial>M2) \<noteq> \<infinity>"
   833     "AE x in M1. (\<integral>\<^isup>+y. ereal (- f (x, y)) \<partial>M2) \<noteq> \<infinity>"
   834     by (auto intro!: M1.positive_integral_PInf_AE )
   835   then have AE: "AE x in M1. \<bar>\<integral>\<^isup>+y. ereal (f (x, y)) \<partial>M2\<bar> \<noteq> \<infinity>"
   836     "AE x in M1. \<bar>\<integral>\<^isup>+y. ereal (- f (x, y)) \<partial>M2\<bar> \<noteq> \<infinity>"
   837     by (auto simp: M2.positive_integral_positive)
   838   from AE_pos show ?AE using assms
   839     by (simp add: measurable_pair_image_snd integrable_def)
   840   { fix f have "(\<integral>\<^isup>+ x. - \<integral>\<^isup>+ y. ereal (f x y) \<partial>M2 \<partial>M1) = (\<integral>\<^isup>+x. 0 \<partial>M1)"
   841       using M2.positive_integral_positive
   842       by (intro M1.positive_integral_cong_pos) (auto simp: ereal_uminus_le_reorder)
   843     then have "(\<integral>\<^isup>+ x. - \<integral>\<^isup>+ y. ereal (f x y) \<partial>M2 \<partial>M1) = 0" by simp }
   844   note this[simp]
   845   { fix f assume borel: "(\<lambda>x. ereal (f x)) \<in> borel_measurable P"
   846       and int: "integral\<^isup>P P (\<lambda>x. ereal (f x)) \<noteq> \<infinity>"
   847       and AE: "M1.almost_everywhere (\<lambda>x. (\<integral>\<^isup>+y. ereal (f (x, y)) \<partial>M2) \<noteq> \<infinity>)"
   848     have "integrable M1 (\<lambda>x. real (\<integral>\<^isup>+y. ereal (f (x, y)) \<partial>M2))" (is "integrable M1 ?f")
   849     proof (intro integrable_def[THEN iffD2] conjI)
   850       show "?f \<in> borel_measurable M1"
   851         using borel by (auto intro!: M1.borel_measurable_real_of_ereal positive_integral_fst_measurable)
   852       have "(\<integral>\<^isup>+x. ereal (?f x) \<partial>M1) = (\<integral>\<^isup>+x. (\<integral>\<^isup>+y. ereal (f (x, y))  \<partial>M2) \<partial>M1)"
   853         using AE M2.positive_integral_positive
   854         by (auto intro!: M1.positive_integral_cong_AE simp: ereal_real)
   855       then show "(\<integral>\<^isup>+x. ereal (?f x) \<partial>M1) \<noteq> \<infinity>"
   856         using positive_integral_fst_measurable[OF borel] int by simp
   857       have "(\<integral>\<^isup>+x. ereal (- ?f x) \<partial>M1) = (\<integral>\<^isup>+x. 0 \<partial>M1)"
   858         by (intro M1.positive_integral_cong_pos)
   859            (simp add: M2.positive_integral_positive real_of_ereal_pos)
   860       then show "(\<integral>\<^isup>+x. ereal (- ?f x) \<partial>M1) \<noteq> \<infinity>" by simp
   861     qed }
   862   with this[OF borel(1) int(1) AE_pos(2)] this[OF borel(2) int(2) AE_pos(1)]
   863   show ?INT
   864     unfolding lebesgue_integral_def[of P] lebesgue_integral_def[of M2]
   865       borel[THEN positive_integral_fst_measurable(2), symmetric]
   866     using AE[THEN M1.integral_real]
   867     by simp
   868 qed
   869 
   870 lemma (in pair_sigma_finite) integrable_snd_measurable:
   871   assumes f: "integrable P f"
   872   shows "M2.almost_everywhere (\<lambda>y. integrable M1 (\<lambda>x. f (x, y)))" (is "?AE")
   873     and "(\<integral>y. (\<integral>x. f (x, y) \<partial>M1) \<partial>M2) = integral\<^isup>L P f" (is "?INT")
   874 proof -
   875   interpret Q: pair_sigma_finite M2 M1 by default
   876   have Q_int: "integrable Q.P (\<lambda>(x, y). f (y, x))"
   877     using f unfolding integrable_product_swap_iff .
   878   show ?INT
   879     using Q.integrable_fst_measurable(2)[OF Q_int]
   880     using integral_product_swap[OF f] by simp
   881   show ?AE
   882     using Q.integrable_fst_measurable(1)[OF Q_int]
   883     by simp
   884 qed
   885 
   886 lemma (in pair_sigma_finite) Fubini_integral:
   887   assumes f: "integrable P f"
   888   shows "(\<integral>y. (\<integral>x. f (x, y) \<partial>M1) \<partial>M2) = (\<integral>x. (\<integral>y. f (x, y) \<partial>M2) \<partial>M1)"
   889   unfolding integrable_snd_measurable[OF assms]
   890   unfolding integrable_fst_measurable[OF assms] ..
   891 
   892 section "Products on finite spaces"
   893 
   894 lemma sigma_sets_pair_measure_generator_finite:
   895   assumes "finite A" and "finite B"
   896   shows "sigma_sets (A \<times> B) { a \<times> b | a b. a \<in> Pow A \<and> b \<in> Pow B} = Pow (A \<times> B)"
   897   (is "sigma_sets ?prod ?sets = _")
   898 proof safe
   899   have fin: "finite (A \<times> B)" using assms by (rule finite_cartesian_product)
   900   fix x assume subset: "x \<subseteq> A \<times> B"
   901   hence "finite x" using fin by (rule finite_subset)
   902   from this subset show "x \<in> sigma_sets ?prod ?sets"
   903   proof (induct x)
   904     case empty show ?case by (rule sigma_sets.Empty)
   905   next
   906     case (insert a x)
   907     hence "{a} \<in> sigma_sets ?prod ?sets"
   908       by (auto simp: pair_measure_generator_def intro!: sigma_sets.Basic)
   909     moreover have "x \<in> sigma_sets ?prod ?sets" using insert by auto
   910     ultimately show ?case unfolding insert_is_Un[of a x] by (rule sigma_sets_Un)
   911   qed
   912 next
   913   fix x a b
   914   assume "x \<in> sigma_sets ?prod ?sets" and "(a, b) \<in> x"
   915   from sigma_sets_into_sp[OF _ this(1)] this(2)
   916   show "a \<in> A" and "b \<in> B" by auto
   917 qed
   918 
   919 locale pair_finite_sigma_algebra = pair_sigma_algebra M1 M2 + M1: finite_sigma_algebra M1 + M2: finite_sigma_algebra M2 for M1 M2
   920 
   921 lemma (in pair_finite_sigma_algebra) finite_pair_sigma_algebra:
   922   shows "P = \<lparr> space = space M1 \<times> space M2, sets = Pow (space M1 \<times> space M2), \<dots> = algebra.more P \<rparr>"
   923 proof -
   924   show ?thesis
   925     using sigma_sets_pair_measure_generator_finite[OF M1.finite_space M2.finite_space]
   926     by (intro algebra.equality) (simp_all add: pair_measure_def pair_measure_generator_def sigma_def)
   927 qed
   928 
   929 sublocale pair_finite_sigma_algebra \<subseteq> finite_sigma_algebra P
   930 proof
   931   show "finite (space P)"
   932     using M1.finite_space M2.finite_space
   933     by (subst finite_pair_sigma_algebra) simp
   934   show "sets P = Pow (space P)"
   935     by (subst (1 2) finite_pair_sigma_algebra) simp
   936 qed
   937 
   938 locale pair_finite_space = pair_sigma_finite M1 M2 + pair_finite_sigma_algebra M1 M2 +
   939   M1: finite_measure_space M1 + M2: finite_measure_space M2 for M1 M2
   940 
   941 lemma (in pair_finite_space) pair_measure_Pair[simp]:
   942   assumes "a \<in> space M1" "b \<in> space M2"
   943   shows "\<mu> {(a, b)} = M1.\<mu> {a} * M2.\<mu> {b}"
   944 proof -
   945   have "\<mu> ({a}\<times>{b}) = M1.\<mu> {a} * M2.\<mu> {b}"
   946     using M1.sets_eq_Pow M2.sets_eq_Pow assms
   947     by (subst pair_measure_times) auto
   948   then show ?thesis by simp
   949 qed
   950 
   951 lemma (in pair_finite_space) pair_measure_singleton[simp]:
   952   assumes "x \<in> space M1 \<times> space M2"
   953   shows "\<mu> {x} = M1.\<mu> {fst x} * M2.\<mu> {snd x}"
   954   using pair_measure_Pair assms by (cases x) auto
   955 
   956 sublocale pair_finite_space \<subseteq> finite_measure_space P
   957 proof unfold_locales
   958   show "measure P (space P) \<noteq> \<infinity>"
   959     by (subst (2) finite_pair_sigma_algebra)
   960        (simp add: pair_measure_times)
   961 qed
   962 
   963 end