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