src/HOL/Probability/Lebesgue_Integration.thy
author hoelzl
Fri, 04 Feb 2011 14:16:55 +0100
changeset 41705 1100512e16d8
parent 41689 3e39b0e730d6
child 41831 91a2b435dd7a
permissions -rw-r--r--
add auto support for AE_mp
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
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
     6
imports Measure Borel_Space
35582
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
lemma sums_If_finite:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
    10
  fixes f :: "nat \<Rightarrow> 'a::real_normed_vector"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    11
  assumes finite: "finite {r. P r}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    12
  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
    13
proof cases
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    14
  assume "{r. P r} = {}" hence "\<And>r. \<not> P r" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    15
  thus ?thesis by (simp add: sums_zero)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    16
next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    17
  assume not_empty: "{r. P r} \<noteq> {}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    18
  have "?F sums (\<Sum>r = 0..< Suc (Max {r. P r}). ?F r)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    19
    by (rule series_zero)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    20
       (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
    21
  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
    22
    by (subst setsum_cases)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    23
       (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
    24
  finally show ?thesis .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    25
qed
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
    26
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    27
lemma sums_single:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
    28
  fixes f :: "nat \<Rightarrow> 'a::real_normed_vector"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
    29
  shows "(\<lambda>r. if r = i then f r else 0) sums f i"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    30
  using sums_If_finite[of "\<lambda>r. r = i" f] by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    31
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    32
section "Simple function"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
    33
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    34
text {*
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    35
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    36
Our simple functions are not restricted to positive real numbers. Instead
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    37
they are just functions with a finite range and are measurable when singleton
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    38
sets are measurable.
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
    39
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    40
*}
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    41
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
    42
definition "simple_function M g \<longleftrightarrow>
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    43
    finite (g ` space M) \<and>
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    44
    (\<forall>x \<in> g ` space M. g -` {x} \<inter> space M \<in> sets M)"
36624
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
    45
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    46
lemma (in sigma_algebra) simple_functionD:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
    47
  assumes "simple_function M g"
40875
9a9d33f6fb46 generalized simple_functionD
hoelzl
parents: 40873
diff changeset
    48
  shows "finite (g ` space M)" and "g -` X \<inter> space M \<in> sets M"
40871
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
    49
proof -
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
    50
  show "finite (g ` space M)"
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
    51
    using assms unfolding simple_function_def by auto
40875
9a9d33f6fb46 generalized simple_functionD
hoelzl
parents: 40873
diff changeset
    52
  have "g -` X \<inter> space M = g -` (X \<inter> g`space M) \<inter> space M" by auto
9a9d33f6fb46 generalized simple_functionD
hoelzl
parents: 40873
diff changeset
    53
  also have "\<dots> = (\<Union>x\<in>X \<inter> g`space M. g-`{x} \<inter> space M)" by auto
9a9d33f6fb46 generalized simple_functionD
hoelzl
parents: 40873
diff changeset
    54
  finally show "g -` X \<inter> space M \<in> sets M" using assms
9a9d33f6fb46 generalized simple_functionD
hoelzl
parents: 40873
diff changeset
    55
    by (auto intro!: finite_UN simp del: UN_simps simp: simple_function_def)
40871
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
    56
qed
36624
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
    57
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    58
lemma (in sigma_algebra) simple_function_indicator_representation:
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40875
diff changeset
    59
  fixes f ::"'a \<Rightarrow> pextreal"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
    60
  assumes f: "simple_function M f" and x: "x \<in> space M"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    61
  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
    62
  (is "?l = ?r")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    63
proof -
38705
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 38656
diff changeset
    64
  have "?r = (\<Sum>y \<in> f ` space M.
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    65
    (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
    66
    by (auto intro!: setsum_cong2)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    67
  also have "... =  f x *  indicator (f -` {f x} \<inter> space M) x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    68
    using assms by (auto dest: simple_functionD simp: setsum_delta)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    69
  also have "... = f x" using x by (auto simp: indicator_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    70
  finally show ?thesis by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    71
qed
36624
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
    72
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    73
lemma (in measure_space) simple_function_notspace:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
    74
  "simple_function M (\<lambda>x. h x * indicator (- space M) x::pextreal)" (is "simple_function M ?h")
35692
f1315bbf1bc9 Moved theorems in Lebesgue to the right places
hoelzl
parents: 35582
diff changeset
    75
proof -
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    76
  have "?h ` space M \<subseteq> {0}" unfolding indicator_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    77
  hence [simp, intro]: "finite (?h ` space M)" by (auto intro: finite_subset)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    78
  have "?h -` {0} \<inter> space M = space M" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    79
  thus ?thesis unfolding simple_function_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    80
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    81
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    82
lemma (in sigma_algebra) simple_function_cong:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    83
  assumes "\<And>t. t \<in> space M \<Longrightarrow> f t = g t"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
    84
  shows "simple_function M f \<longleftrightarrow> simple_function M g"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    85
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    86
  have "f ` space M = g ` space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    87
    "\<And>x. f -` {x} \<inter> space M = g -` {x} \<inter> space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    88
    using assms by (auto intro!: image_eqI)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    89
  thus ?thesis unfolding simple_function_def using assms by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    90
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    91
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
    92
lemma (in sigma_algebra) simple_function_cong_algebra:
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
    93
  assumes "sets N = sets M" "space N = space M"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
    94
  shows "simple_function M f \<longleftrightarrow> simple_function N f"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
    95
  unfolding simple_function_def assms ..
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
    96
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    97
lemma (in sigma_algebra) borel_measurable_simple_function:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
    98
  assumes "simple_function M f"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    99
  shows "f \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   100
proof (rule borel_measurableI)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   101
  fix S
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   102
  let ?I = "f ` (f -` S \<inter> space M)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   103
  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
   104
  have "finite ?I"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   105
    using assms unfolding simple_function_def
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   106
    using finite_subset[of "f ` (f -` S \<inter> space M)" "f ` space M"] by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   107
  hence "?U \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   108
    apply (rule finite_UN)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   109
    using assms unfolding simple_function_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   110
  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
   111
qed
f1315bbf1bc9 Moved theorems in Lebesgue to the right places
hoelzl
parents: 35582
diff changeset
   112
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   113
lemma (in sigma_algebra) simple_function_borel_measurable:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   114
  fixes f :: "'a \<Rightarrow> 'x::t2_space"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   115
  assumes "f \<in> borel_measurable M" and "finite (f ` space M)"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   116
  shows "simple_function M f"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   117
  using assms unfolding simple_function_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   118
  by (auto intro: borel_measurable_vimage)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   119
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   120
lemma (in sigma_algebra) simple_function_const[intro, simp]:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   121
  "simple_function M (\<lambda>x. c)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   122
  by (auto intro: finite_subset simp: simple_function_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   123
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   124
lemma (in sigma_algebra) simple_function_compose[intro, simp]:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   125
  assumes "simple_function M f"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   126
  shows "simple_function M (g \<circ> f)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   127
  unfolding simple_function_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   128
proof safe
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   129
  show "finite ((g \<circ> f) ` space M)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   130
    using assms unfolding simple_function_def by (auto simp: image_compose)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   131
next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   132
  fix x assume "x \<in> space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   133
  let ?G = "g -` {g (f x)} \<inter> (f`space M)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   134
  have *: "(g \<circ> f) -` {(g \<circ> f) x} \<inter> space M =
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   135
    (\<Union>x\<in>?G. f -` {x} \<inter> space M)" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   136
  show "(g \<circ> f) -` {(g \<circ> f) x} \<inter> space M \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   137
    using assms unfolding simple_function_def *
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   138
    by (rule_tac finite_UN) (auto intro!: finite_UN)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   139
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   140
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   141
lemma (in sigma_algebra) simple_function_indicator[intro, simp]:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   142
  assumes "A \<in> sets M"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   143
  shows "simple_function M (indicator A)"
35692
f1315bbf1bc9 Moved theorems in Lebesgue to the right places
hoelzl
parents: 35582
diff changeset
   144
proof -
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   145
  have "indicator A ` space M \<subseteq> {0, 1}" (is "?S \<subseteq> _")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   146
    by (auto simp: indicator_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   147
  hence "finite ?S" by (rule finite_subset) simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   148
  moreover have "- A \<inter> space M = space M - A" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   149
  ultimately show ?thesis unfolding simple_function_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   150
    using assms by (auto simp: indicator_def_raw)
35692
f1315bbf1bc9 Moved theorems in Lebesgue to the right places
hoelzl
parents: 35582
diff changeset
   151
qed
f1315bbf1bc9 Moved theorems in Lebesgue to the right places
hoelzl
parents: 35582
diff changeset
   152
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   153
lemma (in sigma_algebra) simple_function_Pair[intro, simp]:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   154
  assumes "simple_function M f"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   155
  assumes "simple_function M g"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   156
  shows "simple_function M (\<lambda>x. (f x, g x))" (is "simple_function M ?p")
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   157
  unfolding simple_function_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   158
proof safe
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   159
  show "finite (?p ` space M)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   160
    using assms unfolding simple_function_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   161
    by (rule_tac finite_subset[of _ "f`space M \<times> g`space M"]) auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   162
next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   163
  fix x assume "x \<in> space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   164
  have "(\<lambda>x. (f x, g x)) -` {(f x, g x)} \<inter> space M =
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   165
      (f -` {f x} \<inter> space M) \<inter> (g -` {g x} \<inter> space M)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   166
    by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   167
  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
   168
    using assms unfolding simple_function_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   169
qed
35692
f1315bbf1bc9 Moved theorems in Lebesgue to the right places
hoelzl
parents: 35582
diff changeset
   170
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   171
lemma (in sigma_algebra) simple_function_compose1:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   172
  assumes "simple_function M f"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   173
  shows "simple_function M (\<lambda>x. g (f x))"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   174
  using simple_function_compose[OF assms, of g]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   175
  by (simp add: comp_def)
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   176
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   177
lemma (in sigma_algebra) simple_function_compose2:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   178
  assumes "simple_function M f" and "simple_function M g"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   179
  shows "simple_function M (\<lambda>x. h (f x) (g x))"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   180
proof -
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   181
  have "simple_function M ((\<lambda>(x, y). h x y) \<circ> (\<lambda>x. (f x, g x)))"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   182
    using assms by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   183
  thus ?thesis by (simp_all add: comp_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   184
qed
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   185
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   186
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
   187
  and simple_function_diff[intro, simp] = simple_function_compose2[where h="op -"]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   188
  and simple_function_uminus[intro, simp] = simple_function_compose[where g="uminus"]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   189
  and simple_function_mult[intro, simp] = simple_function_compose2[where h="op *"]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   190
  and simple_function_div[intro, simp] = simple_function_compose2[where h="op /"]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   191
  and simple_function_inverse[intro, simp] = simple_function_compose[where g="inverse"]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   192
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   193
lemma (in sigma_algebra) simple_function_setsum[intro, simp]:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   194
  assumes "\<And>i. i \<in> P \<Longrightarrow> simple_function M (f i)"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   195
  shows "simple_function M (\<lambda>x. \<Sum>i\<in>P. f i x)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   196
proof cases
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   197
  assume "finite P" from this assms show ?thesis by induct auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   198
qed auto
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   199
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   200
lemma (in sigma_algebra) simple_function_le_measurable:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   201
  assumes "simple_function M f" "simple_function M g"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   202
  shows "{x \<in> space M. f x \<le> g x} \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   203
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   204
  have *: "{x \<in> space M. f x \<le> g x} =
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   205
    (\<Union>(F, G)\<in>f`space M \<times> g`space M.
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   206
      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
   207
    apply (auto split: split_if_asm)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   208
    apply (rule_tac x=x in bexI)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   209
    apply (rule_tac x=x in bexI)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   210
    by simp_all
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   211
  have **: "\<And>x y. x \<in> space M \<Longrightarrow> y \<in> space M \<Longrightarrow>
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   212
    (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
   213
    using assms unfolding simple_function_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   214
  have "finite (f`space M \<times> g`space M)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   215
    using assms unfolding simple_function_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   216
  thus ?thesis unfolding *
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   217
    apply (rule finite_UN)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   218
    using assms unfolding simple_function_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   219
    by (auto intro!: **)
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   220
qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   221
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   222
lemma (in sigma_algebra) borel_measurable_implies_simple_function_sequence:
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40875
diff changeset
   223
  fixes u :: "'a \<Rightarrow> pextreal"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   224
  assumes u: "u \<in> borel_measurable M"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   225
  shows "\<exists>f. (\<forall>i. simple_function M (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
   226
proof -
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   227
  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
   228
    (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
   229
    (is "\<exists>f. \<forall>x j. ?P x j (f x j)")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   230
  proof(rule choice, rule, rule choice, rule)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   231
    fix x j show "\<exists>n. ?P x j n"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   232
    proof cases
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   233
      assume *: "u x < of_nat j"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   234
      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
   235
      from reals_Archimedean6a[of "r * 2^j"]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   236
      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
   237
        using `0 \<le> r` by (auto simp: zero_le_power zero_le_mult_iff)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   238
      thus ?thesis using r * by (auto intro!: exI[of _ n])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   239
    qed auto
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   240
  qed
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   241
  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
   242
    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
   243
    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
   244
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   245
  { fix j x P
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   246
    assume 1: "of_nat j \<le> u x \<Longrightarrow> P (j * 2^j)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   247
    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
   248
    have "P (f x j)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   249
    proof cases
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   250
      assume "of_nat j \<le> u x" thus "P (f x j)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   251
        using top[of j x] 1 by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   252
    next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   253
      assume "\<not> of_nat j \<le> u x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   254
      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
   255
        using upper lower by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   256
      from 2[OF this] show "P (f x j)" .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   257
    qed }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   258
  note fI = this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   259
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   260
  { fix j x
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   261
    have "f x j = j * 2 ^ j \<longleftrightarrow> of_nat j \<le> u x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   262
      by (rule fI, simp, cases "u x") (auto split: split_if_asm) }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   263
  note f_eq = this
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   264
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   265
  { fix j x
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   266
    have "f x j \<le> j * 2 ^ j"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   267
    proof (rule fI)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   268
      fix k assume *: "u x < of_nat j"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   269
      assume "of_nat k \<le> u x * 2 ^ j"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   270
      also have "\<dots> \<le> of_nat (j * 2^j)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   271
        using * by (cases "u x") (auto simp: zero_le_mult_iff)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   272
      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
   273
    qed simp }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   274
  note f_upper = this
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   275
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40875
diff changeset
   276
  let "?g j x" = "of_nat (f x j) / 2^j :: pextreal"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   277
  show ?thesis unfolding simple_function_def isoton_fun_expand unfolding isoton_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   278
  proof (safe intro!: exI[of _ ?g])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   279
    fix j
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   280
    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
   281
      using f_upper by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   282
    thus "finite (?g j ` space M)" by (rule finite_subset) auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   283
  next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   284
    fix j t assume "t \<in> space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   285
    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
   286
      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
   287
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   288
    show "?g j -` {?g j t} \<inter> space M \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   289
    proof cases
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   290
      assume "of_nat j \<le> u t"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   291
      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
   292
        unfolding ** f_eq[symmetric] by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   293
      thus "?g j -` {?g j t} \<inter> space M \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   294
        using u by auto
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   295
    next
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   296
      assume not_t: "\<not> of_nat j \<le> u t"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   297
      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
   298
      have split_vimage: "?g j -` {?g j t} \<inter> space M =
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   299
          {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
   300
        unfolding **
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   301
      proof safe
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   302
        fix x assume [simp]: "f t j = f x j"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   303
        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
   304
        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
   305
          using upper lower by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   306
        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
   307
          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
   308
        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
   309
      next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   310
        fix x
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   311
        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
   312
        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
   313
          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
   314
        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
   315
        note 2
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   316
        also have "\<dots> \<le> of_nat (j*2^j)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   317
          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
   318
        finally have bound_ux: "u x < of_nat j"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   319
          by (cases "u x") (auto simp: zero_le_mult_iff power_le_zero_eq)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   320
        show "f t j = f x j"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   321
        proof (rule antisym)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   322
          from 1 lower[OF bound_ux]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   323
          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
   324
          from upper[OF bound_ux] 2
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   325
          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
   326
        qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   327
      qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   328
      show ?thesis unfolding split_vimage using u by auto
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   329
    qed
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   330
  next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   331
    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
   332
  next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   333
    fix t
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   334
    { fix i
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   335
      have "f t i * 2 \<le> f t (Suc i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   336
      proof (rule fI)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   337
        assume "of_nat (Suc i) \<le> u t"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   338
        hence "of_nat i \<le> u t" by (cases "u t") auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   339
        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
   340
      next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   341
        fix k
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   342
        assume *: "u t * 2 ^ Suc i < of_nat (Suc k)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   343
        show "f t i * 2 \<le> k"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   344
        proof (rule fI)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   345
          assume "of_nat i \<le> u t"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   346
          hence "of_nat (i * 2^Suc i) \<le> u t * 2^Suc i"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   347
            by (cases "u t") (auto simp: zero_le_mult_iff power_le_zero_eq)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   348
          also have "\<dots> < of_nat (Suc k)" using * by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   349
          finally show "i * 2 ^ i * 2 \<le> k"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   350
            by (auto simp del: real_of_nat_mult)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   351
        next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   352
          fix j assume "of_nat j \<le> u t * 2 ^ i"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   353
          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
   354
        qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   355
      qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   356
      thus "?g i t \<le> ?g (Suc i) t"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   357
        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
   358
    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
   359
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   360
    show "(SUP j. of_nat (f t j) / 2 ^ j) = u t"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40875
diff changeset
   361
    proof (rule pextreal_SUPI)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   362
      fix j show "of_nat (f t j) / 2 ^ j \<le> u t"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   363
      proof (rule fI)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   364
        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
   365
          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
   366
      next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   367
        fix k assume "of_nat k \<le> u t * 2 ^ j"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   368
        thus "of_nat k / 2 ^ j \<le> u t"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   369
          by (cases "u t")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   370
             (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
   371
      qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   372
    next
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40875
diff changeset
   373
      fix y :: pextreal assume *: "\<And>j. j \<in> UNIV \<Longrightarrow> of_nat (f t j) / 2 ^ j \<le> y"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   374
      show "u t \<le> y"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   375
      proof (cases "u t")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   376
        case (preal r)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   377
        show ?thesis
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   378
        proof (rule ccontr)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   379
          assume "\<not> u t \<le> y"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   380
          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
   381
          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
   382
          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
   383
          let ?N = "max n (natfloor r + 1)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   384
          have "u t < of_nat ?N" "n \<le> ?N"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   385
            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
   386
            using real_natfloor_add_one_gt
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 38656
diff changeset
   387
            by (auto simp: max_def real_of_nat_Suc)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   388
          from lower[OF this(1)]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   389
          have "r < (real (f t ?N) + 1) / 2 ^ ?N" unfolding less_divide_eq
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   390
            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
   391
          hence "u t < of_nat (f t ?N) / 2 ^ ?N + 1 / 2 ^ ?N"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   392
            using preal by (auto simp: field_simps divide_real_def[symmetric])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   393
          with n[OF `n \<le> ?N`] p preal *[of ?N]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   394
          show False
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   395
            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
   396
        qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   397
      next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   398
        case infinite
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   399
        { 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
   400
          hence "of_nat j \<le> y" using *[of j]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   401
            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
   402
        note all_less_y = this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   403
        show ?thesis unfolding infinite
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   404
        proof (rule ccontr)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   405
          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
   406
          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
   407
          with all_less_y[of n] r show False by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   408
        qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   409
      qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   410
    qed
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   411
  qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   412
qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   413
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   414
lemma (in sigma_algebra) borel_measurable_implies_simple_function_sequence':
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40875
diff changeset
   415
  fixes u :: "'a \<Rightarrow> pextreal"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   416
  assumes "u \<in> borel_measurable M"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   417
  obtains (x) f where "f \<up> u" "\<And>i. simple_function M (f i)" "\<And>i. \<omega>\<notin>f i`space M"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   418
proof -
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   419
  from borel_measurable_implies_simple_function_sequence[OF assms]
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   420
  obtain f where x: "\<And>i. simple_function M (f i)" "f \<up> u"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   421
    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
   422
  { 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
   423
  with x show thesis by (auto intro!: that[of f])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   424
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   425
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   426
lemma (in sigma_algebra) simple_function_eq_borel_measurable:
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40875
diff changeset
   427
  fixes f :: "'a \<Rightarrow> pextreal"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   428
  shows "simple_function M f \<longleftrightarrow>
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   429
    finite (f`space M) \<and> f \<in> borel_measurable M"
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   430
  using simple_function_borel_measurable[of f]
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   431
    borel_measurable_simple_function[of f]
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   432
  by (fastsimp simp: simple_function_def)
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   433
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   434
lemma (in measure_space) simple_function_restricted:
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40875
diff changeset
   435
  fixes f :: "'a \<Rightarrow> pextreal" assumes "A \<in> sets M"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   436
  shows "simple_function (restricted_space A) f \<longleftrightarrow> simple_function M (\<lambda>x. f x * indicator A x)"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   437
    (is "simple_function ?R f \<longleftrightarrow> simple_function M ?f")
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   438
proof -
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   439
  interpret R: sigma_algebra ?R by (rule restricted_sigma_algebra[OF `A \<in> sets M`])
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   440
  have "finite (f`A) \<longleftrightarrow> finite (?f`space M)"
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   441
  proof cases
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   442
    assume "A = space M"
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   443
    then have "f`A = ?f`space M" by (fastsimp simp: image_iff)
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   444
    then show ?thesis by simp
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   445
  next
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   446
    assume "A \<noteq> space M"
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   447
    then obtain x where x: "x \<in> space M" "x \<notin> A"
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   448
      using sets_into_space `A \<in> sets M` by auto
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   449
    have *: "?f`space M = f`A \<union> {0}"
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   450
    proof (auto simp add: image_iff)
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   451
      show "\<exists>x\<in>space M. f x = 0 \<or> indicator A x = 0"
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   452
        using x by (auto intro!: bexI[of _ x])
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   453
    next
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   454
      fix x assume "x \<in> A"
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   455
      then show "\<exists>y\<in>space M. f x = f y * indicator A y"
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   456
        using `A \<in> sets M` sets_into_space by (auto intro!: bexI[of _ x])
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   457
    next
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   458
      fix x
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40875
diff changeset
   459
      assume "indicator A x \<noteq> (0::pextreal)"
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   460
      then have "x \<in> A" by (auto simp: indicator_def split: split_if_asm)
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   461
      moreover assume "x \<in> space M" "\<forall>y\<in>A. ?f x \<noteq> f y"
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   462
      ultimately show "f x = 0" by auto
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   463
    qed
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   464
    then show ?thesis by auto
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   465
  qed
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   466
  then show ?thesis
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   467
    unfolding simple_function_eq_borel_measurable
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   468
      R.simple_function_eq_borel_measurable
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   469
    unfolding borel_measurable_restricted[OF `A \<in> sets M`]
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   470
    by auto
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   471
qed
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   472
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   473
lemma (in sigma_algebra) simple_function_subalgebra:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   474
  assumes "simple_function N f"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   475
  and N_subalgebra: "sets N \<subseteq> sets M" "space N = space M"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   476
  shows "simple_function M f"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   477
  using assms unfolding simple_function_def by auto
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   478
41661
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41545
diff changeset
   479
lemma (in measure_space) simple_function_vimage:
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41545
diff changeset
   480
  assumes T: "sigma_algebra M'" "T \<in> measurable M M'"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   481
    and f: "simple_function M' f"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   482
  shows "simple_function M (\<lambda>x. f (T x))"
41661
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41545
diff changeset
   483
proof (intro simple_function_def[THEN iffD2] conjI ballI)
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41545
diff changeset
   484
  interpret T: sigma_algebra M' by fact
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41545
diff changeset
   485
  have "(\<lambda>x. f (T x)) ` space M \<subseteq> f ` space M'"
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41545
diff changeset
   486
    using T unfolding measurable_def by auto
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41545
diff changeset
   487
  then show "finite ((\<lambda>x. f (T x)) ` space M)"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   488
    using f unfolding simple_function_def by (auto intro: finite_subset)
41661
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41545
diff changeset
   489
  fix i assume i: "i \<in> (\<lambda>x. f (T x)) ` space M"
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41545
diff changeset
   490
  then have "i \<in> f ` space M'"
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41545
diff changeset
   491
    using T unfolding measurable_def by auto
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41545
diff changeset
   492
  then have "f -` {i} \<inter> space M' \<in> sets M'"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   493
    using f unfolding simple_function_def by auto
41661
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41545
diff changeset
   494
  then have "T -` (f -` {i} \<inter> space M') \<inter> space M \<in> sets M"
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41545
diff changeset
   495
    using T unfolding measurable_def by auto
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41545
diff changeset
   496
  also have "T -` (f -` {i} \<inter> space M') \<inter> space M = (\<lambda>x. f (T x)) -` {i} \<inter> space M"
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41545
diff changeset
   497
    using T unfolding measurable_def by auto
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41545
diff changeset
   498
  finally show "(\<lambda>x. f (T x)) -` {i} \<inter> space M \<in> sets M" .
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   499
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   500
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   501
section "Simple integral"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   502
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   503
definition simple_integral_def:
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   504
  "integral\<^isup>S M f = (\<Sum>x \<in> f ` space M. x * measure M (f -` {x} \<inter> space M))"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   505
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   506
syntax
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   507
  "_simple_integral" :: "'a \<Rightarrow> ('a \<Rightarrow> pextreal) \<Rightarrow> ('a, 'b) measure_space_scheme \<Rightarrow> pextreal" ("\<integral>\<^isup>S _. _ \<partial>_" [60,61] 110)
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   508
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   509
translations
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   510
  "\<integral>\<^isup>S x. f \<partial>M" == "CONST integral\<^isup>S M (%x. f)"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   511
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   512
lemma (in measure_space) simple_integral_cong:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   513
  assumes "\<And>t. t \<in> space M \<Longrightarrow> f t = g t"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   514
  shows "integral\<^isup>S M f = integral\<^isup>S M g"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   515
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   516
  have "f ` space M = g ` space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   517
    "\<And>x. f -` {x} \<inter> space M = g -` {x} \<inter> space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   518
    using assms by (auto intro!: image_eqI)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   519
  thus ?thesis unfolding simple_integral_def by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   520
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   521
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   522
lemma (in measure_space) simple_integral_cong_measure:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   523
  assumes "\<And>A. A \<in> sets M \<Longrightarrow> measure N A = \<mu> A" "sets N = sets M" "space N = space M"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   524
    and "simple_function M f"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   525
  shows "integral\<^isup>S N f = integral\<^isup>S M f"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   526
proof -
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   527
  interpret v: measure_space N
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   528
    by (rule measure_space_cong) fact+
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   529
  from simple_functionD[OF `simple_function M f`] assms show ?thesis
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   530
    by (auto intro!: setsum_cong simp: simple_integral_def)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   531
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   532
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   533
lemma (in measure_space) simple_integral_const[simp]:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   534
  "(\<integral>\<^isup>Sx. c \<partial>M) = c * \<mu> (space M)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   535
proof (cases "space M = {}")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   536
  case True thus ?thesis unfolding simple_integral_def by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   537
next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   538
  case False hence "(\<lambda>x. c) ` space M = {c}" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   539
  thus ?thesis unfolding simple_integral_def by simp
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   540
qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   541
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   542
lemma (in measure_space) simple_function_partition:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   543
  assumes "simple_function M f" and "simple_function M g"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   544
  shows "integral\<^isup>S M f = (\<Sum>A\<in>(\<lambda>x. f -` {f x} \<inter> g -` {g x} \<inter> space M) ` space M. the_elem (f`A) * \<mu> A)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   545
    (is "_ = setsum _ (?p ` space M)")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   546
proof-
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   547
  let "?sub x" = "?p ` (f -` {x} \<inter> space M)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   548
  let ?SIGMA = "Sigma (f`space M) ?sub"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   549
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   550
  have [intro]:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   551
    "finite (f ` space M)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   552
    "finite (g ` space M)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   553
    using assms unfolding simple_function_def by simp_all
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   554
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   555
  { fix A
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   556
    have "?p ` (A \<inter> space M) \<subseteq>
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   557
      (\<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
   558
      by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   559
    hence "finite (?p ` (A \<inter> space M))"
40786
0a54cfc9add3 gave more standard finite set rules simp and intro attribute
nipkow
parents: 39910
diff changeset
   560
      by (rule finite_subset) auto }
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   561
  note this[intro, simp]
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   562
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   563
  { fix x assume "x \<in> space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   564
    have "\<Union>(?sub (f x)) = (f -` {f x} \<inter> space M)" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   565
    moreover {
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   566
      fix x y
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   567
      have *: "f -` {f x} \<inter> g -` {g x} \<inter> space M
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   568
          = (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
   569
      assume "x \<in> space M" "y \<in> space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   570
      hence "f -` {f x} \<inter> g -` {g x} \<inter> space M \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   571
        using assms unfolding simple_function_def * by auto }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   572
    ultimately
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   573
    have "\<mu> (f -` {f x} \<inter> space M) = setsum (\<mu>) (?sub (f x))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   574
      by (subst measure_finitely_additive) auto }
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   575
  hence "integral\<^isup>S M f = (\<Sum>(x,A)\<in>?SIGMA. x * \<mu> A)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   576
    unfolding simple_integral_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   577
    by (subst setsum_Sigma[symmetric],
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   578
       auto intro!: setsum_cong simp: setsum_right_distrib[symmetric])
39910
10097e0a9dbd constant `contents` renamed to `the_elem`
haftmann
parents: 39302
diff changeset
   579
  also have "\<dots> = (\<Sum>A\<in>?p ` space M. the_elem (f`A) * \<mu> A)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   580
  proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   581
    have [simp]: "\<And>x. x \<in> space M \<Longrightarrow> f ` ?p x = {f x}" by (auto intro!: imageI)
39910
10097e0a9dbd constant `contents` renamed to `the_elem`
haftmann
parents: 39302
diff changeset
   582
    have "(\<lambda>A. (the_elem (f ` A), A)) ` ?p ` space M
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   583
      = (\<lambda>x. (f x, ?p x)) ` space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   584
    proof safe
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   585
      fix x assume "x \<in> space M"
39910
10097e0a9dbd constant `contents` renamed to `the_elem`
haftmann
parents: 39302
diff changeset
   586
      thus "(f x, ?p x) \<in> (\<lambda>A. (the_elem (f`A), A)) ` ?p ` space M"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   587
        by (auto intro!: image_eqI[of _ _ "?p x"])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   588
    qed auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   589
    thus ?thesis
39910
10097e0a9dbd constant `contents` renamed to `the_elem`
haftmann
parents: 39302
diff changeset
   590
      apply (auto intro!: setsum_reindex_cong[of "\<lambda>A. (the_elem (f`A), A)"] inj_onI)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   591
      apply (rule_tac x="xa" in image_eqI)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   592
      by simp_all
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   593
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   594
  finally show ?thesis .
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   595
qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   596
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   597
lemma (in measure_space) simple_integral_add[simp]:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   598
  assumes "simple_function M f" and "simple_function M g"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   599
  shows "(\<integral>\<^isup>Sx. f x + g x \<partial>M) = integral\<^isup>S M f + integral\<^isup>S M g"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   600
proof -
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   601
  { fix x let ?S = "g -` {g x} \<inter> f -` {f x} \<inter> space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   602
    assume "x \<in> space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   603
    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
   604
        "(\<lambda>x. (f x, g x)) -` {(f x, g x)} \<inter> space M = ?S"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   605
      by auto }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   606
  thus ?thesis
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   607
    unfolding
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   608
      simple_function_partition[OF simple_function_add[OF assms] simple_function_Pair[OF assms]]
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   609
      simple_function_partition[OF `simple_function M f` `simple_function M g`]
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   610
      simple_function_partition[OF `simple_function M g` `simple_function M f`]
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   611
    apply (subst (3) Int_commute)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   612
    by (auto simp add: field_simps setsum_addf[symmetric] intro!: setsum_cong)
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   613
qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   614
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   615
lemma (in measure_space) simple_integral_setsum[simp]:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   616
  assumes "\<And>i. i \<in> P \<Longrightarrow> simple_function M (f i)"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   617
  shows "(\<integral>\<^isup>Sx. (\<Sum>i\<in>P. f i x) \<partial>M) = (\<Sum>i\<in>P. integral\<^isup>S M (f i))"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   618
proof cases
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   619
  assume "finite P"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   620
  from this assms show ?thesis
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   621
    by induct (auto simp: simple_function_setsum simple_integral_add)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   622
qed auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   623
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   624
lemma (in measure_space) simple_integral_mult[simp]:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   625
  assumes "simple_function M f"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   626
  shows "(\<integral>\<^isup>Sx. c * f x \<partial>M) = c * integral\<^isup>S M f"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   627
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   628
  note mult = simple_function_mult[OF simple_function_const[of c] assms]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   629
  { 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
   630
    assume "x \<in> space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   631
    hence "(\<lambda>x. c * f x) ` ?S = {c * f x}" "f ` ?S = {f x}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   632
      by auto }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   633
  thus ?thesis
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   634
    unfolding simple_function_partition[OF mult assms]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   635
      simple_function_partition[OF assms mult]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   636
    by (auto simp: setsum_right_distrib field_simps intro!: setsum_cong)
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   637
qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   638
40871
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
   639
lemma (in sigma_algebra) simple_function_If:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   640
  assumes sf: "simple_function M f" "simple_function M g" and A: "A \<in> sets M"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   641
  shows "simple_function M (\<lambda>x. if x \<in> A then f x else g x)" (is "simple_function M ?IF")
40871
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
   642
proof -
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
   643
  def F \<equiv> "\<lambda>x. f -` {x} \<inter> space M" and G \<equiv> "\<lambda>x. g -` {x} \<inter> space M"
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
   644
  show ?thesis unfolding simple_function_def
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
   645
  proof safe
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
   646
    have "?IF ` space M \<subseteq> f ` space M \<union> g ` space M" by auto
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
   647
    from finite_subset[OF this] assms
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
   648
    show "finite (?IF ` space M)" unfolding simple_function_def by auto
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
   649
  next
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
   650
    fix x assume "x \<in> space M"
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
   651
    then have *: "?IF -` {?IF x} \<inter> space M = (if x \<in> A
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
   652
      then ((F (f x) \<inter> A) \<union> (G (f x) - (G (f x) \<inter> A)))
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
   653
      else ((F (g x) \<inter> A) \<union> (G (g x) - (G (g x) \<inter> A))))"
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
   654
      using sets_into_space[OF A] by (auto split: split_if_asm simp: G_def F_def)
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
   655
    have [intro]: "\<And>x. F x \<in> sets M" "\<And>x. G x \<in> sets M"
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
   656
      unfolding F_def G_def using sf[THEN simple_functionD(2)] by auto
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
   657
    show "?IF -` {?IF x} \<inter> space M \<in> sets M" unfolding * using A by auto
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
   658
  qed
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
   659
qed
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
   660
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   661
lemma (in measure_space) simple_integral_mono_AE:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   662
  assumes "simple_function M f" and "simple_function M g"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   663
  and mono: "AE x. f x \<le> g x"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   664
  shows "integral\<^isup>S M f \<le> integral\<^isup>S M g"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   665
proof -
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   666
  let "?S x" = "(g -` {g x} \<inter> space M) \<inter> (f -` {f x} \<inter> space M)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   667
  have *: "\<And>x. g -` {g x} \<inter> f -` {f x} \<inter> space M = ?S x"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   668
    "\<And>x. f -` {f x} \<inter> g -` {g x} \<inter> space M = ?S x" by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   669
  show ?thesis
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   670
    unfolding *
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   671
      simple_function_partition[OF `simple_function M f` `simple_function M g`]
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   672
      simple_function_partition[OF `simple_function M g` `simple_function M f`]
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   673
  proof (safe intro!: setsum_mono)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   674
    fix x assume "x \<in> space M"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   675
    then have *: "f ` ?S x = {f x}" "g ` ?S x = {g x}" by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   676
    show "the_elem (f`?S x) * \<mu> (?S x) \<le> the_elem (g`?S x) * \<mu> (?S x)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   677
    proof (cases "f x \<le> g x")
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   678
      case True then show ?thesis using * by (auto intro!: mult_right_mono)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   679
    next
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   680
      case False
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   681
      obtain N where N: "{x\<in>space M. \<not> f x \<le> g x} \<subseteq> N" "N \<in> sets M" "\<mu> N = 0"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   682
        using mono by (auto elim!: AE_E)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   683
      have "?S x \<subseteq> N" using N `x \<in> space M` False by auto
40871
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
   684
      moreover have "?S x \<in> sets M" using assms
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
   685
        by (rule_tac Int) (auto intro!: simple_functionD)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   686
      ultimately have "\<mu> (?S x) \<le> \<mu> N"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   687
        using `N \<in> sets M` by (auto intro!: measure_mono)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   688
      then show ?thesis using `\<mu> N = 0` by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   689
    qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   690
  qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   691
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   692
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   693
lemma (in measure_space) simple_integral_mono:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   694
  assumes "simple_function M f" and "simple_function M g"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   695
  and mono: "\<And> x. x \<in> space M \<Longrightarrow> f x \<le> g x"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   696
  shows "integral\<^isup>S M f \<le> integral\<^isup>S M g"
41705
1100512e16d8 add auto support for AE_mp
hoelzl
parents: 41689
diff changeset
   697
  using assms by (intro simple_integral_mono_AE) auto
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   698
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   699
lemma (in measure_space) simple_integral_cong_AE:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   700
  assumes "simple_function M f" "simple_function M g" and "AE x. f x = g x"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   701
  shows "integral\<^isup>S M f = integral\<^isup>S M g"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   702
  using assms by (auto simp: eq_iff intro!: simple_integral_mono_AE)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   703
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   704
lemma (in measure_space) simple_integral_cong':
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   705
  assumes sf: "simple_function M f" "simple_function M g"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   706
  and mea: "\<mu> {x\<in>space M. f x \<noteq> g x} = 0"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   707
  shows "integral\<^isup>S M f = integral\<^isup>S M g"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   708
proof (intro simple_integral_cong_AE sf AE_I)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   709
  show "\<mu> {x\<in>space M. f x \<noteq> g x} = 0" by fact
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   710
  show "{x \<in> space M. f x \<noteq> g x} \<in> sets M"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   711
    using sf[THEN borel_measurable_simple_function] by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   712
qed simp
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   713
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   714
lemma (in measure_space) simple_integral_indicator:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   715
  assumes "A \<in> sets M"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   716
  assumes "simple_function M f"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   717
  shows "(\<integral>\<^isup>Sx. f x * indicator A x \<partial>M) =
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   718
    (\<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
   719
proof cases
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   720
  assume "A = space M"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   721
  moreover hence "(\<integral>\<^isup>Sx. f x * indicator A x \<partial>M) = integral\<^isup>S M f"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   722
    by (auto intro!: simple_integral_cong)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   723
  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
   724
  ultimately show ?thesis by (simp add: simple_integral_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   725
next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   726
  assume "A \<noteq> space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   727
  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
   728
  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
   729
  proof safe
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   730
    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
   731
  next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   732
    fix y assume "y \<in> A" thus "f y \<in> ?I ` space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   733
      using sets_into_space[OF assms(1)] by (auto intro!: image_eqI[of _ _ y])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   734
  next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   735
    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
   736
  qed
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   737
  have *: "(\<integral>\<^isup>Sx. f x * indicator A x \<partial>M) =
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   738
    (\<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
   739
    unfolding simple_integral_def I
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   740
  proof (rule setsum_mono_zero_cong_left)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   741
    show "finite (f ` space M \<union> {0})"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   742
      using assms(2) unfolding simple_function_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   743
    show "f ` A \<union> {0} \<subseteq> f`space M \<union> {0}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   744
      using sets_into_space[OF assms(1)] by auto
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   745
    have "\<And>x. f x \<notin> f ` A \<Longrightarrow> f -` {f x} \<inter> space M \<inter> A = {}"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   746
      by (auto simp: image_iff)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   747
    thus "\<forall>i\<in>f ` space M \<union> {0} - (f ` A \<union> {0}).
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   748
      i * \<mu> (f -` {i} \<inter> space M \<inter> A) = 0" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   749
  next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   750
    fix x assume "x \<in> f`A \<union> {0}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   751
    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
   752
      by (auto simp: indicator_def split: split_if_asm)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   753
    thus "x * \<mu> (?I -` {x} \<inter> space M) =
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   754
      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
   755
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   756
  show ?thesis unfolding *
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   757
    using assms(2) unfolding simple_function_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   758
    by (auto intro!: setsum_mono_zero_cong_right)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   759
qed
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   760
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   761
lemma (in measure_space) simple_integral_indicator_only[simp]:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   762
  assumes "A \<in> sets M"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   763
  shows "integral\<^isup>S M (indicator A) = \<mu> A"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   764
proof cases
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   765
  assume "space M = {}" hence "A = {}" using sets_into_space[OF assms] by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   766
  thus ?thesis unfolding simple_integral_def using `space M = {}` by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   767
next
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40875
diff changeset
   768
  assume "space M \<noteq> {}" hence "(\<lambda>x. 1) ` space M = {1::pextreal}" by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   769
  thus ?thesis
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   770
    using simple_integral_indicator[OF assms simple_function_const[of 1]]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   771
    using sets_into_space[OF assms]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   772
    by (auto intro!: arg_cong[where f="\<mu>"])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   773
qed
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   774
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   775
lemma (in measure_space) simple_integral_null_set:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   776
  assumes "simple_function M u" "N \<in> null_sets"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   777
  shows "(\<integral>\<^isup>Sx. u x * indicator N x \<partial>M) = 0"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   778
proof -
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40875
diff changeset
   779
  have "AE x. indicator N x = (0 :: pextreal)"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   780
    using `N \<in> null_sets` by (auto simp: indicator_def intro!: AE_I[of _ N])
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   781
  then have "(\<integral>\<^isup>Sx. u x * indicator N x \<partial>M) = (\<integral>\<^isup>Sx. 0 \<partial>M)"
41705
1100512e16d8 add auto support for AE_mp
hoelzl
parents: 41689
diff changeset
   782
    using assms by (intro simple_integral_cong_AE) auto
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   783
  then show ?thesis by simp
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   784
qed
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   785
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   786
lemma (in measure_space) simple_integral_cong_AE_mult_indicator:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   787
  assumes sf: "simple_function M f" and eq: "AE x. x \<in> S" and "S \<in> sets M"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   788
  shows "integral\<^isup>S M f = (\<integral>\<^isup>Sx. f x * indicator S x \<partial>M)"
41705
1100512e16d8 add auto support for AE_mp
hoelzl
parents: 41689
diff changeset
   789
  using assms by (intro simple_integral_cong_AE) auto
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   790
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   791
lemma (in measure_space) simple_integral_restricted:
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   792
  assumes "A \<in> sets M"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   793
  assumes sf: "simple_function M (\<lambda>x. f x * indicator A x)"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   794
  shows "integral\<^isup>S (restricted_space A) f = (\<integral>\<^isup>Sx. f x * indicator A x \<partial>M)"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   795
    (is "_ = integral\<^isup>S M ?f")
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   796
  unfolding simple_integral_def
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   797
proof (simp, safe intro!: setsum_mono_zero_cong_left)
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   798
  from sf show "finite (?f ` space M)"
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   799
    unfolding simple_function_def by auto
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   800
next
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   801
  fix x assume "x \<in> A"
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   802
  then show "f x \<in> ?f ` space M"
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   803
    using sets_into_space `A \<in> sets M` by (auto intro!: image_eqI[of _ _ x])
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   804
next
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   805
  fix x assume "x \<in> space M" "?f x \<notin> f`A"
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   806
  then have "x \<notin> A" by (auto simp: image_iff)
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   807
  then show "?f x * \<mu> (?f -` {?f x} \<inter> space M) = 0" by simp
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   808
next
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   809
  fix x assume "x \<in> A"
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   810
  then have "f x \<noteq> 0 \<Longrightarrow>
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   811
    f -` {f x} \<inter> A = ?f -` {f x} \<inter> space M"
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   812
    using `A \<in> sets M` sets_into_space
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   813
    by (auto simp: indicator_def split: split_if_asm)
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   814
  then show "f x * \<mu> (f -` {f x} \<inter> A) =
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   815
    f x * \<mu> (?f -` {f x} \<inter> space M)"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40875
diff changeset
   816
    unfolding pextreal_mult_cancel_left by auto
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   817
qed
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   818
41545
9c869baf1c66 tuned formalization of subalgebra
hoelzl
parents: 41544
diff changeset
   819
lemma (in measure_space) simple_integral_subalgebra:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   820
  assumes N: "measure_space N" and [simp]: "space N = space M" "measure N = measure M"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   821
  shows "integral\<^isup>S N = integral\<^isup>S M"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   822
  unfolding simple_integral_def_raw by simp
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   823
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   824
lemma (in measure_space) simple_integral_vimage:
41661
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41545
diff changeset
   825
  assumes T: "sigma_algebra M'" "T \<in> measurable M M'"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   826
    and f: "simple_function M' f"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   827
    and \<nu>: "\<And>A. A \<in> sets M' \<Longrightarrow> measure M' A = \<mu> (T -` A \<inter> space M)"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   828
  shows "integral\<^isup>S M' f = (\<integral>\<^isup>S x. f (T x) \<partial>M)"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   829
proof -
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   830
  interpret T: measure_space M' using \<nu> by (rule measure_space_vimage[OF T])
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   831
  show "integral\<^isup>S M' f = (\<integral>\<^isup>S x. f (T x) \<partial>M)"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   832
    unfolding simple_integral_def
41661
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41545
diff changeset
   833
  proof (intro setsum_mono_zero_cong_right ballI)
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41545
diff changeset
   834
    show "(\<lambda>x. f (T x)) ` space M \<subseteq> f ` space M'"
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41545
diff changeset
   835
      using T unfolding measurable_def by auto
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41545
diff changeset
   836
    show "finite (f ` space M')"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   837
      using f unfolding simple_function_def by auto
41661
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41545
diff changeset
   838
  next
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41545
diff changeset
   839
    fix i assume "i \<in> f ` space M' - (\<lambda>x. f (T x)) ` space M"
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41545
diff changeset
   840
    then have "T -` (f -` {i} \<inter> space M') \<inter> space M = {}" by (auto simp: image_iff)
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   841
    with f[THEN T.simple_functionD(2), THEN \<nu>]
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   842
    show "i * T.\<mu> (f -` {i} \<inter> space M') = 0" by simp
41661
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41545
diff changeset
   843
  next
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41545
diff changeset
   844
    fix i assume "i \<in> (\<lambda>x. f (T x)) ` space M"
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41545
diff changeset
   845
    then have "T -` (f -` {i} \<inter> space M') \<inter> space M = (\<lambda>x. f (T x)) -` {i} \<inter> space M"
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41545
diff changeset
   846
      using T unfolding measurable_def by auto
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   847
    with f[THEN T.simple_functionD(2), THEN \<nu>]
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   848
    show "i * T.\<mu> (f -` {i} \<inter> space M') = i * \<mu> ((\<lambda>x. f (T x)) -` {i} \<inter> space M)"
41661
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41545
diff changeset
   849
      by auto
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41545
diff changeset
   850
  qed
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   851
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   852
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   853
section "Continuous positive integration"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   854
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   855
definition positive_integral_def:
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   856
  "integral\<^isup>P M f = (SUP g : {g. simple_function M g \<and> g \<le> f}. integral\<^isup>S M g)"
35692
f1315bbf1bc9 Moved theorems in Lebesgue to the right places
hoelzl
parents: 35582
diff changeset
   857
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   858
syntax
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   859
  "_positive_integral" :: "'a \<Rightarrow> ('a \<Rightarrow> pextreal) \<Rightarrow> ('a, 'b) measure_space_scheme \<Rightarrow> pextreal" ("\<integral>\<^isup>+ _. _ \<partial>_" [60,61] 110)
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   860
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   861
translations
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   862
  "\<integral>\<^isup>+ x. f \<partial>M" == "CONST integral\<^isup>P M (%x. f)"
40872
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
   863
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   864
lemma (in measure_space) positive_integral_alt: "integral\<^isup>P M f =
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   865
    (SUP g : {g. simple_function M g \<and> g \<le> f \<and> \<omega> \<notin> g`space M}. integral\<^isup>S M g)"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   866
  (is "_ = ?alt")
40872
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
   867
proof (rule antisym SUP_leI)
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   868
  show "integral\<^isup>P M f \<le> ?alt" unfolding positive_integral_def
40872
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
   869
  proof (safe intro!: SUP_leI)
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   870
    fix g assume g: "simple_function M g" "g \<le> f"
40872
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
   871
    let ?G = "g -` {\<omega>} \<inter> space M"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   872
    show "integral\<^isup>S M g \<le>
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   873
      (SUP h : {i. simple_function M i \<and> i \<le> f \<and> \<omega> \<notin> i ` space M}. integral\<^isup>S M h)"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   874
      (is "integral\<^isup>S M g \<le> SUPR ?A _")
40872
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
   875
    proof cases
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
   876
      let ?g = "\<lambda>x. indicator (space M - ?G) x * g x"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   877
      have g': "simple_function M ?g"
40872
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
   878
        using g by (auto intro: simple_functionD)
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
   879
      moreover
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
   880
      assume "\<mu> ?G = 0"
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
   881
      then have "AE x. g x = ?g x" using g
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
   882
        by (intro AE_I[where N="?G"])
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
   883
           (auto intro: simple_functionD simp: indicator_def)
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   884
      with g(1) g' have "integral\<^isup>S M g = integral\<^isup>S M ?g"
40872
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
   885
        by (rule simple_integral_cong_AE)
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
   886
      moreover have "?g \<le> g" by (auto simp: le_fun_def indicator_def)
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
   887
      from this `g \<le> f` have "?g \<le> f" by (rule order_trans)
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
   888
      moreover have "\<omega> \<notin> ?g ` space M"
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
   889
        by (auto simp: indicator_def split: split_if_asm)
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
   890
      ultimately show ?thesis by (auto intro!: le_SUPI)
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
   891
    next
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
   892
      assume "\<mu> ?G \<noteq> 0"
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
   893
      then have "?G \<noteq> {}" by auto
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
   894
      then have "\<omega> \<in> g`space M" by force
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
   895
      then have "space M \<noteq> {}" by auto
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   896
      have "SUPR ?A (integral\<^isup>S M) = \<omega>"
40872
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
   897
      proof (intro SUP_\<omega>[THEN iffD2] allI impI)
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
   898
        fix x assume "x < \<omega>"
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
   899
        then guess n unfolding less_\<omega>_Ex_of_nat .. note n = this
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
   900
        then have "0 < n" by (intro neq0_conv[THEN iffD1] notI) simp
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
   901
        let ?g = "\<lambda>x. (of_nat n / (if \<mu> ?G = \<omega> then 1 else \<mu> ?G)) * indicator ?G x"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   902
        show "\<exists>i\<in>?A. x < integral\<^isup>S M i"
40872
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
   903
        proof (intro bexI impI CollectI conjI)
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   904
          show "simple_function M ?g" using g
40872
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
   905
            by (auto intro!: simple_functionD simple_function_add)
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
   906
          have "?g \<le> g" by (auto simp: le_fun_def indicator_def)
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
   907
          from this g(2) show "?g \<le> f" by (rule order_trans)
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
   908
          show "\<omega> \<notin> ?g ` space M"
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
   909
            using `\<mu> ?G \<noteq> 0` by (auto simp: indicator_def split: split_if_asm)
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
   910
          have "x < (of_nat n / (if \<mu> ?G = \<omega> then 1 else \<mu> ?G)) * \<mu> ?G"
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
   911
            using n `\<mu> ?G \<noteq> 0` `0 < n`
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40875
diff changeset
   912
            by (auto simp: pextreal_noteq_omega_Ex field_simps)
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   913
          also have "\<dots> = integral\<^isup>S M ?g" using g `space M \<noteq> {}`
40872
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
   914
            by (subst simple_integral_indicator)
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
   915
               (auto simp: image_constant ac_simps dest: simple_functionD)
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   916
          finally show "x < integral\<^isup>S M ?g" .
40872
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
   917
        qed
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
   918
      qed
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
   919
      then show ?thesis by simp
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
   920
    qed
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   921
  qed
40872
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
   922
qed (auto intro!: SUP_subset simp: positive_integral_def)
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   923
40873
1ef85f4e7097 Shorter definition for positive_integral.
hoelzl
parents: 40872
diff changeset
   924
lemma (in measure_space) positive_integral_cong_measure:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   925
  assumes "\<And>A. A \<in> sets M \<Longrightarrow> measure N A = \<mu> A" "sets N = sets M" "space N = space M"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   926
  shows "integral\<^isup>P N f = integral\<^isup>P M f"
40873
1ef85f4e7097 Shorter definition for positive_integral.
hoelzl
parents: 40872
diff changeset
   927
proof -
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   928
  interpret v: measure_space N
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   929
    by (rule measure_space_cong) fact+
40873
1ef85f4e7097 Shorter definition for positive_integral.
hoelzl
parents: 40872
diff changeset
   930
  with assms show ?thesis
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   931
    unfolding positive_integral_def SUPR_def
40873
1ef85f4e7097 Shorter definition for positive_integral.
hoelzl
parents: 40872
diff changeset
   932
    by (auto intro!: arg_cong[where f=Sup] image_cong
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   933
             simp: simple_integral_cong_measure[OF assms]
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   934
                   simple_function_cong_algebra[OF assms(2,3)])
40873
1ef85f4e7097 Shorter definition for positive_integral.
hoelzl
parents: 40872
diff changeset
   935
qed
1ef85f4e7097 Shorter definition for positive_integral.
hoelzl
parents: 40872
diff changeset
   936
1ef85f4e7097 Shorter definition for positive_integral.
hoelzl
parents: 40872
diff changeset
   937
lemma (in measure_space) positive_integral_alt1:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   938
  "integral\<^isup>P M f =
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   939
    (SUP g : {g. simple_function M g \<and> (\<forall>x\<in>space M. g x \<le> f x \<and> g x \<noteq> \<omega>)}. integral\<^isup>S M g)"
40873
1ef85f4e7097 Shorter definition for positive_integral.
hoelzl
parents: 40872
diff changeset
   940
  unfolding positive_integral_alt SUPR_def
1ef85f4e7097 Shorter definition for positive_integral.
hoelzl
parents: 40872
diff changeset
   941
proof (safe intro!: arg_cong[where f=Sup])
1ef85f4e7097 Shorter definition for positive_integral.
hoelzl
parents: 40872
diff changeset
   942
  fix g let ?g = "\<lambda>x. if x \<in> space M then g x else f x"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   943
  assume "simple_function M g" "\<forall>x\<in>space M. g x \<le> f x \<and> g x \<noteq> \<omega>"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   944
  hence "?g \<le> f" "simple_function M ?g" "integral\<^isup>S M g = integral\<^isup>S M ?g"
40873
1ef85f4e7097 Shorter definition for positive_integral.
hoelzl
parents: 40872
diff changeset
   945
    "\<omega> \<notin> g`space M"
1ef85f4e7097 Shorter definition for positive_integral.
hoelzl
parents: 40872
diff changeset
   946
    unfolding le_fun_def by (auto cong: simple_function_cong simple_integral_cong)
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   947
  thus "integral\<^isup>S M g \<in> integral\<^isup>S M ` {g. simple_function M g \<and> g \<le> f \<and> \<omega> \<notin> g`space M}"
40873
1ef85f4e7097 Shorter definition for positive_integral.
hoelzl
parents: 40872
diff changeset
   948
    by auto
1ef85f4e7097 Shorter definition for positive_integral.
hoelzl
parents: 40872
diff changeset
   949
next
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   950
  fix g assume "simple_function M g" "g \<le> f" "\<omega> \<notin> g`space M"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   951
  hence "simple_function M g" "\<forall>x\<in>space M. g x \<le> f x \<and> g x \<noteq> \<omega>"
40873
1ef85f4e7097 Shorter definition for positive_integral.
hoelzl
parents: 40872
diff changeset
   952
    by (auto simp add: le_fun_def image_iff)
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   953
  thus "integral\<^isup>S M g \<in> integral\<^isup>S M ` {g. simple_function M g \<and> (\<forall>x\<in>space M. g x \<le> f x \<and> g x \<noteq> \<omega>)}"
40873
1ef85f4e7097 Shorter definition for positive_integral.
hoelzl
parents: 40872
diff changeset
   954
    by auto
1ef85f4e7097 Shorter definition for positive_integral.
hoelzl
parents: 40872
diff changeset
   955
qed
1ef85f4e7097 Shorter definition for positive_integral.
hoelzl
parents: 40872
diff changeset
   956
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   957
lemma (in measure_space) positive_integral_cong:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   958
  assumes "\<And>x. x \<in> space M \<Longrightarrow> f x = g x"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   959
  shows "integral\<^isup>P M f = integral\<^isup>P M g"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   960
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   961
  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
   962
    using assms by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   963
  thus ?thesis unfolding positive_integral_alt1 by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   964
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   965
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   966
lemma (in measure_space) positive_integral_eq_simple_integral:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   967
  assumes "simple_function M f"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   968
  shows "integral\<^isup>P M f = integral\<^isup>S M f"
40873
1ef85f4e7097 Shorter definition for positive_integral.
hoelzl
parents: 40872
diff changeset
   969
  unfolding positive_integral_def
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40875
diff changeset
   970
proof (safe intro!: pextreal_SUPI)
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   971
  fix g assume "simple_function M g" "g \<le> f"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   972
  with assms show "integral\<^isup>S M g \<le> integral\<^isup>S M f"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   973
    by (auto intro!: simple_integral_mono simp: le_fun_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   974
next
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   975
  fix y assume "\<forall>x. x\<in>{g. simple_function M g \<and> g \<le> f} \<longrightarrow> integral\<^isup>S M x \<le> y"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   976
  with assms show "integral\<^isup>S M f \<le> y" by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   977
qed
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   978
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   979
lemma (in measure_space) positive_integral_mono_AE:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   980
  assumes ae: "AE x. u x \<le> v x"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   981
  shows "integral\<^isup>P M u \<le> integral\<^isup>P M v"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   982
  unfolding positive_integral_alt1
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   983
proof (safe intro!: SUPR_mono)
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   984
  fix a assume a: "simple_function M a" and mono: "\<forall>x\<in>space M. a x \<le> u x \<and> a x \<noteq> \<omega>"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   985
  from ae obtain N where N: "{x\<in>space M. \<not> u x \<le> v x} \<subseteq> N" "N \<in> sets M" "\<mu> N = 0"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   986
    by (auto elim!: AE_E)
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   987
  have "simple_function M (\<lambda>x. a x * indicator (space M - N) x)"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   988
    using `N \<in> sets M` a by auto
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   989
  with a show "\<exists>b\<in>{g. simple_function M g \<and> (\<forall>x\<in>space M. g x \<le> v x \<and> g x \<noteq> \<omega>)}.
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   990
    integral\<^isup>S M a \<le> integral\<^isup>S M b"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   991
  proof (safe intro!: bexI[of _ "\<lambda>x. a x * indicator (space M - N) x"]
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   992
                      simple_integral_mono_AE)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   993
    show "AE x. a x \<le> a x * indicator (space M - N) x"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   994
    proof (rule AE_I, rule subset_refl)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   995
      have *: "{x \<in> space M. \<not> a x \<le> a x * indicator (space M - N) x} =
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   996
        N \<inter> {x \<in> space M. a x \<noteq> 0}" (is "?N = _")
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   997
        using `N \<in> sets M`[THEN sets_into_space] by (auto simp: indicator_def)
41705
1100512e16d8 add auto support for AE_mp
hoelzl
parents: 41689
diff changeset
   998
      then show "?N \<in> sets M"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   999
        using `N \<in> sets M` `simple_function M a`[THEN borel_measurable_simple_function]
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1000
        by (auto intro!: measure_mono Int)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1001
      then have "\<mu> ?N \<le> \<mu> N"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1002
        unfolding * using `N \<in> sets M` by (auto intro!: measure_mono)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1003
      then show "\<mu> ?N = 0" using `\<mu> N = 0` by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1004
    qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1005
  next
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1006
    fix x assume "x \<in> space M"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1007
    show "a x * indicator (space M - N) x \<le> v x"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1008
    proof (cases "x \<in> N")
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1009
      case True then show ?thesis by simp
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1010
    next
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1011
      case False
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1012
      with N mono have "a x \<le> u x" "u x \<le> v x" using `x \<in> space M` by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1013
      with False `x \<in> space M` show "a x * indicator (space M - N) x \<le> v x" by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1014
    qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1015
    assume "a x * indicator (space M - N) x = \<omega>"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1016
    with mono `x \<in> space M` show False
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1017
      by (simp split: split_if_asm add: indicator_def)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1018
  qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1019
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1020
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1021
lemma (in measure_space) positive_integral_cong_AE:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1022
  "AE x. u x = v x \<Longrightarrow> integral\<^isup>P M u = integral\<^isup>P M v"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1023
  by (auto simp: eq_iff intro!: positive_integral_mono_AE)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1024
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1025
lemma (in measure_space) positive_integral_mono:
41705
1100512e16d8 add auto support for AE_mp
hoelzl
parents: 41689
diff changeset
  1026
  "(\<And>x. x \<in> space M \<Longrightarrow> u x \<le> v x) \<Longrightarrow> integral\<^isup>P M u \<le> integral\<^isup>P M v"
1100512e16d8 add auto support for AE_mp
hoelzl
parents: 41689
diff changeset
  1027
  by (auto intro: positive_integral_mono_AE)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1028
40873
1ef85f4e7097 Shorter definition for positive_integral.
hoelzl
parents: 40872
diff changeset
  1029
lemma image_set_cong:
1ef85f4e7097 Shorter definition for positive_integral.
hoelzl
parents: 40872
diff changeset
  1030
  assumes A: "\<And>x. x \<in> A \<Longrightarrow> \<exists>y\<in>B. f x = g y"
1ef85f4e7097 Shorter definition for positive_integral.
hoelzl
parents: 40872
diff changeset
  1031
  assumes B: "\<And>y. y \<in> B \<Longrightarrow> \<exists>x\<in>A. g y = f x"
1ef85f4e7097 Shorter definition for positive_integral.
hoelzl
parents: 40872
diff changeset
  1032
  shows "f ` A = g ` B"
1ef85f4e7097 Shorter definition for positive_integral.
hoelzl
parents: 40872
diff changeset
  1033
  using assms by blast
1ef85f4e7097 Shorter definition for positive_integral.
hoelzl
parents: 40872
diff changeset
  1034
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1035
lemma (in measure_space) positive_integral_SUP_approx:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1036
  assumes "f \<up> s"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1037
  and f: "\<And>i. f i \<in> borel_measurable M"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1038
  and "simple_function M u"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1039
  and le: "u \<le> s" and real: "\<omega> \<notin> u`space M"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1040
  shows "integral\<^isup>S M u \<le> (SUP i. integral\<^isup>P M (f i))" (is "_ \<le> ?S")
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40875
diff changeset
  1041
proof (rule pextreal_le_mult_one_interval)
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40875
diff changeset
  1042
  fix a :: pextreal assume "0 < a" "a < 1"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1043
  hence "a \<noteq> 0" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1044
  let "?B i" = "{x \<in> space M. a * u x \<le> f i x}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1045
  have B: "\<And>i. ?B i \<in> sets M"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1046
    using f `simple_function M u` by (auto simp: borel_measurable_simple_function)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1047
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1048
  let "?uB i x" = "u x * indicator (?B i) x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1049
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1050
  { fix i have "?B i \<subseteq> ?B (Suc i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1051
    proof safe
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1052
      fix i x assume "a * u x \<le> f i x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1053
      also have "\<dots> \<le> f (Suc i) x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1054
        using `f \<up> s` unfolding isoton_def le_fun_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1055
      finally show "a * u x \<le> f (Suc i) x" .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1056
    qed }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1057
  note B_mono = this
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1058
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1059
  have u: "\<And>x. x \<in> space M \<Longrightarrow> u -` {u x} \<inter> space M \<in> sets M"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1060
    using `simple_function M u` by (auto simp add: simple_function_def)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1061
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1062
  have "\<And>i. (\<lambda>n. (u -` {i} \<inter> space M) \<inter> ?B n) \<up> (u -` {i} \<inter> space M)" using B_mono unfolding isoton_def
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1063
  proof safe
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1064
    fix x i assume "x \<in> space M"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1065
    show "x \<in> (\<Union>i. (u -` {u x} \<inter> space M) \<inter> ?B i)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1066
    proof cases
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1067
      assume "u x = 0" thus ?thesis using `x \<in> space M` by simp
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1068
    next
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1069
      assume "u x \<noteq> 0"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1070
      with `a < 1` real `x \<in> space M`
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40875
diff changeset
  1071
      have "a * u x < 1 * u x" by (rule_tac pextreal_mult_strict_right_mono) (auto simp: image_iff)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1072
      also have "\<dots> \<le> (SUP i. f i x)" using le `f \<up> s`
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1073
        unfolding isoton_fun_expand by (auto simp: isoton_def le_fun_def)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1074
      finally obtain i where "a * u x < f i x" unfolding SUPR_def
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1075
        by (auto simp add: less_Sup_iff)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1076
      hence "a * u x \<le> f i x" by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1077
      thus ?thesis using `x \<in> space M` by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1078
    qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1079
  qed auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1080
  note measure_conv = measure_up[OF Int[OF u B] this]
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1081
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1082
  have "integral\<^isup>S M u = (SUP i. integral\<^isup>S M (?uB i))"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1083
    unfolding simple_integral_indicator[OF B `simple_function M u`]
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40875
diff changeset
  1084
  proof (subst SUPR_pextreal_setsum, safe)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1085
    fix x n assume "x \<in> space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1086
    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
  1087
      \<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
  1088
      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
  1089
    thus "u x * \<mu> (u -` {u x} \<inter> space M \<inter> ?B n)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1090
            \<le> u x * \<mu> (u -` {u x} \<inter> space M \<inter> ?B (Suc n))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1091
      by (auto intro: mult_left_mono)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1092
  next
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1093
    show "integral\<^isup>S M u =
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1094
      (\<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
  1095
      using measure_conv unfolding simple_integral_def isoton_def
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40875
diff changeset
  1096
      by (auto intro!: setsum_cong simp: pextreal_SUP_cmult)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1097
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1098
  moreover
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1099
  have "a * (SUP i. integral\<^isup>S M (?uB i)) \<le> ?S"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40875
diff changeset
  1100
    unfolding pextreal_SUP_cmult[symmetric]
38705
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 38656
diff changeset
  1101
  proof (safe intro!: SUP_mono bexI)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1102
    fix i
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1103
    have "a * integral\<^isup>S M (?uB i) = (\<integral>\<^isup>Sx. a * ?uB i x \<partial>M)"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1104
      using B `simple_function M u`
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1105
      by (subst simple_integral_mult[symmetric]) (auto simp: field_simps)
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1106
    also have "\<dots> \<le> integral\<^isup>P M (f i)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1107
    proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1108
      have "\<And>x. a * ?uB i x \<le> f i x" unfolding indicator_def by auto
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1109
      hence *: "simple_function M (\<lambda>x. a * ?uB i x)" using B assms(3)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1110
        by (auto intro!: simple_integral_mono)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1111
      show ?thesis unfolding positive_integral_eq_simple_integral[OF *, symmetric]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1112
        by (auto intro!: positive_integral_mono simp: indicator_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1113
    qed
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1114
    finally show "a * integral\<^isup>S M (?uB i) \<le> integral\<^isup>P M (f i)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1115
      by auto
38705
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 38656
diff changeset
  1116
  qed simp
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1117
  ultimately show "a * integral\<^isup>S M u \<le> ?S" by simp
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1118
qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1119
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1120
text {* Beppo-Levi monotone convergence theorem *}
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1121
lemma (in measure_space) positive_integral_isoton:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1122
  assumes "f \<up> u" "\<And>i. f i \<in> borel_measurable M"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1123
  shows "(\<lambda>i. integral\<^isup>P M (f i)) \<up> integral\<^isup>P M u"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1124
  unfolding isoton_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1125
proof safe
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1126
  fix i show "integral\<^isup>P M (f i) \<le> integral\<^isup>P M (f (Suc i))"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1127
    apply (rule positive_integral_mono)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1128
    using `f \<up> u` unfolding isoton_def le_fun_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1129
next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1130
  have u: "u = (SUP i. f i)" using `f \<up> u` unfolding isoton_def by auto
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1131
  show "(SUP i. integral\<^isup>P M (f i)) = integral\<^isup>P M u"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1132
  proof (rule antisym)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1133
    from `f \<up> u`[THEN isoton_Sup, unfolded le_fun_def]
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1134
    show "(SUP j. integral\<^isup>P M (f j)) \<le> integral\<^isup>P M u"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1135
      by (auto intro!: SUP_leI positive_integral_mono)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1136
  next
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1137
    show "integral\<^isup>P M u \<le> (SUP i. integral\<^isup>P M (f i))"
40873
1ef85f4e7097 Shorter definition for positive_integral.
hoelzl
parents: 40872
diff changeset
  1138
      unfolding positive_integral_alt[of u]
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1139
      by (auto intro!: SUP_leI positive_integral_SUP_approx[OF assms])
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1140
  qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1141
qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1142
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1143
lemma (in measure_space) positive_integral_monotone_convergence_SUP:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1144
  assumes "\<And>i x. x \<in> space M \<Longrightarrow> f i x \<le> f (Suc i) x"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1145
  assumes "\<And>i. f i \<in> borel_measurable M"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1146
  shows "(SUP i. integral\<^isup>P M (f i)) = (\<integral>\<^isup>+ x. (SUP i. f i x) \<partial>M)"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1147
    (is "_ = integral\<^isup>P M ?u")
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1148
proof -
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1149
  show ?thesis
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1150
  proof (rule antisym)
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1151
    show "(SUP j. integral\<^isup>P M (f j)) \<le> integral\<^isup>P M ?u"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1152
      by (auto intro!: SUP_leI positive_integral_mono le_SUPI)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1153
  next
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1154
    def rf \<equiv> "\<lambda>i. \<lambda>x\<in>space M. f i x" and ru \<equiv> "\<lambda>x\<in>space M. ?u x"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1155
    have "\<And>i. rf i \<in> borel_measurable M" unfolding rf_def
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1156
      using assms by (simp cong: measurable_cong)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1157
    moreover have iso: "rf \<up> ru" using assms unfolding rf_def ru_def
41097
a1abfa4e2b44 use SUPR_ and INFI_apply instead of SUPR_, INFI_fun_expand
hoelzl
parents: 41096
diff changeset
  1158
      unfolding isoton_def le_fun_def fun_eq_iff SUPR_apply
40872
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
  1159
      using SUP_const[OF UNIV_not_empty]
41097
a1abfa4e2b44 use SUPR_ and INFI_apply instead of SUPR_, INFI_fun_expand
hoelzl
parents: 41096
diff changeset
  1160
      by (auto simp: restrict_def le_fun_def fun_eq_iff)
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1161
    ultimately have "integral\<^isup>P M ru \<le> (SUP i. integral\<^isup>P M (rf i))"
40873
1ef85f4e7097 Shorter definition for positive_integral.
hoelzl
parents: 40872
diff changeset
  1162
      unfolding positive_integral_alt[of ru]
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1163
      by (auto simp: le_fun_def intro!: SUP_leI positive_integral_SUP_approx)
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1164
    then show "integral\<^isup>P M ?u \<le> (SUP i. integral\<^isup>P M (f i))"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1165
      unfolding ru_def rf_def by (simp cong: positive_integral_cong)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1166
  qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1167
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1168
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1169
lemma (in measure_space) SUP_simple_integral_sequences:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1170
  assumes f: "f \<up> u" "\<And>i. simple_function M (f i)"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1171
  and g: "g \<up> u" "\<And>i. simple_function M (g i)"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1172
  shows "(SUP i. integral\<^isup>S M (f i)) = (SUP i. integral\<^isup>S M (g i))"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1173
    (is "SUPR _ ?F = SUPR _ ?G")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1174
proof -
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1175
  have "(SUP i. ?F i) = (SUP i. integral\<^isup>P M (f i))"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1176
    using assms by (simp add: positive_integral_eq_simple_integral)
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1177
  also have "\<dots> = integral\<^isup>P M u"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1178
    using positive_integral_isoton[OF f(1) borel_measurable_simple_function[OF f(2)]]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1179
    unfolding isoton_def by simp
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1180
  also have "\<dots> = (SUP i. integral\<^isup>P M (g i))"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1181
    using positive_integral_isoton[OF g(1) borel_measurable_simple_function[OF g(2)]]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1182
    unfolding isoton_def by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1183
  also have "\<dots> = (SUP i. ?G i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1184
    using assms by (simp add: positive_integral_eq_simple_integral)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1185
  finally show ?thesis .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1186
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1187
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1188
lemma (in measure_space) positive_integral_const[simp]:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1189
  "(\<integral>\<^isup>+ x. c \<partial>M) = c * \<mu> (space M)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1190
  by (subst positive_integral_eq_simple_integral) auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1191
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1192
lemma (in measure_space) positive_integral_isoton_simple:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1193
  assumes "f \<up> u" and e: "\<And>i. simple_function M (f i)"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1194
  shows "(\<lambda>i. integral\<^isup>S M (f i)) \<up> integral\<^isup>P M u"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1195
  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
  1196
  unfolding positive_integral_eq_simple_integral[OF e] .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1197
41661
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41545
diff changeset
  1198
lemma (in measure_space) positive_integral_vimage:
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41545
diff changeset
  1199
  assumes T: "sigma_algebra M'" "T \<in> measurable M M'" and f: "f \<in> borel_measurable M'"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1200
    and \<nu>: "\<And>A. A \<in> sets M' \<Longrightarrow> measure M' A = \<mu> (T -` A \<inter> space M)"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1201
  shows "integral\<^isup>P M' f = (\<integral>\<^isup>+ x. f (T x) \<partial>M)"
41661
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41545
diff changeset
  1202
proof -
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1203
  interpret T: measure_space M' using \<nu> by (rule measure_space_vimage[OF T])
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1204
  obtain f' where f': "f' \<up> f" "\<And>i. simple_function M' (f' i)"
41661
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41545
diff changeset
  1205
    using T.borel_measurable_implies_simple_function_sequence[OF f] by blast
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1206
  then have f: "(\<lambda>i x. f' i (T x)) \<up> (\<lambda>x. f (T x))" "\<And>i. simple_function M (\<lambda>x. f' i (T x))"
41661
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41545
diff changeset
  1207
    using simple_function_vimage[OF T] unfolding isoton_fun_expand by auto
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1208
  show "integral\<^isup>P M' f = (\<integral>\<^isup>+ x. f (T x) \<partial>M)"
41661
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41545
diff changeset
  1209
    using positive_integral_isoton_simple[OF f]
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41545
diff changeset
  1210
    using T.positive_integral_isoton_simple[OF f']
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1211
    by (simp add: simple_integral_vimage[OF T f'(2) \<nu>] isoton_def)
41661
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41545
diff changeset
  1212
qed
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41545
diff changeset
  1213
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1214
lemma (in measure_space) positive_integral_linear:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1215
  assumes f: "f \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1216
  and g: "g \<in> borel_measurable M"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1217
  shows "(\<integral>\<^isup>+ x. a * f x + g x \<partial>M) = a * integral\<^isup>P M f + integral\<^isup>P M g"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1218
    (is "integral\<^isup>P M ?L = _")
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1219
proof -
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1220
  from borel_measurable_implies_simple_function_sequence'[OF f] guess u .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1221
  note u = this positive_integral_isoton_simple[OF this(1-2)]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1222
  from borel_measurable_implies_simple_function_sequence'[OF g] guess v .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1223
  note v = this positive_integral_isoton_simple[OF this(1-2)]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1224
  let "?L' i x" = "a * u i x + v i x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1225
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1226
  have "?L \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1227
    using assms by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1228
  from borel_measurable_implies_simple_function_sequence'[OF this] guess l .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1229
  note positive_integral_isoton_simple[OF this(1-2)] and l = this
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1230
  moreover have "(SUP i. integral\<^isup>S M (l i)) = (SUP i. integral\<^isup>S M (?L' i))"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1231
  proof (rule SUP_simple_integral_sequences[OF l(1-2)])
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1232
    show "?L' \<up> ?L" "\<And>i. simple_function M (?L' i)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1233
      using u v by (auto simp: isoton_fun_expand isoton_add isoton_cmult_right)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1234
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1235
  moreover from u v have L'_isoton:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1236
      "(\<lambda>i. integral\<^isup>S M (?L' i)) \<up> a * integral\<^isup>P M f + integral\<^isup>P M g"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1237
    by (simp add: isoton_add isoton_cmult_right)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1238
  ultimately show ?thesis by (simp add: isoton_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1239
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1240
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1241
lemma (in measure_space) positive_integral_cmult:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1242
  assumes "f \<in> borel_measurable M"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1243
  shows "(\<integral>\<^isup>+ x. c * f x \<partial>M) = c * integral\<^isup>P M f"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1244
  using positive_integral_linear[OF assms, of "\<lambda>x. 0" c] by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1245
41096
843c40bbc379 integral over setprod
hoelzl
parents: 41095
diff changeset
  1246
lemma (in measure_space) positive_integral_multc:
843c40bbc379 integral over setprod
hoelzl
parents: 41095
diff changeset
  1247
  assumes "f \<in> borel_measurable M"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1248
  shows "(\<integral>\<^isup>+ x. f x * c \<partial>M) = integral\<^isup>P M f * c"
41096
843c40bbc379 integral over setprod
hoelzl
parents: 41095
diff changeset
  1249
  unfolding mult_commute[of _ c] positive_integral_cmult[OF assms] by simp
843c40bbc379 integral over setprod
hoelzl
parents: 41095
diff changeset
  1250
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1251
lemma (in measure_space) positive_integral_indicator[simp]:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1252
  "A \<in> sets M \<Longrightarrow> (\<integral>\<^isup>+ x. indicator A x\<partial>M) = \<mu> A"
41544
c3b977fee8a3 introduced integral syntax
hoelzl
parents: 41097
diff changeset
  1253
  by (subst positive_integral_eq_simple_integral)
c3b977fee8a3 introduced integral syntax
hoelzl
parents: 41097
diff changeset
  1254
     (auto simp: simple_function_indicator simple_integral_indicator)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1255
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1256
lemma (in measure_space) positive_integral_cmult_indicator:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1257
  "A \<in> sets M \<Longrightarrow> (\<integral>\<^isup>+ x. c * indicator A x \<partial>M) = c * \<mu> A"
41544
c3b977fee8a3 introduced integral syntax
hoelzl
parents: 41097
diff changeset
  1258
  by (subst positive_integral_eq_simple_integral)
c3b977fee8a3 introduced integral syntax
hoelzl
parents: 41097
diff changeset
  1259
     (auto simp: simple_function_indicator simple_integral_indicator)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1260
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1261
lemma (in measure_space) positive_integral_add:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1262
  assumes "f \<in> borel_measurable M" "g \<in> borel_measurable M"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1263
  shows "(\<integral>\<^isup>+ x. f x + g x \<partial>M) = integral\<^isup>P M f + integral\<^isup>P M g"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1264
  using positive_integral_linear[OF assms, of 1] by simp
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) positive_integral_setsum:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1267
  assumes "\<And>i. i\<in>P \<Longrightarrow> f i \<in> borel_measurable M"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1268
  shows "(\<integral>\<^isup>+ x. (\<Sum>i\<in>P. f i x) \<partial>M) = (\<Sum>i\<in>P. integral\<^isup>P M (f i))"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1269
proof cases
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1270
  assume "finite P"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1271
  from this assms show ?thesis
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1272
  proof induct
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1273
    case (insert i P)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1274
    have "f i \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1275
      "(\<lambda>x. \<Sum>i\<in>P. f i x) \<in> borel_measurable M"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40875
diff changeset
  1276
      using insert by (auto intro!: borel_measurable_pextreal_setsum)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1277
    from positive_integral_add[OF this]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1278
    show ?case using insert by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1279
  qed simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1280
qed simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1281
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1282
lemma (in measure_space) positive_integral_diff:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1283
  assumes f: "f \<in> borel_measurable M" and g: "g \<in> borel_measurable M"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1284
  and fin: "integral\<^isup>P M g \<noteq> \<omega>"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1285
  and mono: "\<And>x. x \<in> space M \<Longrightarrow> g x \<le> f x"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1286
  shows "(\<integral>\<^isup>+ x. f x - g x \<partial>M) = integral\<^isup>P M f - integral\<^isup>P M g"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1287
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1288
  have borel: "(\<lambda>x. f x - g x) \<in> borel_measurable M"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40875
diff changeset
  1289
    using f g by (rule borel_measurable_pextreal_diff)
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1290
  have "(\<integral>\<^isup>+x. f x - g x \<partial>M) + integral\<^isup>P M g = integral\<^isup>P M f"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1291
    unfolding positive_integral_add[OF borel g, symmetric]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1292
  proof (rule positive_integral_cong)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1293
    fix x assume "x \<in> space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1294
    from mono[OF this] show "f x - g x + g x = f x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1295
      by (cases "f x", cases "g x", simp, simp, cases "g x", auto)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1296
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1297
  with mono show ?thesis
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40875
diff changeset
  1298
    by (subst minus_pextreal_eq2[OF _ fin]) (auto intro!: positive_integral_mono)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1299
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1300
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1301
lemma (in measure_space) positive_integral_psuminf:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1302
  assumes "\<And>i. f i \<in> borel_measurable M"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1303
  shows "(\<integral>\<^isup>+ x. (\<Sum>\<^isub>\<infinity> i. f i x) \<partial>M) = (\<Sum>\<^isub>\<infinity> i. integral\<^isup>P M (f i))"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1304
proof -
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1305
  have "(\<lambda>i. (\<integral>\<^isup>+x. (\<Sum>i<i. f i x) \<partial>M)) \<up> (\<integral>\<^isup>+x. (\<Sum>\<^isub>\<infinity>i. f i x) \<partial>M)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1306
    by (rule positive_integral_isoton)
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40875
diff changeset
  1307
       (auto intro!: borel_measurable_pextreal_setsum assms positive_integral_mono
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1308
                     arg_cong[where f=Sup]
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 39198
diff changeset
  1309
             simp: isoton_def le_fun_def psuminf_def fun_eq_iff SUPR_def Sup_fun_def)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1310
  thus ?thesis
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1311
    by (auto simp: isoton_def psuminf_def positive_integral_setsum[OF assms])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1312
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1313
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1314
text {* Fatou's lemma: convergence theorem on limes inferior *}
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1315
lemma (in measure_space) positive_integral_lim_INF:
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40875
diff changeset
  1316
  fixes u :: "nat \<Rightarrow> 'a \<Rightarrow> pextreal"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1317
  assumes "\<And>i. u i \<in> borel_measurable M"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1318
  shows "(\<integral>\<^isup>+ x. (SUP n. INF m. u (m + n) x) \<partial>M) \<le>
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1319
    (SUP n. INF m. integral\<^isup>P M (u (m + n)))"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1320
proof -
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1321
  have "(\<integral>\<^isup>+x. (SUP n. INF m. u (m + n) x) \<partial>M)
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1322
      = (SUP n. (\<integral>\<^isup>+x. (INF m. u (m + n) x) \<partial>M))"
41097
a1abfa4e2b44 use SUPR_ and INFI_apply instead of SUPR_, INFI_fun_expand
hoelzl
parents: 41096
diff changeset
  1323
    using assms
a1abfa4e2b44 use SUPR_ and INFI_apply instead of SUPR_, INFI_fun_expand
hoelzl
parents: 41096
diff changeset
  1324
    by (intro positive_integral_monotone_convergence_SUP[symmetric] INF_mono)
a1abfa4e2b44 use SUPR_ and INFI_apply instead of SUPR_, INFI_fun_expand
hoelzl
parents: 41096
diff changeset
  1325
       (auto simp del: add_Suc simp add: add_Suc[symmetric])
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1326
  also have "\<dots> \<le> (SUP n. INF m. integral\<^isup>P M (u (m + n)))"
41097
a1abfa4e2b44 use SUPR_ and INFI_apply instead of SUPR_, INFI_fun_expand
hoelzl
parents: 41096
diff changeset
  1327
    by (auto intro!: SUP_mono bexI le_INFI positive_integral_mono INF_leI)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1328
  finally show ?thesis .
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1329
qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1330
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1331
lemma (in measure_space) measure_space_density:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1332
  assumes borel: "u \<in> borel_measurable M"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1333
    and M'[simp]: "M' = (M\<lparr>measure := \<lambda>A. (\<integral>\<^isup>+ x. u x * indicator A x \<partial>M)\<rparr>)"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1334
  shows "measure_space M'"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1335
proof -
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1336
  interpret M': sigma_algebra M' by (intro sigma_algebra_cong) auto
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1337
  show ?thesis
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1338
  proof
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1339
    show "measure M' {} = 0" unfolding M' by simp
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1340
    show "countably_additive M' (measure M')"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1341
    proof (intro countably_additiveI)
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1342
      fix A :: "nat \<Rightarrow> 'a set" assume "range A \<subseteq> sets M'"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1343
      then have "\<And>i. (\<lambda>x. u x * indicator (A i) x) \<in> borel_measurable M"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1344
        using borel by (auto intro: borel_measurable_indicator)
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1345
      moreover assume "disjoint_family A"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1346
      note psuminf_indicator[OF this]
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1347
      ultimately show "(\<Sum>\<^isub>\<infinity>n. measure M' (A n)) = measure M' (\<Union>x. A x)"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1348
        by (simp add: positive_integral_psuminf[symmetric])
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1349
    qed
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1350
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1351
qed
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1352
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
  1353
lemma (in measure_space) positive_integral_translated_density:
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
  1354
  assumes "f \<in> borel_measurable M" "g \<in> borel_measurable M"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1355
    and M': "M' = (M\<lparr> measure := \<lambda>A. (\<integral>\<^isup>+ x. f x * indicator A x \<partial>M)\<rparr>)"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1356
  shows "integral\<^isup>P M' g = (\<integral>\<^isup>+ x. f x * g x \<partial>M)"
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
  1357
proof -
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1358
  from measure_space_density[OF assms(1) M']
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1359
  interpret T: measure_space M' .
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1360
  have borel[simp]:
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1361
    "borel_measurable M' = borel_measurable M"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1362
    "simple_function M' = simple_function M"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1363
    unfolding measurable_def simple_function_def_raw by (auto simp: M')
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
  1364
  from borel_measurable_implies_simple_function_sequence[OF assms(2)]
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1365
  obtain G where G: "\<And>i. simple_function M (G i)" "G \<up> g" by blast
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
  1366
  note G_borel = borel_measurable_simple_function[OF this(1)]
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1367
  from T.positive_integral_isoton[unfolded borel, OF `G \<up> g` G_borel]
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1368
  have *: "(\<lambda>i. integral\<^isup>P M' (G i)) \<up> integral\<^isup>P M' g" .
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
  1369
  { fix i
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
  1370
    have [simp]: "finite (G i ` space M)"
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
  1371
      using G(1) unfolding simple_function_def by auto
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1372
    have "integral\<^isup>P M' (G i) = integral\<^isup>S M' (G i)"
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
  1373
      using G T.positive_integral_eq_simple_integral by simp
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1374
    also have "\<dots> = (\<integral>\<^isup>+x. f x * (\<Sum>y\<in>G i`space M. y * indicator (G i -` {y} \<inter> space M) x) \<partial>M)"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1375
      apply (simp add: simple_integral_def M')
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
  1376
      apply (subst positive_integral_cmult[symmetric])
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1377
      using G_borel assms(1) apply (fastsimp intro: borel_measurable_vimage)
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
  1378
      apply (subst positive_integral_setsum[symmetric])
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1379
      using G_borel assms(1) apply (fastsimp intro: borel_measurable_vimage)
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
  1380
      by (simp add: setsum_right_distrib field_simps)
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1381
    also have "\<dots> = (\<integral>\<^isup>+x. f x * G i x \<partial>M)"
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
  1382
      by (auto intro!: positive_integral_cong
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
  1383
               simp: indicator_def if_distrib setsum_cases)
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1384
    finally have "integral\<^isup>P M' (G i) = (\<integral>\<^isup>+x. f x * G i x \<partial>M)" . }
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1385
  with * have eq_Tg: "(\<lambda>i. (\<integral>\<^isup>+x. f x * G i x \<partial>M)) \<up> integral\<^isup>P M' g" by simp
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
  1386
  from G(2) have "(\<lambda>i x. f x * G i x) \<up> (\<lambda>x. f x * g x)"
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
  1387
    unfolding isoton_fun_expand by (auto intro!: isoton_cmult_right)
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1388
  then have "(\<lambda>i. (\<integral>\<^isup>+x. f x * G i x \<partial>M)) \<up> (\<integral>\<^isup>+x. f x * g x \<partial>M)"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40875
diff changeset
  1389
    using assms(1) G_borel by (auto intro!: positive_integral_isoton borel_measurable_pextreal_times)
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1390
  with eq_Tg show "integral\<^isup>P M' g = (\<integral>\<^isup>+x. f x * g x \<partial>M)"
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
  1391
    unfolding isoton_def by simp
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
  1392
qed
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
  1393
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1394
lemma (in measure_space) positive_integral_null_set:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1395
  assumes "N \<in> null_sets" shows "(\<integral>\<^isup>+ x. u x * indicator N x \<partial>M) = 0"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1396
proof -
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1397
  have "(\<integral>\<^isup>+ x. u x * indicator N x \<partial>M) = (\<integral>\<^isup>+ x. 0 \<partial>M)"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1398
  proof (intro positive_integral_cong_AE AE_I)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1399
    show "{x \<in> space M. u x * indicator N x \<noteq> 0} \<subseteq> N"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1400
      by (auto simp: indicator_def)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1401
    show "\<mu> N = 0" "N \<in> sets M"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1402
      using assms by auto
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1403
  qed
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1404
  then show ?thesis by simp
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1405
qed
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1406
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1407
lemma (in measure_space) positive_integral_Markov_inequality:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1408
  assumes borel: "u \<in> borel_measurable M" and "A \<in> sets M" and c: "c \<noteq> \<omega>"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1409
  shows "\<mu> ({x\<in>space M. 1 \<le> c * u x} \<inter> A) \<le> c * (\<integral>\<^isup>+ x. u x * indicator A x \<partial>M)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1410
    (is "\<mu> ?A \<le> _ * ?PI")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1411
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1412
  have "?A \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1413
    using `A \<in> sets M` borel by auto
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1414
  hence "\<mu> ?A = (\<integral>\<^isup>+ x. indicator ?A x \<partial>M)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1415
    using positive_integral_indicator by simp
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1416
  also have "\<dots> \<le> (\<integral>\<^isup>+ x. c * (u x * indicator A x) \<partial>M)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1417
  proof (rule positive_integral_mono)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1418
    fix x assume "x \<in> space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1419
    show "indicator ?A x \<le> c * (u x * indicator A x)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1420
      by (cases "x \<in> ?A") auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1421
  qed
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1422
  also have "\<dots> = c * (\<integral>\<^isup>+ x. u x * indicator A x \<partial>M)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1423
    using assms
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1424
    by (auto intro!: positive_integral_cmult borel_measurable_indicator)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1425
  finally show ?thesis .
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1426
qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1427
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1428
lemma (in measure_space) positive_integral_0_iff:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1429
  assumes borel: "u \<in> borel_measurable M"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1430
  shows "integral\<^isup>P M u = 0 \<longleftrightarrow> \<mu> {x\<in>space M. u x \<noteq> 0} = 0"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1431
    (is "_ \<longleftrightarrow> \<mu> ?A = 0")
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1432
proof -
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1433
  have A: "?A \<in> sets M" using borel by auto
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1434
  have u: "(\<integral>\<^isup>+ x. u x * indicator ?A x \<partial>M) = integral\<^isup>P M u"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1435
    by (auto intro!: positive_integral_cong simp: indicator_def)
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1436
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1437
  show ?thesis
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1438
  proof
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1439
    assume "\<mu> ?A = 0"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1440
    hence "?A \<in> null_sets" using `?A \<in> sets M` by auto
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1441
    from positive_integral_null_set[OF this]
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1442
    have "0 = (\<integral>\<^isup>+ x. u x * indicator ?A x \<partial>M)" by simp
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1443
    thus "integral\<^isup>P M u = 0" unfolding u by simp
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1444
  next
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1445
    assume *: "integral\<^isup>P M u = 0"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1446
    let "?M n" = "{x \<in> space M. 1 \<le> of_nat n * u x}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1447
    have "0 = (SUP n. \<mu> (?M n \<inter> ?A))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1448
    proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1449
      { fix n
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1450
        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
  1451
        have "\<mu> (?M n \<inter> ?A) = 0" unfolding * u by simp }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1452
      thus ?thesis by simp
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1453
    qed
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1454
    also have "\<dots> = \<mu> (\<Union>n. ?M n \<inter> ?A)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1455
    proof (safe intro!: continuity_from_below)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1456
      fix n show "?M n \<inter> ?A \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1457
        using borel by (auto intro!: Int)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1458
    next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1459
      fix n x assume "1 \<le> of_nat n * u x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1460
      also have "\<dots> \<le> of_nat (Suc n) * u x"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40875
diff changeset
  1461
        by (subst (1 2) mult_commute) (auto intro!: pextreal_mult_cancel)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1462
      finally show "1 \<le> of_nat (Suc n) * u x" .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1463
    qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1464
    also have "\<dots> = \<mu> ?A"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1465
    proof (safe intro!: arg_cong[where f="\<mu>"])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1466
      fix x assume "u x \<noteq> 0" and [simp, intro]: "x \<in> space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1467
      show "x \<in> (\<Union>n. ?M n \<inter> ?A)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1468
      proof (cases "u x")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1469
        case (preal r)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1470
        obtain j where "1 / r \<le> of_nat j" using ex_le_of_nat ..
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1471
        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
  1472
        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
  1473
        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
  1474
      qed auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1475
    qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1476
    finally show "\<mu> ?A = 0" by simp
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1477
  qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1478
qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1479
41705
1100512e16d8 add auto support for AE_mp
hoelzl
parents: 41689
diff changeset
  1480
lemma (in measure_space) positive_integral_0_iff_AE:
1100512e16d8 add auto support for AE_mp
hoelzl
parents: 41689
diff changeset
  1481
  assumes u: "u \<in> borel_measurable M"
1100512e16d8 add auto support for AE_mp
hoelzl
parents: 41689
diff changeset
  1482
  shows "integral\<^isup>P M u = 0 \<longleftrightarrow> (AE x. u x = 0)"
1100512e16d8 add auto support for AE_mp
hoelzl
parents: 41689
diff changeset
  1483
proof -
1100512e16d8 add auto support for AE_mp
hoelzl
parents: 41689
diff changeset
  1484
  have sets: "{x\<in>space M. u x \<noteq> 0} \<in> sets M"
1100512e16d8 add auto support for AE_mp
hoelzl
parents: 41689
diff changeset
  1485
    using u by auto
1100512e16d8 add auto support for AE_mp
hoelzl
parents: 41689
diff changeset
  1486
  then show ?thesis
1100512e16d8 add auto support for AE_mp
hoelzl
parents: 41689
diff changeset
  1487
    using positive_integral_0_iff[OF u] AE_iff_null_set[OF sets] by auto
1100512e16d8 add auto support for AE_mp
hoelzl
parents: 41689
diff changeset
  1488
qed
1100512e16d8 add auto support for AE_mp
hoelzl
parents: 41689
diff changeset
  1489
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
  1490
lemma (in measure_space) positive_integral_restricted:
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
  1491
  assumes "A \<in> sets M"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1492
  shows "integral\<^isup>P (restricted_space A) f = (\<integral>\<^isup>+ x. f x * indicator A x \<partial>M)"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1493
    (is "integral\<^isup>P ?R f = integral\<^isup>P M ?f")
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
  1494
proof -
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1495
  have msR: "measure_space ?R" by (rule restricted_measure_space[OF `A \<in> sets M`])
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1496
  then interpret R: measure_space ?R .
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
  1497
  have saR: "sigma_algebra ?R" by fact
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1498
  have *: "integral\<^isup>P ?R f = integral\<^isup>P ?R ?f"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1499
    by (intro R.positive_integral_cong) auto
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
  1500
  show ?thesis
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1501
    unfolding * positive_integral_def
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
  1502
    unfolding simple_function_restricted[OF `A \<in> sets M`]
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
  1503
    apply (simp add: SUPR_def)
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
  1504
    apply (rule arg_cong[where f=Sup])
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1505
  proof (auto simp add: image_iff simple_integral_restricted[OF `A \<in> sets M`])
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1506
    fix g assume "simple_function M (\<lambda>x. g x * indicator A x)"
40873
1ef85f4e7097 Shorter definition for positive_integral.
hoelzl
parents: 40872
diff changeset
  1507
      "g \<le> f"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1508
    then show "\<exists>x. simple_function M x \<and> x \<le> (\<lambda>x. f x * indicator A x) \<and>
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1509
      (\<integral>\<^isup>Sx. g x * indicator A x \<partial>M) = integral\<^isup>S M x"
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
  1510
      apply (rule_tac exI[of _ "\<lambda>x. g x * indicator A x"])
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
  1511
      by (auto simp: indicator_def le_fun_def)
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
  1512
  next
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1513
    fix g assume g: "simple_function M g" "g \<le> (\<lambda>x. f x * indicator A x)"
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
  1514
    then have *: "(\<lambda>x. g x * indicator A x) = g"
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
  1515
      "\<And>x. g x * indicator A x = g x"
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
  1516
      "\<And>x. g x \<le> f x"
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 39198
diff changeset
  1517
      by (auto simp: le_fun_def fun_eq_iff indicator_def split: split_if_asm)
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1518
    from g show "\<exists>x. simple_function M (\<lambda>xa. x xa * indicator A xa) \<and> x \<le> f \<and>
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1519
      integral\<^isup>S M g = integral\<^isup>S M (\<lambda>xa. x xa * indicator A xa)"
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
  1520
      using `A \<in> sets M`[THEN sets_into_space]
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
  1521
      apply (rule_tac exI[of _ "\<lambda>x. g x * indicator A x"])
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
  1522
      by (fastsimp simp: le_fun_def *)
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
  1523
  qed
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
  1524
qed
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
  1525
41545
9c869baf1c66 tuned formalization of subalgebra
hoelzl
parents: 41544
diff changeset
  1526
lemma (in measure_space) positive_integral_subalgebra:
9c869baf1c66 tuned formalization of subalgebra
hoelzl
parents: 41544
diff changeset
  1527
  assumes borel: "f \<in> borel_measurable N"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1528
  and N: "sets N \<subseteq> sets M" "space N = space M" "\<And>A. A \<in> sets N \<Longrightarrow> measure N A = \<mu> A"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1529
  and sa: "sigma_algebra N"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1530
  shows "integral\<^isup>P N f = integral\<^isup>P M f"
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
  1531
proof -
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1532
  interpret N: measure_space N using measure_space_subalgebra[OF sa N] .
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
  1533
  from N.borel_measurable_implies_simple_function_sequence[OF borel]
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1534
  obtain fs where Nsf: "\<And>i. simple_function N (fs i)" and "fs \<up> f" by blast
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1535
  then have sf: "\<And>i. simple_function M (fs i)"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1536
    using simple_function_subalgebra[OF _ N(1,2)] by blast
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1537
  from N.positive_integral_isoton_simple[OF `fs \<up> f` Nsf]
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1538
  have "integral\<^isup>P N f = (SUP i. \<Sum>x\<in>fs i ` space M. x * N.\<mu> (fs i -` {x} \<inter> space M))"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1539
    unfolding isoton_def simple_integral_def `space N = space M` by simp
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1540
  also have "\<dots> = (SUP i. \<Sum>x\<in>fs i ` space M. x * \<mu> (fs i -` {x} \<inter> space M))"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1541
    using N N.simple_functionD(2)[OF Nsf] unfolding `space N = space M` by auto
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1542
  also have "\<dots> = integral\<^isup>P M f"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1543
    using positive_integral_isoton_simple[OF `fs \<up> f` sf]
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1544
    unfolding isoton_def simple_integral_def `space N = space M` by simp
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1545
  finally show ?thesis .
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
  1546
qed
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
  1547
35692
f1315bbf1bc9 Moved theorems in Lebesgue to the right places
hoelzl
parents: 35582
diff changeset
  1548
section "Lebesgue Integral"
f1315bbf1bc9 Moved theorems in Lebesgue to the right places
hoelzl
parents: 35582
diff changeset
  1549
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1550
definition integrable where
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1551
  "integrable M f \<longleftrightarrow> f \<in> borel_measurable M \<and>
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1552
    (\<integral>\<^isup>+ x. Real (f x) \<partial>M) \<noteq> \<omega> \<and>
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1553
    (\<integral>\<^isup>+ x. Real (- f x) \<partial>M) \<noteq> \<omega>"
35692
f1315bbf1bc9 Moved theorems in Lebesgue to the right places
hoelzl
parents: 35582
diff changeset
  1554
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1555
lemma integrableD[dest]:
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1556
  assumes "integrable M f"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1557
  shows "f \<in> borel_measurable M" "(\<integral>\<^isup>+ x. Real (f x) \<partial>M) \<noteq> \<omega>" "(\<integral>\<^isup>+ x. Real (- f x) \<partial>M) \<noteq> \<omega>"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1558
  using assms unfolding integrable_def by auto
35692
f1315bbf1bc9 Moved theorems in Lebesgue to the right places
hoelzl
parents: 35582
diff changeset
  1559
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1560
definition lebesgue_integral_def:
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1561
  "integral\<^isup>L M f = real ((\<integral>\<^isup>+ x. Real (f x) \<partial>M)) - real ((\<integral>\<^isup>+ x. Real (- f x) \<partial>M))"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1562
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1563
syntax
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1564
  "_lebesgue_integral" :: "'a \<Rightarrow> ('a \<Rightarrow> real) \<Rightarrow> ('a, 'b) measure_space_scheme \<Rightarrow> real" ("\<integral> _. _ \<partial>_" [60,61] 110)
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1565
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1566
translations
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1567
  "\<integral> x. f \<partial>M" == "CONST integral\<^isup>L M (%x. f)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1568
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1569
lemma (in measure_space) integral_cong:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1570
  assumes "\<And>x. x \<in> space M \<Longrightarrow> f x = g x"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1571
  shows "integral\<^isup>L M f = integral\<^isup>L M g"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1572
  using assms by (simp cong: positive_integral_cong add: lebesgue_integral_def)
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1573
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1574
lemma (in measure_space) integral_cong_measure:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1575
  assumes "\<And>A. A \<in> sets M \<Longrightarrow> measure N A = \<mu> A" "sets N = sets M" "space N = space M"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1576
  shows "integral\<^isup>L N f = integral\<^isup>L M f"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1577
proof -
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1578
  interpret v: measure_space N
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1579
    by (rule measure_space_cong) fact+
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1580
  show ?thesis
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1581
    by (simp add: positive_integral_cong_measure[OF assms] lebesgue_integral_def)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1582
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1583
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1584
lemma (in measure_space) integral_cong_AE:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1585
  assumes cong: "AE x. f x = g x"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1586
  shows "integral\<^isup>L M f = integral\<^isup>L M g"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1587
proof -
41705
1100512e16d8 add auto support for AE_mp
hoelzl
parents: 41689
diff changeset
  1588
  have "AE x. Real (f x) = Real (g x)" using cong by auto
1100512e16d8 add auto support for AE_mp
hoelzl
parents: 41689
diff changeset
  1589
  moreover have "AE x. Real (- f x) = Real (- g x)" using cong by auto
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1590
  ultimately show ?thesis
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1591
    by (simp cong: positive_integral_cong_AE add: lebesgue_integral_def)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1592
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1593
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1594
lemma (in measure_space) integrable_cong:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1595
  "(\<And>x. x \<in> space M \<Longrightarrow> f x = g x) \<Longrightarrow> integrable M f \<longleftrightarrow> integrable M g"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1596
  by (simp cong: positive_integral_cong measurable_cong add: integrable_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1597
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1598
lemma (in measure_space) integral_eq_positive_integral:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1599
  assumes "\<And>x. 0 \<le> f x"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1600
  shows "integral\<^isup>L M f = real (\<integral>\<^isup>+ x. Real (f x) \<partial>M)"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1601
proof -
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1602
  have "\<And>x. Real (- f x) = 0" using assms by simp
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1603
  thus ?thesis by (simp del: Real_eq_0 add: lebesgue_integral_def)
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1604
qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1605
41661
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41545
diff changeset
  1606
lemma (in measure_space) integral_vimage:
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41545
diff changeset
  1607
  assumes T: "sigma_algebra M'" "T \<in> measurable M M'"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1608
    and \<nu>: "\<And>A. A \<in> sets M' \<Longrightarrow> measure M' A = \<mu> (T -` A \<inter> space M)"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1609
  assumes f: "integrable M' f"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1610
  shows "integrable M (\<lambda>x. f (T x))" (is ?P)
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1611
    and "integral\<^isup>L M' f = (\<integral>x. f (T x) \<partial>M)" (is ?I)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1612
proof -
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1613
  interpret T: measure_space M' using \<nu> by (rule measure_space_vimage[OF T])
41661
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41545
diff changeset
  1614
  from measurable_comp[OF T(2), of f borel]
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41545
diff changeset
  1615
  have borel: "(\<lambda>x. Real (f x)) \<in> borel_measurable M'" "(\<lambda>x. Real (- f x)) \<in> borel_measurable M'"
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41545
diff changeset
  1616
    and "(\<lambda>x. f (T x)) \<in> borel_measurable M"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1617
    using f unfolding integrable_def by (auto simp: comp_def)
41661
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41545
diff changeset
  1618
  then show ?P ?I
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1619
    using f unfolding lebesgue_integral_def integrable_def
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1620
    by (auto simp: borel[THEN positive_integral_vimage[OF T], OF \<nu>])
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1621
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1622
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1623
lemma (in measure_space) integral_minus[intro, simp]:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1624
  assumes "integrable M f"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1625
  shows "integrable M (\<lambda>x. - f x)" "(\<integral>x. - f x \<partial>M) = - integral\<^isup>L M f"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1626
  using assms by (auto simp: integrable_def lebesgue_integral_def)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1627
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1628
lemma (in measure_space) integral_of_positive_diff:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1629
  assumes integrable: "integrable M u" "integrable M v"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1630
  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"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1631
  shows "integrable M f" and "integral\<^isup>L M f = integral\<^isup>L M u - integral\<^isup>L M v"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1632
proof -
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1633
  let "?f x" = "Real (f x)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1634
  let "?mf x" = "Real (- f x)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1635
  let "?u x" = "Real (u x)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1636
  let "?v x" = "Real (v x)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1637
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1638
  from borel_measurable_diff[of u v] integrable
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1639
  have f_borel: "?f \<in> borel_measurable M" and
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1640
    mf_borel: "?mf \<in> borel_measurable M" and
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1641
    v_borel: "?v \<in> borel_measurable M" and
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1642
    u_borel: "?u \<in> borel_measurable M" and
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1643
    "f \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1644
    by (auto simp: f_def[symmetric] integrable_def)
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1645
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1646
  have "(\<integral>\<^isup>+ x. Real (u x - v x) \<partial>M) \<le> integral\<^isup>P M ?u"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1647
    using pos by (auto intro!: positive_integral_mono)
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1648
  moreover have "(\<integral>\<^isup>+ x. Real (v x - u x) \<partial>M) \<le> integral\<^isup>P M ?v"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1649
    using pos by (auto intro!: positive_integral_mono)
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1650
  ultimately show f: "integrable M f"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1651
    using `integrable M u` `integrable M v` `f \<in> borel_measurable M`
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1652
    by (auto simp: integrable_def f_def)
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1653
  hence mf: "integrable M (\<lambda>x. - f x)" ..
38656
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 (- v x) = 0" "\<And>x. Real (- u x) = 0"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1656
    using pos by auto
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1657
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1658
  have "\<And>x. ?u x + ?mf x = ?v x + ?f x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1659
    unfolding f_def using pos by simp
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1660
  hence "(\<integral>\<^isup>+ x. ?u x + ?mf x \<partial>M) = (\<integral>\<^isup>+ x. ?v x + ?f x \<partial>M)" by simp
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1661
  hence "real (integral\<^isup>P M ?u + integral\<^isup>P M ?mf) =
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1662
      real (integral\<^isup>P M ?v + integral\<^isup>P M ?f)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1663
    using positive_integral_add[OF u_borel mf_borel]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1664
    using positive_integral_add[OF v_borel f_borel]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1665
    by auto
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1666
  then show "integral\<^isup>L M f = integral\<^isup>L M u - integral\<^isup>L M v"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1667
    using f mf `integrable M u` `integrable M v`
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1668
    unfolding lebesgue_integral_def integrable_def *
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1669
    by (cases "integral\<^isup>P M ?f", cases "integral\<^isup>P M ?mf", cases "integral\<^isup>P M ?v", cases "integral\<^isup>P M ?u")
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1670
       (auto simp add: field_simps)
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1671
qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1672
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1673
lemma (in measure_space) integral_linear:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1674
  assumes "integrable M f" "integrable M g" and "0 \<le> a"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1675
  shows "integrable M (\<lambda>t. a * f t + g t)"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1676
  and "(\<integral> t. a * f t + g t \<partial>M) = a * integral\<^isup>L M f + integral\<^isup>L M g"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1677
proof -
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1678
  let ?PI = "integral\<^isup>P M"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1679
  let "?f x" = "Real (f x)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1680
  let "?g x" = "Real (g x)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1681
  let "?mf x" = "Real (- f x)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1682
  let "?mg x" = "Real (- g x)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1683
  let "?p t" = "max 0 (a * f t) + max 0 (g t)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1684
  let "?n t" = "max 0 (- (a * f t)) + max 0 (- g t)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1685
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1686
  have pos: "?f \<in> borel_measurable M" "?g \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1687
    and neg: "?mf \<in> borel_measurable M" "?mg \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1688
    and p: "?p \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1689
    and n: "?n \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1690
    using assms by (simp_all add: integrable_def)
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1691
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1692
  have *: "\<And>x. Real (?p x) = Real a * ?f x + ?g x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1693
          "\<And>x. Real (?n x) = Real a * ?mf x + ?mg x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1694
          "\<And>x. Real (- ?p x) = 0"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1695
          "\<And>x. Real (- ?n x) = 0"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1696
    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
  1697
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1698
  note linear =
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1699
    positive_integral_linear[OF pos]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1700
    positive_integral_linear[OF neg]
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1701
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1702
  have "integrable M ?p" "integrable M ?n"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1703
      "\<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
  1704
    using assms p n unfolding integrable_def * linear by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1705
  note diff = integral_of_positive_diff[OF this]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1706
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1707
  show "integrable M (\<lambda>t. a * f t + g t)" by (rule diff)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1708
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1709
  from assms show "(\<integral> t. a * f t + g t \<partial>M) = a * integral\<^isup>L M f + integral\<^isup>L M g"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1710
    unfolding diff(2) unfolding lebesgue_integral_def * linear integrable_def
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1711
    by (cases "?PI ?f", cases "?PI ?mf", cases "?PI ?g", cases "?PI ?mg")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1712
       (auto simp add: field_simps zero_le_mult_iff)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1713
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1714
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1715
lemma (in measure_space) integral_add[simp, intro]:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1716
  assumes "integrable M f" "integrable M g"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1717
  shows "integrable M (\<lambda>t. f t + g t)"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1718
  and "(\<integral> t. f t + g t \<partial>M) = integral\<^isup>L M f + integral\<^isup>L M g"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1719
  using assms integral_linear[where a=1] by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1720
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1721
lemma (in measure_space) integral_zero[simp, intro]:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1722
  shows "integrable M (\<lambda>x. 0)" "(\<integral> x.0 \<partial>M) = 0"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1723
  unfolding integrable_def lebesgue_integral_def
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1724
  by (auto simp add: borel_measurable_const)
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1725
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1726
lemma (in measure_space) integral_cmult[simp, intro]:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1727
  assumes "integrable M f"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1728
  shows "integrable M (\<lambda>t. a * f t)" (is ?P)
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1729
  and "(\<integral> t. a * f t \<partial>M) = a * integral\<^isup>L M f" (is ?I)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1730
proof -
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1731
  have "integrable M (\<lambda>t. a * f t) \<and> (\<integral> t. a * f t \<partial>M) = a * integral\<^isup>L M f"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1732
  proof (cases rule: le_cases)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1733
    assume "0 \<le> a" show ?thesis
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1734
      using integral_linear[OF assms integral_zero(1) `0 \<le> a`]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1735
      by (simp add: integral_zero)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1736
  next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1737
    assume "a \<le> 0" hence "0 \<le> - a" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1738
    have *: "\<And>t. - a * t + 0 = (-a) * t" by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1739
    show ?thesis using integral_linear[OF assms integral_zero(1) `0 \<le> - a`]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1740
        integral_minus(1)[of "\<lambda>t. - a * f t"]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1741
      unfolding * integral_zero by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1742
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1743
  thus ?P ?I by auto
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1744
qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1745
41096
843c40bbc379 integral over setprod
hoelzl
parents: 41095
diff changeset
  1746
lemma (in measure_space) integral_multc:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1747
  assumes "integrable M f"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1748
  shows "(\<integral> x. f x * c \<partial>M) = integral\<^isup>L M f * c"
41096
843c40bbc379 integral over setprod
hoelzl
parents: 41095
diff changeset
  1749
  unfolding mult_commute[of _ c] integral_cmult[OF assms] ..
843c40bbc379 integral over setprod
hoelzl
parents: 41095
diff changeset
  1750
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1751
lemma (in measure_space) integral_mono_AE:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1752
  assumes fg: "integrable M f" "integrable M g"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1753
  and mono: "AE t. f t \<le> g t"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1754
  shows "integral\<^isup>L M f \<le> integral\<^isup>L M g"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1755
proof -
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1756
  have "AE x. Real (f x) \<le> Real (g x)"
41705
1100512e16d8 add auto support for AE_mp
hoelzl
parents: 41689
diff changeset
  1757
    using mono by auto
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1758
  moreover have "AE x. Real (- g x) \<le> Real (- f x)"
41705
1100512e16d8 add auto support for AE_mp
hoelzl
parents: 41689
diff changeset
  1759
    using mono by auto
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1760
  ultimately show ?thesis using fg
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1761
    by (auto simp: lebesgue_integral_def integrable_def diff_minus
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40875
diff changeset
  1762
             intro!: add_mono real_of_pextreal_mono positive_integral_mono_AE)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1763
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1764
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1765
lemma (in measure_space) integral_mono:
41705
1100512e16d8 add auto support for AE_mp
hoelzl
parents: 41689
diff changeset
  1766
  assumes "integrable M f" "integrable M g" "\<And>t. t \<in> space M \<Longrightarrow> f t \<le> g t"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1767
  shows "integral\<^isup>L M f \<le> integral\<^isup>L M g"
41705
1100512e16d8 add auto support for AE_mp
hoelzl
parents: 41689
diff changeset
  1768
  using assms by (auto intro: integral_mono_AE)
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1769
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1770
lemma (in measure_space) integral_diff[simp, intro]:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1771
  assumes f: "integrable M f" and g: "integrable M g"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1772
  shows "integrable M (\<lambda>t. f t - g t)"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1773
  and "(\<integral> t. f t - g t \<partial>M) = integral\<^isup>L M f - integral\<^isup>L M g"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1774
  using integral_add[OF f integral_minus(1)[OF g]]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1775
  unfolding diff_minus integral_minus(2)[OF g]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1776
  by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1777
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1778
lemma (in measure_space) integral_indicator[simp, intro]:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1779
  assumes "a \<in> sets M" and "\<mu> a \<noteq> \<omega>"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1780
  shows "integral\<^isup>L M (indicator a) = real (\<mu> a)" (is ?int)
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1781
  and "integrable M (indicator a)" (is ?able)
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1782
proof -
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1783
  have *:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1784
    "\<And>A x. Real (indicator A x) = indicator A x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1785
    "\<And>A x. Real (- indicator A x) = 0" unfolding indicator_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1786
  show ?int ?able
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1787
    using assms unfolding lebesgue_integral_def integrable_def
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1788
    by (auto simp: * positive_integral_indicator borel_measurable_indicator)
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1789
qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1790
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1791
lemma (in measure_space) integral_cmul_indicator:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1792
  assumes "A \<in> sets M" and "c \<noteq> 0 \<Longrightarrow> \<mu> A \<noteq> \<omega>"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1793
  shows "integrable M (\<lambda>x. c * indicator A x)" (is ?P)
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1794
  and "(\<integral>x. c * indicator A x \<partial>M) = c * real (\<mu> A)" (is ?I)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1795
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1796
  show ?P
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1797
  proof (cases "c = 0")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1798
    case False with assms show ?thesis by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1799
  qed simp
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1800
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1801
  show ?I
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1802
  proof (cases "c = 0")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1803
    case False with assms show ?thesis by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1804
  qed simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1805
qed
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1806
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1807
lemma (in measure_space) integral_setsum[simp, intro]:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1808
  assumes "\<And>n. n \<in> S \<Longrightarrow> integrable M (f n)"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1809
  shows "(\<integral>x. (\<Sum> i \<in> S. f i x) \<partial>M) = (\<Sum> i \<in> S. integral\<^isup>L M (f i))" (is "?int S")
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1810
    and "integrable M (\<lambda>x. \<Sum> i \<in> S. f i x)" (is "?I S")
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1811
proof -
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1812
  have "?int S \<and> ?I S"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1813
  proof (cases "finite S")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1814
    assume "finite S"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1815
    from this assms show ?thesis by (induct S) simp_all
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1816
  qed simp
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1817
  thus "?int S" and "?I S" by auto
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1818
qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1819
36624
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
  1820
lemma (in measure_space) integrable_abs:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1821
  assumes "integrable M f"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1822
  shows "integrable M (\<lambda> x. \<bar>f x\<bar>)"
36624
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
  1823
proof -
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1824
  have *:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1825
    "\<And>x. Real \<bar>f x\<bar> = Real (f x) + Real (- f x)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1826
    "\<And>x. Real (- \<bar>f x\<bar>) = 0" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1827
  have abs: "(\<lambda>x. \<bar>f x\<bar>) \<in> borel_measurable M" and
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1828
    f: "(\<lambda>x. Real (f x)) \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1829
        "(\<lambda>x. Real (- f x)) \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1830
    using assms unfolding integrable_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1831
  from abs assms show ?thesis unfolding integrable_def *
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1832
    using positive_integral_linear[OF f, of 1] by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1833
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1834
41545
9c869baf1c66 tuned formalization of subalgebra
hoelzl
parents: 41544
diff changeset
  1835
lemma (in measure_space) integral_subalgebra:
9c869baf1c66 tuned formalization of subalgebra
hoelzl
parents: 41544
diff changeset
  1836
  assumes borel: "f \<in> borel_measurable N"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1837
  and N: "sets N \<subseteq> sets M" "space N = space M" "\<And>A. A \<in> sets N \<Longrightarrow> measure N A = \<mu> A" and sa: "sigma_algebra N"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1838
  shows "integrable N f \<longleftrightarrow> integrable M f" (is ?P)
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1839
    and "integral\<^isup>L N f = integral\<^isup>L M f" (is ?I)
41545
9c869baf1c66 tuned formalization of subalgebra
hoelzl
parents: 41544
diff changeset
  1840
proof -
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1841
  interpret N: measure_space N using measure_space_subalgebra[OF sa N] .
41545
9c869baf1c66 tuned formalization of subalgebra
hoelzl
parents: 41544
diff changeset
  1842
  have "(\<lambda>x. Real (f x)) \<in> borel_measurable N" "(\<lambda>x. Real (- f x)) \<in> borel_measurable N"
9c869baf1c66 tuned formalization of subalgebra
hoelzl
parents: 41544
diff changeset
  1843
    using borel by auto
9c869baf1c66 tuned formalization of subalgebra
hoelzl
parents: 41544
diff changeset
  1844
  note * = this[THEN positive_integral_subalgebra[OF _ N sa]]
9c869baf1c66 tuned formalization of subalgebra
hoelzl
parents: 41544
diff changeset
  1845
  have "f \<in> borel_measurable M \<longleftrightarrow> f \<in> borel_measurable N"
9c869baf1c66 tuned formalization of subalgebra
hoelzl
parents: 41544
diff changeset
  1846
    using assms unfolding measurable_def by auto
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1847
  then show ?P ?I by (auto simp: * integrable_def lebesgue_integral_def)
41545
9c869baf1c66 tuned formalization of subalgebra
hoelzl
parents: 41544
diff changeset
  1848
qed
9c869baf1c66 tuned formalization of subalgebra
hoelzl
parents: 41544
diff changeset
  1849
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1850
lemma (in measure_space) integrable_bound:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1851
  assumes "integrable M f"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1852
  and f: "\<And>x. x \<in> space M \<Longrightarrow> 0 \<le> f x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1853
    "\<And>x. x \<in> space M \<Longrightarrow> \<bar>g x\<bar> \<le> f x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1854
  assumes borel: "g \<in> borel_measurable M"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1855
  shows "integrable M g"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1856
proof -
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1857
  have "(\<integral>\<^isup>+ x. Real (g x) \<partial>M) \<le> (\<integral>\<^isup>+ x. Real \<bar>g x\<bar> \<partial>M)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1858
    by (auto intro!: positive_integral_mono)
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1859
  also have "\<dots> \<le> (\<integral>\<^isup>+ x. Real (f x) \<partial>M)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1860
    using f by (auto intro!: positive_integral_mono)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1861
  also have "\<dots> < \<omega>"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1862
    using `integrable M f` unfolding integrable_def by (auto simp: pextreal_less_\<omega>)
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1863
  finally have pos: "(\<integral>\<^isup>+ x. Real (g x) \<partial>M) < \<omega>" .
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1864
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1865
  have "(\<integral>\<^isup>+ x. Real (- g x) \<partial>M) \<le> (\<integral>\<^isup>+ x. Real (\<bar>g x\<bar>) \<partial>M)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1866
    by (auto intro!: positive_integral_mono)
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1867
  also have "\<dots> \<le> (\<integral>\<^isup>+ x. Real (f x) \<partial>M)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1868
    using f by (auto intro!: positive_integral_mono)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1869
  also have "\<dots> < \<omega>"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1870
    using `integrable M f` unfolding integrable_def by (auto simp: pextreal_less_\<omega>)
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1871
  finally have neg: "(\<integral>\<^isup>+ x. Real (- g x) \<partial>M) < \<omega>" .
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1872
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1873
  from neg pos borel show ?thesis
36624
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
  1874
    unfolding integrable_def by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1875
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1876
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1877
lemma (in measure_space) integrable_abs_iff:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1878
  "f \<in> borel_measurable M \<Longrightarrow> integrable M (\<lambda> x. \<bar>f x\<bar>) \<longleftrightarrow> integrable M f"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1879
  by (auto intro!: integrable_bound[where g=f] integrable_abs)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1880
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1881
lemma (in measure_space) integrable_max:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1882
  assumes int: "integrable M f" "integrable M g"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1883
  shows "integrable M (\<lambda> x. max (f x) (g x))"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1884
proof (rule integrable_bound)
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1885
  show "integrable M (\<lambda>x. \<bar>f x\<bar> + \<bar>g x\<bar>)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1886
    using int by (simp add: integrable_abs)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1887
  show "(\<lambda>x. max (f x) (g x)) \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1888
    using int unfolding integrable_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1889
next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1890
  fix x assume "x \<in> space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1891
  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
  1892
    by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1893
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1894
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1895
lemma (in measure_space) integrable_min:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1896
  assumes int: "integrable M f" "integrable M g"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1897
  shows "integrable M (\<lambda> x. min (f x) (g x))"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1898
proof (rule integrable_bound)
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1899
  show "integrable M (\<lambda>x. \<bar>f x\<bar> + \<bar>g x\<bar>)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1900
    using int by (simp add: integrable_abs)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1901
  show "(\<lambda>x. min (f x) (g x)) \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1902
    using int unfolding integrable_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1903
next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1904
  fix x assume "x \<in> space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1905
  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
  1906
    by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1907
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1908
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1909
lemma (in measure_space) integral_triangle_inequality:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1910
  assumes "integrable M f"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1911
  shows "\<bar>integral\<^isup>L M f\<bar> \<le> (\<integral>x. \<bar>f x\<bar> \<partial>M)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1912
proof -
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1913
  have "\<bar>integral\<^isup>L M f\<bar> = max (integral\<^isup>L M f) (- integral\<^isup>L M f)" by auto
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1914
  also have "\<dots> \<le> (\<integral>x. \<bar>f x\<bar> \<partial>M)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1915
      using assms integral_minus(2)[of f, symmetric]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1916
      by (auto intro!: integral_mono integrable_abs simp del: integral_minus)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1917
  finally show ?thesis .
36624
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
  1918
qed
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
  1919
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1920
lemma (in measure_space) integral_positive:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1921
  assumes "integrable M f" "\<And>x. x \<in> space M \<Longrightarrow> 0 \<le> f x"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1922
  shows "0 \<le> integral\<^isup>L M f"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1923
proof -
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1924
  have "0 = (\<integral>x. 0 \<partial>M)" by (auto simp: integral_zero)
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1925
  also have "\<dots> \<le> integral\<^isup>L M f"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1926
    using assms by (rule integral_mono[OF integral_zero(1)])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1927
  finally show ?thesis .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1928
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1929
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1930
lemma (in measure_space) integral_monotone_convergence_pos:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1931
  assumes i: "\<And>i. integrable M (f i)" and mono: "\<And>x. mono (\<lambda>n. f n x)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1932
  and pos: "\<And>x i. 0 \<le> f i x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1933
  and lim: "\<And>x. (\<lambda>i. f i x) ----> u x"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1934
  and ilim: "(\<lambda>i. integral\<^isup>L M (f i)) ----> x"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1935
  shows "integrable M u"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1936
  and "integral\<^isup>L M u = x"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1937
proof -
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1938
  { fix x have "0 \<le> u x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1939
      using mono pos[of 0 x] incseq_le[OF _ lim, of x 0]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1940
      by (simp add: mono_def incseq_def) }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1941
  note pos_u = this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1942
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1943
  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
  1944
    using pos by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1945
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1946
  have SUP_F: "\<And>x. (SUP n. Real (f n x)) = Real (u x)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1947
    using mono pos pos_u lim by (rule SUP_eq_LIMSEQ[THEN iffD2])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1948
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1949
  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
  1950
    using i unfolding integrable_def by auto
41097
a1abfa4e2b44 use SUPR_ and INFI_apply instead of SUPR_, INFI_fun_expand
hoelzl
parents: 41096
diff changeset
  1951
  hence "(\<lambda>x. SUP i. Real (f i x)) \<in> borel_measurable M"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1952
    by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1953
  hence borel_u: "u \<in> borel_measurable M"
41097
a1abfa4e2b44 use SUPR_ and INFI_apply instead of SUPR_, INFI_fun_expand
hoelzl
parents: 41096
diff changeset
  1954
    using pos_u by (auto simp: borel_measurable_Real_eq SUP_F)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1955
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1956
  have integral_eq: "\<And>n. (\<integral>\<^isup>+ x. Real (f n x) \<partial>M) = Real (integral\<^isup>L M (f n))"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1957
    using i unfolding lebesgue_integral_def integrable_def by (auto simp: Real_real)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1958
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1959
  have pos_integral: "\<And>n. 0 \<le> integral\<^isup>L M (f n)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1960
    using pos i by (auto simp: integral_positive)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1961
  hence "0 \<le> x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1962
    using LIMSEQ_le_const[OF ilim, of 0] by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1963
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1964
  have "(\<lambda>i. (\<integral>\<^isup>+ x. Real (f i x) \<partial>M)) \<up> (\<integral>\<^isup>+ x. Real (u x) \<partial>M)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1965
  proof (rule positive_integral_isoton)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1966
    from SUP_F mono pos
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1967
    show "(\<lambda>i x. Real (f i x)) \<up> (\<lambda>x. Real (u x))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1968
      unfolding isoton_fun_expand by (auto simp: isoton_def mono_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1969
  qed (rule borel_f)
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1970
  hence pI: "(\<integral>\<^isup>+ x. Real (u x) \<partial>M) = (SUP n. (\<integral>\<^isup>+ x. Real (f n x) \<partial>M))"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1971
    unfolding isoton_def by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1972
  also have "\<dots> = Real x" unfolding integral_eq
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1973
  proof (rule SUP_eq_LIMSEQ[THEN iffD2])
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1974
    show "mono (\<lambda>n. integral\<^isup>L M (f n))"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1975
      using mono i by (auto simp: mono_def intro!: integral_mono)
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1976
    show "\<And>n. 0 \<le> integral\<^isup>L M (f n)" using pos_integral .
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1977
    show "0 \<le> x" using `0 \<le> x` .
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1978
    show "(\<lambda>n. integral\<^isup>L M (f n)) ----> x" using ilim .
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1979
  qed
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1980
  finally show  "integrable M u" "integral\<^isup>L M u = x" using borel_u `0 \<le> x`
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1981
    unfolding integrable_def lebesgue_integral_def by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1982
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1983
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1984
lemma (in measure_space) integral_monotone_convergence:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1985
  assumes f: "\<And>i. integrable M (f i)" and "mono f"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1986
  and lim: "\<And>x. (\<lambda>i. f i x) ----> u x"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1987
  and ilim: "(\<lambda>i. integral\<^isup>L M (f i)) ----> x"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1988
  shows "integrable M u"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1989
  and "integral\<^isup>L M u = x"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1990
proof -
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1991
  have 1: "\<And>i. integrable M (\<lambda>x. f i x - f 0 x)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1992
      using f by (auto intro!: integral_diff)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1993
  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
  1994
      unfolding mono_def le_fun_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1995
  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
  1996
      unfolding mono_def le_fun_def by (auto simp: field_simps)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1997
  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
  1998
    using lim by (auto intro!: LIMSEQ_diff)
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1999
  have 5: "(\<lambda>i. (\<integral>x. f i x - f 0 x \<partial>M)) ----> x - integral\<^isup>L M (f 0)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2000
    using f ilim by (auto intro!: LIMSEQ_diff simp: integral_diff)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2001
  note diff = integral_monotone_convergence_pos[OF 1 2 3 4 5]
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2002
  have "integrable M (\<lambda>x. (u x - f 0 x) + f 0 x)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2003
    using diff(1) f by (rule integral_add(1))
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2004
  with diff(2) f show "integrable M u" "integral\<^isup>L M u = x"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2005
    by (auto simp: integral_diff)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2006
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2007
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2008
lemma (in measure_space) integral_0_iff:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2009
  assumes "integrable M f"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2010
  shows "(\<integral>x. \<bar>f x\<bar> \<partial>M) = 0 \<longleftrightarrow> \<mu> {x\<in>space M. f x \<noteq> 0} = 0"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2011
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2012
  have *: "\<And>x. Real (- \<bar>f x\<bar>) = 0" by auto
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2013
  have "integrable M (\<lambda>x. \<bar>f x\<bar>)" using assms by (rule integrable_abs)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2014
  hence "(\<lambda>x. Real (\<bar>f x\<bar>)) \<in> borel_measurable M"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2015
    "(\<integral>\<^isup>+ x. Real \<bar>f x\<bar> \<partial>M) \<noteq> \<omega>" unfolding integrable_def by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2016
  from positive_integral_0_iff[OF this(1)] this(2)
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2017
  show ?thesis unfolding lebesgue_integral_def *
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40875
diff changeset
  2018
    by (simp add: real_of_pextreal_eq_0)
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  2019
qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  2020
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  2021
lemma (in measure_space) positive_integral_omega:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  2022
  assumes "f \<in> borel_measurable M"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2023
  and "integral\<^isup>P M f \<noteq> \<omega>"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  2024
  shows "\<mu> (f -` {\<omega>} \<inter> space M) = 0"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  2025
proof -
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2026
  have "\<omega> * \<mu> (f -` {\<omega>} \<inter> space M) = (\<integral>\<^isup>+ x. \<omega> * indicator (f -` {\<omega>} \<inter> space M) x \<partial>M)"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  2027
    using positive_integral_cmult_indicator[OF borel_measurable_vimage, OF assms(1), of \<omega> \<omega>] by simp
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2028
  also have "\<dots> \<le> integral\<^isup>P M f"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  2029
    by (auto intro!: positive_integral_mono simp: indicator_def)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  2030
  finally show ?thesis
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  2031
    using assms(2) by (cases ?thesis) auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  2032
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  2033
41026
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41023
diff changeset
  2034
lemma (in measure_space) positive_integral_omega_AE:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2035
  assumes "f \<in> borel_measurable M" "integral\<^isup>P M f \<noteq> \<omega>" shows "AE x. f x \<noteq> \<omega>"
41026
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41023
diff changeset
  2036
proof (rule AE_I)
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41023
diff changeset
  2037
  show "\<mu> (f -` {\<omega>} \<inter> space M) = 0"
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41023
diff changeset
  2038
    by (rule positive_integral_omega[OF assms])
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41023
diff changeset
  2039
  show "f -` {\<omega>} \<inter> space M \<in> sets M"
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41023
diff changeset
  2040
    using assms by (auto intro: borel_measurable_vimage)
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41023
diff changeset
  2041
qed auto
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41023
diff changeset
  2042
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  2043
lemma (in measure_space) simple_integral_omega:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2044
  assumes "simple_function M f"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2045
  and "integral\<^isup>S M f \<noteq> \<omega>"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  2046
  shows "\<mu> (f -` {\<omega>} \<inter> space M) = 0"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  2047
proof (rule positive_integral_omega)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  2048
  show "f \<in> borel_measurable M" using assms by (auto intro: borel_measurable_simple_function)
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2049
  show "integral\<^isup>P M f \<noteq> \<omega>"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  2050
    using assms by (simp add: positive_integral_eq_simple_integral)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  2051
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  2052
41026
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41023
diff changeset
  2053
lemma (in measure_space) integral_real:
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41023
diff changeset
  2054
  fixes f :: "'a \<Rightarrow> pextreal"
41705
1100512e16d8 add auto support for AE_mp
hoelzl
parents: 41689
diff changeset
  2055
  assumes [simp]: "AE x. f x \<noteq> \<omega>"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2056
  shows "(\<integral>x. real (f x) \<partial>M) = real (integral\<^isup>P M f)" (is ?plus)
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2057
    and "(\<integral>x. - real (f x) \<partial>M) = - real (integral\<^isup>P M f)" (is ?minus)
41026
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41023
diff changeset
  2058
proof -
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2059
  have "(\<integral>\<^isup>+ x. Real (real (f x)) \<partial>M) = integral\<^isup>P M f"
41705
1100512e16d8 add auto support for AE_mp
hoelzl
parents: 41689
diff changeset
  2060
    by (auto intro!: positive_integral_cong_AE simp: Real_real)
41026
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41023
diff changeset
  2061
  moreover
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2062
  have "(\<integral>\<^isup>+ x. Real (- real (f x)) \<partial>M) = (\<integral>\<^isup>+ x. 0 \<partial>M)"
41026
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41023
diff changeset
  2063
    by (intro positive_integral_cong) auto
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41023
diff changeset
  2064
  ultimately show ?plus ?minus
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2065
    by (auto simp: lebesgue_integral_def integrable_def)
41026
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41023
diff changeset
  2066
qed
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41023
diff changeset
  2067
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2068
lemma (in measure_space) integral_dominated_convergence:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2069
  assumes u: "\<And>i. integrable M (u i)" and bound: "\<And>x j. x\<in>space M \<Longrightarrow> \<bar>u j x\<bar> \<le> w x"
41705
1100512e16d8 add auto support for AE_mp
hoelzl
parents: 41689
diff changeset
  2070
  and w: "integrable M w"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2071
  and u': "\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>i. u i x) ----> u' x"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2072
  shows "integrable M u'"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2073
  and "(\<lambda>i. (\<integral>x. \<bar>u i x - u' x\<bar> \<partial>M)) ----> 0" (is "?lim_diff")
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2074
  and "(\<lambda>i. integral\<^isup>L M (u i)) ----> integral\<^isup>L M u'" (is ?lim)
36624
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
  2075
proof -
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2076
  { fix x j assume x: "x \<in> space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2077
    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
  2078
    from LIMSEQ_le_const2[OF this]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2079
    have "\<bar>u' x\<bar> \<le> w x" using bound[OF x] by auto }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2080
  note u'_bound = this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2081
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2082
  from u[unfolded integrable_def]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2083
  have u'_borel: "u' \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2084
    using u' by (blast intro: borel_measurable_LIMSEQ[of u])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2085
41705
1100512e16d8 add auto support for AE_mp
hoelzl
parents: 41689
diff changeset
  2086
  { fix x assume x: "x \<in> space M"
1100512e16d8 add auto support for AE_mp
hoelzl
parents: 41689
diff changeset
  2087
    then have "0 \<le> \<bar>u 0 x\<bar>" by auto
1100512e16d8 add auto support for AE_mp
hoelzl
parents: 41689
diff changeset
  2088
    also have "\<dots> \<le> w x" using bound[OF x] by auto
1100512e16d8 add auto support for AE_mp
hoelzl
parents: 41689
diff changeset
  2089
    finally have "0 \<le> w x" . }
1100512e16d8 add auto support for AE_mp
hoelzl
parents: 41689
diff changeset
  2090
  note w_pos = this
1100512e16d8 add auto support for AE_mp
hoelzl
parents: 41689
diff changeset
  2091
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2092
  show "integrable M u'"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2093
  proof (rule integrable_bound)
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2094
    show "integrable M w" by fact
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2095
    show "u' \<in> borel_measurable M" by fact
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2096
  next
41705
1100512e16d8 add auto support for AE_mp
hoelzl
parents: 41689
diff changeset
  2097
    fix x assume x: "x \<in> space M" then show "0 \<le> w x" by fact
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2098
    show "\<bar>u' x\<bar> \<le> w x" using u'_bound[OF x] .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2099
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2100
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2101
  let "?diff n x" = "2 * w x - \<bar>u n x - u' x\<bar>"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2102
  have diff: "\<And>n. integrable M (\<lambda>x. \<bar>u n x - u' x\<bar>)"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2103
    using w u `integrable M u'`
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2104
    by (auto intro!: integral_add integral_diff integral_cmult integrable_abs)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2105
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2106
  { fix j x assume x: "x \<in> space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2107
    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
  2108
    also have "\<dots> \<le> w x + w x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2109
      by (rule add_mono[OF bound[OF x] u'_bound[OF x]])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2110
    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
  2111
  note diff_less_2w = this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2112
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2113
  have PI_diff: "\<And>m n. (\<integral>\<^isup>+ x. Real (?diff (m + n) x) \<partial>M) =
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2114
    (\<integral>\<^isup>+ x. Real (2 * w x) \<partial>M) - (\<integral>\<^isup>+ x. Real \<bar>u (m + n) x - u' x\<bar> \<partial>M)"
41705
1100512e16d8 add auto support for AE_mp
hoelzl
parents: 41689
diff changeset
  2115
    using diff w diff_less_2w w_pos
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2116
    by (subst positive_integral_diff[symmetric])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2117
       (auto simp: integrable_def intro!: positive_integral_cong)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2118
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2119
  have "integrable M (\<lambda>x. 2 * w x)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2120
    using w by (auto intro: integral_cmult)
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2121
  hence I2w_fin: "(\<integral>\<^isup>+ x. Real (2 * w x) \<partial>M) \<noteq> \<omega>" and
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2122
    borel_2w: "(\<lambda>x. Real (2 * w x)) \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2123
    unfolding integrable_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2124
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2125
  have "(INF n. SUP m. (\<integral>\<^isup>+ x. Real \<bar>u (m + n) x - u' x\<bar> \<partial>M)) = 0" (is "?lim_SUP = 0")
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2126
  proof cases
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2127
    assume eq_0: "(\<integral>\<^isup>+ x. Real (2 * w x) \<partial>M) = 0"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2128
    have "\<And>i. (\<integral>\<^isup>+ x. Real \<bar>u i x - u' x\<bar> \<partial>M) \<le> (\<integral>\<^isup>+ x. Real (2 * w x) \<partial>M)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2129
    proof (rule positive_integral_mono)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2130
      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
  2131
      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
  2132
    qed
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2133
    hence "\<And>i. (\<integral>\<^isup>+ x. Real \<bar>u i x - u' x\<bar> \<partial>M) = 0" using eq_0 by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2134
    thus ?thesis by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2135
  next
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2136
    assume neq_0: "(\<integral>\<^isup>+ x. Real (2 * w x) \<partial>M) \<noteq> 0"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2137
    have "(\<integral>\<^isup>+ x. Real (2 * w x) \<partial>M) = (\<integral>\<^isup>+ x. (SUP n. INF m. Real (?diff (m + n) x)) \<partial>M)"
41097
a1abfa4e2b44 use SUPR_ and INFI_apply instead of SUPR_, INFI_fun_expand
hoelzl
parents: 41096
diff changeset
  2138
    proof (rule positive_integral_cong, subst add_commute)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2139
      fix x assume x: "x \<in> space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2140
      show "Real (2 * w x) = (SUP n. INF m. Real (?diff (n + m) x))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2141
      proof (rule LIMSEQ_imp_lim_INF[symmetric])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2142
        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
  2143
      next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2144
        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
  2145
          using u'[OF x] by (safe intro!: LIMSEQ_diff LIMSEQ_const LIMSEQ_imp_rabs)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2146
        thus "(\<lambda>i. ?diff i x) ----> 2 * w x" by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2147
      qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2148
    qed
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2149
    also have "\<dots> \<le> (SUP n. INF m. (\<integral>\<^isup>+ x. Real (?diff (m + n) x) \<partial>M))"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2150
      using u'_borel w u unfolding integrable_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2151
      by (auto intro!: positive_integral_lim_INF)
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2152
    also have "\<dots> = (\<integral>\<^isup>+ x. Real (2 * w x) \<partial>M) -
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2153
        (INF n. SUP m. \<integral>\<^isup>+ x. Real \<bar>u (m + n) x - u' x\<bar> \<partial>M)"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40875
diff changeset
  2154
      unfolding PI_diff pextreal_INF_minus[OF I2w_fin] pextreal_SUP_minus ..
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40875
diff changeset
  2155
    finally show ?thesis using neq_0 I2w_fin by (rule pextreal_le_minus_imp_0)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2156
  qed
41705
1100512e16d8 add auto support for AE_mp
hoelzl
parents: 41689
diff changeset
  2157
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2158
  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
  2159
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2160
  have [simp]: "\<And>n m. (\<integral>\<^isup>+ x. Real \<bar>u (m + n) x - u' x\<bar> \<partial>M) =
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2161
    Real ((\<integral>x. \<bar>u (n + m) x - u' x\<bar> \<partial>M))"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2162
    using diff by (subst add_commute) (simp add: lebesgue_integral_def integrable_def Real_real)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2163
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2164
  have "(SUP n. INF m. (\<integral>\<^isup>+ x. Real \<bar>u (m + n) x - u' x\<bar> \<partial>M)) \<le> ?lim_SUP"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2165
    (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
  2166
  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
  2167
  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
  2168
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2169
  show ?lim
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2170
  proof (rule LIMSEQ_I)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2171
    fix r :: real assume "0 < r"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2172
    from LIMSEQ_D[OF `?lim_diff` this]
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2173
    obtain N where N: "\<And>n. N \<le> n \<Longrightarrow> (\<integral>x. \<bar>u n x - u' x\<bar> \<partial>M) < r"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2174
      using diff by (auto simp: integral_positive)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2175
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2176
    show "\<exists>N. \<forall>n\<ge>N. norm (integral\<^isup>L M (u n) - integral\<^isup>L M u') < r"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2177
    proof (safe intro!: exI[of _ N])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2178
      fix n assume "N \<le> n"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2179
      have "\<bar>integral\<^isup>L M (u n) - integral\<^isup>L M u'\<bar> = \<bar>(\<integral>x. u n x - u' x \<partial>M)\<bar>"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2180
        using u `integrable M u'` by (auto simp: integral_diff)
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2181
      also have "\<dots> \<le> (\<integral>x. \<bar>u n x - u' x\<bar> \<partial>M)" using u `integrable M u'`
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2182
        by (rule_tac integral_triangle_inequality) (auto intro!: integral_diff)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2183
      also note N[OF `N \<le> n`]
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2184
      finally show "norm (integral\<^isup>L M (u n) - integral\<^isup>L M u') < r" by simp
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2185
    qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2186
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2187
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2188
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2189
lemma (in measure_space) integral_sums:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2190
  assumes borel: "\<And>i. integrable M (f i)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2191
  and summable: "\<And>x. x \<in> space M \<Longrightarrow> summable (\<lambda>i. \<bar>f i x\<bar>)"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2192
  and sums: "summable (\<lambda>i. (\<integral>x. \<bar>f i x\<bar> \<partial>M))"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2193
  shows "integrable M (\<lambda>x. (\<Sum>i. f i x))" (is "integrable M ?S")
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2194
  and "(\<lambda>i. integral\<^isup>L M (f i)) sums (\<integral>x. (\<Sum>i. f i x) \<partial>M)" (is ?integral)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2195
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2196
  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
  2197
    using summable unfolding summable_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2198
  from bchoice[OF this]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2199
  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
  2200
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2201
  let "?w y" = "if y \<in> space M then w y else 0"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2202
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2203
  obtain x where abs_sum: "(\<lambda>i. (\<integral>x. \<bar>f i x\<bar> \<partial>M)) sums x"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2204
    using sums unfolding summable_def ..
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2205
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2206
  have 1: "\<And>n. integrable M (\<lambda>x. \<Sum>i = 0..<n. f i x)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2207
    using borel by (auto intro!: integral_setsum)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2208
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2209
  { fix j x assume [simp]: "x \<in> space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2210
    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
  2211
    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
  2212
    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
  2213
  note 2 = this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2214
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2215
  have 3: "integrable M ?w"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2216
  proof (rule integral_monotone_convergence(1))
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2217
    let "?F n y" = "(\<Sum>i = 0..<n. \<bar>f i y\<bar>)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2218
    let "?w' n y" = "if y \<in> space M then ?F n y else 0"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2219
    have "\<And>n. integrable M (?F n)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2220
      using borel by (auto intro!: integral_setsum integrable_abs)
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2221
    thus "\<And>n. integrable M (?w' n)" by (simp cong: integrable_cong)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2222
    show "mono ?w'"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2223
      by (auto simp: mono_def le_fun_def intro!: setsum_mono2)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2224
    { fix x show "(\<lambda>n. ?w' n x) ----> ?w x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2225
        using w by (cases "x \<in> space M") (simp_all add: LIMSEQ_const sums_def) }
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2226
    have *: "\<And>n. integral\<^isup>L M (?w' n) = (\<Sum>i = 0..< n. (\<integral>x. \<bar>f i x\<bar> \<partial>M))"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2227
      using borel by (simp add: integral_setsum integrable_abs cong: integral_cong)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2228
    from abs_sum
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2229
    show "(\<lambda>i. integral\<^isup>L M (?w' i)) ----> x" unfolding * sums_def .
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2230
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2231
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2232
  from summable[THEN summable_rabs_cancel]
41705
1100512e16d8 add auto support for AE_mp
hoelzl
parents: 41689
diff changeset
  2233
  have 4: "\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>n. \<Sum>i = 0..<n. f i x) ----> (\<Sum>i. f i x)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2234
    by (auto intro: summable_sumr_LIMSEQ_suminf)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2235
41705
1100512e16d8 add auto support for AE_mp
hoelzl
parents: 41689
diff changeset
  2236
  note int = integral_dominated_convergence(1,3)[OF 1 2 3 4]
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2237
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2238
  from int show "integrable M ?S" by simp
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2239
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2240
  show ?integral unfolding sums_def integral_setsum(1)[symmetric, OF borel]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2241
    using int(2) by simp
36624
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
  2242
qed
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
  2243
35748
5f35613d9a65 Equality of integral and infinite sum.
hoelzl
parents: 35692
diff changeset
  2244
section "Lebesgue integration on countable spaces"
5f35613d9a65 Equality of integral and infinite sum.
hoelzl
parents: 35692
diff changeset
  2245
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2246
lemma (in measure_space) integral_on_countable:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2247
  assumes f: "f \<in> borel_measurable M"
35748
5f35613d9a65 Equality of integral and infinite sum.
hoelzl
parents: 35692
diff changeset
  2248
  and bij: "bij_betw enum S (f ` space M)"
5f35613d9a65 Equality of integral and infinite sum.
hoelzl
parents: 35692
diff changeset
  2249
  and enum_zero: "enum ` (-S) \<subseteq> {0}"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2250
  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
  2251
  and abs_summable: "summable (\<lambda>r. \<bar>enum r * real (\<mu> (f -` {enum r} \<inter> space M))\<bar>)"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2252
  shows "integrable M f"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2253
  and "(\<lambda>r. enum r * real (\<mu> (f -` {enum r} \<inter> space M))) sums integral\<^isup>L M f" (is ?sums)
35748
5f35613d9a65 Equality of integral and infinite sum.
hoelzl
parents: 35692
diff changeset
  2254
proof -
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2255
  let "?A r" = "f -` {enum r} \<inter> space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2256
  let "?F r x" = "enum r * indicator (?A r) x"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2257
  have enum_eq: "\<And>r. enum r * real (\<mu> (?A r)) = integral\<^isup>L M (?F r)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2258
    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
  2259
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2260
  { fix x assume "x \<in> space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2261
    hence "f x \<in> enum ` S" using bij unfolding bij_betw_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2262
    then obtain i where "i\<in>S" "enum i = f x" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2263
    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
  2264
    proof cases
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2265
      fix j assume "j = i"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2266
      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
  2267
    next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2268
      fix j assume "j \<noteq> i"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2269
      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
  2270
        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
  2271
    qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2272
    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
  2273
    have "(\<lambda>i. ?F i x) sums f x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2274
         "(\<lambda>i. \<bar>?F i x\<bar>) sums \<bar>f x\<bar>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2275
      by (auto intro!: sums_single simp: F F_abs) }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2276
  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
  2277
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2278
  have int_f: "integral\<^isup>L M f = (\<integral>x. (\<Sum>r. ?F r x) \<partial>M)" "integrable M f = integrable M (\<lambda>x. \<Sum>r. ?F r x)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2279
    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
  2280
5f35613d9a65 Equality of integral and infinite sum.
hoelzl
parents: 35692
diff changeset
  2281
  { fix r
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2282
    have "(\<integral>x. \<bar>?F r x\<bar> \<partial>M) = (\<integral>x. \<bar>enum r\<bar> * indicator (?A r) x \<partial>M)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2283
      by (auto simp: indicator_def intro!: integral_cong)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2284
    also have "\<dots> = \<bar>enum r\<bar> * real (\<mu> (?A r))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2285
      using f fin by (simp add: borel_measurable_vimage integral_cmul_indicator)
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2286
    finally have "(\<integral>x. \<bar>?F r x\<bar> \<partial>M) = \<bar>enum r * real (\<mu> (?A r))\<bar>"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40875
diff changeset
  2287
      by (simp add: abs_mult_pos real_pextreal_pos) }
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2288
  note int_abs_F = this
35748
5f35613d9a65 Equality of integral and infinite sum.
hoelzl
parents: 35692
diff changeset
  2289
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2290
  have 1: "\<And>i. integrable M (\<lambda>x. ?F i x)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2291
    using f fin by (simp add: borel_measurable_vimage integral_cmul_indicator)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2292
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2293
  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
  2294
    using F_abs_sums_f unfolding sums_iff by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2295
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2296
  from integral_sums(2)[OF 1 2, unfolded int_abs_F, OF _ abs_summable]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2297
  show ?sums unfolding enum_eq int_f by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2298
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2299
  from integral_sums(1)[OF 1 2, unfolded int_abs_F, OF _ abs_summable]
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2300
  show "integrable M f" unfolding int_f by simp
35748
5f35613d9a65 Equality of integral and infinite sum.
hoelzl
parents: 35692
diff changeset
  2301
qed
5f35613d9a65 Equality of integral and infinite sum.
hoelzl
parents: 35692
diff changeset
  2302
35692
f1315bbf1bc9 Moved theorems in Lebesgue to the right places
hoelzl
parents: 35582
diff changeset
  2303
section "Lebesgue integration on finite space"
f1315bbf1bc9 Moved theorems in Lebesgue to the right places
hoelzl
parents: 35582
diff changeset
  2304
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2305
lemma (in measure_space) integral_on_finite:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2306
  assumes f: "f \<in> borel_measurable M" and finite: "finite (f`space M)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2307
  and fin: "\<And>x. x \<noteq> 0 \<Longrightarrow> \<mu> (f -` {x} \<inter> space M) \<noteq> \<omega>"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2308
  shows "integrable M f"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2309
  and "(\<integral>x. f x \<partial>M) =
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2310
    (\<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
  2311
proof -
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2312
  let "?A r" = "f -` {r} \<inter> space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2313
  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
  2314
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2315
  { fix x assume "x \<in> space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2316
    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
  2317
      using finite `x \<in> space M` by (simp add: setsum_cases)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2318
    also have "\<dots> = ?S x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2319
      by (auto intro!: setsum_cong)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2320
    finally have "f x = ?S x" . }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2321
  note f_eq = this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2322
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2323
  have f_eq_S: "integrable M f \<longleftrightarrow> integrable M ?S" "integral\<^isup>L M f = integral\<^isup>L M ?S"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2324
    by (auto intro!: integrable_cong integral_cong simp only: f_eq)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2325
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2326
  show "integrable M f" ?integral using fin f f_eq_S
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2327
    by (simp_all add: integral_cmul_indicator borel_measurable_vimage)
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  2328
qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  2329
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2330
lemma (in finite_measure_space) simple_function_finite[simp, intro]: "simple_function M f"
40875
9a9d33f6fb46 generalized simple_functionD
hoelzl
parents: 40873
diff changeset
  2331
  unfolding simple_function_def using finite_space by auto
35977
30d42bfd0174 Added finite measure space.
hoelzl
parents: 35833
diff changeset
  2332
38705
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 38656
diff changeset
  2333
lemma (in finite_measure_space) borel_measurable_finite[intro, simp]: "f \<in> borel_measurable M"
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
  2334
  by (auto intro: borel_measurable_simple_function)
38705
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 38656
diff changeset
  2335
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 38656
diff changeset
  2336
lemma (in finite_measure_space) positive_integral_finite_eq_setsum:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2337
  "integral\<^isup>P M f = (\<Sum>x \<in> space M. f x * \<mu> {x})"
38705
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 38656
diff changeset
  2338
proof -
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2339
  have *: "integral\<^isup>P M f = (\<integral>\<^isup>+ x. (\<Sum>y\<in>space M. f y * indicator {y} x) \<partial>M)"
38705
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 38656
diff changeset
  2340
    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
  2341
  show ?thesis unfolding * using borel_measurable_finite[of f]
40875
9a9d33f6fb46 generalized simple_functionD
hoelzl
parents: 40873
diff changeset
  2342
    by (simp add: positive_integral_setsum positive_integral_cmult_indicator)
38705
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 38656
diff changeset
  2343
qed
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 38656
diff changeset
  2344
35977
30d42bfd0174 Added finite measure space.
hoelzl
parents: 35833
diff changeset
  2345
lemma (in finite_measure_space) integral_finite_singleton:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2346
  shows "integrable M f"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2347
  and "integral\<^isup>L M f = (\<Sum>x \<in> space M. f x * real (\<mu> {x}))" (is ?I)
35977
30d42bfd0174 Added finite measure space.
hoelzl
parents: 35833
diff changeset
  2348
proof -
38705
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 38656
diff changeset
  2349
  have [simp]:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2350
    "(\<integral>\<^isup>+ x. Real (f x) \<partial>M) = (\<Sum>x \<in> space M. Real (f x) * \<mu> {x})"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2351
    "(\<integral>\<^isup>+ x. Real (- f x) \<partial>M) = (\<Sum>x \<in> space M. Real (- f x) * \<mu> {x})"
38705
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 38656
diff changeset
  2352
    unfolding positive_integral_finite_eq_setsum by auto
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2353
  show "integrable M f" using finite_space finite_measure
40875
9a9d33f6fb46 generalized simple_functionD
hoelzl
parents: 40873
diff changeset
  2354
    by (simp add: setsum_\<omega> integrable_def)
38705
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 38656
diff changeset
  2355
  show ?I using finite_measure
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2356
    apply (simp add: lebesgue_integral_def real_of_pextreal_setsum[symmetric]
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40875
diff changeset
  2357
      real_of_pextreal_mult[symmetric] setsum_subtractf[symmetric])
38705
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 38656
diff changeset
  2358
    by (rule setsum_cong) (simp_all split: split_if)
35977
30d42bfd0174 Added finite measure space.
hoelzl
parents: 35833
diff changeset
  2359
qed
30d42bfd0174 Added finite measure space.
hoelzl
parents: 35833
diff changeset
  2360
35748
5f35613d9a65 Equality of integral and infinite sum.
hoelzl
parents: 35692
diff changeset
  2361
end