src/HOL/Probability/Binary_Product_Measure.thy
author hoelzl
Thu, 04 Dec 2014 17:05:58 +0100
changeset 59088 ff2bd4a14ddb
parent 59048 7dc8ac6f0895
child 59353 f0707dc3d9aa
permissions -rw-r--r--
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
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
58876
1888e3cb8048 modernized header;
wenzelm
parents: 58606
diff changeset
     5
section {*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
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 54863
diff changeset
     8
imports Nonnegative_Lebesgue_Integration
35833
7b7ae5aa396d Added product measure space
hoelzl
parents:
diff changeset
     9
begin
7b7ae5aa396d Added product measure space
hoelzl
parents:
diff changeset
    10
50104
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50003
diff changeset
    11
lemma Pair_vimage_times[simp]: "Pair x -` (A \<times> B) = (if x \<in> A then B else {})"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
    12
  by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
    13
50104
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50003
diff changeset
    14
lemma rev_Pair_vimage_times[simp]: "(\<lambda>x. (x, y)) -` (A \<times> B) = (if y \<in> B then A else {})"
40859
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
56994
8d5e5ec1cac3 fixed document generation for HOL-Probability
hoelzl
parents: 56993
diff changeset
    17
subsection "Binary products"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
    18
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
    19
definition pair_measure (infixr "\<Otimes>\<^sub>M" 80) where
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
    20
  "A \<Otimes>\<^sub>M B = measure_of (space A \<times> space B)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
    21
      {a \<times> b | a b. a \<in> sets A \<and> b \<in> sets B}
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
    22
      (\<lambda>X. \<integral>\<^sup>+x. (\<integral>\<^sup>+y. indicator X (x,y) \<partial>B) \<partial>A)"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
    23
49789
e0a4cb91a8a9 add induction rule for intersection-stable sigma-sets
hoelzl
parents: 49784
diff changeset
    24
lemma pair_measure_closed: "{a \<times> b | a b. a \<in> sets A \<and> b \<in> sets B} \<subseteq> Pow (space A \<times> space B)"
50244
de72bbe42190 qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents: 50104
diff changeset
    25
  using sets.space_closed[of A] sets.space_closed[of B] by auto
49789
e0a4cb91a8a9 add induction rule for intersection-stable sigma-sets
hoelzl
parents: 49784
diff changeset
    26
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
    27
lemma space_pair_measure:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
    28
  "space (A \<Otimes>\<^sub>M B) = space A \<times> space B"
49789
e0a4cb91a8a9 add induction rule for intersection-stable sigma-sets
hoelzl
parents: 49784
diff changeset
    29
  unfolding pair_measure_def using pair_measure_closed[of A B]
e0a4cb91a8a9 add induction rule for intersection-stable sigma-sets
hoelzl
parents: 49784
diff changeset
    30
  by (rule space_measure_of)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
    31
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
    32
lemma SIGMA_Collect_eq: "(SIGMA x:space M. {y\<in>space N. P x y}) = {x\<in>space (M \<Otimes>\<^sub>M N). P (fst x) (snd x)}"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
    33
  by (auto simp: space_pair_measure)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
    34
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
    35
lemma sets_pair_measure:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
    36
  "sets (A \<Otimes>\<^sub>M B) = sigma_sets (space A \<times> space B) {a \<times> b | a b. a \<in> sets A \<and> b \<in> sets B}"
49789
e0a4cb91a8a9 add induction rule for intersection-stable sigma-sets
hoelzl
parents: 49784
diff changeset
    37
  unfolding pair_measure_def using pair_measure_closed[of A B]
e0a4cb91a8a9 add induction rule for intersection-stable sigma-sets
hoelzl
parents: 49784
diff changeset
    38
  by (rule sets_measure_of)
41095
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41026
diff changeset
    39
58606
9c66f7c541fb add Giry monad
hoelzl
parents: 57447
diff changeset
    40
lemma sets_pair_in_sets:
9c66f7c541fb add Giry monad
hoelzl
parents: 57447
diff changeset
    41
  assumes N: "space A \<times> space B = space N"
9c66f7c541fb add Giry monad
hoelzl
parents: 57447
diff changeset
    42
  assumes "\<And>a b. a \<in> sets A \<Longrightarrow> b \<in> sets B \<Longrightarrow> a \<times> b \<in> sets N"
9c66f7c541fb add Giry monad
hoelzl
parents: 57447
diff changeset
    43
  shows "sets (A \<Otimes>\<^sub>M B) \<subseteq> sets N"
9c66f7c541fb add Giry monad
hoelzl
parents: 57447
diff changeset
    44
  using assms by (auto intro!: sets.sigma_sets_subset simp: sets_pair_measure N)
9c66f7c541fb add Giry monad
hoelzl
parents: 57447
diff changeset
    45
59048
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59000
diff changeset
    46
lemma sets_pair_measure_cong[measurable_cong, cong]:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
    47
  "sets M1 = sets M1' \<Longrightarrow> sets M2 = sets M2' \<Longrightarrow> sets (M1 \<Otimes>\<^sub>M M2) = sets (M1' \<Otimes>\<^sub>M M2')"
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
    48
  unfolding sets_pair_measure by (simp cong: sets_eq_imp_space_eq)
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
    49
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
    50
lemma pair_measureI[intro, simp, measurable]:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
    51
  "x \<in> sets A \<Longrightarrow> y \<in> sets B \<Longrightarrow> x \<times> y \<in> sets (A \<Otimes>\<^sub>M B)"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
    52
  by (auto simp: sets_pair_measure)
41095
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41026
diff changeset
    53
58606
9c66f7c541fb add Giry monad
hoelzl
parents: 57447
diff changeset
    54
lemma sets_Pair: "{x} \<in> sets M1 \<Longrightarrow> {y} \<in> sets M2 \<Longrightarrow> {(x, y)} \<in> sets (M1 \<Otimes>\<^sub>M M2)"
9c66f7c541fb add Giry monad
hoelzl
parents: 57447
diff changeset
    55
  using pair_measureI[of "{x}" M1 "{y}" M2] by simp
9c66f7c541fb add Giry monad
hoelzl
parents: 57447
diff changeset
    56
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
    57
lemma measurable_pair_measureI:
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
    58
  assumes 1: "f \<in> space M \<rightarrow> space M1 \<times> space M2"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
    59
  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"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
    60
  shows "f \<in> measurable M (M1 \<Otimes>\<^sub>M M2)"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
    61
  unfolding pair_measure_def using 1 2
50244
de72bbe42190 qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents: 50104
diff changeset
    62
  by (intro measurable_measure_of) (auto dest: sets.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
    63
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
    64
lemma measurable_split_replace[measurable (raw)]:
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
    65
  "(\<lambda>x. f x (fst (g x)) (snd (g x))) \<in> measurable M N \<Longrightarrow> (\<lambda>x. split (f x) (g x)) \<in> measurable M N"
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
    66
  unfolding split_beta' .
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
    67
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
    68
lemma measurable_Pair[measurable (raw)]:
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
    69
  assumes f: "f \<in> measurable M M1" and g: "g \<in> measurable M M2"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
    70
  shows "(\<lambda>x. (f x, g x)) \<in> measurable M (M1 \<Otimes>\<^sub>M M2)"
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
    71
proof (rule measurable_pair_measureI)
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
    72
  show "(\<lambda>x. (f x, g x)) \<in> space M \<rightarrow> space M1 \<times> space M2"
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
    73
    using f g by (auto simp: measurable_def)
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
    74
  fix A B assume *: "A \<in> sets M1" "B \<in> sets M2"
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
    75
  have "(\<lambda>x. (f x, g x)) -` (A \<times> B) \<inter> space M = (f -` A \<inter> space M) \<inter> (g -` B \<inter> space M)"
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
    76
    by auto
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
    77
  also have "\<dots> \<in> sets M"
50244
de72bbe42190 qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents: 50104
diff changeset
    78
    by (rule sets.Int) (auto intro!: measurable_sets * f g)
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
    79
  finally show "(\<lambda>x. (f x, g x)) -` (A \<times> B) \<inter> space M \<in> sets M" .
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
    80
qed
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
    81
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
    82
lemma measurable_Pair_compose_split[measurable_dest]:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
    83
  assumes f: "split f \<in> measurable (M1 \<Otimes>\<^sub>M M2) N"
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
    84
  assumes g: "g \<in> measurable M M1" and h: "h \<in> measurable M M2"
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
    85
  shows "(\<lambda>x. f (g x) (h x)) \<in> measurable M N"
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
    86
  using measurable_compose[OF measurable_Pair f, OF g h] by simp
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
    87
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
    88
lemma measurable_pair:
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
    89
  assumes "(fst \<circ> f) \<in> measurable M M1" "(snd \<circ> f) \<in> measurable M M2"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
    90
  shows "f \<in> measurable M (M1 \<Otimes>\<^sub>M M2)"
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
    91
  using measurable_Pair[OF assms] by simp
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
    92
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
    93
lemma measurable_fst[intro!, simp, measurable]: "fst \<in> measurable (M1 \<Otimes>\<^sub>M M2) M1"
50244
de72bbe42190 qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents: 50104
diff changeset
    94
  by (auto simp: fst_vimage_eq_Times space_pair_measure sets.sets_into_space times_Int_times
de72bbe42190 qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents: 50104
diff changeset
    95
    measurable_def)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
    96
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
    97
lemma measurable_snd[intro!, simp, measurable]: "snd \<in> measurable (M1 \<Otimes>\<^sub>M M2) M2"
50244
de72bbe42190 qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents: 50104
diff changeset
    98
  by (auto simp: snd_vimage_eq_Times space_pair_measure sets.sets_into_space times_Int_times
de72bbe42190 qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents: 50104
diff changeset
    99
    measurable_def)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   100
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   101
lemma 
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
   102
  assumes f[measurable]: "f \<in> measurable M (N \<Otimes>\<^sub>M P)" 
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   103
  shows measurable_fst': "(\<lambda>x. fst (f x)) \<in> measurable M N"
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   104
    and measurable_snd': "(\<lambda>x. snd (f x)) \<in> measurable M P"
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   105
  by simp_all
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   106
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   107
lemma
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   108
  assumes f[measurable]: "f \<in> measurable M N"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
   109
  shows measurable_fst'': "(\<lambda>x. f (fst x)) \<in> measurable (M \<Otimes>\<^sub>M P) N"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
   110
    and measurable_snd'': "(\<lambda>x. f (snd x)) \<in> measurable (P \<Otimes>\<^sub>M M) N"
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   111
  by simp_all
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   112
58606
9c66f7c541fb add Giry monad
hoelzl
parents: 57447
diff changeset
   113
lemma sets_pair_eq_sets_fst_snd:
9c66f7c541fb add Giry monad
hoelzl
parents: 57447
diff changeset
   114
  "sets (A \<Otimes>\<^sub>M B) = sets (Sup_sigma {vimage_algebra (space A \<times> space B) fst A, vimage_algebra (space A \<times> space B) snd B})"
9c66f7c541fb add Giry monad
hoelzl
parents: 57447
diff changeset
   115
    (is "?P = sets (Sup_sigma {?fst, ?snd})")
9c66f7c541fb add Giry monad
hoelzl
parents: 57447
diff changeset
   116
proof -
9c66f7c541fb add Giry monad
hoelzl
parents: 57447
diff changeset
   117
  { fix a b assume ab: "a \<in> sets A" "b \<in> sets B"
9c66f7c541fb add Giry monad
hoelzl
parents: 57447
diff changeset
   118
    then have "a \<times> b = (fst -` a \<inter> (space A \<times> space B)) \<inter> (snd -` b \<inter> (space A \<times> space B))"
9c66f7c541fb add Giry monad
hoelzl
parents: 57447
diff changeset
   119
      by (auto dest: sets.sets_into_space)
9c66f7c541fb add Giry monad
hoelzl
parents: 57447
diff changeset
   120
    also have "\<dots> \<in> sets (Sup_sigma {?fst, ?snd})"
9c66f7c541fb add Giry monad
hoelzl
parents: 57447
diff changeset
   121
      using ab by (auto intro: in_Sup_sigma in_vimage_algebra)
9c66f7c541fb add Giry monad
hoelzl
parents: 57447
diff changeset
   122
    finally have "a \<times> b \<in> sets (Sup_sigma {?fst, ?snd})" . }
9c66f7c541fb add Giry monad
hoelzl
parents: 57447
diff changeset
   123
  moreover have "sets ?fst \<subseteq> sets (A \<Otimes>\<^sub>M B)"
9c66f7c541fb add Giry monad
hoelzl
parents: 57447
diff changeset
   124
    by (rule sets_image_in_sets) (auto simp: space_pair_measure[symmetric])
9c66f7c541fb add Giry monad
hoelzl
parents: 57447
diff changeset
   125
  moreover have "sets ?snd \<subseteq> sets (A \<Otimes>\<^sub>M B)"  
9c66f7c541fb add Giry monad
hoelzl
parents: 57447
diff changeset
   126
    by (rule sets_image_in_sets) (auto simp: space_pair_measure)
9c66f7c541fb add Giry monad
hoelzl
parents: 57447
diff changeset
   127
  ultimately show ?thesis
9c66f7c541fb add Giry monad
hoelzl
parents: 57447
diff changeset
   128
    by (intro antisym[of "sets A" for A] sets_Sup_in_sets sets_pair_in_sets )
9c66f7c541fb add Giry monad
hoelzl
parents: 57447
diff changeset
   129
       (auto simp add: space_Sup_sigma space_pair_measure)
9c66f7c541fb add Giry monad
hoelzl
parents: 57447
diff changeset
   130
qed
9c66f7c541fb add Giry monad
hoelzl
parents: 57447
diff changeset
   131
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   132
lemma measurable_pair_iff:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
   133
  "f \<in> measurable M (M1 \<Otimes>\<^sub>M M2) \<longleftrightarrow> (fst \<circ> f) \<in> measurable M M1 \<and> (snd \<circ> f) \<in> measurable M M2"
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   134
  by (auto intro: measurable_pair[of f M M1 M2]) 
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   135
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   136
lemma measurable_split_conv:
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   137
  "(\<lambda>(x, y). f x y) \<in> measurable A B \<longleftrightarrow> (\<lambda>x. f (fst x) (snd x)) \<in> measurable A B"
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   138
  by (intro arg_cong2[where f="op \<in>"]) auto
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   139
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
   140
lemma measurable_pair_swap': "(\<lambda>(x,y). (y, x)) \<in> measurable (M1 \<Otimes>\<^sub>M M2) (M2 \<Otimes>\<^sub>M M1)"
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   141
  by (auto intro!: measurable_Pair simp: measurable_split_conv)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   142
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   143
lemma measurable_pair_swap:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
   144
  assumes f: "f \<in> measurable (M1 \<Otimes>\<^sub>M M2) M" shows "(\<lambda>(x,y). f (y, x)) \<in> measurable (M2 \<Otimes>\<^sub>M M1) M"
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   145
  using measurable_comp[OF measurable_Pair f] by (auto simp: measurable_split_conv comp_def)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   146
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   147
lemma measurable_pair_swap_iff:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
   148
  "f \<in> measurable (M2 \<Otimes>\<^sub>M M1) M \<longleftrightarrow> (\<lambda>(x,y). f (y,x)) \<in> measurable (M1 \<Otimes>\<^sub>M M2) M"
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   149
  by (auto dest: measurable_pair_swap)
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   150
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
   151
lemma measurable_Pair1': "x \<in> space M1 \<Longrightarrow> Pair x \<in> measurable M2 (M1 \<Otimes>\<^sub>M M2)"
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   152
  by simp
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   153
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   154
lemma sets_Pair1[measurable (raw)]:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
   155
  assumes A: "A \<in> sets (M1 \<Otimes>\<^sub>M M2)" shows "Pair x -` A \<in> sets M2"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   156
proof -
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   157
  have "Pair x -` A = (if x \<in> space M1 then Pair x -` A \<inter> space M2 else {})"
50244
de72bbe42190 qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents: 50104
diff changeset
   158
    using A[THEN sets.sets_into_space] by (auto simp: space_pair_measure)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   159
  also have "\<dots> \<in> sets M2"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   160
    using A by (auto simp add: measurable_Pair1' intro!: measurable_sets split: split_if_asm)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   161
  finally show ?thesis .
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   162
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   163
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
   164
lemma measurable_Pair2': "y \<in> space M2 \<Longrightarrow> (\<lambda>x. (x, y)) \<in> measurable M1 (M1 \<Otimes>\<^sub>M M2)"
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   165
  by (auto intro!: measurable_Pair)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   166
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
   167
lemma sets_Pair2: assumes A: "A \<in> sets (M1 \<Otimes>\<^sub>M M2)" shows "(\<lambda>x. (x, y)) -` A \<in> sets M1"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   168
proof -
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   169
  have "(\<lambda>x. (x, y)) -` A = (if y \<in> space M2 then (\<lambda>x. (x, y)) -` A \<inter> space M1 else {})"
50244
de72bbe42190 qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents: 50104
diff changeset
   170
    using A[THEN sets.sets_into_space] by (auto simp: space_pair_measure)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   171
  also have "\<dots> \<in> sets M1"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   172
    using A by (auto simp add: measurable_Pair2' intro!: measurable_sets split: split_if_asm)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   173
  finally show ?thesis .
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   174
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   175
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   176
lemma measurable_Pair2:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
   177
  assumes f: "f \<in> measurable (M1 \<Otimes>\<^sub>M M2) M" and x: "x \<in> space M1"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   178
  shows "(\<lambda>y. f (x, y)) \<in> measurable M2 M"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   179
  using measurable_comp[OF measurable_Pair1' f, OF x]
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   180
  by (simp add: comp_def)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   181
  
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   182
lemma measurable_Pair1:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
   183
  assumes f: "f \<in> measurable (M1 \<Otimes>\<^sub>M M2) M" and y: "y \<in> space M2"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   184
  shows "(\<lambda>x. f (x, y)) \<in> measurable M1 M"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   185
  using measurable_comp[OF measurable_Pair2' f, OF y]
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   186
  by (simp add: comp_def)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   187
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   188
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
   189
  unfolding Int_stable_def
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   190
  by safe (auto simp add: times_Int_times)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   191
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   192
lemma disjoint_family_vimageI: "disjoint_family F \<Longrightarrow> disjoint_family (\<lambda>i. f -` F i)"
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   193
  by (auto simp: disjoint_family_on_def)
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   194
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   195
lemma (in finite_measure) finite_measure_cut_measurable:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
   196
  assumes [measurable]: "Q \<in> sets (N \<Otimes>\<^sub>M M)"
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   197
  shows "(\<lambda>x. emeasure M (Pair x -` Q)) \<in> borel_measurable N"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   198
    (is "?s Q \<in> _")
49789
e0a4cb91a8a9 add induction rule for intersection-stable sigma-sets
hoelzl
parents: 49784
diff changeset
   199
  using Int_stable_pair_measure_generator pair_measure_closed assms
e0a4cb91a8a9 add induction rule for intersection-stable sigma-sets
hoelzl
parents: 49784
diff changeset
   200
  unfolding sets_pair_measure
e0a4cb91a8a9 add induction rule for intersection-stable sigma-sets
hoelzl
parents: 49784
diff changeset
   201
proof (induct rule: sigma_sets_induct_disjoint)
e0a4cb91a8a9 add induction rule for intersection-stable sigma-sets
hoelzl
parents: 49784
diff changeset
   202
  case (compl A)
50244
de72bbe42190 qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents: 50104
diff changeset
   203
  with sets.sets_into_space have "\<And>x. emeasure M (Pair x -` ((space N \<times> space M) - A)) =
49789
e0a4cb91a8a9 add induction rule for intersection-stable sigma-sets
hoelzl
parents: 49784
diff changeset
   204
      (if x \<in> space N then emeasure M (space M) - ?s A x else 0)"
e0a4cb91a8a9 add induction rule for intersection-stable sigma-sets
hoelzl
parents: 49784
diff changeset
   205
    unfolding sets_pair_measure[symmetric]
e0a4cb91a8a9 add induction rule for intersection-stable sigma-sets
hoelzl
parents: 49784
diff changeset
   206
    by (auto intro!: emeasure_compl simp: vimage_Diff sets_Pair1)
50244
de72bbe42190 qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents: 50104
diff changeset
   207
  with compl sets.top show ?case
49789
e0a4cb91a8a9 add induction rule for intersection-stable sigma-sets
hoelzl
parents: 49784
diff changeset
   208
    by (auto intro!: measurable_If simp: space_pair_measure)
e0a4cb91a8a9 add induction rule for intersection-stable sigma-sets
hoelzl
parents: 49784
diff changeset
   209
next
e0a4cb91a8a9 add induction rule for intersection-stable sigma-sets
hoelzl
parents: 49784
diff changeset
   210
  case (union F)
53374
a14d2a854c02 tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents: 53015
diff changeset
   211
  then have "\<And>x. emeasure M (Pair x -` (\<Union>i. F i)) = (\<Sum>i. ?s (F i) x)"
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   212
    by (simp add: suminf_emeasure disjoint_family_vimageI subset_eq vimage_UN sets_pair_measure[symmetric])
53374
a14d2a854c02 tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents: 53015
diff changeset
   213
  with union show ?case
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   214
    unfolding sets_pair_measure[symmetric] by simp
49789
e0a4cb91a8a9 add induction rule for intersection-stable sigma-sets
hoelzl
parents: 49784
diff changeset
   215
qed (auto simp add: if_distrib Int_def[symmetric] intro!: measurable_If)
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   216
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   217
lemma (in sigma_finite_measure) measurable_emeasure_Pair:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
   218
  assumes Q: "Q \<in> sets (N \<Otimes>\<^sub>M M)" shows "(\<lambda>x. emeasure M (Pair x -` Q)) \<in> borel_measurable N" (is "?s Q \<in> _")
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   219
proof -
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   220
  from sigma_finite_disjoint guess F . note F = this
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   221
  then have F_sets: "\<And>i. F i \<in> sets M" by auto
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   222
  let ?C = "\<lambda>x i. F i \<inter> Pair x -` Q"
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   223
  { fix i
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   224
    have [simp]: "space N \<times> F i \<inter> space N \<times> space M = space N \<times> F i"
50244
de72bbe42190 qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents: 50104
diff changeset
   225
      using F sets.sets_into_space by auto
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   226
    let ?R = "density M (indicator (F i))"
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   227
    have "finite_measure ?R"
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   228
      using F by (intro finite_measureI) (auto simp: emeasure_restricted subset_eq)
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   229
    then have "(\<lambda>x. emeasure ?R (Pair x -` (space N \<times> space ?R \<inter> Q))) \<in> borel_measurable N"
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   230
     by (rule finite_measure.finite_measure_cut_measurable) (auto intro: Q)
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   231
    moreover have "\<And>x. emeasure ?R (Pair x -` (space N \<times> space ?R \<inter> Q))
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   232
        = emeasure M (F i \<inter> Pair x -` (space N \<times> space ?R \<inter> Q))"
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   233
      using Q F_sets by (intro emeasure_restricted) (auto intro: sets_Pair1)
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   234
    moreover have "\<And>x. F i \<inter> Pair x -` (space N \<times> space ?R \<inter> Q) = ?C x i"
50244
de72bbe42190 qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents: 50104
diff changeset
   235
      using sets.sets_into_space[OF Q] by (auto simp: space_pair_measure)
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   236
    ultimately have "(\<lambda>x. emeasure M (?C x i)) \<in> borel_measurable N"
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   237
      by simp }
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   238
  moreover
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   239
  { fix x
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   240
    have "(\<Sum>i. emeasure M (?C x i)) = emeasure M (\<Union>i. ?C x i)"
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   241
    proof (intro suminf_emeasure)
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   242
      show "range (?C x) \<subseteq> sets M"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
   243
        using F `Q \<in> sets (N \<Otimes>\<^sub>M M)` by (auto intro!: sets_Pair1)
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   244
      have "disjoint_family F" using F by auto
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   245
      show "disjoint_family (?C x)"
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   246
        by (rule disjoint_family_on_bisimulation[OF `disjoint_family F`]) auto
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   247
    qed
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   248
    also have "(\<Union>i. ?C x i) = Pair x -` Q"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
   249
      using F sets.sets_into_space[OF `Q \<in> sets (N \<Otimes>\<^sub>M M)`]
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   250
      by (auto simp: space_pair_measure)
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   251
    finally have "emeasure M (Pair x -` Q) = (\<Sum>i. emeasure M (?C x i))"
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   252
      by simp }
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
   253
  ultimately show ?thesis using `Q \<in> sets (N \<Otimes>\<^sub>M M)` F_sets
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   254
    by auto
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   255
qed
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   256
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   257
lemma (in sigma_finite_measure) measurable_emeasure[measurable (raw)]:
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   258
  assumes space: "\<And>x. x \<in> space N \<Longrightarrow> A x \<subseteq> space M"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
   259
  assumes A: "{x\<in>space (N \<Otimes>\<^sub>M M). snd x \<in> A (fst x)} \<in> sets (N \<Otimes>\<^sub>M M)"
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   260
  shows "(\<lambda>x. emeasure M (A x)) \<in> borel_measurable N"
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   261
proof -
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
   262
  from space have "\<And>x. x \<in> space N \<Longrightarrow> Pair x -` {x \<in> space (N \<Otimes>\<^sub>M M). snd x \<in> A (fst x)} = A x"
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   263
    by (auto simp: space_pair_measure)
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   264
  with measurable_emeasure_Pair[OF A] show ?thesis
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   265
    by (auto cong: measurable_cong)
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   266
qed
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   267
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   268
lemma (in sigma_finite_measure) emeasure_pair_measure:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
   269
  assumes "X \<in> sets (N \<Otimes>\<^sub>M M)"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
   270
  shows "emeasure (N \<Otimes>\<^sub>M M) X = (\<integral>\<^sup>+ x. \<integral>\<^sup>+ y. indicator X (x, y) \<partial>M \<partial>N)" (is "_ = ?\<mu> X")
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   271
proof (rule emeasure_measure_of[OF pair_measure_def])
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
   272
  show "positive (sets (N \<Otimes>\<^sub>M M)) ?\<mu>"
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
   273
    by (auto simp: positive_def nn_integral_nonneg)
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   274
  have eq[simp]: "\<And>A x y. indicator A (x, y) = indicator (Pair x -` A) y"
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   275
    by (auto simp: indicator_def)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
   276
  show "countably_additive (sets (N \<Otimes>\<^sub>M M)) ?\<mu>"
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   277
  proof (rule countably_additiveI)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
   278
    fix F :: "nat \<Rightarrow> ('b \<times> 'a) set" assume F: "range F \<subseteq> sets (N \<Otimes>\<^sub>M M)" "disjoint_family F"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
   279
    from F have *: "\<And>i. F i \<in> sets (N \<Otimes>\<^sub>M M)" "(\<Union>i. F i) \<in> sets (N \<Otimes>\<^sub>M M)" by auto
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   280
    moreover from F have "\<And>i. (\<lambda>x. emeasure M (Pair x -` F i)) \<in> borel_measurable N"
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   281
      by (intro measurable_emeasure_Pair) auto
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   282
    moreover have "\<And>x. disjoint_family (\<lambda>i. Pair x -` F i)"
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   283
      by (intro disjoint_family_on_bisimulation[OF F(2)]) auto
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   284
    moreover have "\<And>x. range (\<lambda>i. Pair x -` F i) \<subseteq> sets M"
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   285
      using F by (auto simp: sets_Pair1)
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   286
    ultimately show "(\<Sum>n. ?\<mu> (F n)) = ?\<mu> (\<Union>i. F i)"
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
   287
      by (auto simp add: vimage_UN nn_integral_suminf[symmetric] suminf_emeasure subset_eq emeasure_nonneg sets_Pair1
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
   288
               intro!: nn_integral_cong nn_integral_indicator[symmetric])
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   289
  qed
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   290
  show "{a \<times> b |a b. a \<in> sets N \<and> b \<in> sets M} \<subseteq> Pow (space N \<times> space M)"
50244
de72bbe42190 qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents: 50104
diff changeset
   291
    using sets.space_closed[of N] sets.space_closed[of M] by auto
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   292
qed fact
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   293
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   294
lemma (in sigma_finite_measure) emeasure_pair_measure_alt:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
   295
  assumes X: "X \<in> sets (N \<Otimes>\<^sub>M M)"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
   296
  shows "emeasure (N  \<Otimes>\<^sub>M M) X = (\<integral>\<^sup>+x. emeasure M (Pair x -` X) \<partial>N)"
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   297
proof -
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   298
  have [simp]: "\<And>x y. indicator X (x, y) = indicator (Pair x -` X) y"
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   299
    by (auto simp: indicator_def)
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   300
  show ?thesis
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
   301
    using X by (auto intro!: nn_integral_cong simp: emeasure_pair_measure sets_Pair1)
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   302
qed
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   303
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   304
lemma (in sigma_finite_measure) emeasure_pair_measure_Times:
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   305
  assumes A: "A \<in> sets N" and B: "B \<in> sets M"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
   306
  shows "emeasure (N \<Otimes>\<^sub>M M) (A \<times> B) = emeasure N A * emeasure M B"
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   307
proof -
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
   308
  have "emeasure (N \<Otimes>\<^sub>M M) (A \<times> B) = (\<integral>\<^sup>+x. emeasure M B * indicator A x \<partial>N)"
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
   309
    using A B by (auto intro!: nn_integral_cong simp: emeasure_pair_measure_alt)
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   310
  also have "\<dots> = emeasure M B * emeasure N A"
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
   311
    using A by (simp add: emeasure_nonneg nn_integral_cmult_indicator)
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   312
  finally show ?thesis
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   313
    by (simp add: ac_simps)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   314
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   315
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   316
subsection {* Binary products of $\sigma$-finite emeasure spaces *}
40859
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
locale pair_sigma_finite = M1: sigma_finite_measure M1 + M2: sigma_finite_measure M2
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   319
  for M1 :: "'a measure" and M2 :: "'b measure"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   320
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   321
lemma (in pair_sigma_finite) measurable_emeasure_Pair1:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
   322
  "Q \<in> sets (M1 \<Otimes>\<^sub>M M2) \<Longrightarrow> (\<lambda>x. emeasure M2 (Pair x -` Q)) \<in> borel_measurable M1"
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   323
  using M2.measurable_emeasure_Pair .
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   324
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   325
lemma (in pair_sigma_finite) measurable_emeasure_Pair2:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
   326
  assumes Q: "Q \<in> sets (M1 \<Otimes>\<^sub>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
   327
proof -
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
   328
  have "(\<lambda>(x, y). (y, x)) -` Q \<inter> space (M2 \<Otimes>\<^sub>M M1) \<in> sets (M2 \<Otimes>\<^sub>M M1)"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   329
    using Q measurable_pair_swap' by (auto intro: measurable_sets)
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   330
  note M1.measurable_emeasure_Pair[OF this]
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
   331
  moreover have "\<And>y. Pair y -` ((\<lambda>(x, y). (y, x)) -` Q \<inter> space (M2 \<Otimes>\<^sub>M M1)) = (\<lambda>x. (x, y)) -` Q"
50244
de72bbe42190 qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents: 50104
diff changeset
   332
    using Q[THEN sets.sets_into_space] by (auto simp: space_pair_measure)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   333
  ultimately show ?thesis by simp
39088
ca17017c10e6 Measurable on product space is equiv. to measurable components
hoelzl
parents: 39082
diff changeset
   334
qed
ca17017c10e6 Measurable on product space is equiv. to measurable components
hoelzl
parents: 39082
diff changeset
   335
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
   336
lemma (in pair_sigma_finite) sigma_finite_up_in_pair_measure_generator:
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   337
  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
   338
  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>
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
   339
    (\<forall>i. emeasure (M1 \<Otimes>\<^sub>M M2) (F i) \<noteq> \<infinity>)"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   340
proof -
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   341
  from M1.sigma_finite_incseq guess F1 . note F1 = this
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   342
  from M2.sigma_finite_incseq guess F2 . note F2 = this
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   343
  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
   344
  let ?F = "\<lambda>i. F1 i \<times> F2 i"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   345
  show ?thesis
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   346
  proof (intro exI[of _ ?F] conjI allI)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   347
    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
   348
  next
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   349
    have "space M1 \<times> space M2 \<subseteq> (\<Union>i. ?F i)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   350
    proof (intro subsetI)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   351
      fix x assume "x \<in> space M1 \<times> space M2"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   352
      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
   353
        by (auto simp: space)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   354
      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
   355
        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
   356
        by (force split: split_max)+
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   357
      then have "(fst x, snd x) \<in> F1 (max i j) \<times> F2 (max i j)"
54863
82acc20ded73 prefer more canonical names for lemmas on min/max
haftmann
parents: 53374
diff changeset
   358
        by (intro SigmaI) (auto simp add: max.commute)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   359
      then show "x \<in> (\<Union>i. ?F i)" by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   360
    qed
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   361
    then show "(\<Union>i. ?F i) = space M1 \<times> space M2"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   362
      using space by (auto simp: space)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   363
  next
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   364
    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
   365
      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
   366
  next
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   367
    fix i
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   368
    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
   369
    with F1 F2 emeasure_nonneg[of M1 "F1 i"] emeasure_nonneg[of M2 "F2 i"]
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
   370
    show "emeasure (M1 \<Otimes>\<^sub>M M2) (F1 i \<times> F2 i) \<noteq> \<infinity>"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   371
      by (auto simp add: emeasure_pair_measure_Times)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   372
  qed
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   373
qed
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   374
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
   375
sublocale pair_sigma_finite \<subseteq> P: sigma_finite_measure "M1 \<Otimes>\<^sub>M M2"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   376
proof
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57235
diff changeset
   377
  from M1.sigma_finite_countable guess F1 ..
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57235
diff changeset
   378
  moreover from M2.sigma_finite_countable guess F2 ..
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57235
diff changeset
   379
  ultimately show
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57235
diff changeset
   380
    "\<exists>A. countable A \<and> A \<subseteq> sets (M1 \<Otimes>\<^sub>M M2) \<and> \<Union>A = space (M1 \<Otimes>\<^sub>M M2) \<and> (\<forall>a\<in>A. emeasure (M1 \<Otimes>\<^sub>M M2) a \<noteq> \<infinity>)"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57235
diff changeset
   381
    by (intro exI[of _ "(\<lambda>(a, b). a \<times> b) ` (F1 \<times> F2)"] conjI)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57235
diff changeset
   382
       (auto simp: M2.emeasure_pair_measure_Times space_pair_measure set_eq_iff subset_eq
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57235
diff changeset
   383
             dest: sets.sets_into_space)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   384
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   385
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   386
lemma sigma_finite_pair_measure:
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   387
  assumes A: "sigma_finite_measure A" and B: "sigma_finite_measure B"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
   388
  shows "sigma_finite_measure (A \<Otimes>\<^sub>M B)"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   389
proof -
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   390
  interpret A: sigma_finite_measure A by fact
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   391
  interpret B: sigma_finite_measure B by fact
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   392
  interpret AB: pair_sigma_finite A  B ..
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   393
  show ?thesis ..
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   394
qed
39088
ca17017c10e6 Measurable on product space is equiv. to measurable components
hoelzl
parents: 39082
diff changeset
   395
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   396
lemma sets_pair_swap:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
   397
  assumes "A \<in> sets (M1 \<Otimes>\<^sub>M M2)"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
   398
  shows "(\<lambda>(x, y). (y, x)) -` A \<inter> space (M2 \<Otimes>\<^sub>M M1) \<in> sets (M2 \<Otimes>\<^sub>M M1)"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   399
  using measurable_pair_swap' assms by (rule measurable_sets)
41661
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41659
diff changeset
   400
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   401
lemma (in pair_sigma_finite) distr_pair_swap:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
   402
  "M1 \<Otimes>\<^sub>M M2 = distr (M2 \<Otimes>\<^sub>M M1) (M1 \<Otimes>\<^sub>M M2) (\<lambda>(x, y). (y, x))" (is "?P = ?D")
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   403
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
   404
  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
   405
  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
   406
  show ?thesis
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   407
  proof (rule measure_eqI_generator_eq[OF Int_stable_pair_measure_generator[of M1 M2]])
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   408
    show "?E \<subseteq> Pow (space ?P)"
50244
de72bbe42190 qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents: 50104
diff changeset
   409
      using sets.space_closed[of M1] sets.space_closed[of M2] by (auto simp: space_pair_measure)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   410
    show "sets ?P = sigma_sets (space ?P) ?E"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   411
      by (simp add: sets_pair_measure space_pair_measure)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   412
    then show "sets ?D = sigma_sets (space ?P) ?E"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   413
      by simp
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   414
  next
49784
5e5b2da42a69 remove incseq assumption from measure_eqI_generator_eq
hoelzl
parents: 49776
diff changeset
   415
    show "range F \<subseteq> ?E" "(\<Union>i. F i) = space ?P" "\<And>i. emeasure ?P (F i) \<noteq> \<infinity>"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   416
      using F by (auto simp: space_pair_measure)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   417
  next
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   418
    fix X assume "X \<in> ?E"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   419
    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
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
   420
    have "(\<lambda>(y, x). (x, y)) -` X \<inter> space (M2 \<Otimes>\<^sub>M M1) = B \<times> A"
50244
de72bbe42190 qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents: 50104
diff changeset
   421
      using sets.sets_into_space[OF A] sets.sets_into_space[OF B] by (auto simp: space_pair_measure)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
   422
    with A B show "emeasure (M1 \<Otimes>\<^sub>M M2) X = emeasure ?D X"
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   423
      by (simp add: M2.emeasure_pair_measure_Times M1.emeasure_pair_measure_Times emeasure_distr
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   424
                    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
   425
  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
   426
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
   427
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   428
lemma (in pair_sigma_finite) emeasure_pair_measure_alt2:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
   429
  assumes A: "A \<in> sets (M1 \<Otimes>\<^sub>M M2)"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
   430
  shows "emeasure (M1 \<Otimes>\<^sub>M M2) A = (\<integral>\<^sup>+y. emeasure M1 ((\<lambda>x. (x, y)) -` A) \<partial>M2)"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   431
    (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
   432
proof -
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
   433
  have [simp]: "\<And>y. (Pair y -` ((\<lambda>(x, y). (y, x)) -` A \<inter> space (M2 \<Otimes>\<^sub>M M1))) = (\<lambda>x. (x, y)) -` A"
50244
de72bbe42190 qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents: 50104
diff changeset
   434
    using sets.sets_into_space[OF A] by (auto simp: space_pair_measure)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   435
  show ?thesis using A
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   436
    by (subst distr_pair_swap)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   437
       (simp_all del: vimage_Int add: measurable_sets[OF measurable_pair_swap']
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   438
                 M1.emeasure_pair_measure_alt emeasure_distr[OF measurable_pair_swap' A])
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   439
qed
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   440
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   441
lemma (in pair_sigma_finite) AE_pair:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
   442
  assumes "AE x in (M1 \<Otimes>\<^sub>M M2). Q x"
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   443
  shows "AE x in M1. (AE y in M2. Q (x, y))"
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   444
proof -
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
   445
  obtain N where N: "N \<in> sets (M1 \<Otimes>\<^sub>M M2)" "emeasure (M1 \<Otimes>\<^sub>M M2) N = 0" "{x\<in>space (M1 \<Otimes>\<^sub>M M2). \<not> Q x} \<subseteq> N"
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   446
    using assms unfolding eventually_ae_filter by auto
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   447
  show ?thesis
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   448
  proof (rule AE_I)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
   449
    from N measurable_emeasure_Pair1[OF `N \<in> sets (M1 \<Otimes>\<^sub>M M2)`]
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   450
    show "emeasure M1 {x\<in>space M1. emeasure M2 (Pair x -` N) \<noteq> 0} = 0"
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
   451
      by (auto simp: M2.emeasure_pair_measure_alt nn_integral_0_iff emeasure_nonneg)
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   452
    show "{x \<in> space M1. emeasure M2 (Pair x -` N) \<noteq> 0} \<in> sets M1"
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   453
      by (intro borel_measurable_ereal_neq_const measurable_emeasure_Pair1 N)
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   454
    { fix x assume "x \<in> space M1" "emeasure M2 (Pair x -` N) = 0"
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   455
      have "AE y in M2. Q (x, y)"
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   456
      proof (rule AE_I)
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   457
        show "emeasure M2 (Pair x -` N) = 0" by fact
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   458
        show "Pair x -` N \<in> sets M2" using N(1) by (rule sets_Pair1)
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   459
        show "{y \<in> space M2. \<not> Q (x, y)} \<subseteq> Pair x -` N"
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   460
          using N `x \<in> space M1` unfolding space_pair_measure by auto
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   461
      qed }
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   462
    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}"
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   463
      by auto
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   464
  qed
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   465
qed
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   466
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   467
lemma (in pair_sigma_finite) AE_pair_measure:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
   468
  assumes "{x\<in>space (M1 \<Otimes>\<^sub>M M2). P x} \<in> sets (M1 \<Otimes>\<^sub>M M2)"
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   469
  assumes ae: "AE x in M1. AE y in M2. P (x, y)"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
   470
  shows "AE x in M1 \<Otimes>\<^sub>M M2. P x"
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   471
proof (subst AE_iff_measurable[OF _ refl])
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
   472
  show "{x\<in>space (M1 \<Otimes>\<^sub>M M2). \<not> P x} \<in> sets (M1 \<Otimes>\<^sub>M M2)"
50244
de72bbe42190 qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents: 50104
diff changeset
   473
    by (rule sets.sets_Collect) fact
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
   474
  then have "emeasure (M1 \<Otimes>\<^sub>M M2) {x \<in> space (M1 \<Otimes>\<^sub>M M2). \<not> P x} =
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
   475
      (\<integral>\<^sup>+ x. \<integral>\<^sup>+ y. indicator {x \<in> space (M1 \<Otimes>\<^sub>M M2). \<not> P x} (x, y) \<partial>M2 \<partial>M1)"
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   476
    by (simp add: M2.emeasure_pair_measure)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
   477
  also have "\<dots> = (\<integral>\<^sup>+ x. \<integral>\<^sup>+ y. 0 \<partial>M2 \<partial>M1)"
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   478
    using ae
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
   479
    apply (safe intro!: nn_integral_cong_AE)
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   480
    apply (intro AE_I2)
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
   481
    apply (safe intro!: nn_integral_cong_AE)
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   482
    apply auto
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   483
    done
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
   484
  finally show "emeasure (M1 \<Otimes>\<^sub>M M2) {x \<in> space (M1 \<Otimes>\<^sub>M M2). \<not> P x} = 0" by simp
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   485
qed
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   486
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   487
lemma (in pair_sigma_finite) AE_pair_iff:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
   488
  "{x\<in>space (M1 \<Otimes>\<^sub>M M2). P (fst x) (snd x)} \<in> sets (M1 \<Otimes>\<^sub>M M2) \<Longrightarrow>
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
   489
    (AE x in M1. AE y in M2. P x y) \<longleftrightarrow> (AE x in (M1 \<Otimes>\<^sub>M M2). P (fst x) (snd x))"
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   490
  using AE_pair[of "\<lambda>x. P (fst x) (snd x)"] AE_pair_measure[of "\<lambda>x. P (fst x) (snd x)"] by auto
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   491
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   492
lemma (in pair_sigma_finite) AE_commute:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
   493
  assumes P: "{x\<in>space (M1 \<Otimes>\<^sub>M M2). P (fst x) (snd x)} \<in> sets (M1 \<Otimes>\<^sub>M M2)"
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   494
  shows "(AE x in M1. AE y in M2. P x y) \<longleftrightarrow> (AE y in M2. AE x in M1. P x y)"
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   495
proof -
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   496
  interpret Q: pair_sigma_finite M2 M1 ..
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   497
  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"
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   498
    by auto
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
   499
  have "{x \<in> space (M2 \<Otimes>\<^sub>M M1). P (snd x) (fst x)} =
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
   500
    (\<lambda>(x, y). (y, x)) -` {x \<in> space (M1 \<Otimes>\<^sub>M M2). P (fst x) (snd x)} \<inter> space (M2 \<Otimes>\<^sub>M M1)"
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   501
    by (auto simp: space_pair_measure)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
   502
  also have "\<dots> \<in> sets (M2 \<Otimes>\<^sub>M M1)"
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   503
    by (intro sets_pair_swap P)
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   504
  finally show ?thesis
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   505
    apply (subst AE_pair_iff[OF P])
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   506
    apply (subst distr_pair_swap)
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   507
    apply (subst AE_distr_iff[OF measurable_pair_swap' P])
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   508
    apply (subst Q.AE_pair_iff)
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   509
    apply simp_all
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   510
    done
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   511
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   512
56994
8d5e5ec1cac3 fixed document generation for HOL-Probability
hoelzl
parents: 56993
diff changeset
   513
subsection "Fubinis theorem"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   514
49800
a6678da5692c induction prove for positive_integral_fst
hoelzl
parents: 49789
diff changeset
   515
lemma measurable_compose_Pair1:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
   516
  "x \<in> space M1 \<Longrightarrow> g \<in> measurable (M1 \<Otimes>\<^sub>M M2) L \<Longrightarrow> (\<lambda>y. g (x, y)) \<in> measurable M2 L"
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   517
  by simp
49800
a6678da5692c induction prove for positive_integral_fst
hoelzl
parents: 49789
diff changeset
   518
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
   519
lemma (in sigma_finite_measure) borel_measurable_nn_integral_fst':
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
   520
  assumes f: "f \<in> borel_measurable (M1 \<Otimes>\<^sub>M M)" "\<And>x. 0 \<le> f x"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
   521
  shows "(\<lambda>x. \<integral>\<^sup>+ y. f (x, y) \<partial>M) \<in> borel_measurable M1"
49800
a6678da5692c induction prove for positive_integral_fst
hoelzl
parents: 49789
diff changeset
   522
using f proof induct
a6678da5692c induction prove for positive_integral_fst
hoelzl
parents: 49789
diff changeset
   523
  case (cong u v)
49999
dfb63b9b8908 for the product measure it is enough if only one measure is sigma-finite
hoelzl
parents: 49825
diff changeset
   524
  then have "\<And>w x. w \<in> space M1 \<Longrightarrow> x \<in> space M \<Longrightarrow> u (w, x) = v (w, x)"
49800
a6678da5692c induction prove for positive_integral_fst
hoelzl
parents: 49789
diff changeset
   525
    by (auto simp: space_pair_measure)
a6678da5692c induction prove for positive_integral_fst
hoelzl
parents: 49789
diff changeset
   526
  show ?case
a6678da5692c induction prove for positive_integral_fst
hoelzl
parents: 49789
diff changeset
   527
    apply (subst measurable_cong)
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
   528
    apply (rule nn_integral_cong)
49800
a6678da5692c induction prove for positive_integral_fst
hoelzl
parents: 49789
diff changeset
   529
    apply fact+
a6678da5692c induction prove for positive_integral_fst
hoelzl
parents: 49789
diff changeset
   530
    done
a6678da5692c induction prove for positive_integral_fst
hoelzl
parents: 49789
diff changeset
   531
next
a6678da5692c induction prove for positive_integral_fst
hoelzl
parents: 49789
diff changeset
   532
  case (set Q)
a6678da5692c induction prove for positive_integral_fst
hoelzl
parents: 49789
diff changeset
   533
  have [simp]: "\<And>x y. indicator Q (x, y) = indicator (Pair x -` Q) y"
a6678da5692c induction prove for positive_integral_fst
hoelzl
parents: 49789
diff changeset
   534
    by (auto simp: indicator_def)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
   535
  have "\<And>x. x \<in> space M1 \<Longrightarrow> emeasure M (Pair x -` Q) = \<integral>\<^sup>+ y. indicator Q (x, y) \<partial>M"
49800
a6678da5692c induction prove for positive_integral_fst
hoelzl
parents: 49789
diff changeset
   536
    by (simp add: sets_Pair1[OF set])
49999
dfb63b9b8908 for the product measure it is enough if only one measure is sigma-finite
hoelzl
parents: 49825
diff changeset
   537
  from this measurable_emeasure_Pair[OF set] show ?case
49800
a6678da5692c induction prove for positive_integral_fst
hoelzl
parents: 49789
diff changeset
   538
    by (rule measurable_cong[THEN iffD1])
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
   539
qed (simp_all add: nn_integral_add nn_integral_cmult measurable_compose_Pair1
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
   540
                   nn_integral_monotone_convergence_SUP incseq_def le_fun_def
49800
a6678da5692c induction prove for positive_integral_fst
hoelzl
parents: 49789
diff changeset
   541
              cong: measurable_cong)
a6678da5692c induction prove for positive_integral_fst
hoelzl
parents: 49789
diff changeset
   542
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
   543
lemma (in sigma_finite_measure) nn_integral_fst':
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
   544
  assumes f: "f \<in> borel_measurable (M1 \<Otimes>\<^sub>M M)" "\<And>x. 0 \<le> f x"
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
   545
  shows "(\<integral>\<^sup>+ x. \<integral>\<^sup>+ y. f (x, y) \<partial>M \<partial>M1) = integral\<^sup>N (M1 \<Otimes>\<^sub>M M) f" (is "?I f = _")
49800
a6678da5692c induction prove for positive_integral_fst
hoelzl
parents: 49789
diff changeset
   546
using f proof induct
a6678da5692c induction prove for positive_integral_fst
hoelzl
parents: 49789
diff changeset
   547
  case (cong u v)
53374
a14d2a854c02 tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents: 53015
diff changeset
   548
  then have "?I u = ?I v"
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
   549
    by (intro nn_integral_cong) (auto simp: space_pair_measure)
53374
a14d2a854c02 tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents: 53015
diff changeset
   550
  with cong show ?case
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
   551
    by (simp cong: nn_integral_cong)
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
   552
qed (simp_all add: emeasure_pair_measure nn_integral_cmult nn_integral_add
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
   553
                   nn_integral_monotone_convergence_SUP
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
   554
                   measurable_compose_Pair1 nn_integral_nonneg
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
   555
                   borel_measurable_nn_integral_fst' nn_integral_mono incseq_def le_fun_def
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
   556
              cong: nn_integral_cong)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   557
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
   558
lemma (in sigma_finite_measure) nn_integral_fst:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
   559
  assumes f: "f \<in> borel_measurable (M1 \<Otimes>\<^sub>M M)"
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
   560
  shows "(\<integral>\<^sup>+ x. (\<integral>\<^sup>+ y. f (x, y) \<partial>M) \<partial>M1) = integral\<^sup>N (M1 \<Otimes>\<^sub>M M) f"
49800
a6678da5692c induction prove for positive_integral_fst
hoelzl
parents: 49789
diff changeset
   561
  using f
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
   562
    borel_measurable_nn_integral_fst'[of "\<lambda>x. max 0 (f x)"]
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
   563
    nn_integral_fst'[of "\<lambda>x. max 0 (f x)"]
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
   564
  unfolding nn_integral_max_0 by auto
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   565
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
   566
lemma (in sigma_finite_measure) borel_measurable_nn_integral[measurable (raw)]:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
   567
  "split f \<in> borel_measurable (N \<Otimes>\<^sub>M M) \<Longrightarrow> (\<lambda>x. \<integral>\<^sup>+ y. f x y \<partial>M) \<in> borel_measurable N"
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
   568
  using borel_measurable_nn_integral_fst'[of "\<lambda>x. max 0 (split f x)" N]
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
   569
  by (simp add: nn_integral_max_0)
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   570
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
   571
lemma (in pair_sigma_finite) nn_integral_snd:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
   572
  assumes f: "f \<in> borel_measurable (M1 \<Otimes>\<^sub>M M2)"
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
   573
  shows "(\<integral>\<^sup>+ y. (\<integral>\<^sup>+ x. f (x, y) \<partial>M1) \<partial>M2) = integral\<^sup>N (M1 \<Otimes>\<^sub>M M2) f"
41661
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41659
diff changeset
   574
proof -
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   575
  note measurable_pair_swap[OF f]
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
   576
  from M1.nn_integral_fst[OF this]
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
   577
  have "(\<integral>\<^sup>+ y. (\<integral>\<^sup>+ x. f (x, y) \<partial>M1) \<partial>M2) = (\<integral>\<^sup>+ (x, y). f (y, x) \<partial>(M2 \<Otimes>\<^sub>M M1))"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   578
    by simp
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
   579
  also have "(\<integral>\<^sup>+ (x, y). f (y, x) \<partial>(M2 \<Otimes>\<^sub>M M1)) = integral\<^sup>N (M1 \<Otimes>\<^sub>M M2) f"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   580
    by (subst distr_pair_swap)
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
   581
       (auto simp: nn_integral_distr[OF measurable_pair_swap' f] intro!: nn_integral_cong)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   582
  finally show ?thesis .
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   583
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   584
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   585
lemma (in pair_sigma_finite) Fubini:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
   586
  assumes f: "f \<in> borel_measurable (M1 \<Otimes>\<^sub>M M2)"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
   587
  shows "(\<integral>\<^sup>+ y. (\<integral>\<^sup>+ x. f (x, y) \<partial>M1) \<partial>M2) = (\<integral>\<^sup>+ x. (\<integral>\<^sup>+ y. f (x, y) \<partial>M2) \<partial>M1)"
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
   588
  unfolding nn_integral_snd[OF assms] M2.nn_integral_fst[OF assms] ..
41026
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41023
diff changeset
   589
57235
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   590
lemma (in pair_sigma_finite) Fubini':
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   591
  assumes f: "split f \<in> borel_measurable (M1 \<Otimes>\<^sub>M M2)"
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   592
  shows "(\<integral>\<^sup>+ y. (\<integral>\<^sup>+ x. f x y \<partial>M1) \<partial>M2) = (\<integral>\<^sup>+ x. (\<integral>\<^sup>+ y. f x y \<partial>M2) \<partial>M1)"
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   593
  using Fubini[OF f] by simp
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   594
56994
8d5e5ec1cac3 fixed document generation for HOL-Probability
hoelzl
parents: 56993
diff changeset
   595
subsection {* Products on counting spaces, densities and distributions *}
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   596
59088
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   597
lemma sigma_prod:
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   598
  assumes X_cover: "\<exists>E\<subseteq>A. countable E \<and> X = \<Union>E" and A: "A \<subseteq> Pow X"
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   599
  assumes Y_cover: "\<exists>E\<subseteq>B. countable E \<and> Y = \<Union>E" and B: "B \<subseteq> Pow Y"
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   600
  shows "sigma X A \<Otimes>\<^sub>M sigma Y B = sigma (X \<times> Y) {a \<times> b | a b. a \<in> A \<and> b \<in> B}"
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   601
    (is "?P = ?S")
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   602
proof (rule measure_eqI)
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   603
  have [simp]: "snd \<in> X \<times> Y \<rightarrow> Y" "fst \<in> X \<times> Y \<rightarrow> X"
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   604
    by auto
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   605
  let ?XY = "{{fst -` a \<inter> X \<times> Y | a. a \<in> A}, {snd -` b \<inter> X \<times> Y | b. b \<in> B}}"
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   606
  have "sets ?P = 
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   607
    sets (\<Squnion>\<^sub>\<sigma> xy\<in>?XY. sigma (X \<times> Y) xy)"
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   608
    by (simp add: vimage_algebra_sigma sets_pair_eq_sets_fst_snd A B)
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   609
  also have "\<dots> = sets (sigma (X \<times> Y) (\<Union>?XY))"
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   610
    by (intro Sup_sigma_sigma arg_cong[where f=sets]) auto
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   611
  also have "\<dots> = sets ?S"
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   612
  proof (intro arg_cong[where f=sets] sigma_eqI sigma_sets_eqI) 
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   613
    show "\<Union>?XY \<subseteq> Pow (X \<times> Y)" "{a \<times> b |a b. a \<in> A \<and> b \<in> B} \<subseteq> Pow (X \<times> Y)"
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   614
      using A B by auto
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   615
  next
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   616
    interpret XY: sigma_algebra "X \<times> Y" "sigma_sets (X \<times> Y) {a \<times> b |a b. a \<in> A \<and> b \<in> B}"
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   617
      using A B by (intro sigma_algebra_sigma_sets) auto
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   618
    fix Z assume "Z \<in> \<Union>?XY"
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   619
    then show "Z \<in> sigma_sets (X \<times> Y) {a \<times> b |a b. a \<in> A \<and> b \<in> B}"
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   620
    proof safe
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   621
      fix a assume "a \<in> A"
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   622
      from Y_cover obtain E where E: "E \<subseteq> B" "countable E" and "Y = \<Union>E"
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   623
        by auto
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   624
      with `a \<in> A` A have eq: "fst -` a \<inter> X \<times> Y = (\<Union>e\<in>E. a \<times> e)"
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   625
        by auto
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   626
      show "fst -` a \<inter> X \<times> Y \<in> sigma_sets (X \<times> Y) {a \<times> b |a b. a \<in> A \<and> b \<in> B}"
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   627
        using `a \<in> A` E unfolding eq by (auto intro!: XY.countable_UN')
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   628
    next
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   629
      fix b assume "b \<in> B"
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   630
      from X_cover obtain E where E: "E \<subseteq> A" "countable E" and "X = \<Union>E"
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   631
        by auto
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   632
      with `b \<in> B` B have eq: "snd -` b \<inter> X \<times> Y = (\<Union>e\<in>E. e \<times> b)"
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   633
        by auto
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   634
      show "snd -` b \<inter> X \<times> Y \<in> sigma_sets (X \<times> Y) {a \<times> b |a b. a \<in> A \<and> b \<in> B}"
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   635
        using `b \<in> B` E unfolding eq by (auto intro!: XY.countable_UN')
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   636
    qed
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   637
  next
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   638
    fix Z assume "Z \<in> {a \<times> b |a b. a \<in> A \<and> b \<in> B}"
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   639
    then obtain a b where "Z = a \<times> b" and ab: "a \<in> A" "b \<in> B"
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   640
      by auto
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   641
    then have Z: "Z = (fst -` a \<inter> X \<times> Y) \<inter> (snd -` b \<inter> X \<times> Y)"
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   642
      using A B by auto
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   643
    interpret XY: sigma_algebra "X \<times> Y" "sigma_sets (X \<times> Y) (\<Union>?XY)"
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   644
      by (intro sigma_algebra_sigma_sets) auto
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   645
    show "Z \<in> sigma_sets (X \<times> Y) (\<Union>?XY)"
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   646
      unfolding Z by (rule XY.Int) (blast intro: ab)+
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   647
  qed
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   648
  finally show "sets ?P = sets ?S" .
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   649
next
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   650
  interpret finite_measure "sigma X A" for X A
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   651
    proof qed (simp add: emeasure_sigma)
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   652
  fix A assume "A \<in> sets ?P" then show "emeasure ?P A = emeasure ?S A"
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   653
    by (simp add: emeasure_pair_measure_alt emeasure_sigma)
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   654
qed
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   655
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
   656
lemma sigma_sets_pair_measure_generator_finite:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36649
diff changeset
   657
  assumes "finite A" and "finite B"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   658
  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
   659
  (is "sigma_sets ?prod ?sets = _")
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36649
diff changeset
   660
proof safe
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36649
diff changeset
   661
  have fin: "finite (A \<times> B)" using assms by (rule finite_cartesian_product)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36649
diff changeset
   662
  fix x assume subset: "x \<subseteq> A \<times> B"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36649
diff changeset
   663
  hence "finite x" using fin by (rule finite_subset)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   664
  from this subset show "x \<in> sigma_sets ?prod ?sets"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36649
diff changeset
   665
  proof (induct x)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36649
diff changeset
   666
    case empty show ?case by (rule sigma_sets.Empty)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36649
diff changeset
   667
  next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36649
diff changeset
   668
    case (insert a x)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   669
    hence "{a} \<in> sigma_sets ?prod ?sets" by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36649
diff changeset
   670
    moreover have "x \<in> sigma_sets ?prod ?sets" using insert by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36649
diff changeset
   671
    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
   672
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36649
diff changeset
   673
next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36649
diff changeset
   674
  fix x a b
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   675
  assume "x \<in> sigma_sets ?prod ?sets" and "(a, b) \<in> x"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36649
diff changeset
   676
  from sigma_sets_into_sp[OF _ this(1)] this(2)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   677
  show "a \<in> A" and "b \<in> B" by auto
35833
7b7ae5aa396d Added product measure space
hoelzl
parents:
diff changeset
   678
qed
7b7ae5aa396d Added product measure space
hoelzl
parents:
diff changeset
   679
59088
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   680
lemma borel_prod:
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   681
  "(borel \<Otimes>\<^sub>M borel) = (borel :: ('a::second_countable_topology \<times> 'b::second_countable_topology) measure)"
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   682
  (is "?P = ?B")
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   683
proof -
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   684
  have "?B = sigma UNIV {A \<times> B | A B. open A \<and> open B}"
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   685
    by (rule second_countable_borel_measurable[OF open_prod_generated])
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   686
  also have "\<dots> = ?P"
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   687
    unfolding borel_def
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   688
    by (subst sigma_prod) (auto intro!: exI[of _ "{UNIV}"])
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   689
  finally show ?thesis ..
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   690
qed
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   691
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   692
lemma pair_measure_count_space:
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   693
  assumes A: "finite A" and B: "finite B"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
   694
  shows "count_space A \<Otimes>\<^sub>M count_space B = count_space (A \<times> B)" (is "?P = ?C")
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   695
proof (rule measure_eqI)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   696
  interpret A: finite_measure "count_space A" by (rule finite_measure_count_space) fact
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   697
  interpret B: finite_measure "count_space B" by (rule finite_measure_count_space) fact
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   698
  interpret P: pair_sigma_finite "count_space A" "count_space B" by default
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   699
  show eq: "sets ?P = sets ?C"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   700
    by (simp add: sets_pair_measure sigma_sets_pair_measure_generator_finite A B)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   701
  fix X assume X: "X \<in> sets ?P"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   702
  with eq have X_subset: "X \<subseteq> A \<times> B" by simp
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   703
  with A B have fin_Pair: "\<And>x. finite (Pair x -` X)"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   704
    by (intro finite_subset[OF _ B]) auto
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   705
  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
   706
  show "emeasure ?P X = emeasure ?C X"
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   707
    apply (subst B.emeasure_pair_measure_alt[OF X])
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   708
    apply (subst emeasure_count_space)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   709
    using X_subset apply auto []
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   710
    apply (simp add: fin_Pair emeasure_count_space X_subset fin_X)
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
   711
    apply (subst nn_integral_count_space)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   712
    using A apply simp
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   713
    apply (simp del: real_of_nat_setsum add: real_of_nat_setsum[symmetric])
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   714
    apply (subst card_gt_0_iff)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   715
    apply (simp add: fin_Pair)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   716
    apply (subst card_SigmaI[symmetric])
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   717
    using A apply simp
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   718
    using fin_Pair apply simp
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   719
    using X_subset apply (auto intro!: arg_cong[where f=card])
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   720
    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
   721
qed
35833
7b7ae5aa396d Added product measure space
hoelzl
parents:
diff changeset
   722
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   723
lemma pair_measure_density:
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   724
  assumes f: "f \<in> borel_measurable M1" "AE x in M1. 0 \<le> f x"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   725
  assumes g: "g \<in> borel_measurable M2" "AE x in M2. 0 \<le> g x"
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   726
  assumes "sigma_finite_measure M2" "sigma_finite_measure (density M2 g)"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
   727
  shows "density M1 f \<Otimes>\<^sub>M density M2 g = density (M1 \<Otimes>\<^sub>M M2) (\<lambda>(x,y). f x * g y)" (is "?L = ?R")
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   728
proof (rule measure_eqI)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   729
  interpret M2: sigma_finite_measure M2 by fact
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   730
  interpret D2: sigma_finite_measure "density M2 g" by fact
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   731
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   732
  fix A assume A: "A \<in> sets ?L"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
   733
  with f g have "(\<integral>\<^sup>+ x. f x * \<integral>\<^sup>+ y. g y * indicator A (x, y) \<partial>M2 \<partial>M1) =
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
   734
    (\<integral>\<^sup>+ x. \<integral>\<^sup>+ y. f x * g y * indicator A (x, y) \<partial>M2 \<partial>M1)"
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
   735
    by (intro nn_integral_cong_AE)
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
   736
       (auto simp add: nn_integral_cmult[symmetric] ac_simps)
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   737
  with A f g show "emeasure ?L A = emeasure ?R A"
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
   738
    by (simp add: D2.emeasure_pair_measure emeasure_density nn_integral_density
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
   739
                  M2.nn_integral_fst[symmetric]
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
   740
             cong: nn_integral_cong)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   741
qed simp
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   742
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   743
lemma sigma_finite_measure_distr:
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   744
  assumes "sigma_finite_measure (distr M N f)" and f: "f \<in> measurable M N"
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   745
  shows "sigma_finite_measure M"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   746
proof -
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   747
  interpret sigma_finite_measure "distr M N f" by fact
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57235
diff changeset
   748
  from sigma_finite_countable guess A .. note A = this
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   749
  show ?thesis
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57235
diff changeset
   750
  proof
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57235
diff changeset
   751
    show "\<exists>A. countable A \<and> A \<subseteq> sets M \<and> \<Union>A = space M \<and> (\<forall>a\<in>A. emeasure M a \<noteq> \<infinity>)"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57235
diff changeset
   752
      using A f
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57235
diff changeset
   753
      by (intro exI[of _ "(\<lambda>a. f -` a \<inter> space M) ` A"])
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57235
diff changeset
   754
         (auto simp: emeasure_distr set_eq_iff subset_eq intro: measurable_space)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   755
  qed
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36649
diff changeset
   756
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36649
diff changeset
   757
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   758
lemma pair_measure_distr:
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   759
  assumes f: "f \<in> measurable M S" and g: "g \<in> measurable N T"
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   760
  assumes "sigma_finite_measure (distr N T g)"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
   761
  shows "distr M S f \<Otimes>\<^sub>M distr N T g = distr (M \<Otimes>\<^sub>M N) (S \<Otimes>\<^sub>M T) (\<lambda>(x, y). (f x, g y))" (is "?P = ?D")
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   762
proof (rule measure_eqI)
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   763
  interpret T: sigma_finite_measure "distr N T g" by fact
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   764
  interpret N: sigma_finite_measure N by (rule sigma_finite_measure_distr) fact+
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   765
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46898
diff changeset
   766
  fix A assume A: "A \<in> sets ?P"
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   767
  with f g show "emeasure ?P A = emeasure ?D A"
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   768
    by (auto simp add: N.emeasure_pair_measure_alt space_pair_measure emeasure_distr
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
   769
                       T.emeasure_pair_measure_alt nn_integral_distr
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
   770
             intro!: nn_integral_cong arg_cong[where f="emeasure N"])
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   771
qed simp
39097
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   772
50104
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50003
diff changeset
   773
lemma pair_measure_eqI:
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50003
diff changeset
   774
  assumes "sigma_finite_measure M1" "sigma_finite_measure M2"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
   775
  assumes sets: "sets (M1 \<Otimes>\<^sub>M M2) = sets M"
50104
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50003
diff changeset
   776
  assumes emeasure: "\<And>A B. A \<in> sets M1 \<Longrightarrow> B \<in> sets M2 \<Longrightarrow> emeasure M1 A * emeasure M2 B = emeasure M (A \<times> B)"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
   777
  shows "M1 \<Otimes>\<^sub>M M2 = M"
50104
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50003
diff changeset
   778
proof -
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50003
diff changeset
   779
  interpret M1: sigma_finite_measure M1 by fact
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50003
diff changeset
   780
  interpret M2: sigma_finite_measure M2 by fact
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50003
diff changeset
   781
  interpret pair_sigma_finite M1 M2 by default
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50003
diff changeset
   782
  from sigma_finite_up_in_pair_measure_generator guess F :: "nat \<Rightarrow> ('a \<times> 'b) set" .. note F = this
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50003
diff changeset
   783
  let ?E = "{a \<times> b |a b. a \<in> sets M1 \<and> b \<in> sets M2}"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50244
diff changeset
   784
  let ?P = "M1 \<Otimes>\<^sub>M M2"
50104
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50003
diff changeset
   785
  show ?thesis
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50003
diff changeset
   786
  proof (rule measure_eqI_generator_eq[OF Int_stable_pair_measure_generator[of M1 M2]])
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50003
diff changeset
   787
    show "?E \<subseteq> Pow (space ?P)"
50244
de72bbe42190 qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents: 50104
diff changeset
   788
      using sets.space_closed[of M1] sets.space_closed[of M2] by (auto simp: space_pair_measure)
50104
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50003
diff changeset
   789
    show "sets ?P = sigma_sets (space ?P) ?E"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50003
diff changeset
   790
      by (simp add: sets_pair_measure space_pair_measure)
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50003
diff changeset
   791
    then show "sets M = sigma_sets (space ?P) ?E"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50003
diff changeset
   792
      using sets[symmetric] by simp
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50003
diff changeset
   793
  next
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50003
diff changeset
   794
    show "range F \<subseteq> ?E" "(\<Union>i. F i) = space ?P" "\<And>i. emeasure ?P (F i) \<noteq> \<infinity>"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50003
diff changeset
   795
      using F by (auto simp: space_pair_measure)
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50003
diff changeset
   796
  next
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50003
diff changeset
   797
    fix X assume "X \<in> ?E"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50003
diff changeset
   798
    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
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50003
diff changeset
   799
    then have "emeasure ?P X = emeasure M1 A * emeasure M2 B"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50003
diff changeset
   800
       by (simp add: M2.emeasure_pair_measure_Times)
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50003
diff changeset
   801
    also have "\<dots> = emeasure M (A \<times> B)"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50003
diff changeset
   802
      using A B emeasure by auto
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50003
diff changeset
   803
    finally show "emeasure ?P X = emeasure M X"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50003
diff changeset
   804
      by simp
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50003
diff changeset
   805
  qed
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50003
diff changeset
   806
qed
57025
e7fd64f82876 add various lemmas
hoelzl
parents: 56996
diff changeset
   807
  
e7fd64f82876 add various lemmas
hoelzl
parents: 56996
diff changeset
   808
lemma sets_pair_countable:
e7fd64f82876 add various lemmas
hoelzl
parents: 56996
diff changeset
   809
  assumes "countable S1" "countable S2"
e7fd64f82876 add various lemmas
hoelzl
parents: 56996
diff changeset
   810
  assumes M: "sets M = Pow S1" and N: "sets N = Pow S2"
e7fd64f82876 add various lemmas
hoelzl
parents: 56996
diff changeset
   811
  shows "sets (M \<Otimes>\<^sub>M N) = Pow (S1 \<times> S2)"
e7fd64f82876 add various lemmas
hoelzl
parents: 56996
diff changeset
   812
proof auto
e7fd64f82876 add various lemmas
hoelzl
parents: 56996
diff changeset
   813
  fix x a b assume x: "x \<in> sets (M \<Otimes>\<^sub>M N)" "(a, b) \<in> x"
e7fd64f82876 add various lemmas
hoelzl
parents: 56996
diff changeset
   814
  from sets.sets_into_space[OF x(1)] x(2)
e7fd64f82876 add various lemmas
hoelzl
parents: 56996
diff changeset
   815
    sets_eq_imp_space_eq[of N "count_space S2"] sets_eq_imp_space_eq[of M "count_space S1"] M N
e7fd64f82876 add various lemmas
hoelzl
parents: 56996
diff changeset
   816
  show "a \<in> S1" "b \<in> S2"
e7fd64f82876 add various lemmas
hoelzl
parents: 56996
diff changeset
   817
    by (auto simp: space_pair_measure)
e7fd64f82876 add various lemmas
hoelzl
parents: 56996
diff changeset
   818
next
e7fd64f82876 add various lemmas
hoelzl
parents: 56996
diff changeset
   819
  fix X assume X: "X \<subseteq> S1 \<times> S2"
e7fd64f82876 add various lemmas
hoelzl
parents: 56996
diff changeset
   820
  then have "countable X"
e7fd64f82876 add various lemmas
hoelzl
parents: 56996
diff changeset
   821
    by (metis countable_subset `countable S1` `countable S2` countable_SIGMA)
e7fd64f82876 add various lemmas
hoelzl
parents: 56996
diff changeset
   822
  have "X = (\<Union>(a, b)\<in>X. {a} \<times> {b})" by auto
e7fd64f82876 add various lemmas
hoelzl
parents: 56996
diff changeset
   823
  also have "\<dots> \<in> sets (M \<Otimes>\<^sub>M N)"
e7fd64f82876 add various lemmas
hoelzl
parents: 56996
diff changeset
   824
    using X
e7fd64f82876 add various lemmas
hoelzl
parents: 56996
diff changeset
   825
    by (safe intro!: sets.countable_UN' `countable X` subsetI pair_measureI) (auto simp: M N)
e7fd64f82876 add various lemmas
hoelzl
parents: 56996
diff changeset
   826
  finally show "X \<in> sets (M \<Otimes>\<^sub>M N)" .
e7fd64f82876 add various lemmas
hoelzl
parents: 56996
diff changeset
   827
qed
e7fd64f82876 add various lemmas
hoelzl
parents: 56996
diff changeset
   828
e7fd64f82876 add various lemmas
hoelzl
parents: 56996
diff changeset
   829
lemma pair_measure_countable:
e7fd64f82876 add various lemmas
hoelzl
parents: 56996
diff changeset
   830
  assumes "countable S1" "countable S2"
e7fd64f82876 add various lemmas
hoelzl
parents: 56996
diff changeset
   831
  shows "count_space S1 \<Otimes>\<^sub>M count_space S2 = count_space (S1 \<times> S2)"
e7fd64f82876 add various lemmas
hoelzl
parents: 56996
diff changeset
   832
proof (rule pair_measure_eqI)
e7fd64f82876 add various lemmas
hoelzl
parents: 56996
diff changeset
   833
  show "sigma_finite_measure (count_space S1)" "sigma_finite_measure (count_space S2)"
e7fd64f82876 add various lemmas
hoelzl
parents: 56996
diff changeset
   834
    using assms by (auto intro!: sigma_finite_measure_count_space_countable)
e7fd64f82876 add various lemmas
hoelzl
parents: 56996
diff changeset
   835
  show "sets (count_space S1 \<Otimes>\<^sub>M count_space S2) = sets (count_space (S1 \<times> S2))"
e7fd64f82876 add various lemmas
hoelzl
parents: 56996
diff changeset
   836
    by (subst sets_pair_countable[OF assms]) auto
e7fd64f82876 add various lemmas
hoelzl
parents: 56996
diff changeset
   837
next
e7fd64f82876 add various lemmas
hoelzl
parents: 56996
diff changeset
   838
  fix A B assume "A \<in> sets (count_space S1)" "B \<in> sets (count_space S2)"
e7fd64f82876 add various lemmas
hoelzl
parents: 56996
diff changeset
   839
  then show "emeasure (count_space S1) A * emeasure (count_space S2) B = 
e7fd64f82876 add various lemmas
hoelzl
parents: 56996
diff changeset
   840
    emeasure (count_space (S1 \<times> S2)) (A \<times> B)"
e7fd64f82876 add various lemmas
hoelzl
parents: 56996
diff changeset
   841
    by (subst (1 2 3) emeasure_count_space) (auto simp: finite_cartesian_product_iff)
e7fd64f82876 add various lemmas
hoelzl
parents: 56996
diff changeset
   842
qed
50104
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50003
diff changeset
   843
57235
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   844
subsection {* Product of Borel spaces *}
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   845
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   846
lemma borel_Times:
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   847
  fixes A :: "'a::topological_space set" and B :: "'b::topological_space set"
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   848
  assumes A: "A \<in> sets borel" and B: "B \<in> sets borel"
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   849
  shows "A \<times> B \<in> sets borel"
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   850
proof -
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   851
  have "A \<times> B = (A\<times>UNIV) \<inter> (UNIV \<times> B)"
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   852
    by auto
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   853
  moreover
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   854
  { have "A \<in> sigma_sets UNIV {S. open S}" using A by (simp add: sets_borel)
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   855
    then have "A\<times>UNIV \<in> sets borel"
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   856
    proof (induct A)
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   857
      case (Basic S) then show ?case
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   858
        by (auto intro!: borel_open open_Times)
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   859
    next
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   860
      case (Compl A)
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   861
      moreover have *: "(UNIV - A) \<times> UNIV = UNIV - (A \<times> UNIV)"
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   862
        by auto
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   863
      ultimately show ?case
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   864
        unfolding * by auto
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   865
    next
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   866
      case (Union A)
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   867
      moreover have *: "(UNION UNIV A) \<times> UNIV = UNION UNIV (\<lambda>i. A i \<times> UNIV)"
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   868
        by auto
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   869
      ultimately show ?case
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   870
        unfolding * by auto
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   871
    qed simp }
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   872
  moreover
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   873
  { have "B \<in> sigma_sets UNIV {S. open S}" using B by (simp add: sets_borel)
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   874
    then have "UNIV\<times>B \<in> sets borel"
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   875
    proof (induct B)
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   876
      case (Basic S) then show ?case
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   877
        by (auto intro!: borel_open open_Times)
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   878
    next
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   879
      case (Compl B)
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   880
      moreover have *: "UNIV \<times> (UNIV - B) = UNIV - (UNIV \<times> B)"
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   881
        by auto
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   882
      ultimately show ?case
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   883
        unfolding * by auto
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   884
    next
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   885
      case (Union B)
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   886
      moreover have *: "UNIV \<times> (UNION UNIV B) = UNION UNIV (\<lambda>i. UNIV \<times> B i)"
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   887
        by auto
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   888
      ultimately show ?case
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   889
        unfolding * by auto
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   890
    qed simp }
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   891
  ultimately show ?thesis
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   892
    by auto
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   893
qed
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   894
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   895
lemma finite_measure_pair_measure:
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   896
  assumes "finite_measure M" "finite_measure N"
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   897
  shows "finite_measure (N  \<Otimes>\<^sub>M M)"
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   898
proof (rule finite_measureI)
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   899
  interpret M: finite_measure M by fact
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   900
  interpret N: finite_measure N by fact
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   901
  show "emeasure (N  \<Otimes>\<^sub>M M) (space (N  \<Otimes>\<^sub>M M)) \<noteq> \<infinity>"
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   902
    by (auto simp: space_pair_measure M.emeasure_pair_measure_Times)
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   903
qed
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   904
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   905
end