src/HOL/Probability/Radon_Nikodym.thy
author haftmann
Thu, 18 Nov 2010 17:01:15 +0100
changeset 40602 91e583511113
parent 39097 943c7b348524
child 40859 de0b30e6c2d2
permissions -rw-r--r--
map_fun combinator in theory Fun
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
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
     5
lemma (in sigma_finite_measure) Ex_finite_integrable_function:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
     6
  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
     7
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
     8
  obtain A :: "nat \<Rightarrow> 'a set" where
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
     9
    range: "range A \<subseteq> sets M" and
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    10
    space: "(\<Union>i. A i) = space M" and
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    11
    measure: "\<And>i. \<mu> (A i) \<noteq> \<omega>" and
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    12
    disjoint: "disjoint_family A"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    13
    using disjoint_sigma_finite by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    14
  let "?B i" = "2^Suc i * \<mu> (A i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    15
  have "\<forall>i. \<exists>x. 0 < x \<and> x < inverse (?B i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    16
  proof
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    17
    fix i show "\<exists>x. 0 < x \<and> x < inverse (?B i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    18
    proof cases
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    19
      assume "\<mu> (A i) = 0"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    20
      then show ?thesis by (auto intro!: exI[of _ 1])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    21
    next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    22
      assume not_0: "\<mu> (A i) \<noteq> 0"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    23
      then have "?B i \<noteq> \<omega>" using measure[of i] by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    24
      then have "inverse (?B i) \<noteq> 0" unfolding pinfreal_inverse_eq_0 by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    25
      then show ?thesis using measure[of i] not_0
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    26
        by (auto intro!: exI[of _ "inverse (?B i) / 2"]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    27
                 simp: pinfreal_noteq_omega_Ex zero_le_mult_iff zero_less_mult_iff mult_le_0_iff power_le_zero_eq)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    28
    qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    29
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    30
  from choice[OF this] obtain n where n: "\<And>i. 0 < n i"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    31
    "\<And>i. n i < inverse (2^Suc i * \<mu> (A i))" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    32
  let "?h x" = "\<Sum>\<^isub>\<infinity> i. n i * indicator (A i) x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    33
  show ?thesis
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    34
  proof (safe intro!: bexI[of _ ?h] del: notI)
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38656
diff changeset
    35
    have "\<And>i. A i \<in> sets M"
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38656
diff changeset
    36
      using range by fastsimp+
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38656
diff changeset
    37
    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
    38
      by (simp add: positive_integral_psuminf positive_integral_cmult_indicator)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    39
    also have "\<dots> \<le> (\<Sum>\<^isub>\<infinity> i. Real ((1 / 2)^Suc i))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    40
    proof (rule psuminf_le)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    41
      fix N show "n N * \<mu> (A N) \<le> Real ((1 / 2) ^ Suc N)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    42
        using measure[of N] n[of N]
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38656
diff changeset
    43
        by (cases "n N")
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38656
diff changeset
    44
           (auto simp: pinfreal_noteq_omega_Ex field_simps zero_le_mult_iff
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38656
diff changeset
    45
                       mult_le_0_iff mult_less_0_iff power_less_zero_eq
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38656
diff changeset
    46
                       power_le_zero_eq inverse_eq_divide less_divide_eq
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38656
diff changeset
    47
                       power_divide split: split_if_asm)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    48
    qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    49
    also have "\<dots> = Real 1"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    50
      by (rule suminf_imp_psuminf, rule power_half_series, auto)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    51
    finally show "positive_integral ?h \<noteq> \<omega>" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    52
  next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    53
    fix x assume "x \<in> space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    54
    then obtain i where "x \<in> A i" using space[symmetric] by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    55
    from psuminf_cmult_indicator[OF disjoint, OF this]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    56
    have "?h x = n i" by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    57
    then show "0 < ?h x" and "?h x < \<omega>" using n[of i] by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    58
  next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    59
    show "?h \<in> borel_measurable M" using range
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38656
diff changeset
    60
      by (auto intro!: borel_measurable_psuminf borel_measurable_pinfreal_times)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    61
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    62
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    63
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    64
definition (in measure_space)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    65
  "absolutely_continuous \<nu> = (\<forall>N\<in>null_sets. \<nu> N = (0 :: pinfreal))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    66
39097
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39092
diff changeset
    67
lemma (in finite_measure_space) absolutely_continuousI:
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39092
diff changeset
    68
  assumes "finite_measure_space M \<nu>"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39092
diff changeset
    69
  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
    70
  shows "absolutely_continuous \<nu>"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39092
diff changeset
    71
proof (unfold absolutely_continuous_def sets_eq_Pow, safe)
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39092
diff changeset
    72
  fix N assume "\<mu> N = 0" "N \<subseteq> space M"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39092
diff changeset
    73
  interpret v: finite_measure_space M \<nu> by fact
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39092
diff changeset
    74
  have "\<nu> N = \<nu> (\<Union>x\<in>N. {x})" by simp
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39092
diff changeset
    75
  also have "\<dots> = (\<Sum>x\<in>N. \<nu> {x})"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39092
diff changeset
    76
  proof (rule v.measure_finitely_additive''[symmetric])
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39092
diff changeset
    77
    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
    78
    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
    79
    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
    80
  qed
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39092
diff changeset
    81
  also have "\<dots> = 0"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39092
