src/HOL/Probability/Lebesgue_Integration.thy
author hoelzl
Tue, 24 Aug 2010 14:41:37 +0200
changeset 38705 aaee86c0e237
parent 38656 d5d342611edb
child 39092 98de40859858
permissions -rw-r--r--
moved generic lemmas in Probability to HOL
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
     1
(* Author: Armin Heller, Johannes Hoelzl, TU Muenchen *)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
     2
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
     3
header {*Lebesgue Integration*}
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
     4
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
     5
theory Lebesgue_Integration
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
     6
imports Measure Borel
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
     7
begin
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
     8
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
     9
section "@{text \<mu>}-null sets"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
    10
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    11
abbreviation (in measure_space) "null_sets \<equiv> {N\<in>sets M. \<mu> N = 0}"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
    12
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    13
lemma sums_If_finite:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    14
  assumes finite: "finite {r. P r}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    15
  shows "(\<lambda>r. if P r then f r else 0) sums (\<Sum>r\<in>{r. P r}. f r)" (is "?F sums _")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    16
proof cases
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    17
  assume "{r. P r} = {}" hence "\<And>r. \<not> P r" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    18
  thus ?thesis by (simp add: sums_zero)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    19
next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    20
  assume not_empty: "{r. P r} \<noteq> {}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    21
  have "?F sums (\<Sum>r = 0..< Suc (Max {r. P r}). ?F r)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    22
    by (rule series_zero)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    23
       (auto simp add: Max_less_iff[OF finite not_empty] less_eq_Suc_le[symmetric])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    24
  also have "(\<Sum>r = 0..< Suc (Max {r. P r}). ?F r) = (\<Sum>r\<in>{r. P r}. f r)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    25
    by (subst setsum_cases)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    26
       (auto intro!: setsum_cong simp: Max_ge_iff[OF finite not_empty] less_Suc_eq_le)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    27
  finally show ?thesis .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    28
qed
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
    29
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    30
lemma sums_single:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    31
  "(\<lambda>r. if r = i then f r else 0) sums f i"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    32
  using sums_If_finite[of "\<lambda>r. r = i" f] by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    33
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    34
section "Simple function"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
    35
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    36
text {*
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    37
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    38
Our simple functions are not restricted to positive real numbers. Instead
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    39
they are just functions with a finite range and are measurable when singleton
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    40
sets are measurable.
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
    41
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    42
*}
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    43
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    44
definition (in sigma_algebra) "simple_function g \<longleftrightarrow>
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    45
    finite (g ` space M) \<and>
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    46
    (\<forall>x \<in> g ` space M. g -` {x} \<inter> space M \<in> sets M)"
36624
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
    47
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    48
lemma (in sigma_algebra) simple_functionD:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    49
  assumes "simple_function g"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    50
  shows "finite (g ` space M)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    51
  "x \<in> g ` space M \<Longrightarrow> g -` {x} \<inter> space M \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    52
  using assms unfolding simple_function_def by auto
36624
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
    53
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    54
lemma (in sigma_algebra) simple_function_indicator_representation:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    55
  fixes f ::"'a \<Rightarrow> pinfreal"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    56
  assumes f: "simple_function f" and x: "x \<in> space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    57
  shows "f x = (\<Sum>y \<in> f ` space M. y * indicator (f -` {y} \<inter> space M) x)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    58
  (is "?l = ?r")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    59
