src/HOL/Probability/Radon_Nikodym.thy
author hoelzl
Tue, 19 Jul 2011 14:38:29 +0200
changeset 43923 ab93d0190a5d
parent 43920 cedb5cb948fd
child 44568 e6f291cb5810
permissions -rw-r--r--
add ereal to typeclass infinity
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
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    11
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
    12
  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
    13
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    14
  obtain A :: "nat \<Rightarrow> 'a set" where
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    15
    range: "range A \<subseteq> sets M" and
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    16
    space: "(\<Union>i. A i) = space M" and
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
    17
    measure: "\<And>i. \<mu> (A i) \<noteq> \<infinity>" and
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    18
    disjoint: "disjoint_family A"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    19
    using disjoint_sigma_finite by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    20
  let "?B i" = "2^Suc i * \<mu> (A i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    21
  have "\<forall>i. \<exists>x. 0 < x \<and> x < inverse (?B i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    22
  proof
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
    23
    fix i have Ai: "A i \<in> sets M" using range by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
    24
    from measure positive_measure[OF this]
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
    25
    show "\<exists>x. 0 < x \<and> x < inverse (?B i)"
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43556
diff changeset
    26
      by (auto intro!: ereal_dense simp: ereal_0_gt_inverse)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    27
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    28
  from choice[OF this] obtain n where n: "\<And>i. 0 < n i"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    29
    "\<And>i. n i < inverse (2^Suc i * \<mu> (A i))" by auto
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
    30
  { fix i have "0 \<le> n i" using n(1)[of i] by auto } note pos = this
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
    31
  let "?h x" = "\<Sum>i. n i * indicator (A i) x"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    32
  show ?thesis
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    33
  proof (safe intro!: bexI[of _ ?h] del: notI)
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38656
diff changeset
    34
    have "\<And>i. A i \<in> sets M"
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38656
diff changeset
    35
      using range by fastsimp+
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
    36
    then have "integral\<^isup>P M ?h = (\<Sum>i. n i * \<mu> (A i))" using pos
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
    37
      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
    38
    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
    39
    proof (rule suminf_le_pos)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
    40
      fix N
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
    41
      have "n N * \<mu> (A N) \<le> inverse (2^Suc N * \<mu> (A N)) * \<mu> (A N)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
    42
        using positive_measure[OF `A N \<in> sets M`] n[of N]
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43556
diff changeset
    43
        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
    44
      also have "\<dots> \<le> (1 / 2) ^ Suc N"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    45
        using measure[of N] n[of N]
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43556
diff changeset
    46
        by (cases rule: ereal2_cases[of "n N" "\<mu> (A N)"])
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43556
diff changeset
    47
           (simp_all add: inverse_eq_divide power_divide one_ereal_def ereal_power_divide)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
    48
      finally show "n N * \<mu> (A N) \<le> (1 / 2) ^ Suc N" .
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
    49
      show "0 \<le> n N * \<mu> (A N)" using n[of N] `A N \<in> sets M` by simp
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    50
    qed
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43556
diff changeset
    51
    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
    52
  next
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
    53
    { 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
    54
      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
    55
      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
    56
        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
    57
      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
    58
    note pos = this
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
    59
    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
    60
    proof cases
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
    61
      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
    62
    next
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
    63
      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
    64
      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
    65
    qed
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    66
  next
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
    67
    show "?h \<in> borel_measurable M" using range n
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43556
diff changeset
    68
      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
    69
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    70
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    71
40871
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
    72
subsection "Absolutely continuous"
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
    73
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    74
definition (in measure_space)
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43556
diff changeset
    75
  "absolutely_continuous \<nu> = (\<forall>N\<in>null_sets. \<nu> N = (0 :: ereal))"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    76
41832
27cb9113b1a0 add lemma RN_deriv_vimage
hoelzl
parents: 41705
diff changeset
    77
lemma (in measure_space) absolutely_continuous_AE:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
    78
  assumes "measure_space M'" and [simp]: "sets M' = sets M" "space M' = space M"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
    79
    and "absolutely_continuous (measure M')" "AE x. P x"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
    80
   shows "AE x in M'. P x"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
    81
proof -
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
    82
  interpret \<nu>: measure_space M' by fact
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
    83
  from `AE x. P x` obtain N where N: "N \<in> null_sets" and "{x\<in>space M. \<not> P x} \<subseteq> N"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
    84
    unfolding almost_everywhere_def by auto
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
    85
  show "AE x in M'. P x"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
    86
  proof (rule \<nu>.AE_I')
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
    87
    show "{x\<in>space M'. \<not> P x} \<subseteq> N" by simp fact
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
    88
    from `absolutely_continuous (measure M')` show "N \<in> \<nu>.null_sets"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
    89
      using N unfolding absolutely_continuous_def by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
    90
  qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
    91
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
    92
39097
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39092
diff changeset
    93
lemma (in finite_measure_space) absolutely_continuousI:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
    94
  assumes "finite_measure_space (M\<lparr> measure := \<nu>\<rparr>)" (is "finite_measure_space ?\<nu>")
39097
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39092
diff changeset
    95
  assumes v: "\<And>x. \<lbrakk> x \<in> space M ; \<mu> {x} = 0 \<rbrakk> \<Longrightarrow> \<nu> {x} = 0"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39092
diff changeset
    96
  shows "absolutely_continuous \<nu>"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39092
diff changeset
    97
proof (unfold absolutely_continuous_def sets_eq_Pow, safe)
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39092
diff changeset
    98
  fix N assume "\<mu> N = 0" "N \<subseteq> space M"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
    99
  interpret v: finite_measure_space ?\<nu> by fact
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   100
  have "\<nu> N = measure ?\<nu> (\<Union>x\<in>N. {x})" by simp
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   101
  also have "\<dots> = (\<Sum>x\<in>N. measure ?\<nu> {x})"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   102
  proof (rule v.measure_setsum[symmetric])
39097
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39092
diff changeset
   103
    show "finite N" using `N \<subseteq> space M` finite_space by (auto intro: finite_subset)
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39092
diff changeset
   104
    show "disjoint_family_on (\<lambda>i. {i}) N" unfolding disjoint_family_on_def by auto
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   105
    fix x assume "x \<in> N" thus "{x} \<in> sets ?\<nu>" using `N \<subseteq> space M` sets_eq_Pow by auto
39097
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39092
diff changeset
   106
  qed
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39092
diff changeset
   107
  also have "\<dots> = 0"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39092
diff changeset
   108
  proof (safe intro!: setsum_0')
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39092
diff changeset
   109
    fix x assume "x \<in> N"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   110
    hence "\<mu> {x} \<le> \<mu> N" "0 \<le> \<mu> {x}"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   111
      using sets_eq_Pow `N \<subseteq> space M` positive_measure[of "{x}"]
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   112
      by (auto intro!: measure_mono)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   113
    then have "\<mu> {x} = 0" using `\<mu> N = 0` by simp
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   114
    thus "measure ?\<nu> {x} = 0" using v[of x] `x \<in> N` `N \<subseteq> space M` by auto
39097
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39092
diff changeset
   115
  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
   116
  finally show "\<nu> N = 0" by simp
39097
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39092
diff changeset
   117
qed
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39092
diff changeset
   118
40871
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
   119
lemma (in measure_space) density_is_absolutely_continuous:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   120
  assumes "\<And>A. A \<in> sets M \<Longrightarrow> \<nu> A = (\<integral>\<^isup>+x. f x * indicator A x \<partial>M)"
40871
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
   121
  shows "absolutely_continuous \<nu>"
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
   122
  using assms unfolding absolutely_continuous_def
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
   123
  by (simp add: positive_integral_null_set)
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
   124
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
   125
subsection "Existence of the Radon-Nikodym derivative"
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
   126
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   127
lemma (in finite_measure) Radon_Nikodym_aux_epsilon:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   128
  fixes e :: real assumes "0 < e"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   129
  assumes "finite_measure (M\<lparr>measure := \<nu>\<rparr>)"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   130
  shows "\<exists>A\<in>sets M. \<mu>' (space M) - finite_measure.\<mu>' (M\<lparr>measure := \<nu>\<rparr>) (space M) \<le>
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   131
                    \<mu>' A - finite_measure.\<mu>' (M\<lparr>measure := \<nu>\<rparr>) A \<and>
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   132
                    (\<forall>B\<in>sets M. B \<subseteq> A \<longrightarrow> - e < \<mu>' B - finite_measure.\<mu>' (M\<lparr>measure := \<nu>\<rparr>) B)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   133
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
   134
  interpret M': finite_measure "M\<lparr>measure := \<nu>\<rparr>" by fact
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   135
  let "?d A" = "\<mu>' A - M'.\<mu>' A"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   136
  let "?A A" = "if (\<forall>B\<in>sets M. B \<subseteq> space M - A \<longrightarrow> -e < ?d B)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   137
    then {}
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   138
    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
   139
  def A \<equiv> "\<lambda>n. ((\<lambda>B. B \<union> ?A B) ^^ n) {}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   140
  have A_simps[simp]:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   141
    "A 0 = {}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   142
    "\<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
   143
  { fix A assume "A \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   144
    have "?A A \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   145
      by (auto intro!: someI2[of _ _ "\<lambda>A. A \<in> sets M"] simp: not_less) }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   146
  note A'_in_sets = this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   147
  { fix n have "A n \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   148
    proof (induct n)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   149
      case (Suc n) thus "A (Suc n) \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   150
        using A'_in_sets[of "A n"] by (auto split: split_if_asm)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   151
    qed (simp add: A_def) }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   152
  note A_in_sets = this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   153
  hence "range A \<subseteq> sets M" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   154
  { fix n B
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   155
    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
   156
    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
   157
    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
   158
    proof (rule someI2_ex[OF Ex])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   159
      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
   160
      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
   161
      hence "?d (A n \<union> B) = ?d (A n) + ?d B"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   162
        using `A n \<in> sets M` finite_measure_Union M'.finite_measure_Union by simp
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   163
      also have "\<dots> \<le> ?d (A n) - e" using dB by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   164
      finally show "?d (A n \<union> B) \<le> ?d (A n) - e" .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   165
    qed }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   166
  note dA_epsilon = this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   167
  { fix n have "?d (A (Suc n)) \<le> ?d (A n)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   168
    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
   169
      case True from dA_epsilon[OF this] show ?thesis using `0 < e` by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   170
    next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   171
      case False
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   172
      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
   173
      thus ?thesis by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   174
    qed }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   175
  note dA_mono = this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   176
  show ?thesis
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   177
  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
   178
    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
   179
    show ?thesis
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   180
    proof (safe intro!: bexI[of _ "space M - A n"])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   181
      fix B assume "B \<in> sets M" "B \<subseteq> space M - A n"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   182
      from B[OF this] show "-e < ?d B" .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   183
    next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   184
      show "space M - A n \<in> sets M" by (rule compl_sets) fact
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   185
    next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   186
      show "?d (space M) \<le> ?d (space M - A n)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   187
      proof (induct n)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   188
        fix n assume "?d (space M) \<le> ?d (space M - A n)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   189
        also have "\<dots> \<le> ?d (space M - A (Suc n))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   190
          using A_in_sets sets_into_space dA_mono[of n]
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   191
          by (simp del: A_simps add: finite_measure_Diff M'.finite_measure_Diff)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   192
        finally show "?d (space M) \<le> ?d (space M - A (Suc n))" .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   193
      qed simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   194
    qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   195
  next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   196
    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
   197
      by (auto simp add: not_less)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   198
    { fix n have "?d (A n) \<le> - real n * e"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   199
      proof (induct n)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   200
        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
   201
      next
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43556
diff changeset
   202
        case 0 with M'.empty_measure 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
   203
      qed } note dA_less = this
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   204
    have decseq: "decseq (\<lambda>n. ?d (A n))" unfolding decseq_eq_incseq
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   205
    proof (rule incseq_SucI)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   206
      fix n show "- ?d (A n) \<le> - ?d (A (Suc n))" using dA_mono[of n] by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   207
    qed
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   208
    have A: "incseq A" by (auto intro!: incseq_SucI)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   209
    from finite_continuity_from_below[OF _ A] `range A \<subseteq> sets M`
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   210
      M'.finite_continuity_from_below[OF _ A]
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   211
    have convergent: "(\<lambda>i. ?d (A i)) ----> ?d (\<Union>i. A i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   212
      by (auto intro!: LIMSEQ_diff)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   213
    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
   214
    moreover from order_trans[OF decseq_le[OF decseq convergent] dA_less]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   215
    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
   216
    ultimately show ?thesis by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   217
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   218
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   219
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   220
lemma (in finite_measure) restricted_measure_subset:
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   221
  assumes S: "S \<in> sets M" and X: "X \<subseteq> S"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   222
  shows "finite_measure.\<mu>' (restricted_space S) X = \<mu>' X"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   223
proof cases
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   224
  note r = restricted_finite_measure[OF S]
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   225
  { assume "X \<in> sets M" with S X show ?thesis
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   226
      unfolding finite_measure.\<mu>'_def[OF r] \<mu>'_def by auto }
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   227
  { assume "X \<notin> sets M"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   228
    moreover then have "S \<inter> X \<notin> sets M"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   229
      using X by (simp add: Int_absorb1)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   230
    ultimately show ?thesis
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   231
      unfolding finite_measure.\<mu>'_def[OF r] \<mu>'_def using S by auto }
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   232
qed
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   233
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   234
lemma (in finite_measure) restricted_measure:
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   235
  assumes X: "S \<in> sets M" "X \<in> sets (restricted_space S)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   236
  shows "finite_measure.\<mu>' (restricted_space S) X = \<mu>' X"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   237
proof -
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   238
  from X have "S \<in> sets M" "X \<subseteq> S" by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   239
  from restricted_measure_subset[OF this] show ?thesis .
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   240
qed
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   241
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   242
lemma (in finite_measure) Radon_Nikodym_aux:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   243
  assumes "finite_measure (M\<lparr>measure := \<nu>\<rparr>)" (is "finite_measure ?M'")
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   244
  shows "\<exists>A\<in>sets M. \<mu>' (space M) - finite_measure.\<mu>' (M\<lparr>measure := \<nu>\<rparr>) (space M) \<le>
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   245
                    \<mu>' A - finite_measure.\<mu>' (M\<lparr>measure := \<nu>\<rparr>) A \<and>
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   246
                    (\<forall>B\<in>sets M. B \<subseteq> A \<longrightarrow> 0 \<le> \<mu>' B - finite_measure.\<mu>' (M\<lparr>measure := \<nu>\<rparr>) B)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   247
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
   248
  interpret M': finite_measure ?M' where
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   249
    "space ?M' = space M" and "sets ?M' = sets M" and "measure ?M' = \<nu>" by fact auto
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   250
  let "?d A" = "\<mu>' A - M'.\<mu>' A"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   251
  let "?P A B n" = "A \<in> sets M \<and> A \<subseteq> B \<and> ?d B \<le> ?d A \<and> (\<forall>C\<in>sets M. C \<subseteq> A \<longrightarrow> - 1 / real (Suc n) < ?d C)"
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38656
diff changeset
   252
  let "?r S" = "restricted_space S"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   253
  { fix S n assume S: "S \<in> sets M"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   254
    note r = M'.restricted_finite_measure[of S] restricted_finite_measure[OF S] S
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   255
    then have "finite_measure (?r S)" "0 < 1 / real (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
   256
      "finite_measure (?r S\<lparr>measure := \<nu>\<rparr>)" by auto
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   257
    from finite_measure.Radon_Nikodym_aux_epsilon[OF this] guess X .. note X = this
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   258
    have "?P X S n"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   259
    proof (intro conjI ballI impI)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   260
      show "X \<in> sets M" "X \<subseteq> S" using X(1) `S \<in> sets M` by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   261
      have "S \<in> op \<inter> S ` sets M" using `S \<in> sets M` by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   262
      then show "?d S \<le> ?d X"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   263
        using S X restricted_measure[OF S] M'.restricted_measure[OF S] by simp
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   264
      fix C assume "C \<in> sets M" "C \<subseteq> X"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   265
      then have "C \<in> sets (restricted_space S)" "C \<subseteq> X"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   266
        using `S \<in> sets M` `X \<subseteq> S` by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   267
      with X(2) show "- 1 / real (Suc n) < ?d C"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   268
        using S X restricted_measure[OF S] M'.restricted_measure[OF S] by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   269
    qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   270
    hence "\<exists>A. ?P A S n" by auto }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   271
  note Ex_P = this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   272
  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
   273
  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
   274
  have A_0[simp]: "A 0 = space M" unfolding A_def by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   275
  { fix i have "A i \<in> sets M" unfolding A_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   276
    proof (induct i)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   277
      case (Suc i)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   278
      from Ex_P[OF this, of i] show ?case unfolding nat_rec_Suc
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   279
        by (rule someI2_ex) simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   280
    qed simp }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   281
  note A_in_sets = this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   282
  { fix n have "?P (A (Suc n)) (A n) n"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   283
      using Ex_P[OF A_in_sets] unfolding A_Suc
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   284
      by (rule someI2_ex) simp }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   285
  note P_A = this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   286
  have "range A \<subseteq> sets M" using A_in_sets by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   287
  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
   288
  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
   289
  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
   290
      using P_A by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   291
  show ?thesis
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   292
  proof (safe intro!: bexI[of _ "\<Inter>i. A i"])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   293
    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
   294
    have A: "decseq A" using A_mono by (auto intro!: decseq_SucI)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   295
    from
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   296
      finite_continuity_from_above[OF `range A \<subseteq> sets M` A]
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   297
      M'.finite_continuity_from_above[OF `range A \<subseteq> sets M` A]
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   298
    have "(\<lambda>i. ?d (A i)) ----> ?d (\<Inter>i. A i)" by (intro LIMSEQ_diff)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   299
    thus "?d (space M) \<le> ?d (\<Inter>i. A i)" using mono_dA[THEN monoD, of 0 _]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   300
      by (rule_tac LIMSEQ_le_const) (auto intro!: exI)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   301
  next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   302
    fix B assume B: "B \<in> sets M" "B \<subseteq> (\<Inter>i. A i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   303
    show "0 \<le> ?d B"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   304
    proof (rule ccontr)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   305
      assume "\<not> 0 \<le> ?d B"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   306
      hence "0 < - ?d B" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   307
      from ex_inverse_of_nat_Suc_less[OF this]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   308
      obtain n where *: "?d B < - 1 / real (Suc n)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   309
        by (auto simp: real_eq_of_nat inverse_eq_divide field_simps)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   310
      have "B \<subseteq> A (Suc n)" using B by (auto simp del: nat_rec_Suc)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   311
      from epsilon[OF B(1) this] *
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   312
      show False by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   313
    qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   314
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   315
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   316
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   317
lemma (in finite_measure) Radon_Nikodym_finite_measure:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   318
  assumes "finite_measure (M\<lparr> measure := \<nu>\<rparr>)" (is "finite_measure ?M'")
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   319
  assumes "absolutely_continuous \<nu>"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   320
  shows "\<exists>f \<in> borel_measurable M. \<forall>A\<in>sets M. \<nu> A = (\<integral>\<^isup>+x. f x * indicator A x \<partial>M)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   321
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
   322
  interpret M': finite_measure ?M'
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   323
    where "space ?M' = space M" and "sets ?M' = sets M" and "measure ?M' = \<nu>"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   324
    using assms(1) by auto
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   325
  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> \<nu> A)}"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   326
  have "(\<lambda>x. 0) \<in> G" unfolding G_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   327
  hence "G \<noteq> {}" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   328
  { fix f g assume f: "f \<in> G" and g: "g \<in> G"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   329
    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
   330
    proof safe
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   331
      show "?max \<in> borel_measurable M" using f g unfolding G_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   332
      let ?A = "{x \<in> space M. f x \<le> g x}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   333
      have "?A \<in> sets M" using f g unfolding G_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   334
      fix A assume "A \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   335
      hence sets: "?A \<inter> A \<in> sets M" "(space M - ?A) \<inter> A \<in> sets M" using `?A \<in> sets M` by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   336
      have union: "((?A \<inter> A) \<union> ((space M - ?A) \<inter> A)) = A"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   337
        using sets_into_space[OF `A \<in> sets M`] by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   338
      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
   339
        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
   340
        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
   341
      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
   342
        (\<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
   343
        (\<integral>\<^isup>+x. f x * indicator ((space M - ?A) \<inter> A) x \<partial>M)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   344
        using f g sets unfolding G_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   345
        by (auto cong: positive_integral_cong intro!: positive_integral_add borel_measurable_indicator)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   346
      also have "\<dots> \<le> \<nu> (?A \<inter> A) + \<nu> ((space M - ?A) \<inter> A)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   347
        using f g sets unfolding G_def by (auto intro!: add_mono)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   348
      also have "\<dots> = \<nu> A"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   349
        using M'.measure_additive[OF sets] union by auto
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   350
      finally show "(\<integral>\<^isup>+x. max (g x) (f x) * indicator A x \<partial>M) \<le> \<nu> A" .
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   351
    next
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   352
      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
   353
    qed }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   354
  note max_in_G = this
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   355
  { 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
   356
    have "(\<lambda>x. SUP i. f i x) \<in> G" unfolding G_def
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   357
    proof safe
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   358
      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
   359
        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
   360
      { fix x show "0 \<le> (SUP i. f i x)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   361
          using f by (auto simp: G_def intro: le_SUPI2) }
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   362
    next
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   363
      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
   364
      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
   365
        (\<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
   366
        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
   367
      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
   368
        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
   369
        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
   370
           (auto simp: G_def incseq_Suc_iff le_fun_def split: split_indicator)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   371
      finally show "(\<integral>\<^isup>+x. (SUP i. f i x) * indicator A x \<partial>M) \<le> \<nu> A"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   372
        using f `A \<in> sets M` by (auto intro!: SUP_leI simp: G_def)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   373
    qed }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   374
  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
   375
  let ?y = "SUP g : G. integral\<^isup>P M g"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   376
  have "?y \<le> \<nu> (space M)" unfolding G_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   377
  proof (safe intro!: SUP_leI)
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   378
    fix g assume "\<forall>A\<in>sets M. (\<integral>\<^isup>+x. g x * indicator A x \<partial>M) \<le> \<nu> A"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   379
    from this[THEN bspec, OF top] show "integral\<^isup>P M g \<le> \<nu> (space M)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   380
      by (simp cong: positive_integral_cong)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   381
  qed
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   382
  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
   383
  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
   384
  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
   385
    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
   386
    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
   387
    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
   388
  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
   389
  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
   390
  hence y_eq: "?y = (SUP i. integral\<^isup>P M (gs i))" using ys by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   391
  let "?g 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
   392
  def f \<equiv> "\<lambda>x. SUP i. ?g i x"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   393
  let "?F A x" = "f x * indicator A x"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   394
  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
   395
  { fix i have "?g i \<in> G"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   396
    proof (induct i)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   397
      case 0 thus ?case by simp fact
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   398
    next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   399
      case (Suc i)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   400
      with Suc gs_not_empty `gs (Suc i) \<in> G` show ?case
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   401
        by (auto simp add: atMost_Suc intro!: max_in_G)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   402
    qed }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   403
  note g_in_G = this
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   404
  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
   405
    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
   406
  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
   407
  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
   408
  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
   409
    using g_in_G `incseq ?g`
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   410
    by (auto intro!: positive_integral_monotone_convergence_SUP simp: G_def)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   411
  also have "\<dots> = ?y"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   412
  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
   413
    show "(SUP i. integral\<^isup>P M (?g i)) \<le> ?y"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   414
      using g_in_G by (auto intro!: exI Sup_mono simp: SUPR_def)
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   415
    show "?y \<le> (SUP i. integral\<^isup>P M (?g i))" unfolding y_eq
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   416
      by (auto intro!: SUP_mono positive_integral_mono Max_ge)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   417
  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
   418
  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
   419
  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
   420
    unfolding f_def using `\<And>i. gs i \<in> G`
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   421
    by (auto intro!: le_SUPI2 Max_ge_iff[THEN iffD2] simp: G_def)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   422
  let "?t A" = "\<nu> A - (\<integral>\<^isup>+x. ?F A x \<partial>M)"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   423
  let ?M = "M\<lparr> measure := ?t\<rparr>"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   424
  interpret M: sigma_algebra ?M
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   425
    by (intro sigma_algebra_cong) auto
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   426
  have f_le_\<nu>: "\<And>A. A \<in> sets M \<Longrightarrow> (\<integral>\<^isup>+x. ?F A x \<partial>M) \<le> \<nu> A"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   427
    using `f \<in> G` unfolding G_def by auto
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   428
  have fmM: "finite_measure ?M"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   429
  proof (default, simp_all add: countably_additive_def positive_def, safe del: notI)
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
   430
    fix A :: "nat \<Rightarrow> 'a set"  assume A: "range A \<subseteq> sets M" "disjoint_family A"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   431
    have "(\<Sum>n. (\<integral>\<^isup>+x. ?F (A n) x \<partial>M)) = (\<integral>\<^isup>+x. (\<Sum>n. ?F (A n) x) \<partial>M)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   432
      using `range A \<subseteq> sets M` `\<And>x. 0 \<le> f x`
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   433
      by (intro positive_integral_suminf[symmetric]) auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   434
    also have "\<dots> = (\<integral>\<^isup>+x. ?F (\<Union>n. A n) x \<partial>M)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   435
      using `\<And>x. 0 \<le> f x`
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43556
diff changeset
   436
      by (intro positive_integral_cong) (simp add: suminf_cmult_ereal suminf_indicator[OF `disjoint_family A`])
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   437
    finally have "(\<Sum>n. (\<integral>\<^isup>+x. ?F (A n) x \<partial>M)) = (\<integral>\<^isup>+x. ?F (\<Union>n. A n) x \<partial>M)" .
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   438
    moreover have "(\<Sum>n. \<nu> (A n)) = \<nu> (\<Union>n. A 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
   439
      using M'.measure_countably_additive A by (simp add: comp_def)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   440
    moreover have v_fin: "\<nu> (\<Union>i. A i) \<noteq> \<infinity>" using M'.finite_measure A by (simp add: countable_UN)
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
   441
    moreover {
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   442
      have "(\<integral>\<^isup>+x. ?F (\<Union>i. A i) x \<partial>M) \<le> \<nu> (\<Union>i. A i)"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   443
        using A `f \<in> G` unfolding G_def by (auto simp: countable_UN)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   444
      also have "\<nu> (\<Union>i. A i) < \<infinity>" using v_fin by simp
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   445
      finally have "(\<integral>\<^isup>+x. ?F (\<Union>i. A i) x \<partial>M) \<noteq> \<infinity>" by simp }
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   446
    moreover have "\<And>i. (\<integral>\<^isup>+x. ?F (A i) x \<partial>M) \<le> \<nu> (A i)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   447
      using A by (intro f_le_\<nu>) 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
   448
    ultimately
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   449
    show "(\<Sum>n. ?t (A n)) = ?t (\<Union>i. A i)"
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43556
diff changeset
   450
      by (subst suminf_ereal_minus) (simp_all add: positive_integral_positive)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   451
  next
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   452
    fix A assume A: "A \<in> sets M" show "0 \<le> \<nu> A - \<integral>\<^isup>+ x. ?F A x \<partial>M"
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43556
diff changeset
   453
      using f_le_\<nu>[OF A] `f \<in> G` M'.finite_measure[OF A] by (auto simp: G_def ereal_le_minus_iff)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   454
  next
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   455
    show "\<nu> (space M) - (\<integral>\<^isup>+ x. ?F (space M) x \<partial>M) \<noteq> \<infinity>" (is "?a - ?b \<noteq> _")
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   456
      using positive_integral_positive[of "?F (space M)"]
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43556
diff changeset
   457
      by (cases rule: ereal2_cases[of ?a ?b]) auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   458
  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
   459
  then interpret M: finite_measure ?M
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   460
    where "space ?M = space M" and "sets ?M = sets M" and "measure ?M = ?t"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   461
    by (simp_all add: fmM)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   462
  have ac: "absolutely_continuous ?t" unfolding absolutely_continuous_def
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   463
  proof
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   464
    fix N assume N: "N \<in> null_sets"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   465
    with `absolutely_continuous \<nu>` have "\<nu> N = 0" unfolding absolutely_continuous_def by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   466
    moreover with N have "(\<integral>\<^isup>+ x. ?F N x \<partial>M) \<le> \<nu> N" using `f \<in> G` by (auto simp: G_def)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   467
    ultimately show "\<nu> N - (\<integral>\<^isup>+ x. ?F N x \<partial>M) = 0"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   468
      using positive_integral_positive by (auto intro!: antisym)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   469
  qed
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   470
  have upper_bound: "\<forall>A\<in>sets M. ?t A \<le> 0"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   471
  proof (rule ccontr)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   472
    assume "\<not> ?thesis"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   473
    then obtain A where A: "A \<in> sets M" and pos: "0 < ?t A"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   474
      by (auto simp: not_le)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   475
    note pos
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   476
    also have "?t A \<le> ?t (space M)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   477
      using M.measure_mono[of A "space M"] A sets_into_space by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   478
    finally have pos_t: "0 < ?t (space M)" by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   479
    moreover
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   480
    then have "\<mu> (space M) \<noteq> 0"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   481
      using ac unfolding absolutely_continuous_def by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   482
    then have pos_M: "0 < \<mu> (space M)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   483
      using positive_measure[OF top] by (simp add: le_less)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   484
    moreover
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   485
    have "(\<integral>\<^isup>+x. f x * indicator (space M) x \<partial>M) \<le> \<nu> (space M)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   486
      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
   487
    hence "(\<integral>\<^isup>+x. f x * indicator (space M) x \<partial>M) \<noteq> \<infinity>"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   488
      using M'.finite_measure_of_space by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   489
    moreover
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   490
    def b \<equiv> "?t (space M) / \<mu> (space M) / 2"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   491
    ultimately have b: "b \<noteq> 0 \<and> 0 \<le> b \<and> b \<noteq> \<infinity>"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   492
      using M'.finite_measure_of_space positive_integral_positive[of "?F (space M)"]
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43556
diff changeset
   493
      by (cases rule: ereal3_cases[of "integral\<^isup>P M (?F (space M))" "\<nu> (space M)" "\<mu> (space M)"])
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   494
         (simp_all add: field_simps)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   495
    then have b: "b \<noteq> 0" "0 \<le> b" "0 < b"  "b \<noteq> \<infinity>" 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
   496
    let ?Mb = "?M\<lparr>measure := \<lambda>A. b * \<mu> A\<rparr>"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   497
    interpret b: sigma_algebra ?Mb by (intro sigma_algebra_cong) auto
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   498
    have Mb: "finite_measure ?Mb"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   499
    proof
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   500
      show "positive ?Mb (measure ?Mb)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   501
        using `0 \<le> b` by (auto simp: positive_def)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   502
      show "countably_additive ?Mb (measure ?Mb)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   503
        using `0 \<le> b` measure_countably_additive
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43556
diff changeset
   504
        by (auto simp: countably_additive_def suminf_cmult_ereal subset_eq)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   505
      show "measure ?Mb (space ?Mb) \<noteq> \<infinity>"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   506
        using b by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   507
    qed
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   508
    from M.Radon_Nikodym_aux[OF this]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   509
    obtain A0 where "A0 \<in> sets M" and
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   510
      space_less_A0: "real (?t (space M)) - real (b * \<mu> (space M)) \<le> real (?t A0) - real (b * \<mu> A0)" and
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   511
      *: "\<And>B. \<lbrakk> B \<in> sets M ; B \<subseteq> A0 \<rbrakk> \<Longrightarrow> 0 \<le> real (?t B) - real (b * \<mu> B)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   512
      unfolding M.\<mu>'_def finite_measure.\<mu>'_def[OF Mb] by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   513
    { fix B assume B: "B \<in> sets M" "B \<subseteq> A0"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   514
      with *[OF this] have "b * \<mu> B \<le> ?t B"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   515
        using M'.finite_measure b finite_measure M.positive_measure[OF B(1)]
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43556
diff changeset
   516
        by (cases rule: ereal2_cases[of "?t B" "b * \<mu> B"]) auto }
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   517
    note bM_le_t = this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   518
    let "?f0 x" = "f x + b * indicator A0 x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   519
    { fix A assume A: "A \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   520
      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
   521
      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
   522
        (\<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
   523
        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
   524
      hence "(\<integral>\<^isup>+x. ?f0 x * indicator A x \<partial>M) =
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   525
          (\<integral>\<^isup>+x. f x * indicator A x \<partial>M) + b * \<mu> (A \<inter> A0)"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   526
        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
   527
        by (simp add: G_def positive_integral_add positive_integral_cmult_indicator) }
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   528
    note f0_eq = this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   529
    { fix A assume A: "A \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   530
      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
   531
      have f_le_v: "(\<integral>\<^isup>+x. f x * indicator A x \<partial>M) \<le> \<nu> A"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   532
        using `f \<in> G` A unfolding G_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   533
      note f0_eq[OF A]
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   534
      also have "(\<integral>\<^isup>+x. f x * indicator A x \<partial>M) + b * \<mu> (A \<inter> A0) \<le>
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   535
          (\<integral>\<^isup>+x. f x * indicator A x \<partial>M) + ?t (A \<inter> A0)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   536
        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
   537
        by (auto intro!: add_left_mono)
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   538
      also have "\<dots> \<le> (\<integral>\<^isup>+x. f x * indicator A x \<partial>M) + ?t A"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   539
        using M.measure_mono[simplified, OF _ `A \<inter> A0 \<in> sets M` `A \<in> sets M`]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   540
        by (auto intro!: add_left_mono)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   541
      also have "\<dots> \<le> \<nu> A"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   542
        using f_le_v M'.finite_measure[simplified, OF `A \<in> sets M`] positive_integral_positive[of "?F A"]
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   543
        by (cases "\<integral>\<^isup>+x. ?F A x \<partial>M", cases "\<nu> A") 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
   544
      finally have "(\<integral>\<^isup>+x. ?f0 x * indicator A x \<partial>M) \<le> \<nu> A" . }
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   545
    hence "?f0 \<in> G" using `A0 \<in> sets M` b `f \<in> G` unfolding G_def
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43556
diff changeset
   546
      by (auto intro!: borel_measurable_indicator borel_measurable_ereal_add
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43556
diff changeset
   547
                       borel_measurable_ereal_times ereal_add_nonneg_nonneg)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   548
    have real: "?t (space M) \<noteq> \<infinity>" "?t A0 \<noteq> \<infinity>"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   549
      "b * \<mu> (space M) \<noteq> \<infinity>" "b * \<mu> A0 \<noteq> \<infinity>"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   550
      using `A0 \<in> sets M` b
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   551
        finite_measure[of A0] M.finite_measure[of A0]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   552
        finite_measure_of_space M.finite_measure_of_space
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   553
      by auto
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   554
    have int_f_finite: "integral\<^isup>P M f \<noteq> \<infinity>"
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43556
diff changeset
   555
      using M'.finite_measure_of_space pos_t unfolding ereal_less_minus_iff
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   556
      by (auto cong: positive_integral_cong)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   557
    have  "0 < ?t (space M) - b * \<mu> (space M)" unfolding b_def
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   558
      using finite_measure_of_space M'.finite_measure_of_space pos_t pos_M
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   559
      using positive_integral_positive[of "?F (space M)"]
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43556
diff changeset
   560
      by (cases rule: ereal3_cases[of "\<mu> (space M)" "\<nu> (space M)" "integral\<^isup>P M (?F (space M))"])
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   561
         (auto simp: field_simps mult_less_cancel_left)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   562
    also have "\<dots> \<le> ?t A0 - b * \<mu> A0"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   563
      using space_less_A0 b
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   564
      using
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   565
        `A0 \<in> sets M`[THEN M.real_measure]
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   566
        top[THEN M.real_measure]
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   567
      apply safe
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   568
      apply simp
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   569
      using
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   570
        `A0 \<in> sets M`[THEN real_measure]
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   571
        `A0 \<in> sets M`[THEN M'.real_measure]
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   572
        top[THEN real_measure]
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   573
        top[THEN M'.real_measure]
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   574
      by (cases b) auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   575
    finally have 1: "b * \<mu> A0 < ?t A0"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   576
      using
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   577
        `A0 \<in> sets M`[THEN M.real_measure]
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   578
      apply safe
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   579
      apply simp
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   580
      using
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   581
        `A0 \<in> sets M`[THEN real_measure]
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   582
        `A0 \<in> sets M`[THEN M'.real_measure]
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   583
      by (cases b) auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   584
    have "0 < ?t A0"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   585
      using b `A0 \<in> sets M` by (auto intro!: le_less_trans[OF _ 1])
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   586
    then have "\<mu> A0 \<noteq> 0" using ac unfolding absolutely_continuous_def
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   587
      using `A0 \<in> sets M` by auto
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   588
    then have "0 < \<mu> A0" using positive_measure[OF `A0 \<in> sets M`] by auto
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43556
diff changeset
   589
    hence "0 < b * \<mu> A0" using b by (auto simp: ereal_zero_less_0_iff)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   590
    with int_f_finite have "?y + 0 < integral\<^isup>P M f + b * \<mu> A0" unfolding int_f_eq_y
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   591
      using `f \<in> G`
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43556
diff changeset
   592
      by (intro ereal_add_strict_mono) (auto intro!: le_SUPI2 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
   593
    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
   594
      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
   595
    finally have "?y < integral\<^isup>P M ?f0" by simp
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   596
    moreover from `?f0 \<in> G` have "integral\<^isup>P M ?f0 \<le> ?y" by (auto intro!: le_SUPI)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   597
    ultimately show False by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   598
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   599
  show ?thesis
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   600
  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
   601
    fix A assume A: "A\<in>sets M"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   602
    show "\<nu> A = (\<integral>\<^isup>+x. f x * indicator A x \<partial>M)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   603
    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
   604
      show "(\<integral>\<^isup>+x. f x * indicator A x \<partial>M) \<le> \<nu> A"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   605
        using `f \<in> G` `A \<in> sets M` unfolding G_def by auto
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   606
      show "\<nu> A \<le> (\<integral>\<^isup>+x. f x * indicator A x \<partial>M)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   607
        using upper_bound[THEN bspec, OF `A \<in> sets M`]
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   608
        using M'.real_measure[OF A]
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   609
        by (cases "integral\<^isup>P M (?F A)") auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   610
    qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   611
  qed simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   612
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   613
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   614
lemma (in finite_measure) split_space_into_finite_sets_and_rest:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   615
  assumes "measure_space (M\<lparr>measure := \<nu>\<rparr>)" (is "measure_space ?N")
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   616
  assumes ac: "absolutely_continuous \<nu>"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   617
  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>
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   618
    (\<forall>A\<in>sets M. A \<subseteq> A0 \<longrightarrow> (\<mu> A = 0 \<and> \<nu> A = 0) \<or> (\<mu> A > 0 \<and> \<nu> A = \<infinity>)) \<and>
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   619
    (\<forall>i. \<nu> (B i) \<noteq> \<infinity>)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   620
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
   621
  interpret v: measure_space ?N
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   622
    where "space ?N = space M" and "sets ?N = sets M" and "measure ?N = \<nu>"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   623
    by fact auto
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   624
  let ?Q = "{Q\<in>sets M. \<nu> Q \<noteq> \<infinity>}"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   625
  let ?a = "SUP Q:?Q. \<mu> Q"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   626
  have "{} \<in> ?Q" using v.empty_measure by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   627
  then have Q_not_empty: "?Q \<noteq> {}" by blast
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   628
  have "?a \<le> \<mu> (space M)" using sets_into_space
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   629
    by (auto intro!: SUP_leI measure_mono top)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   630
  then have "?a \<noteq> \<infinity>" using finite_measure_of_space
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   631
    by auto
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   632
  from SUPR_countable_SUPR[OF Q_not_empty, of \<mu>]
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   633
  obtain Q'' where "range Q'' \<subseteq> \<mu> ` ?Q" and a: "?a = (SUP i::nat. Q'' i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   634
    by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   635
  then have "\<forall>i. \<exists>Q'. Q'' i = \<mu> Q' \<and> Q' \<in> ?Q" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   636
  from choice[OF this] obtain Q' where Q': "\<And>i. Q'' i = \<mu> (Q' i)" "\<And>i. Q' i \<in> ?Q"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   637
    by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   638
  then have a_Lim: "?a = (SUP i::nat. \<mu> (Q' i))" using a by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   639
  let "?O n" = "\<Union>i\<le>n. Q' i"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   640
  have Union: "(SUP i. \<mu> (?O i)) = \<mu> (\<Union>i. ?O i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   641
  proof (rule continuity_from_below[of ?O])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   642
    show "range ?O \<subseteq> sets M" using Q' by (auto intro!: finite_UN)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   643
    show "incseq ?O" by (fastsimp intro!: incseq_SucI)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   644
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   645
  have Q'_sets: "\<And>i. Q' i \<in> sets M" using Q' by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   646
  have O_sets: "\<And>i. ?O i \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   647
     using Q' by (auto intro!: finite_UN Un)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   648
  then have O_in_G: "\<And>i. ?O i \<in> ?Q"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   649
  proof (safe del: notI)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   650
    fix i have "Q' ` {..i} \<subseteq> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   651
      using Q' by (auto intro: finite_UN)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   652
    with v.measure_finitely_subadditive[of "{.. i}" Q']
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   653
    have "\<nu> (?O i) \<le> (\<Sum>i\<le>i. \<nu> (Q' i))" by auto
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   654
    also have "\<dots> < \<infinity>" using Q' by (simp add: setsum_Pinfty)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   655
    finally show "\<nu> (?O i) \<noteq> \<infinity>" by simp
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   656
  qed auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   657
  have O_mono: "\<And>n. ?O n \<subseteq> ?O (Suc n)" by fastsimp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   658
  have a_eq: "?a = \<mu> (\<Union>i. ?O i)" unfolding Union[symmetric]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   659
  proof (rule antisym)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   660
    show "?a \<le> (SUP i. \<mu> (?O i))" unfolding a_Lim
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   661
      using Q' by (auto intro!: SUP_mono measure_mono finite_UN)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   662
    show "(SUP i. \<mu> (?O i)) \<le> ?a" unfolding SUPR_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   663
    proof (safe intro!: Sup_mono, unfold bex_simps)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   664
      fix i
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   665
      have *: "(\<Union>Q' ` {..i}) = ?O i" by auto
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   666
      then show "\<exists>x. (x \<in> sets M \<and> \<nu> x \<noteq> \<infinity>) \<and>
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   667
        \<mu> (\<Union>Q' ` {..i}) \<le> \<mu> x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   668
        using O_in_G[of i] by (auto intro!: exI[of _ "?O i"])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   669
    qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   670
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   671
  let "?O_0" = "(\<Union>i. ?O i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   672
  have "?O_0 \<in> sets M" using Q' by auto
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   673
  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
   674
  { 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
   675
  note Q_sets = this
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   676
  show ?thesis
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   677
  proof (intro bexI exI conjI ballI impI allI)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   678
    show "disjoint_family Q"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   679
      by (fastsimp simp: disjoint_family_on_def Q_def
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   680
        split: nat.split_asm)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   681
    show "range Q \<subseteq> sets M"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   682
      using Q_sets by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   683
    { fix A assume A: "A \<in> sets M" "A \<subseteq> space M - ?O_0"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   684
      show "\<mu> A = 0 \<and> \<nu> A = 0 \<or> 0 < \<mu> A \<and> \<nu> A = \<infinity>"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   685
      proof (rule disjCI, simp)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   686
        assume *: "0 < \<mu> A \<longrightarrow> \<nu> A \<noteq> \<infinity>"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   687
        show "\<mu> A = 0 \<and> \<nu> A = 0"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   688
        proof cases
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   689
          assume "\<mu> A = 0" moreover with ac A have "\<nu> A = 0"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   690
            unfolding absolutely_continuous_def by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   691
          ultimately show ?thesis by simp
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   692
        next
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   693
          assume "\<mu> A \<noteq> 0" with * have "\<nu> A \<noteq> \<infinity>" using positive_measure[OF A(1)] by auto
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   694
          with A have "\<mu> ?O_0 + \<mu> A = \<mu> (?O_0 \<union> A)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   695
            using Q' by (auto intro!: measure_additive countable_UN)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   696
          also have "\<dots> = (SUP i. \<mu> (?O i \<union> A))"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   697
          proof (rule continuity_from_below[of "\<lambda>i. ?O i \<union> A", symmetric, simplified])
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   698
            show "range (\<lambda>i. ?O i \<union> A) \<subseteq> sets M"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   699
              using `\<nu> A \<noteq> \<infinity>` O_sets A by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   700
          qed (fastsimp intro!: incseq_SucI)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   701
          also have "\<dots> \<le> ?a"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   702
          proof (safe intro!: SUP_leI)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   703
            fix i have "?O i \<union> A \<in> ?Q"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   704
            proof (safe del: notI)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   705
              show "?O i \<union> A \<in> sets M" using O_sets A by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   706
              from O_in_G[of i]
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   707
              moreover have "\<nu> (?O i \<union> A) \<le> \<nu> (?O i) + \<nu> A"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   708
                using v.measure_subadditive[of "?O i" A] A O_sets by auto
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   709
              ultimately show "\<nu> (?O i \<union> A) \<noteq> \<infinity>"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   710
                using `\<nu> A \<noteq> \<infinity>` by auto
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   711
            qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   712
            then show "\<mu> (?O i \<union> A) \<le> ?a" by (rule le_SUPI)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   713
          qed
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   714
          finally have "\<mu> A = 0"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   715
            unfolding a_eq using real_measure[OF `?O_0 \<in> sets M`] real_measure[OF A(1)] by auto
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   716
          with `\<mu> A \<noteq> 0` show ?thesis by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   717
        qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   718
      qed }
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   719
    { fix i show "\<nu> (Q i) \<noteq> \<infinity>"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   720
      proof (cases i)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   721
        case 0 then show ?thesis
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   722
          unfolding Q_def using Q'[of 0] by simp
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   723
      next
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   724
        case (Suc n)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   725
        then show ?thesis unfolding Q_def
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   726
          using `?O n \<in> ?Q` `?O (Suc n) \<in> ?Q`
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   727
          using v.measure_mono[OF O_mono, of n] v.positive_measure[of "?O n"] v.positive_measure[of "?O (Suc n)"]
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   728
          using v.measure_Diff[of "?O n" "?O (Suc n)", OF _ _ _ O_mono]
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43556
diff changeset
   729
          by (cases rule: ereal2_cases[of "\<nu> (\<Union> x\<le>Suc n. Q' x)" "\<nu> (\<Union> i\<le>n. Q' i)"]) auto
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   730
      qed }
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   731
    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
   732
    { 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
   733
      proof (induct j)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   734
        case 0 then show ?case by (simp add: Q_def)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   735
      next
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   736
        case (Suc j)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   737
        have eq: "\<And>j. (\<Union>i\<le>j. ?O i) = (\<Union>i\<le>j. Q' i)" by fastsimp
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   738
        have "{..j} \<union> {..Suc j} = {..Suc j}" by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   739
        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
   740
          by (simp add: UN_Un[symmetric] Q_def del: UN_Un)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   741
        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
   742
      qed }
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   743
    then have "(\<Union>j. (\<Union>i\<le>j. ?O i)) = (\<Union>j. (\<Union>i\<le>j. Q i))" by simp
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   744
    then show "space M - ?O_0 = space M - (\<Union>i. Q i)" by fastsimp
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   745
  qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   746
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   747
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   748
lemma (in finite_measure) Radon_Nikodym_finite_measure_infinite:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   749
  assumes "measure_space (M\<lparr>measure := \<nu>\<rparr>)" (is "measure_space ?N")
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   750
  assumes "absolutely_continuous \<nu>"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   751
  shows "\<exists>f \<in> borel_measurable M. (\<forall>x. 0 \<le> f x) \<and> (\<forall>A\<in>sets M. \<nu> A = (\<integral>\<^isup>+x. f x * indicator A x \<partial>M))"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   752
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
   753
  interpret v: measure_space ?N
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   754
    where "space ?N = space M" and "sets ?N = sets M" and "measure ?N = \<nu>"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   755
    by fact auto
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   756
  from split_space_into_finite_sets_and_rest[OF assms]
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   757
  obtain Q0 and Q :: "nat \<Rightarrow> 'a set"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   758
    where Q: "disjoint_family Q" "range Q \<subseteq> sets M"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   759
    and Q0: "Q0 \<in> sets M" "Q0 = space M - (\<Union>i. Q i)"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   760
    and in_Q0: "\<And>A. A \<in> sets M \<Longrightarrow> A \<subseteq> Q0 \<Longrightarrow> \<mu> A = 0 \<and> \<nu> A = 0 \<or> 0 < \<mu> A \<and> \<nu> A = \<infinity>"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   761
    and Q_fin: "\<And>i. \<nu> (Q i) \<noteq> \<infinity>" by force
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   762
  from Q have Q_sets: "\<And>i. Q i \<in> sets M" by auto
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   763
  have "\<forall>i. \<exists>f. f\<in>borel_measurable M \<and> (\<forall>x. 0 \<le> f x) \<and> (\<forall>A\<in>sets M.
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   764
    \<nu> (Q i \<inter> A) = (\<integral>\<^isup>+x. f x * indicator (Q i \<inter> A) x \<partial>M))"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   765
  proof
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   766
    fix i
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43556
diff changeset
   767
    have indicator_eq: "\<And>f x A. (f x :: ereal) * indicator (Q i \<inter> A) x * indicator (Q i) x
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   768
      = (f x * indicator (Q i) x) * indicator A x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   769
      unfolding indicator_def by auto
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   770
    have fm: "finite_measure (restricted_space (Q i))"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   771
      (is "finite_measure ?R") by (rule restricted_finite_measure[OF Q_sets[of i]])
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   772
    then interpret R: finite_measure ?R .
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   773
    have fmv: "finite_measure (restricted_space (Q i) \<lparr> measure := \<nu>\<rparr>)" (is "finite_measure ?Q")
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   774
      unfolding finite_measure_def finite_measure_axioms_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   775
    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
   776
      show "measure_space ?Q"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   777
        using v.restricted_measure_space Q_sets[of i] by auto
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   778
      show "measure ?Q (space ?Q) \<noteq> \<infinity>" using Q_fin by simp
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   779
    qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   780
    have "R.absolutely_continuous \<nu>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   781
      using `absolutely_continuous \<nu>` `Q i \<in> sets M`
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   782
      by (auto simp: R.absolutely_continuous_def absolutely_continuous_def)
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   783
    from R.Radon_Nikodym_finite_measure[OF fmv this]
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   784
    obtain f where f: "(\<lambda>x. f x * indicator (Q i) x) \<in> borel_measurable M"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   785
      and f_int: "\<And>A. A\<in>sets M \<Longrightarrow> \<nu> (Q i \<inter> A) = (\<integral>\<^isup>+x. (f x * indicator (Q i) x) * indicator A x \<partial>M)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   786
      unfolding Bex_def borel_measurable_restricted[OF `Q i \<in> sets M`]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   787
        positive_integral_restricted[OF `Q i \<in> sets M`] by (auto simp: indicator_eq)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   788
    then show "\<exists>f. f\<in>borel_measurable M \<and> (\<forall>x. 0 \<le> f x) \<and> (\<forall>A\<in>sets M.
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   789
      \<nu> (Q i \<inter> A) = (\<integral>\<^isup>+x. f x * indicator (Q i \<inter> A) x \<partial>M))"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   790
      by (auto intro!: exI[of _ "\<lambda>x. max 0 (f x * indicator (Q i) x)"] positive_integral_cong_pos
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   791
        split: split_indicator split_if_asm simp: max_def)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   792
  qed
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   793
  from choice[OF this] obtain f where borel: "\<And>i. f i \<in> borel_measurable M" "\<And>i x. 0 \<le> f i x"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   794
    and f: "\<And>A i. A \<in> sets M \<Longrightarrow>
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   795
      \<nu> (Q i \<inter> A) = (\<integral>\<^isup>+x. f i x * indicator (Q i \<inter> A) x \<partial>M)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   796
    by auto
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   797
  let "?f x" = "(\<Sum>i. f i x * indicator (Q i) x) + \<infinity> * indicator Q0 x"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   798
  show ?thesis
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   799
  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
   800
    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
   801
      by (auto intro!: measurable_If)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   802
    show "\<And>x. 0 \<le> ?f x" using borel by (auto intro!: suminf_0_le simp: indicator_def)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   803
    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
   804
    have Qi: "\<And>i. Q i \<in> sets M" using Q by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   805
    have [intro,simp]: "\<And>i. (\<lambda>x. f i x * indicator (Q i \<inter> A) x) \<in> borel_measurable M"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   806
      "\<And>i. AE x. 0 \<le> f i x * indicator (Q i \<inter> A) x"
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43556
diff changeset
   807
      using borel Qi Q0(1) `A \<in> sets M` by (auto intro!: borel_measurable_ereal_times)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   808
    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)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   809
      using borel by (intro positive_integral_cong) (auto simp: indicator_def)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   810
    also have "\<dots> = (\<integral>\<^isup>+x. (\<Sum>i. f i x * indicator (Q i \<inter> A) x) \<partial>M) + \<infinity> * \<mu> (Q0 \<inter> A)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   811
      using borel Qi Q0(1) `A \<in> sets M`
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43556
diff changeset
   812
      by (subst positive_integral_add) (auto simp del: ereal_infty_mult
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   813
          simp add: positive_integral_cmult_indicator Int intro!: suminf_0_le)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   814
    also have "\<dots> = (\<Sum>i. \<nu> (Q i \<inter> A)) + \<infinity> * \<mu> (Q0 \<inter> A)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   815
      by (subst f[OF `A \<in> sets M`], subst positive_integral_suminf) auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   816
    finally have "(\<integral>\<^isup>+x. ?f x * indicator A x \<partial>M) = (\<Sum>i. \<nu> (Q i \<inter> A)) + \<infinity> * \<mu> (Q0 \<inter> A)" .
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   817
    moreover have "(\<Sum>i. \<nu> (Q i \<inter> A)) = \<nu> ((\<Union>i. Q i) \<inter> A)"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   818
      using Q Q_sets `A \<in> sets M`
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   819
      by (intro v.measure_countably_additive[of "\<lambda>i. Q i \<inter> A", unfolded comp_def, simplified])
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   820
         (auto simp: disjoint_family_on_def)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   821
    moreover have "\<infinity> * \<mu> (Q0 \<inter> A) = \<nu> (Q0 \<inter> A)"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   822
    proof -
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   823
      have "Q0 \<inter> A \<in> sets M" using Q0(1) `A \<in> sets M` by blast
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   824
      from in_Q0[OF this] show ?thesis by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   825
    qed
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   826
    moreover have "Q0 \<inter> A \<in> sets M" "((\<Union>i. Q i) \<inter> A) \<in> sets M"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   827
      using Q_sets `A \<in> sets M` Q0(1) by (auto intro!: countable_UN)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   828
    moreover have "((\<Union>i. Q i) \<inter> A) \<union> (Q0 \<inter> A) = A" "((\<Union>i. Q i) \<inter> A) \<inter> (Q0 \<inter> A) = {}"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   829
      using `A \<in> sets M` sets_into_space Q0 by auto
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   830
    ultimately show "\<nu> A = (\<integral>\<^isup>+x. ?f x * indicator A x \<partial>M)"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   831
      using v.measure_additive[simplified, of "(\<Union>i. Q i) \<inter> A" "Q0 \<inter> A"]
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   832
      by simp
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   833
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   834
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   835
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   836
lemma (in sigma_finite_measure) Radon_Nikodym:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   837
  assumes "measure_space (M\<lparr>measure := \<nu>\<rparr>)" (is "measure_space ?N")
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   838
  assumes ac: "absolutely_continuous \<nu>"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   839
  shows "\<exists>f \<in> borel_measurable M. (\<forall>x. 0 \<le> f x) \<and> (\<forall>A\<in>sets M. \<nu> A = (\<integral>\<^isup>+x. f x * indicator A x \<partial>M))"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   840
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   841
  from Ex_finite_integrable_function
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   842
  obtain h where finite: "integral\<^isup>P M h \<noteq> \<infinity>" and
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   843
    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
   844
    nn: "\<And>x. 0 \<le> h x" and
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   845
    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
   846
    "\<And>x. x \<in> space M \<Longrightarrow> h x < \<infinity>" 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
   847
  let "?T A" = "(\<integral>\<^isup>+x. h x * indicator A x \<partial>M)"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   848
  let ?MT = "M\<lparr> measure := ?T \<rparr>"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   849
  interpret T: finite_measure ?MT
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   850
    where "space ?MT = space M" and "sets ?MT = sets M" and "measure ?MT = ?T"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   851
    unfolding finite_measure_def finite_measure_axioms_def using borel finite nn
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   852
    by (auto intro!: measure_space_density cong: positive_integral_cong)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   853
  have "T.absolutely_continuous \<nu>"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   854
  proof (unfold T.absolutely_continuous_def, safe)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   855
    fix N assume "N \<in> sets M" "(\<integral>\<^isup>+x. h x * indicator N x \<partial>M) = 0"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   856
    with borel ac pos have "AE x. x \<notin> N"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   857
      by (subst (asm) positive_integral_0_iff_AE) (auto split: split_indicator simp: not_le)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   858
    then have "N \<in> null_sets" using `N \<in> sets M` sets_into_space
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   859
      by (subst (asm) AE_iff_measurable[OF `N \<in> sets M`]) auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   860
    then show "\<nu> N = 0" using ac by (auto simp: absolutely_continuous_def)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   861
  qed
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   862
  from T.Radon_Nikodym_finite_measure_infinite[simplified, OF assms(1) this]
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   863
  obtain f where f_borel: "f \<in> borel_measurable M" "\<And>x. 0 \<le> f x" and
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   864
    fT: "\<And>A. A \<in> sets M \<Longrightarrow> \<nu> A = (\<integral>\<^isup>+ x. f x * indicator A x \<partial>?MT)"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   865
    by (auto simp: measurable_def)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   866
  show ?thesis
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   867
  proof (safe intro!: bexI[of _ "\<lambda>x. h x * f x"])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   868
    show "(\<lambda>x. h x * f x) \<in> borel_measurable M"
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43556
diff changeset
   869
      using borel f_borel by (auto intro: borel_measurable_ereal_times)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   870
    show "\<And>x. 0 \<le> h x * f x" using nn f_borel by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   871
    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
   872
    then show "\<nu> A = (\<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
   873
      unfolding fT[OF `A \<in> sets M`] mult_assoc using nn borel f_borel
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   874
      by (intro positive_integral_translated_density) auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   875
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   876
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   877
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   878
section "Uniqueness of densities"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   879
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   880
lemma (in measure_space) finite_density_unique:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   881
  assumes borel: "f \<in> borel_measurable M" "g \<in> borel_measurable M"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   882
  assumes pos: "AE x. 0 \<le> f x" "AE x. 0 \<le> g x"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   883
  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
   884
  shows "(\<forall>A\<in>sets M. (\<integral>\<^isup>+x. f x * indicator A x \<partial>M) = (\<integral>\<^isup>+x. g x * indicator A x \<partial>M))
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   885
    \<longleftrightarrow> (AE x. f x = g x)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   886
    (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
   887
proof (intro iffI ballI)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   888
  fix A assume eq: "AE x. f x = g x"
41705
1100512e16d8 add auto support for AE_mp
hoelzl
parents: 41689
diff changeset
   889
  then show "?P f A = ?P g A"
1100512e16d8 add auto support for AE_mp
hoelzl
parents: 41689
diff changeset
   890
    by (auto intro: positive_integral_cong_AE)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   891
next
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   892
  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
   893
  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
   894
  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
   895
  { fix f g assume borel: "f \<in> borel_measurable M" "g \<in> borel_measurable M"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   896
      and pos: "AE x. 0 \<le> f x" "AE x. 0 \<le> g x"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   897
      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
   898
    let ?N = "{x\<in>space M. g x < f x}"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   899
    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
   900
    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
   901
      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
   902
    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
   903
    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
   904
      by (auto intro!: positive_integral_cong simp: indicator_def)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   905
    also have "\<dots> = ?P f ?N - ?P g ?N"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   906
    proof (rule positive_integral_diff)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   907
      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
   908
        using borel N by auto
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   909
      show "AE x. g x * indicator ?N x \<le> f x * indicator ?N x"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   910
           "AE x. 0 \<le> g x * indicator ?N x"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   911
        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
   912
    qed fact
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   913
    also have "\<dots> = 0"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   914
      unfolding eq[THEN bspec, OF N] using positive_integral_positive Pg_fin by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   915
    finally have "AE x. f x \<le> g x"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   916
      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
   917
      by (subst (asm) positive_integral_0_iff_AE)
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43556
diff changeset
   918
         (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
   919
  from this[OF borel pos g_fin eq] this[OF borel(2,1) pos(2,1) fin] eq
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   920
  show "AE x. f x = g x" by auto
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   921
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   922
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   923
lemma (in finite_measure) density_unique_finite_measure:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   924
  assumes borel: "f \<in> borel_measurable M" "f' \<in> borel_measurable M"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   925
  assumes pos: "AE x. 0 \<le> f x" "AE x. 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
   926
  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
   927
    (is "\<And>A. A \<in> sets M \<Longrightarrow> ?P f A = ?P f' A")
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   928
  shows "AE x. f x = f' x"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   929
proof -
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   930
  let "?\<nu> A" = "?P f A" and "?\<nu>' A" = "?P f' A"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   931
  let "?f A x" = "f x * indicator A x" and "?f' A x" = "f' x * indicator A x"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   932
  interpret M: measure_space "M\<lparr> measure := ?\<nu>\<rparr>"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   933
    using borel(1) pos(1) by (rule measure_space_density) simp
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   934
  have ac: "absolutely_continuous ?\<nu>"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   935
    using f by (rule density_is_absolutely_continuous)
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   936
  from split_space_into_finite_sets_and_rest[OF `measure_space (M\<lparr> measure := ?\<nu>\<rparr>)` ac]
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   937
  obtain Q0 and Q :: "nat \<Rightarrow> 'a set"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   938
    where Q: "disjoint_family Q" "range Q \<subseteq> sets M"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   939
    and Q0: "Q0 \<in> sets M" "Q0 = space M - (\<Union>i. Q i)"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   940
    and in_Q0: "\<And>A. A \<in> sets M \<Longrightarrow> A \<subseteq> Q0 \<Longrightarrow> \<mu> A = 0 \<and> ?\<nu> A = 0 \<or> 0 < \<mu> A \<and> ?\<nu> A = \<infinity>"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   941
    and Q_fin: "\<And>i. ?\<nu> (Q i) \<noteq> \<infinity>" by force
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   942
  from Q have Q_sets: "\<And>i. Q i \<in> sets M" by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   943
  let ?N = "{x\<in>space M. f x \<noteq> f' x}"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   944
  have "?N \<in> sets M" using borel by auto
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43556
diff changeset
   945
  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
   946
    unfolding indicator_def by auto
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   947
  have "\<forall>i. AE x. ?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
   948
    by (intro finite_density_unique[THEN iffD1] allI)
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43556
diff changeset
   949
       (auto intro!: borel_measurable_ereal_times f Int simp: *)
41705
1100512e16d8 add auto support for AE_mp
hoelzl
parents: 41689
diff changeset
   950
  moreover have "AE x. ?f Q0 x = ?f' Q0 x"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   951
  proof (rule AE_I')
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43556
diff changeset
   952
    { fix f :: "'a \<Rightarrow> ereal" assume borel: "f \<in> borel_measurable M"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   953
        and eq: "\<And>A. A \<in> sets M \<Longrightarrow> ?\<nu> A = (\<integral>\<^isup>+x. f x * indicator A x \<partial>M)"
43923
ab93d0190a5d add ereal to typeclass infinity
hoelzl
parents: 43920
diff changeset
   954
      let "?A i" = "Q0 \<inter> {x \<in> space M. f x < (i::nat)}"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   955
      have "(\<Union>i. ?A i) \<in> null_sets"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   956
      proof (rule null_sets_UN)
43923
ab93d0190a5d add ereal to typeclass infinity
hoelzl
parents: 43920
diff changeset
   957
        fix i ::nat have "?A i \<in> sets M"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   958
          using borel Q0(1) by auto
43923
ab93d0190a5d add ereal to typeclass infinity
hoelzl
parents: 43920
diff changeset
   959
        have "?\<nu> (?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
   960
          unfolding eq[OF `?A i \<in> sets M`]
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   961
          by (auto intro!: positive_integral_mono simp: indicator_def)
43923
ab93d0190a5d add ereal to typeclass infinity
hoelzl
parents: 43920
diff changeset
   962
        also have "\<dots> = i * \<mu> (?A i)"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   963
          using `?A i \<in> sets M` 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
   964
        also have "\<dots> < \<infinity>"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   965
          using `?A i \<in> sets M`[THEN finite_measure] by auto
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   966
        finally have "?\<nu> (?A i) \<noteq> \<infinity>" by simp
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   967
        then show "?A i \<in> null_sets" using in_Q0[OF `?A i \<in> sets M`] `?A i \<in> sets M` by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   968
      qed
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   969
      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
   970
        by (auto simp: less_PInf_Ex_of_nat real_eq_of_nat)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   971
      finally have "Q0 \<inter> {x\<in>space M. f x \<noteq> \<infinity>} \<in> null_sets" by simp }
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   972
    from this[OF borel(1) refl] this[OF borel(2) f]
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   973
    have "Q0 \<inter> {x\<in>space M. f x \<noteq> \<infinity>} \<in> null_sets" "Q0 \<inter> {x\<in>space M. f' x \<noteq> \<infinity>} \<in> null_sets" by simp_all
42866
b0746bd57a41 the measurable sets with null measure form a ring
hoelzl
parents: 42067
diff changeset
   974
    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" by (rule nullsets.Un)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   975
    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
   976
      (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
   977
  qed
41705
1100512e16d8 add auto support for AE_mp
hoelzl
parents: 41689
diff changeset
   978
  moreover have "\<And>x. (?f Q0 x = ?f' Q0 x) \<longrightarrow> (\<forall>i. ?f (Q i) x = ?f' (Q i) x) \<longrightarrow>
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   979
    ?f (space M) x = ?f' (space M) x"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   980
    by (auto simp: indicator_def Q0)
41705
1100512e16d8 add auto support for AE_mp
hoelzl
parents: 41689
diff changeset
   981
  ultimately have "AE x. ?f (space M) x = ?f' (space M) x"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   982
    by (auto simp: AE_all_countable[symmetric])
41705
1100512e16d8 add auto support for AE_mp
hoelzl
parents: 41689
diff changeset
   983
  then show "AE x. f x = f' x" by auto
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   984
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   985
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   986
lemma (in sigma_finite_measure) density_unique:
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   987
  assumes f: "f \<in> borel_measurable M" "AE x. 0 \<le> f x"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   988
  assumes f': "f' \<in> borel_measurable M" "AE x. 0 \<le> f' x"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   989
  assumes eq: "\<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
   990
    (is "\<And>A. A \<in> sets M \<Longrightarrow> ?P f A = ?P f' A")
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   991
  shows "AE x. f x = f' x"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   992
proof -
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
   993
  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
   994
    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
   995
    using Ex_finite_integrable_function by auto
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   996
  then have h_nn: "AE x. 0 \<le> h x" by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   997
  let ?H = "M\<lparr> measure := \<lambda>A. (\<integral>\<^isup>+x. h x * indicator A x \<partial>M) \<rparr>"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   998
  have H: "measure_space ?H"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
   999
    using h_borel h_nn by (rule measure_space_density) simp
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
  1000
  then interpret h: measure_space ?H .
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
  1001
  interpret h: finite_measure "M\<lparr> measure := \<lambda>A. (\<integral>\<^isup>+x. h x * indicator A x \<partial>M) \<rparr>"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1002
    by default (simp cong: positive_integral_cong add: fin)
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1003
  let ?fM = "M\<lparr>measure := \<lambda>A. (\<integral>\<^isup>+x. f x * indicator A x \<partial>M)\<rparr>"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1004
  interpret f: measure_space ?fM
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
  1005
    using f by (rule measure_space_density) simp
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1006
  let ?f'M = "M\<lparr>measure := \<lambda>A. (\<integral>\<^isup>+x. f' x * indicator A x \<partial>M)\<rparr>"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1007
  interpret f': measure_space ?f'M
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
  1008
    using f' by (rule measure_space_density) simp
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1009
  { 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
  1010
    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
  1011
      using pos(1) sets_into_space by (force simp: indicator_def)
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1012
    then have "(\<integral>\<^isup>+x. h x * indicator A x \<partial>M) = 0 \<longleftrightarrow> A \<in> null_sets"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
  1013
      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
  1014
  note h_null_sets = this
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1015
  { 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
  1016
    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
  1017
      using `A \<in> sets M` h_borel h_nn f f'
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
  1018
      by (intro positive_integral_translated_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
  1019
    also have "\<dots> = (\<integral>\<^isup>+x. h x * indicator A x \<partial>?f'M)"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
  1020
      by (rule f'.positive_integral_cong_measure) (simp_all add: eq)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
  1021
    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
  1022
      using `A \<in> sets M` h_borel h_nn f f'
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
  1023
      by (intro positive_integral_translated_density) auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
  1024
    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
  1025
      by (simp add: ac_simps)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
  1026
    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
  1027
      using `A \<in> sets M` h_borel h_nn f f'
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
  1028
      by (subst (asm) (1 2) positive_integral_translated_density[symmetric]) auto }
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
  1029
  then have "AE x in ?H. f x = f' x" using h_borel h_nn f f'
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
  1030
    by (intro h.density_unique_finite_measure absolutely_continuous_AE[OF H] density_is_absolutely_continuous)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
  1031
       simp_all
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1032
  then show "AE x. f x = f' x"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1033
    unfolding h.almost_everywhere_def almost_everywhere_def
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1034
    by (auto simp add: h_null_sets)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1035
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1036
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1037
lemma (in sigma_finite_measure) sigma_finite_iff_density_finite:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1038
  assumes \<nu>: "measure_space (M\<lparr>measure := \<nu>\<rparr>)" (is "measure_space ?N")
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
  1039
    and f: "f \<in> borel_measurable M" "AE x. 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
  1040
    and eq: "\<And>A. A \<in> sets M \<Longrightarrow> \<nu> A = (\<integral>\<^isup>+x. f x * indicator A x \<partial>M)"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
  1041
  shows "sigma_finite_measure (M\<lparr>measure := \<nu>\<rparr>) \<longleftrightarrow> (AE x. f x \<noteq> \<infinity>)"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1042
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
  1043
  assume "sigma_finite_measure ?N"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1044
  then interpret \<nu>: sigma_finite_measure ?N
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1045
    where "borel_measurable ?N = borel_measurable M" and "measure ?N = \<nu>"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1046
    and "sets ?N = sets M" and "space ?N = space M" by (auto simp: measurable_def)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1047
  from \<nu>.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
  1048
    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
  1049
    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
  1050
    fin: "\<forall>x\<in>space M. 0 < h x \<and> h x < \<infinity>" by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
  1051
  have "AE x. f x * h x \<noteq> \<infinity>"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1052
  proof (rule AE_I')
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
  1053
    have "integral\<^isup>P ?N h = (\<integral>\<^isup>+x. f x * h x \<partial>M)" using f h h_nn
41705
1100512e16d8 add auto support for AE_mp
hoelzl
parents: 41689
diff changeset
  1054
      by (subst \<nu>.positive_integral_cong_measure[symmetric,
1100512e16d8 add auto support for AE_mp
hoelzl
parents: 41689
diff changeset
  1055
          of "M\<lparr> measure := \<lambda> A. \<integral>\<^isup>+x. f x * indicator A x \<partial>M \<rparr>"])
1100512e16d8 add auto support for AE_mp
hoelzl
parents: 41689
diff changeset
  1056
         (auto intro!: positive_integral_translated_density simp: eq)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
  1057
    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
  1058
      using h(2) by simp
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
  1059
    then show "(\<lambda>x. f x * h x) -` {\<infinity>} \<inter> space M \<in> null_sets"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
  1060
      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
  1061
  qed auto
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
  1062
  then show "AE x. f x \<noteq> \<infinity>"
41705
1100512e16d8 add auto support for AE_mp
hoelzl
parents: 41689
diff changeset
  1063
    using fin by (auto elim!: AE_Ball_mp)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1064
next
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
  1065
  assume AE: "AE x. f x \<noteq> \<infinity>"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1066
  from sigma_finite guess Q .. note Q = this
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1067
  interpret \<nu>: measure_space ?N
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1068
    where "borel_measurable ?N = borel_measurable M" and "measure ?N = \<nu>"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1069
    and "sets ?N = sets M" and "space ?N = space M" using \<nu> by (auto simp: measurable_def)
43923
ab93d0190a5d add ereal to typeclass infinity
hoelzl
parents: 43920
diff changeset
  1070
  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
  1071
  { 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
  1072
    unfolding A_def using f Q
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1073
    apply (rule_tac Int)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
  1074
    by (cases i) (auto intro: measurable_sets[OF f(1)]) }
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1075
  note A_in_sets = this
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1076
  let "?A n" = "case prod_decode n of (i,j) \<Rightarrow> A i \<inter> Q j"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1077
  show "sigma_finite_measure ?N"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1078
  proof (default, intro exI conjI subsetI allI)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1079
    fix x assume "x \<in> range ?A"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1080
    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
  1081
    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
  1082
  next
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1083
    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
  1084
    proof safe
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1085
      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
  1086
      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
  1087
        by (intro UN_I[of "prod_encode (i,j)"]) auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1088
    qed auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1089
    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
  1090
    also have "(\<Union>i. A i) = space M"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1091
    proof safe
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1092
      fix x assume x: "x \<in> space M"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1093
      show "x \<in> (\<Union>i. A i)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1094
      proof (cases "f x")
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
  1095
        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
  1096
      next
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
  1097
        case (real r)
43923
ab93d0190a5d add ereal to typeclass infinity
hoelzl
parents: 43920
diff changeset
  1098
        with less_PInf_Ex_of_nat[of "f x"] obtain n :: nat where "f x < n" by (auto simp: real_eq_of_nat)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
  1099
        then show ?thesis using x real unfolding A_def by (auto intro!: exI[of _ "Suc n"])
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
  1100
      next
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
  1101
        case MInf with x show ?thesis
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
  1102
          unfolding A_def by (auto intro!: exI[of _ "Suc 0"])
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1103
      qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1104
    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
  1105
    finally show "(\<Union>i. ?A i) = space ?N" by simp
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1106
  next
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1107
    fix n obtain i j where
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1108
      [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
  1109
    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
  1110
    proof (cases i)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1111
      case 0
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1112
      have "AE x. f x * indicator (A i \<inter> Q j) x = 0"
41705
1100512e16d8 add auto support for AE_mp
hoelzl
parents: 41689
diff changeset
  1113
        using AE by (auto simp: A_def `i = 0`)
1100512e16d8 add auto support for AE_mp
hoelzl
parents: 41689
diff changeset
  1114
      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
  1115
    next
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1116
      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
  1117
      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
  1118
        (\<integral>\<^isup>+x. (Suc n :: ereal) * indicator (Q j) x \<partial>M)"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1119
        by (auto intro!: positive_integral_mono simp: indicator_def A_def)
43923
ab93d0190a5d add ereal to typeclass infinity
hoelzl
parents: 43920
diff changeset
  1120
      also have "\<dots> = Suc n * \<mu> (Q j)"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1121
        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
  1122
      also have "\<dots> < \<infinity>"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
  1123
        using Q by (auto simp: real_eq_of_nat[symmetric])
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1124
      finally show ?thesis by simp
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1125
    qed
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
  1126
    then show "measure ?N (?A n) \<noteq> \<infinity>"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1127
      using A_in_sets Q eq by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1128
  qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1129
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1130
40871
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
  1131
section "Radon-Nikodym derivative"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1132
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
  1133
definition
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
  1134
  "RN_deriv M \<nu> \<equiv> SOME f. f \<in> borel_measurable M \<and> (\<forall>x. 0 \<le> f x) \<and>
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1135
    (\<forall>A \<in> sets M. \<nu> A = (\<integral>\<^isup>+x. f x * indicator A x \<partial>M))"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1136
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1137
lemma (in sigma_finite_measure) RN_deriv_cong:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1138
  assumes cong: "\<And>A. A \<in> sets M \<Longrightarrow> measure M' A = \<mu> A" "sets M' = sets M" "space M' = space M"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1139
    and \<nu>: "\<And>A. A \<in> sets M \<Longrightarrow> \<nu>' A = \<nu> A"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1140
  shows "RN_deriv M' \<nu>' x = RN_deriv M \<nu> x"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1141
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
  1142
  interpret \<mu>': sigma_finite_measure M'
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1143
    using cong by (rule sigma_finite_measure_cong)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1144
  show ?thesis
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1145
    unfolding RN_deriv_def
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1146
    by (simp add: cong \<nu> positive_integral_cong_measure[OF cong] measurable_def)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1147
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1148
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1149
lemma (in sigma_finite_measure) RN_deriv:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1150
  assumes "measure_space (M\<lparr>measure := \<nu>\<rparr>)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1151
  assumes "absolutely_continuous \<nu>"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1152
  shows "RN_deriv M \<nu> \<in> borel_measurable M" (is ?borel)
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1153
  and "\<And>A. A \<in> sets M \<Longrightarrow> \<nu> A = (\<integral>\<^isup>+x. RN_deriv M \<nu> x * indicator A x \<partial>M)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1154
    (is "\<And>A. _ \<Longrightarrow> ?int A")
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
  1155
  and "0 \<le> RN_deriv M \<nu> x"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1156
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1157
  note Ex = Radon_Nikodym[OF assms, unfolded Bex_def]
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
  1158
  then show ?borel unfolding RN_deriv_def by (rule someI2_ex) auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
  1159
  from Ex show "0 \<le> RN_deriv M \<nu> x" unfolding RN_deriv_def
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
  1160
    by (rule someI2_ex) simp
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1161
  fix A assume "A \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1162
  from Ex show "?int A" unfolding RN_deriv_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1163
    by (rule someI2_ex) (simp add: `A \<in> sets M`)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1164
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1165
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1166
lemma (in sigma_finite_measure) RN_deriv_positive_integral:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1167
  assumes \<nu>: "measure_space (M\<lparr>measure := \<nu>\<rparr>)" "absolutely_continuous \<nu>"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1168
    and f: "f \<in> borel_measurable M"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1169
  shows "integral\<^isup>P (M\<lparr>measure := \<nu>\<rparr>) f = (\<integral>\<^isup>+x. RN_deriv M \<nu> x * f x \<partial>M)"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1170
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
  1171
  interpret \<nu>: measure_space "M\<lparr>measure := \<nu>\<rparr>" by fact
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
  1172
  note RN = RN_deriv[OF \<nu>]
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
  1173
  have "integral\<^isup>P (M\<lparr>measure := \<nu>\<rparr>) f = (\<integral>\<^isup>+x. max 0 (f x) \<partial>M\<lparr>measure := \<nu>\<rparr>)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
  1174
    unfolding positive_integral_max_0 ..
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
  1175
  also have "(\<integral>\<^isup>+x. max 0 (f x) \<partial>M\<lparr>measure := \<nu>\<rparr>) =
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
  1176
    (\<integral>\<^isup>+x. max 0 (f x) \<partial>M\<lparr>measure := \<lambda>A. (\<integral>\<^isup>+x. RN_deriv M \<nu> x * indicator A x \<partial>M)\<rparr>)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
  1177
    by (intro \<nu>.positive_integral_cong_measure[symmetric]) (simp_all add: RN(2))
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
  1178
  also have "\<dots> = (\<integral>\<^isup>+x. RN_deriv M \<nu> x * max 0 (f x) \<partial>M)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
  1179
    by (intro positive_integral_translated_density) (auto simp add: RN f)
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
  1180
  also have "\<dots> = (\<integral>\<^isup>+x. RN_deriv M \<nu> x * f x \<partial>M)"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
  1181
    using RN_deriv(3)[OF \<nu>]
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
  1182
    by (auto intro!: positive_integral_cong_pos split: split_if_asm
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43556
diff changeset
  1183
             simp: max_def ereal_mult_le_0_iff)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1184
  finally show ?thesis .
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1185
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1186
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1187
lemma (in sigma_finite_measure) RN_deriv_unique:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1188
  assumes \<nu>: "measure_space (M\<lparr>measure := \<nu>\<rparr>)" "absolutely_continuous \<nu>"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
  1189
  and f: "f \<in> borel_measurable M" "AE x. 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
  1190
  and eq: "\<And>A. A \<in> sets M \<Longrightarrow> \<nu> A = (\<integral>\<^isup>+x. f x * indicator A x \<partial>M)"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1191
  shows "AE x. f x = RN_deriv M \<nu> x"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1192
proof (rule density_unique[OF f RN_deriv(1)[OF \<nu>]])
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
  1193
  show "AE x. 0 \<le> RN_deriv M \<nu> x" using RN_deriv[OF \<nu>] by auto
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1194
  fix A assume A: "A \<in> sets M"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1195
  show "(\<integral>\<^isup>+x. f x * indicator A x \<partial>M) = (\<integral>\<^isup>+x. RN_deriv M \<nu> x * indicator A x \<partial>M)"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1196
    unfolding eq[OF A, symmetric] RN_deriv(2)[OF \<nu> A, symmetric] ..
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1197
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1198
41832
27cb9113b1a0 add lemma RN_deriv_vimage
hoelzl
parents: 41705
diff changeset
  1199
lemma (in sigma_finite_measure) RN_deriv_vimage:
27cb9113b1a0 add lemma RN_deriv_vimage
hoelzl
parents: 41705
diff changeset
  1200
  assumes T: "T \<in> measure_preserving M M'"
27cb9113b1a0 add lemma RN_deriv_vimage
hoelzl
parents: 41705
diff changeset
  1201
    and T': "T' \<in> measure_preserving (M'\<lparr> measure := \<nu>' \<rparr>) (M\<lparr> measure := \<nu> \<rparr>)"
27cb9113b1a0 add lemma RN_deriv_vimage
hoelzl
parents: 41705
diff changeset
  1202
    and inv: "\<And>x. x \<in> space M \<Longrightarrow> T' (T x) = x"
27cb9113b1a0 add lemma RN_deriv_vimage
hoelzl
parents: 41705
diff changeset
  1203
  and \<nu>': "measure_space (M'\<lparr>measure := \<nu>'\<rparr>)" "measure_space.absolutely_continuous M' \<nu>'"
27cb9113b1a0 add lemma RN_deriv_vimage
hoelzl
parents: 41705
diff changeset
  1204
  shows "AE x. RN_deriv M' \<nu>' (T x) = RN_deriv M \<nu> x"
27cb9113b1a0 add lemma RN_deriv_vimage
hoelzl
parents: 41705
diff changeset
  1205
proof (rule RN_deriv_unique)
27cb9113b1a0 add lemma RN_deriv_vimage
hoelzl
parents: 41705
diff changeset
  1206
  interpret \<nu>': measure_space "M'\<lparr>measure := \<nu>'\<rparr>" by fact
27cb9113b1a0 add lemma RN_deriv_vimage
hoelzl
parents: 41705
diff changeset
  1207
  show "measure_space (M\<lparr> measure := \<nu>\<rparr>)"
27cb9113b1a0 add lemma RN_deriv_vimage
hoelzl
parents: 41705
diff changeset
  1208
    by (rule \<nu>'.measure_space_vimage[OF _ T'], simp) default
27cb9113b1a0 add lemma RN_deriv_vimage
hoelzl
parents: 41705
diff changeset
  1209
  interpret M': measure_space M'
27cb9113b1a0 add lemma RN_deriv_vimage
hoelzl
parents: 41705
diff changeset
  1210
  proof (rule measure_space_vimage)
27cb9113b1a0 add lemma RN_deriv_vimage
hoelzl
parents: 41705
diff changeset
  1211
    have "sigma_algebra (M'\<lparr> measure := \<nu>'\<rparr>)" by default
27cb9113b1a0 add lemma RN_deriv_vimage
hoelzl
parents: 41705
diff changeset
  1212
    then show "sigma_algebra M'" by simp
27cb9113b1a0 add lemma RN_deriv_vimage
hoelzl
parents: 41705
diff changeset
  1213
  qed fact
27cb9113b1a0 add lemma RN_deriv_vimage
hoelzl
parents: 41705
diff changeset
  1214
  show "absolutely_continuous \<nu>" unfolding absolutely_continuous_def
27cb9113b1a0 add lemma RN_deriv_vimage
hoelzl
parents: 41705
diff changeset
  1215
  proof safe
27cb9113b1a0 add lemma RN_deriv_vimage
hoelzl
parents: 41705
diff changeset
  1216
    fix N assume N: "N \<in> sets M" and N_0: "\<mu> N = 0"
27cb9113b1a0 add lemma RN_deriv_vimage
hoelzl
parents: 41705
diff changeset
  1217
    then have N': "T' -` N \<inter> space M' \<in> sets M'"
27cb9113b1a0 add lemma RN_deriv_vimage
hoelzl
parents: 41705
diff changeset
  1218
      using T' by (auto simp: measurable_def measure_preserving_def)
27cb9113b1a0 add lemma RN_deriv_vimage
hoelzl
parents: 41705
diff changeset
  1219
    have "T -` (T' -` N \<inter> space M') \<inter> space M = N"
27cb9113b1a0 add lemma RN_deriv_vimage
hoelzl
parents: 41705
diff changeset
  1220
      using inv T N sets_into_space[OF N] by (auto simp: measurable_def measure_preserving_def)
27cb9113b1a0 add lemma RN_deriv_vimage
hoelzl
parents: 41705
diff changeset
  1221
    then have "measure M' (T' -` N \<inter> space M') = 0"
27cb9113b1a0 add lemma RN_deriv_vimage
hoelzl
parents: 41705
diff changeset
  1222
      using measure_preservingD[OF T N'] N_0 by auto
27cb9113b1a0 add lemma RN_deriv_vimage
hoelzl
parents: 41705
diff changeset
  1223
    with \<nu>'(2) N' show "\<nu> N = 0" using measure_preservingD[OF T', of N] N
27cb9113b1a0 add lemma RN_deriv_vimage
hoelzl
parents: 41705
diff changeset
  1224
      unfolding M'.absolutely_continuous_def measurable_def by auto
27cb9113b1a0 add lemma RN_deriv_vimage
hoelzl
parents: 41705
diff changeset
  1225
  qed
27cb9113b1a0 add lemma RN_deriv_vimage
hoelzl
parents: 41705
diff changeset
  1226
  interpret M': sigma_finite_measure M'
27cb9113b1a0 add lemma RN_deriv_vimage
hoelzl
parents: 41705
diff changeset
  1227
  proof
27cb9113b1a0 add lemma RN_deriv_vimage
hoelzl
parents: 41705
diff changeset
  1228
    from sigma_finite guess F .. note F = this
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
  1229
    show "\<exists>A::nat \<Rightarrow> 'c set. range A \<subseteq> sets M' \<and> (\<Union>i. A i) = space M' \<and> (\<forall>i. M'.\<mu> (A i) \<noteq> \<infinity>)"
41832
27cb9113b1a0 add lemma RN_deriv_vimage
hoelzl
parents: 41705
diff changeset
  1230
    proof (intro exI conjI allI)
27cb9113b1a0 add lemma RN_deriv_vimage
hoelzl
parents: 41705
diff changeset
  1231
      show *: "range (\<lambda>i. T' -` F i \<inter> space M') \<subseteq> sets M'"
27cb9113b1a0 add lemma RN_deriv_vimage
hoelzl
parents: 41705
diff changeset
  1232
        using F T' by (auto simp: measurable_def measure_preserving_def)
27cb9113b1a0 add lemma RN_deriv_vimage
hoelzl
parents: 41705
diff changeset
  1233
      show "(\<Union>i. T' -` F i \<inter> space M') = space M'"
27cb9113b1a0 add lemma RN_deriv_vimage
hoelzl
parents: 41705
diff changeset
  1234
        using F T' by (force simp: measurable_def measure_preserving_def)
27cb9113b1a0 add lemma RN_deriv_vimage
hoelzl
parents: 41705
diff changeset
  1235
      fix i
27cb9113b1a0 add lemma RN_deriv_vimage
hoelzl
parents: 41705
diff changeset
  1236
      have "T' -` F i \<inter> space M' \<in> sets M'" using * by auto
27cb9113b1a0 add lemma RN_deriv_vimage
hoelzl
parents: 41705
diff changeset
  1237
      note measure_preservingD[OF T this, symmetric]
27cb9113b1a0 add lemma RN_deriv_vimage
hoelzl
parents: 41705
diff changeset
  1238
      moreover
27cb9113b1a0 add lemma RN_deriv_vimage
hoelzl
parents: 41705
diff changeset
  1239
      have Fi: "F i \<in> sets M" using F by auto
27cb9113b1a0 add lemma RN_deriv_vimage
hoelzl
parents: 41705
diff changeset
  1240
      then have "T -` (T' -` F i \<inter> space M') \<inter> space M = F i"
27cb9113b1a0 add lemma RN_deriv_vimage
hoelzl
parents: 41705
diff changeset
  1241
        using T inv sets_into_space[OF Fi]
27cb9113b1a0 add lemma RN_deriv_vimage
hoelzl
parents: 41705
diff changeset
  1242
        by (auto simp: measurable_def measure_preserving_def)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
  1243
      ultimately show "measure M' (T' -` F i \<inter> space M') \<noteq> \<infinity>"
41832
27cb9113b1a0 add lemma RN_deriv_vimage
hoelzl
parents: 41705
diff changeset
  1244
        using F by simp
27cb9113b1a0 add lemma RN_deriv_vimage
hoelzl
parents: 41705
diff changeset
  1245
    qed
27cb9113b1a0 add lemma RN_deriv_vimage
hoelzl
parents: 41705
diff changeset
  1246
  qed
27cb9113b1a0 add lemma RN_deriv_vimage
hoelzl
parents: 41705
diff changeset
  1247
  have "(RN_deriv M' \<nu>') \<circ> T \<in> borel_measurable M"
27cb9113b1a0 add lemma RN_deriv_vimage
hoelzl
parents: 41705
diff changeset
  1248
    by (intro measurable_comp[where b=M'] M'.RN_deriv(1) measure_preservingD2[OF T]) fact+
27cb9113b1a0 add lemma RN_deriv_vimage
hoelzl
parents: 41705
diff changeset
  1249
  then show "(\<lambda>x. RN_deriv M' \<nu>' (T x)) \<in> borel_measurable M"
27cb9113b1a0 add lemma RN_deriv_vimage
hoelzl
parents: 41705
diff changeset
  1250
    by (simp add: comp_def)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
  1251
  show "AE x. 0 \<le> RN_deriv M' \<nu>' (T x)" using M'.RN_deriv(3)[OF \<nu>'] by auto
41832
27cb9113b1a0 add lemma RN_deriv_vimage
hoelzl
parents: 41705
diff changeset
  1252
  fix A let ?A = "T' -` A \<inter> space M'"
27cb9113b1a0 add lemma RN_deriv_vimage
hoelzl
parents: 41705
diff changeset
  1253
  assume A: "A \<in> sets M"
27cb9113b1a0 add lemma RN_deriv_vimage
hoelzl
parents: 41705
diff changeset
  1254
  then have A': "?A \<in> sets M'" using T' unfolding measurable_def measure_preserving_def
27cb9113b1a0 add lemma RN_deriv_vimage
hoelzl
parents: 41705
diff changeset
  1255
    by auto
27cb9113b1a0 add lemma RN_deriv_vimage
hoelzl
parents: 41705
diff changeset
  1256
  from A have "\<nu> A = \<nu>' ?A" using T'[THEN measure_preservingD] by simp
27cb9113b1a0 add lemma RN_deriv_vimage
hoelzl
parents: 41705
diff changeset
  1257
  also have "\<dots> = \<integral>\<^isup>+ x. RN_deriv M' \<nu>' x * indicator ?A x \<partial>M'"
27cb9113b1a0 add lemma RN_deriv_vimage
hoelzl
parents: 41705
diff changeset
  1258
    using A' by (rule M'.RN_deriv(2)[OF \<nu>'])
27cb9113b1a0 add lemma RN_deriv_vimage
hoelzl
parents: 41705
diff changeset
  1259
  also have "\<dots> = \<integral>\<^isup>+ x. RN_deriv M' \<nu>' (T x) * indicator ?A (T x) \<partial>M"
27cb9113b1a0 add lemma RN_deriv_vimage
hoelzl
parents: 41705
diff changeset
  1260
  proof (rule positive_integral_vimage)
27cb9113b1a0 add lemma RN_deriv_vimage
hoelzl
parents: 41705
diff changeset
  1261
    show "sigma_algebra M'" by default
27cb9113b1a0 add lemma RN_deriv_vimage
hoelzl
parents: 41705
diff changeset
  1262
    show "(\<lambda>x. RN_deriv M' \<nu>' x * indicator (T' -` A \<inter> space M') x) \<in> borel_measurable M'"
27cb9113b1a0 add lemma RN_deriv_vimage
hoelzl
parents: 41705
diff changeset
  1263
      by (auto intro!: A' M'.RN_deriv(1)[OF \<nu>'])
27cb9113b1a0 add lemma RN_deriv_vimage
hoelzl
parents: 41705
diff changeset
  1264
  qed fact
27cb9113b1a0 add lemma RN_deriv_vimage
hoelzl
parents: 41705
diff changeset
  1265
  also have "\<dots> = \<integral>\<^isup>+ x. RN_deriv M' \<nu>' (T x) * indicator A x \<partial>M"
27cb9113b1a0 add lemma RN_deriv_vimage
hoelzl
parents: 41705
diff changeset
  1266
    using T inv by (auto intro!: positive_integral_cong simp: measure_preserving_def measurable_def indicator_def)
27cb9113b1a0 add lemma RN_deriv_vimage
hoelzl
parents: 41705
diff changeset
  1267
  finally show "\<nu> A = \<integral>\<^isup>+ x. RN_deriv M' \<nu>' (T x) * indicator A x \<partial>M" .
27cb9113b1a0 add lemma RN_deriv_vimage
hoelzl
parents: 41705
diff changeset
  1268
qed
27cb9113b1a0 add lemma RN_deriv_vimage
hoelzl
parents: 41705
diff changeset
  1269
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1270
lemma (in sigma_finite_measure) RN_deriv_finite:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1271
  assumes sfm: "sigma_finite_measure (M\<lparr>measure := \<nu>\<rparr>)" and ac: "absolutely_continuous \<nu>"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
  1272
  shows "AE x. RN_deriv M \<nu> x \<noteq> \<infinity>"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1273
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
  1274
  interpret \<nu>: sigma_finite_measure "M\<lparr>measure := \<nu>\<rparr>" by fact
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1275
  have \<nu>: "measure_space (M\<lparr>measure := \<nu>\<rparr>)" by default
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1276
  from sfm show ?thesis
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
  1277
    using sigma_finite_iff_density_finite[OF \<nu> RN_deriv(1)[OF \<nu> ac]] RN_deriv(2,3)[OF \<nu> ac] by simp
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1278
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1279
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1280
lemma (in sigma_finite_measure)
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1281
  assumes \<nu>: "sigma_finite_measure (M\<lparr>measure := \<nu>\<rparr>)" "absolutely_continuous \<nu>"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1282
    and f: "f \<in> borel_measurable M"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1283
  shows RN_deriv_integrable: "integrable (M\<lparr>measure := \<nu>\<rparr>) f \<longleftrightarrow>
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1284
      integrable M (\<lambda>x. real (RN_deriv M \<nu> x) * f x)" (is ?integrable)
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1285
    and RN_deriv_integral: "integral\<^isup>L (M\<lparr>measure := \<nu>\<rparr>) f =
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1286
      (\<integral>x. real (RN_deriv M \<nu> x) * f x \<partial>M)" (is ?integral)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1287
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
  1288
  interpret \<nu>: sigma_finite_measure "M\<lparr>measure := \<nu>\<rparr>" by fact
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1289
  have ms: "measure_space (M\<lparr>measure := \<nu>\<rparr>)" by default
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43556
diff changeset
  1290
  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
  1291
  have f': "(\<lambda>x. - f x) \<in> borel_measurable M" using f by auto
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
  1292
  have Nf: "f \<in> borel_measurable (M\<lparr>measure := \<nu>\<rparr>)" using f by simp
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1293
  { fix f :: "'a \<Rightarrow> real"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
  1294
    { fix x assume *: "RN_deriv M \<nu> x \<noteq> \<infinity>"
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43556
diff changeset
  1295
      have "ereal (real (RN_deriv M \<nu> x)) * ereal (f x) = ereal (real (RN_deriv M \<nu> x) * f x)"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1296
        by (simp add: mult_le_0_iff)
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43556
diff changeset
  1297
      then have "RN_deriv M \<nu> x * ereal (f x) = ereal (real (RN_deriv M \<nu> x) * f x)"
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43556
diff changeset
  1298
        using RN_deriv(3)[OF ms \<nu>(2)] * by (auto simp add: ereal_real split: split_if_asm) }
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43556
diff changeset
  1299
    then have "(\<integral>\<^isup>+x. ereal (real (RN_deriv M \<nu> x) * f x) \<partial>M) = (\<integral>\<^isup>+x. RN_deriv M \<nu> x * ereal (f x) \<partial>M)"
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43556
diff changeset
  1300
              "(\<integral>\<^isup>+x. ereal (- (real (RN_deriv M \<nu> x) * f x)) \<partial>M) = (\<integral>\<^isup>+x. RN_deriv M \<nu> x * ereal (- f x) \<partial>M)"
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43556
diff changeset
  1301
      using RN_deriv_finite[OF \<nu>] 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
  1302
      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
  1303
  note * = this
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1304
  show ?integral ?integrable
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
  1305
    unfolding lebesgue_integral_def integrable_def *
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
  1306
    using f RN_deriv(1)[OF ms \<nu>(2)]
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
  1307
    by (auto simp: RN_deriv_positive_integral[OF ms \<nu>(2)])
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1308
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39097
diff changeset
  1309
43340
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42866
diff changeset
  1310
lemma (in sigma_finite_measure) real_RN_deriv:
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42866
diff changeset
  1311
  assumes "finite_measure (M\<lparr>measure := \<nu>\<rparr>)" (is "finite_measure ?\<nu>")
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42866
diff changeset
  1312
  assumes ac: "absolutely_continuous \<nu>"
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42866
diff changeset
  1313
  obtains D where "D \<in> borel_measurable M"
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43556
diff changeset
  1314
    and "AE x. RN_deriv M \<nu> x = ereal (D x)"
43340
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42866
diff changeset
  1315
    and "AE x in M\<lparr>measure := \<nu>\<rparr>. 0 < D x"
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42866
diff changeset
  1316
    and "\<And>x. 0 \<le> D x"
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42866
diff changeset
  1317
proof
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42866
diff changeset
  1318
  interpret \<nu>: finite_measure ?\<nu> by fact
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42866
diff changeset
  1319
  have ms: "measure_space ?\<nu>" by default
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42866
diff changeset
  1320
  note RN = RN_deriv[OF ms ac]
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42866
diff changeset
  1321
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42866
diff changeset
  1322
  let ?RN = "\<lambda>t. {x \<in> space M. RN_deriv M \<nu> x = t}"
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42866
diff changeset
  1323
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42866
diff changeset
  1324
  show "(\<lambda>x. real (RN_deriv M \<nu> x)) \<in> borel_measurable M"
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42866
diff changeset
  1325
    using RN by auto
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42866
diff changeset
  1326
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42866
diff changeset
  1327
  have "\<nu> (?RN \<infinity>) = (\<integral>\<^isup>+ x. RN_deriv M \<nu> x * indicator (?RN \<infinity>) x \<partial>M)"
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42866
diff changeset
  1328
    using RN by auto
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42866
diff changeset
  1329
  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
  1330
    by (intro positive_integral_cong) (auto simp: indicator_def)
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42866
diff changeset
  1331
  also have "\<dots> = \<infinity> * \<mu> (?RN \<infinity>)"
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42866
diff changeset
  1332
    using RN by (intro positive_integral_cmult_indicator) auto
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42866
diff changeset
  1333
  finally have eq: "\<nu> (?RN \<infinity>) = \<infinity> * \<mu> (?RN \<infinity>)" .
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42866
diff changeset
  1334
  moreover
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42866
diff changeset
  1335
  have "\<mu> (?RN \<infinity>) = 0"
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42866
diff changeset
  1336
  proof (rule ccontr)
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42866
diff changeset
  1337
    assume "\<mu> {x \<in> space M. RN_deriv M \<nu> x = \<infinity>} \<noteq> 0"
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42866
diff changeset
  1338
    moreover from RN have "0 \<le> \<mu> {x \<in> space M. RN_deriv M \<nu> x = \<infinity>}" by auto
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42866
diff changeset
  1339
    ultimately have "0 < \<mu> {x \<in> space M. RN_deriv M \<nu> x = \<infinity>}" by auto
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42866
diff changeset
  1340
    with eq have "\<nu> (?RN \<infinity>) = \<infinity>" by simp
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42866
diff changeset
  1341
    with \<nu>.finite_measure[of "?RN \<infinity>"] RN show False by auto
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42866
diff changeset
  1342
  qed
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42866
diff changeset
  1343
  ultimately have "AE x. RN_deriv M \<nu> x < \<infinity>"
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42866
diff changeset
  1344
    using RN by (intro AE_iff_measurable[THEN iffD2]) auto
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43556
diff changeset
  1345
  then show "AE x. RN_deriv M \<nu> x = ereal (real (RN_deriv M \<nu> x))"
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43556
diff changeset
  1346
    using RN(3) by (auto simp: ereal_real)
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43556
diff changeset
  1347
  then have eq: "AE x in (M\<lparr>measure := \<nu>\<rparr>) . RN_deriv M \<nu> x = ereal (real (RN_deriv M \<nu> x))"
43340
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42866
diff changeset
  1348
    using ac absolutely_continuous_AE[OF ms] by auto
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42866
diff changeset
  1349
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42866
diff changeset
  1350
  show "\<And>x. 0 \<le> real (RN_deriv M \<nu> x)"
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43556
diff changeset
  1351
    using RN by (auto intro: real_of_ereal_pos)
43340
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42866
diff changeset
  1352
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42866
diff changeset
  1353
  have "\<nu> (?RN 0) = (\<integral>\<^isup>+ x. RN_deriv M \<nu> x * indicator (?RN 0) x \<partial>M)"
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42866
diff changeset
  1354
    using RN by auto
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42866
diff changeset
  1355
  also have "\<dots> = (\<integral>\<^isup>+ x. 0 \<partial>M)"
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42866
diff changeset
  1356
    by (intro positive_integral_cong) (auto simp: indicator_def)
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42866
diff changeset
  1357
  finally have "AE x in (M\<lparr>measure := \<nu>\<rparr>). RN_deriv M \<nu> x \<noteq> 0"
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42866
diff changeset
  1358
    using RN by (intro \<nu>.AE_iff_measurable[THEN iffD2]) auto
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42866
diff changeset
  1359
  with RN(3) eq show "AE x in (M\<lparr>measure := \<nu>\<rparr>). 0 < real (RN_deriv M \<nu> x)"
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43556
diff changeset
  1360
    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
  1361
qed
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42866
diff changeset
  1362
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1363
lemma (in sigma_finite_measure) RN_deriv_singleton:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1364
  assumes "measure_space (M\<lparr>measure := \<nu>\<rparr>)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1365
  and ac: "absolutely_continuous \<nu>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1366
  and "{x} \<in> sets M"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1367
  shows "\<nu> {x} = RN_deriv M \<nu> x * \<mu> {x}"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1368
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1369
  note deriv = RN_deriv[OF assms(1, 2)]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1370
  from deriv(2)[OF `{x} \<in> sets M`]
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1371
  have "\<nu> {x} = (\<integral>\<^isup>+w. RN_deriv M \<nu> x * indicator {x} w \<partial>M)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1372
    by (auto simp: indicator_def intro!: positive_integral_cong)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41832
diff changeset
  1373
  thus ?thesis using positive_integral_cmult_indicator[OF _ `{x} \<in> sets M`] deriv(3)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1374
    by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1375
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1376
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1377
theorem (in finite_measure_space) RN_deriv_finite_measure:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1378
  assumes "measure_space (M\<lparr>measure := \<nu>\<rparr>)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1379
  and ac: "absolutely_continuous \<nu>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1380
  and "x \<in> space M"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1381
  shows "\<nu> {x} = RN_deriv M \<nu> x * \<mu> {x}"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1382
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1383
  have "{x} \<in> sets M" using sets_eq_Pow `x \<in> space M` by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1384
  from RN_deriv_singleton[OF assms(1,2) this] show ?thesis .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1385
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1386
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1387
end