diff changeset
    82
  proof (safe intro!: setsum_0')
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39092
diff changeset
    83
    fix x assume "x \<in> N"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39092
diff changeset
    84
    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
    85
    hence "\<mu> {x} = 0" using `\<mu> N = 0` by simp
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39092
diff changeset
    86
    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
    87
  qed
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39092
diff changeset
    88
  finally show "\<nu> N = 0" .
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39092
diff changeset
    89
qed
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39092
diff changeset
    90
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    91
lemma (in finite_measure) Radon_Nikodym_aux_epsilon:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    92
  fixes e :: real assumes "0 < e"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    93
  assumes "finite_measure M s"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    94
  shows "\<exists>A\<in>sets M. real (\<mu> (space M)) - real (s (space M)) \<le>
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    95
                    real (\<mu> A) - real (s A) \<and>
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    96
                    (\<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
    97
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    98
  let "?d A" = "real (\<mu> A) - real (s A)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    99
  interpret M': finite_measure M s by fact
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   100
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   101
  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
   102
    then {}
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   103
    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
   104
  def A \<equiv> "\<lambda>n. ((\<lambda>B. B \<union> ?A B) ^^ n) {}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   105
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   106
  have A_simps[simp]:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   107
    "A 0 = {}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   108
    "\<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
   109
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   110
  { fix A assume "A \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   111
    have "?A A \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   112
      by (auto intro!: someI2[of _ _ "\<lambda>A. A \<in> sets M"] simp: not_less) }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   113
  note A'_in_sets = this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   114
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   115
  { fix n have "A n \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   116
    proof (induct n)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   117
      case (Suc n) thus "A (Suc n) \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   118
        using A'_in_sets[of "A n"] by (auto split: split_if_asm)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   119
    qed (simp add: A_def) }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   120
  note A_in_sets = this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   121
  hence "range A \<subseteq> sets M" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   122
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   123
  { fix n B
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   124
    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
   125
    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
   126
    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
   127
    proof (rule someI2_ex[OF Ex])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   128
      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
   129
      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
   130
      hence "?d (A n \<union> B) = ?d (A n) + ?d B"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   131
        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
   132
      also have "\<dots> \<le> ?d (A n) - e" using dB by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   133
      finally show "?d (A n \<union> B) \<le> ?d (A n) - e" .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   134
    qed }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   135
  note dA_epsilon = this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   136
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   137
  { fix n have "?d (A (Suc n)) \<le> ?d (A n)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   138
    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
   139
      case True from dA_epsilon[OF this] show ?thesis using `0 < e` by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   140
    next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   141
      case False
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   142
      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
   143
      thus ?thesis by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   144
    qed }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   145
  note dA_mono = this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   146
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   147
  show ?thesis
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   148
  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
   149
    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
   150
    show ?thesis
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   151
    proof (safe intro!: bexI[of _ "space M - A n"])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   152
      fix B assume "B \<in> sets M" "B \<subseteq> space M - A n"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   153
      from B[OF this] show "-e < ?d B" .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   154
    next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   155
      show "space M - A n \<in> sets M" by (rule compl_sets) fact
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   156
    next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   157
      show "?d (space M) \<le> ?d (space M - A n)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   158
      proof (induct n)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   159
        fix n assume "?d (space M) \<le> ?d (space M - A n)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   160
        also have "\<dots> \<le> ?d (space M - A (Suc n))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   161
          using A_in_sets sets_into_space dA_mono[of n]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   162
            real_finite_measure_Diff[of "space M"]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   163
            real_finite_measure_Diff[of "space M"]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   164
            M'.real_finite_measure_Diff[of "space M"]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   165
            M'.real_finite_measure_Diff[of "space M"]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   166
          by (simp del: A_simps)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   167
        finally show "?d (space M) \<le> ?d (space M - A (Suc n))" .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   168
      qed simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   169
    qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   170
  next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   171
    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
   172
      by (auto simp add: not_less)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   173
    { fix n have "?d (A n) \<le> - real n * e"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   174
      proof (induct n)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   175
        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
   176
      qed simp } note dA_less = this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   177
    have decseq: "decseq (\<lambda>n. ?d (A n))" unfolding decseq_eq_incseq
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   178
    proof (rule incseq_SucI)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   179
      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
   180
    qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   181
    from real_finite_continuity_from_below[of A] `range A \<subseteq> sets M`
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   182
      M'.real_finite_continuity_from_below[of A]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   183
    have convergent: "(\<lambda>i. ?d (A i)) ----> ?d (\<Union>i. A i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   184
      by (auto intro!: LIMSEQ_diff)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   185
    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
   186
    moreover from order_trans[OF decseq_le[OF decseq convergent] dA_less]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   187
    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
   188
    ultimately show ?thesis by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   189
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   190
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   191
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   192
lemma (in finite_measure) Radon_Nikodym_aux:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   193
  assumes "finite_measure M s"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   194
  shows "\<exists>A\<in>sets M. real (\<mu> (space M)) - real (s (space M)) \<le>
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   195
                    real (\<mu> A) - real (s A) \<and>
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   196
                    (\<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
   197
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   198
  let "?d A" = "real (\<mu> A) - real (s A)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   199
  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
   200
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   201
  interpret M': finite_measure M s by fact
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   202
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38656
diff changeset
   203
  let "?r S" = "restricted_space S"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   204
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   205
  { fix S n
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   206
    assume S: "S \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   207
    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
   208
    from M'.restricted_finite_measure[of S] restricted_finite_measure[of S] S
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   209
    have "finite_measure (?r S) \<mu>" "0 < 1 / real (Suc n)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   210
      "finite_measure (?r S) s" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   211
    from finite_measure.Radon_Nikodym_aux_epsilon[OF this] guess X ..
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   212
    hence "?P X S n"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   213
    proof (simp add: **, safe)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   214
      fix C assume C: "C \<in> sets M" "C \<subseteq> X" "X \<subseteq> S" and
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   215
        *: "\<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
   216
      hence "C \<subseteq> S" "C \<subseteq> X" "S \<inter> C = C" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   217
      with *[THEN bspec, OF `C \<in> sets M`]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   218
      show "- (1 / real (Suc n)) < ?d C" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   219
    qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   220
    hence "\<exists>A. ?P A S n" by auto }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   221
  note Ex_P = this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   222
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   223
  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
   224
  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
   225
  have A_0[simp]: "A 0 = space M" unfolding A_def by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   226
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   227
  { fix i have "A i \<in> sets M" unfolding A_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   228
    proof (induct i)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   229
      case (Suc i)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   230
      from Ex_P[OF this, of i] show ?case unfolding nat_rec_Suc
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   231
        by (rule someI2_ex) simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   232
    qed simp }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   233
  note A_in_sets = this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   234
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   235
  { fix n have "?P (A (Suc n)) (A n) n"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   236
      using Ex_P[OF A_in_sets] unfolding A_Suc
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   237
      by (rule someI2_ex) simp }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   238
  note P_A = this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   239
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   240
  have "range A \<subseteq> sets M" using A_in_sets by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   241
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   242
  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
   243
  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
   244
  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
   245
      using P_A by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   246
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   247
  show ?thesis
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   248
  proof (safe intro!: bexI[of _ "\<Inter>i. A i"])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   249
    show "(\<Inter>i. A i) \<in> sets M" using A_in_sets by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   250
    from `range A \<subseteq> sets M` A_mono
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   251
      real_finite_continuity_from_above[of A]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   252
      M'.real_finite_continuity_from_above[of A]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   253
    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
   254
    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
   255
      by (rule_tac LIMSEQ_le_const) (auto intro!: exI)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   256
  next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   257
    fix B assume B: "B \<in> sets M" "B \<subseteq> (\<Inter>i. A i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   258
    show "0 \<le> ?d B"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   259
    proof (rule ccontr)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   260
      assume "\<not> 0 \<le> ?d B"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   261
      hence "0 < - ?d B" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   262
      from ex_inverse_of_nat_Suc_less[OF this]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   263
      obtain n where *: "?d B < - 1 / real (Suc n)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   264
        by (auto simp: real_eq_of_nat inverse_eq_divide field_simps)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   265
      have "B \<subseteq> A (Suc n)" using B by (auto simp del: nat_rec_Suc)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   266
      from epsilon[OF B(1) this] *
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   267
      show False by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   268
    qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   269
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   270
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   271
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   272
lemma (in finite_measure) Radon_Nikodym_finite_measure:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   273
  assumes "finite_measure M \<nu>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   274
  assumes "absolutely_continuous \<nu>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   275
  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
   276
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   277
  interpret M': finite_measure M \<nu> using assms(1) .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   278
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   279
  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
   280
  have "(\<lambda>x. 0) \<in> G" unfolding G_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   281
  hence "G \<noteq> {}" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   282
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   283
  { fix f g assume f: "f \<in> G" and g: "g \<in> G"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   284
    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
   285
    proof safe
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   286
      show "?max \<in> borel_measurable M" using f g unfolding G_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   287
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   288
      let ?A = "{x \<in> space M. f x \<le> g x}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   289
      have "?A \<in> sets M" using f g unfolding G_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   290
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   291
      fix A assume "A \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   292
      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
   293
      have union: "((?A \<inter> A) \<union> ((space M - ?A) \<inter> A)) = A"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   294
        using sets_into_space[OF `A \<in> sets M`] by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   295
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   296
      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
   297
        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
   298
        by (auto simp: indicator_def max_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   299
      hence "positive_integral (\<lambda>x. max (g x) (f x) * indicator A x) =
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   300
        positive_integral (\<lambda>x. g x * indicator (?A \<inter> A) x) +
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   301
        positive_integral (\<lambda>x. f x * indicator ((space M - ?A) \<inter> A) x)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   302
        using f g sets unfolding G_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   303
        by (auto cong: positive_integral_cong intro!: positive_integral_add borel_measurable_indicator)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   304
      also have "\<dots> \<le> \<nu> (?A \<inter> A) + \<nu> ((space M - ?A) \<inter> A)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   305
        using f g sets unfolding G_def by (auto intro!: add_mono)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   306
      also have "\<dots> = \<nu> A"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   307
        using M'.measure_additive[OF sets] union by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   308
      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
   309
    qed }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   310
  note max_in_G = this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   311
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   312
  { fix f g assume  "f \<up> g" and f: "\<And>i. f i \<in> G"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   313
    have "g \<in> G" unfolding G_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   314
    proof safe
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   315
      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
   316
      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
   317
      thus "g \<in> borel_measurable M" by (auto intro!: borel_measurable_SUP)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   318
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   319
      fix A assume "A \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   320
      hence "\<And>i. (\<lambda>x. f i x * indicator A x) \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   321
        using f_borel by (auto intro!: borel_measurable_indicator)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   322
      from positive_integral_isoton[OF isoton_indicator[OF `f \<up> g`] this]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   323
      have SUP: "positive_integral (\<lambda>x. g x * indicator A x) =
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   324
          (SUP i. positive_integral (\<lambda>x. f i x * indicator A x))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   325
        unfolding isoton_def by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   326
      show "positive_integral (\<lambda>x. g x * indicator A x) \<le> \<nu> A" unfolding SUP
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   327
        using f `A \<in> sets M` unfolding G_def by (auto intro!: SUP_leI)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   328
    qed }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   329
  note SUP_in_G = this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   330
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   331
  let ?y = "SUP g : G. positive_integral g"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   332
  have "?y \<le> \<nu> (space M)" unfolding G_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   333
  proof (safe intro!: SUP_leI)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   334
    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
   335
    from this[THEN bspec, OF top] show "positive_integral g \<le> \<nu> (space M)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   336
      by (simp cong: positive_integral_cong)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   337
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   338
  hence "?y \<noteq> \<omega>" using M'.finite_measure_of_space by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   339
  from SUPR_countable_SUPR[OF this `G \<noteq> {}`] guess ys .. note ys = this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   340
  hence "\<forall>n. \<exists>g. g\<in>G \<and> positive_integral g = ys n"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   341
  proof safe
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   342
    fix n assume "range ys \<subseteq> positive_integral ` G"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   343
    hence "ys n \<in> positive_integral ` G" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   344
    thus "\<exists>g. g\<in>G \<and> positive_integral g = ys n" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   345
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   346
  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
   347
  hence y_eq: "?y = (SUP i. positive_integral (gs i))" using ys by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   348
  let "?g i x" = "Max ((\<lambda>n. gs n x) ` {..i})"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   349
  def f \<equiv> "SUP i. ?g i"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   350
  have gs_not_empty: "\<And>i. (\<lambda>n. gs n x) ` {..i} \<noteq> {}" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   351
  { fix i have "?g i \<in> G"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   352
    proof (induct i)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   353
      case 0 thus ?case by simp fact
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   354
    next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   355
      case (Suc i)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   356
      with Suc gs_not_empty `gs (Suc i) \<in> G` show ?case
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   357
        by (auto simp add: atMost_Suc intro!: max_in_G)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   358
    qed }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   359
  note g_in_G = this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   360
  have "\<And>x. \<forall>i. ?g i x \<le> ?g (Suc i) x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   361
    using gs_not_empty by (simp add: atMost_Suc)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   362
  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
   363
  from SUP_in_G[OF this g_in_G] have "f \<in> G" .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   364
  hence [simp, intro]: "f \<in> borel_measurable M" unfolding G_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   365
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   366
  have "(\<lambda>i. positive_integral (?g i)) \<up> positive_integral f"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   367
    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
   368
  hence "positive_integral f = (SUP i. positive_integral (?g i))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   369
    unfolding isoton_def by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   370
  also have "\<dots> = ?y"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   371
  proof (rule antisym)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   372
    show "(SUP i. positive_integral (?g i)) \<le> ?y"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   373
      using g_in_G by (auto intro!: exI Sup_mono simp: SUPR_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   374
    show "?y \<le> (SUP i. positive_integral (?g i))" unfolding y_eq
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   375
      by (auto intro!: SUP_mono positive_integral_mono Max_ge)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   376
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   377
  finally have int_f_eq_y: "positive_integral f = ?y" .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   378
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   379
  let "?t A" = "\<nu> A - positive_integral (\<lambda>x. f x * indicator A x)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   380
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   381
  have "finite_measure M ?t"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   382
  proof
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   383
    show "?t {} = 0" by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   384
    show "?t (space M) \<noteq> \<omega>" using M'.finite_measure by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   385
    show "countably_additive M ?t" unfolding countably_additive_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   386
    proof safe
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   387
      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
   388
      have "(\<Sum>\<^isub>\<infinity> n. positive_integral (\<lambda>x. f x * indicator (A n) x))
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   389
        = positive_integral (\<lambda>x. (\<Sum>\<^isub>\<infinity>n. f x * indicator (A n) x))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   390
        using `range A \<subseteq> sets M`
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   391
        by (rule_tac positive_integral_psuminf[symmetric]) (auto intro!: borel_measurable_indicator)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   392
      also have "\<dots> = positive_integral (\<lambda>x. f x * indicator (\<Union>n. A n) x)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   393
        apply (rule positive_integral_cong)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   394
        apply (subst psuminf_cmult_right)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   395
        unfolding psuminf_indicator[OF `disjoint_family A`] ..
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   396
      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
   397
        = positive_integral (\<lambda>x. f x * indicator (\<Union>n. A n) x)" .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   398
      moreover have "(\<Sum>\<^isub>\<infinity>n. \<nu> (A n)) = \<nu> (\<Union>n. A n)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   399
        using M'.measure_countably_additive A by (simp add: comp_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   400
      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
   401
          using A `f \<in> G` unfolding G_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   402
      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
   403
      moreover {
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   404
        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
   405
          using A `f \<in> G` unfolding G_def by (auto simp: countable_UN)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   406
        also have "\<nu> (\<Union>i. A i) < \<omega>" using v_fin by (simp add: pinfreal_less_\<omega>)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   407
        finally have "positive_integral (\<lambda>x. f x * indicator (\<Union>i. A i) x) \<noteq> \<omega>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   408
          by (simp add: pinfreal_less_\<omega>) }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   409
      ultimately
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   410
      show "(\<Sum>\<^isub>\<infinity> n. ?t (A n)) = ?t (\<Union>i. A i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   411
        apply (subst psuminf_minus) by simp_all
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   412
    qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   413
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   414
  then interpret M: finite_measure M ?t .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   415
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   416
  have ac: "absolutely_continuous ?t" using `absolutely_continuous \<nu>` unfolding absolutely_continuous_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   417
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   418
  have upper_bound: "\<forall>A\<in>sets M. ?t A \<le> 0"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   419
  proof (rule ccontr)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   420
    assume "\<not> ?thesis"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   421
    then obtain A where A: "A \<in> sets M" and pos: "0 < ?t A"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   422
      by (auto simp: not_le)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   423
    note pos
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   424
    also have "?t A \<le> ?t (space M)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   425
      using M.measure_mono[of A "space M"] A sets_into_space by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   426
    finally have pos_t: "0 < ?t (space M)" by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   427
    moreover
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   428
    hence pos_M: "0 < \<mu> (space M)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   429
      using ac top unfolding absolutely_continuous_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   430
    moreover
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   431
    have "positive_integral (\<lambda>x. f x * indicator (space M) x) \<le> \<nu> (space M)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   432
      using `f \<in> G` unfolding G_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   433
    hence "positive_integral (\<lambda>x. f x * indicator (space M) x) \<noteq> \<omega>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   434
      using M'.finite_measure_of_space by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   435
    moreover
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   436
    def b \<equiv> "?t (space M) / \<mu> (space M) / 2"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   437
    ultimately have b: "b \<noteq> 0" "b \<noteq> \<omega>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   438
      using M'.finite_measure_of_space
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   439
      by (auto simp: pinfreal_inverse_eq_0 finite_measure_of_space)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   440
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   441
    have "finite_measure M (\<lambda>A. b * \<mu> A)" (is "finite_measure M ?b")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   442
    proof
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   443
      show "?b {} = 0" by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   444
      show "?b (space M) \<noteq> \<omega>" using finite_measure_of_space b by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   445
      show "countably_additive M ?b"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   446
        unfolding countably_additive_def psuminf_cmult_right
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   447
        using measure_countably_additive by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   448
    qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   449
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   450
    from M.Radon_Nikodym_aux[OF this]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   451
    obtain A0 where "A0 \<in> sets M" and
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   452
      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
   453
      *: "\<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
   454
    { fix B assume "B \<in> sets M" "B \<subseteq> A0"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   455
      with *[OF this] have "b * \<mu> B \<le> ?t B"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   456
        using M'.finite_measure b finite_measure
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   457
        by (cases "b * \<mu> B", cases "?t B") (auto simp: field_simps) }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   458
    note bM_le_t = this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   459
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   460
    let "?f0 x" = "f x + b * indicator A0 x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   461
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
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   473
    { fix A assume A: "A \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   474
      hence "A \<inter> A0 \<in> sets M" using `A0 \<in> sets M` by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   475
      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
   476
        using `f \<in> G` A unfolding G_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   477
      note f0_eq[OF A]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   478
      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
   479
          positive_integral (\<lambda>x. f x * indicator A x) + ?t (A \<inter> A0)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   480
        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
   481
        by (auto intro!: add_left_mono)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   482
      also have "\<dots> \<le> positive_integral (\<lambda>x. f x * indicator A x) + ?t A"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   483
        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
   484
        by (auto intro!: add_left_mono)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   485
      also have "\<dots> \<le> \<nu> A"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   486
        using f_le_v M'.finite_measure[simplified, OF `A \<in> sets M`]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   487
        by (cases "positive_integral (\<lambda>x. f x * indicator A x)", cases "\<nu> A", auto)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   488
      finally have "positive_integral (\<lambda>x. ?f0 x * indicator A x) \<le> \<nu> A" . }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   489
    hence "?f0 \<in> G" using `A0 \<in> sets M` unfolding G_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   490
      by (auto intro!: borel_measurable_indicator borel_measurable_pinfreal_add borel_measurable_pinfreal_times)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   491
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   492
    have real: "?t (space M) \<noteq> \<omega>" "?t A0 \<noteq> \<omega>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   493
      "b * \<mu> (space M) \<noteq> \<omega>" "b * \<mu> A0 \<noteq> \<omega>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   494
      using `A0 \<in> sets M` b
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   495
        finite_measure[of A0] M.finite_measure[of A0]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   496
        finite_measure_of_space M.finite_measure_of_space
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   497
      by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   498
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   499
    have int_f_finite: "positive_integral f \<noteq> \<omega>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   500
      using M'.finite_measure_of_space pos_t unfolding pinfreal_zero_less_diff_iff
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   501
      by (auto cong: positive_integral_cong)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   502
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   503
    have "?t (space M) > b * \<mu> (space M)" unfolding b_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   504
      apply (simp add: field_simps)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   505
      apply (subst mult_assoc[symmetric])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   506
      apply (subst pinfreal_mult_inverse)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   507
      using finite_measure_of_space M'.finite_measure_of_space pos_t pos_M
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   508
      using pinfreal_mult_strict_right_mono[of "Real (1/2)" 1 "?t (space M)"]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   509
      by simp_all
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   510
    hence  "0 < ?t (space M) - b * \<mu> (space M)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   511
      by (simp add: pinfreal_zero_less_diff_iff)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   512
    also have "\<dots> \<le> ?t A0 - b * \<mu> A0"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   513
      using space_less_A0 pos_M pos_t b real[unfolded pinfreal_noteq_omega_Ex] by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   514
    finally have "b * \<mu> A0 < ?t A0" unfolding pinfreal_zero_less_diff_iff .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   515
    hence "0 < ?t A0" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   516
    hence "0 < \<mu> A0" using ac unfolding absolutely_continuous_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   517
      using `A0 \<in> sets M` by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   518
    hence "0 < b * \<mu> A0" using b by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   519
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   520
    from int_f_finite this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   521
    have "?y + 0 < positive_integral f + b * \<mu> A0" unfolding int_f_eq_y
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   522
      by (rule pinfreal_less_add)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   523
    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
   524
      by (simp cong: positive_integral_cong)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   525
    finally have "?y < positive_integral ?f0" by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   526
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   527
    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
   528
    ultimately show False by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   529
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   530
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   531
  show ?thesis
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   532
  proof (safe intro!: bexI[of _ f])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   533
    fix A assume "A\<in>sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   534
    show "\<nu> A = positive_integral (\<lambda>x. f x * indicator A x)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   535
    proof (rule antisym)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   536
      show "positive_integral (\<lambda>x. f x * indicator A x) \<le> \<nu> A"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   537
        using `f \<in> G` `A \<in> sets M` unfolding G_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   538
      show "\<nu> A \<le> positive_integral (\<lambda>x. f x * indicator A x)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   539
        using upper_bound[THEN bspec, OF `A \<in> sets M`]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   540
         by (simp add: pinfreal_zero_le_diff)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   541
    qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   542
  qed simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   543
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   544
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   545
lemma (in finite_measure) Radon_Nikodym_finite_measure_infinite:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   546
  assumes "measure_space M \<nu>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   547
  assumes "absolutely_continuous \<nu>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   548
  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
   549
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   550
  interpret v: measure_space M \<nu> by fact
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   551
  let ?Q = "{Q\<in>sets M. \<nu> Q \<noteq> \<omega>}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   552
  let ?a = "SUP Q:?Q. \<mu> Q"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   553
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   554
  have "{} \<in> ?Q" using v.empty_measure by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   555
  then have Q_not_empty: "?Q \<noteq> {}" by blast
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   556
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   557
  have "?a \<le> \<mu> (space M)" using sets_into_space
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   558
    by (auto intro!: SUP_leI measure_mono top)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   559
  then have "?a \<noteq> \<omega>" using finite_measure_of_space
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   560
    by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   561
  from SUPR_countable_SUPR[OF this Q_not_empty]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   562
  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
   563
    by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   564
  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
   565
  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
   566
    by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   567
  then have a_Lim: "?a = (SUP i::nat. \<mu> (Q' i))" using a by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   568
  let "?O n" = "\<Union>i\<le>n. Q' i"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   569
  have Union: "(SUP i. \<mu> (?O i)) = \<mu> (\<Union>i. ?O i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   570
  proof (rule continuity_from_below[of ?O])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   571
    show "range ?O \<subseteq> sets M" using Q' by (auto intro!: finite_UN)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   572
    show "\<And>i. ?O i \<subseteq> ?O (Suc i)" by fastsimp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   573
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   574
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   575
  have Q'_sets: "\<And>i. Q' i \<in> sets M" using Q' by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   576
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   577
  have O_sets: "\<And>i. ?O i \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   578
     using Q' by (auto intro!: finite_UN Un)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   579
  then have O_in_G: "\<And>i. ?O i \<in> ?Q"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   580
  proof (safe del: notI)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   581
    fix i have "Q' ` {..i} \<subseteq> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   582
      using Q' by (auto intro: finite_UN)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   583
    with v.measure_finitely_subadditive[of "{.. i}" Q']
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   584
    have "\<nu> (?O i) \<le> (\<Sum>i\<le>i. \<nu> (Q' i))" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   585
    also have "\<dots> < \<omega>" unfolding setsum_\<omega> pinfreal_less_\<omega> using Q' by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   586
    finally show "\<nu> (?O i) \<noteq> \<omega>" unfolding pinfreal_less_\<omega> by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   587
  qed auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   588
  have O_mono: "\<And>n. ?O n \<subseteq> ?O (Suc n)" by fastsimp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   589
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   590
  have a_eq: "?a = \<mu> (\<Union>i. ?O i)" unfolding Union[symmetric]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   591
  proof (rule antisym)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   592
    show "?a \<le> (SUP i. \<mu> (?O i))" unfolding a_Lim
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   593
      using Q' by (auto intro!: SUP_mono measure_mono finite_UN)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   594
    show "(SUP i. \<mu> (?O i)) \<le> ?a" unfolding SUPR_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   595
    proof (safe intro!: Sup_mono, unfold bex_simps)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   596
      fix i
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   597
      have *: "(\<Union>Q' ` {..i}) = ?O i" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   598
      then show "\<exists>x. (x \<in> sets M \<and> \<nu> x \<noteq> \<omega>) \<and>
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   599
        \<mu> (\<Union>Q' ` {..i}) \<le> \<mu> x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   600
        using O_in_G[of i] by (auto intro!: exI[of _ "?O i"])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   601
    qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   602
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   603
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   604
  let "?O_0" = "(\<Union>i. ?O i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   605
  have "?O_0 \<in> sets M" using Q' by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   606
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   607
  { fix A assume *: "A \<in> ?Q" "A \<subseteq> space M - ?O_0"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   608
    then have "\<mu> ?O_0 + \<mu> A = \<mu> (?O_0 \<union> A)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   609
      using Q' by (auto intro!: measure_additive countable_UN)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   610
    also have "\<dots> = (SUP i. \<mu> (?O i \<union> A))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   611
    proof (rule continuity_from_below[of "\<lambda>i. ?O i \<union> A", symmetric, simplified])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   612
      show "range (\<lambda>i. ?O i \<union> A) \<subseteq> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   613
        using * O_sets by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   614
    qed fastsimp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   615
    also have "\<dots> \<le> ?a"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   616
    proof (safe intro!: SUPR_bound)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   617
      fix i have "?O i \<union> A \<in> ?Q"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   618
      proof (safe del: notI)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   619
        show "?O i \<union> A \<in> sets M" using O_sets * by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   620
        from O_in_G[of i]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   621
        moreover have "\<nu> (?O i \<union> A) \<le> \<nu> (?O i) + \<nu> A"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   622
          using v.measure_subadditive[of "?O i" A] * O_sets by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   623
        ultimately show "\<nu> (?O i \<union> A) \<noteq> \<omega>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   624
          using * by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   625
      qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   626
      then show "\<mu> (?O i \<union> A) \<le> ?a" by (rule le_SUPI)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   627
    qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   628
    finally have "\<mu> A = 0" unfolding a_eq using finite_measure[OF `?O_0 \<in> sets M`]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   629
      by (cases "\<mu> A") (auto simp: pinfreal_noteq_omega_Ex) }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   630
  note stetic = this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   631
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   632
  def Q \<equiv> "\<lambda>i. case i of 0 \<Rightarrow> ?O 0 | Suc n \<Rightarrow> ?O (Suc n) - ?O n"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   633
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   634
  { 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
   635
  note Q_sets = this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   636
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   637
  { fix i have "\<nu> (Q i) \<noteq> \<omega>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   638
    proof (cases i)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   639
      case 0 then show ?thesis
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   640
        unfolding Q_def using Q'[of 0] by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   641
    next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   642
      case (Suc n)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   643
      then show ?thesis unfolding Q_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   644
        using `?O n \<in> ?Q` `?O (Suc n) \<in> ?Q` O_mono
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   645
        using v.measure_Diff[of "?O n" "?O (Suc n)"] by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   646
    qed }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   647
  note Q_omega = this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   648
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   649
  { fix j have "(\<Union>i\<le>j. ?O i) = (\<Union>i\<le>j. Q i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   650
    proof (induct j)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   651
      case 0 then show ?case by (simp add: Q_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   652
    next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   653
      case (Suc j)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   654
      have eq: "\<And>j. (\<Union>i\<le>j. ?O i) = (\<Union>i\<le>j. Q' i)" by fastsimp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   655
      have "{..j} \<union> {..Suc j} = {..Suc j}" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   656
      then have "(\<Union>i\<le>Suc j. Q' i) = (\<Union>i\<le>j. Q' i) \<union> Q (Suc j)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   657
        by (simp add: UN_Un[symmetric] Q_def del: UN_Un)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   658
      then show ?case using Suc by (auto simp add: eq atMost_Suc)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   659
    qed }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   660
  then have "(\<Union>j. (\<Union>i\<le>j. ?O i)) = (\<Union>j. (\<Union>i\<le>j. Q i))" by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   661
  then have O_0_eq_Q: "?O_0 = (\<Union>j. Q j)" by fastsimp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   662
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   663
  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
   664
    \<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
   665
  proof
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   666
    fix i
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   667
    have indicator_eq: "\<And>f x A. (f x :: pinfreal) * indicator (Q i \<inter> A) x * indicator (Q i) x
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   668
      = (f x * indicator (Q i) x) * indicator A x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   669
      unfolding indicator_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   670
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38656
diff changeset
   671
    have fm: "finite_measure (restricted_space (Q i)) \<mu>"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   672
      (is "finite_measure ?R \<mu>") by (rule restricted_finite_measure[OF Q_sets[of i]])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   673
    then interpret R: finite_measure ?R .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   674
    have fmv: "finite_measure ?R \<nu>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   675
      unfolding finite_measure_def finite_measure_axioms_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   676
    proof
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   677
      show "measure_space ?R \<nu>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   678
        using v.restricted_measure_space Q_sets[of i] by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   679
      show "\<nu>  (space ?R) \<noteq> \<omega>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   680
        using Q_omega by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   681
    qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   682
    have "R.absolutely_continuous \<nu>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   683
      using `absolutely_continuous \<nu>` `Q i \<in> sets M`
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   684
      by (auto simp: R.absolutely_continuous_def absolutely_continuous_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   685
    from finite_measure.Radon_Nikodym_finite_measure[OF fm fmv this]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   686
    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
   687
      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
   688
      unfolding Bex_def borel_measurable_restricted[OF `Q i \<in> sets M`]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   689
        positive_integral_restricted[OF `Q i \<in> sets M`] by (auto simp: indicator_eq)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   690
    then show "\<exists>f. f\<in>borel_measurable M \<and> (\<forall>A\<in>sets M.
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   691
      \<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
   692
      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
   693
          simp: indicator_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   694
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   695
  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
   696
    and f: "\<And>A i. A \<in> sets M \<Longrightarrow>
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   697
      \<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
   698
    by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   699
  let "?f x" =
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   700
    "(\<Sum>\<^isub>\<infinity> i. f i x * indicator (Q i) x) + \<omega> * indicator (space M - ?O_0) x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   701
  show ?thesis
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   702
  proof (safe intro!: bexI[of _ ?f])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   703
    show "?f \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   704
      by (safe intro!: borel_measurable_psuminf borel_measurable_pinfreal_times
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   705
        borel_measurable_pinfreal_add borel_measurable_indicator
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   706
        borel_measurable_const borel Q_sets O_sets Diff countable_UN)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   707
    fix A assume "A \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   708
    let ?C = "(space M - (\<Union>i. Q i)) \<inter> A"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   709
    have *: 
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   710
      "\<And>x i. indicator A x * (f i x * indicator (Q i) x) =
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   711
        f i x * indicator (Q i \<inter> A) x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   712
      "\<And>x i. (indicator A x * indicator (space M - (\<Union>i. UNION {..i} Q')) x :: pinfreal) =
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   713
        indicator ?C x" unfolding O_0_eq_Q by (auto simp: indicator_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   714
    have "positive_integral (\<lambda>x. ?f x * indicator A x) =
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   715
      (\<Sum>\<^isub>\<infinity> i. \<nu> (Q i \<inter> A)) + \<omega> * \<mu> ?C"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   716
      unfolding f[OF `A \<in> sets M`]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   717
      apply (simp del: pinfreal_times(2) add: field_simps)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   718
      apply (subst positive_integral_add)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   719
      apply (safe intro!: borel_measurable_pinfreal_times Diff borel_measurable_const
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   720
        borel_measurable_psuminf borel_measurable_indicator `A \<in> sets M` Q_sets borel countable_UN Q'_sets)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   721
      unfolding psuminf_cmult_right[symmetric]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   722
      apply (subst positive_integral_psuminf)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   723
      apply (safe intro!: borel_measurable_pinfreal_times Diff borel_measurable_const
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   724
        borel_measurable_psuminf borel_measurable_indicator `A \<in> sets M` Q_sets borel countable_UN Q'_sets)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   725
      apply (subst positive_integral_cmult)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   726
      apply (safe intro!: borel_measurable_pinfreal_times Diff borel_measurable_const
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   727
        borel_measurable_psuminf borel_measurable_indicator `A \<in> sets M` Q_sets borel countable_UN Q'_sets)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   728
      unfolding *
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   729
      apply (subst positive_integral_indicator)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   730
      apply (safe intro!: borel_measurable_pinfreal_times Diff borel_measurable_const Int
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   731
        borel_measurable_psuminf borel_measurable_indicator `A \<in> sets M` Q_sets borel countable_UN Q'_sets)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   732
      by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   733
    moreover have "(\<Sum>\<^isub>\<infinity>i. \<nu> (Q i \<inter> A)) = \<nu> ((\<Union>i. Q i) \<inter> A)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   734
    proof (rule v.measure_countably_additive[of "\<lambda>i. Q i \<inter> A", unfolded comp_def, simplified])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   735
      show "range (\<lambda>i. Q i \<inter> A) \<subseteq> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   736
        using Q_sets `A \<in> sets M` by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   737
      show "disjoint_family (\<lambda>i. Q i \<inter> A)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   738
        by (fastsimp simp: disjoint_family_on_def Q_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   739
          split: nat.split_asm)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   740
    qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   741
    moreover have "\<omega> * \<mu> ?C = \<nu> ?C"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   742
    proof cases
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   743
      assume null: "\<mu> ?C = 0"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   744
      hence "?C \<in> null_sets" using Q_sets `A \<in> sets M` by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   745
      with `absolutely_continuous \<nu>` and null
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   746
      show ?thesis by (simp add: absolutely_continuous_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   747
    next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   748
      assume not_null: "\<mu> ?C \<noteq> 0"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   749
      have "\<nu> ?C = \<omega>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   750
      proof (rule ccontr)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   751
        assume "\<nu> ?C \<noteq> \<omega>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   752
        then have "?C \<in> ?Q"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   753
          using Q_sets `A \<in> sets M` by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   754
        from stetic[OF this] not_null
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   755
        show False unfolding O_0_eq_Q by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   756
      qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   757
      then show ?thesis using not_null by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   758
    qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   759
    moreover have "?C \<in> sets M" "((\<Union>i. Q i) \<inter> A) \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   760
      using Q_sets `A \<in> sets M` by (auto intro!: countable_UN)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   761
    moreover have "((\<Union>i. Q i) \<inter> A) \<union> ?C = A" "((\<Union>i. Q i) \<inter> A) \<inter> ?C = {}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   762
      using `A \<in> sets M` sets_into_space by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   763
    ultimately show "\<nu> A = positive_integral (\<lambda>x. ?f x * indicator A x)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   764
      using v.measure_additive[simplified, of "(\<Union>i. Q i) \<inter> A" ?C] by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   765
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   766
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   767
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   768
lemma (in sigma_finite_measure) Radon_Nikodym:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   769
  assumes "measure_space M \<nu>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   770
  assumes "absolutely_continuous \<nu>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   771
  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
   772
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   773
  from Ex_finite_integrable_function
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   774
  obtain h where finite: "positive_integral h \<noteq> \<omega>" and
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   775
    borel: "h \<in> borel_measurable M" and
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   776
    pos: "\<And>x. x \<in> space M \<Longrightarrow> 0 < h x" and
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   777
    "\<And>x. x \<in> space M \<Longrightarrow> h x < \<omega>" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   778
  let "?T A" = "positive_integral (\<lambda>x. h x * indicator A x)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   779
  from measure_space_density[OF borel] finite
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   780
  interpret T: finite_measure M ?T
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   781
    unfolding finite_measure_def finite_measure_axioms_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   782
    by (simp cong: positive_integral_cong)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   783
  have "\<And>N. N \<in> sets M \<Longrightarrow> {x \<in> space M. h x \<noteq> 0 \<and> indicator N x \<noteq> (0::pinfreal)} = N"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   784
    using sets_into_space pos by (force simp: indicator_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   785
  then have "T.absolutely_continuous \<nu>" using assms(2) borel
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   786
    unfolding T.absolutely_continuous_def absolutely_continuous_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   787
    by (fastsimp simp: borel_measurable_indicator positive_integral_0_iff)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   788
  from T.Radon_Nikodym_finite_measure_infinite[simplified, OF assms(1) this]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   789
  obtain f where f_borel: "f \<in> borel_measurable M" and
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   790
    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
   791
  show ?thesis
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   792
  proof (safe intro!: bexI[of _ "\<lambda>x. h x * f x"])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   793
    show "(\<lambda>x. h x * f x) \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   794
      using borel f_borel by (auto intro: borel_measurable_pinfreal_times)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   795
    fix A assume "A \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   796
    then have "(\<lambda>x. f x * indicator A x) \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   797
      using f_borel by (auto intro: borel_measurable_pinfreal_times borel_measurable_indicator)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   798
    from positive_integral_translated_density[OF borel this]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   799
    show "\<nu> A = positive_integral (\<lambda>x. h x * f x * indicator A x)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   800
      unfolding fT[OF `A \<in> sets M`] by (simp add: field_simps)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   801
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   802
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   803
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   804
section "Radon Nikodym derivative"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   805
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   806
definition (in sigma_finite_measure)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   807
  "RN_deriv \<nu> \<equiv> SOME f. f \<in> borel_measurable M \<and>
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   808
    (\<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
   809
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   810
lemma (in sigma_finite_measure) RN_deriv:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   811
  assumes "measure_space M \<nu>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   812
  assumes "absolutely_continuous \<nu>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   813
  shows "RN_deriv \<nu> \<in> borel_measurable M" (is ?borel)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   814
  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
   815
    (is "\<And>A. _ \<Longrightarrow> ?int A")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   816
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   817
  note Ex = Radon_Nikodym[OF assms, unfolded Bex_def]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   818
  thus ?borel unfolding RN_deriv_def by (rule someI2_ex) auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   819
  fix A assume "A \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   820
  from Ex show "?int A" unfolding RN_deriv_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   821
    by (rule someI2_ex) (simp add: `A \<in> sets M`)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   822
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   823
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   824
lemma (in sigma_finite_measure) RN_deriv_singleton:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   825
  assumes "measure_space M \<nu>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   826
  and ac: "absolutely_continuous \<nu>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   827
  and "{x} \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   828
  shows "\<nu> {x} = RN_deriv \<nu> x * \<mu> {x}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   829
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   830
  note deriv = RN_deriv[OF assms(1, 2)]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   831
  from deriv(2)[OF `{x} \<in> sets M`]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   832
  have "\<nu> {x} = positive_integral (\<lambda>w. RN_deriv \<nu> x * indicator {x} w)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   833
    by (auto simp: indicator_def intro!: positive_integral_cong)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   834
  thus ?thesis using positive_integral_cmult_indicator[OF `{x} \<in> sets M`]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   835
    by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   836
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   837
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   838
theorem (in finite_measure_space) RN_deriv_finite_measure:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   839
  assumes "measure_space M \<nu>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   840
  and ac: "absolutely_continuous \<nu>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   841
  and "x \<in> space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   842
  shows "\<nu> {x} = RN_deriv \<nu> x * \<mu> {x}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   843
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   844
  have "{x} \<in> sets M" using sets_eq_Pow `x \<in> space M` by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   845
  from RN_deriv_singleton[OF assms(1,2) this] show ?thesis .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   846
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   847
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   848
end