proof -
38705
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 38656
diff changeset
    60
  have "?r = (\<Sum>y \<in> f ` space M.
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    61
    (if y = f x then y * indicator (f -` {y} \<inter> space M) x else 0))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    62
    by (auto intro!: setsum_cong2)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    63
  also have "... =  f x *  indicator (f -` {f x} \<inter> space M) x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    64
    using assms by (auto dest: simple_functionD simp: setsum_delta)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    65
  also have "... = f x" using x by (auto simp: indicator_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    66
  finally show ?thesis by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    67
qed
36624
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
    68
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    69
lemma (in measure_space) simple_function_notspace:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    70
  "simple_function (\<lambda>x. h x * indicator (- space M) x::pinfreal)" (is "simple_function ?h")
35692
f1315bbf1bc9 Moved theorems in Lebesgue to the right places
hoelzl
parents: 35582
diff changeset
    71
proof -
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    72
  have "?h ` space M \<subseteq> {0}" unfolding indicator_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    73
  hence [simp, intro]: "finite (?h ` space M)" by (auto intro: finite_subset)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    74
  have "?h -` {0} \<inter> space M = space M" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    75
  thus ?thesis unfolding simple_function_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    76
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    77
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    78
lemma (in sigma_algebra) simple_function_cong:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    79
  assumes "\<And>t. t \<in> space M \<Longrightarrow> f t = g t"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    80
  shows "simple_function f \<longleftrightarrow> simple_function g"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    81
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    82
  have "f ` space M = g ` space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    83
    "\<And>x. f -` {x} \<inter> space M = g -` {x} \<inter> space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    84
    using assms by (auto intro!: image_eqI)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    85
  thus ?thesis unfolding simple_function_def using assms by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    86
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    87
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    88
lemma (in sigma_algebra) borel_measurable_simple_function:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    89
  assumes "simple_function f"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    90
  shows "f \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    91
proof (rule borel_measurableI)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    92
  fix S
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    93
  let ?I = "f ` (f -` S \<inter> space M)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    94
  have *: "(\<Union>x\<in>?I. f -` {x} \<inter> space M) = f -` S \<inter> space M" (is "?U = _") by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    95
  have "finite ?I"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    96
    using assms unfolding simple_function_def by (auto intro: finite_subset)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    97
  hence "?U \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    98
    apply (rule finite_UN)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    99
    using assms unfolding simple_function_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   100
  thus "f -` S \<inter> space M \<in> sets M" unfolding * .
35692
f1315bbf1bc9 Moved theorems in Lebesgue to the right places
hoelzl
parents: 35582
diff changeset
   101
qed
f1315bbf1bc9 Moved theorems in Lebesgue to the right places
hoelzl
parents: 35582
diff changeset
   102
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   103
lemma (in sigma_algebra) simple_function_borel_measurable:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   104
  fixes f :: "'a \<Rightarrow> 'x::t2_space"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   105
  assumes "f \<in> borel_measurable M" and "finite (f ` space M)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   106
  shows "simple_function f"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   107
  using assms unfolding simple_function_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   108
  by (auto intro: borel_measurable_vimage)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   109
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   110
lemma (in sigma_algebra) simple_function_const[intro, simp]:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   111
  "simple_function (\<lambda>x. c)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   112
  by (auto intro: finite_subset simp: simple_function_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   113
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   114
lemma (in sigma_algebra) simple_function_compose[intro, simp]:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   115
  assumes "simple_function f"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   116
  shows "simple_function (g \<circ> f)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   117
  unfolding simple_function_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   118
proof safe
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   119
  show "finite ((g \<circ> f) ` space M)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   120
    using assms unfolding simple_function_def by (auto simp: image_compose)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   121
next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   122
  fix x assume "x \<in> space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   123
  let ?G = "g -` {g (f x)} \<inter> (f`space M)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   124
  have *: "(g \<circ> f) -` {(g \<circ> f) x} \<inter> space M =
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   125
    (\<Union>x\<in>?G. f -` {x} \<inter> space M)" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   126
  show "(g \<circ> f) -` {(g \<circ> f) x} \<inter> space M \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   127
    using assms unfolding simple_function_def *
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   128
    by (rule_tac finite_UN) (auto intro!: finite_UN)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   129
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   130
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   131
lemma (in sigma_algebra) simple_function_indicator[intro, simp]:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   132
  assumes "A \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   133
  shows "simple_function (indicator A)"
35692
f1315bbf1bc9 Moved theorems in Lebesgue to the right places
hoelzl
parents: 35582
diff changeset
   134
proof -
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   135
  have "indicator A ` space M \<subseteq> {0, 1}" (is "?S \<subseteq> _")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   136
    by (auto simp: indicator_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   137
  hence "finite ?S" by (rule finite_subset) simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   138
  moreover have "- A \<inter> space M = space M - A" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   139
  ultimately show ?thesis unfolding simple_function_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   140
    using assms by (auto simp: indicator_def_raw)
35692
f1315bbf1bc9 Moved theorems in Lebesgue to the right places
hoelzl
parents: 35582
diff changeset
   141
qed
f1315bbf1bc9 Moved theorems in Lebesgue to the right places
hoelzl
parents: 35582
diff changeset
   142
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   143
lemma (in sigma_algebra) simple_function_Pair[intro, simp]:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   144
  assumes "simple_function f"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   145
  assumes "simple_function g"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   146
  shows "simple_function (\<lambda>x. (f x, g x))" (is "simple_function ?p")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   147
  unfolding simple_function_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   148
proof safe
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   149
  show "finite (?p ` space M)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   150
    using assms unfolding simple_function_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   151
    by (rule_tac finite_subset[of _ "f`space M \<times> g`space M"]) auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   152
next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   153
  fix x assume "x \<in> space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   154
  have "(\<lambda>x. (f x, g x)) -` {(f x, g x)} \<inter> space M =
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   155
      (f -` {f x} \<inter> space M) \<inter> (g -` {g x} \<inter> space M)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   156
    by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   157
  with `x \<in> space M` show "(\<lambda>x. (f x, g x)) -` {(f x, g x)} \<inter> space M \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   158
    using assms unfolding simple_function_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   159
qed
35692
f1315bbf1bc9 Moved theorems in Lebesgue to the right places
hoelzl
parents: 35582
diff changeset
   160
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   161
lemma (in sigma_algebra) simple_function_compose1:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   162
  assumes "simple_function f"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   163
  shows "simple_function (\<lambda>x. g (f x))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   164
  using simple_function_compose[OF assms, of g]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   165
  by (simp add: comp_def)
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   166
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   167
lemma (in sigma_algebra) simple_function_compose2:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   168
  assumes "simple_function f" and "simple_function g"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   169
  shows "simple_function (\<lambda>x. h (f x) (g x))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   170
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   171
  have "simple_function ((\<lambda>(x, y). h x y) \<circ> (\<lambda>x. (f x, g x)))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   172
    using assms by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   173
  thus ?thesis by (simp_all add: comp_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   174
qed
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   175
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   176
lemmas (in sigma_algebra) simple_function_add[intro, simp] = simple_function_compose2[where h="op +"]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   177
  and simple_function_diff[intro, simp] = simple_function_compose2[where h="op -"]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   178
  and simple_function_uminus[intro, simp] = simple_function_compose[where g="uminus"]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   179
  and simple_function_mult[intro, simp] = simple_function_compose2[where h="op *"]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   180
  and simple_function_div[intro, simp] = simple_function_compose2[where h="op /"]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   181
  and simple_function_inverse[intro, simp] = simple_function_compose[where g="inverse"]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   182
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   183
lemma (in sigma_algebra) simple_function_setsum[intro, simp]:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   184
  assumes "\<And>i. i \<in> P \<Longrightarrow> simple_function (f i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   185
  shows "simple_function (\<lambda>x. \<Sum>i\<in>P. f i x)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   186
proof cases
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   187
  assume "finite P" from this assms show ?thesis by induct auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   188
qed auto
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   189
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   190
lemma (in sigma_algebra) simple_function_le_measurable:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   191
  assumes "simple_function f" "simple_function g"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   192
  shows "{x \<in> space M. f x \<le> g x} \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   193
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   194
  have *: "{x \<in> space M. f x \<le> g x} =
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   195
    (\<Union>(F, G)\<in>f`space M \<times> g`space M.
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   196
      if F \<le> G then (f -` {F} \<inter> space M) \<inter> (g -` {G} \<inter> space M) else {})"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   197
    apply (auto split: split_if_asm)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   198
    apply (rule_tac x=x in bexI)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   199
    apply (rule_tac x=x in bexI)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   200
    by simp_all
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   201
  have **: "\<And>x y. x \<in> space M \<Longrightarrow> y \<in> space M \<Longrightarrow>
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   202
    (f -` {f x} \<inter> space M) \<inter> (g -` {g y} \<inter> space M) \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   203
    using assms unfolding simple_function_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   204
  have "finite (f`space M \<times> g`space M)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   205
    using assms unfolding simple_function_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   206
  thus ?thesis unfolding *
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   207
    apply (rule finite_UN)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   208
    using assms unfolding simple_function_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   209
    by (auto intro!: **)
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   210
qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   211
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   212
lemma setsum_indicator_disjoint_family:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   213
  fixes f :: "'d \<Rightarrow> 'e::semiring_1"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   214
  assumes d: "disjoint_family_on A P" and "x \<in> A j" and "finite P" and "j \<in> P"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   215
  shows "(\<Sum>i\<in>P. f i * indicator (A i) x) = f j"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   216
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   217
  have "P \<inter> {i. x \<in> A i} = {j}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   218
    using d `x \<in> A j` `j \<in> P` unfolding disjoint_family_on_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   219
    by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   220
  thus ?thesis
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   221
    unfolding indicator_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   222
    by (simp add: if_distrib setsum_cases[OF `finite P`])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   223
qed
35692
f1315bbf1bc9 Moved theorems in Lebesgue to the right places
hoelzl
parents: 35582
diff changeset
   224
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   225
lemma (in sigma_algebra) borel_measurable_implies_simple_function_sequence:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   226
  fixes u :: "'a \<Rightarrow> pinfreal"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   227
  assumes u: "u \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   228
  shows "\<exists>f. (\<forall>i. simple_function (f i) \<and> (\<forall>x\<in>space M. f i x \<noteq> \<omega>)) \<and> f \<up> u"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   229
proof -
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   230
  have "\<exists>f. \<forall>x j. (of_nat j \<le> u x \<longrightarrow> f x j = j*2^j) \<and>
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   231
    (u x < of_nat j \<longrightarrow> of_nat (f x j) \<le> u x * 2^j \<and> u x * 2^j < of_nat (Suc (f x j)))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   232
    (is "\<exists>f. \<forall>x j. ?P x j (f x j)")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   233
  proof(rule choice, rule, rule choice, rule)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   234
    fix x j show "\<exists>n. ?P x j n"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   235
    proof cases
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   236
      assume *: "u x < of_nat j"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   237
      then obtain r where r: "u x = Real r" "0 \<le> r" by (cases "u x") auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   238
      from reals_Archimedean6a[of "r * 2^j"]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   239
      obtain n :: nat where "real n \<le> r * 2 ^ j" "r * 2 ^ j < real (Suc n)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   240
        using `0 \<le> r` by (auto simp: zero_le_power zero_le_mult_iff)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   241
      thus ?thesis using r * by (auto intro!: exI[of _ n])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   242
    qed auto
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   243
  qed
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   244
  then obtain f where top: "\<And>j x. of_nat j \<le> u x \<Longrightarrow> f x j = j*2^j" and
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   245
    upper: "\<And>j x. u x < of_nat j \<Longrightarrow> of_nat (f x j) \<le> u x * 2^j" and
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   246
    lower: "\<And>j x. u x < of_nat j \<Longrightarrow> u x * 2^j < of_nat (Suc (f x j))" by blast
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   247
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   248
  { fix j x P
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   249
    assume 1: "of_nat j \<le> u x \<Longrightarrow> P (j * 2^j)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   250
    assume 2: "\<And>k. \<lbrakk> u x < of_nat j ; of_nat k \<le> u x * 2^j ; u x * 2^j < of_nat (Suc k) \<rbrakk> \<Longrightarrow> P k"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   251
    have "P (f x j)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   252
    proof cases
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   253
      assume "of_nat j \<le> u x" thus "P (f x j)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   254
        using top[of j x] 1 by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   255
    next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   256
      assume "\<not> of_nat j \<le> u x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   257
      hence "u x < of_nat j" "of_nat (f x j) \<le> u x * 2^j" "u x * 2^j < of_nat (Suc (f x j))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   258
        using upper lower by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   259
      from 2[OF this] show "P (f x j)" .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   260
    qed }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   261
  note fI = this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   262
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   263
  { fix j x
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   264
    have "f x j = j * 2 ^ j \<longleftrightarrow> of_nat j \<le> u x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   265
      by (rule fI, simp, cases "u x") (auto split: split_if_asm) }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   266
  note f_eq = this
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   267
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   268
  { fix j x
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   269
    have "f x j \<le> j * 2 ^ j"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   270
    proof (rule fI)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   271
      fix k assume *: "u x < of_nat j"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   272
      assume "of_nat k \<le> u x * 2 ^ j"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   273
      also have "\<dots> \<le> of_nat (j * 2^j)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   274
        using * by (cases "u x") (auto simp: zero_le_mult_iff)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   275
      finally show "k \<le> j*2^j" by (auto simp del: real_of_nat_mult)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   276
    qed simp }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   277
  note f_upper = this
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   278
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   279
  let "?g j x" = "of_nat (f x j) / 2^j :: pinfreal"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   280
  show ?thesis unfolding simple_function_def isoton_fun_expand unfolding isoton_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   281
  proof (safe intro!: exI[of _ ?g])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   282
    fix j
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   283
    have *: "?g j ` space M \<subseteq> (\<lambda>x. of_nat x / 2^j) ` {..j * 2^j}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   284
      using f_upper by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   285
    thus "finite (?g j ` space M)" by (rule finite_subset) auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   286
  next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   287
    fix j t assume "t \<in> space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   288
    have **: "?g j -` {?g j t} \<inter> space M = {x \<in> space M. f t j = f x j}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   289
      by (auto simp: power_le_zero_eq Real_eq_Real mult_le_0_iff)
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   290
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   291
    show "?g j -` {?g j t} \<inter> space M \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   292
    proof cases
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   293
      assume "of_nat j \<le> u t"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   294
      hence "?g j -` {?g j t} \<inter> space M = {x\<in>space M. of_nat j \<le> u x}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   295
        unfolding ** f_eq[symmetric] by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   296
      thus "?g j -` {?g j t} \<inter> space M \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   297
        using u by auto
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   298
    next
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   299
      assume not_t: "\<not> of_nat j \<le> u t"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   300
      hence less: "f t j < j*2^j" using f_eq[symmetric] `f t j \<le> j*2^j` by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   301
      have split_vimage: "?g j -` {?g j t} \<inter> space M =
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   302
          {x\<in>space M. of_nat (f t j)/2^j \<le> u x} \<inter> {x\<in>space M. u x < of_nat (Suc (f t j))/2^j}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   303
        unfolding **
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   304
      proof safe
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   305
        fix x assume [simp]: "f t j = f x j"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   306
        have *: "\<not> of_nat j \<le> u x" using not_t f_eq[symmetric] by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   307
        hence "of_nat (f t j) \<le> u x * 2^j \<and> u x * 2^j < of_nat (Suc (f t j))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   308
          using upper lower by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   309
        hence "of_nat (f t j) / 2^j \<le> u x \<and> u x < of_nat (Suc (f t j))/2^j" using *
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   310
          by (cases "u x") (auto simp: zero_le_mult_iff power_le_zero_eq divide_real_def[symmetric] field_simps)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   311
        thus "of_nat (f t j) / 2^j \<le> u x" "u x < of_nat (Suc (f t j))/2^j" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   312
      next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   313
        fix x
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   314
        assume "of_nat (f t j) / 2^j \<le> u x" "u x < of_nat (Suc (f t j))/2^j"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   315
        hence "of_nat (f t j) \<le> u x * 2 ^ j \<and> u x * 2 ^ j < of_nat (Suc (f t j))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   316
          by (cases "u x") (auto simp: zero_le_mult_iff power_le_zero_eq divide_real_def[symmetric] field_simps)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   317
        hence 1: "of_nat (f t j) \<le> u x * 2 ^ j" and 2: "u x * 2 ^ j < of_nat (Suc (f t j))" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   318
        note 2
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   319
        also have "\<dots> \<le> of_nat (j*2^j)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   320
          using less by (auto simp: zero_le_mult_iff simp del: real_of_nat_mult)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   321
        finally have bound_ux: "u x < of_nat j"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   322
          by (cases "u x") (auto simp: zero_le_mult_iff power_le_zero_eq)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   323
        show "f t j = f x j"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   324
        proof (rule antisym)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   325
          from 1 lower[OF bound_ux]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   326
          show "f t j \<le> f x j" by (cases "u x") (auto split: split_if_asm)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   327
          from upper[OF bound_ux] 2
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   328
          show "f x j \<le> f t j" by (cases "u x") (auto split: split_if_asm)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   329
        qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   330
      qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   331
      show ?thesis unfolding split_vimage using u by auto
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   332
    qed
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   333
  next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   334
    fix j t assume "?g j t = \<omega>" thus False by (auto simp: power_le_zero_eq)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   335
  next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   336
    fix t
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   337
    { fix i
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   338
      have "f t i * 2 \<le> f t (Suc i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   339
      proof (rule fI)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   340
        assume "of_nat (Suc i) \<le> u t"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   341
        hence "of_nat i \<le> u t" by (cases "u t") auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   342
        thus "f t i * 2 \<le> Suc i * 2 ^ Suc i" unfolding f_eq[symmetric] by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   343
      next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   344
        fix k
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   345
        assume *: "u t * 2 ^ Suc i < of_nat (Suc k)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   346
        show "f t i * 2 \<le> k"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   347
        proof (rule fI)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   348
          assume "of_nat i \<le> u t"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   349
          hence "of_nat (i * 2^Suc i) \<le> u t * 2^Suc i"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   350
            by (cases "u t") (auto simp: zero_le_mult_iff power_le_zero_eq)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   351
          also have "\<dots> < of_nat (Suc k)" using * by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   352
          finally show "i * 2 ^ i * 2 \<le> k"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   353
            by (auto simp del: real_of_nat_mult)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   354
        next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   355
          fix j assume "of_nat j \<le> u t * 2 ^ i"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   356
          with * show "j * 2 \<le> k" by (cases "u t") (auto simp: zero_le_mult_iff power_le_zero_eq)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   357
        qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   358
      qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   359
      thus "?g i t \<le> ?g (Suc i) t"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   360
        by (auto simp: zero_le_mult_iff power_le_zero_eq divide_real_def[symmetric] field_simps) }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   361
    hence mono: "mono (\<lambda>i. ?g i t)" unfolding mono_iff_le_Suc by auto
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   362
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   363
    show "(SUP j. of_nat (f t j) / 2 ^ j) = u t"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   364
    proof (rule pinfreal_SUPI)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   365
      fix j show "of_nat (f t j) / 2 ^ j \<le> u t"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   366
      proof (rule fI)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   367
        assume "of_nat j \<le> u t" thus "of_nat (j * 2 ^ j) / 2 ^ j \<le> u t"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   368
          by (cases "u t") (auto simp: power_le_zero_eq divide_real_def[symmetric] field_simps)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   369
      next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   370
        fix k assume "of_nat k \<le> u t * 2 ^ j"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   371
        thus "of_nat k / 2 ^ j \<le> u t"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   372
          by (cases "u t")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   373
             (auto simp: power_le_zero_eq divide_real_def[symmetric] field_simps zero_le_mult_iff)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   374
      qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   375
    next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   376
      fix y :: pinfreal assume *: "\<And>j. j \<in> UNIV \<Longrightarrow> of_nat (f t j) / 2 ^ j \<le> y"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   377
      show "u t \<le> y"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   378
      proof (cases "u t")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   379
        case (preal r)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   380
        show ?thesis
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   381
        proof (rule ccontr)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   382
          assume "\<not> u t \<le> y"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   383
          then obtain p where p: "y = Real p" "0 \<le> p" "p < r" using preal by (cases y) auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   384
          with LIMSEQ_inverse_realpow_zero[of 2, unfolded LIMSEQ_iff, rule_format, of "r - p"]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   385
          obtain n where n: "\<And>N. n \<le> N \<Longrightarrow> inverse (2^N) < r - p" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   386
          let ?N = "max n (natfloor r + 1)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   387
          have "u t < of_nat ?N" "n \<le> ?N"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   388
            using ge_natfloor_plus_one_imp_gt[of r n] preal
38705
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 38656
diff changeset
   389
            using real_natfloor_add_one_gt
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 38656
diff changeset
   390
            by (auto simp: max_def real_of_nat_Suc)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   391
          from lower[OF this(1)]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   392
          have "r < (real (f t ?N) + 1) / 2 ^ ?N" unfolding less_divide_eq
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   393
            using preal by (auto simp add: power_le_zero_eq zero_le_mult_iff real_of_nat_Suc split: split_if_asm)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   394
          hence "u t < of_nat (f t ?N) / 2 ^ ?N + 1 / 2 ^ ?N"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   395
            using preal by (auto simp: field_simps divide_real_def[symmetric])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   396
          with n[OF `n \<le> ?N`] p preal *[of ?N]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   397
          show False
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   398
            by (cases "f t ?N") (auto simp: power_le_zero_eq split: split_if_asm)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   399
        qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   400
      next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   401
        case infinite
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   402
        { fix j have "f t j = j*2^j" using top[of j t] infinite by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   403
          hence "of_nat j \<le> y" using *[of j]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   404
            by (cases y) (auto simp: power_le_zero_eq zero_le_mult_iff field_simps) }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   405
        note all_less_y = this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   406
        show ?thesis unfolding infinite
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   407
        proof (rule ccontr)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   408
          assume "\<not> \<omega> \<le> y" then obtain r where r: "y = Real r" "0 \<le> r" by (cases y) auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   409
          moreover obtain n ::nat where "r < real n" using ex_less_of_nat by (auto simp: real_eq_of_nat)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   410
          with all_less_y[of n] r show False by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   411
        qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   412
      qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   413
    qed
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   414
  qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   415
qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   416
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   417
lemma (in sigma_algebra) borel_measurable_implies_simple_function_sequence':
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   418
  fixes u :: "'a \<Rightarrow> pinfreal"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   419
  assumes "u \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   420
  obtains (x) f where "f \<up> u" "\<And>i. simple_function (f i)" "\<And>i. \<omega>\<notin>f i`space M"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   421
proof -
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   422
  from borel_measurable_implies_simple_function_sequence[OF assms]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   423
  obtain f where x: "\<And>i. simple_function (f i)" "f \<up> u"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   424
    and fin: "\<And>i. \<And>x. x\<in>space M \<Longrightarrow> f i x \<noteq> \<omega>" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   425
  { fix i from fin[of _ i] have "\<omega> \<notin> f i`space M" by fastsimp }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   426
  with x show thesis by (auto intro!: that[of f])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   427
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   428
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   429
section "Simple integral"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   430
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   431
definition (in measure_space)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   432
  "simple_integral f = (\<Sum>x \<in> f ` space M. x * \<mu> (f -` {x} \<inter> space M))"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   433
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   434
lemma (in measure_space) simple_integral_cong:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   435
  assumes "\<And>t. t \<in> space M \<Longrightarrow> f t = g t"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   436
  shows "simple_integral f = simple_integral g"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   437
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   438
  have "f ` space M = g ` space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   439
    "\<And>x. f -` {x} \<inter> space M = g -` {x} \<inter> space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   440
    using assms by (auto intro!: image_eqI)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   441
  thus ?thesis unfolding simple_integral_def by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   442
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   443
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   444
lemma (in measure_space) simple_integral_const[simp]:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   445
  "simple_integral (\<lambda>x. c) = c * \<mu> (space M)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   446
proof (cases "space M = {}")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   447
  case True thus ?thesis unfolding simple_integral_def by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   448
next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   449
  case False hence "(\<lambda>x. c) ` space M = {c}" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   450
  thus ?thesis unfolding simple_integral_def by simp
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   451
qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   452
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   453
lemma (in measure_space) simple_function_partition:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   454
  assumes "simple_function f" and "simple_function g"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   455
  shows "simple_integral f = (\<Sum>A\<in>(\<lambda>x. f -` {f x} \<inter> g -` {g x} \<inter> space M) ` space M. contents (f`A) * \<mu> A)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   456
    (is "_ = setsum _ (?p ` space M)")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   457
proof-
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   458
  let "?sub x" = "?p ` (f -` {x} \<inter> space M)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   459
  let ?SIGMA = "Sigma (f`space M) ?sub"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   460
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   461
  have [intro]:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   462
    "finite (f ` space M)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   463
    "finite (g ` space M)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   464
    using assms unfolding simple_function_def by simp_all
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   465
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   466
  { fix A
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   467
    have "?p ` (A \<inter> space M) \<subseteq>
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   468
      (\<lambda>(x,y). f -` {x} \<inter> g -` {y} \<inter> space M) ` (f`space M \<times> g`space M)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   469
      by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   470
    hence "finite (?p ` (A \<inter> space M))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   471
      by (rule finite_subset) (auto intro: finite_SigmaI finite_imageI) }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   472
  note this[intro, simp]
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   473
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   474
  { fix x assume "x \<in> space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   475
    have "\<Union>(?sub (f x)) = (f -` {f x} \<inter> space M)" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   476
    moreover {
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   477
      fix x y
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   478
      have *: "f -` {f x} \<inter> g -` {g x} \<inter> space M
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   479
          = (f -` {f x} \<inter> space M) \<inter> (g -` {g x} \<inter> space M)" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   480
      assume "x \<in> space M" "y \<in> space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   481
      hence "f -` {f x} \<inter> g -` {g x} \<inter> space M \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   482
        using assms unfolding simple_function_def * by auto }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   483
    ultimately
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   484
    have "\<mu> (f -` {f x} \<inter> space M) = setsum (\<mu>) (?sub (f x))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   485
      by (subst measure_finitely_additive) auto }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   486
  hence "simple_integral f = (\<Sum>(x,A)\<in>?SIGMA. x * \<mu> A)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   487
    unfolding simple_integral_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   488
    by (subst setsum_Sigma[symmetric],
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   489
       auto intro!: setsum_cong simp: setsum_right_distrib[symmetric])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   490
  also have "\<dots> = (\<Sum>A\<in>?p ` space M. contents (f`A) * \<mu> A)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   491
  proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   492
    have [simp]: "\<And>x. x \<in> space M \<Longrightarrow> f ` ?p x = {f x}" by (auto intro!: imageI)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   493
    have "(\<lambda>A. (contents (f ` A), A)) ` ?p ` space M
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   494
      = (\<lambda>x. (f x, ?p x)) ` space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   495
    proof safe
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   496
      fix x assume "x \<in> space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   497
      thus "(f x, ?p x) \<in> (\<lambda>A. (contents (f`A), A)) ` ?p ` space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   498
        by (auto intro!: image_eqI[of _ _ "?p x"])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   499
    qed auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   500
    thus ?thesis
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   501
      apply (auto intro!: setsum_reindex_cong[of "\<lambda>A. (contents (f`A), A)"] inj_onI)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   502
      apply (rule_tac x="xa" in image_eqI)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   503
      by simp_all
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   504
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   505
  finally show ?thesis .
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   506
qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   507
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   508
lemma (in measure_space) simple_integral_add[simp]:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   509
  assumes "simple_function f" and "simple_function g"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   510
  shows "simple_integral (\<lambda>x. f x + g x) = simple_integral f + simple_integral g"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   511
proof -
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   512
  { fix x let ?S = "g -` {g x} \<inter> f -` {f x} \<inter> space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   513
    assume "x \<in> space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   514
    hence "(\<lambda>a. f a + g a) ` ?S = {f x + g x}" "f ` ?S = {f x}" "g ` ?S = {g x}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   515
        "(\<lambda>x. (f x, g x)) -` {(f x, g x)} \<inter> space M = ?S"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   516
      by auto }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   517
  thus ?thesis
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   518
    unfolding
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   519
      simple_function_partition[OF simple_function_add[OF assms] simple_function_Pair[OF assms]]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   520
      simple_function_partition[OF `simple_function f` `simple_function g`]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   521
      simple_function_partition[OF `simple_function g` `simple_function f`]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   522
    apply (subst (3) Int_commute)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   523
    by (auto simp add: field_simps setsum_addf[symmetric] intro!: setsum_cong)
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   524
qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   525
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   526
lemma (in measure_space) simple_integral_setsum[simp]:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   527
  assumes "\<And>i. i \<in> P \<Longrightarrow> simple_function (f i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   528
  shows "simple_integral (\<lambda>x. \<Sum>i\<in>P. f i x) = (\<Sum>i\<in>P. simple_integral (f i))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   529
proof cases
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   530
  assume "finite P"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   531
  from this assms show ?thesis
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   532
    by induct (auto simp: simple_function_setsum simple_integral_add)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   533
qed auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   534
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   535
lemma (in measure_space) simple_integral_mult[simp]:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   536
  assumes "simple_function f"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   537
  shows "simple_integral (\<lambda>x. c * f x) = c * simple_integral f"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   538
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   539
  note mult = simple_function_mult[OF simple_function_const[of c] assms]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   540
  { fix x let ?S = "f -` {f x} \<inter> (\<lambda>x. c * f x) -` {c * f x} \<inter> space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   541
    assume "x \<in> space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   542
    hence "(\<lambda>x. c * f x) ` ?S = {c * f x}" "f ` ?S = {f x}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   543
      by auto }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   544
  thus ?thesis
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   545
    unfolding simple_function_partition[OF mult assms]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   546
      simple_function_partition[OF assms mult]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   547
    by (auto simp: setsum_right_distrib field_simps intro!: setsum_cong)
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   548
qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   549
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   550
lemma (in measure_space) simple_integral_mono:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   551
  assumes "simple_function f" and "simple_function g"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   552
  and mono: "\<And> x. x \<in> space M \<Longrightarrow> f x \<le> g x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   553
  shows "simple_integral f \<le> simple_integral g"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   554
  unfolding
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   555
    simple_function_partition[OF `simple_function f` `simple_function g`]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   556
    simple_function_partition[OF `simple_function g` `simple_function f`]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   557
  apply (subst Int_commute)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   558
proof (safe intro!: setsum_mono)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   559
  fix x let ?S = "g -` {g x} \<inter> f -` {f x} \<inter> space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   560
  assume "x \<in> space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   561
  hence "f ` ?S = {f x}" "g ` ?S = {g x}" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   562
  thus "contents (f`?S) * \<mu> ?S \<le> contents (g`?S) * \<mu> ?S"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   563
    using mono[OF `x \<in> space M`] by (auto intro!: mult_right_mono)
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   564
qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   565
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   566
lemma (in measure_space) simple_integral_indicator:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   567
  assumes "A \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   568
  assumes "simple_function f"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   569
  shows "simple_integral (\<lambda>x. f x * indicator A x) =
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   570
    (\<Sum>x \<in> f ` space M. x * \<mu> (f -` {x} \<inter> space M \<inter> A))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   571
proof cases
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   572
  assume "A = space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   573
  moreover hence "simple_integral (\<lambda>x. f x * indicator A x) = simple_integral f"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   574
    by (auto intro!: simple_integral_cong)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   575
  moreover have "\<And>X. X \<inter> space M \<inter> space M = X \<inter> space M" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   576
  ultimately show ?thesis by (simp add: simple_integral_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   577
next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   578
  assume "A \<noteq> space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   579
  then obtain x where x: "x \<in> space M" "x \<notin> A" using sets_into_space[OF assms(1)] by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   580
  have I: "(\<lambda>x. f x * indicator A x) ` space M = f ` A \<union> {0}" (is "?I ` _ = _")
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   581
  proof safe
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   582
    fix y assume "?I y \<notin> f ` A" hence "y \<notin> A" by auto thus "?I y = 0" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   583
  next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   584
    fix y assume "y \<in> A" thus "f y \<in> ?I ` space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   585
      using sets_into_space[OF assms(1)] by (auto intro!: image_eqI[of _ _ y])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   586
  next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   587
    show "0 \<in> ?I ` space M" using x by (auto intro!: image_eqI[of _ _ x])
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   588
  qed
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   589
  have *: "simple_integral (\<lambda>x. f x * indicator A x) =
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   590
    (\<Sum>x \<in> f ` space M \<union> {0}. x * \<mu> (f -` {x} \<inter> space M \<inter> A))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   591
    unfolding simple_integral_def I
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   592
  proof (rule setsum_mono_zero_cong_left)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   593
    show "finite (f ` space M \<union> {0})"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   594
      using assms(2) unfolding simple_function_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   595
    show "f ` A \<union> {0} \<subseteq> f`space M \<union> {0}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   596
      using sets_into_space[OF assms(1)] by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   597
    have "\<And>x. f x \<notin> f ` A \<Longrightarrow> f -` {f x} \<inter> space M \<inter> A = {}" by (auto simp: image_iff)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   598
    thus "\<forall>i\<in>f ` space M \<union> {0} - (f ` A \<union> {0}).
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   599
      i * \<mu> (f -` {i} \<inter> space M \<inter> A) = 0" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   600
  next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   601
    fix x assume "x \<in> f`A \<union> {0}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   602
    hence "x \<noteq> 0 \<Longrightarrow> ?I -` {x} \<inter> space M = f -` {x} \<inter> space M \<inter> A"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   603
      by (auto simp: indicator_def split: split_if_asm)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   604
    thus "x * \<mu> (?I -` {x} \<inter> space M) =
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   605
      x * \<mu> (f -` {x} \<inter> space M \<inter> A)" by (cases "x = 0") simp_all
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   606
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   607
  show ?thesis unfolding *
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   608
    using assms(2) unfolding simple_function_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   609
    by (auto intro!: setsum_mono_zero_cong_right)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   610
qed
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   611
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   612
lemma (in measure_space) simple_integral_indicator_only[simp]:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   613
  assumes "A \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   614
  shows "simple_integral (indicator A) = \<mu> A"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   615
proof cases
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   616
  assume "space M = {}" hence "A = {}" using sets_into_space[OF assms] by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   617
  thus ?thesis unfolding simple_integral_def using `space M = {}` by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   618
next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   619
  assume "space M \<noteq> {}" hence "(\<lambda>x. 1) ` space M = {1::pinfreal}" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   620
  thus ?thesis
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   621
    using simple_integral_indicator[OF assms simple_function_const[of 1]]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   622
    using sets_into_space[OF assms]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   623
    by (auto intro!: arg_cong[where f="\<mu>"])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   624
qed
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   625
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   626
lemma (in measure_space) simple_integral_null_set:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   627
  assumes "simple_function u" "N \<in> null_sets"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   628
  shows "simple_integral (\<lambda>x. u x * indicator N x) = 0"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   629
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   630
  have "simple_integral (\<lambda>x. u x * indicator N x) \<le>
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   631
    simple_integral (\<lambda>x. \<omega> * indicator N x)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   632
    using assms
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   633
    by (safe intro!: simple_integral_mono simple_function_mult simple_function_indicator simple_function_const) simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   634
  also have "... = 0" apply(subst simple_integral_mult)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   635
    using assms(2) by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   636
  finally show ?thesis by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   637
qed
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   638
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   639
lemma (in measure_space) simple_integral_cong':
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   640
  assumes f: "simple_function f" and g: "simple_function g"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   641
  and mea: "\<mu> {x\<in>space M. f x \<noteq> g x} = 0"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   642
  shows "simple_integral f = simple_integral g"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   643
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   644
  let ?h = "\<lambda>h. \<lambda>x. (h x * indicator {x\<in>space M. f x = g x} x
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   645
    + h x * indicator {x\<in>space M. f x \<noteq> g x} x
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   646
    + h x * indicator (-space M) x::pinfreal)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   647
  have *:"\<And>h. h = ?h h" unfolding indicator_def apply rule by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   648
  have mea_neq:"{x \<in> space M. f x \<noteq> g x} \<in> sets M" using f g by (auto simp: borel_measurable_simple_function)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   649
  then have mea_nullset: "{x \<in> space M. f x \<noteq> g x} \<in> null_sets" using mea by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   650
  have h1:"\<And>h::_=>pinfreal. simple_function h \<Longrightarrow>
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   651
    simple_function (\<lambda>x. h x * indicator {x\<in>space M. f x = g x} x)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   652
    apply(safe intro!: simple_function_add simple_function_mult simple_function_indicator)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   653
    using f g by (auto simp: borel_measurable_simple_function)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   654
  have h2:"\<And>h::_\<Rightarrow>pinfreal. simple_function h \<Longrightarrow>
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   655
    simple_function (\<lambda>x. h x * indicator {x\<in>space M. f x \<noteq> g x} x)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   656
    apply(safe intro!: simple_function_add simple_function_mult simple_function_indicator)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   657
    by(rule mea_neq)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   658
  have **:"\<And>a b c d e f. a = b \<Longrightarrow> c = d \<Longrightarrow> e = f \<Longrightarrow> a+c+e = b+d+f" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   659
  note *** = simple_integral_add[OF simple_function_add[OF h1 h2] simple_function_notspace]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   660
    simple_integral_add[OF h1 h2]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   661
  show ?thesis apply(subst *[of g]) apply(subst *[of f])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   662
    unfolding ***[OF f f] ***[OF g g]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   663
  proof(rule **) case goal1 show ?case apply(rule arg_cong[where f=simple_integral]) apply rule 
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   664
      unfolding indicator_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   665
  next note * = simple_integral_null_set[OF _ mea_nullset]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   666
    case goal2 show ?case unfolding *[OF f] *[OF g] ..
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   667
  next case goal3 show ?case apply(rule simple_integral_cong) by auto
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   668
  qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   669
qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   670
35692
f1315bbf1bc9 Moved theorems in Lebesgue to the right places
hoelzl
parents: 35582
diff changeset
   671
section "Continuous posititve integration"
f1315bbf1bc9 Moved theorems in Lebesgue to the right places
hoelzl
parents: 35582
diff changeset
   672
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   673
definition (in measure_space)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   674
  "positive_integral f =
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   675
    (SUP g : {g. simple_function g \<and> g \<le> f \<and> \<omega> \<notin> g`space M}. simple_integral g)"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   676
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   677
lemma (in measure_space) positive_integral_alt1:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   678
  "positive_integral f =
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   679
    (SUP g : {g. simple_function g \<and> (\<forall>x\<in>space M. g x \<le> f x \<and> g x \<noteq> \<omega>)}. simple_integral g)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   680
  unfolding positive_integral_def SUPR_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   681
proof (safe intro!: arg_cong[where f=Sup])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   682
  fix g let ?g = "\<lambda>x. if x \<in> space M then g x else f x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   683
  assume "simple_function g" "\<forall>x\<in>space M. g x \<le> f x \<and> g x \<noteq> \<omega>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   684
  hence "?g \<le> f" "simple_function ?g" "simple_integral g = simple_integral ?g"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   685
    "\<omega> \<notin> g`space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   686
    unfolding le_fun_def by (auto cong: simple_function_cong simple_integral_cong)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   687
  thus "simple_integral g \<in> simple_integral ` {g. simple_function g \<and> g \<le> f \<and> \<omega> \<notin> g`space M}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   688
    by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   689
next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   690
  fix g assume "simple_function g" "g \<le> f" "\<omega> \<notin> g`space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   691
  hence "simple_function g" "\<forall>x\<in>space M. g x \<le> f x \<and> g x \<noteq> \<omega>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   692
    by (auto simp add: le_fun_def image_iff)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   693
  thus "simple_integral g \<in> simple_integral ` {g. simple_function g \<and> (\<forall>x\<in>space M. g x \<le> f x \<and> g x \<noteq> \<omega>)}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   694
    by auto
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   695
qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   696
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   697
lemma (in measure_space) positive_integral_alt:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   698
  "positive_integral f =
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   699
    (SUP g : {g. simple_function g \<and> g \<le> f}. simple_integral g)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   700
  apply(rule order_class.antisym) unfolding positive_integral_def 
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   701
  apply(rule SUPR_subset) apply blast apply(rule SUPR_mono_lim)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   702
proof safe fix u assume u:"simple_function u" and uf:"u \<le> f"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   703
  let ?u = "\<lambda>n x. if u x = \<omega> then Real (real (n::nat)) else u x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   704
  have su:"\<And>n. simple_function (?u n)" using simple_function_compose1[OF u] .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   705
  show "\<exists>b. \<forall>n. b n \<in> {g. simple_function g \<and> g \<le> f \<and> \<omega> \<notin> g ` space M} \<and>
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   706
    (\<lambda>n. simple_integral (b n)) ----> simple_integral u"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   707
    apply(rule_tac x="?u" in exI, safe) apply(rule su)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   708
  proof- fix n::nat have "?u n \<le> u" unfolding le_fun_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   709
    also note uf finally show "?u n \<le> f" .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   710
    let ?s = "{x \<in> space M. u x = \<omega>}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   711
    show "(\<lambda>n. simple_integral (?u n)) ----> simple_integral u"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   712
    proof(cases "\<mu> ?s = 0")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   713
      case True have *:"\<And>n. {x\<in>space M. ?u n x \<noteq> u x} = {x\<in>space M. u x = \<omega>}" by auto 
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   714
      have *:"\<And>n. simple_integral (?u n) = simple_integral u"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   715
        apply(rule simple_integral_cong'[OF su u]) unfolding * True ..
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   716
      show ?thesis unfolding * by auto 
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   717
    next case False note m0=this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   718
      have s:"{x \<in> space M. u x = \<omega>} \<in> sets M" using u  by (auto simp: borel_measurable_simple_function)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   719
      have "\<omega> = simple_integral (\<lambda>x. \<omega> * indicator {x\<in>space M. u x = \<omega>} x)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   720
        apply(subst simple_integral_mult) using s
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   721
        unfolding simple_integral_indicator_only[OF s] using False by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   722
      also have "... \<le> simple_integral u"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   723
        apply (rule simple_integral_mono)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   724
        apply (rule simple_function_mult)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   725
        apply (rule simple_function_const)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   726
        apply(rule ) prefer 3 apply(subst indicator_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   727
        using s u by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   728
      finally have *:"simple_integral u = \<omega>" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   729
      show ?thesis unfolding * Lim_omega_pos
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   730
      proof safe case goal1
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   731
        from real_arch_simple[of "B / real (\<mu> ?s)"] guess N0 .. note N=this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   732
        def N \<equiv> "Suc N0" have N:"real N \<ge> B / real (\<mu> ?s)" "N > 0"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   733
          unfolding N_def using N by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   734
        show ?case apply-apply(rule_tac x=N in exI,safe) 
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   735
        proof- case goal1
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   736
          have "Real B \<le> Real (real N) * \<mu> ?s"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   737
          proof(cases "\<mu> ?s = \<omega>")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   738
            case True thus ?thesis using `B>0` N by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   739
          next case False
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   740
            have *:"B \<le> real N * real (\<mu> ?s)" 
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   741
              using N(1) apply-apply(subst (asm) pos_divide_le_eq)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   742
              apply rule using m0 False by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   743
            show ?thesis apply(subst Real_real'[THEN sym,OF False])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   744
              apply(subst pinfreal_times,subst if_P) defer
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   745
              apply(subst pinfreal_less_eq,subst if_P) defer
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   746
              using * N `B>0` by(auto intro: mult_nonneg_nonneg)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   747
          qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   748
          also have "... \<le> Real (real n) * \<mu> ?s"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   749
            apply(rule mult_right_mono) using goal1 by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   750
          also have "... = simple_integral (\<lambda>x. Real (real n) * indicator ?s x)" 
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   751
            apply(subst simple_integral_mult) apply(rule simple_function_indicator[OF s])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   752
            unfolding simple_integral_indicator_only[OF s] ..
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   753
          also have "... \<le> simple_integral (\<lambda>x. if u x = \<omega> then Real (real n) else u x)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   754
            apply(rule simple_integral_mono) apply(rule simple_function_mult)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   755
            apply(rule simple_function_const)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   756
            apply(rule simple_function_indicator) apply(rule s su)+ by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   757
          finally show ?case .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   758
        qed qed qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   759
    fix x assume x:"\<omega> = (if u x = \<omega> then Real (real n) else u x)" "x \<in> space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   760
    hence "u x = \<omega>" apply-apply(rule ccontr) by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   761
    hence "\<omega> = Real (real n)" using x by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   762
    thus False by auto
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   763
  qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   764
qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   765
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   766
lemma (in measure_space) positive_integral_cong:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   767
  assumes "\<And>x. x \<in> space M \<Longrightarrow> f x = g x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   768
  shows "positive_integral f = positive_integral g"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   769
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   770
  have "\<And>h. (\<forall>x\<in>space M. h x \<le> f x \<and> h x \<noteq> \<omega>) = (\<forall>x\<in>space M. h x \<le> g x \<and> h x \<noteq> \<omega>)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   771
    using assms by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   772
  thus ?thesis unfolding positive_integral_alt1 by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   773
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   774
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   775
lemma (in measure_space) positive_integral_eq_simple_integral:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   776
  assumes "simple_function f"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   777
  shows "positive_integral f = simple_integral f"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   778
  unfolding positive_integral_alt
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   779
proof (safe intro!: pinfreal_SUPI)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   780
  fix g assume "simple_function g" "g \<le> f"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   781
  with assms show "simple_integral g \<le> simple_integral f"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   782
    by (auto intro!: simple_integral_mono simp: le_fun_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   783
next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   784
  fix y assume "\<forall>x. x\<in>{g. simple_function g \<and> g \<le> f} \<longrightarrow> simple_integral x \<le> y"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   785
  with assms show "simple_integral f \<le> y" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   786
qed
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   787
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   788
lemma (in measure_space) positive_integral_mono:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   789
  assumes mono: "\<And>x. x \<in> space M \<Longrightarrow> u x \<le> v x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   790
  shows "positive_integral u \<le> positive_integral v"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   791
  unfolding positive_integral_alt1
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   792
proof (safe intro!: SUPR_mono)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   793
  fix a assume a: "simple_function a" and "\<forall>x\<in>space M. a x \<le> u x \<and> a x \<noteq> \<omega>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   794
  with mono have "\<forall>x\<in>space M. a x \<le> v x \<and> a x \<noteq> \<omega>" by fastsimp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   795
  with a show "\<exists>b\<in>{g. simple_function g \<and> (\<forall>x\<in>space M. g x \<le> v x \<and> g x \<noteq> \<omega>)}. simple_integral a \<le> simple_integral b"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   796
    by (auto intro!: bexI[of _ a])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   797
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   798
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   799
lemma (in measure_space) positive_integral_SUP_approx:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   800
  assumes "f \<up> s"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   801
  and f: "\<And>i. f i \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   802
  and "simple_function u"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   803
  and le: "u \<le> s" and real: "\<omega> \<notin> u`space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   804
  shows "simple_integral u \<le> (SUP i. positive_integral (f i))" (is "_ \<le> ?S")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   805
proof (rule pinfreal_le_mult_one_interval)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   806
  fix a :: pinfreal assume "0 < a" "a < 1"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   807
  hence "a \<noteq> 0" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   808
  let "?B i" = "{x \<in> space M. a * u x \<le> f i x}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   809
  have B: "\<And>i. ?B i \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   810
    using f `simple_function u` by (auto simp: borel_measurable_simple_function)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   811
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   812
  let "?uB i x" = "u x * indicator (?B i) x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   813
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   814
  { fix i have "?B i \<subseteq> ?B (Suc i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   815
    proof safe
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   816
      fix i x assume "a * u x \<le> f i x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   817
      also have "\<dots> \<le> f (Suc i) x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   818
        using `f \<up> s` unfolding isoton_def le_fun_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   819
      finally show "a * u x \<le> f (Suc i) x" .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   820
    qed }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   821
  note B_mono = this
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   822
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   823
  have u: "\<And>x. x \<in> space M \<Longrightarrow> u -` {u x} \<inter> space M \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   824
    using `simple_function u` by (auto simp add: simple_function_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   825
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   826
  { fix i
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   827
    have "(\<lambda>n. (u -` {i} \<inter> space M) \<inter> ?B n) \<up> (u -` {i} \<inter> space M)" using B_mono unfolding isoton_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   828
    proof safe
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   829
      fix x assume "x \<in> space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   830
      show "x \<in> (\<Union>i. (u -` {u x} \<inter> space M) \<inter> ?B i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   831
      proof cases
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   832
        assume "u x = 0" thus ?thesis using `x \<in> space M` by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   833
      next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   834
        assume "u x \<noteq> 0"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   835
        with `a < 1` real `x \<in> space M`
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   836
        have "a * u x < 1 * u x" by (rule_tac pinfreal_mult_strict_right_mono) (auto simp: image_iff)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   837
        also have "\<dots> \<le> (SUP i. f i x)" using le `f \<up> s`
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   838
          unfolding isoton_fun_expand by (auto simp: isoton_def le_fun_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   839
        finally obtain i where "a * u x < f i x" unfolding SUPR_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   840
          by (auto simp add: less_Sup_iff)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   841
        hence "a * u x \<le> f i x" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   842
        thus ?thesis using `x \<in> space M` by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   843
      qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   844
    qed auto }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   845
  note measure_conv = measure_up[OF u Int[OF u B] this]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   846
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   847
  have "simple_integral u = (SUP i. simple_integral (?uB i))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   848
    unfolding simple_integral_indicator[OF B `simple_function u`]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   849
  proof (subst SUPR_pinfreal_setsum, safe)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   850
    fix x n assume "x \<in> space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   851
    have "\<mu> (u -` {u x} \<inter> space M \<inter> {x \<in> space M. a * u x \<le> f n x})
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   852
      \<le> \<mu> (u -` {u x} \<inter> space M \<inter> {x \<in> space M. a * u x \<le> f (Suc n) x})"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   853
      using B_mono Int[OF u[OF `x \<in> space M`] B] by (auto intro!: measure_mono)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   854
    thus "u x * \<mu> (u -` {u x} \<inter> space M \<inter> ?B n)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   855
            \<le> u x * \<mu> (u -` {u x} \<inter> space M \<inter> ?B (Suc n))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   856
      by (auto intro: mult_left_mono)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   857
  next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   858
    show "simple_integral u =
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   859
      (\<Sum>i\<in>u ` space M. SUP n. i * \<mu> (u -` {i} \<inter> space M \<inter> ?B n))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   860
      using measure_conv unfolding simple_integral_def isoton_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   861
      by (auto intro!: setsum_cong simp: pinfreal_SUP_cmult)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   862
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   863
  moreover
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   864
  have "a * (SUP i. simple_integral (?uB i)) \<le> ?S"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   865
    unfolding pinfreal_SUP_cmult[symmetric]
38705
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 38656
diff changeset
   866
  proof (safe intro!: SUP_mono bexI)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   867
    fix i
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   868
    have "a * simple_integral (?uB i) = simple_integral (\<lambda>x. a * ?uB i x)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   869
      using B `simple_function u`
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   870
      by (subst simple_integral_mult[symmetric]) (auto simp: field_simps)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   871
    also have "\<dots> \<le> positive_integral (f i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   872
    proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   873
      have "\<And>x. a * ?uB i x \<le> f i x" unfolding indicator_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   874
      hence *: "simple_function (\<lambda>x. a * ?uB i x)" using B assms(3)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   875
        by (auto intro!: simple_integral_mono)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   876
      show ?thesis unfolding positive_integral_eq_simple_integral[OF *, symmetric]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   877
        by (auto intro!: positive_integral_mono simp: indicator_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   878
    qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   879
    finally show "a * simple_integral (?uB i) \<le> positive_integral (f i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   880
      by auto
38705
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 38656
diff changeset
   881
  qed simp
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   882
  ultimately show "a * simple_integral u \<le> ?S" by simp
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   883
qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   884
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   885
text {* Beppo-Levi monotone convergence theorem *}
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   886
lemma (in measure_space) positive_integral_isoton:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   887
  assumes "f \<up> u" "\<And>i. f i \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   888
  shows "(\<lambda>i. positive_integral (f i)) \<up> positive_integral u"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   889
  unfolding isoton_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   890
proof safe
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   891
  fix i show "positive_integral (f i) \<le> positive_integral (f (Suc i))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   892
    apply (rule positive_integral_mono)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   893
    using `f \<up> u` unfolding isoton_def le_fun_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   894
next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   895
  have "u \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   896
    using borel_measurable_SUP[of UNIV f] assms by (auto simp: isoton_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   897
  have u: "u = (SUP i. f i)" using `f \<up> u` unfolding isoton_def by auto
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   898
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   899
  show "(SUP i. positive_integral (f i)) = positive_integral u"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   900
  proof (rule antisym)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   901
    from `f \<up> u`[THEN isoton_Sup, unfolded le_fun_def]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   902
    show "(SUP j. positive_integral (f j)) \<le> positive_integral u"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   903
      by (auto intro!: SUP_leI positive_integral_mono)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   904
  next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   905
    show "positive_integral u \<le> (SUP i. positive_integral (f i))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   906
      unfolding positive_integral_def[of u]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   907
      by (auto intro!: SUP_leI positive_integral_SUP_approx[OF assms])
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   908
  qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   909
qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   910
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   911
lemma (in measure_space) SUP_simple_integral_sequences:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   912
  assumes f: "f \<up> u" "\<And>i. simple_function (f i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   913
  and g: "g \<up> u" "\<And>i. simple_function (g i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   914
  shows "(SUP i. simple_integral (f i)) = (SUP i. simple_integral (g i))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   915
    (is "SUPR _ ?F = SUPR _ ?G")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   916
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   917
  have "(SUP i. ?F i) = (SUP i. positive_integral (f i))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   918
    using assms by (simp add: positive_integral_eq_simple_integral)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   919
  also have "\<dots> = positive_integral u"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   920
    using positive_integral_isoton[OF f(1) borel_measurable_simple_function[OF f(2)]]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   921
    unfolding isoton_def by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   922
  also have "\<dots> = (SUP i. positive_integral (g i))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   923
    using positive_integral_isoton[OF g(1) borel_measurable_simple_function[OF g(2)]]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   924
    unfolding isoton_def by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   925
  also have "\<dots> = (SUP i. ?G i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   926
    using assms by (simp add: positive_integral_eq_simple_integral)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   927
  finally show ?thesis .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   928
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   929
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   930
lemma (in measure_space) positive_integral_const[simp]:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   931
  "positive_integral (\<lambda>x. c) = c * \<mu> (space M)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   932
  by (subst positive_integral_eq_simple_integral) auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   933
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   934
lemma (in measure_space) positive_integral_isoton_simple:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   935
  assumes "f \<up> u" and e: "\<And>i. simple_function (f i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   936
  shows "(\<lambda>i. simple_integral (f i)) \<up> positive_integral u"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   937
  using positive_integral_isoton[OF `f \<up> u` e(1)[THEN borel_measurable_simple_function]]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   938
  unfolding positive_integral_eq_simple_integral[OF e] .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   939
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   940
lemma (in measure_space) positive_integral_linear:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   941
  assumes f: "f \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   942
  and g: "g \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   943
  shows "positive_integral (\<lambda>x. a * f x + g x) =
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   944
      a * positive_integral f + positive_integral g"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   945
    (is "positive_integral ?L = _")
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   946
proof -
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   947
  from borel_measurable_implies_simple_function_sequence'[OF f] guess u .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   948
  note u = this positive_integral_isoton_simple[OF this(1-2)]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   949
  from borel_measurable_implies_simple_function_sequence'[OF g] guess v .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   950
  note v = this positive_integral_isoton_simple[OF this(1-2)]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   951
  let "?L' i x" = "a * u i x + v i x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   952
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   953
  have "?L \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   954
    using assms by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   955
  from borel_measurable_implies_simple_function_sequence'[OF this] guess l .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   956
  note positive_integral_isoton_simple[OF this(1-2)] and l = this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   957
  moreover have
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   958
      "(SUP i. simple_integral (l i)) = (SUP i. simple_integral (?L' i))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   959
  proof (rule SUP_simple_integral_sequences[OF l(1-2)])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   960
    show "?L' \<up> ?L" "\<And>i. simple_function (?L' i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   961
      using u v by (auto simp: isoton_fun_expand isoton_add isoton_cmult_right)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   962
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   963
  moreover from u v have L'_isoton:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   964
      "(\<lambda>i. simple_integral (?L' i)) \<up> a * positive_integral f + positive_integral g"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   965
    by (simp add: isoton_add isoton_cmult_right)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   966
  ultimately show ?thesis by (simp add: isoton_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   967
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   968
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   969
lemma (in measure_space) positive_integral_cmult:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   970
  assumes "f \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   971
  shows "positive_integral (\<lambda>x. c * f x) = c * positive_integral f"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   972
  using positive_integral_linear[OF assms, of "\<lambda>x. 0" c] by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   973
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   974
lemma (in measure_space) positive_integral_indicator[simp]:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   975
  "A \<in> sets M \<Longrightarrow> positive_integral (\<lambda>x. indicator A x) = \<mu> A"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   976
by (subst positive_integral_eq_simple_integral) (auto simp: simple_function_indicator simple_integral_indicator)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   977
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   978
lemma (in measure_space) positive_integral_cmult_indicator:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   979
  "A \<in> sets M \<Longrightarrow> positive_integral (\<lambda>x. c * indicator A x) = c * \<mu> A"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   980
by (subst positive_integral_eq_simple_integral) (auto simp: simple_function_indicator simple_integral_indicator)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   981
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   982
lemma (in measure_space) positive_integral_add:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   983
  assumes "f \<in> borel_measurable M" "g \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   984
  shows "positive_integral (\<lambda>x. f x + g x) = positive_integral f + positive_integral g"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   985
  using positive_integral_linear[OF assms, of 1] by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   986
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   987
lemma (in measure_space) positive_integral_setsum:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   988
  assumes "\<And>i. i\<in>P \<Longrightarrow> f i \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   989
  shows "positive_integral (\<lambda>x. \<Sum>i\<in>P. f i x) = (\<Sum>i\<in>P. positive_integral (f i))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   990
proof cases
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   991
  assume "finite P"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   992
  from this assms show ?thesis
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   993
  proof induct
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   994
    case (insert i P)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   995
    have "f i \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   996
      "(\<lambda>x. \<Sum>i\<in>P. f i x) \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   997
      using insert by (auto intro!: borel_measurable_pinfreal_setsum)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   998
    from positive_integral_add[OF this]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   999
    show ?case using insert by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1000
  qed simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1001
qed simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1002
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1003
lemma (in measure_space) positive_integral_diff:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1004
  assumes f: "f \<in> borel_measurable M" and g: "g \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1005
  and fin: "positive_integral g \<noteq> \<omega>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1006
  and mono: "\<And>x. x \<in> space M \<Longrightarrow> g x \<le> f x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1007
  shows "positive_integral (\<lambda>x. f x - g x) = positive_integral f - positive_integral g"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1008
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1009
  have borel: "(\<lambda>x. f x - g x) \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1010
    using f g by (rule borel_measurable_pinfreal_diff)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1011
  have "positive_integral (\<lambda>x. f x - g x) + positive_integral g =
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1012
    positive_integral f"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1013
    unfolding positive_integral_add[OF borel g, symmetric]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1014
  proof (rule positive_integral_cong)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1015
    fix x assume "x \<in> space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1016
    from mono[OF this] show "f x - g x + g x = f x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1017
      by (cases "f x", cases "g x", simp, simp, cases "g x", auto)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1018
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1019
  with mono show ?thesis
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1020
    by (subst minus_pinfreal_eq2[OF _ fin]) (auto intro!: positive_integral_mono)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1021
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1022
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1023
lemma (in measure_space) positive_integral_psuminf:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1024
  assumes "\<And>i. f i \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1025
  shows "positive_integral (\<lambda>x. \<Sum>\<^isub>\<infinity> i. f i x) = (\<Sum>\<^isub>\<infinity> i. positive_integral (f i))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1026
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1027
  have "(\<lambda>i. positive_integral (\<lambda>x. \<Sum>i<i. f i x)) \<up> positive_integral (\<lambda>x. \<Sum>\<^isub>\<infinity>i. f i x)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1028
    by (rule positive_integral_isoton)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1029
       (auto intro!: borel_measurable_pinfreal_setsum assms positive_integral_mono
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1030
                     arg_cong[where f=Sup]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1031
             simp: isoton_def le_fun_def psuminf_def expand_fun_eq SUPR_def Sup_fun_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1032
  thus ?thesis
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1033
    by (auto simp: isoton_def psuminf_def positive_integral_setsum[OF assms])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1034
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1035
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1036
text {* Fatou's lemma: convergence theorem on limes inferior *}
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1037
lemma (in measure_space) positive_integral_lim_INF:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1038
  fixes u :: "nat \<Rightarrow> 'a \<Rightarrow> pinfreal"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1039
  assumes "\<And>i. u i \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1040
  shows "positive_integral (SUP n. INF m. u (m + n)) \<le>
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1041
    (SUP n. INF m. positive_integral (u (m + n)))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1042
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1043
  have "(SUP n. INF m. u (m + n)) \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1044
    by (auto intro!: borel_measurable_SUP borel_measurable_INF assms)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1045
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1046
  have "(\<lambda>n. INF m. u (m + n)) \<up> (SUP n. INF m. u (m + n))"
38705
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 38656
diff changeset
  1047
  proof (unfold isoton_def, safe intro!: INF_mono bexI)
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 38656
diff changeset
  1048
    fix i m show "u (Suc m + i) \<le> u (m + Suc i)" by simp
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 38656
diff changeset
  1049
  qed simp
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1050
  from positive_integral_isoton[OF this] assms
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1051
  have "positive_integral (SUP n. INF m. u (m + n)) =
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1052
    (SUP n. positive_integral (INF m. u (m + n)))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1053
    unfolding isoton_def by (simp add: borel_measurable_INF)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1054
  also have "\<dots> \<le> (SUP n. INF m. positive_integral (u (m + n)))"
38705
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 38656
diff changeset
  1055
    apply (rule SUP_mono)
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 38656
diff changeset
  1056
    apply (rule_tac x=n in bexI)
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 38656
diff changeset
  1057
    by (auto intro!: positive_integral_mono INFI_bound INF_leI exI simp: INFI_fun_expand)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1058
  finally show ?thesis .
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1059
qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1060
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1061
lemma (in measure_space) measure_space_density:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1062
  assumes borel: "u \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1063
  shows "measure_space M (\<lambda>A. positive_integral (\<lambda>x. u x * indicator A x))" (is "measure_space M ?v")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1064
proof
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1065
  show "?v {} = 0" by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1066
  show "countably_additive M ?v"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1067
    unfolding countably_additive_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1068
  proof safe
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1069
    fix A :: "nat \<Rightarrow> 'a set"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1070
    assume "range A \<subseteq> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1071
    hence "\<And>i. (\<lambda>x. u x * indicator (A i) x) \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1072
      using borel by (auto intro: borel_measurable_indicator)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1073
    moreover assume "disjoint_family A"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1074
    note psuminf_indicator[OF this]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1075
    ultimately show "(\<Sum>\<^isub>\<infinity>n. ?v (A n)) = ?v (\<Union>x. A x)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1076
      by (simp add: positive_integral_psuminf[symmetric])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1077
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1078
qed
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1079
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1080
lemma (in measure_space) positive_integral_null_set:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1081
  assumes borel: "u \<in> borel_measurable M" and "N \<in> null_sets"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1082
  shows "positive_integral (\<lambda>x. u x * indicator N x) = 0" (is "?I = 0")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1083
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1084
  have "N \<in> sets M" using `N \<in> null_sets` by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1085
  have "(\<lambda>i x. min (of_nat i) (u x) * indicator N x) \<up> (\<lambda>x. u x * indicator N x)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1086
    unfolding isoton_fun_expand
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1087
  proof (safe intro!: isoton_cmult_left, unfold isoton_def, safe)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1088
    fix j i show "min (of_nat j) (u i) \<le> min (of_nat (Suc j)) (u i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1089
      by (rule min_max.inf_mono) auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1090
  next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1091
    fix i show "(SUP j. min (of_nat j) (u i)) = u i"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1092
    proof (cases "u i")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1093
      case infinite
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1094
      moreover hence "\<And>j. min (of_nat j) (u i) = of_nat j"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1095
        by (auto simp: min_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1096
      ultimately show ?thesis by (simp add: Sup_\<omega>)
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1097
    next
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1098
      case (preal r)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1099
      obtain j where "r \<le> of_nat j" using ex_le_of_nat ..
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1100
      hence "u i \<le> of_nat j" using preal by (auto simp: real_of_nat_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1101
      show ?thesis
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1102
      proof (rule pinfreal_SUPI)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1103
        fix y assume "\<And>j. j \<in> UNIV \<Longrightarrow> min (of_nat j) (u i) \<le> y"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1104
        note this[of j]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1105
        moreover have "min (of_nat j) (u i) = u i"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1106
          using `u i \<le> of_nat j` by (auto simp: min_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1107
        ultimately show "u i \<le> y" by simp
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1108
      qed simp
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1109
    qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1110
  qed
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1111
  from positive_integral_isoton[OF this]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1112
  have "?I = (SUP i. positive_integral (\<lambda>x. min (of_nat i) (u x) * indicator N x))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1113
    unfolding isoton_def using borel `N \<in> sets M` by (simp add: borel_measurable_indicator)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1114
  also have "\<dots> \<le> (SUP i. positive_integral (\<lambda>x. of_nat i * indicator N x))"
38705
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 38656
diff changeset
  1115
  proof (rule SUP_mono, rule bexI, rule positive_integral_mono)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1116
    fix x i show "min (of_nat i) (u x) * indicator N x \<le> of_nat i * indicator N x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1117
      by (cases "x \<in> N") auto
38705
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 38656
diff changeset
  1118
  qed simp
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1119
  also have "\<dots> = 0"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1120
    using `N \<in> null_sets` by (simp add: positive_integral_cmult_indicator)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1121
  finally show ?thesis by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1122
qed
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1123
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1124
lemma (in measure_space) positive_integral_Markov_inequality:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1125
  assumes borel: "u \<in> borel_measurable M" and "A \<in> sets M" and c: "c \<noteq> \<omega>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1126
  shows "\<mu> ({x\<in>space M. 1 \<le> c * u x} \<inter> A) \<le> c * positive_integral (\<lambda>x. u x * indicator A x)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1127
    (is "\<mu> ?A \<le> _ * ?PI")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1128
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1129
  have "?A \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1130
    using `A \<in> sets M` borel by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1131
  hence "\<mu> ?A = positive_integral (\<lambda>x. indicator ?A x)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1132
    using positive_integral_indicator by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1133
  also have "\<dots> \<le> positive_integral (\<lambda>x. c * (u x * indicator A x))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1134
  proof (rule positive_integral_mono)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1135
    fix x assume "x \<in> space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1136
    show "indicator ?A x \<le> c * (u x * indicator A x)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1137
      by (cases "x \<in> ?A") auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1138
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1139
  also have "\<dots> = c * positive_integral (\<lambda>x. u x * indicator A x)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1140
    using assms
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1141
    by (auto intro!: positive_integral_cmult borel_measurable_indicator)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1142
  finally show ?thesis .
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1143
qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1144
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1145
lemma (in measure_space) positive_integral_0_iff:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1146
  assumes borel: "u \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1147
  shows "positive_integral u = 0 \<longleftrightarrow> \<mu> {x\<in>space M. u x \<noteq> 0} = 0"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1148
    (is "_ \<longleftrightarrow> \<mu> ?A = 0")
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1149
proof -
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1150
  have A: "?A \<in> sets M" using borel by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1151
  have u: "positive_integral (\<lambda>x. u x * indicator ?A x) = positive_integral u"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1152
    by (auto intro!: positive_integral_cong simp: indicator_def)
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1153
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1154
  show ?thesis
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1155
  proof
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1156
    assume "\<mu> ?A = 0"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1157
    hence "?A \<in> null_sets" using `?A \<in> sets M` by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1158
    from positive_integral_null_set[OF borel this]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1159
    have "0 = positive_integral (\<lambda>x. u x * indicator ?A x)" by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1160
    thus "positive_integral u = 0" unfolding u by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1161
  next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1162
    assume *: "positive_integral u = 0"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1163
    let "?M n" = "{x \<in> space M. 1 \<le> of_nat n * u x}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1164
    have "0 = (SUP n. \<mu> (?M n \<inter> ?A))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1165
    proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1166
      { fix n
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1167
        from positive_integral_Markov_inequality[OF borel `?A \<in> sets M`, of "of_nat n"]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1168
        have "\<mu> (?M n \<inter> ?A) = 0" unfolding * u by simp }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1169
      thus ?thesis by simp
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1170
    qed
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1171
    also have "\<dots> = \<mu> (\<Union>n. ?M n \<inter> ?A)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1172
    proof (safe intro!: continuity_from_below)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1173
      fix n show "?M n \<inter> ?A \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1174
        using borel by (auto intro!: Int)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1175
    next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1176
      fix n x assume "1 \<le> of_nat n * u x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1177
      also have "\<dots> \<le> of_nat (Suc n) * u x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1178
        by (subst (1 2) mult_commute) (auto intro!: pinfreal_mult_cancel)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1179
      finally show "1 \<le> of_nat (Suc n) * u x" .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1180
    qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1181
    also have "\<dots> = \<mu> ?A"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1182
    proof (safe intro!: arg_cong[where f="\<mu>"])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1183
      fix x assume "u x \<noteq> 0" and [simp, intro]: "x \<in> space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1184
      show "x \<in> (\<Union>n. ?M n \<inter> ?A)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1185
      proof (cases "u x")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1186
        case (preal r)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1187
        obtain j where "1 / r \<le> of_nat j" using ex_le_of_nat ..
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1188
        hence "1 / r * r \<le> of_nat j * r" using preal unfolding mult_le_cancel_right by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1189
        hence "1 \<le> of_nat j * r" using preal `u x \<noteq> 0` by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1190
        thus ?thesis using `u x \<noteq> 0` preal by (auto simp: real_of_nat_def[symmetric])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1191
      qed auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1192
    qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1193
    finally show "\<mu> ?A = 0" by simp
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1194
  qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1195
qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1196
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1197
lemma (in measure_space) positive_integral_cong_on_null_sets:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1198
  assumes f: "f \<in> borel_measurable M" and g: "g \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1199
  and measure: "\<mu> {x\<in>space M. f x \<noteq> g x} = 0"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1200
  shows "positive_integral f = positive_integral g"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1201
proof -
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1202
  let ?N = "{x\<in>space M. f x \<noteq> g x}" and ?E = "{x\<in>space M. f x = g x}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1203
  let "?A h x" = "h x * indicator ?E x :: pinfreal"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1204
  let "?B h x" = "h x * indicator ?N x :: pinfreal"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1205
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1206
  have A: "positive_integral (?A f) = positive_integral (?A g)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1207
    by (auto intro!: positive_integral_cong simp: indicator_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1208
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1209
  have [intro]: "?N \<in> sets M" "?E \<in> sets M" using f g by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1210
  hence "?N \<in> null_sets" using measure by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1211
  hence B: "positive_integral (?B f) = positive_integral (?B g)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1212
    using f g by (simp add: positive_integral_null_set)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1213
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1214
  have "positive_integral f = positive_integral (\<lambda>x. ?A f x + ?B f x)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1215
    by (auto intro!: positive_integral_cong simp: indicator_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1216
  also have "\<dots> = positive_integral (?A f) + positive_integral (?B f)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1217
    using f g by (auto intro!: positive_integral_add borel_measurable_indicator)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1218
  also have "\<dots> = positive_integral (\<lambda>x. ?A g x + ?B g x)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1219
    unfolding A B using f g by (auto intro!: positive_integral_add[symmetric] borel_measurable_indicator)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1220
  also have "\<dots> = positive_integral g"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1221
    by (auto intro!: positive_integral_cong simp: indicator_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1222
  finally show ?thesis by simp
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1223
qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1224
35692
f1315bbf1bc9 Moved theorems in Lebesgue to the right places
hoelzl
parents: 35582
diff changeset
  1225
section "Lebesgue Integral"
f1315bbf1bc9 Moved theorems in Lebesgue to the right places
hoelzl
parents: 35582
diff changeset
  1226
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1227
definition (in measure_space) integrable where
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1228
  "integrable f \<longleftrightarrow> f \<in> borel_measurable M \<and>
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1229
    positive_integral (\<lambda>x. Real (f x)) \<noteq> \<omega> \<and>
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1230
    positive_integral (\<lambda>x. Real (- f x)) \<noteq> \<omega>"
35692
f1315bbf1bc9 Moved theorems in Lebesgue to the right places
hoelzl
parents: 35582
diff changeset
  1231
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1232
lemma (in measure_space) integrableD[dest]:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1233
  assumes "integrable f"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1234
  shows "f \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1235
  "positive_integral (\<lambda>x. Real (f x)) \<noteq> \<omega>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1236
  "positive_integral (\<lambda>x. Real (- f x)) \<noteq> \<omega>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1237
  using assms unfolding integrable_def by auto
35692
f1315bbf1bc9 Moved theorems in Lebesgue to the right places
hoelzl
parents: 35582
diff changeset
  1238
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1239
definition (in measure_space) integral where
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1240
  "integral f =
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1241
    real (positive_integral (\<lambda>x. Real (f x))) -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1242
    real (positive_integral (\<lambda>x. Real (- f x)))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1243
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1244
lemma (in measure_space) integral_cong:
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1245
  assumes cong: "\<And>x. x \<in> space M \<Longrightarrow> f x = g x"
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1246
  shows "integral f = integral g"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1247
  using assms by (simp cong: positive_integral_cong add: integral_def)
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1248
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1249
lemma (in measure_space) integrable_cong:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1250
  "(\<And>x. x \<in> space M \<Longrightarrow> f x = g x) \<Longrightarrow> integrable f \<longleftrightarrow> integrable g"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1251
  by (simp cong: positive_integral_cong measurable_cong add: integrable_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1252
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1253
lemma (in measure_space) integral_eq_positive_integral:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1254
  assumes "\<And>x. 0 \<le> f x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1255
  shows "integral f = real (positive_integral (\<lambda>x. Real (f x)))"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1256
proof -
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1257
  have "\<And>x. Real (- f x) = 0" using assms by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1258
  thus ?thesis by (simp del: Real_eq_0 add: integral_def)
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1259
qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1260
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1261
lemma (in measure_space) integral_minus[intro, simp]:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1262
  assumes "integrable f"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1263
  shows "integrable (\<lambda>x. - f x)" "integral (\<lambda>x. - f x) = - integral f"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1264
  using assms by (auto simp: integrable_def integral_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1265
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1266
lemma (in measure_space) integral_of_positive_diff:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1267
  assumes integrable: "integrable u" "integrable v"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1268
  and f_def: "\<And>x. f x = u x - v x" and pos: "\<And>x. 0 \<le> u x" "\<And>x. 0 \<le> v x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1269
  shows "integrable f" and "integral f = integral u - integral v"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1270
proof -
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1271
  let ?PI = positive_integral
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1272
  let "?f x" = "Real (f x)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1273
  let "?mf x" = "Real (- f x)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1274
  let "?u x" = "Real (u x)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1275
  let "?v x" = "Real (v x)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1276
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1277
  from borel_measurable_diff[of u v] integrable
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1278
  have f_borel: "?f \<in> borel_measurable M" and
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1279
    mf_borel: "?mf \<in> borel_measurable M" and
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1280
    v_borel: "?v \<in> borel_measurable M" and
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1281
    u_borel: "?u \<in> borel_measurable M" and
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1282
    "f \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1283
    by (auto simp: f_def[symmetric] integrable_def)
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1284
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1285
  have "?PI (\<lambda>x. Real (u x - v x)) \<le> ?PI ?u"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1286
    using pos by (auto intro!: positive_integral_mono)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1287
  moreover have "?PI (\<lambda>x. Real (v x - u x)) \<le> ?PI ?v"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1288
    using pos by (auto intro!: positive_integral_mono)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1289
  ultimately show f: "integrable f"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1290
    using `integrable u` `integrable v` `f \<in> borel_measurable M`
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1291
    by (auto simp: integrable_def f_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1292
  hence mf: "integrable (\<lambda>x. - f x)" ..
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1293
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1294
  have *: "\<And>x. Real (- v x) = 0" "\<And>x. Real (- u x) = 0"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1295
    using pos by auto
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1296
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1297
  have "\<And>x. ?u x + ?mf x = ?v x + ?f x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1298
    unfolding f_def using pos by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1299
  hence "?PI (\<lambda>x. ?u x + ?mf x) = ?PI (\<lambda>x. ?v x + ?f x)" by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1300
  hence "real (?PI ?u + ?PI ?mf) = real (?PI ?v + ?PI ?f)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1301
    using positive_integral_add[OF u_borel mf_borel]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1302
    using positive_integral_add[OF v_borel f_borel]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1303
    by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1304
  then show "integral f = integral u - integral v"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1305
    using f mf `integrable u` `integrable v`
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1306
    unfolding integral_def integrable_def *
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1307
    by (cases "?PI ?f", cases "?PI ?mf", cases "?PI ?v", cases "?PI ?u")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1308
       (auto simp add: field_simps)
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1309
qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1310
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1311
lemma (in measure_space) integral_linear:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1312
  assumes "integrable f" "integrable g" and "0 \<le> a"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1313
  shows "integrable (\<lambda>t. a * f t + g t)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1314
  and "integral (\<lambda>t. a * f t + g t) = a * integral f + integral g"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1315
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1316
  let ?PI = positive_integral
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1317
  let "?f x" = "Real (f x)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1318
  let "?g x" = "Real (g x)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1319
  let "?mf x" = "Real (- f x)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1320
  let "?mg x" = "Real (- g x)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1321
  let "?p t" = "max 0 (a * f t) + max 0 (g t)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1322
  let "?n t" = "max 0 (- (a * f t)) + max 0 (- g t)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1323
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1324
  have pos: "?f \<in> borel_measurable M" "?g \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1325
    and neg: "?mf \<in> borel_measurable M" "?mg \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1326
    and p: "?p \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1327
    and n: "?n \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1328
    using assms by (simp_all add: integrable_def)
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1329
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1330
  have *: "\<And>x. Real (?p x) = Real a * ?f x + ?g x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1331
          "\<And>x. Real (?n x) = Real a * ?mf x + ?mg x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1332
          "\<And>x. Real (- ?p x) = 0"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1333
          "\<And>x. Real (- ?n x) = 0"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1334
    using `0 \<le> a` by (auto simp: max_def min_def zero_le_mult_iff mult_le_0_iff add_nonpos_nonpos)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1335
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1336
  note linear =
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1337
    positive_integral_linear[OF pos]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1338
    positive_integral_linear[OF neg]
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1339
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1340
  have "integrable ?p" "integrable ?n"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1341
      "\<And>t. a * f t + g t = ?p t - ?n t" "\<And>t. 0 \<le> ?p t" "\<And>t. 0 \<le> ?n t"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1342
    using assms p n unfolding integrable_def * linear by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1343
  note diff = integral_of_positive_diff[OF this]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1344
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1345
  show "integrable (\<lambda>t. a * f t + g t)" by (rule diff)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1346
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1347
  from assms show "integral (\<lambda>t. a * f t + g t) = a * integral f + integral g"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1348
    unfolding diff(2) unfolding integral_def * linear integrable_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1349
    by (cases "?PI ?f", cases "?PI ?mf", cases "?PI ?g", cases "?PI ?mg")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1350
       (auto simp add: field_simps zero_le_mult_iff)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1351
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1352
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1353
lemma (in measure_space) integral_add[simp, intro]:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1354
  assumes "integrable f" "integrable g"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1355
  shows "integrable (\<lambda>t. f t + g t)"
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1356
  and "integral (\<lambda>t. f t + g t) = integral f + integral g"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1357
  using assms integral_linear[where a=1] by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1358
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1359
lemma (in measure_space) integral_zero[simp, intro]:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1360
  shows "integrable (\<lambda>x. 0)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1361
  and "integral (\<lambda>x. 0) = 0"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1362
  unfolding integrable_def integral_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1363
  by (auto simp add: borel_measurable_const)
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1364
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1365
lemma (in measure_space) integral_cmult[simp, intro]:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1366
  assumes "integrable f"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1367
  shows "integrable (\<lambda>t. a * f t)" (is ?P)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1368
  and "integral (\<lambda>t. a * f t) = a * integral f" (is ?I)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1369
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1370
  have "integrable (\<lambda>t. a * f t) \<and> integral (\<lambda>t. a * f t) = a * integral f"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1371
  proof (cases rule: le_cases)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1372
    assume "0 \<le> a" show ?thesis
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1373
      using integral_linear[OF assms integral_zero(1) `0 \<le> a`]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1374
      by (simp add: integral_zero)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1375
  next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1376
    assume "a \<le> 0" hence "0 \<le> - a" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1377
    have *: "\<And>t. - a * t + 0 = (-a) * t" by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1378
    show ?thesis using integral_linear[OF assms integral_zero(1) `0 \<le> - a`]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1379
        integral_minus(1)[of "\<lambda>t. - a * f t"]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1380
      unfolding * integral_zero by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1381
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1382
  thus ?P ?I by auto
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1383
qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1384
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1385
lemma (in measure_space) integral_mono:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1386
  assumes fg: "integrable f" "integrable g"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1387
  and mono: "\<And>t. t \<in> space M \<Longrightarrow> f t \<le> g t"
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1388
  shows "integral f \<le> integral g"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1389
  using fg unfolding integral_def integrable_def diff_minus
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1390
proof (safe intro!: add_mono real_of_pinfreal_mono le_imp_neg_le positive_integral_mono)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1391
  fix x assume "x \<in> space M" from mono[OF this]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1392
  show "Real (f x) \<le> Real (g x)" "Real (- g x) \<le> Real (- f x)" by auto
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1393
qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1394
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1395
lemma (in measure_space) integral_diff[simp, intro]:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1396
  assumes f: "integrable f" and g: "integrable g"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1397
  shows "integrable (\<lambda>t. f t - g t)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1398
  and "integral (\<lambda>t. f t - g t) = integral f - integral g"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1399
  using integral_add[OF f integral_minus(1)[OF g]]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1400
  unfolding diff_minus integral_minus(2)[OF g]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1401
  by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1402
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1403
lemma (in measure_space) integral_indicator[simp, intro]:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1404
  assumes "a \<in> sets M" and "\<mu> a \<noteq> \<omega>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1405
  shows "integral (indicator a) = real (\<mu> a)" (is ?int)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1406
  and "integrable (indicator a)" (is ?able)
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1407
proof -
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1408
  have *:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1409
    "\<And>A x. Real (indicator A x) = indicator A x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1410
    "\<And>A x. Real (- indicator A x) = 0" unfolding indicator_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1411
  show ?int ?able
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1412
    using assms unfolding integral_def integrable_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1413
    by (auto simp: * positive_integral_indicator borel_measurable_indicator)
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1414
qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1415
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1416
lemma (in measure_space) integral_cmul_indicator:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1417
  assumes "A \<in> sets M" and "c \<noteq> 0 \<Longrightarrow> \<mu> A \<noteq> \<omega>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1418
  shows "integrable (\<lambda>x. c * indicator A x)" (is ?P)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1419
  and "integral (\<lambda>x. c * indicator A x) = c * real (\<mu> A)" (is ?I)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1420
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1421
  show ?P
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1422
  proof (cases "c = 0")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1423
    case False with assms show ?thesis by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1424
  qed simp
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1425
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1426
  show ?I
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1427
  proof (cases "c = 0")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1428
    case False with assms show ?thesis by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1429
  qed simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1430
qed
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1431
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1432
lemma (in measure_space) integral_setsum[simp, intro]:
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1433
  assumes "\<And>n. n \<in> S \<Longrightarrow> integrable (f n)"
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1434
  shows "integral (\<lambda>x. \<Sum> i \<in> S. f i x) = (\<Sum> i \<in> S. integral (f i))" (is "?int S")
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1435
    and "integrable (\<lambda>x. \<Sum> i \<in> S. f i x)" (is "?I S")
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1436
proof -
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1437
  have "?int S \<and> ?I S"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1438
  proof (cases "finite S")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1439
    assume "finite S"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1440
    from this assms show ?thesis by (induct S) simp_all
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1441
  qed simp
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1442
  thus "?int S" and "?I S" by auto
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1443
qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1444
36624
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
  1445
lemma (in measure_space) integrable_abs:
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
  1446
  assumes "integrable f"
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
  1447
  shows "integrable (\<lambda> x. \<bar>f x\<bar>)"
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
  1448
proof -
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1449
  have *:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1450
    "\<And>x. Real \<bar>f x\<bar> = Real (f x) + Real (- f x)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1451
    "\<And>x. Real (- \<bar>f x\<bar>) = 0" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1452
  have abs: "(\<lambda>x. \<bar>f x\<bar>) \<in> borel_measurable M" and
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1453
    f: "(\<lambda>x. Real (f x)) \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1454
        "(\<lambda>x. Real (- f x)) \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1455
    using assms unfolding integrable_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1456
  from abs assms show ?thesis unfolding integrable_def *
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1457
    using positive_integral_linear[OF f, of 1] by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1458
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1459
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1460
lemma (in measure_space) integrable_bound:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1461
  assumes "integrable f"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1462
  and f: "\<And>x. x \<in> space M \<Longrightarrow> 0 \<le> f x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1463
    "\<And>x. x \<in> space M \<Longrightarrow> \<bar>g x\<bar> \<le> f x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1464
  assumes borel: "g \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1465
  shows "integrable g"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1466
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1467
  have "positive_integral (\<lambda>x. Real (g x)) \<le> positive_integral (\<lambda>x. Real \<bar>g x\<bar>)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1468
    by (auto intro!: positive_integral_mono)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1469
  also have "\<dots> \<le> positive_integral (\<lambda>x. Real (f x))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1470
    using f by (auto intro!: positive_integral_mono)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1471
  also have "\<dots> < \<omega>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1472
    using `integrable f` unfolding integrable_def by (auto simp: pinfreal_less_\<omega>)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1473
  finally have pos: "positive_integral (\<lambda>x. Real (g x)) < \<omega>" .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1474
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1475
  have "positive_integral (\<lambda>x. Real (- g x)) \<le> positive_integral (\<lambda>x. Real (\<bar>g x\<bar>))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1476
    by (auto intro!: positive_integral_mono)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1477
  also have "\<dots> \<le> positive_integral (\<lambda>x. Real (f x))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1478
    using f by (auto intro!: positive_integral_mono)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1479
  also have "\<dots> < \<omega>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1480
    using `integrable f` unfolding integrable_def by (auto simp: pinfreal_less_\<omega>)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1481
  finally have neg: "positive_integral (\<lambda>x. Real (- g x)) < \<omega>" .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1482
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1483
  from neg pos borel show ?thesis
36624
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
  1484
    unfolding integrable_def by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1485
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1486
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1487
lemma (in measure_space) integrable_abs_iff:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1488
  "f \<in> borel_measurable M \<Longrightarrow> integrable (\<lambda> x. \<bar>f x\<bar>) \<longleftrightarrow> integrable f"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1489
  by (auto intro!: integrable_bound[where g=f] integrable_abs)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1490
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1491
lemma (in measure_space) integrable_max:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1492
  assumes int: "integrable f" "integrable g"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1493
  shows "integrable (\<lambda> x. max (f x) (g x))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1494
proof (rule integrable_bound)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1495
  show "integrable (\<lambda>x. \<bar>f x\<bar> + \<bar>g x\<bar>)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1496
    using int by (simp add: integrable_abs)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1497
  show "(\<lambda>x. max (f x) (g x)) \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1498
    using int unfolding integrable_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1499
next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1500
  fix x assume "x \<in> space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1501
  show "0 \<le> \<bar>f x\<bar> + \<bar>g x\<bar>" "\<bar>max (f x) (g x)\<bar> \<le> \<bar>f x\<bar> + \<bar>g x\<bar>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1502
    by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1503
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1504
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1505
lemma (in measure_space) integrable_min:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1506
  assumes int: "integrable f" "integrable g"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1507
  shows "integrable (\<lambda> x. min (f x) (g x))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1508
proof (rule integrable_bound)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1509
  show "integrable (\<lambda>x. \<bar>f x\<bar> + \<bar>g x\<bar>)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1510
    using int by (simp add: integrable_abs)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1511
  show "(\<lambda>x. min (f x) (g x)) \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1512
    using int unfolding integrable_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1513
next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1514
  fix x assume "x \<in> space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1515
  show "0 \<le> \<bar>f x\<bar> + \<bar>g x\<bar>" "\<bar>min (f x) (g x)\<bar> \<le> \<bar>f x\<bar> + \<bar>g x\<bar>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1516
    by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1517
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1518
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1519
lemma (in measure_space) integral_triangle_inequality:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1520
  assumes "integrable f"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1521
  shows "\<bar>integral f\<bar> \<le> integral (\<lambda>x. \<bar>f x\<bar>)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1522
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1523
  have "\<bar>integral f\<bar> = max (integral f) (- integral f)" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1524
  also have "\<dots> \<le> integral (\<lambda>x. \<bar>f x\<bar>)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1525
      using assms integral_minus(2)[of f, symmetric]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1526
      by (auto intro!: integral_mono integrable_abs simp del: integral_minus)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1527
  finally show ?thesis .
36624
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
  1528
qed
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
  1529
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1530
lemma (in measure_space) integral_positive:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1531
  assumes "integrable f" "\<And>x. x \<in> space M \<Longrightarrow> 0 \<le> f x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1532
  shows "0 \<le> integral f"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1533
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1534
  have "0 = integral (\<lambda>x. 0)" by (auto simp: integral_zero)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1535
  also have "\<dots> \<le> integral f"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1536
    using assms by (rule integral_mono[OF integral_zero(1)])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1537
  finally show ?thesis .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1538
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1539
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1540
lemma (in measure_space) integral_monotone_convergence_pos:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1541
  assumes i: "\<And>i. integrable (f i)" and mono: "\<And>x. mono (\<lambda>n. f n x)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1542
  and pos: "\<And>x i. 0 \<le> f i x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1543
  and lim: "\<And>x. (\<lambda>i. f i x) ----> u x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1544
  and ilim: "(\<lambda>i. integral (f i)) ----> x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1545
  shows "integrable u"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1546
  and "integral u = x"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1547
proof -
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1548
  { fix x have "0 \<le> u x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1549
      using mono pos[of 0 x] incseq_le[OF _ lim, of x 0]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1550
      by (simp add: mono_def incseq_def) }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1551
  note pos_u = this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1552
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1553
  hence [simp]: "\<And>i x. Real (- f i x) = 0" "\<And>x. Real (- u x) = 0"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1554
    using pos by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1555
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1556
  have SUP_F: "\<And>x. (SUP n. Real (f n x)) = Real (u x)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1557
    using mono pos pos_u lim by (rule SUP_eq_LIMSEQ[THEN iffD2])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1558
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1559
  have borel_f: "\<And>i. (\<lambda>x. Real (f i x)) \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1560
    using i unfolding integrable_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1561
  hence "(SUP i. (\<lambda>x. Real (f i x))) \<in> borel_measurable M"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1562
    by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1563
  hence borel_u: "u \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1564
    using pos_u by (auto simp: borel_measurable_Real_eq SUPR_fun_expand SUP_F)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1565
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1566
  have integral_eq: "\<And>n. positive_integral (\<lambda>x. Real (f n x)) = Real (integral (f n))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1567
    using i unfolding integral_def integrable_def by (auto simp: Real_real)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1568
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1569
  have pos_integral: "\<And>n. 0 \<le> integral (f n)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1570
    using pos i by (auto simp: integral_positive)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1571
  hence "0 \<le> x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1572
    using LIMSEQ_le_const[OF ilim, of 0] by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1573
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1574
  have "(\<lambda>i. positive_integral (\<lambda>x. Real (f i x))) \<up> positive_integral (\<lambda>x. Real (u x))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1575
  proof (rule positive_integral_isoton)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1576
    from SUP_F mono pos
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1577
    show "(\<lambda>i x. Real (f i x)) \<up> (\<lambda>x. Real (u x))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1578
      unfolding isoton_fun_expand by (auto simp: isoton_def mono_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1579
  qed (rule borel_f)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1580
  hence pI: "positive_integral (\<lambda>x. Real (u x)) =
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1581
      (SUP n. positive_integral (\<lambda>x. Real (f n x)))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1582
    unfolding isoton_def by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1583
  also have "\<dots> = Real x" unfolding integral_eq
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1584
  proof (rule SUP_eq_LIMSEQ[THEN iffD2])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1585
    show "mono (\<lambda>n. integral (f n))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1586
      using mono i by (auto simp: mono_def intro!: integral_mono)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1587
    show "\<And>n. 0 \<le> integral (f n)" using pos_integral .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1588
    show "0 \<le> x" using `0 \<le> x` .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1589
    show "(\<lambda>n. integral (f n)) ----> x" using ilim .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1590
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1591
  finally show  "integrable u" "integral u = x" using borel_u `0 \<le> x`
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1592
    unfolding integrable_def integral_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1593
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1594
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1595
lemma (in measure_space) integral_monotone_convergence:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1596
  assumes f: "\<And>i. integrable (f i)" and "mono f"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1597
  and lim: "\<And>x. (\<lambda>i. f i x) ----> u x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1598
  and ilim: "(\<lambda>i. integral (f i)) ----> x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1599
  shows "integrable u"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1600
  and "integral u = x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1601
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1602
  have 1: "\<And>i. integrable (\<lambda>x. f i x - f 0 x)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1603
      using f by (auto intro!: integral_diff)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1604
  have 2: "\<And>x. mono (\<lambda>n. f n x - f 0 x)" using `mono f`
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1605
      unfolding mono_def le_fun_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1606
  have 3: "\<And>x n. 0 \<le> f n x - f 0 x" using `mono f`
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1607
      unfolding mono_def le_fun_def by (auto simp: field_simps)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1608
  have 4: "\<And>x. (\<lambda>i. f i x - f 0 x) ----> u x - f 0 x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1609
    using lim by (auto intro!: LIMSEQ_diff)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1610
  have 5: "(\<lambda>i. integral (\<lambda>x. f i x - f 0 x)) ----> x - integral (f 0)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1611
    using f ilim by (auto intro!: LIMSEQ_diff simp: integral_diff)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1612
  note diff = integral_monotone_convergence_pos[OF 1 2 3 4 5]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1613
  have "integrable (\<lambda>x. (u x - f 0 x) + f 0 x)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1614
    using diff(1) f by (rule integral_add(1))
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1615
  with diff(2) f show "integrable u" "integral u = x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1616
    by (auto simp: integral_diff)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1617
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1618
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1619
lemma (in measure_space) integral_0_iff:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1620
  assumes "integrable f"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1621
  shows "integral (\<lambda>x. \<bar>f x\<bar>) = 0 \<longleftrightarrow> \<mu> {x\<in>space M. f x \<noteq> 0} = 0"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1622
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1623
  have *: "\<And>x. Real (- \<bar>f x\<bar>) = 0" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1624
  have "integrable (\<lambda>x. \<bar>f x\<bar>)" using assms by (rule integrable_abs)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1625
  hence "(\<lambda>x. Real (\<bar>f x\<bar>)) \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1626
    "positive_integral (\<lambda>x. Real \<bar>f x\<bar>) \<noteq> \<omega>" unfolding integrable_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1627
  from positive_integral_0_iff[OF this(1)] this(2)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1628
  show ?thesis unfolding integral_def *
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1629
    by (simp add: real_of_pinfreal_eq_0)
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1630
qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1631
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1632
lemma LIMSEQ_max:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1633
  "u ----> (x::real) \<Longrightarrow> (\<lambda>i. max (u i) 0) ----> max x 0"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1634
  by (fastsimp intro!: LIMSEQ_I dest!: LIMSEQ_D)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1635
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1636
lemma (in sigma_algebra) borel_measurable_LIMSEQ:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1637
  fixes u :: "nat \<Rightarrow> 'a \<Rightarrow> real"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1638
  assumes u': "\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>i. u i x) ----> u' x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1639
  and u: "\<And>i. u i \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1640
  shows "u' \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1641
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1642
  let "?pu x i" = "max (u i x) 0"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1643
  let "?nu x i" = "max (- u i x) 0"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1644
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1645
  { fix x assume x: "x \<in> space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1646
    have "(?pu x) ----> max (u' x) 0"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1647
      "(?nu x) ----> max (- u' x) 0"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1648
      using u'[OF x] by (auto intro!: LIMSEQ_max LIMSEQ_minus)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1649
    from LIMSEQ_imp_lim_INF[OF _ this(1)] LIMSEQ_imp_lim_INF[OF _ this(2)]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1650
    have "(SUP n. INF m. Real (u (n + m) x)) = Real (u' x)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1651
      "(SUP n. INF m. Real (- u (n + m) x)) = Real (- u' x)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1652
      by (simp_all add: Real_max'[symmetric]) }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1653
  note eq = this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1654
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1655
  have *: "\<And>x. real (Real (u' x)) - real (Real (- u' x)) = u' x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1656
    by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1657
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1658
  have "(SUP n. INF m. (\<lambda>x. Real (u (n + m) x))) \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1659
       "(SUP n. INF m. (\<lambda>x. Real (- u (n + m) x))) \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1660
    using u by (auto intro: borel_measurable_SUP borel_measurable_INF borel_measurable_Real)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1661
  with eq[THEN measurable_cong, of M "\<lambda>x. x" borel_space]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1662
  have "(\<lambda>x. Real (u' x)) \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1663
       "(\<lambda>x. Real (- u' x)) \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1664
    unfolding SUPR_fun_expand INFI_fun_expand by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1665
  note this[THEN borel_measurable_real]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1666
  from borel_measurable_diff[OF this]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1667
  show ?thesis unfolding * .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1668
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1669
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1670
lemma (in measure_space) integral_dominated_convergence:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1671
  assumes u: "\<And>i. integrable (u i)" and bound: "\<And>x j. x\<in>space M \<Longrightarrow> \<bar>u j x\<bar> \<le> w x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1672
  and w: "integrable w" "\<And>x. x \<in> space M \<Longrightarrow> 0 \<le> w x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1673
  and u': "\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>i. u i x) ----> u' x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1674
  shows "integrable u'"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1675
  and "(\<lambda>i. integral (\<lambda>x. \<bar>u i x - u' x\<bar>)) ----> 0" (is "?lim_diff")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1676
  and "(\<lambda>i. integral (u i)) ----> integral u'" (is ?lim)
36624
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
  1677
proof -
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1678
  { fix x j assume x: "x \<in> space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1679
    from u'[OF x] have "(\<lambda>i. \<bar>u i x\<bar>) ----> \<bar>u' x\<bar>" by (rule LIMSEQ_imp_rabs)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1680
    from LIMSEQ_le_const2[OF this]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1681
    have "\<bar>u' x\<bar> \<le> w x" using bound[OF x] by auto }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1682
  note u'_bound = this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1683
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1684
  from u[unfolded integrable_def]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1685
  have u'_borel: "u' \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1686
    using u' by (blast intro: borel_measurable_LIMSEQ[of u])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1687
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1688
  show "integrable u'"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1689
  proof (rule integrable_bound)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1690
    show "integrable w" by fact
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1691
    show "u' \<in> borel_measurable M" by fact
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1692
  next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1693
    fix x assume x: "x \<in> space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1694
    thus "0 \<le> w x" by fact
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1695
    show "\<bar>u' x\<bar> \<le> w x" using u'_bound[OF x] .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1696
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1697
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1698
  let "?diff n x" = "2 * w x - \<bar>u n x - u' x\<bar>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1699
  have diff: "\<And>n. integrable (\<lambda>x. \<bar>u n x - u' x\<bar>)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1700
    using w u `integrable u'`
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1701
    by (auto intro!: integral_add integral_diff integral_cmult integrable_abs)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1702
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1703
  { fix j x assume x: "x \<in> space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1704
    have "\<bar>u j x - u' x\<bar> \<le> \<bar>u j x\<bar> + \<bar>u' x\<bar>" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1705
    also have "\<dots> \<le> w x + w x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1706
      by (rule add_mono[OF bound[OF x] u'_bound[OF x]])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1707
    finally have "\<bar>u j x - u' x\<bar> \<le> 2 * w x" by simp }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1708
  note diff_less_2w = this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1709
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1710
  have PI_diff: "\<And>m n. positive_integral (\<lambda>x. Real (?diff (m + n) x)) =
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1711
    positive_integral (\<lambda>x. Real (2 * w x)) - positive_integral (\<lambda>x. Real \<bar>u (m + n) x - u' x\<bar>)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1712
    using diff w diff_less_2w
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1713
    by (subst positive_integral_diff[symmetric])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1714
       (auto simp: integrable_def intro!: positive_integral_cong)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1715
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1716
  have "integrable (\<lambda>x. 2 * w x)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1717
    using w by (auto intro: integral_cmult)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1718
  hence I2w_fin: "positive_integral (\<lambda>x. Real (2 * w x)) \<noteq> \<omega>" and
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1719
    borel_2w: "(\<lambda>x. Real (2 * w x)) \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1720
    unfolding integrable_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1721
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1722
  have "(INF n. SUP m. positive_integral (\<lambda>x. Real \<bar>u (m + n) x - u' x\<bar>)) = 0" (is "?lim_SUP = 0")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1723
  proof cases
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1724
    assume eq_0: "positive_integral (\<lambda>x. Real (2 * w x)) = 0"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1725
    have "\<And>i. positive_integral (\<lambda>x. Real \<bar>u i x - u' x\<bar>) \<le> positive_integral (\<lambda>x. Real (2 * w x))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1726
    proof (rule positive_integral_mono)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1727
      fix i x assume "x \<in> space M" from diff_less_2w[OF this, of i]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1728
      show "Real \<bar>u i x - u' x\<bar> \<le> Real (2 * w x)" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1729
    qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1730
    hence "\<And>i. positive_integral (\<lambda>x. Real \<bar>u i x - u' x\<bar>) = 0" using eq_0 by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1731
    thus ?thesis by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1732
  next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1733
    assume neq_0: "positive_integral (\<lambda>x. Real (2 * w x)) \<noteq> 0"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1734
    have "positive_integral (\<lambda>x. Real (2 * w x)) = positive_integral (SUP n. INF m. (\<lambda>x. Real (?diff (m + n) x)))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1735
    proof (rule positive_integral_cong, unfold SUPR_fun_expand INFI_fun_expand, subst add_commute)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1736
      fix x assume x: "x \<in> space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1737
      show "Real (2 * w x) = (SUP n. INF m. Real (?diff (n + m) x))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1738
      proof (rule LIMSEQ_imp_lim_INF[symmetric])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1739
        fix j show "0 \<le> ?diff j x" using diff_less_2w[OF x, of j] by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1740
      next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1741
        have "(\<lambda>i. ?diff i x) ----> 2 * w x - \<bar>u' x - u' x\<bar>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1742
          using u'[OF x] by (safe intro!: LIMSEQ_diff LIMSEQ_const LIMSEQ_imp_rabs)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1743
        thus "(\<lambda>i. ?diff i x) ----> 2 * w x" by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1744
      qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1745
    qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1746
    also have "\<dots> \<le> (SUP n. INF m. positive_integral (\<lambda>x. Real (?diff (m + n) x)))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1747
      using u'_borel w u unfolding integrable_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1748
      by (auto intro!: positive_integral_lim_INF)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1749
    also have "\<dots> = positive_integral (\<lambda>x. Real (2 * w x)) -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1750
        (INF n. SUP m. positive_integral (\<lambda>x. Real \<bar>u (m + n) x - u' x\<bar>))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1751
      unfolding PI_diff pinfreal_INF_minus[OF I2w_fin] pinfreal_SUP_minus ..
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1752
    finally show ?thesis using neq_0 I2w_fin by (rule pinfreal_le_minus_imp_0)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1753
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1754
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1755
  have [simp]: "\<And>n m x. Real (- \<bar>u (m + n) x - u' x\<bar>) = 0" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1756
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1757
  have [simp]: "\<And>n m. positive_integral (\<lambda>x. Real \<bar>u (m + n) x - u' x\<bar>) =
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1758
    Real (integral (\<lambda>x. \<bar>u (n + m) x - u' x\<bar>))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1759
    using diff by (subst add_commute) (simp add: integral_def integrable_def Real_real)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1760
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1761
  have "(SUP n. INF m. positive_integral (\<lambda>x. Real \<bar>u (m + n) x - u' x\<bar>)) \<le> ?lim_SUP"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1762
    (is "?lim_INF \<le> _") by (subst (1 2) add_commute) (rule lim_INF_le_lim_SUP)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1763
  hence "?lim_INF = Real 0" "?lim_SUP = Real 0" using `?lim_SUP = 0` by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1764
  thus ?lim_diff using diff by (auto intro!: integral_positive lim_INF_eq_lim_SUP)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1765
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1766
  show ?lim
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1767
  proof (rule LIMSEQ_I)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1768
    fix r :: real assume "0 < r"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1769
    from LIMSEQ_D[OF `?lim_diff` this]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1770
    obtain N where N: "\<And>n. N \<le> n \<Longrightarrow> integral (\<lambda>x. \<bar>u n x - u' x\<bar>) < r"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1771
      using diff by (auto simp: integral_positive)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1772
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1773
    show "\<exists>N. \<forall>n\<ge>N. norm (integral (u n) - integral u') < r"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1774
    proof (safe intro!: exI[of _ N])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1775
      fix n assume "N \<le> n"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1776
      have "\<bar>integral (u n) - integral u'\<bar> = \<bar>integral (\<lambda>x. u n x - u' x)\<bar>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1777
        using u `integrable u'` by (auto simp: integral_diff)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1778
      also have "\<dots> \<le> integral (\<lambda>x. \<bar>u n x - u' x\<bar>)" using u `integrable u'`
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1779
        by (rule_tac integral_triangle_inequality) (auto intro!: integral_diff)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1780
      also note N[OF `N \<le> n`]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1781
      finally show "norm (integral (u n) - integral u') < r" by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1782
    qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1783
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1784
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1785
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1786
lemma (in measure_space) integral_sums:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1787
  assumes borel: "\<And>i. integrable (f i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1788
  and summable: "\<And>x. x \<in> space M \<Longrightarrow> summable (\<lambda>i. \<bar>f i x\<bar>)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1789
  and sums: "summable (\<lambda>i. integral (\<lambda>x. \<bar>f i x\<bar>))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1790
  shows "integrable (\<lambda>x. (\<Sum>i. f i x))" (is "integrable ?S")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1791
  and "(\<lambda>i. integral (f i)) sums integral (\<lambda>x. (\<Sum>i. f i x))" (is ?integral)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1792
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1793
  have "\<forall>x\<in>space M. \<exists>w. (\<lambda>i. \<bar>f i x\<bar>) sums w"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1794
    using summable unfolding summable_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1795
  from bchoice[OF this]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1796
  obtain w where w: "\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>i. \<bar>f i x\<bar>) sums w x" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1797
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1798
  let "?w y" = "if y \<in> space M then w y else 0"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1799
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1800
  obtain x where abs_sum: "(\<lambda>i. integral (\<lambda>x. \<bar>f i x\<bar>)) sums x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1801
    using sums unfolding summable_def ..
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1802
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1803
  have 1: "\<And>n. integrable (\<lambda>x. \<Sum>i = 0..<n. f i x)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1804
    using borel by (auto intro!: integral_setsum)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1805
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1806
  { fix j x assume [simp]: "x \<in> space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1807
    have "\<bar>\<Sum>i = 0..< j. f i x\<bar> \<le> (\<Sum>i = 0..< j. \<bar>f i x\<bar>)" by (rule setsum_abs)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1808
    also have "\<dots> \<le> w x" using w[of x] series_pos_le[of "\<lambda>i. \<bar>f i x\<bar>"] unfolding sums_iff by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1809
    finally have "\<bar>\<Sum>i = 0..<j. f i x\<bar> \<le> ?w x" by simp }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1810
  note 2 = this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1811
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1812
  have 3: "integrable ?w"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1813
  proof (rule integral_monotone_convergence(1))
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1814
    let "?F n y" = "(\<Sum>i = 0..<n. \<bar>f i y\<bar>)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1815
    let "?w' n y" = "if y \<in> space M then ?F n y else 0"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1816
    have "\<And>n. integrable (?F n)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1817
      using borel by (auto intro!: integral_setsum integrable_abs)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1818
    thus "\<And>n. integrable (?w' n)" by (simp cong: integrable_cong)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1819
    show "mono ?w'"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1820
      by (auto simp: mono_def le_fun_def intro!: setsum_mono2)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1821
    { fix x show "(\<lambda>n. ?w' n x) ----> ?w x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1822
        using w by (cases "x \<in> space M") (simp_all add: LIMSEQ_const sums_def) }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1823
    have *: "\<And>n. integral (?w' n) = (\<Sum>i = 0..< n. integral (\<lambda>x. \<bar>f i x\<bar>))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1824
      using borel by (simp add: integral_setsum integrable_abs cong: integral_cong)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1825
    from abs_sum
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1826
    show "(\<lambda>i. integral (?w' i)) ----> x" unfolding * sums_def .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1827
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1828
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1829
  have 4: "\<And>x. x \<in> space M \<Longrightarrow> 0 \<le> ?w x" using 2[of _ 0] by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1830
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1831
  from summable[THEN summable_rabs_cancel]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1832
  have 5: "\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>n. \<Sum>i = 0..<n. f i x) ----> (\<Sum>i. f i x)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1833
    by (auto intro: summable_sumr_LIMSEQ_suminf)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1834
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1835
  note int = integral_dominated_convergence(1,3)[OF 1 2 3 4 5]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1836
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1837
  from int show "integrable ?S" by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1838
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1839
  show ?integral unfolding sums_def integral_setsum(1)[symmetric, OF borel]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1840
    using int(2) by simp
36624
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
  1841
qed
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
  1842
35748
5f35613d9a65 Equality of integral and infinite sum.
hoelzl
parents: 35692
diff changeset
  1843
section "Lebesgue integration on countable spaces"
5f35613d9a65 Equality of integral and infinite sum.
hoelzl
parents: 35692
diff changeset
  1844
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1845
lemma (in measure_space) integral_on_countable:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1846
  assumes f: "f \<in> borel_measurable M"
35748
5f35613d9a65 Equality of integral and infinite sum.
hoelzl
parents: 35692
diff changeset
  1847
  and bij: "bij_betw enum S (f ` space M)"
5f35613d9a65 Equality of integral and infinite sum.
hoelzl
parents: 35692
diff changeset
  1848
  and enum_zero: "enum ` (-S) \<subseteq> {0}"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1849
  and fin: "\<And>x. x \<noteq> 0 \<Longrightarrow> \<mu> (f -` {x} \<inter> space M) \<noteq> \<omega>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1850
  and abs_summable: "summable (\<lambda>r. \<bar>enum r * real (\<mu> (f -` {enum r} \<inter> space M))\<bar>)"
35748
5f35613d9a65 Equality of integral and infinite sum.
hoelzl
parents: 35692
diff changeset
  1851
  shows "integrable f"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1852
  and "(\<lambda>r. enum r * real (\<mu> (f -` {enum r} \<inter> space M))) sums integral f" (is ?sums)
35748
5f35613d9a65 Equality of integral and infinite sum.
hoelzl
parents: 35692
diff changeset
  1853
proof -
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1854
  let "?A r" = "f -` {enum r} \<inter> space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1855
  let "?F r x" = "enum r * indicator (?A r) x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1856
  have enum_eq: "\<And>r. enum r * real (\<mu> (?A r)) = integral (?F r)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1857
    using f fin by (simp add: borel_measurable_vimage integral_cmul_indicator)
35748
5f35613d9a65 Equality of integral and infinite sum.
hoelzl
parents: 35692
diff changeset
  1858
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1859
  { fix x assume "x \<in> space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1860
    hence "f x \<in> enum ` S" using bij unfolding bij_betw_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1861
    then obtain i where "i\<in>S" "enum i = f x" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1862
    have F: "\<And>j. ?F j x = (if j = i then f x else 0)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1863
    proof cases
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1864
      fix j assume "j = i"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1865
      thus "?thesis j" using `x \<in> space M` `enum i = f x` by (simp add: indicator_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1866
    next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1867
      fix j assume "j \<noteq> i"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1868
      show "?thesis j" using bij `i \<in> S` `j \<noteq> i` `enum i = f x` enum_zero
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1869
        by (cases "j \<in> S") (auto simp add: indicator_def bij_betw_def inj_on_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1870
    qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1871
    hence F_abs: "\<And>j. \<bar>if j = i then f x else 0\<bar> = (if j = i then \<bar>f x\<bar> else 0)" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1872
    have "(\<lambda>i. ?F i x) sums f x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1873
         "(\<lambda>i. \<bar>?F i x\<bar>) sums \<bar>f x\<bar>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1874
      by (auto intro!: sums_single simp: F F_abs) }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1875
  note F_sums_f = this(1) and F_abs_sums_f = this(2)
35748
5f35613d9a65 Equality of integral and infinite sum.
hoelzl
parents: 35692
diff changeset
  1876
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1877
  have int_f: "integral f = integral (\<lambda>x. \<Sum>r. ?F r x)" "integrable f = integrable (\<lambda>x. \<Sum>r. ?F r x)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1878
    using F_sums_f by (auto intro!: integral_cong integrable_cong simp: sums_iff)
35748
5f35613d9a65 Equality of integral and infinite sum.
hoelzl
parents: 35692
diff changeset
  1879
5f35613d9a65 Equality of integral and infinite sum.
hoelzl
parents: 35692
diff changeset
  1880
  { fix r
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1881
    have "integral (\<lambda>x. \<bar>?F r x\<bar>) = integral (\<lambda>x. \<bar>enum r\<bar> * indicator (?A r) x)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1882
      by (auto simp: indicator_def intro!: integral_cong)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1883
    also have "\<dots> = \<bar>enum r\<bar> * real (\<mu> (?A r))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1884
      using f fin by (simp add: borel_measurable_vimage integral_cmul_indicator)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1885
    finally have "integral (\<lambda>x. \<bar>?F r x\<bar>) = \<bar>enum r * real (\<mu> (?A r))\<bar>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1886
      by (simp add: abs_mult_pos real_pinfreal_pos) }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1887
  note int_abs_F = this
35748
5f35613d9a65 Equality of integral and infinite sum.
hoelzl
parents: 35692
diff changeset
  1888
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1889
  have 1: "\<And>i. integrable (\<lambda>x. ?F i x)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1890
    using f fin by (simp add: borel_measurable_vimage integral_cmul_indicator)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1891
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1892
  have 2: "\<And>x. x \<in> space M \<Longrightarrow> summable (\<lambda>i. \<bar>?F i x\<bar>)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1893
    using F_abs_sums_f unfolding sums_iff by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1894
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1895
  from integral_sums(2)[OF 1 2, unfolded int_abs_F, OF _ abs_summable]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1896
  show ?sums unfolding enum_eq int_f by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1897
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1898
  from integral_sums(1)[OF 1 2, unfolded int_abs_F, OF _ abs_summable]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1899
  show "integrable f" unfolding int_f by simp
35748
5f35613d9a65 Equality of integral and infinite sum.
hoelzl
parents: 35692
diff changeset
  1900
qed
5f35613d9a65 Equality of integral and infinite sum.
hoelzl
parents: 35692
diff changeset
  1901
35692
f1315bbf1bc9 Moved theorems in Lebesgue to the right places
hoelzl
parents: 35582
diff changeset
  1902
section "Lebesgue integration on finite space"
f1315bbf1bc9 Moved theorems in Lebesgue to the right places
hoelzl
parents: 35582
diff changeset
  1903
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1904
lemma (in measure_space) integral_on_finite:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1905
  assumes f: "f \<in> borel_measurable M" and finite: "finite (f`space M)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1906
  and fin: "\<And>x. x \<noteq> 0 \<Longrightarrow> \<mu> (f -` {x} \<inter> space M) \<noteq> \<omega>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1907
  shows "integrable f"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1908
  and "integral (\<lambda>x. f x) =
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1909
    (\<Sum> r \<in> f`space M. r * real (\<mu> (f -` {r} \<inter> space M)))" (is "?integral")
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1910
proof -
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1911
  let "?A r" = "f -` {r} \<inter> space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1912
  let "?S x" = "\<Sum>r\<in>f`space M. r * indicator (?A r) x"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1913
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1914
  { fix x assume "x \<in> space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1915
    have "f x = (\<Sum>r\<in>f`space M. if x \<in> ?A r then r else 0)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1916
      using finite `x \<in> space M` by (simp add: setsum_cases)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1917
    also have "\<dots> = ?S x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1918
      by (auto intro!: setsum_cong)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1919
    finally have "f x = ?S x" . }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1920
  note f_eq = this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1921
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1922
  have f_eq_S: "integrable f \<longleftrightarrow> integrable ?S" "integral f = integral ?S"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1923
    by (auto intro!: integrable_cong integral_cong simp only: f_eq)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1924
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1925
  show "integrable f" ?integral using fin f f_eq_S
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1926
    by (simp_all add: integral_cmul_indicator borel_measurable_vimage)
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1927
qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1928
35977
30d42bfd0174 Added finite measure space.
hoelzl
parents: 35833
diff changeset
  1929
lemma sigma_algebra_cong:
30d42bfd0174 Added finite measure space.
hoelzl
parents: 35833
diff changeset
  1930
  fixes M :: "('a, 'b) algebra_scheme" and M' :: "('a, 'c) algebra_scheme"
30d42bfd0174 Added finite measure space.
hoelzl
parents: 35833
diff changeset
  1931
  assumes *: "sigma_algebra M"
30d42bfd0174 Added finite measure space.
hoelzl
parents: 35833
diff changeset
  1932
  and cong: "space M = space M'" "sets M = sets M'"
30d42bfd0174 Added finite measure space.
hoelzl
parents: 35833
diff changeset
  1933
  shows "sigma_algebra M'"
30d42bfd0174 Added finite measure space.
hoelzl
parents: 35833
diff changeset
  1934
using * unfolding sigma_algebra_def algebra_def sigma_algebra_axioms_def unfolding cong .
30d42bfd0174 Added finite measure space.
hoelzl
parents: 35833
diff changeset
  1935
30d42bfd0174 Added finite measure space.
hoelzl
parents: 35833
diff changeset
  1936
lemma finite_Pow_additivity_sufficient:
30d42bfd0174 Added finite measure space.
hoelzl
parents: 35833
diff changeset
  1937
  assumes "finite (space M)" and "sets M = Pow (space M)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1938
  and "positive \<mu>" and "additive M \<mu>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1939
  and "\<And>x. x \<in> space M \<Longrightarrow> \<mu> {x} \<noteq> \<omega>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1940
  shows "finite_measure_space M \<mu>"
35977
30d42bfd0174 Added finite measure space.
hoelzl
parents: 35833
diff changeset
  1941
proof -
30d42bfd0174 Added finite measure space.
hoelzl
parents: 35833
diff changeset
  1942
  have "sigma_algebra M"
30d42bfd0174 Added finite measure space.
hoelzl
parents: 35833
diff changeset
  1943
    using assms by (auto intro!: sigma_algebra_cong[OF sigma_algebra_Pow])
30d42bfd0174 Added finite measure space.
hoelzl
parents: 35833
diff changeset
  1944
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1945
  have "measure_space M \<mu>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1946
    by (rule sigma_algebra.finite_additivity_sufficient) (fact+)
35977
30d42bfd0174 Added finite measure space.
hoelzl
parents: 35833
diff changeset
  1947
  thus ?thesis
30d42bfd0174 Added finite measure space.
hoelzl
parents: 35833
diff changeset
  1948
    unfolding finite_measure_space_def finite_measure_space_axioms_def
30d42bfd0174 Added finite measure space.
hoelzl
parents: 35833
diff changeset
  1949
    using assms by simp
30d42bfd0174 Added finite measure space.
hoelzl
parents: 35833
diff changeset
  1950
qed
30d42bfd0174 Added finite measure space.
hoelzl
parents: 35833
diff changeset
  1951
30d42bfd0174 Added finite measure space.
hoelzl
parents: 35833
diff changeset
  1952
lemma finite_measure_spaceI:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1953
  assumes "measure_space M \<mu>" and "finite (space M)" and "sets M = Pow (space M)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1954
  and "\<And>x. x \<in> space M \<Longrightarrow> \<mu> {x} \<noteq> \<omega>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1955
  shows "finite_measure_space M \<mu>"
35977
30d42bfd0174 Added finite measure space.
hoelzl
parents: 35833
diff changeset
  1956
  unfolding finite_measure_space_def finite_measure_space_axioms_def
30d42bfd0174 Added finite measure space.
hoelzl
parents: 35833
diff changeset
  1957
  using assms by simp
30d42bfd0174 Added finite measure space.
hoelzl
parents: 35833
diff changeset
  1958
38705
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 38656
diff changeset
  1959
lemma (in finite_measure_space) borel_measurable_finite[intro, simp]: "f \<in> borel_measurable M"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1960
  unfolding measurable_def sets_eq_Pow by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1961
38705
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 38656
diff changeset
  1962
lemma (in finite_measure_space) simple_function_finite: "simple_function f"
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 38656
diff changeset
  1963
  unfolding simple_function_def sets_eq_Pow using finite_space by auto
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 38656
diff changeset
  1964
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 38656
diff changeset
  1965
lemma (in finite_measure_space) positive_integral_finite_eq_setsum:
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 38656
diff changeset
  1966
  "positive_integral f = (\<Sum>x \<in> space M. f x * \<mu> {x})"
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 38656
diff changeset
  1967
proof -
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 38656
diff changeset
  1968
  have *: "positive_integral f = positive_integral (\<lambda>x. \<Sum>y\<in>space M. f y * indicator {y} x)"
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 38656
diff changeset
  1969
    by (auto intro!: positive_integral_cong simp add: indicator_def if_distrib setsum_cases[OF finite_space])
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 38656
diff changeset
  1970
  show ?thesis unfolding * using borel_measurable_finite[of f]
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 38656
diff changeset
  1971
    by (simp add: positive_integral_setsum positive_integral_cmult_indicator sets_eq_Pow)
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 38656
diff changeset
  1972
qed
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 38656
diff changeset
  1973
35977
30d42bfd0174 Added finite measure space.
hoelzl
parents: 35833
diff changeset
  1974
lemma (in finite_measure_space) integral_finite_singleton:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1975
  shows "integrable f"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1976
  and "integral f = (\<Sum>x \<in> space M. f x * real (\<mu> {x}))" (is ?I)
35977
30d42bfd0174 Added finite measure space.
hoelzl
parents: 35833
diff changeset
  1977
proof -
38705
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 38656
diff changeset
  1978
  have [simp]:
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 38656
diff changeset
  1979
    "positive_integral (\<lambda>x. Real (f x)) = (\<Sum>x \<in> space M. Real (f x) * \<mu> {x})"
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 38656
diff changeset
  1980
    "positive_integral (\<lambda>x. Real (- f x)) = (\<Sum>x \<in> space M. Real (- f x) * \<mu> {x})"
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 38656
diff changeset
  1981
    unfolding positive_integral_finite_eq_setsum by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1982
38705
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 38656
diff changeset
  1983
  show "integrable f" using finite_space finite_measure
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 38656
diff changeset
  1984
    by (simp add: setsum_\<omega> integrable_def sets_eq_Pow)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1985
38705
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 38656
diff changeset
  1986
  show ?I using finite_measure
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 38656
diff changeset
  1987
    apply (simp add: integral_def sets_eq_Pow real_of_pinfreal_setsum[symmetric]
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 38656
diff changeset
  1988
      real_of_pinfreal_mult[symmetric] setsum_subtractf[symmetric])
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 38656
diff changeset
  1989
    by (rule setsum_cong) (simp_all split: split_if)
35977
30d42bfd0174 Added finite measure space.
hoelzl
parents: 35833
diff changeset
  1990
qed
30d42bfd0174 Added finite measure space.
hoelzl
parents: 35833
diff changeset
  1991
35748
5f35613d9a65 Equality of integral and infinite sum.
hoelzl
parents: 35692
diff changeset
  1992
end