src/HOL/Probability/Radon_Nikodym.thy
author hoelzl
Wed, 08 Dec 2010 16:15:14 +0100
changeset 41095 c335d880ff82
parent 41023 9118eb4eb8dc
child 41097 a1abfa4e2b44
permissions -rw-r--r--
cleanup bijectivity btw. product spaces and pairs
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
     1
theory Radon_Nikodym
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
     2
imports Lebesgue_Integration
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
     3
begin
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
     4
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
     5
lemma less_\<omega>_Ex_of_nat: "x < \<omega> \<longleftrightarrow> (\<exists>n. x < of_nat n)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
     6
proof safe
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
     7
  assume "x < \<omega>"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
     8
  then obtain r where "0 \<le> r" "x = Real r" by (cases x) auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
     9
  moreover obtain n where "r < of_nat n" using ex_less_of_nat by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
    10
  ultimately show "\<exists>n. x < of_nat n" by (auto simp: real_eq_of_nat)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
    11
qed auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
    12
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    13
lemma (in sigma_finite_measure) Ex_finite_integrable_function:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    14
  shows "\<exists>h\<in>borel_measurable M. positive_integral h \<noteq> \<omega> \<and> (\<forall>x\<in>space M. 0 < h x \<and> h x < \<omega>)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    15
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    16
  obtain A :: "nat \<Rightarrow> 'a set" where
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    17
    range: "range A \<subseteq> sets M" and
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    18
    space: "(\<Union>i. A i) = space M" and
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    19
    measure: "\<And>i. \<mu> (A i) \<noteq> \<omega>" and
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    20
    disjoint: "disjoint_family A"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    21
    using disjoint_sigma_finite by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    22
  let "?B i" = "2^Suc i * \<mu> (A i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    23
  have "\<forall>i. \<exists>x. 0 < x \<and> x < inverse (?B i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    24
  proof
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    25
    fix i show "\<exists>x. 0 < x \<and> x < inverse (?B i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    26
    proof cases
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    27
      assume "\<mu> (A i) = 0"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    28
      then show ?thesis by (auto intro!: exI[of _ 1])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    29
    next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    30
      assume not_0: "\<mu> (A i) \<noteq> 0"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    31
      then have "?B i \<noteq> \<omega>" using measure[of i] by auto
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40871
diff changeset
    32
      then have "inverse (?B i) \<noteq> 0" unfolding pextreal_inverse_eq_0 by simp
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    33
      then show ?thesis using measure[of i] not_0
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    34
        by (auto intro!: exI[of _ "inverse (?B i) / 2"]
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40871
diff changeset
    35
                 simp: pextreal_noteq_omega_Ex zero_le_mult_iff zero_less_mult_iff mult_le_0_iff power_le_zero_eq)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    36
    qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    37
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    38
  from choice[OF this] obtain n where n: "\<And>i. 0 < n i"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    39
    "\<And>i. n i < inverse (2^Suc i * \<mu> (A i))" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    40
  let "?h x" = "\<Sum>\<^isub>\<infinity> i. n i * indicator (A i) x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    41
  show ?thesis
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    42
  proof (safe intro!: bexI[of _ ?h] del: notI)
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38656
diff changeset
    43
    have "\<And>i. A i \<in> sets M"
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38656
diff changeset
    44
      using range by fastsimp+
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38656
diff changeset
    45
    then have "positive_integral ?h = (\<Sum>\<^isub>\<infinity> i. n i * \<mu> (A i))"
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38656
diff changeset
    46
      by (simp add: positive_integral_psuminf positive_integral_cmult_indicator)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    47
    also have "\<dots> \<le> (\<Sum>\<^isub>\<infinity> i. Real ((1 / 2)^Suc i))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    48
    proof (rule psuminf_le)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    49
      fix N show "n N * \<mu> (A N) \<le> Real ((1 / 2) ^ Suc N)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    50
        using measure[of N] n[of N]
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38656
diff changeset
    51
        by (cases "n N")
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40871
diff changeset
    52
           (auto simp: pextreal_noteq_omega_Ex field_simps zero_le_mult_iff
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38656
diff changeset
    53
                       mult_le_0_iff mult_less_0_iff power_less_zero_eq
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38656
diff changeset
    54
                       power_le_zero_eq inverse_eq_divide less_divide_eq
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38656
diff changeset
    55
                       power_divide split: split_if_asm)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    56
    qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    57
    also have "\<dots> = Real 1"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    58
      by (rule suminf_imp_psuminf, rule power_half_series, auto)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    59
    finally show "positive_integral ?h \<noteq> \<omega>" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    60
  next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    61
    fix x assume "x \<in> space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    62
    then obtain i where "x \<in> A i" using space[symmetric] by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    63
    from psuminf_cmult_indicator[OF disjoint, OF this]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    64
    have "?h x = n i" by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    65
    then show "0 < ?h x" and "?h x < \<omega>" using n[of i] by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    66
  next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    67
    show "?h \<in> borel_measurable M" using range
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40871
diff changeset
    68
      by (auto intro!: borel_measurable_psuminf borel_measurable_pextreal_times)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    69
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    70
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    71
40871
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
    72
subsection "Absolutely continuous"
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
    73
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    74
definition (in measure_space)
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40871
diff changeset
    75
  "absolutely_continuous \<nu> = (\<forall>N\<in>null_sets. \<nu> N = (0 :: pextreal))"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    76
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
    77
lemma (in sigma_finite_measure) absolutely_continuous_AE:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
    78
  assumes "measure_space M \<nu>" "absolutely_continuous \<nu>" "AE x. P x"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
    79
  shows "measure_space.almost_everywhere M \<nu> P"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
    80
