src/HOL/Probability/Radon_Nikodym.thy
author blanchet
Tue, 22 May 2012 16:59:27 +0200
changeset 47953 a2c3706c4cb1
parent 47694 05663f75964c
child 49778 bbbc0f492780
permissions -rw-r--r--
added "ext_cong_neq" lemma (not used yet); tuning
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
42067
66c8281349ec standardized headers
hoelzl
parents: 41981
diff changeset
     1
(*  Title:      HOL/Probability/Radon_Nikodym.thy
66c8281349ec standardized headers
hoelzl
parents: 41981
diff changeset
     2
    Author:     Johannes Hölzl, TU München
66c8281349ec standardized headers
hoelzl
parents: 41981
diff changeset
     3
*)
66c8281349ec standardized headers
hoelzl
parents: 41981
diff changeset
     4
66c8281349ec standardized headers
hoelzl
parents: 41981
diff changeset
     5
header {*Radon-Nikod{\'y}m derivative*}
66c8281349ec standardized headers
hoelzl
parents: 41981
diff changeset
     6
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
     7
theory Radon_Nikodym
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
     8
imports Lebesgue_Integration
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
     9
begin
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    10
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
    11
definition "diff_measure M N =
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
    12
  measure_of (space M) (sets M) (\<lambda>A. emeasure M A - emeasure N A)"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
    13
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
    14
lemma 
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
    15
  shows space_diff_measure[simp]: "space (diff_measure M N) = space M"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
    16
    and sets_diff_measure[simp]: "sets (diff_measure M N) = sets M"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
    17
  by (auto simp: diff_measure_def)
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
    18
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
    19
lemma emeasure_diff_measure:
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
    20
  assumes fin: "finite_measure M" "finite_measure N" and sets_eq: "sets M = sets N"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
    21
  assumes pos: "\<And>A. A \<in> sets M \<Longrightarrow> emeasure N A \<le> emeasure M A" and A: "A \<in> sets M"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
    22
  shows "emeasure (diff_measure M N) A = emeasure M A - emeasure N A" (is "_ = ?\<mu> A")
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
    23
  unfolding diff_measure_def
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
    24
proof (rule emeasure_measure_of_sigma)
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
    25
  show "sigma_algebra (space M) (sets M)" ..
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
    26
  show "positive (sets M) ?\<mu>"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
    27
    using pos by (simp add: positive_def ereal_diff_positive)
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
    28
  show "countably_additive (sets M) ?\<mu>"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
    29
  proof (rule countably_additiveI)
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
    30
    fix A :: "nat \<Rightarrow> _"  assume A: "range A \<subseteq> sets M" and "disjoint_family A"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
    31
    then have suminf:
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
    32
      "(\<Sum>i. emeasure M (A i)) = emeasure M (\<Union>i. A i)"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
    33
      "(\<Sum>i. emeasure N (A i)) = emeasure N (\<Union>i. A i)"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
    34
      by (simp_all add: suminf_emeasure sets_eq)
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
    35
    with A have "(\<Sum>i. emeasure M (A i) - emeasure N (A i)) =
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
    36
      (\<Sum>i. emeasure M (A i)) - (\<Sum>i. emeasure N (A i))"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
    37
      using fin
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
    38
      by (intro suminf_ereal_minus pos emeasure_nonneg)
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
    39
         (auto simp: sets_eq finite_measure.emeasure_eq_measure suminf_emeasure)
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
    40
    then show "(\<Sum>i. emeasure M (A i) - emeasure N (A i)) =
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
    41
      emeasure M (\<Union>i. A i) - emeasure N (\<Union>i. A i) "
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
    42
      by (simp add: suminf)
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
    43
  qed
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
    44
qed fact
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
    45
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    46
lemma (in sigma_finite_measure) Ex_finite_integrable_function:
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
    47
  shows "\<exists>h\<in>borel_measurable M. integral\<^isup>P M h \<noteq> \<infinity> \<and> (\<forall>x\<in>space M. 0 < h x \<and> h x < \<infinity>) \<and> (\<forall>x. 0 \<le> h x)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    48
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    49
  obtain A :: "nat \<Rightarrow> 'a set" where
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    50
    range: "range A \<subseteq> sets M" and
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    51
    space: "(\<Union>i. A i) = space M" and
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
    52
    measure: "\<And>i. emeasure M (A i) \<noteq> \<infinity>" and
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    53
    disjoint: "disjoint_family A"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
    54
    using sigma_finite_disjoint by auto
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
    55
  let ?B = "\<lambda>i. 2^Suc i * emeasure M (A i)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    56
  have "\<forall>i. \<exists>x. 0 < x \<and> x < inverse (?B i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    57
  proof
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
    58
    fix i show "\<exists>x. 0 < x \<and> x < inverse (?B i)"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
    59
      using measure[of i] emeasure_nonneg[of M "A i"]
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
    60
      by (auto intro!: ereal_dense simp: ereal_0_gt_inverse ereal_zero_le_0_iff)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    61
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    62
  from choice[OF this] obtain n where n: "\<And>i. 0 < n i"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
    63
    "\<And>i. n i < inverse (2^Suc i * emeasure M (A i))" by auto
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
    64
  { fix i have "0 \<le> n i" using n(1)[of i] by auto } note pos = this
46731
5302e932d1e5 avoid undeclared variables in let bindings;
wenzelm
parents: 45777
diff changeset
    65
  let ?h = "\<lambda>x. \<Sum>i. n i * indicator (A i) x"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    66
  show ?thesis
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    67
  proof (safe intro!: bexI[of _ ?h] del: notI)
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38656
diff changeset
    68
    have "\<And>i. A i \<in> sets M"
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 44568
diff changeset
    69
      using range by fastforce+
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
    70
    then have "integral\<^isup>P M ?h = (\<Sum>i. n i * emeasure M (A i))" using pos
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
    71
      by (simp add: positive_integral_suminf positive_integral_cmult_indicator)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
    72
    also have "\<dots> \<le> (\<Sum>i. (1 / 2)^Suc i)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
    73
    proof (rule suminf_le_pos)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
    74
      fix N
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
    75
      have "n N * emeasure M (A N) \<le> inverse (2^Suc N * emeasure M (A N)) * emeasure M (A N)"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
    76
        using n[of N]
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43556
diff changeset
    77
        by (intro ereal_mult_right_mono) auto
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
    78
      also have "\<dots> \<le> (1 / 2) ^ Suc N"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    79
        using measure[of N] n[of N]
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
    80
        by (cases rule: ereal2_cases[of "n N" "emeasure M (A N)"])
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43556
diff changeset
    81
           (simp_all add: inverse_eq_divide power_divide one_ereal_def ereal_power_divide)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
    82
      finally show "n N * emeasure M (A N) \<le> (1 / 2) ^ Suc N" .
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
    83
      show "0 \<le> n N * emeasure M (A N)" using n[of N] `A N \<in> sets M` by (simp add: emeasure_nonneg)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    84
    qed
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43556
diff changeset
    85
    finally show "integral\<^isup>P M ?h \<noteq> \<infinity>" unfolding suminf_half_series_ereal by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    86
  next
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
    87
    { fix x assume "x \<in> space M"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
    88
      then obtain i where "x \<in> A i" using space[symmetric] by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
    89
      with disjoint n have "?h x = n i"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
    90
        by (auto intro!: suminf_cmult_indicator intro: less_imp_le)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
    91
      then show "0 < ?h x" and "?h x < \<infinity>" using n[of i] by auto }
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
    92
    note pos = this
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
    93
    fix x show "0 \<le> ?h x"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
    94
    proof cases
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
    95
      assume "x \<in> space M" then show "0 \<le> ?h x" using pos by (auto intro: less_imp_le)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
    96
    next
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
    97
      assume "x \<notin> space M" then have "\<And>i. x \<notin> A i" using space by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
    98
      then show "0 \<le> ?h x" by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
    99
    qed
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   100
  next
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   101
    show "?h \<in> borel_measurable M" using range n
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43556
diff changeset
   102
      by (auto intro!: borel_measurable_psuminf borel_measurable_ereal_times ereal_0_le_mult intro: less_imp_le)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   103
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   104
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   105
40871
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
   106
subsection "Absolutely continuous"
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
   107
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   108
definition absolutely_continuous :: "'a measure \<Rightarrow> 'a measure \<Rightarrow> bool" where
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   109
  "absolutely_continuous M N \<longleftrightarrow> null_sets M \<subseteq> null_sets N"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   110
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   111
lemma absolutely_continuousI_count_space: "absolutely_continuous (count_space A) M"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   112
  unfolding absolutely_continuous_def by (auto simp: null_sets_count_space)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   113
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   114
lemma absolutely_continuousI_density:
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   115
  "f \<in> borel_measurable M \<Longrightarrow> absolutely_continuous M (density M f)"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   116
  by (force simp add: absolutely_continuous_def null_sets_density_iff dest: AE_not_in)
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   117
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   118
lemma absolutely_continuousI_point_measure_finite:
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   119
  "(\<And>x. \<lbrakk> x \<in> A ; f x \<le> 0 \<rbrakk> \<Longrightarrow> g x \<le> 0) \<Longrightarrow> absolutely_continuous (point_measure A f) (point_measure A g)"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   120
  unfolding absolutely_continuous_def by (force simp: null_sets_point_measure_iff)
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   121
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   122
lemma absolutely_continuous_AE:
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   123
  assumes sets_eq: "sets M' = sets M"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   124
    and "absolutely_continuous M M'" "AE x in M. P x"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   125
   shows "AE x in M'. P x"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   126
proof -
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   127
  from `AE x in M. P x` obtain N where N: "N \<in> null_sets M" "{x\<in>space M. \<not> P x} \<subseteq> N"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   128
    unfolding eventually_ae_filter by auto
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   129
  show "AE x in M'. P x"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   130
  proof (rule AE_I')
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   131
    show "{x\<in>space M'. \<not> P x} \<subseteq> N" using sets_eq_imp_space_eq[OF sets_eq] N(2) by simp
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   132
    from `absolutely_continuous M M'` show "N \<in> null_sets M'"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   133
      using N unfolding absolutely_continuous_def sets_eq null_sets_def by auto
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   134
  qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   135
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   136
40871
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
   137
subsection "Existence of the Radon-Nikodym derivative"
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
   138
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   139
lemma (in finite_measure) Radon_Nikodym_aux_epsilon:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   140
  fixes e :: real assumes "0 < e"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   141
  assumes "finite_measure N" and sets_eq: "sets N = sets M"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   142
  shows "\<exists>A\<in>sets M. measure M (space M) - measure N (space M) \<le> measure M A - measure N A \<and>
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   143
                    (\<forall>B\<in>sets M. B \<subseteq> A \<longrightarrow> - e < measure M B - measure N B)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   144
proof -
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   145
  interpret M': finite_measure N by fact
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   146
  let ?d = "\<lambda>A. measure M A - measure N A"
46731
5302e932d1e5 avoid undeclared variables in let bindings;
wenzelm
parents: 45777
diff changeset
   147
  let ?A = "\<lambda>A. if (\<forall>B\<in>sets M. B \<subseteq> space M - A \<longrightarrow> -e < ?d B)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   148
    then {}
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   149
    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
   150
  def A \<equiv> "\<lambda>n. ((\<lambda>B. B \<union> ?A B) ^^ n) {}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   151
  have A_simps[simp]:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   152
    "A 0 = {}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   153
    "\<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
   154
  { fix A assume "A \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   155
    have "?A A \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   156
      by (auto intro!: someI2[of _ _ "\<lambda>A. A \<in> sets M"] simp: not_less) }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   157
  note A'_in_sets = this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   158
  { fix n have "A n \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   159
    proof (induct n)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   160
      case (Suc n) thus "A (Suc n) \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   161
        using A'_in_sets[of "A n"] by (auto split: split_if_asm)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   162
    qed (simp add: A_def) }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   163
  note A_in_sets = this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   164
  hence "range A \<subseteq> sets M" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   165
  { fix n B
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   166
    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
   167
    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
   168
    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
   169
    proof (rule someI2_ex[OF Ex])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   170
      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
   171
      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
   172
      hence "?d (A n \<union> B) = ?d (A n) + ?d B"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   173
        using `A n \<in> sets M` finite_measure_Union M'.finite_measure_Union by (simp add: sets_eq)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   174
      also have "\<dots> \<le> ?d (A n) - e" using dB by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   175
      finally show "?d (A n \<union> B) \<le> ?d (A n) - e" .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   176
    qed }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   177
  note dA_epsilon = this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   178
  { fix n have "?d (A (Suc n)) \<le> ?d (A n)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   179
    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
   180
      case True from dA_epsilon[OF this] show ?thesis using `0 < e` by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   181
    next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   182
      case False
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   183
      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
   184
      thus ?thesis by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   185
    qed }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   186
  note dA_mono = this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   187
  show ?thesis
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   188
  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
   189
    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
   190
    show ?thesis
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   191
    proof (safe intro!: bexI[of _ "space M - A n"])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   192
      fix B assume "B \<in> sets M" "B \<subseteq> space M - A n"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   193
      from B[OF this] show "-e < ?d B" .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   194
    next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   195
      show "space M - A n \<in> sets M" by (rule compl_sets) fact
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   196
    next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   197
      show "?d (space M) \<le> ?d (space M - A n)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   198
      proof (induct n)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   199
        fix n assume "?d (space M) \<le> ?d (space M - A n)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   200
        also have "\<dots> \<le> ?d (space M - A (Suc n))"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   201
          using A_in_sets sets_into_space dA_mono[of n] finite_measure_compl M'.finite_measure_compl
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   202
          by (simp del: A_simps add: sets_eq sets_eq_imp_space_eq[OF sets_eq])
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   203
        finally show "?d (space M) \<le> ?d (space M - A (Suc n))" .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   204
      qed simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   205
    qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   206
  next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   207
    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
   208
      by (auto simp add: not_less)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   209
    { fix n have "?d (A n) \<le> - real n * e"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   210
      proof (induct n)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   211
        case (Suc n) with dA_epsilon[of n, OF B] show ?case by (simp del: A_simps add: real_of_nat_Suc field_simps)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   212
      next
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   213
        case 0 with measure_empty show ?case by (simp add: zero_ereal_def)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   214
      qed } note dA_less = this
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   215
    have decseq: "decseq (\<lambda>n. ?d (A n))" unfolding decseq_eq_incseq
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   216
    proof (rule incseq_SucI)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   217
      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
   218
    qed
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   219
    have A: "incseq A" by (auto intro!: incseq_SucI)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   220
    from finite_Lim_measure_incseq[OF _ A] `range A \<subseteq> sets M`
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   221
      M'.finite_Lim_measure_incseq[OF _ A]
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   222
    have convergent: "(\<lambda>i. ?d (A i)) ----> ?d (\<Union>i. A i)"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   223
      by (auto intro!: tendsto_diff simp: sets_eq)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   224
    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
   225
    moreover from order_trans[OF decseq_le[OF decseq convergent] dA_less]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   226
    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
   227
    ultimately show ?thesis by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   228
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   229
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   230
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   231
lemma (in finite_measure) Radon_Nikodym_aux:
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   232
  assumes "finite_measure N" and sets_eq: "sets N = sets M"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   233
  shows "\<exists>A\<in>sets M. measure M (space M) - measure N (space M) \<le>
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   234
                    measure M A - measure N A \<and>
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   235
                    (\<forall>B\<in>sets M. B \<subseteq> A \<longrightarrow> 0 \<le> measure M B - measure N B)"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   236
proof -
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   237
  interpret N: finite_measure N by fact
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   238
  let ?d = "\<lambda>A. measure M A - measure N A"
46731
5302e932d1e5 avoid undeclared variables in let bindings;
wenzelm
parents: 45777
diff changeset
   239
  let ?P = "\<lambda>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)"
5302e932d1e5 avoid undeclared variables in let bindings;
wenzelm
parents: 45777
diff changeset
   240
  let ?r = "\<lambda>S. restricted_space S"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   241
  { fix S n assume S: "S \<in> sets M"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   242
    then have "finite_measure (density M (indicator S))" "0 < 1 / real (Suc n)"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   243
         "finite_measure (density N (indicator S))" "sets (density N (indicator S)) = sets (density M (indicator S))"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   244
      by (auto simp: finite_measure_restricted N.finite_measure_restricted sets_eq)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   245
    from finite_measure.Radon_Nikodym_aux_epsilon[OF this] guess X .. note X = this
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   246
    with S have "?P (S \<inter> X) S n"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   247
      by (simp add: measure_restricted sets_eq Int) (metis inf_absorb2)
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   248
    hence "\<exists>A. ?P A S n" .. }
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   249
  note Ex_P = this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   250
  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
   251
  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
   252
  have A_0[simp]: "A 0 = space M" unfolding A_def by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   253
  { fix i have "A i \<in> sets M" unfolding A_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   254
    proof (induct i)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   255
      case (Suc i)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   256
      from Ex_P[OF this, of i] show ?case unfolding nat_rec_Suc
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   257
        by (rule someI2_ex) simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   258
    qed simp }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   259
  note A_in_sets = this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   260
  { fix n have "?P (A (Suc n)) (A n) n"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   261
      using Ex_P[OF A_in_sets] unfolding A_Suc
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   262
      by (rule someI2_ex) simp }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   263
  note P_A = this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   264
  have "range A \<subseteq> sets M" using A_in_sets by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   265
  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
   266
  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
   267
  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
   268
      using P_A by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   269
  show ?thesis
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   270
  proof (safe intro!: bexI[of _ "\<Inter>i. A i"])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   271
    show "(\<Inter>i. A i) \<in> sets M" using A_in_sets by auto
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   272
    have A: "decseq A" using A_mono by (auto intro!: decseq_SucI)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   273
    from `range A \<subseteq> sets M`
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   274
      finite_Lim_measure_decseq[OF _ A] N.finite_Lim_measure_decseq[OF _ A]
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   275
    have "(\<lambda>i. ?d (A i)) ----> ?d (\<Inter>i. A i)" by (auto intro!: tendsto_diff simp: sets_eq)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   276
    thus "?d (space M) \<le> ?d (\<Inter>i. A i)" using mono_dA[THEN monoD, of 0 _]
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   277
      by (rule_tac LIMSEQ_le_const) auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   278
  next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   279
    fix B assume B: "B \<in> sets M" "B \<subseteq> (\<Inter>i. A i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   280
    show "0 \<le> ?d B"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   281
    proof (rule ccontr)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   282
      assume "\<not> 0 \<le> ?d B"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   283
      hence "0 < - ?d B" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   284
      from ex_inverse_of_nat_Suc_less[OF this]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   285
      obtain n where *: "?d B < - 1 / real (Suc n)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   286
        by (auto simp: real_eq_of_nat inverse_eq_divide field_simps)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   287
      have "B \<subseteq> A (Suc n)" using B by (auto simp del: nat_rec_Suc)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   288
      from epsilon[OF B(1) this] *
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   289
      show False by auto
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
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   293
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   294
lemma (in finite_measure) Radon_Nikodym_finite_measure:
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   295
  assumes "finite_measure N" and sets_eq: "sets N = sets M"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   296
  assumes "absolutely_continuous M N"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   297
  shows "\<exists>f \<in> borel_measurable M. (\<forall>x. 0 \<le> f x) \<and> density M f = N"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   298
proof -
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   299
  interpret N: finite_measure N by fact
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   300
  def G \<equiv> "{g \<in> borel_measurable M. (\<forall>x. 0 \<le> g x) \<and> (\<forall>A\<in>sets M. (\<integral>\<^isup>+x. g x * indicator A x \<partial>M) \<le> N A)}"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   301
  have "(\<lambda>x. 0) \<in> G" unfolding G_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   302
  hence "G \<noteq> {}" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   303
  { fix f g assume f: "f \<in> G" and g: "g \<in> G"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   304
    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
   305
    proof safe
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   306
      show "?max \<in> borel_measurable M" using f g unfolding G_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   307
      let ?A = "{x \<in> space M. f x \<le> g x}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   308
      have "?A \<in> sets M" using f g unfolding G_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   309
      fix A assume "A \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   310
      hence sets: "?A \<inter> A \<in> sets M" "(space M - ?A) \<inter> A \<in> sets M" using `?A \<in> sets M` by auto
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   311
      hence sets': "?A \<inter> A \<in> sets N" "(space M - ?A) \<inter> A \<in> sets N" by (auto simp: sets_eq)
38656
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
46731
5302e932d1e5 avoid undeclared variables in let bindings;
wenzelm
parents: 45777
diff changeset
   321
        by (auto cong: positive_integral_cong intro!: positive_integral_add)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   322
      also have "\<dots> \<le> N (?A \<inter> A) + N ((space M - ?A) \<inter> A)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   323
        using f g sets unfolding G_def by (auto intro!: add_mono)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   324
      also have "\<dots> = N A"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   325
        using plus_emeasure[OF sets'] union by auto
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   326
      finally show "(\<integral>\<^isup>+x. max (g x) (f x) * indicator A x \<partial>M) \<le> N A" .
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   327
    next
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   328
      fix x show "0 \<le> max (g x) (f x)" using f g by (auto simp: G_def split: split_max)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   329
    qed }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   330
  note max_in_G = this
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   331
  { fix f assume  "incseq f" and f: "\<And>i. f i \<in> G"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   332
    have "(\<lambda>x. SUP i. f i x) \<in> G" unfolding G_def
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   333
    proof safe
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   334
      show "(\<lambda>x. SUP i. f i x) \<in> borel_measurable M"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   335
        using f by (auto simp: G_def)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   336
      { fix x show "0 \<le> (SUP i. f i x)"
44928
7ef6505bde7f renamed Complete_Lattices lemmas, removed legacy names
hoelzl
parents: 44918
diff changeset
   337
          using f by (auto simp: G_def intro: SUP_upper2) }
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   338
    next
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   339
      fix A assume "A \<in> sets M"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   340
      have "(\<integral>\<^isup>+x. (SUP i. f i x) * indicator A x \<partial>M) =
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   341
        (\<integral>\<^isup>+x. (SUP i. f i x * indicator A x) \<partial>M)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   342
        by (intro positive_integral_cong) (simp split: split_indicator)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   343
      also have "\<dots> = (SUP i. (\<integral>\<^isup>+x. f i x * indicator A x \<partial>M))"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   344
        using `incseq f` f `A \<in> sets M`
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   345
        by (intro positive_integral_monotone_convergence_SUP)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   346
           (auto simp: G_def incseq_Suc_iff le_fun_def split: split_indicator)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   347
      finally show "(\<integral>\<^isup>+x. (SUP i. f i x) * indicator A x \<partial>M) \<le> N A"
44928
7ef6505bde7f renamed Complete_Lattices lemmas, removed legacy names
hoelzl
parents: 44918
diff changeset
   348
        using f `A \<in> sets M` by (auto intro!: SUP_least simp: G_def)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   349
    qed }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   350
  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
   351
  let ?y = "SUP g : G. integral\<^isup>P M g"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   352
  have y_le: "?y \<le> N (space M)" unfolding G_def
44928
7ef6505bde7f renamed Complete_Lattices lemmas, removed legacy names
hoelzl
parents: 44918
diff changeset
   353
  proof (safe intro!: SUP_least)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   354
    fix g assume "\<forall>A\<in>sets M. (\<integral>\<^isup>+x. g x * indicator A x \<partial>M) \<le> N A"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   355
    from this[THEN bspec, OF top] show "integral\<^isup>P M g \<le> N (space M)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   356
      by (simp cong: positive_integral_cong)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   357
  qed
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   358
  from SUPR_countable_SUPR[OF `G \<noteq> {}`, of "integral\<^isup>P M"] guess ys .. note ys = this
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   359
  then have "\<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
   360
  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
   361
    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
   362
    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
   363
    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
   364
  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
   365
  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
   366
  hence y_eq: "?y = (SUP i. integral\<^isup>P M (gs i))" using ys by auto
46731
5302e932d1e5 avoid undeclared variables in let bindings;
wenzelm
parents: 45777
diff changeset
   367
  let ?g = "\<lambda>i x. Max ((\<lambda>n. gs n x) ` {..i})"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   368
  def f \<equiv> "\<lambda>x. SUP i. ?g i x"
46731
5302e932d1e5 avoid undeclared variables in let bindings;
wenzelm
parents: 45777
diff changeset
   369
  let ?F = "\<lambda>A x. f x * indicator A x"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   370
  have gs_not_empty: "\<And>i x. (\<lambda>n. gs n x) ` {..i} \<noteq> {}" by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   371
  { fix i have "?g i \<in> G"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   372
    proof (induct i)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   373
      case 0 thus ?case by simp fact
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   374
    next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   375
      case (Suc i)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   376
      with Suc gs_not_empty `gs (Suc i) \<in> G` show ?case
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   377
        by (auto simp add: atMost_Suc intro!: max_in_G)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   378
    qed }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   379
  note g_in_G = this
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   380
  have "incseq ?g" using gs_not_empty
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   381
    by (auto intro!: incseq_SucI le_funI simp add: atMost_Suc)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   382
  from SUP_in_G[OF this g_in_G] have "f \<in> G" unfolding f_def .
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   383
  then have [simp, intro]: "f \<in> borel_measurable M" unfolding G_def by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   384
  have "integral\<^isup>P M f = (SUP i. integral\<^isup>P M (?g i))" unfolding f_def
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   385
    using g_in_G `incseq ?g`
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   386
    by (auto intro!: positive_integral_monotone_convergence_SUP simp: G_def)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   387
  also have "\<dots> = ?y"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   388
  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
   389
    show "(SUP i. integral\<^isup>P M (?g i)) \<le> ?y"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   390
      using g_in_G by (auto intro: Sup_mono simp: SUP_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
   391
    show "?y \<le> (SUP i. integral\<^isup>P M (?g i))" unfolding y_eq
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   392
      by (auto intro!: SUP_mono positive_integral_mono Max_ge)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   393
  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
   394
  finally have int_f_eq_y: "integral\<^isup>P M f = ?y" .
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   395
  have "\<And>x. 0 \<le> f x"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   396
    unfolding f_def using `\<And>i. gs i \<in> G`
44928
7ef6505bde7f renamed Complete_Lattices lemmas, removed legacy names
hoelzl
parents: 44918
diff changeset
   397
    by (auto intro!: SUP_upper2 Max_ge_iff[THEN iffD2] simp: G_def)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   398
  let ?t = "\<lambda>A. N A - (\<integral>\<^isup>+x. ?F A x \<partial>M)"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   399
  let ?M = "diff_measure N (density M f)"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   400
  have f_le_N: "\<And>A. A \<in> sets M \<Longrightarrow> (\<integral>\<^isup>+x. ?F A x \<partial>M) \<le> N A"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   401
    using `f \<in> G` unfolding G_def by auto
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   402
  have emeasure_M: "\<And>A. A \<in> sets M \<Longrightarrow> emeasure ?M A = ?t A"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   403
  proof (subst emeasure_diff_measure)
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   404
    from f_le_N[of "space M"] show "finite_measure N" "finite_measure (density M f)"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   405
      by (auto intro!: finite_measureI simp: emeasure_density cong: positive_integral_cong)
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   406
  next
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   407
    fix B assume "B \<in> sets N" with f_le_N[of B] show "emeasure (density M f) B \<le> emeasure N B"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   408
      by (auto simp: sets_eq emeasure_density cong: positive_integral_cong)
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   409
  qed (auto simp: sets_eq emeasure_density)
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   410
  from emeasure_M[of "space M"] N.finite_emeasure_space positive_integral_positive[of M "?F (space M)"]
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   411
  interpret M': finite_measure ?M
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   412
    by (auto intro!: finite_measureI simp: sets_eq_imp_space_eq[OF sets_eq] N.emeasure_eq_measure ereal_minus_eq_PInfty_iff)
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   413
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   414
  have ac: "absolutely_continuous M ?M" unfolding absolutely_continuous_def
45777
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45769
diff changeset
   415
  proof
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   416
    fix A assume A: "A \<in> null_sets M"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   417
    with `absolutely_continuous M N` have "A \<in> null_sets N"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   418
      unfolding absolutely_continuous_def by auto
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   419
    moreover with A have "(\<integral>\<^isup>+ x. ?F A x \<partial>M) \<le> N A" using `f \<in> G` by (auto simp: G_def)
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   420
    ultimately have "N A - (\<integral>\<^isup>+ x. ?F A x \<partial>M) = 0"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   421
      using positive_integral_positive[of M] by (auto intro!: antisym)
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   422
    then show "A \<in> null_sets ?M"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   423
      using A by (simp add: emeasure_M null_sets_def sets_eq)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   424
  qed
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   425
  have upper_bound: "\<forall>A\<in>sets M. ?M A \<le> 0"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   426
  proof (rule ccontr)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   427
    assume "\<not> ?thesis"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   428
    then obtain A where A: "A \<in> sets M" and pos: "0 < ?M A"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   429
      by (auto simp: not_le)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   430
    note pos
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   431
    also have "?M A \<le> ?M (space M)"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   432
      using emeasure_space[of ?M A] by (simp add: sets_eq[THEN sets_eq_imp_space_eq])
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   433
    finally have pos_t: "0 < ?M (space M)" by simp
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   434
    moreover
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   435
    then have "emeasure M (space M) \<noteq> 0"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   436
      using ac unfolding absolutely_continuous_def by (auto simp: null_sets_def)
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   437
    then have pos_M: "0 < emeasure M (space M)"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   438
      using emeasure_nonneg[of M "space M"] by (simp add: le_less)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   439
    moreover
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   440
    have "(\<integral>\<^isup>+x. f x * indicator (space M) x \<partial>M) \<le> N (space M)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   441
      using `f \<in> G` unfolding G_def by auto
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   442
    hence "(\<integral>\<^isup>+x. f x * indicator (space M) x \<partial>M) \<noteq> \<infinity>"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   443
      using M'.finite_emeasure_space by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   444
    moreover
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   445
    def b \<equiv> "?M (space M) / emeasure M (space M) / 2"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   446
    ultimately have b: "b \<noteq> 0 \<and> 0 \<le> b \<and> b \<noteq> \<infinity>"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   447
      by (auto simp: ereal_divide_eq)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   448
    then have b: "b \<noteq> 0" "0 \<le> b" "0 < b"  "b \<noteq> \<infinity>" by auto
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   449
    let ?Mb = "density M (\<lambda>_. b)"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   450
    have Mb: "finite_measure ?Mb" "sets ?Mb = sets ?M"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   451
        using b by (auto simp: emeasure_density_const sets_eq intro!: finite_measureI)
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   452
    from M'.Radon_Nikodym_aux[OF this] guess A0 ..
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   453
    then have "A0 \<in> sets M"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   454
      and space_less_A0: "measure ?M (space M) - real b * measure M (space M) \<le> measure ?M A0 - real b * measure M A0"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   455
      and *: "\<And>B. B \<in> sets M \<Longrightarrow> B \<subseteq> A0 \<Longrightarrow> 0 \<le> measure ?M B - real b * measure M B"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   456
      using b by (simp_all add: measure_density_const sets_eq_imp_space_eq[OF sets_eq] sets_eq)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   457
    { fix B assume B: "B \<in> sets M" "B \<subseteq> A0"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   458
      with *[OF this] have "b * emeasure M B \<le> ?M B"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   459
        using b unfolding M'.emeasure_eq_measure emeasure_eq_measure by (cases b) auto }
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   460
    note bM_le_t = this
46731
5302e932d1e5 avoid undeclared variables in let bindings;
wenzelm
parents: 45777
diff changeset
   461
    let ?f0 = "\<lambda>x. f x + b * indicator A0 x"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   462
    { fix A assume A: "A \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   463
      hence "A \<inter> A0 \<in> sets M" using `A0 \<in> sets M` by auto
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   464
      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
   465
        (\<integral>\<^isup>+x. f x * indicator A x + b * indicator (A \<inter> A0) x \<partial>M)"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   466
        by (auto intro!: positive_integral_cong split: split_indicator)
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
   467
      hence "(\<integral>\<^isup>+x. ?f0 x * indicator A x \<partial>M) =
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   468
          (\<integral>\<^isup>+x. f x * indicator A x \<partial>M) + b * emeasure M (A \<inter> A0)"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   469
        using `A0 \<in> sets M` `A \<inter> A0 \<in> sets M` A b `f \<in> G`
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   470
        by (simp add: G_def positive_integral_add positive_integral_cmult_indicator) }
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   471
    note f0_eq = this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   472
    { fix A assume A: "A \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   473
      hence "A \<inter> A0 \<in> sets M" using `A0 \<in> sets M` by auto
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   474
      have f_le_v: "(\<integral>\<^isup>+x. ?F A x \<partial>M) \<le> N A" using `f \<in> G` A unfolding G_def by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   475
      note f0_eq[OF A]
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   476
      also have "(\<integral>\<^isup>+x. ?F A x \<partial>M) + b * emeasure M (A \<inter> A0) \<le> (\<integral>\<^isup>+x. ?F A x \<partial>M) + ?M (A \<inter> A0)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   477
        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
   478
        by (auto intro!: add_left_mono)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   479
      also have "\<dots> \<le> (\<integral>\<^isup>+x. f x * indicator A x \<partial>M) + ?M A"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   480
        using emeasure_mono[of "A \<inter> A0" A ?M] `A \<in> sets M` `A0 \<in> sets M`
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   481
        by (auto intro!: add_left_mono simp: sets_eq)
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   482
      also have "\<dots> \<le> N A"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   483
        unfolding emeasure_M[OF `A \<in> sets M`]
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   484
        using f_le_v N.emeasure_eq_measure[of A] positive_integral_positive[of M "?F A"]
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   485
        by (cases "\<integral>\<^isup>+x. ?F A x \<partial>M", cases "N A") auto
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   486
      finally have "(\<integral>\<^isup>+x. ?f0 x * indicator A x \<partial>M) \<le> N A" . }
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   487
    hence "?f0 \<in> G" using `A0 \<in> sets M` b `f \<in> G` unfolding G_def
46731
5302e932d1e5 avoid undeclared variables in let bindings;
wenzelm
parents: 45777
diff changeset
   488
      by (auto intro!: ereal_add_nonneg_nonneg)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   489
    have int_f_finite: "integral\<^isup>P M f \<noteq> \<infinity>"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   490
      by (metis N.emeasure_finite ereal_infty_less_eq2(1) int_f_eq_y y_le)
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   491
    have  "0 < ?M (space M) - emeasure ?Mb (space M)"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   492
      using pos_t
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   493
      by (simp add: b emeasure_density_const)
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   494
         (simp add: M'.emeasure_eq_measure emeasure_eq_measure pos_M b_def)
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   495
    also have "\<dots> \<le> ?M A0 - b * emeasure M A0"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   496
      using space_less_A0 `A0 \<in> sets M` b
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   497
      by (cases b) (auto simp add: b emeasure_density_const sets_eq M'.emeasure_eq_measure emeasure_eq_measure)
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   498
    finally have 1: "b * emeasure M A0 < ?M A0"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   499
      by (metis M'.emeasure_real `A0 \<in> sets M` bM_le_t diff_self ereal_less(1) ereal_minus(1)
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   500
                less_eq_ereal_def mult_zero_left not_square_less_zero subset_refl zero_ereal_def)
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   501
    with b have "0 < ?M A0"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   502
      by (metis M'.emeasure_real MInfty_neq_PInfty(1) emeasure_real ereal_less_eq(5) ereal_zero_times
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   503
               ereal_mult_eq_MInfty ereal_mult_eq_PInfty ereal_zero_less_0_iff less_eq_ereal_def)
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   504
    then have "emeasure M A0 \<noteq> 0" using ac `A0 \<in> sets M`
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   505
      by (auto simp: absolutely_continuous_def null_sets_def)
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   506
    then have "0 < emeasure M A0" using emeasure_nonneg[of M A0] by auto
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   507
    hence "0 < b * emeasure M A0" using b by (auto simp: ereal_zero_less_0_iff)
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   508
    with int_f_finite have "?y + 0 < integral\<^isup>P M f + b * emeasure M A0" unfolding int_f_eq_y
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   509
      using `f \<in> G`
44928
7ef6505bde7f renamed Complete_Lattices lemmas, removed legacy names
hoelzl
parents: 44918
diff changeset
   510
      by (intro ereal_add_strict_mono) (auto intro!: SUP_upper2 positive_integral_positive)
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
   511
    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
   512
      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
   513
    finally have "?y < integral\<^isup>P M ?f0" by simp
44928
7ef6505bde7f renamed Complete_Lattices lemmas, removed legacy names
hoelzl
parents: 44918
diff changeset
   514
    moreover from `?f0 \<in> G` have "integral\<^isup>P M ?f0 \<le> ?y" by (auto intro!: SUP_upper)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   515
    ultimately show False by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   516
  qed
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   517
  let ?f = "\<lambda>x. max 0 (f x)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   518
  show ?thesis
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   519
  proof (intro bexI[of _ ?f] measure_eqI conjI)
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   520
    show "sets (density M ?f) = sets N"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   521
      by (simp add: sets_eq)
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   522
    fix A assume A: "A\<in>sets (density M ?f)"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   523
    then show "emeasure (density M ?f) A = emeasure N A"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   524
      using `f \<in> G` A upper_bound[THEN bspec, of A] N.emeasure_eq_measure[of A]
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   525
      by (cases "integral\<^isup>P M (?F A)")
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   526
         (auto intro!: antisym simp add: emeasure_density G_def emeasure_M density_max_0[symmetric])
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   527
  qed auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   528
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   529
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   530
lemma (in finite_measure) split_space_into_finite_sets_and_rest:
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   531
  assumes ac: "absolutely_continuous M N" and sets_eq: "sets N = sets M"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   532
  shows "\<exists>A0\<in>sets M. \<exists>B::nat\<Rightarrow>'a set. disjoint_family B \<and> range B \<subseteq> sets M \<and> A0 = space M - (\<Union>i. B i) \<and>
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   533
    (\<forall>A\<in>sets M. A \<subseteq> A0 \<longrightarrow> (emeasure M A = 0 \<and> N A = 0) \<or> (emeasure M A > 0 \<and> N A = \<infinity>)) \<and>
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   534
    (\<forall>i. N (B i) \<noteq> \<infinity>)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   535
proof -
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   536
  let ?Q = "{Q\<in>sets M. N Q \<noteq> \<infinity>}"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   537
  let ?a = "SUP Q:?Q. emeasure M Q"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   538
  have "{} \<in> ?Q" by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   539
  then have Q_not_empty: "?Q \<noteq> {}" by blast
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   540
  have "?a \<le> emeasure M (space M)" using sets_into_space
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   541
    by (auto intro!: SUP_least emeasure_mono)
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   542
  then have "?a \<noteq> \<infinity>" using finite_emeasure_space
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   543
    by auto
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   544
  from SUPR_countable_SUPR[OF Q_not_empty, of "emeasure M"]
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   545
  obtain Q'' where "range Q'' \<subseteq> emeasure M ` ?Q" and a: "?a = (SUP i::nat. Q'' i)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   546
    by auto
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   547
  then have "\<forall>i. \<exists>Q'. Q'' i = emeasure M Q' \<and> Q' \<in> ?Q" by auto
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   548
  from choice[OF this] obtain Q' where Q': "\<And>i. Q'' i = emeasure M (Q' i)" "\<And>i. Q' i \<in> ?Q"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   549
    by auto
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   550
  then have a_Lim: "?a = (SUP i::nat. emeasure M (Q' i))" using a by simp
46731
5302e932d1e5 avoid undeclared variables in let bindings;
wenzelm
parents: 45777
diff changeset
   551
  let ?O = "\<lambda>n. \<Union>i\<le>n. Q' i"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   552
  have Union: "(SUP i. emeasure M (?O i)) = emeasure M (\<Union>i. ?O i)"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   553
  proof (rule SUP_emeasure_incseq[of ?O])
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   554
    show "range ?O \<subseteq> sets M" using Q' by auto
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 44568
diff changeset
   555
    show "incseq ?O" by (fastforce intro!: incseq_SucI)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   556
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   557
  have Q'_sets: "\<And>i. Q' i \<in> sets M" using Q' by auto
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   558
  have O_sets: "\<And>i. ?O i \<in> sets M" using Q' by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   559
  then have O_in_G: "\<And>i. ?O i \<in> ?Q"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   560
  proof (safe del: notI)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   561
    fix i have "Q' ` {..i} \<subseteq> sets M" using Q' by auto
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   562
    then have "N (?O i) \<le> (\<Sum>i\<le>i. N (Q' i))"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   563
      by (simp add: sets_eq emeasure_subadditive_finite)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   564
    also have "\<dots> < \<infinity>" using Q' by (simp add: setsum_Pinfty)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   565
    finally show "N (?O i) \<noteq> \<infinity>" by simp
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   566
  qed auto
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 44568
diff changeset
   567
  have O_mono: "\<And>n. ?O n \<subseteq> ?O (Suc n)" by fastforce
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   568
  have a_eq: "?a = emeasure M (\<Union>i. ?O i)" unfolding Union[symmetric]
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   569
  proof (rule antisym)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   570
    show "?a \<le> (SUP i. emeasure M (?O i))" unfolding a_Lim
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   571
      using Q' by (auto intro!: SUP_mono emeasure_mono)
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   572
    show "(SUP i. emeasure M (?O i)) \<le> ?a" unfolding SUP_def
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   573
    proof (safe intro!: Sup_mono, unfold bex_simps)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   574
      fix i
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   575
      have *: "(\<Union>Q' ` {..i}) = ?O i" by auto
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   576
      then show "\<exists>x. (x \<in> sets M \<and> N x \<noteq> \<infinity>) \<and>
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   577
        emeasure M (\<Union>Q' ` {..i}) \<le> emeasure M x"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   578
        using O_in_G[of i] by (auto intro!: exI[of _ "?O i"])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   579
    qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   580
  qed
46731
5302e932d1e5 avoid undeclared variables in let bindings;
wenzelm
parents: 45777
diff changeset
   581
  let ?O_0 = "(\<Union>i. ?O i)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   582
  have "?O_0 \<in> sets M" using Q' by auto
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   583
  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
   584
  { 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
   585
  note Q_sets = this
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   586
  show ?thesis
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   587
  proof (intro bexI exI conjI ballI impI allI)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   588
    show "disjoint_family Q"
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 44568
diff changeset
   589
      by (fastforce simp: disjoint_family_on_def Q_def
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   590
        split: nat.split_asm)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   591
    show "range Q \<subseteq> sets M"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   592
      using Q_sets by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   593
    { fix A assume A: "A \<in> sets M" "A \<subseteq> space M - ?O_0"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   594
      show "emeasure M A = 0 \<and> N A = 0 \<or> 0 < emeasure M A \<and> N A = \<infinity>"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   595
      proof (rule disjCI, simp)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   596
        assume *: "0 < emeasure M A \<longrightarrow> N A \<noteq> \<infinity>"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   597
        show "emeasure M A = 0 \<and> N A = 0"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   598
        proof cases
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   599
          assume "emeasure M A = 0" moreover with ac A have "N A = 0"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   600
            unfolding absolutely_continuous_def by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   601
          ultimately show ?thesis by simp
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   602
        next
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   603
          assume "emeasure M A \<noteq> 0" with * have "N A \<noteq> \<infinity>" using emeasure_nonneg[of M A] by auto
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   604
          with A have "emeasure M ?O_0 + emeasure M A = emeasure M (?O_0 \<union> A)"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   605
            using Q' by (auto intro!: plus_emeasure countable_UN)
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   606
          also have "\<dots> = (SUP i. emeasure M (?O i \<union> A))"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   607
          proof (rule SUP_emeasure_incseq[of "\<lambda>i. ?O i \<union> A", symmetric, simplified])
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   608
            show "range (\<lambda>i. ?O i \<union> A) \<subseteq> sets M"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   609
              using `N A \<noteq> \<infinity>` O_sets A by auto
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 44568
diff changeset
   610
          qed (fastforce intro!: incseq_SucI)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   611
          also have "\<dots> \<le> ?a"
44928
7ef6505bde7f renamed Complete_Lattices lemmas, removed legacy names
hoelzl
parents: 44918
diff changeset
   612
          proof (safe intro!: SUP_least)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   613
            fix i have "?O i \<union> A \<in> ?Q"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   614
            proof (safe del: notI)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   615
              show "?O i \<union> A \<in> sets M" using O_sets A by auto
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   616
              from O_in_G[of i] have "N (?O i \<union> A) \<le> N (?O i) + N A"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   617
                using emeasure_subadditive[of "?O i" N A] A O_sets by (auto simp: sets_eq)
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   618
              with O_in_G[of i] show "N (?O i \<union> A) \<noteq> \<infinity>"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   619
                using `N A \<noteq> \<infinity>` by auto
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   620
            qed
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   621
            then show "emeasure M (?O i \<union> A) \<le> ?a" by (rule SUP_upper)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   622
          qed
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   623
          finally have "emeasure M A = 0"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   624
            unfolding a_eq using measure_nonneg[of M A] by (simp add: emeasure_eq_measure)
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   625
          with `emeasure M A \<noteq> 0` show ?thesis by auto
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   626
        qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   627
      qed }
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   628
    { fix i show "N (Q i) \<noteq> \<infinity>"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   629
      proof (cases i)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   630
        case 0 then show ?thesis
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   631
          unfolding Q_def using Q'[of 0] by simp
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   632
      next
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   633
        case (Suc n)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   634
        with `?O n \<in> ?Q` `?O (Suc n) \<in> ?Q`
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   635
            emeasure_Diff[OF _ _ _ O_mono, of N n] emeasure_nonneg[of N "(\<Union> x\<le>n. Q' x)"]
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   636
        show ?thesis
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   637
          by (auto simp: sets_eq ereal_minus_eq_PInfty_iff Q_def)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   638
      qed }
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   639
    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
   640
    { 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
   641
      proof (induct j)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   642
        case 0 then show ?case by (simp add: Q_def)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   643
      next
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   644
        case (Suc j)
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 44568
diff changeset
   645
        have eq: "\<And>j. (\<Union>i\<le>j. ?O i) = (\<Union>i\<le>j. Q' i)" by fastforce
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   646
        have "{..j} \<union> {..Suc j} = {..Suc j}" by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   647
        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
   648
          by (simp add: UN_Un[symmetric] Q_def del: UN_Un)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   649
        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
   650
      qed }
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   651
    then have "(\<Union>j. (\<Union>i\<le>j. ?O i)) = (\<Union>j. (\<Union>i\<le>j. Q i))" by simp
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 44568
diff changeset
   652
    then show "space M - ?O_0 = space M - (\<Union>i. Q i)" by fastforce
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   653
  qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   654
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   655
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   656
lemma (in finite_measure) Radon_Nikodym_finite_measure_infinite:
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   657
  assumes "absolutely_continuous M N" and sets_eq: "sets N = sets M"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   658
  shows "\<exists>f\<in>borel_measurable M. (\<forall>x. 0 \<le> f x) \<and> density M f = N"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   659
proof -
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   660
  from split_space_into_finite_sets_and_rest[OF assms]
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   661
  obtain Q0 and Q :: "nat \<Rightarrow> 'a set"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   662
    where Q: "disjoint_family Q" "range Q \<subseteq> sets M"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   663
    and Q0: "Q0 \<in> sets M" "Q0 = space M - (\<Union>i. Q i)"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   664
    and in_Q0: "\<And>A. A \<in> sets M \<Longrightarrow> A \<subseteq> Q0 \<Longrightarrow> emeasure M A = 0 \<and> N A = 0 \<or> 0 < emeasure M A \<and> N A = \<infinity>"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   665
    and Q_fin: "\<And>i. N (Q i) \<noteq> \<infinity>" by force
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   666
  from Q have Q_sets: "\<And>i. Q i \<in> sets M" by auto
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   667
  let ?N = "\<lambda>i. density N (indicator (Q i))" and ?M = "\<lambda>i. density M (indicator (Q i))"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   668
  have "\<forall>i. \<exists>f\<in>borel_measurable (?M i). (\<forall>x. 0 \<le> f x) \<and> density (?M i) f = ?N i"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   669
  proof (intro allI finite_measure.Radon_Nikodym_finite_measure)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   670
    fix i
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   671
    from Q show "finite_measure (?M i)"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   672
      by (auto intro!: finite_measureI cong: positive_integral_cong
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   673
               simp add: emeasure_density subset_eq sets_eq)
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   674
    from Q have "emeasure (?N i) (space N) = emeasure N (Q i)"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   675
      by (simp add: sets_eq[symmetric] emeasure_density subset_eq cong: positive_integral_cong)
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   676
    with Q_fin show "finite_measure (?N i)"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   677
      by (auto intro!: finite_measureI)
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   678
    show "sets (?N i) = sets (?M i)" by (simp add: sets_eq)
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   679
    show "absolutely_continuous (?M i) (?N i)"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   680
      using `absolutely_continuous M N` `Q i \<in> sets M`
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   681
      by (auto simp: absolutely_continuous_def null_sets_density_iff sets_eq
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   682
               intro!: absolutely_continuous_AE[OF sets_eq])
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   683
  qed
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   684
  from choice[OF this[unfolded Bex_def]]
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   685
  obtain f where borel: "\<And>i. f i \<in> borel_measurable M" "\<And>i x. 0 \<le> f i x"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   686
    and f_density: "\<And>i. density (?M i) (f i) = ?N i"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   687
    by auto
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   688
  { fix A i assume A: "A \<in> sets M"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   689
    with Q borel have "(\<integral>\<^isup>+x. f i x * indicator (Q i \<inter> A) x \<partial>M) = emeasure (density (?M i) (f i)) A"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   690
      by (auto simp add: emeasure_density positive_integral_density subset_eq
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   691
               intro!: positive_integral_cong split: split_indicator)
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   692
    also have "\<dots> = emeasure N (Q i \<inter> A)"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   693
      using A Q by (simp add: f_density emeasure_restricted subset_eq sets_eq)
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   694
    finally have "emeasure N (Q i \<inter> A) = (\<integral>\<^isup>+x. f i x * indicator (Q i \<inter> A) x \<partial>M)" .. }
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   695
  note integral_eq = this
46731
5302e932d1e5 avoid undeclared variables in let bindings;
wenzelm
parents: 45777
diff changeset
   696
  let ?f = "\<lambda>x. (\<Sum>i. f i x * indicator (Q i) x) + \<infinity> * indicator Q0 x"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   697
  show ?thesis
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   698
  proof (safe intro!: bexI[of _ ?f])
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   699
    show "?f \<in> borel_measurable M" using Q0 borel Q_sets
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   700
      by (auto intro!: measurable_If)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   701
    show "\<And>x. 0 \<le> ?f x" using borel by (auto intro!: suminf_0_le simp: indicator_def)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   702
    show "density M ?f = N"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   703
    proof (rule measure_eqI)
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   704
      fix A assume "A \<in> sets (density M ?f)"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   705
      then have "A \<in> sets M" by simp
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   706
      have Qi: "\<And>i. Q i \<in> sets M" using Q by auto
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   707
      have [intro,simp]: "\<And>i. (\<lambda>x. f i x * indicator (Q i \<inter> A) x) \<in> borel_measurable M"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   708
        "\<And>i. AE x in M. 0 \<le> f i x * indicator (Q i \<inter> A) x"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   709
        using borel Qi Q0(1) `A \<in> sets M` by (auto intro!: borel_measurable_ereal_times)
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   710
      have "(\<integral>\<^isup>+x. ?f x * indicator A x \<partial>M) = (\<integral>\<^isup>+x. (\<Sum>i. f i x * indicator (Q i \<inter> A) x) + \<infinity> * indicator (Q0 \<inter> A) x \<partial>M)"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   711
        using borel by (intro positive_integral_cong) (auto simp: indicator_def)
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   712
      also have "\<dots> = (\<integral>\<^isup>+x. (\<Sum>i. f i x * indicator (Q i \<inter> A) x) \<partial>M) + \<infinity> * emeasure M (Q0 \<inter> A)"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   713
        using borel Qi Q0(1) `A \<in> sets M`
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   714
        by (subst positive_integral_add) (auto simp del: ereal_infty_mult
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   715
            simp add: positive_integral_cmult_indicator Int intro!: suminf_0_le)
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   716
      also have "\<dots> = (\<Sum>i. N (Q i \<inter> A)) + \<infinity> * emeasure M (Q0 \<inter> A)"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   717
        by (subst integral_eq[OF `A \<in> sets M`], subst positive_integral_suminf) auto
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   718
      finally have "(\<integral>\<^isup>+x. ?f x * indicator A x \<partial>M) = (\<Sum>i. N (Q i \<inter> A)) + \<infinity> * emeasure M (Q0 \<inter> A)" .
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   719
      moreover have "(\<Sum>i. N (Q i \<inter> A)) = N ((\<Union>i. Q i) \<inter> A)"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   720
        using Q Q_sets `A \<in> sets M`
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   721
        by (subst suminf_emeasure) (auto simp: disjoint_family_on_def sets_eq)
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   722
      moreover have "\<infinity> * emeasure M (Q0 \<inter> A) = N (Q0 \<inter> A)"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   723
      proof -
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   724
        have "Q0 \<inter> A \<in> sets M" using Q0(1) `A \<in> sets M` by blast
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   725
        from in_Q0[OF this] show ?thesis by auto
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   726
      qed
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   727
      moreover have "Q0 \<inter> A \<in> sets M" "((\<Union>i. Q i) \<inter> A) \<in> sets M"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   728
        using Q_sets `A \<in> sets M` Q0(1) by auto
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   729
      moreover have "((\<Union>i. Q i) \<inter> A) \<union> (Q0 \<inter> A) = A" "((\<Union>i. Q i) \<inter> A) \<inter> (Q0 \<inter> A) = {}"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   730
        using `A \<in> sets M` sets_into_space Q0 by auto
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   731
      ultimately have "N A = (\<integral>\<^isup>+x. ?f x * indicator A x \<partial>M)"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   732
        using plus_emeasure[of "(\<Union>i. Q i) \<inter> A" N "Q0 \<inter> A"] by (simp add: sets_eq)
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   733
      with `A \<in> sets M` borel Q Q0(1) show "emeasure (density M ?f) A = N A"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   734
        by (subst emeasure_density)
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   735
           (auto intro!: borel_measurable_ereal_add borel_measurable_psuminf measurable_If
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   736
                 simp: subset_eq)
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   737
    qed (simp add: sets_eq)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   738
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   739
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   740
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   741
lemma (in sigma_finite_measure) Radon_Nikodym:
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   742
  assumes ac: "absolutely_continuous M N" assumes sets_eq: "sets N = sets M"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   743
  shows "\<exists>f \<in> borel_measurable M. (\<forall>x. 0 \<le> f x) \<and> density M f = N"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   744
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   745
  from Ex_finite_integrable_function
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   746
  obtain h where finite: "integral\<^isup>P M h \<noteq> \<infinity>" and
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   747
    borel: "h \<in> borel_measurable M" and
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   748
    nn: "\<And>x. 0 \<le> h x" and
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   749
    pos: "\<And>x. x \<in> space M \<Longrightarrow> 0 < h x" and
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   750
    "\<And>x. x \<in> space M \<Longrightarrow> h x < \<infinity>" by auto
46731
5302e932d1e5 avoid undeclared variables in let bindings;
wenzelm
parents: 45777
diff changeset
   751
  let ?T = "\<lambda>A. (\<integral>\<^isup>+x. h x * indicator A x \<partial>M)"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   752
  let ?MT = "density M h"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   753
  from borel finite nn interpret T: finite_measure ?MT
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   754
    by (auto intro!: finite_measureI cong: positive_integral_cong simp: emeasure_density)
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   755
  have "absolutely_continuous ?MT N" "sets N = sets ?MT"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   756
  proof (unfold absolutely_continuous_def, safe)
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   757
    fix A assume "A \<in> null_sets ?MT"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   758
    with borel have "A \<in> sets M" "AE x in M. x \<in> A \<longrightarrow> h x \<le> 0"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   759
      by (auto simp add: null_sets_density_iff)
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   760
    with pos sets_into_space have "AE x in M. x \<notin> A"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   761
      by (elim eventually_elim1) (auto simp: not_le[symmetric])
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   762
    then have "A \<in> null_sets M"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   763
      using `A \<in> sets M` by (simp add: AE_iff_null_sets)
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   764
    with ac show "A \<in> null_sets N"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   765
      by (auto simp: absolutely_continuous_def)
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   766
  qed (auto simp add: sets_eq)
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   767
  from T.Radon_Nikodym_finite_measure_infinite[OF this]
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   768
  obtain f where f_borel: "f \<in> borel_measurable M" "\<And>x. 0 \<le> f x" "density ?MT f = N" by auto
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   769
  with nn borel show ?thesis
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   770
    by (auto intro!: bexI[of _ "\<lambda>x. h x * f x"] simp: density_density_eq)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   771
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   772
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   773
section "Uniqueness of densities"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   774
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   775
lemma finite_density_unique:
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   776
  assumes borel: "f \<in> borel_measurable M" "g \<in> borel_measurable M"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   777
  assumes pos: "AE x in M. 0 \<le> f x" "AE x in M. 0 \<le> g x"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   778
  and fin: "integral\<^isup>P M f \<noteq> \<infinity>"
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
   779
  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))
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   780
    \<longleftrightarrow> (AE x in M. f x = g x)"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   781
    (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
   782
proof (intro iffI ballI)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   783
  fix A assume eq: "AE x in M. f x = g x"
41705
1100512e16d8 add auto support for AE_mp
hoelzl
parents: 41689
diff changeset
   784
  then show "?P f A = ?P g A"
1100512e16d8 add auto support for AE_mp
hoelzl
parents: 41689
diff changeset
   785
    by (auto intro: positive_integral_cong_AE)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   786
next
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   787
  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
   788
  from this[THEN bspec, OF top] fin
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   789
  have g_fin: "integral\<^isup>P M g \<noteq> \<infinity>" by (simp cong: positive_integral_cong)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   790
  { fix f g assume borel: "f \<in> borel_measurable M" "g \<in> borel_measurable M"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   791
      and pos: "AE x in M. 0 \<le> f x" "AE x in M. 0 \<le> g x"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   792
      and g_fin: "integral\<^isup>P M g \<noteq> \<infinity>" 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
   793
    let ?N = "{x\<in>space M. g x < f x}"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   794
    have N: "?N \<in> sets M" using borel by simp
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   795
    have "?P g ?N \<le> integral\<^isup>P M g" using pos
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   796
      by (intro positive_integral_mono_AE) (auto split: split_indicator)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   797
    then have Pg_fin: "?P g ?N \<noteq> \<infinity>" using g_fin 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
   798
    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
   799
      by (auto intro!: positive_integral_cong simp: indicator_def)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   800
    also have "\<dots> = ?P f ?N - ?P g ?N"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   801
    proof (rule positive_integral_diff)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   802
      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
   803
        using borel N by auto
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   804
      show "AE x in M. g x * indicator ?N x \<le> f x * indicator ?N x"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   805
           "AE x in M. 0 \<le> g x * indicator ?N x"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   806
        using pos by (auto split: split_indicator)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   807
    qed fact
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   808
    also have "\<dots> = 0"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   809
      unfolding eq[THEN bspec, OF N] using positive_integral_positive[of M] Pg_fin by auto
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   810
    finally have "AE x in M. f x \<le> g x"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   811
      using pos borel positive_integral_PInf_AE[OF borel(2) g_fin]
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   812
      by (subst (asm) positive_integral_0_iff_AE)
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43556
diff changeset
   813
         (auto split: split_indicator simp: not_less ereal_minus_le_iff) }
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   814
  from this[OF borel pos g_fin eq] this[OF borel(2,1) pos(2,1) fin] eq
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   815
  show "AE x in M. f x = g x" by auto
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   816
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   817
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   818
lemma (in finite_measure) density_unique_finite_measure:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   819
  assumes borel: "f \<in> borel_measurable M" "f' \<in> borel_measurable M"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   820
  assumes pos: "AE x in M. 0 \<le> f x" "AE x in M. 0 \<le> f' x"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   821
  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
   822
    (is "\<And>A. A \<in> sets M \<Longrightarrow> ?P f A = ?P f' A")
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   823
  shows "AE x in M. f x = f' x"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   824
proof -
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   825
  let ?D = "\<lambda>f. density M f"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   826
  let ?N = "\<lambda>A. ?P f A" and ?N' = "\<lambda>A. ?P f' A"
46731
5302e932d1e5 avoid undeclared variables in let bindings;
wenzelm
parents: 45777
diff changeset
   827
  let ?f = "\<lambda>A x. f x * indicator A x" and ?f' = "\<lambda>A x. f' x * indicator A x"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   828
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   829
  have ac: "absolutely_continuous M (density M f)" "sets (density M f) = sets M"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   830
    using borel by (auto intro!: absolutely_continuousI_density) 
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   831
  from split_space_into_finite_sets_and_rest[OF this]
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   832
  obtain Q0 and Q :: "nat \<Rightarrow> 'a set"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   833
    where Q: "disjoint_family Q" "range Q \<subseteq> sets M"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   834
    and Q0: "Q0 \<in> sets M" "Q0 = space M - (\<Union>i. Q i)"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   835
    and in_Q0: "\<And>A. A \<in> sets M \<Longrightarrow> A \<subseteq> Q0 \<Longrightarrow> emeasure M A = 0 \<and> ?D f A = 0 \<or> 0 < emeasure M A \<and> ?D f A = \<infinity>"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   836
    and Q_fin: "\<And>i. ?D f (Q i) \<noteq> \<infinity>" by force
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   837
  with borel pos have in_Q0: "\<And>A. A \<in> sets M \<Longrightarrow> A \<subseteq> Q0 \<Longrightarrow> emeasure M A = 0 \<and> ?N A = 0 \<or> 0 < emeasure M A \<and> ?N A = \<infinity>"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   838
    and Q_fin: "\<And>i. ?N (Q i) \<noteq> \<infinity>" by (auto simp: emeasure_density subset_eq)
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   839
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   840
  from Q have Q_sets: "\<And>i. Q i \<in> sets M" by auto
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   841
  let ?D = "{x\<in>space M. f x \<noteq> f' x}"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   842
  have "?D \<in> sets M" using borel by auto
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43556
diff changeset
   843
  have *: "\<And>i x A. \<And>y::ereal. 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
   844
    unfolding indicator_def by auto
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   845
  have "\<forall>i. AE x in M. ?f (Q i) x = ?f' (Q i) x" using borel Q_fin Q pos
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   846
    by (intro finite_density_unique[THEN iffD1] allI)
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43556
diff changeset
   847
       (auto intro!: borel_measurable_ereal_times f Int simp: *)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   848
  moreover have "AE x in M. ?f Q0 x = ?f' Q0 x"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   849
  proof (rule AE_I')
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43556
diff changeset
   850
    { fix f :: "'a \<Rightarrow> ereal" assume borel: "f \<in> borel_measurable M"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   851
        and eq: "\<And>A. A \<in> sets M \<Longrightarrow> ?N A = (\<integral>\<^isup>+x. f x * indicator A x \<partial>M)"
46731
5302e932d1e5 avoid undeclared variables in let bindings;
wenzelm
parents: 45777
diff changeset
   852
      let ?A = "\<lambda>i. Q0 \<inter> {x \<in> space M. f x < (i::nat)}"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   853
      have "(\<Union>i. ?A i) \<in> null_sets M"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   854
      proof (rule null_sets_UN)
43923
ab93d0190a5d add ereal to typeclass infinity
hoelzl
parents: 43920
diff changeset
   855
        fix i ::nat have "?A i \<in> sets M"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   856
          using borel Q0(1) by auto
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   857
        have "?N (?A i) \<le> (\<integral>\<^isup>+x. (i::ereal) * indicator (?A i) x \<partial>M)"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   858
          unfolding eq[OF `?A i \<in> sets M`]
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   859
          by (auto intro!: positive_integral_mono simp: indicator_def)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   860
        also have "\<dots> = i * emeasure M (?A i)"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   861
          using `?A i \<in> sets M` by (auto intro!: positive_integral_cmult_indicator)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   862
        also have "\<dots> < \<infinity>" using emeasure_real[of "?A i"] by simp
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   863
        finally have "?N (?A i) \<noteq> \<infinity>" by simp
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   864
        then show "?A i \<in> null_sets M" using in_Q0[OF `?A i \<in> sets M`] `?A i \<in> sets M` by auto
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   865
      qed
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   866
      also have "(\<Union>i. ?A i) = Q0 \<inter> {x\<in>space M. f x \<noteq> \<infinity>}"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   867
        by (auto simp: less_PInf_Ex_of_nat real_eq_of_nat)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   868
      finally have "Q0 \<inter> {x\<in>space M. f x \<noteq> \<infinity>} \<in> null_sets M" by simp }
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   869
    from this[OF borel(1) refl] this[OF borel(2) f]
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   870
    have "Q0 \<inter> {x\<in>space M. f x \<noteq> \<infinity>} \<in> null_sets M" "Q0 \<inter> {x\<in>space M. f' x \<noteq> \<infinity>} \<in> null_sets M" by simp_all
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   871
    then show "(Q0 \<inter> {x\<in>space M. f x \<noteq> \<infinity>}) \<union> (Q0 \<inter> {x\<in>space M. f' x \<noteq> \<infinity>}) \<in> null_sets M" by (rule null_sets.Un)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   872
    show "{x \<in> space M. ?f Q0 x \<noteq> ?f' Q0 x} \<subseteq>
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   873
      (Q0 \<inter> {x\<in>space M. f x \<noteq> \<infinity>}) \<union> (Q0 \<inter> {x\<in>space M. f' x \<noteq> \<infinity>})" by (auto simp: indicator_def)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   874
  qed
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   875
  moreover have "AE x in M. (?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
   876
    ?f (space M) x = ?f' (space M) x"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   877
    by (auto simp: indicator_def Q0)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   878
  ultimately have "AE x in M. ?f (space M) x = ?f' (space M) x"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   879
    unfolding AE_all_countable[symmetric]
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   880
    by eventually_elim (auto intro!: AE_I2 split: split_if_asm simp: indicator_def)
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   881
  then show "AE x in M. f x = f' x" by auto
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   882
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   883
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   884
lemma (in sigma_finite_measure) density_unique:
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   885
  assumes f: "f \<in> borel_measurable M" "AE x in M. 0 \<le> f x"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   886
  assumes f': "f' \<in> borel_measurable M" "AE x in M. 0 \<le> f' x"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   887
  assumes density_eq: "density M f = density M f'"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   888
  shows "AE x in M. f x = f' x"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   889
proof -
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   890
  obtain h where h_borel: "h \<in> borel_measurable M"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   891
    and fin: "integral\<^isup>P M h \<noteq> \<infinity>" and pos: "\<And>x. x \<in> space M \<Longrightarrow> 0 < h x \<and> h x < \<infinity>" "\<And>x. 0 \<le> h x"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   892
    using Ex_finite_integrable_function by auto
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   893
  then have h_nn: "AE x in M. 0 \<le> h x" by auto
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   894
  let ?H = "density M h"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   895
  interpret h: finite_measure ?H
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   896
    using fin h_borel pos
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   897
    by (intro finite_measureI) (simp cong: positive_integral_cong emeasure_density add: fin)
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   898
  let ?fM = "density M f"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   899
  let ?f'M = "density M f'"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   900
  { fix A assume "A \<in> sets M"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   901
    then have "{x \<in> space M. h x * indicator A x \<noteq> 0} = A"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   902
      using pos(1) sets_into_space by (force simp: indicator_def)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   903
    then have "(\<integral>\<^isup>+x. h x * indicator A x \<partial>M) = 0 \<longleftrightarrow> A \<in> null_sets M"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   904
      using h_borel `A \<in> sets M` h_nn by (subst positive_integral_0_iff) auto }
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   905
  note h_null_sets = this
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   906
  { fix A assume "A \<in> sets M"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   907
    have "(\<integral>\<^isup>+x. f x * (h x * indicator A x) \<partial>M) = (\<integral>\<^isup>+x. h x * indicator A x \<partial>?fM)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   908
      using `A \<in> sets M` h_borel h_nn f f'
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   909
      by (intro positive_integral_density[symmetric]) 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
   910
    also have "\<dots> = (\<integral>\<^isup>+x. h x * indicator A x \<partial>?f'M)"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   911
      by (simp_all add: density_eq)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   912
    also have "\<dots> = (\<integral>\<^isup>+x. f' x * (h x * indicator A x) \<partial>M)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   913
      using `A \<in> sets M` h_borel h_nn f f'
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   914
      by (intro positive_integral_density) auto
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   915
    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)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   916
      by (simp add: ac_simps)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   917
    then have "(\<integral>\<^isup>+x. (f x * indicator A x) \<partial>?H) = (\<integral>\<^isup>+x. (f' x * indicator A x) \<partial>?H)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   918
      using `A \<in> sets M` h_borel h_nn f f'
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   919
      by (subst (asm) (1 2) positive_integral_density[symmetric]) auto }
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   920
  then have "AE x in ?H. f x = f' x" using h_borel h_nn f f'
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   921
    by (intro h.density_unique_finite_measure absolutely_continuous_AE[of M])
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   922
       (auto simp add: AE_density)
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   923
  then show "AE x in M. f x = f' x"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   924
    unfolding eventually_ae_filter using h_borel pos
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   925
    by (auto simp add: h_null_sets null_sets_density_iff not_less[symmetric]
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   926
                          AE_iff_null_sets[symmetric])
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   927
       blast
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   928
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   929
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   930
lemma (in sigma_finite_measure) density_unique_iff:
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   931
  assumes f: "f \<in> borel_measurable M" and "AE x in M. 0 \<le> f x"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   932
  assumes f': "f' \<in> borel_measurable M" and "AE x in M. 0 \<le> f' x"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   933
  shows "density M f = density M f' \<longleftrightarrow> (AE x in M. f x = f' x)"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   934
  using density_unique[OF assms] density_cong[OF f f'] by auto
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   935
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   936
lemma (in sigma_finite_measure) sigma_finite_iff_density_finite:
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   937
  assumes f: "f \<in> borel_measurable M" "AE x in M. 0 \<le> f x"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   938
  shows "sigma_finite_measure (density M f) \<longleftrightarrow> (AE x in M. f x \<noteq> \<infinity>)"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   939
    (is "sigma_finite_measure ?N \<longleftrightarrow> _")
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   940
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
   941
  assume "sigma_finite_measure ?N"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   942
  then interpret N: sigma_finite_measure ?N .
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   943
  from N.Ex_finite_integrable_function obtain h where
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   944
    h: "h \<in> borel_measurable M" "integral\<^isup>P ?N h \<noteq> \<infinity>" and
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   945
    h_nn: "\<And>x. 0 \<le> h x" and
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   946
    fin: "\<forall>x\<in>space M. 0 < h x \<and> h x < \<infinity>" by auto
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   947
  have "AE x in M. f x * h x \<noteq> \<infinity>"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   948
  proof (rule AE_I')
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   949
    have "integral\<^isup>P ?N h = (\<integral>\<^isup>+x. f x * h x \<partial>M)" using f h h_nn
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   950
      by (auto intro!: positive_integral_density)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   951
    then have "(\<integral>\<^isup>+x. f x * h x \<partial>M) \<noteq> \<infinity>"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   952
      using h(2) by simp
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   953
    then show "(\<lambda>x. f x * h x) -` {\<infinity>} \<inter> space M \<in> null_sets M"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   954
      using f h(1) by (auto intro!: positive_integral_PInf borel_measurable_vimage)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   955
  qed auto
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   956
  then show "AE x in M. f x \<noteq> \<infinity>"
41705
1100512e16d8 add auto support for AE_mp
hoelzl
parents: 41689
diff changeset
   957
    using fin by (auto elim!: AE_Ball_mp)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   958
next
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   959
  assume AE: "AE x in M. f x \<noteq> \<infinity>"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   960
  from sigma_finite guess Q .. note Q = this
43923
ab93d0190a5d add ereal to typeclass infinity
hoelzl
parents: 43920
diff changeset
   961
  def A \<equiv> "\<lambda>i. f -` (case i of 0 \<Rightarrow> {\<infinity>} | Suc n \<Rightarrow> {.. ereal(of_nat (Suc n))}) \<inter> space M"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   962
  { 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
   963
    unfolding A_def using f Q
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   964
    apply (rule_tac Int)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   965
    by (cases i) (auto intro: measurable_sets[OF f(1)]) }
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   966
  note A_in_sets = this
46731
5302e932d1e5 avoid undeclared variables in let bindings;
wenzelm
parents: 45777
diff changeset
   967
  let ?A = "\<lambda>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
   968
  show "sigma_finite_measure ?N"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   969
  proof (default, intro exI conjI subsetI allI)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   970
    fix x assume "x \<in> range ?A"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   971
    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
   972
    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
   973
  next
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   974
    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
   975
    proof safe
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   976
      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
   977
      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
   978
        by (intro UN_I[of "prod_encode (i,j)"]) auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   979
    qed auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   980
    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
   981
    also have "(\<Union>i. A i) = space M"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   982
    proof safe
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   983
      fix x assume x: "x \<in> space M"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   984
      show "x \<in> (\<Union>i. A i)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   985
      proof (cases "f x")
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   986
        case PInf with x show ?thesis unfolding A_def by (auto intro: exI[of _ 0])
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   987
      next
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   988
        case (real r)
43923
ab93d0190a5d add ereal to typeclass infinity
hoelzl
parents: 43920
diff changeset
   989
        with less_PInf_Ex_of_nat[of "f x"] obtain n :: nat where "f x < n" by (auto simp: real_eq_of_nat)
45769
2d5b1af2426a real is better supported than real_of_nat, use it in the nat => ereal coercion
hoelzl
parents: 44941
diff changeset
   990
        then show ?thesis using x real unfolding A_def by (auto intro!: exI[of _ "Suc n"] simp: real_eq_of_nat)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   991
      next
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   992
        case MInf with x show ?thesis
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   993
          unfolding A_def by (auto intro!: exI[of _ "Suc 0"])
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   994
      qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   995
    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
   996
    finally show "(\<Union>i. ?A i) = space ?N" by simp
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   997
  next
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   998
    fix n obtain i j where
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   999
      [simp]: "prod_decode n = (i, j)" by (cases "prod_decode n") auto
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
  1000
    have "(\<integral>\<^isup>+x. f x * indicator (A i \<inter> Q j) x \<partial>M) \<noteq> \<infinity>"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1001
    proof (cases i)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1002
      case 0
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1003
      have "AE x in M. f x * indicator (A i \<inter> Q j) x = 0"
41705
1100512e16d8 add auto support for AE_mp
hoelzl
parents: 41689
diff changeset
  1004
        using AE by (auto simp: A_def `i = 0`)
1100512e16d8 add auto support for AE_mp
hoelzl
parents: 41689
diff changeset
  1005
      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
  1006
    next
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1007
      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
  1008
      then have "(\<integral>\<^isup>+x. f x * indicator (A i \<inter> Q j) x \<partial>M) \<le>
43923
ab93d0190a5d add ereal to typeclass infinity
hoelzl
parents: 43920
diff changeset
  1009
        (\<integral>\<^isup>+x. (Suc n :: ereal) * indicator (Q j) x \<partial>M)"
45769
2d5b1af2426a real is better supported than real_of_nat, use it in the nat => ereal coercion
hoelzl
parents: 44941
diff changeset
  1010
        by (auto intro!: positive_integral_mono simp: indicator_def A_def real_eq_of_nat)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1011
      also have "\<dots> = Suc n * emeasure M (Q j)"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1012
        using Q by (auto intro!: positive_integral_cmult_indicator)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
  1013
      also have "\<dots> < \<infinity>"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
  1014
        using Q by (auto simp: real_eq_of_nat[symmetric])
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1015
      finally show ?thesis by simp
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1016
    qed
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1017
    then show "emeasure ?N (?A n) \<noteq> \<infinity>"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1018
      using A_in_sets Q f by (auto simp: emeasure_density)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1019
  qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1020
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1021
40871
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
  1022
section "Radon-Nikodym derivative"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1023
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
  1024
definition
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1025
  "RN_deriv M N \<equiv> SOME f. f \<in> borel_measurable M \<and> (\<forall>x. 0 \<le> f x) \<and> density M f = N"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1026
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1027
lemma
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1028
  assumes f: "f \<in> borel_measurable M" "AE x in M. 0 \<le> f x"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1029
  shows borel_measurable_RN_deriv_density: "RN_deriv M (density M f) \<in> borel_measurable M" (is ?borel)
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1030
    and density_RN_deriv_density: "density M (RN_deriv M (density M f)) = density M f" (is ?density)
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1031
    and RN_deriv_density_nonneg: "0 \<le> RN_deriv M (density M f) x" (is ?pos)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1032
proof -
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1033
  let ?f = "\<lambda>x. max 0 (f x)"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1034
  let ?P = "\<lambda>g. g \<in> borel_measurable M \<and> (\<forall>x. 0 \<le> g x) \<and> density M g = density M f"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1035
  from f have "?P ?f" using f by (auto intro!: density_cong simp: split: split_max)
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1036
  then have "?P (RN_deriv M (density M f))"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1037
    unfolding RN_deriv_def by (rule someI[where P="?P"])
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1038
  then show ?borel ?density ?pos by auto
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1039
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1040
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1041
lemma (in sigma_finite_measure) RN_deriv:
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1042
  assumes "absolutely_continuous M N" "sets N = sets M"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1043
  shows borel_measurable_RN_deriv: "RN_deriv M N \<in> borel_measurable M" (is ?borel)
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1044
    and density_RN_deriv: "density M (RN_deriv M N) = N" (is ?density)
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1045
    and RN_deriv_nonneg: "0 \<le> RN_deriv M N x" (is ?pos)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1046
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1047
  note Ex = Radon_Nikodym[OF assms, unfolded Bex_def]
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1048
  from Ex show ?borel unfolding RN_deriv_def by (rule someI2_ex) simp
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1049
  from Ex show ?density unfolding RN_deriv_def by (rule someI2_ex) simp
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1050
  from Ex show ?pos unfolding RN_deriv_def by (rule someI2_ex) simp
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1051
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1052
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1053
lemma (in sigma_finite_measure) RN_deriv_positive_integral:
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1054
  assumes N: "absolutely_continuous M N" "sets N = sets M"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1055
    and f: "f \<in> borel_measurable M"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1056
  shows "integral\<^isup>P N f = (\<integral>\<^isup>+x. RN_deriv M N x * f x \<partial>M)"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1057
proof -
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1058
  have "integral\<^isup>P N f = integral\<^isup>P (density M (RN_deriv M N)) f"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1059
    using N by (simp add: density_RN_deriv)
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1060
  also have "\<dots> = (\<integral>\<^isup>+x. RN_deriv M N x * f x \<partial>M)"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1061
    using RN_deriv(1,3)[OF N] f by (simp add: positive_integral_density)
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1062
  finally show ?thesis by simp
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1063
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1064
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1065
lemma null_setsD_AE: "N \<in> null_sets M \<Longrightarrow> AE x in M. x \<notin> N"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1066
  using AE_iff_null_sets[of N M] by auto
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1067
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1068
lemma (in sigma_finite_measure) RN_deriv_unique:
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1069
  assumes f: "f \<in> borel_measurable M" "AE x in M. 0 \<le> f x"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1070
  and eq: "density M f = N"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1071
  shows "AE x in M. f x = RN_deriv M N x"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1072
proof (rule density_unique)
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1073
  have ac: "absolutely_continuous M N"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1074
    using f(1) unfolding eq[symmetric] by (rule absolutely_continuousI_density)
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1075
  have eq2: "sets N = sets M"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1076
    unfolding eq[symmetric] by simp
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1077
  show "RN_deriv M N \<in> borel_measurable M" "AE x in M. 0 \<le> RN_deriv M N x"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1078
    "density M f = density M (RN_deriv M N)"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1079
    using RN_deriv[OF ac eq2] eq by auto
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1080
qed fact+
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1081
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1082
lemma (in sigma_finite_measure) RN_deriv_distr:
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1083
  fixes T :: "'a \<Rightarrow> 'b"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1084
  assumes T: "T \<in> measurable M M'" and T': "T' \<in> measurable M' M"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1085
    and inv: "\<forall>x\<in>space M. T' (T x) = x"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1086
  and ac: "absolutely_continuous (distr M M' T) (distr N M' T)"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1087
  and N: "sets N = sets M"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1088
  shows "AE x in M. RN_deriv (distr M M' T) (distr N M' T) (T x) = RN_deriv M N x"
41832
27cb9113b1a0 add lemma RN_deriv_vimage
hoelzl
parents: 41705
diff changeset
  1089
proof (rule RN_deriv_unique)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1090
  have [simp]: "sets N = sets M" by fact
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1091
  note sets_eq_imp_space_eq[OF N, simp]
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1092
  have measurable_N[simp]: "\<And>M'. measurable N M' = measurable M M'" by (auto simp: measurable_def)
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1093
  { fix A assume "A \<in> sets M"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1094
    with inv T T' sets_into_space[OF this]
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1095
    have "T -` T' -` A \<inter> T -` space M' \<inter> space M = A"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1096
      by (auto simp: measurable_def) }
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1097
  note eq = this[simp]
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1098
  { fix A assume "A \<in> sets M"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1099
    with inv T T' sets_into_space[OF this]
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1100
    have "(T' \<circ> T) -` A \<inter> space M = A"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1101
      by (auto simp: measurable_def) }
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1102
  note eq2 = this[simp]
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1103
  let ?M' = "distr M M' T" and ?N' = "distr N M' T"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1104
  interpret M': sigma_finite_measure ?M'
41832
27cb9113b1a0 add lemma RN_deriv_vimage
hoelzl
parents: 41705
diff changeset
  1105
  proof
27cb9113b1a0 add lemma RN_deriv_vimage
hoelzl
parents: 41705
diff changeset
  1106
    from sigma_finite guess F .. note F = this
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1107
    show "\<exists>A::nat \<Rightarrow> 'b set. range A \<subseteq> sets ?M' \<and> (\<Union>i. A i) = space ?M' \<and> (\<forall>i. emeasure ?M' (A i) \<noteq> \<infinity>)"
41832
27cb9113b1a0 add lemma RN_deriv_vimage
hoelzl
parents: 41705
diff changeset
  1108
    proof (intro exI conjI allI)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1109
      show *: "range (\<lambda>i. T' -` F i \<inter> space ?M') \<subseteq> sets ?M'"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1110
        using F T' by (auto simp: measurable_def)
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1111
      show "(\<Union>i. T' -` F i \<inter> space ?M') = space ?M'"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1112
        using F T' by (force simp: measurable_def)
41832
27cb9113b1a0 add lemma RN_deriv_vimage
hoelzl
parents: 41705
diff changeset
  1113
      fix i
27cb9113b1a0 add lemma RN_deriv_vimage
hoelzl
parents: 41705
diff changeset
  1114
      have "T' -` F i \<inter> space M' \<in> sets M'" using * by auto
27cb9113b1a0 add lemma RN_deriv_vimage
hoelzl
parents: 41705
diff changeset
  1115
      moreover
27cb9113b1a0 add lemma RN_deriv_vimage
hoelzl
parents: 41705
diff changeset
  1116
      have Fi: "F i \<in> sets M" using F by auto
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1117
      ultimately show "emeasure ?M' (T' -` F i \<inter> space ?M') \<noteq> \<infinity>"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1118
        using F T T' by (simp add: emeasure_distr)
41832
27cb9113b1a0 add lemma RN_deriv_vimage
hoelzl
parents: 41705
diff changeset
  1119
    qed
27cb9113b1a0 add lemma RN_deriv_vimage
hoelzl
parents: 41705
diff changeset
  1120
  qed
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1121
  have "(RN_deriv ?M' ?N') \<circ> T \<in> borel_measurable M"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1122
    using T ac by (intro measurable_comp[where b="?M'"] M'.borel_measurable_RN_deriv) simp_all
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1123
  then show "(\<lambda>x. RN_deriv ?M' ?N' (T x)) \<in> borel_measurable M"
41832
27cb9113b1a0 add lemma RN_deriv_vimage
hoelzl
parents: 41705
diff changeset
  1124
    by (simp add: comp_def)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1125
  show "AE x in M. 0 \<le> RN_deriv ?M' ?N' (T x)" using M'.RN_deriv_nonneg[OF ac] by auto
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1126
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1127
  have "N = distr N M (T' \<circ> T)"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1128
    by (subst measure_of_of_measure[of N, symmetric])
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1129
       (auto simp add: distr_def sigma_sets_eq intro!: measure_of_eq space_closed)
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1130
  also have "\<dots> = distr (distr N M' T) M T'"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1131
    using T T' by (simp add: distr_distr)
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1132
  also have "\<dots> = distr (density (distr M M' T) (RN_deriv (distr M M' T) (distr N M' T))) M T'"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1133
    using ac by (simp add: M'.density_RN_deriv)
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1134
  also have "\<dots> = density M (RN_deriv (distr M M' T) (distr N M' T) \<circ> T)"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1135
    using M'.borel_measurable_RN_deriv[OF ac] by (simp add: distr_density_distr[OF T T', OF inv])
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1136
  finally show "density M (\<lambda>x. RN_deriv (distr M M' T) (distr N M' T) (T x)) = N"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1137
    by (simp add: comp_def)
41832
27cb9113b1a0 add lemma RN_deriv_vimage
hoelzl
parents: 41705
diff changeset
  1138
qed
27cb9113b1a0 add lemma RN_deriv_vimage
hoelzl
parents: 41705
diff changeset
  1139
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1140
lemma (in sigma_finite_measure) RN_deriv_finite:
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1141
  assumes N: "sigma_finite_measure N" and ac: "absolutely_continuous M N" "sets N = sets M"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1142
  shows "AE x in M. RN_deriv M N x \<noteq> \<infinity>"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1143
proof -
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1144
  interpret N: sigma_finite_measure N by fact
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1145
  from N show ?thesis
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1146
    using sigma_finite_iff_density_finite[OF RN_deriv(1)[OF ac]] RN_deriv(2,3)[OF ac] by simp
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1147
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1148
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1149
lemma (in sigma_finite_measure)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1150
  assumes N: "sigma_finite_measure N" and ac: "absolutely_continuous M N" "sets N = sets M"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1151
    and f: "f \<in> borel_measurable M"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1152
  shows RN_deriv_integrable: "integrable N f \<longleftrightarrow>
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1153
      integrable M (\<lambda>x. real (RN_deriv M N x) * f x)" (is ?integrable)
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1154
    and RN_deriv_integral: "integral\<^isup>L N f =
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1155
      (\<integral>x. real (RN_deriv M N x) * f x \<partial>M)" (is ?integral)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1156
proof -
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1157
  note ac(2)[simp] and sets_eq_imp_space_eq[OF ac(2), simp]
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1158
  interpret N: sigma_finite_measure N by fact
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43556
diff changeset
  1159
  have minus_cong: "\<And>A B A' B'::ereal. 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
  1160
  have f': "(\<lambda>x. - f x) \<in> borel_measurable M" using f by auto
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1161
  have Nf: "f \<in> borel_measurable N" using f by (simp add: measurable_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
  1162
  { fix f :: "'a \<Rightarrow> real"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1163
    { fix x assume *: "RN_deriv M N x \<noteq> \<infinity>"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1164
      have "ereal (real (RN_deriv M N x)) * ereal (f x) = ereal (real (RN_deriv M N x) * f x)"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1165
        by (simp add: mult_le_0_iff)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1166
      then have "RN_deriv M N x * ereal (f x) = ereal (real (RN_deriv M N x) * f x)"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1167
        using RN_deriv(3)[OF ac] * by (auto simp add: ereal_real split: split_if_asm) }
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1168
    then have "(\<integral>\<^isup>+x. ereal (real (RN_deriv M N x) * f x) \<partial>M) = (\<integral>\<^isup>+x. RN_deriv M N x * ereal (f x) \<partial>M)"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1169
              "(\<integral>\<^isup>+x. ereal (- (real (RN_deriv M N x) * f x)) \<partial>M) = (\<integral>\<^isup>+x. RN_deriv M N x * ereal (- f x) \<partial>M)"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1170
      using RN_deriv_finite[OF N ac] unfolding ereal_mult_minus_right uminus_ereal.simps(1)[symmetric]
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
  1171
      by (auto intro!: positive_integral_cong_AE) }
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
  1172
  note * = this
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1173
  show ?integral ?integrable
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
  1174
    unfolding lebesgue_integral_def integrable_def *
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1175
    using Nf f RN_deriv(1)[OF ac]
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1176
    by (auto simp: RN_deriv_positive_integral[OF ac])
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1177
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1178
43340
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42866
diff changeset
  1179
lemma (in sigma_finite_measure) real_RN_deriv:
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1180
  assumes "finite_measure N"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1181
  assumes ac: "absolutely_continuous M N" "sets N = sets M"
43340
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42866
diff changeset
  1182
  obtains D where "D \<in> borel_measurable M"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1183
    and "AE x in M. RN_deriv M N x = ereal (D x)"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1184
    and "AE x in N. 0 < D x"
43340
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42866
diff changeset
  1185
    and "\<And>x. 0 \<le> D x"
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42866
diff changeset
  1186
proof
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1187
  interpret N: finite_measure N by fact
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1188
  
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1189
  note RN = RN_deriv[OF ac]
43340
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42866
diff changeset
  1190
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1191
  let ?RN = "\<lambda>t. {x \<in> space M. RN_deriv M N x = t}"
43340
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42866
diff changeset
  1192
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1193
  show "(\<lambda>x. real (RN_deriv M N x)) \<in> borel_measurable M"
43340
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42866
diff changeset
  1194
    using RN by auto
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42866
diff changeset
  1195
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1196
  have "N (?RN \<infinity>) = (\<integral>\<^isup>+ x. RN_deriv M N x * indicator (?RN \<infinity>) x \<partial>M)"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1197
    using RN(1,3) by (subst RN(2)[symmetric]) (auto simp: emeasure_density)
43340
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42866
diff changeset
  1198
  also have "\<dots> = (\<integral>\<^isup>+ x. \<infinity> * indicator (?RN \<infinity>) x \<partial>M)"
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42866
diff changeset
  1199
    by (intro positive_integral_cong) (auto simp: indicator_def)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1200
  also have "\<dots> = \<infinity> * emeasure M (?RN \<infinity>)"
43340
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42866
diff changeset
  1201
    using RN by (intro positive_integral_cmult_indicator) auto
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1202
  finally have eq: "N (?RN \<infinity>) = \<infinity> * emeasure M (?RN \<infinity>)" .
43340
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42866
diff changeset
  1203
  moreover
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1204
  have "emeasure M (?RN \<infinity>) = 0"
43340
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42866
diff changeset
  1205
  proof (rule ccontr)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1206
    assume "emeasure M {x \<in> space M. RN_deriv M N x = \<infinity>} \<noteq> 0"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1207
    moreover from RN have "0 \<le> emeasure M {x \<in> space M. RN_deriv M N x = \<infinity>}" by auto
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1208
    ultimately have "0 < emeasure M {x \<in> space M. RN_deriv M N x = \<infinity>}" by auto
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1209
    with eq have "N (?RN \<infinity>) = \<infinity>" by simp
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1210
    with N.emeasure_finite[of "?RN \<infinity>"] RN show False by auto
43340
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42866
diff changeset
  1211
  qed
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1212
  ultimately have "AE x in M. RN_deriv M N x < \<infinity>"
43340
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42866
diff changeset
  1213
    using RN by (intro AE_iff_measurable[THEN iffD2]) auto
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1214
  then show "AE x in M. RN_deriv M N x = ereal (real (RN_deriv M N x))"
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43556
diff changeset
  1215
    using RN(3) by (auto simp: ereal_real)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1216
  then have eq: "AE x in N. RN_deriv M N x = ereal (real (RN_deriv M N x))"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1217
    using ac absolutely_continuous_AE by auto
43340
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42866
diff changeset
  1218
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1219
  show "\<And>x. 0 \<le> real (RN_deriv M N x)"
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43556
diff changeset
  1220
    using RN by (auto intro: real_of_ereal_pos)
43340
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42866
diff changeset
  1221
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1222
  have "N (?RN 0) = (\<integral>\<^isup>+ x. RN_deriv M N x * indicator (?RN 0) x \<partial>M)"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1223
    using RN(1,3) by (subst RN(2)[symmetric]) (auto simp: emeasure_density)
43340
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42866
diff changeset
  1224
  also have "\<dots> = (\<integral>\<^isup>+ x. 0 \<partial>M)"
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42866
diff changeset
  1225
    by (intro positive_integral_cong) (auto simp: indicator_def)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1226
  finally have "AE x in N. RN_deriv M N x \<noteq> 0"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1227
    using RN by (subst AE_iff_measurable[OF _ refl]) (auto simp: ac cong: sets_eq_imp_space_eq)
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1228
  with RN(3) eq show "AE x in N. 0 < real (RN_deriv M N x)"
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43556
diff changeset
  1229
    by (auto simp: zero_less_real_of_ereal le_less)
43340
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42866
diff changeset
  1230
qed
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42866
diff changeset
  1231
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1232
lemma (in sigma_finite_measure) RN_deriv_singleton:
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1233
  assumes ac: "absolutely_continuous M N" "sets N = sets M"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1234
  and x: "{x} \<in> sets M"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1235
  shows "N {x} = RN_deriv M N x * emeasure M {x}"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1236
proof -
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1237
  note deriv = RN_deriv[OF ac]
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1238
  from deriv(1,3) `{x} \<in> sets M`
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1239
  have "density M (RN_deriv M N) {x} = (\<integral>\<^isup>+w. RN_deriv M N x * indicator {x} w \<partial>M)"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1240
    by (auto simp: indicator_def emeasure_density intro!: positive_integral_cong)
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1241
  with x deriv show ?thesis
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
  1242
    by (auto simp: positive_integral_cmult_indicator)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1243
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1244
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1245
end