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