proof -
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
    81
  interpret \<nu>: measure_space M \<nu> by fact
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
    82
  from `AE x. P x` obtain N where N: "N \<in> null_sets" and "{x\<in>space M. \<not> P x} \<subseteq> N"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
    83
    unfolding almost_everywhere_def by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
    84
  show "\<nu>.almost_everywhere P"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
    85
  proof (rule \<nu>.AE_I')
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
    86
    show "{x\<in>space M. \<not> P x} \<subseteq> N" by fact
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
    87
    from `absolutely_continuous \<nu>` show "N \<in> \<nu>.null_sets"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
    88
      using N unfolding absolutely_continuous_def by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
    89
  qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
    90
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
    91
39097
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39092
diff changeset
    92
lemma (in finite_measure_space) absolutely_continuousI:
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39092
diff changeset
    93
  assumes "finite_measure_space M \<nu>"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39092
diff changeset
    94
  assumes v: "\<And>x. \<lbrakk> x \<in> space M ; \<mu> {x} = 0 \<rbrakk> \<Longrightarrow> \<nu> {x} = 0"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39092
diff changeset
    95
  shows "absolutely_continuous \<nu>"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39092
diff changeset
    96
proof (unfold absolutely_continuous_def sets_eq_Pow, safe)
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39092
diff changeset
    97
  fix N assume "\<mu> N = 0" "N \<subseteq> space M"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39092
diff changeset
    98
  interpret v: finite_measure_space M \<nu> by fact
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39092
diff changeset
    99
  have "\<nu> N = \<nu> (\<Union>x\<in>N. {x})" by simp
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39092
diff changeset
   100
  also have "\<dots> = (\<Sum>x\<in>N. \<nu> {x})"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39092
diff changeset
   101
  proof (rule v.measure_finitely_additive''[symmetric])
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39092
diff changeset
   102
    show "finite N" using `N \<subseteq> space M` finite_space by (auto intro: finite_subset)
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39092
diff changeset
   103
    show "disjoint_family_on (\<lambda>i. {i}) N" unfolding disjoint_family_on_def by auto
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39092
diff changeset
   104
    fix x assume "x \<in> N" thus "{x} \<in> sets M" using `N \<subseteq> space M` sets_eq_Pow by auto
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39092
diff changeset
   105
  qed
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39092
diff changeset
   106
  also have "\<dots> = 0"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39092
diff changeset
   107
  proof (safe intro!: setsum_0')
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39092
diff changeset
   108
    fix x assume "x \<in> N"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39092
diff changeset
   109
    hence "\<mu> {x} \<le> \<mu> N" using sets_eq_Pow `N \<subseteq> space M` by (auto intro!: measure_mono)
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39092
diff changeset
   110
    hence "\<mu> {x} = 0" using `\<mu> N = 0` by simp
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39092
diff changeset
   111
    thus "\<nu> {x} = 0" using v[of x] `x \<in> N` `N \<subseteq> space M` by auto
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39092
diff changeset
   112
  qed
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39092
diff changeset
   113
  finally show "\<nu> N = 0" .
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39092
diff changeset
   114
qed
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39092
diff changeset
   115
40871
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
   116
lemma (in measure_space) density_is_absolutely_continuous:
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
   117
  assumes "\<And>A. A \<in> sets M \<Longrightarrow> \<nu> A = positive_integral (\<lambda>x. f x * indicator A x)"
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
   118
  shows "absolutely_continuous \<nu>"
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
   119
  using assms unfolding absolutely_continuous_def
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
   120
  by (simp add: positive_integral_null_set)
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
   121
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
   122
subsection "Existence of the Radon-Nikodym derivative"
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
   123
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   124
lemma (in finite_measure) Radon_Nikodym_aux_epsilon:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   125
  fixes e :: real assumes "0 < e"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   126
  assumes "finite_measure M s"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   127
  shows "\<exists>A\<in>sets M. real (\<mu> (space M)) - real (s (space M)) \<le>
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   128
                    real (\<mu> A) - real (s A) \<and>
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   129
                    (\<forall>B\<in>sets M. B \<subseteq> A \<longrightarrow> - e < real (\<mu> B) - real (s B))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   130
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   131
  let "?d A" = "real (\<mu> A) - real (s A)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   132
  interpret M': finite_measure M s by fact
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   133
  let "?A A" = "if (\<forall>B\<in>sets M. B \<subseteq> space M - A \<longrightarrow> -e < ?d B)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   134
    then {}
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   135
    else (SOME B. B \<in> sets M \<and> B \<subseteq> space M - A \<and> ?d B \<le> -e)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   136
  def A \<equiv> "\<lambda>n. ((\<lambda>B. B \<union> ?A B) ^^ n) {}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   137
  have A_simps[simp]:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   138
    "A 0 = {}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   139
    "\<And>n. A (Suc n) = (A n \<union> ?A (A n))" unfolding A_def by simp_all
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   140
  { fix A assume "A \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   141
    have "?A A \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   142
      by (auto intro!: someI2[of _ _ "\<lambda>A. A \<in> sets M"] simp: not_less) }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   143
  note A'_in_sets = this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   144
  { fix n have "A n \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   145
    proof (induct n)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   146
      case (Suc n) thus "A (Suc n) \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   147
        using A'_in_sets[of "A n"] by (auto split: split_if_asm)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   148
    qed (simp add: A_def) }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   149
  note A_in_sets = this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   150
  hence "range A \<subseteq> sets M" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   151
  { fix n B
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   152
    assume Ex: "\<exists>B. B \<in> sets M \<and> B \<subseteq> space M - A n \<and> ?d B \<le> -e"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   153
    hence False: "\<not> (\<forall>B\<in>sets M. B \<subseteq> space M - A n \<longrightarrow> -e < ?d B)" by (auto simp: not_less)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   154
    have "?d (A (Suc n)) \<le> ?d (A n) - e" unfolding A_simps if_not_P[OF False]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   155
    proof (rule someI2_ex[OF Ex])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   156
      fix B assume "B \<in> sets M \<and> B \<subseteq> space M - A n \<and> ?d B \<le> - e"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   157
      hence "A n \<inter> B = {}" "B \<in> sets M" and dB: "?d B \<le> -e" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   158
      hence "?d (A n \<union> B) = ?d (A n) + ?d B"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   159
        using `A n \<in> sets M` real_finite_measure_Union M'.real_finite_measure_Union by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   160
      also have "\<dots> \<le> ?d (A n) - e" using dB by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   161
      finally show "?d (A n \<union> B) \<le> ?d (A n) - e" .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   162
    qed }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   163
  note dA_epsilon = this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   164
  { fix n have "?d (A (Suc n)) \<le> ?d (A n)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   165
    proof (cases "\<exists>B. B\<in>sets M \<and> B \<subseteq> space M - A n \<and> ?d B \<le> - e")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   166
      case True from dA_epsilon[OF this] show ?thesis using `0 < e` by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   167
    next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   168
      case False
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   169
      hence "\<forall>B\<in>sets M. B \<subseteq> space M - A n \<longrightarrow> -e < ?d B" by (auto simp: not_le)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   170
      thus ?thesis by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   171
    qed }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   172
  note dA_mono = this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   173
  show ?thesis
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   174
  proof (cases "\<exists>n. \<forall>B\<in>sets M. B \<subseteq> space M - A n \<longrightarrow> -e < ?d B")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   175
    case True then obtain n where B: "\<And>B. \<lbrakk> B \<in> sets M; B \<subseteq> space M - A n\<rbrakk> \<Longrightarrow> -e < ?d B" by blast
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   176
    show ?thesis
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   177
    proof (safe intro!: bexI[of _ "space M - A n"])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   178
      fix B assume "B \<in> sets M" "B \<subseteq> space M - A n"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   179
      from B[OF this] show "-e < ?d B" .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   180
    next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   181
      show "space M - A n \<in> sets M" by (rule compl_sets) fact
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   182
    next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   183
      show "?d (space M) \<le> ?d (space M - A n)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   184
      proof (induct n)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   185
        fix n assume "?d (space M) \<le> ?d (space M - A n)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   186
        also have "\<dots> \<le> ?d (space M - A (Suc n))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   187
          using A_in_sets sets_into_space dA_mono[of n]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   188
            real_finite_measure_Diff[of "space M"]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   189
            real_finite_measure_Diff[of "space M"]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   190
            M'.real_finite_measure_Diff[of "space M"]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   191
            M'.real_finite_measure_Diff[of "space M"]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   192
          by (simp del: A_simps)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   193
        finally show "?d (space M) \<le> ?d (space M - A (Suc n))" .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   194
      qed simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   195
    qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   196
  next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   197
    case False hence B: "\<And>n. \<exists>B. B\<in>sets M \<and> B \<subseteq> space M - A n \<and> ?d B \<le> - e"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   198
      by (auto simp add: not_less)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   199
    { fix n have "?d (A n) \<le> - real n * e"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   200
      proof (induct n)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   201
        case (Suc n) with dA_epsilon[of n, OF B] show ?case by (simp del: A_simps add: real_of_nat_Suc field_simps)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   202
      qed simp } note dA_less = this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   203
    have decseq: "decseq (\<lambda>n. ?d (A n))" unfolding decseq_eq_incseq
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   204
    proof (rule incseq_SucI)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   205
      fix n show "- ?d (A n) \<le> - ?d (A (Suc n))" using dA_mono[of n] by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   206
    qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   207
    from real_finite_continuity_from_below[of A] `range A \<subseteq> sets M`
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   208
      M'.real_finite_continuity_from_below[of A]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   209
    have convergent: "(\<lambda>i. ?d (A i)) ----> ?d (\<Union>i. A i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   210
      by (auto intro!: LIMSEQ_diff)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   211
    obtain n :: nat where "- ?d (\<Union>i. A i) / e < real n" using reals_Archimedean2 by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   212
    moreover from order_trans[OF decseq_le[OF decseq convergent] dA_less]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   213
    have "real n \<le> - ?d (\<Union>i. A i) / e" using `0<e` by (simp add: field_simps)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   214
    ultimately show ?thesis by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   215
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   216
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   217
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   218
lemma (in finite_measure) Radon_Nikodym_aux:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   219
  assumes "finite_measure M s"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   220
  shows "\<exists>A\<in>sets M. real (\<mu> (space M)) - real (s (space M)) \<le>
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   221
                    real (\<mu> A) - real (s A) \<and>
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   222
                    (\<forall>B\<in>sets M. B \<subseteq> A \<longrightarrow> 0 \<le> real (\<mu> B) - real (s B))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   223
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   224
  let "?d A" = "real (\<mu> A) - real (s A)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   225
  let "?P A B n" = "A \<in> sets M \<and> A \<subseteq> B \<and> ?d B \<le> ?d A \<and> (\<forall>C\<in>sets M. C \<subseteq> A \<longrightarrow> - 1 / real (Suc n) < ?d C)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   226
  interpret M': finite_measure M s by fact
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38656
diff changeset
   227
  let "?r S" = "restricted_space S"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   228
  { fix S n
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   229
    assume S: "S \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   230
    hence **: "\<And>X. X \<in> op \<inter> S ` sets M \<longleftrightarrow> X \<in> sets M \<and> X \<subseteq> S" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   231
    from M'.restricted_finite_measure[of S] restricted_finite_measure[of S] S
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   232
    have "finite_measure (?r S) \<mu>" "0 < 1 / real (Suc n)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   233
      "finite_measure (?r S) s" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   234
    from finite_measure.Radon_Nikodym_aux_epsilon[OF this] guess X ..
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   235
    hence "?P X S n"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   236
    proof (simp add: **, safe)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   237
      fix C assume C: "C \<in> sets M" "C \<subseteq> X" "X \<subseteq> S" and
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   238
        *: "\<forall>B\<in>sets M. S \<inter> B \<subseteq> X \<longrightarrow> - (1 / real (Suc n)) < ?d (S \<inter> B)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   239
      hence "C \<subseteq> S" "C \<subseteq> X" "S \<inter> C = C" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   240
      with *[THEN bspec, OF `C \<in> sets M`]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   241
      show "- (1 / real (Suc n)) < ?d C" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   242
    qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   243
    hence "\<exists>A. ?P A S n" by auto }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   244
  note Ex_P = this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   245
  def A \<equiv> "nat_rec (space M) (\<lambda>n A. SOME B. ?P B A n)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   246
  have A_Suc: "\<And>n. A (Suc n) = (SOME B. ?P B (A n) n)" by (simp add: A_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   247
  have A_0[simp]: "A 0 = space M" unfolding A_def by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   248
  { fix i have "A i \<in> sets M" unfolding A_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   249
    proof (induct i)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   250
      case (Suc i)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   251
      from Ex_P[OF this, of i] show ?case unfolding nat_rec_Suc
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   252
        by (rule someI2_ex) simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   253
    qed simp }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   254
  note A_in_sets = this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   255
  { fix n have "?P (A (Suc n)) (A n) n"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   256
      using Ex_P[OF A_in_sets] unfolding A_Suc
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   257
      by (rule someI2_ex) simp }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   258
  note P_A = this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   259
  have "range A \<subseteq> sets M" using A_in_sets by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   260
  have A_mono: "\<And>i. A (Suc i) \<subseteq> A i" using P_A by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   261
  have mono_dA: "mono (\<lambda>i. ?d (A i))" using P_A by (simp add: mono_iff_le_Suc)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   262
  have epsilon: "\<And>C i. \<lbrakk> C \<in> sets M; C \<subseteq> A (Suc i) \<rbrakk> \<Longrightarrow> - 1 / real (Suc i) < ?d C"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   263
      using P_A by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   264
  show ?thesis
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   265
  proof (safe intro!: bexI[of _ "\<Inter>i. A i"])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   266
    show "(\<Inter>i. A i) \<in> sets M" using A_in_sets by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   267
    from `range A \<subseteq> sets M` A_mono
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   268
      real_finite_continuity_from_above[of A]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   269
      M'.real_finite_continuity_from_above[of A]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   270
    have "(\<lambda>i. ?d (A i)) ----> ?d (\<Inter>i. A i)" by (auto intro!: LIMSEQ_diff)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   271
    thus "?d (space M) \<le> ?d (\<Inter>i. A i)" using mono_dA[THEN monoD, of 0 _]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   272
      by (rule_tac LIMSEQ_le_const) (auto intro!: exI)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   273
  next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   274
    fix B assume B: "B \<in> sets M" "B \<subseteq> (\<Inter>i. A i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   275
    show "0 \<le> ?d B"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   276
    proof (rule ccontr)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   277
      assume "\<not> 0 \<le> ?d B"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   278
      hence "0 < - ?d B" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   279
      from ex_inverse_of_nat_Suc_less[OF this]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   280
      obtain n where *: "?d B < - 1 / real (Suc n)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   281
        by (auto simp: real_eq_of_nat inverse_eq_divide field_simps)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   282
      have "B \<subseteq> A (Suc n)" using B by (auto simp del: nat_rec_Suc)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   283
      from epsilon[OF B(1) this] *
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   284
      show False by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   285
    qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   286
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   287
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   288
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   289
lemma (in finite_measure) Radon_Nikodym_finite_measure:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   290
  assumes "finite_measure M \<nu>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   291
  assumes "absolutely_continuous \<nu>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   292
  shows "\<exists>f \<in> borel_measurable M. \<forall>A\<in>sets M. \<nu> A = positive_integral (\<lambda>x. f x * indicator A x)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   293
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   294
  interpret M': finite_measure M \<nu> using assms(1) .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   295
  def G \<equiv> "{g \<in> borel_measurable M. \<forall>A\<in>sets M. positive_integral (\<lambda>x. g x * indicator A x) \<le> \<nu> A}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   296
  have "(\<lambda>x. 0) \<in> G" unfolding G_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   297
  hence "G \<noteq> {}" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   298
  { fix f g assume f: "f \<in> G" and g: "g \<in> G"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   299
    have "(\<lambda>x. max (g x) (f x)) \<in> G" (is "?max \<in> G") unfolding G_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   300
    proof safe
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   301
      show "?max \<in> borel_measurable M" using f g unfolding G_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   302
      let ?A = "{x \<in> space M. f x \<le> g x}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   303
      have "?A \<in> sets M" using f g unfolding G_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   304
      fix A assume "A \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   305
      hence sets: "?A \<inter> A \<in> sets M" "(space M - ?A) \<inter> A \<in> sets M" using `?A \<in> sets M` by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   306
      have union: "((?A \<inter> A) \<union> ((space M - ?A) \<inter> A)) = A"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   307
        using sets_into_space[OF `A \<in> sets M`] by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   308
      have "\<And>x. x \<in> space M \<Longrightarrow> max (g x) (f x) * indicator A x =
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   309
        g x * indicator (?A \<inter> A) x + f x * indicator ((space M - ?A) \<inter> A) x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   310
        by (auto simp: indicator_def max_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   311
      hence "positive_integral (\<lambda>x. max (g x) (f x) * indicator A x) =
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   312
        positive_integral (\<lambda>x. g x * indicator (?A \<inter> A) x) +
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   313
        positive_integral (\<lambda>x. f x * indicator ((space M - ?A) \<inter> A) x)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   314
        using f g sets unfolding G_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   315
        by (auto cong: positive_integral_cong intro!: positive_integral_add borel_measurable_indicator)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   316
      also have "\<dots> \<le> \<nu> (?A \<inter> A) + \<nu> ((space M - ?A) \<inter> A)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   317
        using f g sets unfolding G_def by (auto intro!: add_mono)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   318
      also have "\<dots> = \<nu> A"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   319
        using M'.measure_additive[OF sets] union by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   320
      finally show "positive_integral (\<lambda>x. max (g x) (f x) * indicator A x) \<le> \<nu> A" .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   321
    qed }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   322
  note max_in_G = this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   323
  { fix f g assume  "f \<up> g" and f: "\<And>i. f i \<in> G"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   324
    have "g \<in> G" unfolding G_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   325
    proof safe
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   326
      from `f \<up> g` have [simp]: "g = (SUP i. f i)" unfolding isoton_def by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   327
      have f_borel: "\<And>i. f i \<in> borel_measurable M" using f unfolding G_def by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   328
      thus "g \<in> borel_measurable M" by (auto intro!: borel_measurable_SUP)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   329
      fix A assume "A \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   330
      hence "\<And>i. (\<lambda>x. f i x * indicator A x) \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   331
        using f_borel by (auto intro!: borel_measurable_indicator)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   332
      from positive_integral_isoton[OF isoton_indicator[OF `f \<up> g`] this]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   333
      have SUP: "positive_integral (\<lambda>x. g x * indicator A x) =
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   334
          (SUP i. positive_integral (\<lambda>x. f i x * indicator A x))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   335
        unfolding isoton_def by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   336
      show "positive_integral (\<lambda>x. g x * indicator A x) \<le> \<nu> A" unfolding SUP
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   337
        using f `A \<in> sets M` unfolding G_def by (auto intro!: SUP_leI)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   338
    qed }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   339
  note SUP_in_G = this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   340
  let ?y = "SUP g : G. positive_integral g"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   341
  have "?y \<le> \<nu> (space M)" unfolding G_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   342
  proof (safe intro!: SUP_leI)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   343
    fix g assume "\<forall>A\<in>sets M. positive_integral (\<lambda>x. g x * indicator A x) \<le> \<nu> A"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   344
    from this[THEN bspec, OF top] show "positive_integral g \<le> \<nu> (space M)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   345
      by (simp cong: positive_integral_cong)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   346
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   347
  hence "?y \<noteq> \<omega>" using M'.finite_measure_of_space by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   348
  from SUPR_countable_SUPR[OF this `G \<noteq> {}`] guess ys .. note ys = this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   349
  hence "\<forall>n. \<exists>g. g\<in>G \<and> positive_integral g = ys n"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   350
  proof safe
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   351
    fix n assume "range ys \<subseteq> positive_integral ` G"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   352
    hence "ys n \<in> positive_integral ` G" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   353
    thus "\<exists>g. g\<in>G \<and> positive_integral g = ys n" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   354
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   355
  from choice[OF this] obtain gs where "\<And>i. gs i \<in> G" "\<And>n. positive_integral (gs n) = ys n" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   356
  hence y_eq: "?y = (SUP i. positive_integral (gs i))" using ys by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   357
  let "?g i x" = "Max ((\<lambda>n. gs n x) ` {..i})"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   358
  def f \<equiv> "SUP i. ?g i"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   359
  have gs_not_empty: "\<And>i. (\<lambda>n. gs n x) ` {..i} \<noteq> {}" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   360
  { fix i have "?g i \<in> G"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   361
    proof (induct i)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   362
      case 0 thus ?case by simp fact
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   363
    next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   364
      case (Suc i)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   365
      with Suc gs_not_empty `gs (Suc i) \<in> G` show ?case
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   366
        by (auto simp add: atMost_Suc intro!: max_in_G)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   367
    qed }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   368
  note g_in_G = this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   369
  have "\<And>x. \<forall>i. ?g i x \<le> ?g (Suc i) x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   370
    using gs_not_empty by (simp add: atMost_Suc)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   371
  hence isoton_g: "?g \<up> f" by (simp add: isoton_def le_fun_def f_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   372
  from SUP_in_G[OF this g_in_G] have "f \<in> G" .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   373
  hence [simp, intro]: "f \<in> borel_measurable M" unfolding G_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   374
  have "(\<lambda>i. positive_integral (?g i)) \<up> positive_integral f"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   375
    using isoton_g g_in_G by (auto intro!: positive_integral_isoton simp: G_def f_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   376
  hence "positive_integral f = (SUP i. positive_integral (?g i))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   377
    unfolding isoton_def by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   378
  also have "\<dots> = ?y"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   379
  proof (rule antisym)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   380
    show "(SUP i. positive_integral (?g i)) \<le> ?y"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   381
      using g_in_G by (auto intro!: exI Sup_mono simp: SUPR_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   382
    show "?y \<le> (SUP i. positive_integral (?g i))" unfolding y_eq
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   383
      by (auto intro!: SUP_mono positive_integral_mono Max_ge)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   384
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   385
  finally have int_f_eq_y: "positive_integral f = ?y" .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   386
  let "?t A" = "\<nu> A - positive_integral (\<lambda>x. f x * indicator A x)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   387
  have "finite_measure M ?t"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   388
  proof
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   389
    show "?t {} = 0" by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   390
    show "?t (space M) \<noteq> \<omega>" using M'.finite_measure by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   391
    show "countably_additive M ?t" unfolding countably_additive_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   392
    proof safe
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   393
      fix A :: "nat \<Rightarrow> 'a set"  assume A: "range A \<subseteq> sets M" "disjoint_family A"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   394
      have "(\<Sum>\<^isub>\<infinity> n. positive_integral (\<lambda>x. f x * indicator (A n) x))
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   395
        = positive_integral (\<lambda>x. (\<Sum>\<^isub>\<infinity>n. f x * indicator (A n) x))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   396
        using `range A \<subseteq> sets M`
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   397
        by (rule_tac positive_integral_psuminf[symmetric]) (auto intro!: borel_measurable_indicator)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   398
      also have "\<dots> = positive_integral (\<lambda>x. f x * indicator (\<Union>n. A n) x)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   399
        apply (rule positive_integral_cong)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   400
        apply (subst psuminf_cmult_right)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   401
        unfolding psuminf_indicator[OF `disjoint_family A`] ..
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   402
      finally have "(\<Sum>\<^isub>\<infinity> n. positive_integral (\<lambda>x. f x * indicator (A n) x))
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   403
        = positive_integral (\<lambda>x. f x * indicator (\<Union>n. A n) x)" .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   404
      moreover have "(\<Sum>\<^isub>\<infinity>n. \<nu> (A n)) = \<nu> (\<Union>n. A n)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   405
        using M'.measure_countably_additive A by (simp add: comp_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   406
      moreover have "\<And>i. positive_integral (\<lambda>x. f x * indicator (A i) x) \<le> \<nu> (A i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   407
          using A `f \<in> G` unfolding G_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   408
      moreover have v_fin: "\<nu> (\<Union>i. A i) \<noteq> \<omega>" using M'.finite_measure A by (simp add: countable_UN)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   409
      moreover {
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   410
        have "positive_integral (\<lambda>x. f x * indicator (\<Union>i. A i) x) \<le> \<nu> (\<Union>i. A i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   411
          using A `f \<in> G` unfolding G_def by (auto simp: countable_UN)
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40871
diff changeset
   412
        also have "\<nu> (\<Union>i. A i) < \<omega>" using v_fin by (simp add: pextreal_less_\<omega>)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   413
        finally have "positive_integral (\<lambda>x. f x * indicator (\<Union>i. A i) x) \<noteq> \<omega>"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40871
diff changeset
   414
          by (simp add: pextreal_less_\<omega>) }
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   415
      ultimately
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   416
      show "(\<Sum>\<^isub>\<infinity> n. ?t (A n)) = ?t (\<Union>i. A i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   417
        apply (subst psuminf_minus) by simp_all
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   418
    qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   419
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   420
  then interpret M: finite_measure M ?t .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   421
  have ac: "absolutely_continuous ?t" using `absolutely_continuous \<nu>` unfolding absolutely_continuous_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   422
  have upper_bound: "\<forall>A\<in>sets M. ?t A \<le> 0"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   423
  proof (rule ccontr)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   424
    assume "\<not> ?thesis"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   425
    then obtain A where A: "A \<in> sets M" and pos: "0 < ?t A"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   426
      by (auto simp: not_le)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   427
    note pos
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   428
    also have "?t A \<le> ?t (space M)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   429
      using M.measure_mono[of A "space M"] A sets_into_space by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   430
    finally have pos_t: "0 < ?t (space M)" by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   431
    moreover
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   432
    hence pos_M: "0 < \<mu> (space M)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   433
      using ac top unfolding absolutely_continuous_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   434
    moreover
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   435
    have "positive_integral (\<lambda>x. f x * indicator (space M) x) \<le> \<nu> (space M)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   436
      using `f \<in> G` unfolding G_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   437
    hence "positive_integral (\<lambda>x. f x * indicator (space M) x) \<noteq> \<omega>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   438
      using M'.finite_measure_of_space by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   439
    moreover
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   440
    def b \<equiv> "?t (space M) / \<mu> (space M) / 2"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   441
    ultimately have b: "b \<noteq> 0" "b \<noteq> \<omega>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   442
      using M'.finite_measure_of_space
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40871
diff changeset
   443
      by (auto simp: pextreal_inverse_eq_0 finite_measure_of_space)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   444
    have "finite_measure M (\<lambda>A. b * \<mu> A)" (is "finite_measure M ?b")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   445
    proof
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   446
      show "?b {} = 0" by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   447
      show "?b (space M) \<noteq> \<omega>" using finite_measure_of_space b by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   448
      show "countably_additive M ?b"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   449
        unfolding countably_additive_def psuminf_cmult_right
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   450
        using measure_countably_additive by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   451
    qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   452
    from M.Radon_Nikodym_aux[OF this]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   453
    obtain A0 where "A0 \<in> sets M" and
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   454
      space_less_A0: "real (?t (space M)) - real (b * \<mu> (space M)) \<le> real (?t A0) - real (b * \<mu> A0)" and
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   455
      *: "\<And>B. \<lbrakk> B \<in> sets M ; B \<subseteq> A0 \<rbrakk> \<Longrightarrow> 0 \<le> real (?t B) - real (b * \<mu> B)" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   456
    { fix B assume "B \<in> sets M" "B \<subseteq> A0"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   457
      with *[OF this] have "b * \<mu> B \<le> ?t B"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   458
        using M'.finite_measure b finite_measure
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   459
        by (cases "b * \<mu> B", cases "?t B") (auto simp: field_simps) }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   460
    note bM_le_t = this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   461
    let "?f0 x" = "f x + b * indicator A0 x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   462
    { fix A assume A: "A \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   463
      hence "A \<inter> A0 \<in> sets M" using `A0 \<in> sets M` by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   464
      have "positive_integral (\<lambda>x. ?f0 x  * indicator A x) =
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   465
        positive_integral (\<lambda>x. f x * indicator A x + b * indicator (A \<inter> A0) x)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   466
        by (auto intro!: positive_integral_cong simp: field_simps indicator_inter_arith)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   467
      hence "positive_integral (\<lambda>x. ?f0 x * indicator A x) =
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   468
          positive_integral (\<lambda>x. f x * indicator A x) + b * \<mu> (A \<inter> A0)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   469
        using `A0 \<in> sets M` `A \<inter> A0 \<in> sets M` A
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   470
        by (simp add: borel_measurable_indicator positive_integral_add positive_integral_cmult_indicator) }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   471
    note f0_eq = this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   472
    { fix A assume A: "A \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   473
      hence "A \<inter> A0 \<in> sets M" using `A0 \<in> sets M` by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   474
      have f_le_v: "positive_integral (\<lambda>x. f x * indicator A x) \<le> \<nu> A"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   475
        using `f \<in> G` A unfolding G_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   476
      note f0_eq[OF A]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   477
      also have "positive_integral (\<lambda>x. f x * indicator A x) + b * \<mu> (A \<inter> A0) \<le>
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   478
          positive_integral (\<lambda>x. f x * indicator A x) + ?t (A \<inter> A0)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   479
        using bM_le_t[OF `A \<inter> A0 \<in> sets M`] `A \<in> sets M` `A0 \<in> sets M`
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   480
        by (auto intro!: add_left_mono)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   481
      also have "\<dots> \<le> positive_integral (\<lambda>x. f x * indicator A x) + ?t A"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   482
        using M.measure_mono[simplified, OF _ `A \<inter> A0 \<in> sets M` `A \<in> sets M`]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   483
        by (auto intro!: add_left_mono)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   484
      also have "\<dots> \<le> \<nu> A"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   485
        using f_le_v M'.finite_measure[simplified, OF `A \<in> sets M`]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   486
        by (cases "positive_integral (\<lambda>x. f x * indicator A x)", cases "\<nu> A", auto)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   487
      finally have "positive_integral (\<lambda>x. ?f0 x * indicator A x) \<le> \<nu> A" . }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   488
    hence "?f0 \<in> G" using `A0 \<in> sets M` unfolding G_def
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40871
diff changeset
   489
      by (auto intro!: borel_measurable_indicator borel_measurable_pextreal_add borel_measurable_pextreal_times)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   490
    have real: "?t (space M) \<noteq> \<omega>" "?t A0 \<noteq> \<omega>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   491
      "b * \<mu> (space M) \<noteq> \<omega>" "b * \<mu> A0 \<noteq> \<omega>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   492
      using `A0 \<in> sets M` b
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   493
        finite_measure[of A0] M.finite_measure[of A0]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   494
        finite_measure_of_space M.finite_measure_of_space
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   495
      by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   496
    have int_f_finite: "positive_integral f \<noteq> \<omega>"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40871
diff changeset
   497
      using M'.finite_measure_of_space pos_t unfolding pextreal_zero_less_diff_iff
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   498
      by (auto cong: positive_integral_cong)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   499
    have "?t (space M) > b * \<mu> (space M)" unfolding b_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   500
      apply (simp add: field_simps)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   501
      apply (subst mult_assoc[symmetric])
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40871
diff changeset
   502
      apply (subst pextreal_mult_inverse)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   503
      using finite_measure_of_space M'.finite_measure_of_space pos_t pos_M
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40871
diff changeset
   504
      using pextreal_mult_strict_right_mono[of "Real (1/2)" 1 "?t (space M)"]
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   505
      by simp_all
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   506
    hence  "0 < ?t (space M) - b * \<mu> (space M)"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40871
diff changeset
   507
      by (simp add: pextreal_zero_less_diff_iff)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   508
    also have "\<dots> \<le> ?t A0 - b * \<mu> A0"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40871
diff changeset
   509
      using space_less_A0 pos_M pos_t b real[unfolded pextreal_noteq_omega_Ex] by auto
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40871
diff changeset
   510
    finally have "b * \<mu> A0 < ?t A0" unfolding pextreal_zero_less_diff_iff .
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   511
    hence "0 < ?t A0" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   512
    hence "0 < \<mu> A0" using ac unfolding absolutely_continuous_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   513
      using `A0 \<in> sets M` by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   514
    hence "0 < b * \<mu> A0" using b by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   515
    from int_f_finite this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   516
    have "?y + 0 < positive_integral f + b * \<mu> A0" unfolding int_f_eq_y
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40871
diff changeset
   517
      by (rule pextreal_less_add)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   518
    also have "\<dots> = positive_integral ?f0" using f0_eq[OF top] `A0 \<in> sets M` sets_into_space
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   519
      by (simp cong: positive_integral_cong)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   520
    finally have "?y < positive_integral ?f0" by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   521
    moreover from `?f0 \<in> G` have "positive_integral ?f0 \<le> ?y" by (auto intro!: le_SUPI)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   522
    ultimately show False by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   523
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   524
  show ?thesis
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   525
  proof (safe intro!: bexI[of _ f])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   526
    fix A assume "A\<in>sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   527
    show "\<nu> A = positive_integral (\<lambda>x. f x * indicator A x)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   528
    proof (rule antisym)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   529
      show "positive_integral (\<lambda>x. f x * indicator A x) \<le> \<nu> A"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   530
        using `f \<in> G` `A \<in> sets M` unfolding G_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   531
      show "\<nu> A \<le> positive_integral (\<lambda>x. f x * indicator A x)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   532
        using upper_bound[THEN bspec, OF `A \<in> sets M`]
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40871
diff changeset
   533
         by (simp add: pextreal_zero_le_diff)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   534
    qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   535
  qed simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   536
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   537
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   538
lemma (in finite_measure) split_space_into_finite_sets_and_rest:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   539
  assumes "measure_space M \<nu>"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   540
  assumes ac: "absolutely_continuous \<nu>"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   541
  shows "\<exists>\<Omega>0\<in>sets M. \<exists>\<Omega>::nat\<Rightarrow>'a set. disjoint_family \<Omega> \<and> range \<Omega> \<subseteq> sets M \<and> \<Omega>0 = space M - (\<Union>i. \<Omega> i) \<and>
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   542
    (\<forall>A\<in>sets M. A \<subseteq> \<Omega>0 \<longrightarrow> (\<mu> A = 0 \<and> \<nu> A = 0) \<or> (\<mu> A > 0 \<and> \<nu> A = \<omega>)) \<and>
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   543
    (\<forall>i. \<nu> (\<Omega> i) \<noteq> \<omega>)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   544
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   545
  interpret v: measure_space M \<nu> by fact
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   546
  let ?Q = "{Q\<in>sets M. \<nu> Q \<noteq> \<omega>}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   547
  let ?a = "SUP Q:?Q. \<mu> Q"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   548
  have "{} \<in> ?Q" using v.empty_measure by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   549
  then have Q_not_empty: "?Q \<noteq> {}" by blast
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   550
  have "?a \<le> \<mu> (space M)" using sets_into_space
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   551
    by (auto intro!: SUP_leI measure_mono top)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   552
  then have "?a \<noteq> \<omega>" using finite_measure_of_space
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   553
    by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   554
  from SUPR_countable_SUPR[OF this Q_not_empty]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   555
  obtain Q'' where "range Q'' \<subseteq> \<mu> ` ?Q" and a: "?a = (SUP i::nat. Q'' i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   556
    by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   557
  then have "\<forall>i. \<exists>Q'. Q'' i = \<mu> Q' \<and> Q' \<in> ?Q" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   558
  from choice[OF this] obtain Q' where Q': "\<And>i. Q'' i = \<mu> (Q' i)" "\<And>i. Q' i \<in> ?Q"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   559
    by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   560
  then have a_Lim: "?a = (SUP i::nat. \<mu> (Q' i))" using a by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   561
  let "?O n" = "\<Union>i\<le>n. Q' i"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   562
  have Union: "(SUP i. \<mu> (?O i)) = \<mu> (\<Union>i. ?O i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   563
  proof (rule continuity_from_below[of ?O])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   564
    show "range ?O \<subseteq> sets M" using Q' by (auto intro!: finite_UN)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   565
    show "\<And>i. ?O i \<subseteq> ?O (Suc i)" by fastsimp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   566
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   567
  have Q'_sets: "\<And>i. Q' i \<in> sets M" using Q' by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   568
  have O_sets: "\<And>i. ?O i \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   569
     using Q' by (auto intro!: finite_UN Un)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   570
  then have O_in_G: "\<And>i. ?O i \<in> ?Q"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   571
  proof (safe del: notI)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   572
    fix i have "Q' ` {..i} \<subseteq> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   573
      using Q' by (auto intro: finite_UN)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   574
    with v.measure_finitely_subadditive[of "{.. i}" Q']
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   575
    have "\<nu> (?O i) \<le> (\<Sum>i\<le>i. \<nu> (Q' i))" by auto
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40871
diff changeset
   576
    also have "\<dots> < \<omega>" unfolding setsum_\<omega> pextreal_less_\<omega> using Q' by auto
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40871
diff changeset
   577
    finally show "\<nu> (?O i) \<noteq> \<omega>" unfolding pextreal_less_\<omega> by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   578
  qed auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   579
  have O_mono: "\<And>n. ?O n \<subseteq> ?O (Suc n)" by fastsimp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   580
  have a_eq: "?a = \<mu> (\<Union>i. ?O i)" unfolding Union[symmetric]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   581
  proof (rule antisym)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   582
    show "?a \<le> (SUP i. \<mu> (?O i))" unfolding a_Lim
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   583
      using Q' by (auto intro!: SUP_mono measure_mono finite_UN)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   584
    show "(SUP i. \<mu> (?O i)) \<le> ?a" unfolding SUPR_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   585
    proof (safe intro!: Sup_mono, unfold bex_simps)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   586
      fix i
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   587
      have *: "(\<Union>Q' ` {..i}) = ?O i" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   588
      then show "\<exists>x. (x \<in> sets M \<and> \<nu> x \<noteq> \<omega>) \<and>
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   589
        \<mu> (\<Union>Q' ` {..i}) \<le> \<mu> x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   590
        using O_in_G[of i] by (auto intro!: exI[of _ "?O i"])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   591
    qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   592
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   593
  let "?O_0" = "(\<Union>i. ?O i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   594
  have "?O_0 \<in> sets M" using Q' by auto
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   595
  def Q \<equiv> "\<lambda>i. case i of 0 \<Rightarrow> Q' 0 | Suc n \<Rightarrow> ?O (Suc n) - ?O n"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   596
  { fix i have "Q i \<in> sets M" unfolding Q_def using Q'[of 0] by (cases i) (auto intro: O_sets) }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   597
  note Q_sets = this
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   598
  show ?thesis
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   599
  proof (intro bexI exI conjI ballI impI allI)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   600
    show "disjoint_family Q"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   601
      by (fastsimp simp: disjoint_family_on_def Q_def
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   602
        split: nat.split_asm)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   603
    show "range Q \<subseteq> sets M"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   604
      using Q_sets by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   605
    { fix A assume A: "A \<in> sets M" "A \<subseteq> space M - ?O_0"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   606
      show "\<mu> A = 0 \<and> \<nu> A = 0 \<or> 0 < \<mu> A \<and> \<nu> A = \<omega>"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   607
      proof (rule disjCI, simp)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   608
        assume *: "0 < \<mu> A \<longrightarrow> \<nu> A \<noteq> \<omega>"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   609
        show "\<mu> A = 0 \<and> \<nu> A = 0"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   610
        proof cases
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   611
          assume "\<mu> A = 0" moreover with ac A have "\<nu> A = 0"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   612
            unfolding absolutely_continuous_def by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   613
          ultimately show ?thesis by simp
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   614
        next
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   615
          assume "\<mu> A \<noteq> 0" with * have "\<nu> A \<noteq> \<omega>" by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   616
          with A have "\<mu> ?O_0 + \<mu> A = \<mu> (?O_0 \<union> A)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   617
            using Q' by (auto intro!: measure_additive countable_UN)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   618
          also have "\<dots> = (SUP i. \<mu> (?O i \<union> A))"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   619
          proof (rule continuity_from_below[of "\<lambda>i. ?O i \<union> A", symmetric, simplified])
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   620
            show "range (\<lambda>i. ?O i \<union> A) \<subseteq> sets M"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   621
              using `\<nu> A \<noteq> \<omega>` O_sets A by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   622
          qed fastsimp
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   623
          also have "\<dots> \<le> ?a"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   624
          proof (safe intro!: SUPR_bound)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   625
            fix i have "?O i \<union> A \<in> ?Q"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   626
            proof (safe del: notI)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   627
              show "?O i \<union> A \<in> sets M" using O_sets A by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   628
              from O_in_G[of i]
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   629
              moreover have "\<nu> (?O i \<union> A) \<le> \<nu> (?O i) + \<nu> A"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   630
                using v.measure_subadditive[of "?O i" A] A O_sets by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   631
              ultimately show "\<nu> (?O i \<union> A) \<noteq> \<omega>"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   632
                using `\<nu> A \<noteq> \<omega>` by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   633
            qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   634
            then show "\<mu> (?O i \<union> A) \<le> ?a" by (rule le_SUPI)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   635
          qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   636
          finally have "\<mu> A = 0" unfolding a_eq using finite_measure[OF `?O_0 \<in> sets M`]
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40871
diff changeset
   637
            by (cases "\<mu> A") (auto simp: pextreal_noteq_omega_Ex)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   638
          with `\<mu> A \<noteq> 0` show ?thesis by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   639
        qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   640
      qed }
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   641
    { fix i show "\<nu> (Q i) \<noteq> \<omega>"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   642
      proof (cases i)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   643
        case 0 then show ?thesis
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   644
          unfolding Q_def using Q'[of 0] by simp
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   645
      next
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   646
        case (Suc n)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   647
        then show ?thesis unfolding Q_def
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   648
          using `?O n \<in> ?Q` `?O (Suc n) \<in> ?Q` O_mono
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   649
          using v.measure_Diff[of "?O n" "?O (Suc n)"] by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   650
      qed }
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   651
    show "space M - ?O_0 \<in> sets M" using Q'_sets by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   652
    { fix j have "(\<Union>i\<le>j. ?O i) = (\<Union>i\<le>j. Q i)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   653
      proof (induct j)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   654
        case 0 then show ?case by (simp add: Q_def)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   655
      next
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   656
        case (Suc j)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   657
        have eq: "\<And>j. (\<Union>i\<le>j. ?O i) = (\<Union>i\<le>j. Q' i)" by fastsimp
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   658
        have "{..j} \<union> {..Suc j} = {..Suc j}" by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   659
        then have "(\<Union>i\<le>Suc j. Q' i) = (\<Union>i\<le>j. Q' i) \<union> Q (Suc j)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   660
          by (simp add: UN_Un[symmetric] Q_def del: UN_Un)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   661
        then show ?case using Suc by (auto simp add: eq atMost_Suc)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   662
      qed }
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   663
    then have "(\<Union>j. (\<Union>i\<le>j. ?O i)) = (\<Union>j. (\<Union>i\<le>j. Q i))" by simp
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   664
    then show "space M - ?O_0 = space M - (\<Union>i. Q i)" by fastsimp
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   665
  qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   666
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   667
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   668
lemma (in finite_measure) Radon_Nikodym_finite_measure_infinite:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   669
  assumes "measure_space M \<nu>"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   670
  assumes "absolutely_continuous \<nu>"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   671
  shows "\<exists>f \<in> borel_measurable M. \<forall>A\<in>sets M. \<nu> A = positive_integral (\<lambda>x. f x * indicator A x)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   672
proof -
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   673
  interpret v: measure_space M \<nu> by fact
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   674
  from split_space_into_finite_sets_and_rest[OF assms]
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   675
  obtain Q0 and Q :: "nat \<Rightarrow> 'a set"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   676
    where Q: "disjoint_family Q" "range Q \<subseteq> sets M"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   677
    and Q0: "Q0 \<in> sets M" "Q0 = space M - (\<Union>i. Q i)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   678
    and in_Q0: "\<And>A. A \<in> sets M \<Longrightarrow> A \<subseteq> Q0 \<Longrightarrow> \<mu> A = 0 \<and> \<nu> A = 0 \<or> 0 < \<mu> A \<and> \<nu> A = \<omega>"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   679
    and Q_fin: "\<And>i. \<nu> (Q i) \<noteq> \<omega>" by force
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   680
  from Q have Q_sets: "\<And>i. Q i \<in> sets M" by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   681
  have "\<forall>i. \<exists>f. f\<in>borel_measurable M \<and> (\<forall>A\<in>sets M.
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   682
    \<nu> (Q i \<inter> A) = positive_integral (\<lambda>x. f x * indicator (Q i \<inter> A) x))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   683
  proof
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   684
    fix i
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40871
diff changeset
   685
    have indicator_eq: "\<And>f x A. (f x :: pextreal) * indicator (Q i \<inter> A) x * indicator (Q i) x
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   686
      = (f x * indicator (Q i) x) * indicator A x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   687
      unfolding indicator_def by auto
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38656
diff changeset
   688
    have fm: "finite_measure (restricted_space (Q i)) \<mu>"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   689
      (is "finite_measure ?R \<mu>") by (rule restricted_finite_measure[OF Q_sets[of i]])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   690
    then interpret R: finite_measure ?R .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   691
    have fmv: "finite_measure ?R \<nu>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   692
      unfolding finite_measure_def finite_measure_axioms_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   693
    proof
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   694
      show "measure_space ?R \<nu>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   695
        using v.restricted_measure_space Q_sets[of i] by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   696
      show "\<nu>  (space ?R) \<noteq> \<omega>"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   697
        using Q_fin by simp
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   698
    qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   699
    have "R.absolutely_continuous \<nu>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   700
      using `absolutely_continuous \<nu>` `Q i \<in> sets M`
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   701
      by (auto simp: R.absolutely_continuous_def absolutely_continuous_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   702
    from finite_measure.Radon_Nikodym_finite_measure[OF fm fmv this]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   703
    obtain f where f: "(\<lambda>x. f x * indicator (Q i) x) \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   704
      and f_int: "\<And>A. A\<in>sets M \<Longrightarrow> \<nu> (Q i \<inter> A) = positive_integral (\<lambda>x. (f x * indicator (Q i) x) * indicator A x)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   705
      unfolding Bex_def borel_measurable_restricted[OF `Q i \<in> sets M`]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   706
        positive_integral_restricted[OF `Q i \<in> sets M`] by (auto simp: indicator_eq)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   707
    then show "\<exists>f. f\<in>borel_measurable M \<and> (\<forall>A\<in>sets M.
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   708
      \<nu> (Q i \<inter> A) = positive_integral (\<lambda>x. f x * indicator (Q i \<inter> A) x))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   709
      by (fastsimp intro!: exI[of _ "\<lambda>x. f x * indicator (Q i) x"] positive_integral_cong
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   710
          simp: indicator_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   711
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   712
  from choice[OF this] obtain f where borel: "\<And>i. f i \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   713
    and f: "\<And>A i. A \<in> sets M \<Longrightarrow>
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   714
      \<nu> (Q i \<inter> A) = positive_integral (\<lambda>x. f i x * indicator (Q i \<inter> A) x)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   715
    by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   716
  let "?f x" =
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   717
    "(\<Sum>\<^isub>\<infinity> i. f i x * indicator (Q i) x) + \<omega> * indicator Q0 x"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   718
  show ?thesis
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   719
  proof (safe intro!: bexI[of _ ?f])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   720
    show "?f \<in> borel_measurable M"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40871
diff changeset
   721
      by (safe intro!: borel_measurable_psuminf borel_measurable_pextreal_times
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40871
diff changeset
   722
        borel_measurable_pextreal_add borel_measurable_indicator
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   723
        borel_measurable_const borel Q_sets Q0 Diff countable_UN)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   724
    fix A assume "A \<in> sets M"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   725
    have *:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   726
      "\<And>x i. indicator A x * (f i x * indicator (Q i) x) =
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   727
        f i x * indicator (Q i \<inter> A) x"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40871
diff changeset
   728
      "\<And>x i. (indicator A x * indicator Q0 x :: pextreal) =
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   729
        indicator (Q0 \<inter> A) x" by (auto simp: indicator_def)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   730
    have "positive_integral (\<lambda>x. ?f x * indicator A x) =
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   731
      (\<Sum>\<^isub>\<infinity> i. \<nu> (Q i \<inter> A)) + \<omega> * \<mu> (Q0 \<inter> A)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   732
      unfolding f[OF `A \<in> sets M`]
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40871
diff changeset
   733
      apply (simp del: pextreal_times(2) add: field_simps *)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   734
      apply (subst positive_integral_add)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   735
      apply (fastsimp intro: Q0 `A \<in> sets M`)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   736
      apply (fastsimp intro: Q_sets `A \<in> sets M` borel_measurable_psuminf borel)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   737
      apply (subst positive_integral_cmult_indicator)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   738
      apply (fastsimp intro: Q0 `A \<in> sets M`)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   739
      unfolding psuminf_cmult_right[symmetric]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   740
      apply (subst positive_integral_psuminf)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   741
      apply (fastsimp intro: `A \<in> sets M` Q_sets borel)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   742
      apply (simp add: *)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   743
      done
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   744
    moreover have "(\<Sum>\<^isub>\<infinity>i. \<nu> (Q i \<inter> A)) = \<nu> ((\<Union>i. Q i) \<inter> A)"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   745
      using Q Q_sets `A \<in> sets M`
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   746
      by (intro v.measure_countably_additive[of "\<lambda>i. Q i \<inter> A", unfolded comp_def, simplified])
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   747
         (auto simp: disjoint_family_on_def)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   748
    moreover have "\<omega> * \<mu> (Q0 \<inter> A) = \<nu> (Q0 \<inter> A)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   749
    proof -
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   750
      have "Q0 \<inter> A \<in> sets M" using Q0(1) `A \<in> sets M` by blast
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   751
      from in_Q0[OF this] show ?thesis by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   752
    qed
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   753
    moreover have "Q0 \<inter> A \<in> sets M" "((\<Union>i. Q i) \<inter> A) \<in> sets M"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   754
      using Q_sets `A \<in> sets M` Q0(1) by (auto intro!: countable_UN)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   755
    moreover have "((\<Union>i. Q i) \<inter> A) \<union> (Q0 \<inter> A) = A" "((\<Union>i. Q i) \<inter> A) \<inter> (Q0 \<inter> A) = {}"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   756
      using `A \<in> sets M` sets_into_space Q0 by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   757
    ultimately show "\<nu> A = positive_integral (\<lambda>x. ?f x * indicator A x)"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   758
      using v.measure_additive[simplified, of "(\<Union>i. Q i) \<inter> A" "Q0 \<inter> A"]
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   759
      by simp
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   760
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   761
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   762
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   763
lemma (in sigma_finite_measure) Radon_Nikodym:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   764
  assumes "measure_space M \<nu>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   765
  assumes "absolutely_continuous \<nu>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   766
  shows "\<exists>f \<in> borel_measurable M. \<forall>A\<in>sets M. \<nu> A = positive_integral (\<lambda>x. f x * indicator A x)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   767
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   768
  from Ex_finite_integrable_function
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   769
  obtain h where finite: "positive_integral h \<noteq> \<omega>" and
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   770
    borel: "h \<in> borel_measurable M" and
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   771
    pos: "\<And>x. x \<in> space M \<Longrightarrow> 0 < h x" and
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   772
    "\<And>x. x \<in> space M \<Longrightarrow> h x < \<omega>" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   773
  let "?T A" = "positive_integral (\<lambda>x. h x * indicator A x)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   774
  from measure_space_density[OF borel] finite
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   775
  interpret T: finite_measure M ?T
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   776
    unfolding finite_measure_def finite_measure_axioms_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   777
    by (simp cong: positive_integral_cong)
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40871
diff changeset
   778
  have "\<And>N. N \<in> sets M \<Longrightarrow> {x \<in> space M. h x \<noteq> 0 \<and> indicator N x \<noteq> (0::pextreal)} = N"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   779
    using sets_into_space pos by (force simp: indicator_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   780
  then have "T.absolutely_continuous \<nu>" using assms(2) borel
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   781
    unfolding T.absolutely_continuous_def absolutely_continuous_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   782
    by (fastsimp simp: borel_measurable_indicator positive_integral_0_iff)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   783
  from T.Radon_Nikodym_finite_measure_infinite[simplified, OF assms(1) this]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   784
  obtain f where f_borel: "f \<in> borel_measurable M" and
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   785
    fT: "\<And>A. A \<in> sets M \<Longrightarrow> \<nu> A = T.positive_integral (\<lambda>x. f x * indicator A x)" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   786
  show ?thesis
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   787
  proof (safe intro!: bexI[of _ "\<lambda>x. h x * f x"])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   788
    show "(\<lambda>x. h x * f x) \<in> borel_measurable M"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40871
diff changeset
   789
      using borel f_borel by (auto intro: borel_measurable_pextreal_times)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   790
    fix A assume "A \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   791
    then have "(\<lambda>x. f x * indicator A x) \<in> borel_measurable M"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40871
diff changeset
   792
      using f_borel by (auto intro: borel_measurable_pextreal_times borel_measurable_indicator)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   793
    from positive_integral_translated_density[OF borel this]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   794
    show "\<nu> A = positive_integral (\<lambda>x. h x * f x * indicator A x)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   795
      unfolding fT[OF `A \<in> sets M`] by (simp add: field_simps)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   796
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   797
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   798
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   799
section "Uniqueness of densities"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   800
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   801
lemma (in measure_space) finite_density_unique:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   802
  assumes borel: "f \<in> borel_measurable M" "g \<in> borel_measurable M"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   803
  and fin: "positive_integral f < \<omega>"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   804
  shows "(\<forall>A\<in>sets M. positive_integral (\<lambda>x. f x * indicator A x) = positive_integral (\<lambda>x. g x * indicator A x))
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   805
    \<longleftrightarrow> (AE x. f x = g x)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   806
    (is "(\<forall>A\<in>sets M. ?P f A = ?P g A) \<longleftrightarrow> _")
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   807
proof (intro iffI ballI)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   808
  fix A assume eq: "AE x. f x = g x"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   809
  show "?P f A = ?P g A"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   810
    by (rule positive_integral_cong_AE[OF AE_mp[OF eq]]) simp
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   811
next
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   812
  assume eq: "\<forall>A\<in>sets M. ?P f A = ?P g A"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   813
  from this[THEN bspec, OF top] fin
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   814
  have g_fin: "positive_integral g < \<omega>" by (simp cong: positive_integral_cong)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   815
  { fix f g assume borel: "f \<in> borel_measurable M" "g \<in> borel_measurable M"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   816
      and g_fin: "positive_integral g < \<omega>" and eq: "\<forall>A\<in>sets M. ?P f A = ?P g A"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   817
    let ?N = "{x\<in>space M. g x < f x}"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   818
    have N: "?N \<in> sets M" using borel by simp
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   819
    have "?P (\<lambda>x. (f x - g x)) ?N = positive_integral (\<lambda>x. f x * indicator ?N x - g x * indicator ?N x)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   820
      by (auto intro!: positive_integral_cong simp: indicator_def)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   821
    also have "\<dots> = ?P f ?N - ?P g ?N"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   822
    proof (rule positive_integral_diff)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   823
      show "(\<lambda>x. f x * indicator ?N x) \<in> borel_measurable M" "(\<lambda>x. g x * indicator ?N x) \<in> borel_measurable M"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   824
        using borel N by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   825
      have "?P g ?N \<le> positive_integral g"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   826
        by (auto intro!: positive_integral_mono simp: indicator_def)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   827
      then show "?P g ?N \<noteq> \<omega>" using g_fin by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   828
      fix x assume "x \<in> space M"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   829
      show "g x * indicator ?N x \<le> f x * indicator ?N x"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   830
        by (auto simp: indicator_def)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   831
    qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   832
    also have "\<dots> = 0"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   833
      using eq[THEN bspec, OF N] by simp
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   834
    finally have "\<mu> {x\<in>space M. (f x - g x) * indicator ?N x \<noteq> 0} = 0"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   835
      using borel N by (subst (asm) positive_integral_0_iff) auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   836
    moreover have "{x\<in>space M. (f x - g x) * indicator ?N x \<noteq> 0} = ?N"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40871
diff changeset
   837
      by (auto simp: pextreal_zero_le_diff)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   838
    ultimately have "?N \<in> null_sets" using N by simp }
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   839
  from this[OF borel g_fin eq] this[OF borel(2,1) fin]
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   840
  have "{x\<in>space M. g x < f x} \<union> {x\<in>space M. f x < g x} \<in> null_sets"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   841
    using eq by (intro null_sets_Un) auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   842
  also have "{x\<in>space M. g x < f x} \<union> {x\<in>space M. f x < g x} = {x\<in>space M. f x \<noteq> g x}"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   843
    by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   844
  finally show "AE x. f x = g x"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   845
    unfolding almost_everywhere_def by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   846
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   847
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   848
lemma (in finite_measure) density_unique_finite_measure:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   849
  assumes borel: "f \<in> borel_measurable M" "f' \<in> borel_measurable M"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   850
  assumes f: "\<And>A. A \<in> sets M \<Longrightarrow> positive_integral (\<lambda>x. f x * indicator A x) = positive_integral (\<lambda>x. f' x * indicator A x)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   851
    (is "\<And>A. A \<in> sets M \<Longrightarrow> ?P f A = ?P f' A")
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   852
  shows "AE x. f x = f' x"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   853
proof -
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   854
  let "?\<nu> A" = "?P f A" and "?\<nu>' A" = "?P f' A"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   855
  let "?f A x" = "f x * indicator A x" and "?f' A x" = "f' x * indicator A x"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   856
  interpret M: measure_space M ?\<nu>
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   857
    using borel(1) by (rule measure_space_density)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   858
  have ac: "absolutely_continuous ?\<nu>"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   859
    using f by (rule density_is_absolutely_continuous)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   860
  from split_space_into_finite_sets_and_rest[OF `measure_space M ?\<nu>` ac]
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   861
  obtain Q0 and Q :: "nat \<Rightarrow> 'a set"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   862
    where Q: "disjoint_family Q" "range Q \<subseteq> sets M"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   863
    and Q0: "Q0 \<in> sets M" "Q0 = space M - (\<Union>i. Q i)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   864
    and in_Q0: "\<And>A. A \<in> sets M \<Longrightarrow> A \<subseteq> Q0 \<Longrightarrow> \<mu> A = 0 \<and> ?\<nu> A = 0 \<or> 0 < \<mu> A \<and> ?\<nu> A = \<omega>"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   865
    and Q_fin: "\<And>i. ?\<nu> (Q i) \<noteq> \<omega>" by force
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   866
  from Q have Q_sets: "\<And>i. Q i \<in> sets M" by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   867
  let ?N = "{x\<in>space M. f x \<noteq> f' x}"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   868
  have "?N \<in> sets M" using borel by auto
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40871
diff changeset
   869
  have *: "\<And>i x A. \<And>y::pextreal. y * indicator (Q i) x * indicator A x = y * indicator (Q i \<inter> A) x"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   870
    unfolding indicator_def by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   871
  have 1: "\<forall>i. AE x. ?f (Q i) x = ?f' (Q i) x"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   872
    using borel Q_fin Q
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   873
    by (intro finite_density_unique[THEN iffD1] allI)
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40871
diff changeset
   874
       (auto intro!: borel_measurable_pextreal_times f Int simp: *)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   875
  have 2: "AE x. ?f Q0 x = ?f' Q0 x"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   876
  proof (rule AE_I')
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40871
diff changeset
   877
    { fix f :: "'a \<Rightarrow> pextreal" assume borel: "f \<in> borel_measurable M"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   878
        and eq: "\<And>A. A \<in> sets M \<Longrightarrow> ?\<nu> A = positive_integral (\<lambda>x. f x * indicator A x)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   879
      let "?A i" = "Q0 \<inter> {x \<in> space M. f x < of_nat i}"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   880
      have "(\<Union>i. ?A i) \<in> null_sets"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   881
      proof (rule null_sets_UN)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   882
        fix i have "?A i \<in> sets M"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   883
          using borel Q0(1) by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   884
        have "?\<nu> (?A i) \<le> positive_integral (\<lambda>x. of_nat i * indicator (?A i) x)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   885
          unfolding eq[OF `?A i \<in> sets M`]
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   886
          by (auto intro!: positive_integral_mono simp: indicator_def)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   887
        also have "\<dots> = of_nat i * \<mu> (?A i)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   888
          using `?A i \<in> sets M` by (auto intro!: positive_integral_cmult_indicator)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   889
        also have "\<dots> < \<omega>"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   890
          using `?A i \<in> sets M`[THEN finite_measure] by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   891
        finally have "?\<nu> (?A i) \<noteq> \<omega>" by simp
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   892
        then show "?A i \<in> null_sets" using in_Q0[OF `?A i \<in> sets M`] `?A i \<in> sets M` by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   893
      qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   894
      also have "(\<Union>i. ?A i) = Q0 \<inter> {x\<in>space M. f x < \<omega>}"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   895
        by (auto simp: less_\<omega>_Ex_of_nat)
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40871
diff changeset
   896
      finally have "Q0 \<inter> {x\<in>space M. f x \<noteq> \<omega>} \<in> null_sets" by (simp add: pextreal_less_\<omega>) }
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   897
    from this[OF borel(1) refl] this[OF borel(2) f]
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   898
    have "Q0 \<inter> {x\<in>space M. f x \<noteq> \<omega>} \<in> null_sets" "Q0 \<inter> {x\<in>space M. f' x \<noteq> \<omega>} \<in> null_sets" by simp_all
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   899
    then show "(Q0 \<inter> {x\<in>space M. f x \<noteq> \<omega>}) \<union> (Q0 \<inter> {x\<in>space M. f' x \<noteq> \<omega>}) \<in> null_sets" by (rule null_sets_Un)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   900
    show "{x \<in> space M. ?f Q0 x \<noteq> ?f' Q0 x} \<subseteq>
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   901
      (Q0 \<inter> {x\<in>space M. f x \<noteq> \<omega>}) \<union> (Q0 \<inter> {x\<in>space M. f' x \<noteq> \<omega>})" by (auto simp: indicator_def)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   902
  qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   903
  have **: "\<And>x. (?f Q0 x = ?f' Q0 x) \<longrightarrow> (\<forall>i. ?f (Q i) x = ?f' (Q i) x) \<longrightarrow>
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   904
    ?f (space M) x = ?f' (space M) x"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   905
    by (auto simp: indicator_def Q0)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   906
  have 3: "AE x. ?f (space M) x = ?f' (space M) x"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   907
    by (rule AE_mp[OF 1[unfolded all_AE_countable] AE_mp[OF 2]]) (simp add: **)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   908
  then show "AE x. f x = f' x"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   909
    by (rule AE_mp) (auto intro!: AE_cong simp: indicator_def)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   910
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   911
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   912
lemma (in sigma_finite_measure) density_unique:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   913
  assumes borel: "f \<in> borel_measurable M" "f' \<in> borel_measurable M"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   914
  assumes f: "\<And>A. A \<in> sets M \<Longrightarrow> positive_integral (\<lambda>x. f x * indicator A x) = positive_integral (\<lambda>x. f' x * indicator A x)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   915
    (is "\<And>A. A \<in> sets M \<Longrightarrow> ?P f A = ?P f' A")
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   916
  shows "AE x. f x = f' x"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   917
proof -
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   918
  obtain h where h_borel: "h \<in> borel_measurable M"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   919
    and fin: "positive_integral h \<noteq> \<omega>" and pos: "\<And>x. x \<in> space M \<Longrightarrow> 0 < h x \<and> h x < \<omega>"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   920
    using Ex_finite_integrable_function by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   921
  interpret h: measure_space M "\<lambda>A. positive_integral (\<lambda>x. h x * indicator A x)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   922
    using h_borel by (rule measure_space_density)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   923
  interpret h: finite_measure M "\<lambda>A. positive_integral (\<lambda>x. h x * indicator A x)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   924
    by default (simp cong: positive_integral_cong add: fin)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   925
  interpret f: measure_space M "\<lambda>A. positive_integral (\<lambda>x. f x * indicator A x)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   926
    using borel(1) by (rule measure_space_density)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   927
  interpret f': measure_space M "\<lambda>A. positive_integral (\<lambda>x. f' x * indicator A x)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   928
    using borel(2) by (rule measure_space_density)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   929
  { fix A assume "A \<in> sets M"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40871
diff changeset
   930
    then have " {x \<in> space M. h x \<noteq> 0 \<and> indicator A x \<noteq> (0::pextreal)} = A"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   931
      using pos sets_into_space by (force simp: indicator_def)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   932
    then have "positive_integral (\<lambda>xa. h xa * indicator A xa) = 0 \<longleftrightarrow> A \<in> null_sets"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   933
      using h_borel `A \<in> sets M` by (simp add: positive_integral_0_iff) }
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   934
  note h_null_sets = this
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   935
  { fix A assume "A \<in> sets M"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   936
    have "positive_integral (\<lambda>x. h x * (f x * indicator A x)) =
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   937
      f.positive_integral (\<lambda>x. h x * indicator A x)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   938
      using `A \<in> sets M` h_borel borel
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   939
      by (simp add: positive_integral_translated_density ac_simps cong: positive_integral_cong)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   940
    also have "\<dots> = f'.positive_integral (\<lambda>x. h x * indicator A x)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   941
      by (rule f'.positive_integral_cong_measure) (rule f)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   942
    also have "\<dots> = positive_integral (\<lambda>x. h x * (f' x * indicator A x))"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   943
      using `A \<in> sets M` h_borel borel
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   944
      by (simp add: positive_integral_translated_density ac_simps cong: positive_integral_cong)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   945
    finally have "positive_integral (\<lambda>x. h x * (f x * indicator A x)) = positive_integral (\<lambda>x. h x * (f' x * indicator A x))" . }
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   946
  then have "h.almost_everywhere (\<lambda>x. f x = f' x)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   947
    using h_borel borel
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   948
    by (intro h.density_unique_finite_measure[OF borel])
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   949
       (simp add: positive_integral_translated_density)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   950
  then show "AE x. f x = f' x"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   951
    unfolding h.almost_everywhere_def almost_everywhere_def
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   952
    by (auto simp add: h_null_sets)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   953
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   954
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   955
lemma (in sigma_finite_measure) sigma_finite_iff_density_finite:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   956
  assumes \<nu>: "measure_space M \<nu>" and f: "f \<in> borel_measurable M"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   957
    and eq: "\<And>A. A \<in> sets M \<Longrightarrow> \<nu> A = positive_integral (\<lambda>x. f x * indicator A x)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   958
  shows "sigma_finite_measure M \<nu> \<longleftrightarrow> (AE x. f x \<noteq> \<omega>)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   959
proof
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   960
  assume "sigma_finite_measure M \<nu>"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   961
  then interpret \<nu>: sigma_finite_measure M \<nu> .
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   962
  from \<nu>.Ex_finite_integrable_function obtain h where
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   963
    h: "h \<in> borel_measurable M" "\<nu>.positive_integral h \<noteq> \<omega>"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   964
    and fin: "\<And>x. x \<in> space M \<Longrightarrow> 0 < h x \<and> h x < \<omega>" by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   965
  have "AE x. f x * h x \<noteq> \<omega>"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   966
  proof (rule AE_I')
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   967
    have "\<nu>.positive_integral h = positive_integral (\<lambda>x. f x * h x)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   968
      by (simp add: \<nu>.positive_integral_cong_measure[symmetric, OF eq[symmetric]])
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   969
         (intro positive_integral_translated_density f h)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   970
    then have "positive_integral (\<lambda>x. f x * h x) \<noteq> \<omega>"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   971
      using h(2) by simp
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   972
    then show "(\<lambda>x. f x * h x) -` {\<omega>} \<inter> space M \<in> null_sets"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   973
      using f h(1) by (auto intro!: positive_integral_omega borel_measurable_vimage)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   974
  qed auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   975
  then show "AE x. f x \<noteq> \<omega>"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   976
  proof (rule AE_mp, intro AE_cong)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   977
    fix x assume "x \<in> space M" from this[THEN fin]
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   978
    show "f x * h x \<noteq> \<omega> \<longrightarrow> f x \<noteq> \<omega>" by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   979
  qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   980
next
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   981
  assume AE: "AE x. f x \<noteq> \<omega>"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   982
  from sigma_finite guess Q .. note Q = this
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   983
  interpret \<nu>: measure_space M \<nu> by fact
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   984
  def A \<equiv> "\<lambda>i. f -` (case i of 0 \<Rightarrow> {\<omega>} | Suc n \<Rightarrow> {.. of_nat (Suc n)}) \<inter> space M"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   985
  { fix i j have "A i \<inter> Q j \<in> sets M"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   986
    unfolding A_def using f Q
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   987
    apply (rule_tac Int)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   988
    by (cases i) (auto intro: measurable_sets[OF f]) }
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   989
  note A_in_sets = this
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   990
  let "?A n" = "case prod_decode n of (i,j) \<Rightarrow> A i \<inter> Q j"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   991
  show "sigma_finite_measure M \<nu>"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   992
  proof (default, intro exI conjI subsetI allI)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   993
    fix x assume "x \<in> range ?A"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   994
    then obtain n where n: "x = ?A n" by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   995
    then show "x \<in> sets M" using A_in_sets by (cases "prod_decode n") auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   996
  next
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   997
    have "(\<Union>i. ?A i) = (\<Union>i j. A i \<inter> Q j)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   998
    proof safe
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   999
      fix x i j assume "x \<in> A i" "x \<in> Q j"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1000
      then show "x \<in> (\<Union>i. case prod_decode i of (i, j) \<Rightarrow> A i \<inter> Q j)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1001
        by (intro UN_I[of "prod_encode (i,j)"]) auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1002
    qed auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1003
    also have "\<dots> = (\<Union>i. A i) \<inter> space M" using Q by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1004
    also have "(\<Union>i. A i) = space M"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1005
    proof safe
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1006
      fix x assume x: "x \<in> space M"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1007
      show "x \<in> (\<Union>i. A i)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1008
      proof (cases "f x")
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1009
        case infinite then show ?thesis using x unfolding A_def by (auto intro: exI[of _ 0])
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1010
      next
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1011
        case (preal r)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1012
        with less_\<omega>_Ex_of_nat[of "f x"] obtain n where "f x < of_nat n" by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1013
        then show ?thesis using x preal unfolding A_def by (auto intro!: exI[of _ "Suc n"])
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1014
      qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1015
    qed (auto simp: A_def)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1016
    finally show "(\<Union>i. ?A i) = space M" by simp
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1017
  next
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1018
    fix n obtain i j where
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1019
      [simp]: "prod_decode n = (i, j)" by (cases "prod_decode n") auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1020
    have "positive_integral (\<lambda>x. f x * indicator (A i \<inter> Q j) x) \<noteq> \<omega>"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1021
    proof (cases i)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1022
      case 0
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1023
      have "AE x. f x * indicator (A i \<inter> Q j) x = 0"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1024
        using AE by (rule AE_mp) (auto intro!: AE_cong simp: A_def `i = 0`)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1025
      then have "positive_integral (\<lambda>x. f x * indicator (A i \<inter> Q j) x) = 0"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1026
        using A_in_sets f
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1027
        apply (subst positive_integral_0_iff)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1028
        apply fast
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1029
        apply (subst (asm) AE_iff_null_set)
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40871
diff changeset
  1030
        apply (intro borel_measurable_pextreal_neq_const)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1031
        apply fast
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1032
        by simp
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1033
      then show ?thesis by simp
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1034
    next
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1035
      case (Suc n)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1036
      then have "positive_integral (\<lambda>x. f x * indicator (A i \<inter> Q j) x) \<le>
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1037
        positive_integral (\<lambda>x. of_nat (Suc n) * indicator (Q j) x)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1038
        by (auto intro!: positive_integral_mono simp: indicator_def A_def)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1039
      also have "\<dots> = of_nat (Suc n) * \<mu> (Q j)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1040
        using Q by (auto intro!: positive_integral_cmult_indicator)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1041
      also have "\<dots> < \<omega>"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1042
        using Q by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1043
      finally show ?thesis by simp
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1044
    qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1045
    then show "\<nu> (?A n) \<noteq> \<omega>"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1046
      using A_in_sets Q eq by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1047
  qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1048
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1049
40871
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
  1050
section "Radon-Nikodym derivative"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1051
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1052
definition (in sigma_finite_measure)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1053
  "RN_deriv \<nu> \<equiv> SOME f. f \<in> borel_measurable M \<and>
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1054
    (\<forall>A \<in> sets M. \<nu> A = positive_integral (\<lambda>x. f x * indicator A x))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1055
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1056
lemma (in sigma_finite_measure) RN_deriv_cong:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1057
  assumes cong: "\<And>A. A \<in> sets M \<Longrightarrow> \<mu>' A = \<mu> A" "\<And>A. A \<in> sets M \<Longrightarrow> \<nu>' A = \<nu> A"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1058
  shows "sigma_finite_measure.RN_deriv M \<mu>' \<nu>' x = RN_deriv \<nu> x"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1059
proof -
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1060
  interpret \<mu>': sigma_finite_measure M \<mu>'
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1061
    using cong(1) by (rule sigma_finite_measure_cong)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1062
  show ?thesis
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1063
    unfolding RN_deriv_def \<mu>'.RN_deriv_def
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1064
    by (simp add: cong positive_integral_cong_measure[OF cong(1)])
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1065
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1066
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1067
lemma (in sigma_finite_measure) RN_deriv:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1068
  assumes "measure_space M \<nu>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1069
  assumes "absolutely_continuous \<nu>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1070
  shows "RN_deriv \<nu> \<in> borel_measurable M" (is ?borel)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1071
  and "\<And>A. A \<in> sets M \<Longrightarrow> \<nu> A = positive_integral (\<lambda>x. RN_deriv \<nu> x * indicator A x)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1072
    (is "\<And>A. _ \<Longrightarrow> ?int A")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1073
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1074
  note Ex = Radon_Nikodym[OF assms, unfolded Bex_def]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1075
  thus ?borel unfolding RN_deriv_def by (rule someI2_ex) auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1076
  fix A assume "A \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1077
  from Ex show "?int A" unfolding RN_deriv_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1078
    by (rule someI2_ex) (simp add: `A \<in> sets M`)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1079
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1080
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1081
lemma (in sigma_finite_measure) RN_deriv_positive_integral:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1082
  assumes \<nu>: "measure_space M \<nu>" "absolutely_continuous \<nu>"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1083
    and f: "f \<in> borel_measurable M"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1084
  shows "measure_space.positive_integral M \<nu> f = positive_integral (\<lambda>x. RN_deriv \<nu> x * f x)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1085
proof -
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1086
  interpret \<nu>: measure_space M \<nu> by fact
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1087
  have "\<nu>.positive_integral f =
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1088
    measure_space.positive_integral M (\<lambda>A. positive_integral (\<lambda>x. RN_deriv \<nu> x * indicator A x)) f"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1089
    by (intro \<nu>.positive_integral_cong_measure[symmetric] RN_deriv(2)[OF \<nu>, symmetric])
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1090
  also have "\<dots> = positive_integral (\<lambda>x. RN_deriv \<nu> x * f x)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1091
    by (intro positive_integral_translated_density RN_deriv[OF \<nu>] f)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1092
  finally show ?thesis .
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1093
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1094
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1095
lemma (in sigma_finite_measure) RN_deriv_unique:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1096
  assumes \<nu>: "measure_space M \<nu>" "absolutely_continuous \<nu>"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1097
  and f: "f \<in> borel_measurable M"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1098
  and eq: "\<And>A. A \<in> sets M \<Longrightarrow> \<nu> A = positive_integral (\<lambda>x. f x * indicator A x)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1099
  shows "AE x. f x = RN_deriv \<nu> x"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1100
proof (rule density_unique[OF f RN_deriv(1)[OF \<nu>]])
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1101
  fix A assume A: "A \<in> sets M"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1102
  show "positive_integral (\<lambda>x. f x * indicator A x) = positive_integral (\<lambda>x. RN_deriv \<nu> x * indicator A x)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1103
    unfolding eq[OF A, symmetric] RN_deriv(2)[OF \<nu> A, symmetric] ..
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1104
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1105
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1106
lemma (in sigma_finite_measure) RN_deriv_vimage:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1107
  fixes f :: "'b \<Rightarrow> 'a"
41095
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
  1108
  assumes f: "bij_inv S (space M) f g"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1109
  assumes \<nu>: "measure_space M \<nu>" "absolutely_continuous \<nu>"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1110
  shows "AE x.
41095
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
  1111
    sigma_finite_measure.RN_deriv (vimage_algebra S f) (\<lambda>A. \<mu> (f ` A)) (\<lambda>A. \<nu> (f ` A)) (g x) = RN_deriv \<nu> x"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1112
proof (rule RN_deriv_unique[OF \<nu>])
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1113
  interpret sf: sigma_finite_measure "vimage_algebra S f" "\<lambda>A. \<mu> (f ` A)"
41095
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
  1114
    using f by (rule sigma_finite_measure_isomorphic[OF bij_inv_bij_betw(1)])
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1115
  interpret \<nu>: measure_space M \<nu> using \<nu>(1) .
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1116
  have \<nu>': "measure_space (vimage_algebra S f) (\<lambda>A. \<nu> (f ` A))"
41095
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
  1117
    using f by (rule \<nu>.measure_space_isomorphic[OF bij_inv_bij_betw(1)])
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1118
  { fix A assume "A \<in> sets M" then have "f ` (f -` A \<inter> S) = A"
41095
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
  1119
      using sets_into_space f[THEN bij_inv_bij_betw(1), unfolded bij_betw_def]
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1120
      by (intro image_vimage_inter_eq[where T="space M"]) auto }
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1121
  note A_f = this
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1122
  then have ac: "sf.absolutely_continuous (\<lambda>A. \<nu> (f ` A))"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1123
    using \<nu>(2) by (auto simp: sf.absolutely_continuous_def absolutely_continuous_def)
41095
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
  1124
  show "(\<lambda>x. sf.RN_deriv (\<lambda>A. \<nu> (f ` A)) (g x)) \<in> borel_measurable M"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1125
    using sf.RN_deriv(1)[OF \<nu>' ac]
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1126
    unfolding measurable_vimage_iff_inv[OF f] comp_def .
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1127
  fix A assume "A \<in> sets M"
41095
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
  1128
  then have *: "\<And>x. x \<in> space M \<Longrightarrow> indicator (f -` A \<inter> S) (g x) = (indicator A x :: pextreal)"
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
  1129
    using f by (auto simp: bij_inv_def indicator_def)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1130
  have "\<nu> (f ` (f -` A \<inter> S)) = sf.positive_integral (\<lambda>x. sf.RN_deriv (\<lambda>A. \<nu> (f ` A)) x * indicator (f -` A \<inter> S) x)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1131
    using `A \<in> sets M` by (force intro!: sf.RN_deriv(2)[OF \<nu>' ac])
41095
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
  1132
  also have "\<dots> = positive_integral (\<lambda>x. sf.RN_deriv (\<lambda>A. \<nu> (f ` A)) (g x) * indicator A x)"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1133
    unfolding positive_integral_vimage_inv[OF f]
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1134
    by (simp add: * cong: positive_integral_cong)
41095
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
  1135
  finally show "\<nu> A = positive_integral (\<lambda>x. sf.RN_deriv (\<lambda>A. \<nu> (f ` A)) (g x) * indicator A x)"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1136
    unfolding A_f[OF `A \<in> sets M`] .
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1137
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1138
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1139
lemma (in sigma_finite_measure) RN_deriv_finite:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1140
  assumes sfm: "sigma_finite_measure M \<nu>" and ac: "absolutely_continuous \<nu>"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1141
  shows "AE x. RN_deriv \<nu> x \<noteq> \<omega>"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1142
proof -
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1143
  interpret \<nu>: sigma_finite_measure M \<nu> by fact
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1144
  have \<nu>: "measure_space M \<nu>" by default
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1145
  from sfm show ?thesis
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1146
    using sigma_finite_iff_density_finite[OF \<nu> RN_deriv[OF \<nu> ac]] by simp
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1147
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1148
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1149
lemma (in sigma_finite_measure)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1150
  assumes \<nu>: "sigma_finite_measure M \<nu>" "absolutely_continuous \<nu>"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1151
    and f: "f \<in> borel_measurable M"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1152
  shows RN_deriv_integral: "measure_space.integral M \<nu> f = integral (\<lambda>x. real (RN_deriv \<nu> x) * f x)" (is ?integral)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1153
    and RN_deriv_integrable: "measure_space.integrable M \<nu> f \<longleftrightarrow> integrable (\<lambda>x. real (RN_deriv \<nu> x) * f x)" (is ?integrable)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1154
proof -
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1155
  interpret \<nu>: sigma_finite_measure M \<nu> by fact
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1156
  have ms: "measure_space M \<nu>" by default
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40871
diff changeset
  1157
  have minus_cong: "\<And>A B A' B'::pextreal. A = A' \<Longrightarrow> B = B' \<Longrightarrow> real A - real B = real A' - real B'" by simp
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1158
  have f': "(\<lambda>x. - f x) \<in> borel_measurable M" using f by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1159
  { fix f :: "'a \<Rightarrow> real" assume "f \<in> borel_measurable M"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1160
    { fix x assume *: "RN_deriv \<nu> x \<noteq> \<omega>"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1161
      have "Real (real (RN_deriv \<nu> x)) * Real (f x) = Real (real (RN_deriv \<nu> x) * f x)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1162
        by (simp add: mult_le_0_iff)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1163
      then have "RN_deriv \<nu> x * Real (f x) = Real (real (RN_deriv \<nu> x) * f x)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1164
        using * by (simp add: Real_real) }
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1165
    note * = this
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1166
    have "positive_integral (\<lambda>x. RN_deriv \<nu> x * Real (f x)) = positive_integral (\<lambda>x. Real (real (RN_deriv \<nu> x) * f x))"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1167
      apply (rule positive_integral_cong_AE)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1168
      apply (rule AE_mp[OF RN_deriv_finite[OF \<nu>]])
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1169
      by (auto intro!: AE_cong simp: *) }
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1170
  with this[OF f] this[OF f'] f f'
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1171
  show ?integral ?integrable
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1172
    unfolding \<nu>.integral_def integral_def \<nu>.integrable_def integrable_def
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1173
    by (auto intro!: RN_deriv(1)[OF ms \<nu>(2)] minus_cong simp: RN_deriv_positive_integral[OF ms \<nu>(2)])
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1174
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1175
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1176
lemma (in sigma_finite_measure) RN_deriv_singleton:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1177
  assumes "measure_space M \<nu>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1178
  and ac: "absolutely_continuous \<nu>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1179
  and "{x} \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1180
  shows "\<nu> {x} = RN_deriv \<nu> x * \<mu> {x}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1181
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1182
  note deriv = RN_deriv[OF assms(1, 2)]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1183
  from deriv(2)[OF `{x} \<in> sets M`]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1184
  have "\<nu> {x} = positive_integral (\<lambda>w. RN_deriv \<nu> x * indicator {x} w)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1185
    by (auto simp: indicator_def intro!: positive_integral_cong)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1186
  thus ?thesis using positive_integral_cmult_indicator[OF `{x} \<in> sets M`]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1187
    by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1188
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1189
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1190
theorem (in finite_measure_space) RN_deriv_finite_measure:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1191
  assumes "measure_space M \<nu>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1192
  and ac: "absolutely_continuous \<nu>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1193
  and "x \<in> space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1194
  shows "\<nu> {x} = RN_deriv \<nu> x * \<mu> {x}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1195
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1196
  have "{x} \<in> sets M" using sets_eq_Pow `x \<in> space M` by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1197
  from RN_deriv_singleton[OF assms(1,2) this] show ?thesis .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1198
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1199
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1200
end