src/HOL/Probability/Binary_Product_Measure.thy
author blanchet
Mon, 21 May 2012 10:39:31 +0200
changeset 47944 e6b51fab96f7
parent 47694 05663f75964c
child 49776 199d1d5bb17e
permissions -rw-r--r--
added helper -- cf. SET616^5
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
42146
5b52c6a9c627 split Product_Measure into Binary_Product_Measure and Finite_Product_Measure
hoelzl
parents: 42067
diff changeset
     1
(*  Title:      HOL/Probability/Binary_Product_Measure.thy
42067
66c8281349ec standardized headers
hoelzl
parents: 41981
diff changeset
     2
    Author:     Johannes Hölzl, TU München
66c8281349ec standardized headers
hoelzl
parents: 41981
diff changeset
     3
*)
66c8281349ec standardized headers
hoelzl
parents: 41981
diff changeset
     4
42146
5b52c6a9c627 split Product_Measure into Binary_Product_Measure and Finite_Product_Measure
hoelzl
parents: 42067
diff changeset
     5
header {*Binary product measures*}
42067
66c8281349ec standardized headers
hoelzl
parents: 41981
diff changeset
     6
42146
5b52c6a9c627 split Product_Measure into Binary_Product_Measure and Finite_Product_Measure
hoelzl
parents: 42067
diff changeset
     7
theory Binary_Product_Measure
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36649
diff changeset
     8
imports Lebesgue_Integration
35833
7b7ae5aa396d Added product measure space
hoelzl
parents:
diff changeset
     9
begin
7b7ae5aa396d Added product measure space
hoelzl
parents:
diff changeset
    10
42950
6e5c2a3c69da move lemmas to Extended_Reals and Extended_Real_Limits
hoelzl
parents: 42146
diff changeset
    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 = {}))"
6e5c2a3c69da move lemmas to Extended_Reals and Extended_Real_Limits
hoelzl
parents: 42146
diff changeset
    12
  by auto
6e5c2a3c69da move lemmas to Extended_Reals and Extended_Real_Limits
hoelzl
parents: 42146
diff changeset
    13
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
    14
lemma times_Int_times: "A \<times> B \<inter> C \<times> D = (A \<inter> C) \<times> (B \<inter> D)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
    15
  by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
    16
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
    17
lemma Pair_vimage_times[simp]: "\<And>A B x. Pair x -` (A \<times> B) = (if x \<in> A then B else {})"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
    18
  by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
    19
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
    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 {})"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
    21
  by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
    22
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
    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))"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
    24
  by (cases x) simp
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
    25
41026
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41023
diff changeset
    26
lemma split_beta': "(\<lambda>(x,y). f x y) = (\<lambda>x. f (fst x) (snd x))"
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41023
diff changeset
    27
  by (auto simp: fun_eq_iff)
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41023
diff changeset
    28
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
    29
section "Binary products"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
    30
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
    31
definition pair_measure (infixr "\<Otimes>\<^isub>M" 80) where
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
    32
  "A \<Otimes>\<^isub>M B = measure_of (space A \<times> space B)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
    33
      {a \<times> b | a b. a \<in> sets A \<and> b \<in> sets B}
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
    34
      (\<lambda>X. \<integral>\<^isup>+x. (\<integral>\<^isup>+y. indicator X (x,y) \<partial>B) \<partial>A)"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
    35
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
    36
lemma space_pair_measure:
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
    37
  "space (A \<Otimes>\<^isub>M B) = space A \<times> space B"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
    38
  unfolding pair_measure_def using space_closed[of A] space_closed[of B]
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
    39
  by (intro space_measure_of) auto
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
    40
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
    41
lemma sets_pair_measure:
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
    42
  "sets (A \<Otimes>\<^isub>M B) = sigma_sets (space A \<times> space B) {a \<times> b | a b. a \<in> sets A \<and> b \<in> sets B}"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
    43
  unfolding pair_measure_def using space_closed[of A] space_closed[of B]
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
    44
  by (intro sets_measure_of) auto
41095
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41026
diff changeset
    45
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
    46
lemma pair_measureI[intro, simp]:
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
    47
  "x \<in> sets A \<Longrightarrow> y \<in> sets B \<Longrightarrow> x \<times> y \<in> sets (A \<Otimes>\<^isub>M B)"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
    48
  by (auto simp: sets_pair_measure)
41095
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41026
diff changeset
    49
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
    50
lemma measurable_pair_measureI:
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
    51
  assumes 1: "f \<in> space M \<rightarrow> space M1 \<times> space M2"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
    52
  assumes 2: "\<And>A B. A \<in> sets M1 \<Longrightarrow> B \<in> sets M2 \<Longrightarrow> f -` (A \<times> B) \<inter> space M \<in> sets M"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
    53
  shows "f \<in> measurable M (M1 \<Otimes>\<^isub>M M2)"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
    54
  unfolding pair_measure_def using 1 2
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
    55
  by (intro measurable_measure_of) (auto dest: sets_into_space)
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
    56
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
    57
lemma measurable_fst[intro!, simp]: "fst \<in> measurable (M1 \<Otimes>\<^isub>M M2) M1"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
    58
  unfolding measurable_def
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
    59
proof safe
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
    60
  fix A assume A: "A \<in> sets M1"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
    61
  from this[THEN sets_into_space] have "fst -` A \<inter> space M1 \<times> space M2 = A \<times> space M2" by auto
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
    62
  with A show "fst -` A \<inter> space (M1 \<Otimes>\<^isub>M M2) \<in> sets (M1 \<Otimes>\<^isub>M M2)" by (simp add: space_pair_measure)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
    63
qed (simp add: space_pair_measure)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
    64
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
    65
lemma measurable_snd[intro!, simp]: "snd \<in> measurable (M1 \<Otimes>\<^isub>M M2) M2"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
    66
  unfolding measurable_def
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
    67
proof safe
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
    68
  fix A assume A: "A \<in> sets M2"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
    69
  from this[THEN sets_into_space] have "snd -` A \<inter> space M1 \<times> space M2 = space M1 \<times> A" by auto
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
    70
  with A show "snd -` A \<inter> space (M1 \<Otimes>\<^isub>M M2) \<in> sets (M1 \<Otimes>\<^isub>M M2)" by (simp add: space_pair_measure)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
    71
qed (simp add: space_pair_measure)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
    72
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
    73
lemma measurable_fst': "f \<in> measurable M (N \<Otimes>\<^isub>M P) \<Longrightarrow> (\<lambda>x. fst (f x)) \<in> measurable M N"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
    74
  using measurable_comp[OF _ measurable_fst] by (auto simp: comp_def)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
    75
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
    76
lemma measurable_snd': "f \<in> measurable M (N \<Otimes>\<^isub>M P) \<Longrightarrow> (\<lambda>x. snd (f x)) \<in> measurable M P"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
    77
    using measurable_comp[OF _ measurable_snd] by (auto simp: comp_def)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
    78
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
    79
lemma measurable_fst'': "f \<in> measurable M N \<Longrightarrow> (\<lambda>x. f (fst x)) \<in> measurable (M \<Otimes>\<^isub>M P) N"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
    80
  using measurable_comp[OF measurable_fst _] by (auto simp: comp_def)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
    81
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
    82
lemma measurable_snd'': "f \<in> measurable M N \<Longrightarrow> (\<lambda>x. f (snd x)) \<in> measurable (P \<Otimes>\<^isub>M M) N"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
    83
  using measurable_comp[OF measurable_snd _] by (auto simp: comp_def)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
    84
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
    85
lemma measurable_pair_iff:
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
    86
  "f \<in> measurable M (M1 \<Otimes>\<^isub>M M2) \<longleftrightarrow> (fst \<circ> f) \<in> measurable M M1 \<and> (snd \<circ> f) \<in> measurable M M2"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
    87
proof safe
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
    88
  assume f: "(fst \<circ> f) \<in> measurable M M1" and s: "(snd \<circ> f) \<in> measurable M M2"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
    89
  show "f \<in> measurable M (M1 \<Otimes>\<^isub>M M2)"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
    90
  proof (rule measurable_pair_measureI)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
    91
    show "f \<in> space M \<rightarrow> space M1 \<times> space M2"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
    92
      using f s by (auto simp: mem_Times_iff measurable_def comp_def)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
    93
    fix A B assume "A \<in> sets M1" "B \<in> sets M2"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
    94
    moreover have "(fst \<circ> f) -` A \<inter> space M \<in> sets M" "(snd \<circ> f) -` B \<inter> space M \<in> sets M"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
    95
      using f `A \<in> sets M1` s `B \<in> sets M2` by (auto simp: measurable_sets)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
    96
    moreover have "f -` (A \<times> B) \<inter> space M = ((fst \<circ> f) -` A \<inter> space M) \<inter> ((snd \<circ> f) -` B \<inter> space M)"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
    97
      by (auto simp: vimage_Times)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
    98
    ultimately show "f -` (A \<times> B) \<inter> space M \<in> sets M" by auto
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
    99
  qed
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   100
qed auto
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   101
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   102
lemma measurable_pair:
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   103
  "(fst \<circ> f) \<in> measurable M M1 \<Longrightarrow> (snd \<circ> f) \<in> measurable M M2 \<Longrightarrow> f \<in> measurable M (M1 \<Otimes>\<^isub>M M2)"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   104
  unfolding measurable_pair_iff by simp
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   105
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   106
lemma measurable_pair_swap': "(\<lambda>(x,y). (y, x)) \<in> measurable (M1 \<Otimes>\<^isub>M M2) (M2 \<Otimes>\<^isub>M M1)"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   107
proof (rule measurable_pair_measureI)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   108
  fix A B assume "A \<in> sets M2" "B \<in> sets M1"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   109
  moreover then have "(\<lambda>(x, y). (y, x)) -` (A \<times> B) \<inter> space (M1 \<Otimes>\<^isub>M M2) = B \<times> A"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   110
    by (auto dest: sets_into_space simp: space_pair_measure)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   111
  ultimately show "(\<lambda>(x, y). (y, x)) -` (A \<times> B) \<inter> space (M1 \<Otimes>\<^isub>M M2) \<in> sets (M1 \<Otimes>\<^isub>M M2)"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   112
    by auto
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   113
qed (auto simp add: space_pair_measure)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   114
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   115
lemma measurable_pair_swap:
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   116
  assumes f: "f \<in> measurable (M1 \<Otimes>\<^isub>M M2) M" shows "(\<lambda>(x,y). f (y, x)) \<in> measurable (M2 \<Otimes>\<^isub>M M1) M"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   117
proof -
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   118
  have "(\<lambda>x. f (case x of (x, y) \<Rightarrow> (y, x))) = (\<lambda>(x, y). f (y, x))" by auto
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   119
  then show ?thesis
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   120
    using measurable_comp[OF measurable_pair_swap' f] by (simp add: comp_def)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   121
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   122
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   123
lemma measurable_pair_swap_iff:
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   124
  "f \<in> measurable (M2 \<Otimes>\<^isub>M M1) M \<longleftrightarrow> (\<lambda>(x,y). f (y,x)) \<in> measurable (M1 \<Otimes>\<^isub>M M2) M"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   125
  using measurable_pair_swap[of "\<lambda>(x,y). f (y, x)"]
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   126
  by (auto intro!: measurable_pair_swap)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   127
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   128
lemma measurable_Pair1': "x \<in> space M1 \<Longrightarrow> Pair x \<in> measurable M2 (M1 \<Otimes>\<^isub>M M2)"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   129
proof (rule measurable_pair_measureI)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   130
  fix A B assume "A \<in> sets M1" "B \<in> sets M2"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   131
  moreover then have "Pair x -` (A \<times> B) \<inter> space M2 = (if x \<in> A then B else {})"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   132
    by (auto dest: sets_into_space simp: space_pair_measure)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   133
  ultimately show "Pair x -` (A \<times> B) \<inter> space M2 \<in> sets M2"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   134
    by auto
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   135
qed (auto simp add: space_pair_measure)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   136
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   137
lemma sets_Pair1: assumes A: "A \<in> sets (M1 \<Otimes>\<^isub>M M2)" shows "Pair x -` A \<in> sets M2"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   138
proof -
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   139
  have "Pair x -` A = (if x \<in> space M1 then Pair x -` A \<inter> space M2 else {})"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   140
    using A[THEN sets_into_space] by (auto simp: space_pair_measure)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   141
  also have "\<dots> \<in> sets M2"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   142
    using A by (auto simp add: measurable_Pair1' intro!: measurable_sets split: split_if_asm)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   143
  finally show ?thesis .
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   144
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   145
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   146
lemma measurable_Pair2': "y \<in> space M2 \<Longrightarrow> (\<lambda>x. (x, y)) \<in> measurable M1 (M1 \<Otimes>\<^isub>M M2)"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   147
  using measurable_comp[OF measurable_Pair1' measurable_pair_swap', of y M2 M1]
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   148
  by (simp add: comp_def)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   149
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   150
lemma sets_Pair2: assumes A: "A \<in> sets (M1 \<Otimes>\<^isub>M M2)" shows "(\<lambda>x. (x, y)) -` A \<in> sets M1"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   151
proof -
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   152
  have "(\<lambda>x. (x, y)) -` A = (if y \<in> space M2 then (\<lambda>x. (x, y)) -` A \<inter> space M1 else {})"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   153
    using A[THEN sets_into_space] by (auto simp: space_pair_measure)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   154
  also have "\<dots> \<in> sets M1"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   155
    using A by (auto simp add: measurable_Pair2' intro!: measurable_sets split: split_if_asm)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   156
  finally show ?thesis .
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   157
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   158
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   159
lemma measurable_Pair2:
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   160
  assumes f: "f \<in> measurable (M1 \<Otimes>\<^isub>M M2) M" and x: "x \<in> space M1"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   161
  shows "(\<lambda>y. f (x, y)) \<in> measurable M2 M"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   162
  using measurable_comp[OF measurable_Pair1' f, OF x]
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   163
  by (simp add: comp_def)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   164
  
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   165
lemma measurable_Pair1:
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   166
  assumes f: "f \<in> measurable (M1 \<Otimes>\<^isub>M M2) M" and y: "y \<in> space M2"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   167
  shows "(\<lambda>x. f (x, y)) \<in> measurable M1 M"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   168
  using measurable_comp[OF measurable_Pair2' f, OF y]
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   169
  by (simp add: comp_def)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   170
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   171
lemma Int_stable_pair_measure_generator: "Int_stable {a \<times> b | a b. a \<in> sets A \<and> b \<in> sets B}"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   172
  unfolding Int_stable_def
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   173
  by safe (auto simp add: times_Int_times)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   174
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   175
lemma finite_measure_cut_measurable:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   176
  assumes "sigma_finite_measure M1" "finite_measure M2"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   177
  assumes "Q \<in> sets (M1 \<Otimes>\<^isub>M M2)"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   178
  shows "(\<lambda>x. emeasure M2 (Pair x -` Q)) \<in> borel_measurable M1"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   179
    (is "?s Q \<in> _")
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   180
proof -
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   181
  interpret M1: sigma_finite_measure M1 by fact
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   182
  interpret M2: finite_measure M2 by fact
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   183
  let ?\<Omega> = "space M1 \<times> space M2" and ?D = "{A\<in>sets (M1 \<Otimes>\<^isub>M M2). ?s A \<in> borel_measurable M1}"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   184
  note space_pair_measure[simp]
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   185
  interpret dynkin_system ?\<Omega> ?D
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   186
  proof (intro dynkin_systemI)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   187
    fix A assume "A \<in> ?D" then show "A \<subseteq> ?\<Omega>"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   188
      using sets_into_space[of A "M1 \<Otimes>\<^isub>M M2"] by simp
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   189
  next
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   190
    from top show "?\<Omega> \<in> ?D"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   191
      by (auto simp add: if_distrib intro!: measurable_If)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   192
  next
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   193
    fix A assume "A \<in> ?D"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   194
    with sets_into_space have "\<And>x. emeasure M2 (Pair x -` (?\<Omega> - A)) =
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   195
        (if x \<in> space M1 then emeasure M2 (space M2) - ?s A x else 0)"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   196
      by (auto intro!: emeasure_compl simp: vimage_Diff sets_Pair1)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   197
    with `A \<in> ?D` top show "?\<Omega> - A \<in> ?D"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   198
      by (auto intro!: measurable_If)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   199
  next
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   200
    fix F :: "nat \<Rightarrow> ('a\<times>'b) set" assume "disjoint_family F" "range F \<subseteq> ?D"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   201
    moreover then have "\<And>x. emeasure M2 (\<Union>i. Pair x -` F i) = (\<Sum>i. ?s (F i) x)"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   202
      by (intro suminf_emeasure[symmetric]) (auto simp: disjoint_family_on_def sets_Pair1)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   203
    ultimately show "(\<Union>i. F i) \<in> ?D"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   204
      by (auto simp: vimage_UN intro!: borel_measurable_psuminf)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   205
  qed
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   206
  let ?G = "{a \<times> b | a b. a \<in> sets M1 \<and> b \<in> sets M2}"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   207
  have "sigma_sets ?\<Omega> ?G = ?D"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   208
  proof (rule dynkin_lemma)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   209
    show "?G \<subseteq> ?D"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   210
      by (auto simp: if_distrib Int_def[symmetric] intro!: measurable_If)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   211
  qed (auto simp: sets_pair_measure  Int_stable_pair_measure_generator)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   212
  with `Q \<in> sets (M1 \<Otimes>\<^isub>M M2)` have "Q \<in> ?D"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   213
    by (simp add: sets_pair_measure[symmetric])
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   214
  then show "?s Q \<in> borel_measurable M1" by simp
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   215
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   216
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   217
subsection {* Binary products of $\sigma$-finite emeasure spaces *}
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   218
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   219
locale pair_sigma_finite = M1: sigma_finite_measure M1 + M2: sigma_finite_measure M2
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   220
  for M1 :: "'a measure" and M2 :: "'b measure"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   221
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   222
lemma sets_pair_measure_cong[cong]:
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   223
  "sets M1 = sets M1' \<Longrightarrow> sets M2 = sets M2' \<Longrightarrow> sets (M1 \<Otimes>\<^isub>M M2) = sets (M1' \<Otimes>\<^isub>M M2')"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   224
  unfolding sets_pair_measure by (simp cong: sets_eq_imp_space_eq)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   225
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   226
lemma (in pair_sigma_finite) measurable_emeasure_Pair1:
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   227
  assumes Q: "Q \<in> sets (M1 \<Otimes>\<^isub>M M2)" shows "(\<lambda>x. emeasure M2 (Pair x -` Q)) \<in> borel_measurable M1" (is "?s Q \<in> _")
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   228
proof -
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   229
  from M2.sigma_finite_disjoint guess F . note F = this
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   230
  then have F_sets: "\<And>i. F i \<in> sets M2" by auto
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   231
  have M1: "sigma_finite_measure M1" ..
46731
5302e932d1e5 avoid undeclared variables in let bindings;
wenzelm
parents: 45777
diff changeset
   232
  let ?C = "\<lambda>x i. F i \<inter> Pair x -` Q"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   233
  { fix i
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   234
    have [simp]: "space M1 \<times> F i \<inter> space M1 \<times> space M2 = space M1 \<times> F i"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   235
      using F sets_into_space by auto
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   236
    let ?R = "density M2 (indicator (F i))"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   237
    have "(\<lambda>x. emeasure ?R (Pair x -` (space M1 \<times> space ?R \<inter> Q))) \<in> borel_measurable M1"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   238
    proof (intro finite_measure_cut_measurable[OF M1])
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   239
      show "finite_measure ?R"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   240
        using F by (intro finite_measureI) (auto simp: emeasure_restricted subset_eq)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   241
      show "(space M1 \<times> space ?R) \<inter> Q \<in> sets (M1 \<Otimes>\<^isub>M ?R)"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   242
        using Q by (simp add: Int)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   243
    qed
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   244
    moreover have "\<And>x. emeasure ?R (Pair x -` (space M1 \<times> space ?R \<inter> Q))
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   245
        = emeasure M2 (F i \<inter> Pair x -` (space M1 \<times> space ?R \<inter> Q))"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   246
      using Q F_sets by (intro emeasure_restricted) (auto intro: sets_Pair1)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   247
    moreover have "\<And>x. F i \<inter> Pair x -` (space M1 \<times> space ?R \<inter> Q) = ?C x i"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   248
      using sets_into_space[OF Q] by (auto simp: space_pair_measure)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   249
    ultimately have "(\<lambda>x. emeasure M2 (?C x i)) \<in> borel_measurable M1"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   250
      by simp }
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   251
  moreover
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   252
  { fix x
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   253
    have "(\<Sum>i. emeasure M2 (?C x i)) = emeasure M2 (\<Union>i. ?C x i)"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   254
    proof (intro suminf_emeasure)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   255
      show "range (?C x) \<subseteq> sets M2"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   256
        using F `Q \<in> sets (M1 \<Otimes>\<^isub>M M2)` by (auto intro!: sets_Pair1)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   257
      have "disjoint_family F" using F by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   258
      show "disjoint_family (?C x)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   259
        by (rule disjoint_family_on_bisimulation[OF `disjoint_family F`]) auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   260
    qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   261
    also have "(\<Union>i. ?C x i) = Pair x -` Q"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   262
      using F sets_into_space[OF `Q \<in> sets (M1 \<Otimes>\<^isub>M M2)`]
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   263
      by (auto simp: space_pair_measure)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   264
    finally have "emeasure M2 (Pair x -` Q) = (\<Sum>i. emeasure M2 (?C x i))"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   265
      by simp }
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   266
  ultimately show ?thesis using `Q \<in> sets (M1 \<Otimes>\<^isub>M M2)` F_sets
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   267
    by auto
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   268
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   269
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   270
lemma (in pair_sigma_finite) measurable_emeasure_Pair2:
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   271
  assumes Q: "Q \<in> sets (M1 \<Otimes>\<^isub>M M2)" shows "(\<lambda>y. emeasure M1 ((\<lambda>x. (x, y)) -` Q)) \<in> borel_measurable M2"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   272
proof -
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   273
  interpret Q: pair_sigma_finite M2 M1 by default
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   274
  have "(\<lambda>(x, y). (y, x)) -` Q \<inter> space (M2 \<Otimes>\<^isub>M M1) \<in> sets (M2 \<Otimes>\<^isub>M M1)"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   275
    using Q measurable_pair_swap' by (auto intro: measurable_sets)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   276
  note Q.measurable_emeasure_Pair1[OF this]
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   277
  moreover have "\<And>y. Pair y -` ((\<lambda>(x, y). (y, x)) -` Q \<inter> space (M2 \<Otimes>\<^isub>M M1)) = (\<lambda>x. (x, y)) -` Q"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   278
    using Q[THEN sets_into_space] by (auto simp: space_pair_measure)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   279
  ultimately show ?thesis by simp
39088
ca17017c10e6 Measurable on product space is equiv. to measurable components
hoelzl
parents: 39082
diff changeset
   280
qed
ca17017c10e6 Measurable on product space is equiv. to measurable components
hoelzl
parents: 39082
diff changeset
   281
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   282
lemma (in pair_sigma_finite) emeasure_pair_measure:
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   283
  assumes "X \<in> sets (M1 \<Otimes>\<^isub>M M2)"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   284
  shows "emeasure (M1 \<Otimes>\<^isub>M M2) X = (\<integral>\<^isup>+ x. \<integral>\<^isup>+ y. indicator X (x, y) \<partial>M2 \<partial>M1)" (is "_ = ?\<mu> X")
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   285
proof (rule emeasure_measure_of[OF pair_measure_def])
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   286
  show "positive (sets (M1 \<Otimes>\<^isub>M M2)) ?\<mu>"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   287
    by (auto simp: positive_def positive_integral_positive)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   288
  have eq[simp]: "\<And>A x y. indicator A (x, y) = indicator (Pair x -` A) y"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   289
    by (auto simp: indicator_def)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   290
  show "countably_additive (sets (M1 \<Otimes>\<^isub>M M2)) ?\<mu>"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   291
  proof (rule countably_additiveI)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   292
    fix F :: "nat \<Rightarrow> ('a \<times> 'b) set" assume F: "range F \<subseteq> sets (M1 \<Otimes>\<^isub>M M2)" "disjoint_family F"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   293
    from F have *: "\<And>i. F i \<in> sets (M1 \<Otimes>\<^isub>M M2)" "(\<Union>i. F i) \<in> sets (M1 \<Otimes>\<^isub>M M2)" by auto
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   294
    moreover from F have "\<And>i. (\<lambda>x. emeasure M2 (Pair x -` F i)) \<in> borel_measurable M1"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   295
      by (intro measurable_emeasure_Pair1) auto
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   296
    moreover have "\<And>x. disjoint_family (\<lambda>i. Pair x -` F i)"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   297
      by (intro disjoint_family_on_bisimulation[OF F(2)]) auto
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   298
    moreover have "\<And>x. range (\<lambda>i. Pair x -` F i) \<subseteq> sets M2"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   299
      using F by (auto simp: sets_Pair1)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   300
    ultimately show "(\<Sum>n. ?\<mu> (F n)) = ?\<mu> (\<Union>i. F i)"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   301
      by (auto simp add: vimage_UN positive_integral_suminf[symmetric] suminf_emeasure subset_eq emeasure_nonneg sets_Pair1
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   302
               intro!: positive_integral_cong positive_integral_indicator[symmetric])
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   303
  qed
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   304
  show "{a \<times> b |a b. a \<in> sets M1 \<and> b \<in> sets M2} \<subseteq> Pow (space M1 \<times> space M2)"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   305
    using space_closed[of M1] space_closed[of M2] by auto
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   306
qed fact
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   307
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   308
lemma (in pair_sigma_finite) emeasure_pair_measure_alt:
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   309
  assumes X: "X \<in> sets (M1 \<Otimes>\<^isub>M M2)"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   310
  shows "emeasure (M1  \<Otimes>\<^isub>M M2) X = (\<integral>\<^isup>+x. emeasure M2 (Pair x -` X) \<partial>M1)"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   311
proof -
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   312
  have [simp]: "\<And>x y. indicator X (x, y) = indicator (Pair x -` X) y"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   313
    by (auto simp: indicator_def)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   314
  show ?thesis
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   315
    using X by (auto intro!: positive_integral_cong simp: emeasure_pair_measure sets_Pair1)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   316
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   317
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   318
lemma (in pair_sigma_finite) emeasure_pair_measure_Times:
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   319
  assumes A: "A \<in> sets M1" and B: "B \<in> sets M2"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   320
  shows "emeasure (M1 \<Otimes>\<^isub>M M2) (A \<times> B) = emeasure M1 A * emeasure M2 B"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   321
proof -
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   322
  have "emeasure (M1 \<Otimes>\<^isub>M M2) (A \<times> B) = (\<integral>\<^isup>+x. emeasure M2 B * indicator A x \<partial>M1)"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   323
    using A B by (auto intro!: positive_integral_cong simp: emeasure_pair_measure_alt)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   324
  also have "\<dots> = emeasure M2 B * emeasure M1 A"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   325
    using A by (simp add: emeasure_nonneg positive_integral_cmult_indicator)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   326
  finally show ?thesis
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   327
    by (simp add: ac_simps)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   328
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   329
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   330
lemma (in pair_sigma_finite) sigma_finite_up_in_pair_measure_generator:
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   331
  defines "E \<equiv> {A \<times> B | A B. A \<in> sets M1 \<and> B \<in> sets M2}"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   332
  shows "\<exists>F::nat \<Rightarrow> ('a \<times> 'b) set. range F \<subseteq> E \<and> incseq F \<and> (\<Union>i. F i) = space M1 \<times> space M2 \<and>
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   333
    (\<forall>i. emeasure (M1 \<Otimes>\<^isub>M M2) (F i) \<noteq> \<infinity>)"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   334
proof -
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   335
  from M1.sigma_finite_incseq guess F1 . note F1 = this
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   336
  from M2.sigma_finite_incseq guess F2 . note F2 = this
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   337
  from F1 F2 have space: "space M1 = (\<Union>i. F1 i)" "space M2 = (\<Union>i. F2 i)" by auto
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   338
  let ?F = "\<lambda>i. F1 i \<times> F2 i"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   339
  show ?thesis
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   340
  proof (intro exI[of _ ?F] conjI allI)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   341
    show "range ?F \<subseteq> E" using F1 F2 by (auto simp: E_def) (metis range_subsetD)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   342
  next
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   343
    have "space M1 \<times> space M2 \<subseteq> (\<Union>i. ?F i)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   344
    proof (intro subsetI)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   345
      fix x assume "x \<in> space M1 \<times> space M2"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   346
      then obtain i j where "fst x \<in> F1 i" "snd x \<in> F2 j"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   347
        by (auto simp: space)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   348
      then have "fst x \<in> F1 (max i j)" "snd x \<in> F2 (max j i)"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   349
        using `incseq F1` `incseq F2` unfolding incseq_def
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   350
        by (force split: split_max)+
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   351
      then have "(fst x, snd x) \<in> F1 (max i j) \<times> F2 (max i j)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   352
        by (intro SigmaI) (auto simp add: min_max.sup_commute)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   353
      then show "x \<in> (\<Union>i. ?F i)" by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   354
    qed
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   355
    then show "(\<Union>i. ?F i) = space M1 \<times> space M2"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   356
      using space by (auto simp: space)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   357
  next
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   358
    fix i show "incseq (\<lambda>i. F1 i \<times> F2 i)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   359
      using `incseq F1` `incseq F2` unfolding incseq_Suc_iff by auto
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   360
  next
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   361
    fix i
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   362
    from F1 F2 have "F1 i \<in> sets M1" "F2 i \<in> sets M2" by auto
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   363
    with F1 F2 emeasure_nonneg[of M1 "F1 i"] emeasure_nonneg[of M2 "F2 i"]
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   364
    show "emeasure (M1 \<Otimes>\<^isub>M M2) (F1 i \<times> F2 i) \<noteq> \<infinity>"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   365
      by (auto simp add: emeasure_pair_measure_Times)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   366
  qed
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   367
qed
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   368
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   369
sublocale pair_sigma_finite \<subseteq> sigma_finite_measure "M1 \<Otimes>\<^isub>M M2"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   370
proof
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   371
  from sigma_finite_up_in_pair_measure_generator guess F :: "nat \<Rightarrow> ('a \<times> 'b) set" .. note F = this
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   372
  show "\<exists>F::nat \<Rightarrow> ('a \<times> 'b) set. range F \<subseteq> sets (M1 \<Otimes>\<^isub>M M2) \<and> (\<Union>i. F i) = space (M1 \<Otimes>\<^isub>M M2) \<and> (\<forall>i. emeasure (M1 \<Otimes>\<^isub>M M2) (F i) \<noteq> \<infinity>)"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   373
  proof (rule exI[of _ F], intro conjI)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   374
    show "range F \<subseteq> sets (M1 \<Otimes>\<^isub>M M2)" using F by (auto simp: pair_measure_def)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   375
    show "(\<Union>i. F i) = space (M1 \<Otimes>\<^isub>M M2)"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   376
      using F by (auto simp: space_pair_measure)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   377
    show "\<forall>i. emeasure (M1 \<Otimes>\<^isub>M M2) (F i) \<noteq> \<infinity>" using F by auto
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   378
  qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   379
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   380
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   381
lemma sigma_finite_pair_measure:
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   382
  assumes A: "sigma_finite_measure A" and B: "sigma_finite_measure B"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   383
  shows "sigma_finite_measure (A \<Otimes>\<^isub>M B)"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   384
proof -
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   385
  interpret A: sigma_finite_measure A by fact
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   386
  interpret B: sigma_finite_measure B by fact
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   387
  interpret AB: pair_sigma_finite A  B ..
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   388
  show ?thesis ..
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   389
qed
39088
ca17017c10e6 Measurable on product space is equiv. to measurable components
hoelzl
parents: 39082
diff changeset
   390
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   391
lemma sets_pair_swap:
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   392
  assumes "A \<in> sets (M1 \<Otimes>\<^isub>M M2)"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   393
  shows "(\<lambda>(x, y). (y, x)) -` A \<inter> space (M2 \<Otimes>\<^isub>M M1) \<in> sets (M2 \<Otimes>\<^isub>M M1)"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   394
  using measurable_pair_swap' assms by (rule measurable_sets)
41661
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41659
diff changeset
   395
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   396
lemma (in pair_sigma_finite) distr_pair_swap:
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   397
  "M1 \<Otimes>\<^isub>M M2 = distr (M2 \<Otimes>\<^isub>M M1) (M1 \<Otimes>\<^isub>M M2) (\<lambda>(x, y). (y, x))" (is "?P = ?D")
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   398
proof -
41706
a207a858d1f6 prefer p2e before e2p; use measure_unique_Int_stable_vimage;
hoelzl
parents: 41705
diff changeset
   399
  interpret Q: pair_sigma_finite M2 M1 by default
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   400
  from sigma_finite_up_in_pair_measure_generator guess F :: "nat \<Rightarrow> ('a \<times> 'b) set" .. note F = this
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   401
  let ?E = "{a \<times> b |a b. a \<in> sets M1 \<and> b \<in> sets M2}"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   402
  show ?thesis
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   403
  proof (rule measure_eqI_generator_eq[OF Int_stable_pair_measure_generator[of M1 M2]])
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   404
    show "?E \<subseteq> Pow (space ?P)"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   405
      using space_closed[of M1] space_closed[of M2] by (auto simp: space_pair_measure)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   406
    show "sets ?P = sigma_sets (space ?P) ?E"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   407
      by (simp add: sets_pair_measure space_pair_measure)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   408
    then show "sets ?D = sigma_sets (space ?P) ?E"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   409
      by simp
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   410
  next
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   411
    show "range F \<subseteq> ?E" "incseq F" "(\<Union>i. F i) = space ?P" "\<And>i. emeasure ?P (F i) \<noteq> \<infinity>"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   412
      using F by (auto simp: space_pair_measure)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   413
  next
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   414
    fix X assume "X \<in> ?E"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   415
    then obtain A B where X[simp]: "X = A \<times> B" and A: "A \<in> sets M1" and B: "B \<in> sets M2" by auto
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   416
    have "(\<lambda>(y, x). (x, y)) -` X \<inter> space (M2 \<Otimes>\<^isub>M M1) = B \<times> A"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   417
      using sets_into_space[OF A] sets_into_space[OF B] by (auto simp: space_pair_measure)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   418
    with A B show "emeasure (M1 \<Otimes>\<^isub>M M2) X = emeasure ?D X"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   419
      by (simp add: emeasure_pair_measure_Times Q.emeasure_pair_measure_Times emeasure_distr
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   420
                    measurable_pair_swap' ac_simps)
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   421
  qed
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   422
qed
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   423
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   424
lemma (in pair_sigma_finite) emeasure_pair_measure_alt2:
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   425
  assumes A: "A \<in> sets (M1 \<Otimes>\<^isub>M M2)"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   426
  shows "emeasure (M1 \<Otimes>\<^isub>M M2) A = (\<integral>\<^isup>+y. emeasure M1 ((\<lambda>x. (x, y)) -` A) \<partial>M2)"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   427
    (is "_ = ?\<nu> A")
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   428
proof -
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   429
  interpret Q: pair_sigma_finite M2 M1 by default
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   430
  have [simp]: "\<And>y. (Pair y -` ((\<lambda>(x, y). (y, x)) -` A \<inter> space (M2 \<Otimes>\<^isub>M M1))) = (\<lambda>x. (x, y)) -` A"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   431
    using sets_into_space[OF A] by (auto simp: space_pair_measure)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   432
  show ?thesis using A
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   433
    by (subst distr_pair_swap)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   434
       (simp_all del: vimage_Int add: measurable_sets[OF measurable_pair_swap']
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   435
                 Q.emeasure_pair_measure_alt emeasure_distr[OF measurable_pair_swap' A])
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   436
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   437
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   438
section "Fubinis theorem"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   439
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   440
lemma (in pair_sigma_finite) simple_function_cut:
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   441
  assumes f: "simple_function (M1 \<Otimes>\<^isub>M M2) f" "\<And>x. 0 \<le> f x"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   442
  shows "(\<lambda>x. \<integral>\<^isup>+y. f (x, y) \<partial>M2) \<in> borel_measurable M1"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   443
    and "(\<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. f (x, y) \<partial>M2) \<partial>M1) = integral\<^isup>P (M1 \<Otimes>\<^isub>M M2) f"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   444
proof -
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   445
  have f_borel: "f \<in> borel_measurable (M1 \<Otimes>\<^isub>M M2)"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   446
    using f(1) by (rule borel_measurable_simple_function)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   447
  let ?F = "\<lambda>z. f -` {z} \<inter> space (M1 \<Otimes>\<^isub>M M2)"
46731
5302e932d1e5 avoid undeclared variables in let bindings;
wenzelm
parents: 45777
diff changeset
   448
  let ?F' = "\<lambda>x z. Pair x -` ?F z"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   449
  { fix x assume "x \<in> space M1"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   450
    have [simp]: "\<And>z y. indicator (?F z) (x, y) = indicator (?F' x z) y"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   451
      by (auto simp: indicator_def)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   452
    have "\<And>y. y \<in> space M2 \<Longrightarrow> (x, y) \<in> space (M1 \<Otimes>\<^isub>M M2)" using `x \<in> space M1`
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   453
      by (simp add: space_pair_measure)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   454
    moreover have "\<And>x z. ?F' x z \<in> sets M2" using f_borel
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   455
      by (rule sets_Pair1[OF measurable_sets]) auto
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   456
    ultimately have "simple_function M2 (\<lambda> y. f (x, y))"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   457
      apply (rule_tac simple_function_cong[THEN iffD2, OF _])
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   458
      apply (rule simple_function_indicator_representation[OF f(1)])
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   459
      using `x \<in> space M1` by auto }
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   460
  note M2_sf = this
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   461
  { fix x assume x: "x \<in> space M1"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   462
    then have "(\<integral>\<^isup>+y. f (x, y) \<partial>M2) = (\<Sum>z\<in>f ` space (M1 \<Otimes>\<^isub>M M2). z * emeasure M2 (?F' x z))"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   463
      unfolding positive_integral_eq_simple_integral[OF M2_sf[OF x] f(2)]
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   464
      unfolding simple_integral_def
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   465
    proof (safe intro!: setsum_mono_zero_cong_left)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   466
      from f(1) show "finite (f ` space (M1 \<Otimes>\<^isub>M M2))" by (rule simple_functionD)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   467
    next
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   468
      fix y assume "y \<in> space M2" then show "f (x, y) \<in> f ` space (M1 \<Otimes>\<^isub>M M2)"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   469
        using `x \<in> space M1` by (auto simp: space_pair_measure)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   470
    next
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   471
      fix x' y assume "(x', y) \<in> space (M1 \<Otimes>\<^isub>M M2)"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   472
        "f (x', y) \<notin> (\<lambda>y. f (x, y)) ` space M2"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   473
      then have *: "?F' x (f (x', y)) = {}"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   474
        by (force simp: space_pair_measure)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   475
      show  "f (x', y) * emeasure M2 (?F' x (f (x', y))) = 0"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   476
        unfolding * by simp
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   477
    qed (simp add: vimage_compose[symmetric] comp_def
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   478
                   space_pair_measure) }
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   479
  note eq = this
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   480
  moreover have "\<And>z. ?F z \<in> sets (M1 \<Otimes>\<^isub>M M2)"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   481
    by (auto intro!: f_borel borel_measurable_vimage)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   482
  moreover then have "\<And>z. (\<lambda>x. emeasure M2 (?F' x z)) \<in> borel_measurable M1"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   483
    by (auto intro!: measurable_emeasure_Pair1 simp del: vimage_Int)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   484
  moreover have *: "\<And>i x. 0 \<le> emeasure M2 (Pair x -` (f -` {i} \<inter> space (M1 \<Otimes>\<^isub>M M2)))"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   485
    using f(1)[THEN simple_functionD(2)] f(2) by (intro emeasure_nonneg)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   486
  moreover { fix i assume "i \<in> f`space (M1 \<Otimes>\<^isub>M M2)"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   487
    with * have "\<And>x. 0 \<le> i * emeasure M2 (Pair x -` (f -` {i} \<inter> space (M1 \<Otimes>\<^isub>M M2)))"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   488
      using f(2) by auto }
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   489
  ultimately
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   490
  show "(\<lambda>x. \<integral>\<^isup>+y. f (x, y) \<partial>M2) \<in> borel_measurable M1"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   491
    and "(\<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. f (x, y) \<partial>M2) \<partial>M1) = integral\<^isup>P (M1 \<Otimes>\<^isub>M M2) f" using f(2)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   492
    by (auto simp del: vimage_Int cong: measurable_cong intro!: setsum_cong
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   493
             simp add: positive_integral_setsum simple_integral_def
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   494
                       positive_integral_cmult
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   495
                       positive_integral_cong[OF eq]
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   496
                       positive_integral_eq_simple_integral[OF f]
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   497
                       emeasure_pair_measure_alt[symmetric])
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   498
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   499
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   500
lemma (in pair_sigma_finite) positive_integral_fst_measurable:
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   501
  assumes f: "f \<in> borel_measurable (M1 \<Otimes>\<^isub>M M2)"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   502
  shows "(\<lambda>x. \<integral>\<^isup>+ y. f (x, y) \<partial>M2) \<in> borel_measurable M1"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   503
      (is "?C f \<in> borel_measurable M1")
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   504
    and "(\<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. f (x, y) \<partial>M2) \<partial>M1) = integral\<^isup>P (M1 \<Otimes>\<^isub>M M2) f"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   505
proof -
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   506
  from borel_measurable_implies_simple_function_sequence'[OF f] guess F . note F = this
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   507
  then have F_borel: "\<And>i. F i \<in> borel_measurable (M1 \<Otimes>\<^isub>M M2)"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   508
    by (auto intro: borel_measurable_simple_function)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   509
  note sf = simple_function_cut[OF F(1,5)]
41097
a1abfa4e2b44 use SUPR_ and INFI_apply instead of SUPR_, INFI_fun_expand
hoelzl
parents: 41096
diff changeset
   510
  then have "(\<lambda>x. SUP i. ?C (F i) x) \<in> borel_measurable M1"
a1abfa4e2b44 use SUPR_ and INFI_apply instead of SUPR_, INFI_fun_expand
hoelzl
parents: 41096
diff changeset
   511
    using F(1) by auto
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   512
  moreover
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   513
  { fix x assume "x \<in> space M1"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   514
    from F measurable_Pair2[OF F_borel `x \<in> space M1`]
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   515
    have "(\<integral>\<^isup>+y. (SUP i. F i (x, y)) \<partial>M2) = (SUP i. ?C (F i) x)"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   516
      by (intro positive_integral_monotone_convergence_SUP)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   517
         (auto simp: incseq_Suc_iff le_fun_def)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   518
    then have "(SUP i. ?C (F i) x) = ?C f x"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   519
      unfolding F(4) positive_integral_max_0 by simp }
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   520
  note SUPR_C = this
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   521
  ultimately show "?C f \<in> borel_measurable M1"
41097
a1abfa4e2b44 use SUPR_ and INFI_apply instead of SUPR_, INFI_fun_expand
hoelzl
parents: 41096
diff changeset
   522
    by (simp cong: measurable_cong)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   523
  have "(\<integral>\<^isup>+x. (SUP i. F i x) \<partial>(M1 \<Otimes>\<^isub>M M2)) = (SUP i. integral\<^isup>P (M1 \<Otimes>\<^isub>M M2) (F i))"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   524
    using F_borel F
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   525
    by (intro positive_integral_monotone_convergence_SUP) auto
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   526
  also have "(SUP i. integral\<^isup>P (M1 \<Otimes>\<^isub>M M2) (F i)) = (SUP i. \<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. F i (x, y) \<partial>M2) \<partial>M1)"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   527
    unfolding sf(2) by simp
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   528
  also have "\<dots> = \<integral>\<^isup>+ x. (SUP i. \<integral>\<^isup>+ y. F i (x, y) \<partial>M2) \<partial>M1" using F sf(1)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   529
    by (intro positive_integral_monotone_convergence_SUP[symmetric])
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   530
       (auto intro!: positive_integral_mono positive_integral_positive
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   531
             simp: incseq_Suc_iff le_fun_def)
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   532
  also have "\<dots> = \<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. (SUP i. F i (x, y)) \<partial>M2) \<partial>M1"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   533
    using F_borel F(2,5)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   534
    by (auto intro!: positive_integral_cong positive_integral_monotone_convergence_SUP[symmetric] measurable_Pair2
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   535
             simp: incseq_Suc_iff le_fun_def)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   536
  finally show "(\<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. f (x, y) \<partial>M2) \<partial>M1) = integral\<^isup>P (M1 \<Otimes>\<^isub>M M2) f"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   537
    using F by (simp add: positive_integral_max_0)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   538
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   539
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   540
lemma (in pair_sigma_finite) positive_integral_snd_measurable:
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   541
  assumes f: "f \<in> borel_measurable (M1 \<Otimes>\<^isub>M M2)"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   542
  shows "(\<integral>\<^isup>+ y. (\<integral>\<^isup>+ x. f (x, y) \<partial>M1) \<partial>M2) = integral\<^isup>P (M1 \<Otimes>\<^isub>M M2) f"
41661
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41659
diff changeset
   543
proof -
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   544
  interpret Q: pair_sigma_finite M2 M1 by default
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   545
  note measurable_pair_swap[OF f]
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   546
  from Q.positive_integral_fst_measurable[OF this]
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   547
  have "(\<integral>\<^isup>+ y. (\<integral>\<^isup>+ x. f (x, y) \<partial>M1) \<partial>M2) = (\<integral>\<^isup>+ (x, y). f (y, x) \<partial>(M2 \<Otimes>\<^isub>M M1))"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   548
    by simp
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   549
  also have "(\<integral>\<^isup>+ (x, y). f (y, x) \<partial>(M2 \<Otimes>\<^isub>M M1)) = integral\<^isup>P (M1 \<Otimes>\<^isub>M M2) f"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   550
    by (subst distr_pair_swap)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   551
       (auto simp: positive_integral_distr[OF measurable_pair_swap' f] intro!: positive_integral_cong)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   552
  finally show ?thesis .
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   553
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   554
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   555
lemma (in pair_sigma_finite) Fubini:
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   556
  assumes f: "f \<in> borel_measurable (M1 \<Otimes>\<^isub>M M2)"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   557
  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)"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   558
  unfolding positive_integral_snd_measurable[OF assms]
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   559
  unfolding positive_integral_fst_measurable[OF assms] ..
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   560
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   561
lemma (in pair_sigma_finite) AE_pair:
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   562
  assumes "AE x in (M1 \<Otimes>\<^isub>M M2). Q x"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   563
  shows "AE x in M1. (AE y in M2. Q (x, y))"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   564
proof -
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   565
  obtain N where N: "N \<in> sets (M1 \<Otimes>\<^isub>M M2)" "emeasure (M1 \<Otimes>\<^isub>M M2) N = 0" "{x\<in>space (M1 \<Otimes>\<^isub>M M2). \<not> Q x} \<subseteq> N"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   566
    using assms unfolding eventually_ae_filter by auto
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   567
  show ?thesis
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   568
  proof (rule AE_I)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   569
    from N measurable_emeasure_Pair1[OF `N \<in> sets (M1 \<Otimes>\<^isub>M M2)`]
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   570
    show "emeasure M1 {x\<in>space M1. emeasure M2 (Pair x -` N) \<noteq> 0} = 0"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   571
      by (auto simp: emeasure_pair_measure_alt positive_integral_0_iff emeasure_nonneg)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   572
    show "{x \<in> space M1. emeasure M2 (Pair x -` N) \<noteq> 0} \<in> sets M1"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   573
      by (intro borel_measurable_ereal_neq_const measurable_emeasure_Pair1 N)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   574
    { fix x assume "x \<in> space M1" "emeasure M2 (Pair x -` N) = 0"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   575
      have "AE y in M2. Q (x, y)"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   576
      proof (rule AE_I)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   577
        show "emeasure M2 (Pair x -` N) = 0" by fact
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   578
        show "Pair x -` N \<in> sets M2" using N(1) by (rule sets_Pair1)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   579
        show "{y \<in> space M2. \<not> Q (x, y)} \<subseteq> Pair x -` N"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   580
          using N `x \<in> space M1` unfolding space_pair_measure by auto
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   581
      qed }
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   582
    then show "{x \<in> space M1. \<not> (AE y in M2. Q (x, y))} \<subseteq> {x \<in> space M1. emeasure M2 (Pair x -` N) \<noteq> 0}"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   583
      by auto
39088
ca17017c10e6 Measurable on product space is equiv. to measurable components
hoelzl
parents: 39082
diff changeset
   584
  qed
ca17017c10e6 Measurable on product space is equiv. to measurable components
hoelzl
parents: 39082
diff changeset
   585
qed
35833
7b7ae5aa396d Added product measure space
hoelzl
parents:
diff changeset
   586
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   587
lemma (in pair_sigma_finite) AE_pair_measure:
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   588
  assumes "{x\<in>space (M1 \<Otimes>\<^isub>M M2). P x} \<in> sets (M1 \<Otimes>\<^isub>M M2)"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   589
  assumes ae: "AE x in M1. AE y in M2. P (x, y)"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   590
  shows "AE x in M1 \<Otimes>\<^isub>M M2. P x"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   591
proof (subst AE_iff_measurable[OF _ refl])
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   592
  show "{x\<in>space (M1 \<Otimes>\<^isub>M M2). \<not> P x} \<in> sets (M1 \<Otimes>\<^isub>M M2)"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   593
    by (rule sets_Collect) fact
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   594
  then have "emeasure (M1 \<Otimes>\<^isub>M M2) {x \<in> space (M1 \<Otimes>\<^isub>M M2). \<not> P x} =
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   595
      (\<integral>\<^isup>+ x. \<integral>\<^isup>+ y. indicator {x \<in> space (M1 \<Otimes>\<^isub>M M2). \<not> P x} (x, y) \<partial>M2 \<partial>M1)"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   596
    by (simp add: emeasure_pair_measure)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   597
  also have "\<dots> = (\<integral>\<^isup>+ x. \<integral>\<^isup>+ y. 0 \<partial>M2 \<partial>M1)"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   598
    using ae
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   599
    apply (safe intro!: positive_integral_cong_AE)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   600
    apply (intro AE_I2)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   601
    apply (safe intro!: positive_integral_cong_AE)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   602
    apply auto
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   603
    done
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   604
  finally show "emeasure (M1 \<Otimes>\<^isub>M M2) {x \<in> space (M1 \<Otimes>\<^isub>M M2). \<not> P x} = 0" by simp
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   605
qed
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   606
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   607
lemma (in pair_sigma_finite) AE_pair_iff:
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   608
  "{x\<in>space (M1 \<Otimes>\<^isub>M M2). P (fst x) (snd x)} \<in> sets (M1 \<Otimes>\<^isub>M M2) \<Longrightarrow>
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   609
    (AE x in M1. AE y in M2. P x y) \<longleftrightarrow> (AE x in (M1 \<Otimes>\<^isub>M M2). P (fst x) (snd x))"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   610
  using AE_pair[of "\<lambda>x. P (fst x) (snd x)"] AE_pair_measure[of "\<lambda>x. P (fst x) (snd x)"] by auto
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   611
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   612
lemma AE_distr_iff:
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   613
  assumes f: "f \<in> measurable M N" and P: "{x \<in> space N. P x} \<in> sets N"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   614
  shows "(AE x in distr M N f. P x) \<longleftrightarrow> (AE x in M. P (f x))"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   615
proof (subst (1 2) AE_iff_measurable[OF _ refl])
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   616
  from P show "{x \<in> space (distr M N f). \<not> P x} \<in> sets (distr M N f)"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   617
    by (auto intro!: sets_Collect_neg)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   618
  moreover
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   619
  have "f -` {x \<in> space N. P x} \<inter> space M = {x \<in> space M. P (f x)}"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   620
    using f by (auto dest: measurable_space)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   621
  then show "{x \<in> space M. \<not> P (f x)} \<in> sets M"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   622
    using measurable_sets[OF f P] by (auto intro!: sets_Collect_neg)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   623
  moreover have "f -` {x\<in>space N. \<not> P x} \<inter> space M = {x \<in> space M. \<not> P (f x)}"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   624
    using f by (auto dest: measurable_space)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   625
  ultimately show "(emeasure (distr M N f) {x \<in> space (distr M N f). \<not> P x} = 0) =
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   626
    (emeasure M {x \<in> space M. \<not> P (f x)} = 0)"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   627
    using f by (simp add: emeasure_distr)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   628
qed
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   629
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   630
lemma (in pair_sigma_finite) AE_commute:
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   631
  assumes P: "{x\<in>space (M1 \<Otimes>\<^isub>M M2). P (fst x) (snd x)} \<in> sets (M1 \<Otimes>\<^isub>M M2)"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   632
  shows "(AE x in M1. AE y in M2. P x y) \<longleftrightarrow> (AE y in M2. AE x in M1. P x y)"
41026
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41023
diff changeset
   633
proof -
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   634
  interpret Q: pair_sigma_finite M2 M1 ..
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   635
  have [simp]: "\<And>x. (fst (case x of (x, y) \<Rightarrow> (y, x))) = snd x" "\<And>x. (snd (case x of (x, y) \<Rightarrow> (y, x))) = fst x"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   636
    by auto
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   637
  have "{x \<in> space (M2 \<Otimes>\<^isub>M M1). P (snd x) (fst x)} =
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   638
    (\<lambda>(x, y). (y, x)) -` {x \<in> space (M1 \<Otimes>\<^isub>M M2). P (fst x) (snd x)} \<inter> space (M2 \<Otimes>\<^isub>M M1)"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   639
    by (auto simp: space_pair_measure)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   640
  also have "\<dots> \<in> sets (M2 \<Otimes>\<^isub>M M1)"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   641
    by (intro sets_pair_swap P)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   642
  finally show ?thesis
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   643
    apply (subst AE_pair_iff[OF P])
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   644
    apply (subst distr_pair_swap)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   645
    apply (subst AE_distr_iff[OF measurable_pair_swap' P])
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   646
    apply (subst Q.AE_pair_iff)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   647
    apply simp_all
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   648
    done
41026
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41023
diff changeset
   649
qed
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41023
diff changeset
   650
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41023
diff changeset
   651
lemma (in pair_sigma_finite) integrable_product_swap:
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   652
  assumes "integrable (M1 \<Otimes>\<^isub>M M2) f"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   653
  shows "integrable (M2 \<Otimes>\<^isub>M M1) (\<lambda>(x,y). f (y,x))"
41026
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41023
diff changeset
   654
proof -
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   655
  interpret Q: pair_sigma_finite M2 M1 by default
41661
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41659
diff changeset
   656
  have *: "(\<lambda>(x,y). f (y,x)) = (\<lambda>x. f (case x of (x,y)\<Rightarrow>(y,x)))" by (auto simp: fun_eq_iff)
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41659
diff changeset
   657
  show ?thesis unfolding *
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   658
    by (rule integrable_distr[OF measurable_pair_swap'])
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   659
       (simp add: distr_pair_swap[symmetric] assms)
41661
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41659
diff changeset
   660
qed
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41659
diff changeset
   661
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41659
diff changeset
   662
lemma (in pair_sigma_finite) integrable_product_swap_iff:
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   663
  "integrable (M2 \<Otimes>\<^isub>M M1) (\<lambda>(x,y). f (y,x)) \<longleftrightarrow> integrable (M1 \<Otimes>\<^isub>M M2) f"
41661
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41659
diff changeset
   664
proof -
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   665
  interpret Q: pair_sigma_finite M2 M1 by default
41661
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41659
diff changeset
   666
  from Q.integrable_product_swap[of "\<lambda>(x,y). f (y,x)"] integrable_product_swap[of f]
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41659
diff changeset
   667
  show ?thesis by auto
41026
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41023
diff changeset
   668
qed
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41023
diff changeset
   669
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41023
diff changeset
   670
lemma (in pair_sigma_finite) integral_product_swap:
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   671
  assumes f: "f \<in> borel_measurable (M1 \<Otimes>\<^isub>M M2)"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   672
  shows "(\<integral>(x,y). f (y,x) \<partial>(M2 \<Otimes>\<^isub>M M1)) = integral\<^isup>L (M1 \<Otimes>\<^isub>M M2) f"
41026
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41023
diff changeset
   673
proof -
41661
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41659
diff changeset
   674
  have *: "(\<lambda>(x,y). f (y,x)) = (\<lambda>x. f (case x of (x,y)\<Rightarrow>(y,x)))" by (auto simp: fun_eq_iff)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   675
  show ?thesis unfolding *
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   676
    by (simp add: integral_distr[symmetric, OF measurable_pair_swap' f] distr_pair_swap[symmetric])
41026
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41023
diff changeset
   677
qed
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41023
diff changeset
   678
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41023
diff changeset
   679
lemma (in pair_sigma_finite) integrable_fst_measurable:
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   680
  assumes f: "integrable (M1 \<Otimes>\<^isub>M M2) f"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   681
  shows "AE x in M1. integrable M2 (\<lambda> y. f (x, y))" (is "?AE")
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   682
    and "(\<integral>x. (\<integral>y. f (x, y) \<partial>M2) \<partial>M1) = integral\<^isup>L (M1 \<Otimes>\<^isub>M M2) f" (is "?INT")
41026
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41023
diff changeset
   683
proof -
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   684
  have f_borel: "f \<in> borel_measurable (M1 \<Otimes>\<^isub>M M2)"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   685
    using f by auto
46731
5302e932d1e5 avoid undeclared variables in let bindings;
wenzelm
parents: 45777
diff changeset
   686
  let ?pf = "\<lambda>x. ereal (f x)" and ?nf = "\<lambda>x. ereal (- f x)"
41026
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41023
diff changeset
   687
  have
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   688
    borel: "?nf \<in> borel_measurable (M1 \<Otimes>\<^isub>M M2)""?pf \<in> borel_measurable (M1 \<Otimes>\<^isub>M M2)" and
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   689
    int: "integral\<^isup>P (M1 \<Otimes>\<^isub>M M2) ?nf \<noteq> \<infinity>" "integral\<^isup>P (M1 \<Otimes>\<^isub>M M2) ?pf \<noteq> \<infinity>"
41026
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41023
diff changeset
   690
    using assms by auto
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42984
diff changeset
   691
  have "(\<integral>\<^isup>+x. (\<integral>\<^isup>+y. ereal (f (x, y)) \<partial>M2) \<partial>M1) \<noteq> \<infinity>"
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42984
diff changeset
   692
     "(\<integral>\<^isup>+x. (\<integral>\<^isup>+y. ereal (- f (x, y)) \<partial>M2) \<partial>M1) \<noteq> \<infinity>"
41026
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41023
diff changeset
   693
    using borel[THEN positive_integral_fst_measurable(1)] int
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41023
diff changeset
   694
    unfolding borel[THEN positive_integral_fst_measurable(2)] by simp_all
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41023
diff changeset
   695
  with borel[THEN positive_integral_fst_measurable(1)]
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42984
diff changeset
   696
  have AE_pos: "AE x in M1. (\<integral>\<^isup>+y. ereal (f (x, y)) \<partial>M2) \<noteq> \<infinity>"
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42984
diff changeset
   697
    "AE x in M1. (\<integral>\<^isup>+y. ereal (- f (x, y)) \<partial>M2) \<noteq> \<infinity>"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   698
    by (auto intro!: positive_integral_PInf_AE )
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42984
diff changeset
   699
  then have AE: "AE x in M1. \<bar>\<integral>\<^isup>+y. ereal (f (x, y)) \<partial>M2\<bar> \<noteq> \<infinity>"
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42984
diff changeset
   700
    "AE x in M1. \<bar>\<integral>\<^isup>+y. ereal (- f (x, y)) \<partial>M2\<bar> \<noteq> \<infinity>"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   701
    by (auto simp: positive_integral_positive)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   702
  from AE_pos show ?AE using assms
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   703
    by (simp add: measurable_Pair2[OF f_borel] integrable_def)
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42984
diff changeset
   704
  { fix f have "(\<integral>\<^isup>+ x. - \<integral>\<^isup>+ y. ereal (f x y) \<partial>M2 \<partial>M1) = (\<integral>\<^isup>+x. 0 \<partial>M1)"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   705
      using positive_integral_positive
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   706
      by (intro positive_integral_cong_pos) (auto simp: ereal_uminus_le_reorder)
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42984
diff changeset
   707
    then have "(\<integral>\<^isup>+ x. - \<integral>\<^isup>+ y. ereal (f x y) \<partial>M2 \<partial>M1) = 0" by simp }
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   708
  note this[simp]
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   709
  { fix f assume borel: "(\<lambda>x. ereal (f x)) \<in> borel_measurable (M1 \<Otimes>\<^isub>M M2)"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   710
      and int: "integral\<^isup>P (M1 \<Otimes>\<^isub>M M2) (\<lambda>x. ereal (f x)) \<noteq> \<infinity>"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   711
      and AE: "AE x in M1. (\<integral>\<^isup>+y. ereal (f (x, y)) \<partial>M2) \<noteq> \<infinity>"
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42984
diff changeset
   712
    have "integrable M1 (\<lambda>x. real (\<integral>\<^isup>+y. ereal (f (x, y)) \<partial>M2))" (is "integrable M1 ?f")
41705
1100512e16d8 add auto support for AE_mp
hoelzl
parents: 41689
diff changeset
   713
    proof (intro integrable_def[THEN iffD2] conjI)
1100512e16d8 add auto support for AE_mp
hoelzl
parents: 41689
diff changeset
   714
      show "?f \<in> borel_measurable M1"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   715
        using borel by (auto intro!: positive_integral_fst_measurable)
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42984
diff changeset
   716
      have "(\<integral>\<^isup>+x. ereal (?f x) \<partial>M1) = (\<integral>\<^isup>+x. (\<integral>\<^isup>+y. ereal (f (x, y))  \<partial>M2) \<partial>M1)"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   717
        using AE positive_integral_positive[of M2]
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   718
        by (auto intro!: positive_integral_cong_AE simp: ereal_real)
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42984
diff changeset
   719
      then show "(\<integral>\<^isup>+x. ereal (?f x) \<partial>M1) \<noteq> \<infinity>"
41705
1100512e16d8 add auto support for AE_mp
hoelzl
parents: 41689
diff changeset
   720
        using positive_integral_fst_measurable[OF borel] int by simp
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42984
diff changeset
   721
      have "(\<integral>\<^isup>+x. ereal (- ?f x) \<partial>M1) = (\<integral>\<^isup>+x. 0 \<partial>M1)"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   722
        by (intro positive_integral_cong_pos)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   723
           (simp add: positive_integral_positive real_of_ereal_pos)
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42984
diff changeset
   724
      then show "(\<integral>\<^isup>+x. ereal (- ?f x) \<partial>M1) \<noteq> \<infinity>" by simp
41705
1100512e16d8 add auto support for AE_mp
hoelzl
parents: 41689
diff changeset
   725
    qed }
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   726
  with this[OF borel(1) int(1) AE_pos(2)] this[OF borel(2) int(2) AE_pos(1)]
41705
1100512e16d8 add auto support for AE_mp
hoelzl
parents: 41689
diff changeset
   727
  show ?INT
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   728
    unfolding lebesgue_integral_def[of "M1 \<Otimes>\<^isub>M M2"] lebesgue_integral_def[of M2]
41026
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41023
diff changeset
   729
      borel[THEN positive_integral_fst_measurable(2), symmetric]
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   730
    using AE[THEN integral_real]
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   731
    by simp
41026
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41023
diff changeset
   732
qed
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41023
diff changeset
   733
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41023
diff changeset
   734
lemma (in pair_sigma_finite) integrable_snd_measurable:
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   735
  assumes f: "integrable (M1 \<Otimes>\<^isub>M M2) f"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   736
  shows "AE y in M2. integrable M1 (\<lambda>x. f (x, y))" (is "?AE")
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   737
    and "(\<integral>y. (\<integral>x. f (x, y) \<partial>M1) \<partial>M2) = integral\<^isup>L (M1 \<Otimes>\<^isub>M M2) f" (is "?INT")
41026
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41023
diff changeset
   738
proof -
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   739
  interpret Q: pair_sigma_finite M2 M1 by default
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   740
  have Q_int: "integrable (M2 \<Otimes>\<^isub>M M1) (\<lambda>(x, y). f (y, x))"
41661
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41659
diff changeset
   741
    using f unfolding integrable_product_swap_iff .
41026
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41023
diff changeset
   742
  show ?INT
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41023
diff changeset
   743
    using Q.integrable_fst_measurable(2)[OF Q_int]
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   744
    using integral_product_swap[of f] f by auto
41026
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41023
diff changeset
   745
  show ?AE
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41023
diff changeset
   746
    using Q.integrable_fst_measurable(1)[OF Q_int]
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41023
diff changeset
   747
    by simp
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41023
diff changeset
   748
qed
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41023
diff changeset
   749
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   750
lemma (in pair_sigma_finite) positive_integral_fst_measurable':
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   751
  assumes f: "(\<lambda>x. f (fst x) (snd x)) \<in> borel_measurable (M1 \<Otimes>\<^isub>M M2)"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   752
  shows "(\<lambda>x. \<integral>\<^isup>+ y. f x y \<partial>M2) \<in> borel_measurable M1"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   753
  using positive_integral_fst_measurable(1)[OF f] by simp
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   754
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   755
lemma (in pair_sigma_finite) integral_fst_measurable:
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   756
  "(\<lambda>x. f (fst x) (snd x)) \<in> borel_measurable (M1 \<Otimes>\<^isub>M M2) \<Longrightarrow> (\<lambda>x. \<integral> y. f x y \<partial>M2) \<in> borel_measurable M1"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   757
  by (auto simp: lebesgue_integral_def intro!: borel_measurable_diff positive_integral_fst_measurable')
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   758
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   759
lemma (in pair_sigma_finite) positive_integral_snd_measurable':
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   760
  assumes f: "(\<lambda>x. f (fst x) (snd x)) \<in> borel_measurable (M1 \<Otimes>\<^isub>M M2)"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   761
  shows "(\<lambda>y. \<integral>\<^isup>+ x. f x y \<partial>M1) \<in> borel_measurable M2"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   762
proof -
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   763
  interpret Q: pair_sigma_finite M2 M1 ..
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   764
  show ?thesis
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   765
    using measurable_pair_swap[OF f]
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   766
    by (intro Q.positive_integral_fst_measurable') (simp add: split_beta')
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   767
qed
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   768
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   769
lemma (in pair_sigma_finite) integral_snd_measurable:
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   770
  "(\<lambda>x. f (fst x) (snd x)) \<in> borel_measurable (M1 \<Otimes>\<^isub>M M2) \<Longrightarrow> (\<lambda>y. \<integral> x. f x y \<partial>M1) \<in> borel_measurable M2"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   771
  by (auto simp: lebesgue_integral_def intro!: borel_measurable_diff positive_integral_snd_measurable')
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   772
41026
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41023
diff changeset
   773
lemma (in pair_sigma_finite) Fubini_integral:
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   774
  assumes f: "integrable (M1 \<Otimes>\<^isub>M M2) f"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   775
  shows "(\<integral>y. (\<integral>x. f (x, y) \<partial>M1) \<partial>M2) = (\<integral>x. (\<integral>y. f (x, y) \<partial>M2) \<partial>M1)"
41026
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41023
diff changeset
   776
  unfolding integrable_snd_measurable[OF assms]
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41023
diff changeset
   777
  unfolding integrable_fst_measurable[OF assms] ..
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41023
diff changeset
   778
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   779
section {* Products on counting spaces, densities and distributions *}
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   780
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   781
lemma sigma_sets_pair_measure_generator_finite:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36649
diff changeset
   782
  assumes "finite A" and "finite B"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   783
  shows "sigma_sets (A \<times> B) { a \<times> b | a b. a \<subseteq> A \<and> b \<subseteq> B} = Pow (A \<times> B)"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   784
  (is "sigma_sets ?prod ?sets = _")
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36649
diff changeset
   785
proof safe
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36649
diff changeset
   786
  have fin: "finite (A \<times> B)" using assms by (rule finite_cartesian_product)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36649
diff changeset
   787
  fix x assume subset: "x \<subseteq> A \<times> B"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36649
diff changeset
   788
  hence "finite x" using fin by (rule finite_subset)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   789
  from this subset show "x \<in> sigma_sets ?prod ?sets"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36649
diff changeset
   790
  proof (induct x)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36649
diff changeset
   791
    case empty show ?case by (rule sigma_sets.Empty)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36649
diff changeset
   792
  next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36649
diff changeset
   793
    case (insert a x)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   794
    hence "{a} \<in> sigma_sets ?prod ?sets" by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36649
diff changeset
   795
    moreover have "x \<in> sigma_sets ?prod ?sets" using insert by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36649
diff changeset
   796
    ultimately show ?case unfolding insert_is_Un[of a x] by (rule sigma_sets_Un)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36649
diff changeset
   797
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36649
diff changeset
   798
next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36649
diff changeset
   799
  fix x a b
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   800
  assume "x \<in> sigma_sets ?prod ?sets" and "(a, b) \<in> x"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36649
diff changeset
   801
  from sigma_sets_into_sp[OF _ this(1)] this(2)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   802
  show "a \<in> A" and "b \<in> B" by auto
35833
7b7ae5aa396d Added product measure space
hoelzl
parents:
diff changeset
   803
qed
7b7ae5aa396d Added product measure space
hoelzl
parents:
diff changeset
   804
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   805
lemma pair_measure_count_space:
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   806
  assumes A: "finite A" and B: "finite B"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   807
  shows "count_space A \<Otimes>\<^isub>M count_space B = count_space (A \<times> B)" (is "?P = ?C")
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   808
proof (rule measure_eqI)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   809
  interpret A: finite_measure "count_space A" by (rule finite_measure_count_space) fact
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   810
  interpret B: finite_measure "count_space B" by (rule finite_measure_count_space) fact
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   811
  interpret P: pair_sigma_finite "count_space A" "count_space B" by default
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   812
  show eq: "sets ?P = sets ?C"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   813
    by (simp add: sets_pair_measure sigma_sets_pair_measure_generator_finite A B)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   814
  fix X assume X: "X \<in> sets ?P"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   815
  with eq have X_subset: "X \<subseteq> A \<times> B" by simp
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   816
  with A B have fin_Pair: "\<And>x. finite (Pair x -` X)"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   817
    by (intro finite_subset[OF _ B]) auto
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   818
  have fin_X: "finite X" using X_subset by (rule finite_subset) (auto simp: A B)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   819
  show "emeasure ?P X = emeasure ?C X"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   820
    apply (subst P.emeasure_pair_measure_alt[OF X])
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   821
    apply (subst emeasure_count_space)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   822
    using X_subset apply auto []
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   823
    apply (simp add: fin_Pair emeasure_count_space X_subset fin_X)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   824
    apply (subst positive_integral_count_space)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   825
    using A apply simp
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   826
    apply (simp del: real_of_nat_setsum add: real_of_nat_setsum[symmetric])
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   827
    apply (subst card_gt_0_iff)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   828
    apply (simp add: fin_Pair)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   829
    apply (subst card_SigmaI[symmetric])
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   830
    using A apply simp
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   831
    using fin_Pair apply simp
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   832
    using X_subset apply (auto intro!: arg_cong[where f=card])
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   833
    done
45777
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 44890
diff changeset
   834
qed
35833
7b7ae5aa396d Added product measure space
hoelzl
parents:
diff changeset
   835
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   836
lemma pair_measure_density:
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   837
  assumes f: "f \<in> borel_measurable M1" "AE x in M1. 0 \<le> f x"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   838
  assumes g: "g \<in> borel_measurable M2" "AE x in M2. 0 \<le> g x"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   839
  assumes "sigma_finite_measure M1" "sigma_finite_measure M2"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   840
  assumes "sigma_finite_measure (density M1 f)" "sigma_finite_measure (density M2 g)"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   841
  shows "density M1 f \<Otimes>\<^isub>M density M2 g = density (M1 \<Otimes>\<^isub>M M2) (\<lambda>(x,y). f x * g y)" (is "?L = ?R")
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   842
proof (rule measure_eqI)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   843
  interpret M1: sigma_finite_measure M1 by fact
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   844
  interpret M2: sigma_finite_measure M2 by fact
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   845
  interpret D1: sigma_finite_measure "density M1 f" by fact
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   846
  interpret D2: sigma_finite_measure "density M2 g" by fact
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   847
  interpret L: pair_sigma_finite "density M1 f" "density M2 g" ..
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   848
  interpret R: pair_sigma_finite M1 M2 ..
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   849
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   850
  fix A assume A: "A \<in> sets ?L"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   851
  then have indicator_eq: "\<And>x y. indicator A (x, y) = indicator (Pair x -` A) y"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   852
   and Pair_A: "\<And>x. Pair x -` A \<in> sets M2"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   853
    by (auto simp: indicator_def sets_Pair1)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   854
  have f_fst: "(\<lambda>p. f (fst p)) \<in> borel_measurable (M1 \<Otimes>\<^isub>M M2)"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   855
    using measurable_comp[OF measurable_fst f(1)] by (simp add: comp_def)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   856
  have g_snd: "(\<lambda>p. g (snd p)) \<in> borel_measurable (M1 \<Otimes>\<^isub>M M2)"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   857
    using measurable_comp[OF measurable_snd g(1)] by (simp add: comp_def)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   858
  have "(\<lambda>x. \<integral>\<^isup>+ y. g (snd (x, y)) * indicator A (x, y) \<partial>M2) \<in> borel_measurable M1"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   859
    using g_snd Pair_A A by (intro R.positive_integral_fst_measurable) auto
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   860
  then have int_g: "(\<lambda>x. \<integral>\<^isup>+ y. g y * indicator A (x, y) \<partial>M2) \<in> borel_measurable M1"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   861
    by simp
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36649
diff changeset
   862
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   863
  show "emeasure ?L A = emeasure ?R A"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   864
    apply (subst L.emeasure_pair_measure[OF A])
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   865
    apply (subst emeasure_density)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   866
        using f_fst g_snd apply (simp add: split_beta')
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   867
      using A apply simp
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   868
    apply (subst positive_integral_density[OF g])
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   869
      apply (simp add: indicator_eq Pair_A)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   870
    apply (subst positive_integral_density[OF f])
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   871
      apply (rule int_g)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   872
    apply (subst R.positive_integral_fst_measurable(2)[symmetric])
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   873
      using f g A Pair_A f_fst g_snd
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   874
      apply (auto intro!: positive_integral_cong_AE R.measurable_emeasure_Pair1
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   875
                  simp: positive_integral_cmult indicator_eq split_beta')
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   876
    apply (intro AE_I2 impI)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   877
    apply (subst mult_assoc)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   878
    apply (subst positive_integral_cmult)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   879
          apply auto
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   880
    done
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   881
qed simp
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   882
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   883
lemma sigma_finite_measure_distr:
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   884
  assumes "sigma_finite_measure (distr M N f)" and f: "f \<in> measurable M N"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   885
  shows "sigma_finite_measure M"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   886
proof -
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   887
  interpret sigma_finite_measure "distr M N f" by fact
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   888
  from sigma_finite_disjoint guess A . note A = this
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   889
  show ?thesis
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   890
  proof (unfold_locales, intro conjI exI allI)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   891
    show "range (\<lambda>i. f -` A i \<inter> space M) \<subseteq> sets M"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   892
      using A f by (auto intro!: measurable_sets)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   893
    show "(\<Union>i. f -` A i \<inter> space M) = space M"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   894
      using A(1) A(2)[symmetric] f by (auto simp: measurable_def Pi_def)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   895
    fix i show "emeasure M (f -` A i \<inter> space M) \<noteq> \<infinity>"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   896
      using f A(1,2) A(3)[of i] by (simp add: emeasure_distr subset_eq)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   897
  qed
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36649
diff changeset
   898
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36649
diff changeset
   899
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   900
lemma measurable_cong':
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   901
  assumes sets: "sets M = sets M'" "sets N = sets N'"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   902
  shows "measurable M N = measurable M' N'"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   903
  using sets[THEN sets_eq_imp_space_eq] sets by (simp add: measurable_def)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36649
diff changeset
   904
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   905
lemma pair_measure_distr:
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   906
  assumes f: "f \<in> measurable M S" and g: "g \<in> measurable N T"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   907
  assumes "sigma_finite_measure (distr M S f)" "sigma_finite_measure (distr N T g)"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   908
  shows "distr M S f \<Otimes>\<^isub>M distr N T g = distr (M \<Otimes>\<^isub>M N) (S \<Otimes>\<^isub>M T) (\<lambda>(x, y). (f x, g y))" (is "?P = ?D")
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   909
proof (rule measure_eqI)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   910
  show "sets ?P = sets ?D"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   911
    by simp
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   912
  interpret S: sigma_finite_measure "distr M S f" by fact
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   913
  interpret T: sigma_finite_measure "distr N T g" by fact
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   914
  interpret ST: pair_sigma_finite "distr M S f"  "distr N T g" ..
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   915
  interpret M: sigma_finite_measure M by (rule sigma_finite_measure_distr) fact+
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   916
  interpret N: sigma_finite_measure N by (rule sigma_finite_measure_distr) fact+
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   917
  interpret MN: pair_sigma_finite M N ..
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   918
  interpret SN: pair_sigma_finite "distr M S f" N ..
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   919
  have [simp]: 
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   920
    "\<And>f g. fst \<circ> (\<lambda>(x, y). (f x, g y)) = f \<circ> fst" "\<And>f g. snd \<circ> (\<lambda>(x, y). (f x, g y)) = g \<circ> snd"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   921
    by auto
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   922
  then have fg: "(\<lambda>(x, y). (f x, g y)) \<in> measurable (M \<Otimes>\<^isub>M N) (S \<Otimes>\<^isub>M T)"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   923
    using measurable_comp[OF measurable_fst f] measurable_comp[OF measurable_snd g]
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   924
    by (auto simp: measurable_pair_iff)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   925
  fix A assume A: "A \<in> sets ?P"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   926
  then have "emeasure ?P A = (\<integral>\<^isup>+x. emeasure (distr N T g) (Pair x -` A) \<partial>distr M S f)"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   927
    by (rule ST.emeasure_pair_measure_alt)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   928
  also have "\<dots> = (\<integral>\<^isup>+x. emeasure N (g -` (Pair x -` A) \<inter> space N) \<partial>distr M S f)"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   929
    using g A by (simp add: sets_Pair1 emeasure_distr)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   930
  also have "\<dots> = (\<integral>\<^isup>+x. emeasure N (g -` (Pair (f x) -` A) \<inter> space N) \<partial>M)"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   931
    using f g A ST.measurable_emeasure_Pair1[OF A]
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   932
    by (intro positive_integral_distr) (auto simp add: sets_Pair1 emeasure_distr)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   933
  also have "\<dots> = (\<integral>\<^isup>+x. emeasure N (Pair x -` ((\<lambda>(x, y). (f x, g y)) -` A \<inter> space (M \<Otimes>\<^isub>M N))) \<partial>M)"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   934
    by (intro positive_integral_cong arg_cong2[where f=emeasure]) (auto simp: space_pair_measure)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   935
  also have "\<dots> = emeasure (M \<Otimes>\<^isub>M N) ((\<lambda>(x, y). (f x, g y)) -` A \<inter> space (M \<Otimes>\<^isub>M N))"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   936
    using fg by (intro MN.emeasure_pair_measure_alt[symmetric] measurable_sets[OF _ A])
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   937
                (auto cong: measurable_cong')
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   938
  also have "\<dots> = emeasure ?D A"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   939
    using fg A by (subst emeasure_distr) auto
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   940
  finally show "emeasure ?P A = emeasure ?D A" .
45777
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 44890
diff changeset
   941
qed
39097
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   942
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   943
end