src/HOL/Probability/Lebesgue_Integration.thy
author bulwahn
Sat, 25 Feb 2012 09:07:41 +0100
changeset 46671 3a40ea076230
parent 45342 5c760e1692b3
child 46731 5302e932d1e5
permissions -rw-r--r--
removing unnecessary assumptions in RComplete; simplifying proof in Probability
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
42067
66c8281349ec standardized headers
hoelzl
parents: 41981
diff changeset
     1
(*  Title:      HOL/Probability/Lebesgue_Integration.thy
66c8281349ec standardized headers
hoelzl
parents: 41981
diff changeset
     2
    Author:     Johannes Hölzl, TU München
66c8281349ec standardized headers
hoelzl
parents: 41981
diff changeset
     3
    Author:     Armin Heller, TU München
66c8281349ec standardized headers
hoelzl
parents: 41981
diff changeset
     4
*)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
     5
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
     6
header {*Lebesgue Integration*}
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
     7
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
     8
theory Lebesgue_Integration
43941
481566bc20e4 ereal is a complete_linorder instance
haftmann
parents: 43920
diff changeset
     9
  imports Measure Borel_Space
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
    10
begin
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
    11
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
    12
lemma real_ereal_1[simp]: "real (1::ereal) = 1"
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
    13
  unfolding one_ereal_def by simp
42991
3fa22920bf86 integral strong monotone; finite subadditivity for measure
hoelzl
parents: 42950
diff changeset
    14
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
    15
lemma ereal_indicator_pos[simp,intro]: "0 \<le> (indicator A x ::ereal)"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    16
  unfolding indicator_def by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    17
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    18
lemma tendsto_real_max:
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    19
  fixes x y :: real
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    20
  assumes "(X ---> x) net"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    21
  assumes "(Y ---> y) net"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    22
  shows "((\<lambda>x. max (X x) (Y x)) ---> max x y) net"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    23
proof -
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    24
  have *: "\<And>x y :: real. max x y = y + ((x - y) + norm (x - y)) / 2"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    25
    by (auto split: split_max simp: field_simps)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    26
  show ?thesis
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    27
    unfolding *
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    28
    by (intro tendsto_add assms tendsto_divide tendsto_norm tendsto_diff) auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    29
qed
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    30
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    31
lemma (in measure_space) measure_Union:
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    32
  assumes "finite S" "S \<subseteq> sets M" "\<And>A B. A \<in> S \<Longrightarrow> B \<in> S \<Longrightarrow> A \<noteq> B \<Longrightarrow> A \<inter> B = {}"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    33
  shows "setsum \<mu> S = \<mu> (\<Union>S)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    34
proof -
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    35
  have "setsum \<mu> S = \<mu> (\<Union>i\<in>S. i)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    36
    using assms by (intro measure_setsum[OF `finite S`]) (auto simp: disjoint_family_on_def)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    37
  also have "\<dots> = \<mu> (\<Union>S)" by (auto intro!: arg_cong[where f=\<mu>])
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    38
  finally show ?thesis .
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    39
qed
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    40
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    41
lemma (in sigma_algebra) measurable_sets2[intro]:
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    42
  assumes "f \<in> measurable M M'" "g \<in> measurable M M''"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    43
  and "A \<in> sets M'" "B \<in> sets M''"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    44
  shows "f -` A \<inter> g -` B \<inter> space M \<in> sets M"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    45
proof -
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    46
  have "f -` A \<inter> g -` B \<inter> space M = (f -` A \<inter> space M) \<inter> (g -` B \<inter> space M)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    47
    by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    48
  then show ?thesis using assms by (auto intro: measurable_sets)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    49
qed
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    50
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    51
lemma incseq_Suc_iff: "incseq f \<longleftrightarrow> (\<forall>n. f n \<le> f (Suc n))"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    52
proof
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    53
  assume "\<forall>n. f n \<le> f (Suc n)" then show "incseq f" by (auto intro!: incseq_SucI)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    54
qed (auto simp: incseq_def)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    55
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    56
lemma borel_measurable_real_floor:
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    57
  "(\<lambda>x::real. real \<lfloor>x\<rfloor>) \<in> borel_measurable borel"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    58
  unfolding borel.borel_measurable_iff_ge
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    59
proof (intro allI)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    60
  fix a :: real
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    61
  { fix x have "a \<le> real \<lfloor>x\<rfloor> \<longleftrightarrow> real \<lceil>a\<rceil> \<le> x"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    62
      using le_floor_eq[of "\<lceil>a\<rceil>" x] ceiling_le_iff[of a "\<lfloor>x\<rfloor>"]
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    63
      unfolding real_eq_of_int by simp }
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    64
  then have "{w::real \<in> space borel. a \<le> real \<lfloor>w\<rfloor>} = {real \<lceil>a\<rceil>..}" by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    65
  then show "{w::real \<in> space borel. a \<le> real \<lfloor>w\<rfloor>} \<in> sets borel" by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    66
qed
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    67
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    68
lemma measure_preservingD2:
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    69
  "f \<in> measure_preserving A B \<Longrightarrow> f \<in> measurable A B"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    70
  unfolding measure_preserving_def by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    71
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    72
lemma measure_preservingD3:
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    73
  "f \<in> measure_preserving A B \<Longrightarrow> f \<in> space A \<rightarrow> space B"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    74
  unfolding measure_preserving_def measurable_def by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    75
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    76
lemma measure_preservingD:
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    77
  "T \<in> measure_preserving A B \<Longrightarrow> X \<in> sets B \<Longrightarrow> measure A (T -` X \<inter> space A) = measure B X"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    78
  unfolding measure_preserving_def by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    79
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    80
lemma (in sigma_algebra) borel_measurable_real_natfloor[intro, simp]:
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    81
  assumes "f \<in> borel_measurable M"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    82
  shows "(\<lambda>x. real (natfloor (f x))) \<in> borel_measurable M"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    83
proof -
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    84
  have "\<And>x. real (natfloor (f x)) = max 0 (real \<lfloor>f x\<rfloor>)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    85
    by (auto simp: max_def natfloor_def)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    86
  with borel_measurable_max[OF measurable_comp[OF assms borel_measurable_real_floor] borel_measurable_const]
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    87
  show ?thesis by (simp add: comp_def)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    88
qed
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    89
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    90
lemma (in measure_space) AE_not_in:
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    91
  assumes N: "N \<in> null_sets" shows "AE x. x \<notin> N"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    92
  using N by (rule AE_I') auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    93
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    94
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
    95
  fixes f :: "nat \<Rightarrow> 'a::real_normed_vector"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    96
  assumes finite: "finite {r. P r}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    97
  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
    98
proof cases
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
    99
  assume "{r. P r} = {}" hence "\<And>r. \<not> P r" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   100
  thus ?thesis by (simp add: sums_zero)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   101
next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   102
  assume not_empty: "{r. P r} \<noteq> {}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   103
  have "?F sums (\<Sum>r = 0..< Suc (Max {r. P r}). ?F r)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   104
    by (rule series_zero)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   105
       (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
   106
  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
   107
    by (subst setsum_cases)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   108
       (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
   109
  finally show ?thesis .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   110
qed
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   111
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   112
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
   113
  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
   114
  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
   115
  using sums_If_finite[of "\<lambda>r. r = i" f] by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   116
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   117
section "Simple function"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   118
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   119
text {*
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   120
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   121
Our simple functions are not restricted to positive real numbers. Instead
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   122
they are just functions with a finite range and are measurable when singleton
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   123
sets are measurable.
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   124
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   125
*}
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   126
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   127
definition "simple_function M g \<longleftrightarrow>
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   128
    finite (g ` space M) \<and>
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   129
    (\<forall>x \<in> g ` space M. g -` {x} \<inter> space M \<in> sets M)"
36624
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   130
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   131
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
   132
  assumes "simple_function M g"
40875
9a9d33f6fb46 generalized simple_functionD
hoelzl
parents: 40873
diff changeset
   133
  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
   134
proof -
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
   135
  show "finite (g ` space M)"
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
   136
    using assms unfolding simple_function_def by auto
40875
9a9d33f6fb46 generalized simple_functionD
hoelzl
parents: 40873
diff changeset
   137
  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
   138
  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
   139
  finally show "g -` X \<inter> space M \<in> sets M" using assms
9a9d33f6fb46 generalized simple_functionD
hoelzl
parents: 40873
diff changeset
   140
    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
   141
qed
36624
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   142
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   143
lemma (in sigma_algebra) simple_function_measurable2[intro]:
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   144
  assumes "simple_function M f" "simple_function M g"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   145
  shows "f -` A \<inter> g -` B \<inter> space M \<in> sets M"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   146
proof -
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   147
  have "f -` A \<inter> g -` B \<inter> space M = (f -` A \<inter> space M) \<inter> (g -` B \<inter> space M)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   148
    by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   149
  then show ?thesis using assms[THEN simple_functionD(2)] by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   150
qed
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   151
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   152
lemma (in sigma_algebra) simple_function_indicator_representation:
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
   153
  fixes f ::"'a \<Rightarrow> ereal"
41689
3e39b0e730d6 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 f: "simple_function M f" and x: "x \<in> space M"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   155
  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
   156
  (is "?l = ?r")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   157
proof -
38705
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 38656
diff changeset
   158
  have "?r = (\<Sum>y \<in> f ` space M.
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   159
    (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
   160
    by (auto intro!: setsum_cong2)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   161
  also have "... =  f x *  indicator (f -` {f x} \<inter> space M) x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   162
    using assms by (auto dest: simple_functionD simp: setsum_delta)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   163
  also have "... = f x" using x by (auto simp: indicator_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   164
  finally show ?thesis by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   165
qed
36624
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   166
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   167
lemma (in measure_space) simple_function_notspace:
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
   168
  "simple_function M (\<lambda>x. h x * indicator (- space M) x::ereal)" (is "simple_function M ?h")
35692
f1315bbf1bc9 Moved theorems in Lebesgue to the right places
hoelzl
parents: 35582
diff changeset
   169
proof -
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   170
  have "?h ` space M \<subseteq> {0}" unfolding indicator_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   171
  hence [simp, intro]: "finite (?h ` space M)" by (auto intro: finite_subset)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   172
  have "?h -` {0} \<inter> space M = space M" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   173
  thus ?thesis unfolding simple_function_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   174
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   175
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   176
lemma (in sigma_algebra) simple_function_cong:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   177
  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
   178
  shows "simple_function M f \<longleftrightarrow> simple_function M g"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   179
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   180
  have "f ` space M = g ` space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   181
    "\<And>x. f -` {x} \<inter> space M = g -` {x} \<inter> space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   182
    using assms by (auto intro!: image_eqI)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   183
  thus ?thesis unfolding simple_function_def using assms by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   184
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   185
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   186
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
   187
  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
   188
  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
   189
  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
   190
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   191
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
   192
  assumes "simple_function M f"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   193
  shows "f \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   194
proof (rule borel_measurableI)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   195
  fix S
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   196
  let ?I = "f ` (f -` S \<inter> space M)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   197
  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
   198
  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
   199
    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
   200
    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
   201
  hence "?U \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   202
    apply (rule finite_UN)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   203
    using assms unfolding simple_function_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   204
  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
   205
qed
f1315bbf1bc9 Moved theorems in Lebesgue to the right places
hoelzl
parents: 35582
diff changeset
   206
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   207
lemma (in sigma_algebra) simple_function_borel_measurable:
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   208
  fixes f :: "'a \<Rightarrow> 'x::{t2_space}"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   209
  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
   210
  shows "simple_function M f"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   211
  using assms unfolding simple_function_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   212
  by (auto intro: borel_measurable_vimage)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   213
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   214
lemma (in sigma_algebra) simple_function_eq_borel_measurable:
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
   215
  fixes f :: "'a \<Rightarrow> ereal"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   216
  shows "simple_function M f \<longleftrightarrow> finite (f`space M) \<and> f \<in> borel_measurable M"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   217
  using simple_function_borel_measurable[of f]
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   218
    borel_measurable_simple_function[of f]
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 44666
diff changeset
   219
  by (fastforce simp: simple_function_def)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   220
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   221
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
   222
  "simple_function M (\<lambda>x. c)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   223
  by (auto intro: finite_subset simp: simple_function_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   224
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
   225
  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
   226
  shows "simple_function M (g \<circ> f)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   227
  unfolding simple_function_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   228
proof safe
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   229
  show "finite ((g \<circ> f) ` space M)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   230
    using assms unfolding simple_function_def by (auto simp: image_compose)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   231
next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   232
  fix x assume "x \<in> space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   233
  let ?G = "g -` {g (f x)} \<inter> (f`space M)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   234
  have *: "(g \<circ> f) -` {(g \<circ> f) x} \<inter> space M =
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   235
    (\<Union>x\<in>?G. f -` {x} \<inter> space M)" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   236
  show "(g \<circ> f) -` {(g \<circ> f) x} \<inter> space M \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   237
    using assms unfolding simple_function_def *
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   238
    by (rule_tac finite_UN) (auto intro!: finite_UN)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   239
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   240
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   241
lemma (in sigma_algebra) simple_function_indicator[intro, simp]:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   242
  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
   243
  shows "simple_function M (indicator A)"
35692
f1315bbf1bc9 Moved theorems in Lebesgue to the right places
hoelzl
parents: 35582
diff changeset
   244
proof -
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   245
  have "indicator A ` space M \<subseteq> {0, 1}" (is "?S \<subseteq> _")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   246
    by (auto simp: indicator_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   247
  hence "finite ?S" by (rule finite_subset) simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   248
  moreover have "- A \<inter> space M = space M - A" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   249
  ultimately show ?thesis unfolding simple_function_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   250
    using assms by (auto simp: indicator_def_raw)
35692
f1315bbf1bc9 Moved theorems in Lebesgue to the right places
hoelzl
parents: 35582
diff changeset
   251
qed
f1315bbf1bc9 Moved theorems in Lebesgue to the right places
hoelzl
parents: 35582
diff changeset
   252
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   253
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
   254
  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
   255
  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
   256
  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
   257
  unfolding simple_function_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   258
proof safe
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   259
  show "finite (?p ` space M)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   260
    using assms unfolding simple_function_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   261
    by (rule_tac finite_subset[of _ "f`space M \<times> g`space M"]) auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   262
next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   263
  fix x assume "x \<in> space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   264
  have "(\<lambda>x. (f x, g x)) -` {(f x, g x)} \<inter> space M =
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   265
      (f -` {f x} \<inter> space M) \<inter> (g -` {g x} \<inter> space M)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   266
    by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   267
  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
   268
    using assms unfolding simple_function_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   269
qed
35692
f1315bbf1bc9 Moved theorems in Lebesgue to the right places
hoelzl
parents: 35582
diff changeset
   270
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   271
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
   272
  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
   273
  shows "simple_function M (\<lambda>x. g (f x))"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   274
  using simple_function_compose[OF assms, of g]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   275
  by (simp add: comp_def)
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   276
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   277
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
   278
  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
   279
  shows "simple_function M (\<lambda>x. h (f x) (g x))"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   280
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
   281
  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
   282
    using assms by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   283
  thus ?thesis by (simp_all add: comp_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   284
qed
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   285
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   286
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
   287
  and simple_function_diff[intro, simp] = simple_function_compose2[where h="op -"]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   288
  and simple_function_uminus[intro, simp] = simple_function_compose[where g="uminus"]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   289
  and simple_function_mult[intro, simp] = simple_function_compose2[where h="op *"]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   290
  and simple_function_div[intro, simp] = simple_function_compose2[where h="op /"]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   291
  and simple_function_inverse[intro, simp] = simple_function_compose[where g="inverse"]
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   292
  and simple_function_max[intro, simp] = simple_function_compose2[where h=max]
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   293
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   294
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
   295
  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
   296
  shows "simple_function M (\<lambda>x. \<Sum>i\<in>P. f i x)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   297
proof cases
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   298
  assume "finite P" from this assms show ?thesis by induct auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   299
qed auto
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   300
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   301
lemma (in sigma_algebra)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   302
  fixes f g :: "'a \<Rightarrow> real" assumes sf: "simple_function M f"
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
   303
  shows simple_function_ereal[intro, simp]: "simple_function M (\<lambda>x. ereal (f x))"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   304
  by (auto intro!: simple_function_compose1[OF sf])
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   305
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   306
lemma (in sigma_algebra)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   307
  fixes f g :: "'a \<Rightarrow> nat" assumes sf: "simple_function M f"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   308
  shows simple_function_real_of_nat[intro, simp]: "simple_function M (\<lambda>x. real (f x))"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   309
  by (auto intro!: simple_function_compose1[OF sf])
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   310
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   311
lemma (in sigma_algebra) borel_measurable_implies_simple_function_sequence:
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
   312
  fixes u :: "'a \<Rightarrow> ereal"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   313
  assumes u: "u \<in> borel_measurable M"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   314
  shows "\<exists>f. incseq f \<and> (\<forall>i. \<infinity> \<notin> range (f i) \<and> simple_function M (f i)) \<and>
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   315
             (\<forall>x. (SUP i. f i x) = max 0 (u x)) \<and> (\<forall>i x. 0 \<le> f i x)"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   316
proof -
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   317
  def f \<equiv> "\<lambda>x i. if real i \<le> u x then i * 2 ^ i else natfloor (real (u x) * 2 ^ i)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   318
  { fix x j have "f x j \<le> j * 2 ^ j" unfolding f_def
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   319
    proof (split split_if, intro conjI impI)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   320
      assume "\<not> real j \<le> u x"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   321
      then have "natfloor (real (u x) * 2 ^ j) \<le> natfloor (j * 2 ^ j)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   322
         by (cases "u x") (auto intro!: natfloor_mono simp: mult_nonneg_nonneg)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   323
      moreover have "real (natfloor (j * 2 ^ j)) \<le> j * 2^j"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   324
        by (intro real_natfloor_le) (auto simp: mult_nonneg_nonneg)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   325
      ultimately show "natfloor (real (u x) * 2 ^ j) \<le> j * 2 ^ j"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   326
        unfolding real_of_nat_le_iff by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   327
    qed auto }
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   328
  note f_upper = this
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   329
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   330
  have real_f:
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   331
    "\<And>i x. real (f x i) = (if real i \<le> u x then i * 2 ^ i else real (natfloor (real (u x) * 2 ^ i)))"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   332
    unfolding f_def by auto
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   333
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
   334
  let "?g j x" = "real (f x j) / 2^j :: ereal"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   335
  show ?thesis
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   336
  proof (intro exI[of _ ?g] conjI allI ballI)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   337
    fix i
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   338
    have "simple_function M (\<lambda>x. real (f x i))"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   339
    proof (intro simple_function_borel_measurable)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   340
      show "(\<lambda>x. real (f x i)) \<in> borel_measurable M"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   341
        using u by (auto intro!: measurable_If simp: real_f)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   342
      have "(\<lambda>x. real (f x i))`space M \<subseteq> real`{..i*2^i}"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   343
        using f_upper[of _ i] by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   344
      then show "finite ((\<lambda>x. real (f x i))`space M)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   345
        by (rule finite_subset) auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   346
    qed
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   347
    then show "simple_function M (?g i)"
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
   348
      by (auto intro: simple_function_ereal simple_function_div)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   349
  next
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   350
    show "incseq ?g"
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
   351
    proof (intro incseq_ereal incseq_SucI le_funI)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   352
      fix x and i :: nat
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   353
      have "f x i * 2 \<le> f x (Suc i)" unfolding f_def
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   354
      proof ((split split_if)+, intro conjI impI)
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
   355
        assume "ereal (real i) \<le> u x" "\<not> ereal (real (Suc i)) \<le> u x"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   356
        then show "i * 2 ^ i * 2 \<le> natfloor (real (u x) * 2 ^ Suc i)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   357
          by (cases "u x") (auto intro!: le_natfloor)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   358
      next
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
   359
        assume "\<not> ereal (real i) \<le> u x" "ereal (real (Suc i)) \<le> u x"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   360
        then show "natfloor (real (u x) * 2 ^ i) * 2 \<le> Suc i * 2 ^ Suc i"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   361
          by (cases "u x") auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   362
      next
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
   363
        assume "\<not> ereal (real i) \<le> u x" "\<not> ereal (real (Suc i)) \<le> u x"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   364
        have "natfloor (real (u x) * 2 ^ i) * 2 = natfloor (real (u x) * 2 ^ i) * natfloor 2"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   365
          by simp
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   366
        also have "\<dots> \<le> natfloor (real (u x) * 2 ^ i * 2)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   367
        proof cases
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   368
          assume "0 \<le> u x" then show ?thesis
46671
3a40ea076230 removing unnecessary assumptions in RComplete;
bulwahn
parents: 45342
diff changeset
   369
            by (intro le_mult_natfloor) 
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   370
        next
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   371
          assume "\<not> 0 \<le> u x" then show ?thesis
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   372
            by (cases "u x") (auto simp: natfloor_neg mult_nonpos_nonneg)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   373
        qed
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   374
        also have "\<dots> = natfloor (real (u x) * 2 ^ Suc i)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   375
          by (simp add: ac_simps)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   376
        finally show "natfloor (real (u x) * 2 ^ i) * 2 \<le> natfloor (real (u x) * 2 ^ Suc i)" .
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   377
      qed simp
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   378
      then show "?g i x \<le> ?g (Suc i) x"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   379
        by (auto simp: field_simps)
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   380
    qed
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   381
  next
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   382
    fix x show "(SUP i. ?g i x) = max 0 (u x)"
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
   383
    proof (rule ereal_SUPI)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   384
      fix i show "?g i x \<le> max 0 (u x)" unfolding max_def f_def
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   385
        by (cases "u x") (auto simp: field_simps real_natfloor_le natfloor_neg
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   386
                                     mult_nonpos_nonneg mult_nonneg_nonneg)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   387
    next
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   388
      fix y assume *: "\<And>i. i \<in> UNIV \<Longrightarrow> ?g i x \<le> y"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   389
      have "\<And>i. 0 \<le> ?g i x" by (auto simp: divide_nonneg_pos)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   390
      from order_trans[OF this *] have "0 \<le> y" by simp
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   391
      show "max 0 (u x) \<le> y"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   392
      proof (cases y)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   393
        case (real r)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   394
        with * have *: "\<And>i. f x i \<le> r * 2^i" by (auto simp: divide_le_eq)
44666
8670a39d4420 remove more duplicate lemmas
huffman
parents: 44568
diff changeset
   395
        from reals_Archimedean2[of r] * have "u x \<noteq> \<infinity>" by (auto simp: f_def) (metis less_le_not_le)
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
   396
        then have "\<exists>p. max 0 (u x) = ereal p \<and> 0 \<le> p" by (cases "u x") (auto simp: max_def)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   397
        then guess p .. note ux = this
44666
8670a39d4420 remove more duplicate lemmas
huffman
parents: 44568
diff changeset
   398
        obtain m :: nat where m: "p < real m" using reals_Archimedean2 ..
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   399
        have "p \<le> r"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   400
        proof (rule ccontr)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   401
          assume "\<not> p \<le> r"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   402
          with LIMSEQ_inverse_realpow_zero[unfolded LIMSEQ_iff, rule_format, of 2 "p - r"]
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   403
          obtain N where "\<forall>n\<ge>N. r * 2^n < p * 2^n - 1" by (auto simp: inverse_eq_divide field_simps)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   404
          then have "r * 2^max N m < p * 2^max N m - 1" by simp
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   405
          moreover
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   406
          have "real (natfloor (p * 2 ^ max N m)) \<le> r * 2 ^ max N m"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   407
            using *[of "max N m"] m unfolding real_f using ux
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   408
            by (cases "0 \<le> u x") (simp_all add: max_def mult_nonneg_nonneg split: split_if_asm)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   409
          then have "p * 2 ^ max N m - 1 < r * 2 ^ max N m"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   410
            by (metis real_natfloor_gt_diff_one less_le_trans)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   411
          ultimately show False by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   412
        qed
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   413
        then show "max 0 (u x) \<le> y" using real ux by simp
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   414
      qed (insert `0 \<le> y`, auto)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   415
    qed
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   416
  qed (auto simp: divide_nonneg_pos)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   417
qed
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   418
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   419
lemma (in sigma_algebra) borel_measurable_implies_simple_function_sequence':
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
   420
  fixes u :: "'a \<Rightarrow> ereal"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   421
  assumes u: "u \<in> borel_measurable M"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   422
  obtains f where "\<And>i. simple_function M (f i)" "incseq f" "\<And>i. \<infinity> \<notin> range (f i)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   423
    "\<And>x. (SUP i. f i x) = max 0 (u x)" "\<And>i x. 0 \<le> f i x"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   424
  using borel_measurable_implies_simple_function_sequence[OF u] by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   425
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   426
lemma (in sigma_algebra) simple_function_If_set:
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   427
  assumes sf: "simple_function M f" "simple_function M g" and A: "A \<inter> space M \<in> sets M"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   428
  shows "simple_function M (\<lambda>x. if x \<in> A then f x else g x)" (is "simple_function M ?IF")
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   429
proof -
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   430
  def F \<equiv> "\<lambda>x. f -` {x} \<inter> space M" and G \<equiv> "\<lambda>x. g -` {x} \<inter> space M"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   431
  show ?thesis unfolding simple_function_def
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   432
  proof safe
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   433
    have "?IF ` space M \<subseteq> f ` space M \<union> g ` space M" by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   434
    from finite_subset[OF this] assms
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   435
    show "finite (?IF ` space M)" unfolding simple_function_def by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   436
  next
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   437
    fix x assume "x \<in> space M"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   438
    then have *: "?IF -` {?IF x} \<inter> space M = (if x \<in> A
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   439
      then ((F (f x) \<inter> (A \<inter> space M)) \<union> (G (f x) - (G (f x) \<inter> (A \<inter> space M))))
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   440
      else ((F (g x) \<inter> (A \<inter> space M)) \<union> (G (g x) - (G (g x) \<inter> (A \<inter> space M)))))"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   441
      using sets_into_space[OF A] by (auto split: split_if_asm simp: G_def F_def)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   442
    have [intro]: "\<And>x. F x \<in> sets M" "\<And>x. G x \<in> sets M"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   443
      unfolding F_def G_def using sf[THEN simple_functionD(2)] by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   444
    show "?IF -` {?IF x} \<inter> space M \<in> sets M" unfolding * using A by auto
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   445
  qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   446
qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   447
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   448
lemma (in sigma_algebra) simple_function_If:
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   449
  assumes sf: "simple_function M f" "simple_function M g" and P: "{x\<in>space M. P x} \<in> sets M"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   450
  shows "simple_function M (\<lambda>x. if P x then f x else g x)"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   451
proof -
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   452
  have "{x\<in>space M. P x} = {x. P x} \<inter> space M" by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   453
  with simple_function_If_set[OF sf, of "{x. P x}"] P show ?thesis by simp
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   454
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   455
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   456
lemma (in measure_space) simple_function_restricted:
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
   457
  fixes f :: "'a \<Rightarrow> ereal" 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
   458
  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
   459
    (is "simple_function ?R f \<longleftrightarrow> simple_function M ?f")
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   460
proof -
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   461
  interpret R: sigma_algebra ?R by (rule restricted_sigma_algebra[OF `A \<in> sets M`])
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   462
  have f: "finite (f`A) \<longleftrightarrow> finite (?f`space M)"
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   463
  proof cases
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   464
    assume "A = space M"
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 44666
diff changeset
   465
    then have "f`A = ?f`space M" by (fastforce simp: image_iff)
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   466
    then show ?thesis by simp
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   467
  next
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   468
    assume "A \<noteq> space M"
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   469
    then obtain x where x: "x \<in> space M" "x \<notin> A"
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   470
      using sets_into_space `A \<in> sets M` by auto
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   471
    have *: "?f`space M = f`A \<union> {0}"
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   472
    proof (auto simp add: image_iff)
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   473
      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
   474
        using x by (auto intro!: bexI[of _ x])
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   475
    next
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   476
      fix x assume "x \<in> A"
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   477
      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
   478
        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
   479
    next
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   480
      fix x
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
   481
      assume "indicator A x \<noteq> (0::ereal)"
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   482
      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
   483
      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
   484
      ultimately show "f x = 0" by auto
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   485
    qed
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   486
    then show ?thesis by auto
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   487
  qed
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   488
  then show ?thesis
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   489
    unfolding simple_function_eq_borel_measurable
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   490
      R.simple_function_eq_borel_measurable
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   491
    unfolding borel_measurable_restricted[OF `A \<in> sets M`]
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   492
    using assms(1)[THEN sets_into_space]
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   493
    by (auto simp: indicator_def)
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   494
qed
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   495
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   496
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
   497
  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
   498
  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
   499
  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
   500
  using assms unfolding simple_function_def by auto
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   501
41661
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41545
diff changeset
   502
lemma (in measure_space) simple_function_vimage:
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41545
diff changeset
   503
  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
   504
    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
   505
  shows "simple_function M (\<lambda>x. f (T x))"
41661
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41545
diff changeset
   506
proof (intro simple_function_def[THEN iffD2] conjI ballI)
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41545
diff changeset
   507
  interpret T: sigma_algebra M' by fact
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41545
diff changeset
   508
  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
   509
    using T unfolding measurable_def by auto
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41545
diff changeset
   510
  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
   511
    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
   512
  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
   513
  then have "i \<in> f ` space M'"
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41545
diff changeset
   514
    using T unfolding measurable_def by auto
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41545
diff changeset
   515
  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
   516
    using f unfolding simple_function_def by auto
41661
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41545
diff changeset
   517
  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
   518
    using T unfolding measurable_def by auto
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41545
diff changeset
   519
  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
   520
    using T unfolding measurable_def by auto
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41545
diff changeset
   521
  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
   522
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   523
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   524
section "Simple integral"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   525
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   526
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
   527
  "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
   528
3e39b0e730d6 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
syntax
45342
5c760e1692b3 proper syntactic category for abstraction syntax, to avoid low-level exception for malformed "\<integral> x y. f \<partial>M", for example;
wenzelm
parents: 44937
diff changeset
   530
  "_simple_integral" :: "pttrn \<Rightarrow> ereal \<Rightarrow> ('a, 'b) measure_space_scheme \<Rightarrow> ereal" ("\<integral>\<^isup>S _. _ \<partial>_" [60,61] 110)
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   531
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   532
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
   533
  "\<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
   534
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   535
lemma (in measure_space) simple_integral_cong:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   536
  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
   537
  shows "integral\<^isup>S M f = integral\<^isup>S M g"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   538
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   539
  have "f ` space M = g ` space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   540
    "\<And>x. f -` {x} \<inter> space M = g -` {x} \<inter> space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   541
    using assms by (auto intro!: image_eqI)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   542
  thus ?thesis unfolding simple_integral_def by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   543
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   544
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   545
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
   546
  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
   547
    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
   548
  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
   549
proof -
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   550
  interpret v: measure_space N
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   551
    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
   552
  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
   553
    by (auto intro!: setsum_cong simp: simple_integral_def)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   554
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   555
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   556
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
   557
  "(\<integral>\<^isup>Sx. c \<partial>M) = c * \<mu> (space M)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   558
proof (cases "space M = {}")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   559
  case True thus ?thesis unfolding simple_integral_def by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   560
next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   561
  case False hence "(\<lambda>x. c) ` space M = {c}" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   562
  thus ?thesis unfolding simple_integral_def by simp
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   563
qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   564
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   565
lemma (in measure_space) simple_function_partition:
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   566
  assumes f: "simple_function M f" and g: "simple_function M g"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   567
  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
   568
    (is "_ = setsum _ (?p ` space M)")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   569
proof-
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   570
  let "?sub x" = "?p ` (f -` {x} \<inter> space M)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   571
  let ?SIGMA = "Sigma (f`space M) ?sub"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   572
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   573
  have [intro]:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   574
    "finite (f ` space M)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   575
    "finite (g ` space M)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   576
    using assms unfolding simple_function_def by simp_all
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   577
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   578
  { fix A
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   579
    have "?p ` (A \<inter> space M) \<subseteq>
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   580
      (\<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
   581
      by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   582
    hence "finite (?p ` (A \<inter> space M))"
40786
0a54cfc9add3 gave more standard finite set rules simp and intro attribute
nipkow
parents: 39910
diff changeset
   583
      by (rule finite_subset) auto }
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   584
  note this[intro, simp]
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   585
  note sets = simple_function_measurable2[OF f g]
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   586
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   587
  { fix x assume "x \<in> space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   588
    have "\<Union>(?sub (f x)) = (f -` {f x} \<inter> space M)" by auto
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   589
    with sets have "\<mu> (f -` {f x} \<inter> space M) = setsum \<mu> (?sub (f x))"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   590
      by (subst measure_Union) 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
   591
  hence "integral\<^isup>S M f = (\<Sum>(x,A)\<in>?SIGMA. x * \<mu> A)"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   592
    unfolding simple_integral_def using f sets
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   593
    by (subst setsum_Sigma[symmetric])
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
   594
       (auto intro!: setsum_cong setsum_ereal_right_distrib)
39910
10097e0a9dbd constant `contents` renamed to `the_elem`
haftmann
parents: 39302
diff changeset
   595
  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
   596
  proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   597
    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
   598
    have "(\<lambda>A. (the_elem (f ` A), A)) ` ?p ` space M
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   599
      = (\<lambda>x. (f x, ?p x)) ` space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   600
    proof safe
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   601
      fix x assume "x \<in> space M"
39910
10097e0a9dbd constant `contents` renamed to `the_elem`
haftmann
parents: 39302
diff changeset
   602
      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
   603
        by (auto intro!: image_eqI[of _ _ "?p x"])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   604
    qed auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   605
    thus ?thesis
39910
10097e0a9dbd constant `contents` renamed to `the_elem`
haftmann
parents: 39302
diff changeset
   606
      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
   607
      apply (rule_tac x="xa" in image_eqI)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   608
      by simp_all
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   609
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   610
  finally show ?thesis .
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   611
qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   612
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   613
lemma (in measure_space) simple_integral_add[simp]:
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   614
  assumes f: "simple_function M f" and "\<And>x. 0 \<le> f x" and g: "simple_function M g" and "\<And>x. 0 \<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
   615
  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
   616
proof -
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   617
  { fix x let ?S = "g -` {g x} \<inter> f -` {f x} \<inter> space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   618
    assume "x \<in> space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   619
    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
   620
        "(\<lambda>x. (f x, g x)) -` {(f x, g x)} \<inter> space M = ?S"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   621
      by auto }
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   622
  with assms show ?thesis
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   623
    unfolding
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   624
      simple_function_partition[OF simple_function_add[OF f g] simple_function_Pair[OF f g]]
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   625
      simple_function_partition[OF f g]
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   626
      simple_function_partition[OF g f]
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   627
    by (subst (3) Int_commute)
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
   628
       (auto simp add: ereal_left_distrib setsum_addf[symmetric] intro!: setsum_cong)
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   629
qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   630
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   631
lemma (in measure_space) simple_integral_setsum[simp]:
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   632
  assumes "\<And>i x. i \<in> P \<Longrightarrow> 0 \<le> f i 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
   633
  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
   634
  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
   635
proof cases
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   636
  assume "finite P"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   637
  from this assms show ?thesis
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   638
    by induct (auto simp: simple_function_setsum simple_integral_add setsum_nonneg)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   639
qed auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   640
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   641
lemma (in measure_space) simple_integral_mult[simp]:
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   642
  assumes f: "simple_function M f" "\<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
   643
  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
   644
proof -
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   645
  note mult = simple_function_mult[OF simple_function_const[of c] f(1)]
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   646
  { 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
   647
    assume "x \<in> space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   648
    hence "(\<lambda>x. c * f x) ` ?S = {c * f x}" "f ` ?S = {f x}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   649
      by auto }
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   650
  with assms show ?thesis
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   651
    unfolding simple_function_partition[OF mult f(1)]
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   652
              simple_function_partition[OF f(1) mult]
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
   653
    by (subst setsum_ereal_right_distrib)
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
   654
       (auto intro!: ereal_0_le_mult setsum_cong simp: mult_assoc)
40871
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
   655
qed
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
   656
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   657
lemma (in measure_space) simple_integral_mono_AE:
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   658
  assumes f: "simple_function M f" and g: "simple_function M g"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   659
  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
   660
  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
   661
proof -
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   662
  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
   663
  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
   664
    "\<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
   665
  show ?thesis
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   666
    unfolding *
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   667
      simple_function_partition[OF f g]
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   668
      simple_function_partition[OF g f]
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   669
  proof (safe intro!: setsum_mono)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   670
    fix x assume "x \<in> space M"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   671
    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
   672
    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
   673
    proof (cases "f x \<le> g x")
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   674
      case True then show ?thesis
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   675
        using * assms(1,2)[THEN simple_functionD(2)]
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
   676
        by (auto intro!: ereal_mult_right_mono)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   677
    next
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   678
      case False
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   679
      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
   680
        using mono by (auto elim!: AE_E)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   681
      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
   682
      moreover have "?S x \<in> sets M" using assms
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
   683
        by (rule_tac Int) (auto intro!: simple_functionD)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   684
      ultimately have "\<mu> (?S x) \<le> \<mu> N"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   685
        using `N \<in> sets M` by (auto intro!: measure_mono)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   686
      moreover have "0 \<le> \<mu> (?S x)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   687
        using assms(1,2)[THEN simple_functionD(2)] by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   688
      ultimately have "\<mu> (?S x) = 0" using `\<mu> N = 0` by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   689
      then show ?thesis by simp
40859
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
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   693
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   694
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
   695
  assumes "simple_function M f" and "simple_function M g"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   696
  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
   697
  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
   698
  using assms by (intro simple_integral_mono_AE) auto
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   699
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   700
lemma (in measure_space) simple_integral_cong_AE:
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   701
  assumes "simple_function M f" and "simple_function M g"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   702
  and "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
   703
  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
   704
  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
   705
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   706
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
   707
  assumes sf: "simple_function M f" "simple_function M g"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   708
  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
   709
  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
   710
proof (intro simple_integral_cong_AE sf AE_I)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   711
  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
   712
  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
   713
    using sf[THEN borel_measurable_simple_function] by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   714
qed simp
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   715
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   716
lemma (in measure_space) simple_integral_indicator:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   717
  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
   718
  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
   719
  shows "(\<integral>\<^isup>Sx. f x * indicator A x \<partial>M) =
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   720
    (\<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
   721
proof cases
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   722
  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
   723
  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
   724
    by (auto intro!: simple_integral_cong)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   725
  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
   726
  ultimately show ?thesis by (simp add: simple_integral_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   727
next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   728
  assume "A \<noteq> space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   729
  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
   730
  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
   731
  proof safe
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   732
    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
   733
  next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   734
    fix y assume "y \<in> A" thus "f y \<in> ?I ` space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   735
      using sets_into_space[OF assms(1)] by (auto intro!: image_eqI[of _ _ y])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   736
  next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   737
    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
   738
  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
   739
  have *: "(\<integral>\<^isup>Sx. f x * indicator A x \<partial>M) =
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   740
    (\<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
   741
    unfolding simple_integral_def I
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   742
  proof (rule setsum_mono_zero_cong_left)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   743
    show "finite (f ` space M \<union> {0})"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   744
      using assms(2) unfolding simple_function_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   745
    show "f ` A \<union> {0} \<subseteq> f`space M \<union> {0}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   746
      using sets_into_space[OF assms(1)] by auto
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   747
    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
   748
      by (auto simp: image_iff)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   749
    thus "\<forall>i\<in>f ` space M \<union> {0} - (f ` A \<union> {0}).
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   750
      i * \<mu> (f -` {i} \<inter> space M \<inter> A) = 0" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   751
  next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   752
    fix x assume "x \<in> f`A \<union> {0}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   753
    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
   754
      by (auto simp: indicator_def split: split_if_asm)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   755
    thus "x * \<mu> (?I -` {x} \<inter> space M) =
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   756
      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
   757
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   758
  show ?thesis unfolding *
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   759
    using assms(2) unfolding simple_function_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   760
    by (auto intro!: setsum_mono_zero_cong_right)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   761
qed
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   762
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   763
lemma (in measure_space) simple_integral_indicator_only[simp]:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   764
  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
   765
  shows "integral\<^isup>S M (indicator A) = \<mu> A"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   766
proof cases
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   767
  assume "space M = {}" hence "A = {}" using sets_into_space[OF assms] by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   768
  thus ?thesis unfolding simple_integral_def using `space M = {}` by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   769
next
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
   770
  assume "space M \<noteq> {}" hence "(\<lambda>x. 1) ` space M = {1::ereal}" by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   771
  thus ?thesis
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   772
    using simple_integral_indicator[OF assms simple_function_const[of 1]]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   773
    using sets_into_space[OF assms]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   774
    by (auto intro!: arg_cong[where f="\<mu>"])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   775
qed
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   776
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   777
lemma (in measure_space) simple_integral_null_set:
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   778
  assumes "simple_function M u" "\<And>x. 0 \<le> u x" and "N \<in> null_sets"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   779
  shows "(\<integral>\<^isup>Sx. u x * indicator N x \<partial>M) = 0"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   780
proof -
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
   781
  have "AE x. indicator N x = (0 :: ereal)"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   782
    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
   783
  then have "(\<integral>\<^isup>Sx. u x * indicator N x \<partial>M) = (\<integral>\<^isup>Sx. 0 \<partial>M)"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   784
    using assms apply (intro simple_integral_cong_AE) by auto
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   785
  then show ?thesis by simp
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   786
qed
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   787
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   788
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
   789
  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
   790
  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
   791
  using assms by (intro simple_integral_cong_AE) auto
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   792
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   793
lemma (in measure_space) simple_integral_restricted:
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   794
  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
   795
  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
   796
  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
   797
    (is "_ = integral\<^isup>S M ?f")
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   798
  unfolding simple_integral_def
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   799
proof (simp, safe intro!: setsum_mono_zero_cong_left)
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   800
  from sf show "finite (?f ` space M)"
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   801
    unfolding simple_function_def by auto
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   802
next
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   803
  fix x assume "x \<in> A"
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   804
  then show "f x \<in> ?f ` space M"
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   805
    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
   806
next
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   807
  fix x assume "x \<in> space M" "?f x \<notin> f`A"
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   808
  then have "x \<notin> A" by (auto simp: image_iff)
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   809
  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
   810
next
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   811
  fix x assume "x \<in> A"
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   812
  then have "f x \<noteq> 0 \<Longrightarrow>
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   813
    f -` {f x} \<inter> A = ?f -` {f x} \<inter> space M"
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   814
    using `A \<in> sets M` sets_into_space
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   815
    by (auto simp: indicator_def split: split_if_asm)
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   816
  then show "f x * \<mu> (f -` {f x} \<inter> A) =
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   817
    f x * \<mu> (?f -` {f x} \<inter> space M)"
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
   818
    unfolding ereal_mult_cancel_left by auto
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   819
qed
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   820
41545
9c869baf1c66 tuned formalization of subalgebra
hoelzl
parents: 41544
diff changeset
   821
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
   822
  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
   823
  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
   824
  unfolding simple_integral_def_raw by simp
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   825
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   826
lemma (in measure_space) simple_integral_vimage:
41831
91a2b435dd7a use measure_preserving in ..._vimage lemmas
hoelzl
parents: 41705
diff changeset
   827
  assumes T: "sigma_algebra M'" "T \<in> measure_preserving 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
   828
    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
   829
  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
   830
proof -
41831
91a2b435dd7a use measure_preserving in ..._vimage lemmas
hoelzl
parents: 41705
diff changeset
   831
  interpret T: measure_space M' by (rule measure_space_vimage[OF 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
   832
  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
   833
    unfolding simple_integral_def
41661
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41545
diff changeset
   834
  proof (intro setsum_mono_zero_cong_right ballI)
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41545
diff changeset
   835
    show "(\<lambda>x. f (T x)) ` space M \<subseteq> f ` space M'"
41831
91a2b435dd7a use measure_preserving in ..._vimage lemmas
hoelzl
parents: 41705
diff changeset
   836
      using T unfolding measurable_def measure_preserving_def by auto
41661
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41545
diff changeset
   837
    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
   838
      using f unfolding simple_function_def by auto
41661
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41545
diff changeset
   839
  next
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41545
diff changeset
   840
    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
   841
    then have "T -` (f -` {i} \<inter> space M') \<inter> space M = {}" by (auto simp: image_iff)
41831
91a2b435dd7a use measure_preserving in ..._vimage lemmas
hoelzl
parents: 41705
diff changeset
   842
    with f[THEN T.simple_functionD(2), THEN measure_preservingD[OF T(2)], of "{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
   843
    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
   844
  next
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41545
diff changeset
   845
    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
   846
    then have "T -` (f -` {i} \<inter> space M') \<inter> space M = (\<lambda>x. f (T x)) -` {i} \<inter> space M"
41831
91a2b435dd7a use measure_preserving in ..._vimage lemmas
hoelzl
parents: 41705
diff changeset
   847
      using T unfolding measurable_def measure_preserving_def by auto
91a2b435dd7a use measure_preserving in ..._vimage lemmas
hoelzl
parents: 41705
diff changeset
   848
    with f[THEN T.simple_functionD(2), THEN measure_preservingD[OF T(2)], of "{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
   849
    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
   850
      by auto
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41545
diff changeset
   851
  qed
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   852
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   853
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   854
lemma (in measure_space) simple_integral_cmult_indicator:
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   855
  assumes A: "A \<in> sets M"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   856
  shows "(\<integral>\<^isup>Sx. c * indicator A x \<partial>M) = c * \<mu> A"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   857
  using simple_integral_mult[OF simple_function_indicator[OF A]]
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   858
  unfolding simple_integral_indicator_only[OF A] by simp
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   859
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   860
lemma (in measure_space) simple_integral_positive:
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   861
  assumes f: "simple_function M f" and ae: "AE x. 0 \<le> f x"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   862
  shows "0 \<le> integral\<^isup>S M f"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   863
proof -
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   864
  have "integral\<^isup>S M (\<lambda>x. 0) \<le> integral\<^isup>S M f"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   865
    using simple_integral_mono_AE[OF _ f ae] by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   866
  then show ?thesis by simp
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   867
qed
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   868
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   869
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
   870
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   871
definition positive_integral_def:
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   872
  "integral\<^isup>P M f = (SUP g : {g. simple_function M g \<and> g \<le> max 0 \<circ> f}. integral\<^isup>S M g)"
35692
f1315bbf1bc9 Moved theorems in Lebesgue to the right places
hoelzl
parents: 35582
diff changeset
   873
41689
3e39b0e730d6 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
syntax
45342
5c760e1692b3 proper syntactic category for abstraction syntax, to avoid low-level exception for malformed "\<integral> x y. f \<partial>M", for example;
wenzelm
parents: 44937
diff changeset
   875
  "_positive_integral" :: "pttrn \<Rightarrow> ereal \<Rightarrow> ('a, 'b) measure_space_scheme \<Rightarrow> ereal" ("\<integral>\<^isup>+ _. _ \<partial>_" [60,61] 110)
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   876
3e39b0e730d6 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
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
   878
  "\<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
   879
40873
1ef85f4e7097 Shorter definition for positive_integral.
hoelzl
parents: 40872
diff changeset
   880
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
   881
  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
   882
  shows "integral\<^isup>P N f = integral\<^isup>P M f"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   883
  unfolding positive_integral_def
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   884
  unfolding simple_function_cong_algebra[OF assms(2,3), symmetric]
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   885
  using AE_cong_measure[OF assms]
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   886
  using simple_integral_cong_measure[OF assms]
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   887
  by (auto intro!: SUP_cong)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   888
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   889
lemma (in measure_space) positive_integral_positive:
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   890
  "0 \<le> integral\<^isup>P M f"
44928
7ef6505bde7f renamed Complete_Lattices lemmas, removed legacy names
hoelzl
parents: 44890
diff changeset
   891
  by (auto intro!: SUP_upper2[of "\<lambda>x. 0"] simp: positive_integral_def le_fun_def)
40873
1ef85f4e7097 Shorter definition for positive_integral.
hoelzl
parents: 40872
diff changeset
   892
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   893
lemma (in measure_space) positive_integral_def_finite:
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   894
  "integral\<^isup>P M f = (SUP g : {g. simple_function M g \<and> g \<le> max 0 \<circ> f \<and> range g \<subseteq> {0 ..< \<infinity>}}. integral\<^isup>S M g)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   895
    (is "_ = SUPR ?A ?f")
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   896
  unfolding positive_integral_def
44928
7ef6505bde7f renamed Complete_Lattices lemmas, removed legacy names
hoelzl
parents: 44890
diff changeset
   897
proof (safe intro!: antisym SUP_least)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   898
  fix g assume g: "simple_function M g" "g \<le> max 0 \<circ> f"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   899
  let ?G = "{x \<in> space M. \<not> g x \<noteq> \<infinity>}"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   900
  note gM = g(1)[THEN borel_measurable_simple_function]
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   901
  have \<mu>G_pos: "0 \<le> \<mu> ?G" using gM by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   902
  let "?g y x" = "if g x = \<infinity> then y else max 0 (g x)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   903
  from g gM have g_in_A: "\<And>y. 0 \<le> y \<Longrightarrow> y \<noteq> \<infinity> \<Longrightarrow> ?g y \<in> ?A"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   904
    apply (safe intro!: simple_function_max simple_function_If)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   905
    apply (force simp: max_def le_fun_def split: split_if_asm)+
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   906
    done
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   907
  show "integral\<^isup>S M g \<le> SUPR ?A ?f"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   908
  proof cases
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   909
    have g0: "?g 0 \<in> ?A" by (intro g_in_A) auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   910
    assume "\<mu> ?G = 0"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   911
    with gM have "AE x. x \<notin> ?G" by (simp add: AE_iff_null_set)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   912
    with gM g show ?thesis
44928
7ef6505bde7f renamed Complete_Lattices lemmas, removed legacy names
hoelzl
parents: 44890
diff changeset
   913
      by (intro SUP_upper2[OF g0] simple_integral_mono_AE)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   914
         (auto simp: max_def intro!: simple_function_If)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   915
  next
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   916
    assume \<mu>G: "\<mu> ?G \<noteq> 0"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   917
    have "SUPR ?A (integral\<^isup>S M) = \<infinity>"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   918
    proof (intro SUP_PInfty)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   919
      fix n :: nat
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
   920
      let ?y = "ereal (real n) / (if \<mu> ?G = \<infinity> then 1 else \<mu> ?G)"
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
   921
      have "0 \<le> ?y" "?y \<noteq> \<infinity>" using \<mu>G \<mu>G_pos by (auto simp: ereal_divide_eq)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   922
      then have "?g ?y \<in> ?A" by (rule g_in_A)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   923
      have "real n \<le> ?y * \<mu> ?G"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   924
        using \<mu>G \<mu>G_pos by (cases "\<mu> ?G") (auto simp: field_simps)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   925
      also have "\<dots> = (\<integral>\<^isup>Sx. ?y * indicator ?G x \<partial>M)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   926
        using `0 \<le> ?y` `?g ?y \<in> ?A` gM
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   927
        by (subst simple_integral_cmult_indicator) auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   928
      also have "\<dots> \<le> integral\<^isup>S M (?g ?y)" using `?g ?y \<in> ?A` gM
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   929
        by (intro simple_integral_mono) auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   930
      finally show "\<exists>i\<in>?A. real n \<le> integral\<^isup>S M i"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   931
        using `?g ?y \<in> ?A` by blast
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   932
    qed
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   933
    then show ?thesis by simp
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   934
  qed
44928
7ef6505bde7f renamed Complete_Lattices lemmas, removed legacy names
hoelzl
parents: 44890
diff changeset
   935
qed (auto intro: SUP_upper)
40873
1ef85f4e7097 Shorter definition for positive_integral.
hoelzl
parents: 40872
diff changeset
   936
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   937
lemma (in measure_space) positive_integral_mono_AE:
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   938
  assumes ae: "AE x. u x \<le> v x" shows "integral\<^isup>P M u \<le> integral\<^isup>P M v"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   939
  unfolding positive_integral_def
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   940
proof (safe intro!: SUP_mono)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   941
  fix n assume n: "simple_function M n" "n \<le> max 0 \<circ> u"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   942
  from ae[THEN AE_E] guess N . note N = this
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   943
  then have ae_N: "AE x. x \<notin> N" by (auto intro: AE_not_in)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   944
  let "?n x" = "n x * indicator (space M - N) x"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   945
  have "AE x. n x \<le> ?n x" "simple_function M ?n"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   946
    using n N ae_N by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   947
  moreover
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   948
  { fix x have "?n x \<le> max 0 (v x)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   949
    proof cases
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   950
      assume x: "x \<in> space M - N"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   951
      with N have "u x \<le> v x" by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   952
      with n(2)[THEN le_funD, of x] x show ?thesis
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   953
        by (auto simp: max_def split: split_if_asm)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   954
    qed simp }
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   955
  then have "?n \<le> max 0 \<circ> v" by (auto simp: le_funI)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   956
  moreover have "integral\<^isup>S M n \<le> integral\<^isup>S M ?n"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   957
    using ae_N N n by (auto intro!: simple_integral_mono_AE)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   958
  ultimately show "\<exists>m\<in>{g. simple_function M g \<and> g \<le> max 0 \<circ> v}. integral\<^isup>S M n \<le> integral\<^isup>S M m"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   959
    by force
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   960
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   961
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   962
lemma (in measure_space) positive_integral_mono:
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   963
  "(\<And>x. x \<in> space M \<Longrightarrow> u x \<le> v x) \<Longrightarrow> integral\<^isup>P M u \<le> integral\<^isup>P M v"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   964
  by (auto intro: positive_integral_mono_AE)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   965
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   966
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
   967
  "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
   968
  by (auto simp: eq_iff intro!: positive_integral_mono_AE)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   969
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   970
lemma (in measure_space) positive_integral_cong:
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   971
  "(\<And>x. x \<in> space M \<Longrightarrow> u x = v x) \<Longrightarrow> integral\<^isup>P M u = integral\<^isup>P M v"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   972
  by (auto intro: positive_integral_cong_AE)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
   973
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   974
lemma (in measure_space) positive_integral_eq_simple_integral:
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   975
  assumes f: "simple_function M f" "\<And>x. 0 \<le> f x" shows "integral\<^isup>P M f = integral\<^isup>S M f"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   976
proof -
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   977
  let "?f x" = "f x * indicator (space M) x"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   978
  have f': "simple_function M ?f" using f by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   979
  with f(2) have [simp]: "max 0 \<circ> ?f = ?f"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   980
    by (auto simp: fun_eq_iff max_def split: split_indicator)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   981
  have "integral\<^isup>P M ?f \<le> integral\<^isup>S M ?f" using f'
44928
7ef6505bde7f renamed Complete_Lattices lemmas, removed legacy names
hoelzl
parents: 44890
diff changeset
   982
    by (force intro!: SUP_least simple_integral_mono simp: le_fun_def positive_integral_def)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   983
  moreover have "integral\<^isup>S M ?f \<le> integral\<^isup>P M ?f"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   984
    unfolding positive_integral_def
44928
7ef6505bde7f renamed Complete_Lattices lemmas, removed legacy names
hoelzl
parents: 44890
diff changeset
   985
    using f' by (auto intro!: SUP_upper)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   986
  ultimately show ?thesis
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   987
    by (simp cong: positive_integral_cong simple_integral_cong)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   988
qed
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   989
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   990
lemma (in measure_space) positive_integral_eq_simple_integral_AE:
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   991
  assumes f: "simple_function M f" "AE x. 0 \<le> f x" shows "integral\<^isup>P M f = integral\<^isup>S M f"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   992
proof -
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   993
  have "AE x. f x = max 0 (f x)" using f by (auto split: split_max)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   994
  with f have "integral\<^isup>P M f = integral\<^isup>S M (\<lambda>x. max 0 (f x))"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   995
    by (simp cong: positive_integral_cong_AE simple_integral_cong_AE
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   996
             add: positive_integral_eq_simple_integral)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   997
  with assms show ?thesis
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   998
    by (auto intro!: simple_integral_cong_AE split: split_max)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   999
qed
40873
1ef85f4e7097 Shorter definition for positive_integral.
hoelzl
parents: 40872
diff changeset
  1000
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1001
lemma (in measure_space) positive_integral_SUP_approx:
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1002
  assumes f: "incseq f" "\<And>i. f i \<in> borel_measurable M" "\<And>i x. 0 \<le> f i x"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1003
  and u: "simple_function M u" "u \<le> (SUP i. f i)" "u`space M \<subseteq> {0..<\<infinity>}"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1004
  shows "integral\<^isup>S M u \<le> (SUP i. integral\<^isup>P M (f i))" (is "_ \<le> ?S")
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1005
proof (rule ereal_le_mult_one_interval)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1006
  have "0 \<le> (SUP i. integral\<^isup>P M (f i))"
44928
7ef6505bde7f renamed Complete_Lattices lemmas, removed legacy names
hoelzl
parents: 44890
diff changeset
  1007
    using f(3) by (auto intro!: SUP_upper2 positive_integral_positive)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1008
  then show "(SUP i. integral\<^isup>P M (f i)) \<noteq> -\<infinity>" by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1009
  have u_range: "\<And>x. x \<in> space M \<Longrightarrow> 0 \<le> u x \<and> u x \<noteq> \<infinity>"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1010
    using u(3) by auto
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1011
  fix a :: ereal assume "0 < a" "a < 1"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1012
  hence "a \<noteq> 0" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1013
  let "?B i" = "{x \<in> space M. a * u x \<le> f i x}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1014
  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
  1015
    using f `simple_function M u` by (auto simp: borel_measurable_simple_function)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1016
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1017
  let "?uB i x" = "u x * indicator (?B i) x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1018
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1019
  { fix i have "?B i \<subseteq> ?B (Suc i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1020
    proof safe
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1021
      fix i x assume "a * u x \<le> f i x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1022
      also have "\<dots> \<le> f (Suc i) x"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1023
        using `incseq f`[THEN incseq_SucD] unfolding le_fun_def by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1024
      finally show "a * u x \<le> f (Suc i) x" .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1025
    qed }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1026
  note B_mono = this
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1027
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1028
  note B_u = Int[OF u(1)[THEN simple_functionD(2)] B]
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1029
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1030
  let "?B' i n" = "(u -` {i} \<inter> space M) \<inter> ?B n"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1031
  have measure_conv: "\<And>i. \<mu> (u -` {i} \<inter> space M) = (SUP n. \<mu> (?B' i n))"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1032
  proof -
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1033
    fix i
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1034
    have 1: "range (?B' i) \<subseteq> sets M" using B_u by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1035
    have 2: "incseq (?B' i)" using B_mono by (auto intro!: incseq_SucI)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1036
    have "(\<Union>n. ?B' i n) = u -` {i} \<inter> space M"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1037
    proof safe
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1038
      fix x i assume x: "x \<in> space M"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1039
      show "x \<in> (\<Union>i. ?B' (u x) i)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1040
      proof cases
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1041
        assume "u x = 0" thus ?thesis using `x \<in> space M` f(3) by simp
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1042
      next
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1043
        assume "u x \<noteq> 0"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1044
        with `a < 1` u_range[OF `x \<in> space M`]
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1045
        have "a * u x < 1 * u x"
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1046
          by (intro ereal_mult_strict_right_mono) (auto simp: image_iff)
44928
7ef6505bde7f renamed Complete_Lattices lemmas, removed legacy names
hoelzl
parents: 44890
diff changeset
  1047
        also have "\<dots> \<le> (SUP i. f i x)" using u(2) by (auto simp: le_fun_def SUP_apply)
7ef6505bde7f renamed Complete_Lattices lemmas, removed legacy names
hoelzl
parents: 44890
diff changeset
  1048
        finally obtain i where "a * u x < f i x" unfolding SUP_def
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1049
          by (auto simp add: less_Sup_iff)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1050
        hence "a * u x \<le> f i x" by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1051
        thus ?thesis using `x \<in> space M` by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1052
      qed
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1053
    qed
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1054
    then show "?thesis i" using continuity_from_below[OF 1 2] by simp
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1055
  qed
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1056
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1057
  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
  1058
    unfolding simple_integral_indicator[OF B `simple_function M u`]
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1059
  proof (subst SUPR_ereal_setsum, safe)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1060
    fix x n assume "x \<in> space M"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1061
    with u_range show "incseq (\<lambda>i. u x * \<mu> (?B' (u x) i))" "\<And>i. 0 \<le> u x * \<mu> (?B' (u x) i)"
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1062
      using B_mono B_u by (auto intro!: measure_mono ereal_mult_left_mono incseq_SucI simp: ereal_zero_le_0_iff)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1063
  next
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1064
    show "integral\<^isup>S M u = (\<Sum>i\<in>u ` space M. SUP n. i * \<mu> (?B' i n))"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1065
      using measure_conv u_range B_u unfolding simple_integral_def
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1066
      by (auto intro!: setsum_cong SUPR_ereal_cmult[symmetric])
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1067
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1068
  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
  1069
  have "a * (SUP i. integral\<^isup>S M (?uB i)) \<le> ?S"
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1070
    apply (subst SUPR_ereal_cmult[symmetric])
38705
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 38656
diff changeset
  1071
  proof (safe intro!: SUP_mono bexI)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1072
    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
  1073
    have "a * integral\<^isup>S M (?uB i) = (\<integral>\<^isup>Sx. a * ?uB i x \<partial>M)"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1074
      using B `simple_function M u` u_range
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1075
      by (subst simple_integral_mult) (auto split: split_indicator)
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1076
    also have "\<dots> \<le> integral\<^isup>P M (f i)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1077
    proof -
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1078
      have *: "simple_function M (\<lambda>x. a * ?uB i x)" using B `0 < a` u(1) by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1079
      show ?thesis using f(3) * u_range `0 < a`
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1080
        by (subst positive_integral_eq_simple_integral[symmetric])
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1081
           (auto intro!: positive_integral_mono split: split_indicator)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1082
    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
  1083
    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
  1084
      by auto
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1085
  next
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1086
    fix i show "0 \<le> \<integral>\<^isup>S x. ?uB i x \<partial>M" using B `0 < a` u(1) u_range
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1087
      by (intro simple_integral_positive) (auto split: split_indicator)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1088
  qed (insert `0 < a`, auto)
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1089
  ultimately show "a * integral\<^isup>S M u \<le> ?S" by simp
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1090
qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1091
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1092
lemma (in measure_space) incseq_positive_integral:
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1093
  assumes "incseq f" shows "incseq (\<lambda>i. integral\<^isup>P M (f i))"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1094
proof -
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1095
  have "\<And>i x. f i x \<le> f (Suc i) x"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1096
    using assms by (auto dest!: incseq_SucD simp: le_fun_def)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1097
  then show ?thesis
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1098
    by (auto intro!: incseq_SucI positive_integral_mono)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1099
qed
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1100
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1101
text {* Beppo-Levi monotone convergence theorem *}
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1102
lemma (in measure_space) positive_integral_monotone_convergence_SUP:
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1103
  assumes f: "incseq f" "\<And>i. f i \<in> borel_measurable M" "\<And>i x. 0 \<le> f i x"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1104
  shows "(\<integral>\<^isup>+ x. (SUP i. f i x) \<partial>M) = (SUP i. integral\<^isup>P M (f i))"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1105
proof (rule antisym)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1106
  show "(SUP j. integral\<^isup>P M (f j)) \<le> (\<integral>\<^isup>+ x. (SUP i. f i x) \<partial>M)"
44928
7ef6505bde7f renamed Complete_Lattices lemmas, removed legacy names
hoelzl
parents: 44890
diff changeset
  1107
    by (auto intro!: SUP_least SUP_upper positive_integral_mono)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1108
next
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1109
  show "(\<integral>\<^isup>+ x. (SUP i. f i x) \<partial>M) \<le> (SUP j. integral\<^isup>P M (f j))"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1110
    unfolding positive_integral_def_finite[of "\<lambda>x. SUP i. f i x"]
44928
7ef6505bde7f renamed Complete_Lattices lemmas, removed legacy names
hoelzl
parents: 44890
diff changeset
  1111
  proof (safe intro!: SUP_least)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1112
    fix g assume g: "simple_function M g"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1113
      and "g \<le> max 0 \<circ> (\<lambda>x. SUP i. f i x)" "range g \<subseteq> {0..<\<infinity>}"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1114
    moreover then have "\<And>x. 0 \<le> (SUP i. f i x)" and g': "g`space M \<subseteq> {0..<\<infinity>}"
44928
7ef6505bde7f renamed Complete_Lattices lemmas, removed legacy names
hoelzl
parents: 44890
diff changeset
  1115
      using f by (auto intro!: SUP_upper2)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1116
    ultimately show "integral\<^isup>S M g \<le> (SUP j. integral\<^isup>P M (f j))"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1117
      by (intro  positive_integral_SUP_approx[OF f g _ g'])
44928
7ef6505bde7f renamed Complete_Lattices lemmas, removed legacy names
hoelzl
parents: 44890
diff changeset
  1118
         (auto simp: le_fun_def max_def SUP_apply)
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1119
  qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1120
qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1121
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1122
lemma (in measure_space) positive_integral_monotone_convergence_SUP_AE:
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1123
  assumes f: "\<And>i. AE x. f i x \<le> f (Suc i) x \<and> 0 \<le> f i x" "\<And>i. f i \<in> borel_measurable M"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1124
  shows "(\<integral>\<^isup>+ x. (SUP i. f i x) \<partial>M) = (SUP i. integral\<^isup>P M (f i))"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1125
proof -
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1126
  from f have "AE x. \<forall>i. f i x \<le> f (Suc i) x \<and> 0 \<le> f i x"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1127
    by (simp add: AE_all_countable)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1128
  from this[THEN AE_E] guess N . note N = this
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1129
  let "?f i x" = "if x \<in> space M - N then f i x else 0"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1130
  have f_eq: "AE x. \<forall>i. ?f i x = f i x" using N by (auto intro!: AE_I[of _ N])
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1131
  then have "(\<integral>\<^isup>+ x. (SUP i. f i x) \<partial>M) = (\<integral>\<^isup>+ x. (SUP i. ?f i x) \<partial>M)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1132
    by (auto intro!: positive_integral_cong_AE)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1133
  also have "\<dots> = (SUP i. (\<integral>\<^isup>+ x. ?f i x \<partial>M))"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1134
  proof (rule positive_integral_monotone_convergence_SUP)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1135
    show "incseq ?f" using N(1) by (force intro!: incseq_SucI le_funI)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1136
    { fix i show "(\<lambda>x. if x \<in> space M - N then f i x else 0) \<in> borel_measurable M"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1137
        using f N(3) by (intro measurable_If_set) auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1138
      fix x show "0 \<le> ?f i x"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1139
        using N(1) by auto }
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1140
  qed
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1141
  also have "\<dots> = (SUP i. (\<integral>\<^isup>+ x. f i x \<partial>M))"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1142
    using f_eq by (force intro!: arg_cong[where f="SUPR UNIV"] positive_integral_cong_AE ext)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1143
  finally show ?thesis .
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1144
qed
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1145
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1146
lemma (in measure_space) positive_integral_monotone_convergence_SUP_AE_incseq:
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1147
  assumes f: "incseq f" "\<And>i. AE x. 0 \<le> f i x" and borel: "\<And>i. f i \<in> borel_measurable M"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1148
  shows "(\<integral>\<^isup>+ x. (SUP i. f i x) \<partial>M) = (SUP i. integral\<^isup>P M (f i))"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1149
  using f[unfolded incseq_Suc_iff le_fun_def]
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1150
  by (intro positive_integral_monotone_convergence_SUP_AE[OF _ borel])
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1151
     auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1152
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1153
lemma (in measure_space) positive_integral_monotone_convergence_simple:
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1154
  assumes f: "incseq f" "\<And>i x. 0 \<le> f i x" "\<And>i. simple_function M (f i)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1155
  shows "(SUP i. integral\<^isup>S M (f i)) = (\<integral>\<^isup>+x. (SUP i. f i x) \<partial>M)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1156
  using assms unfolding positive_integral_monotone_convergence_SUP[OF f(1)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1157
    f(3)[THEN borel_measurable_simple_function] f(2)]
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1158
  by (auto intro!: positive_integral_eq_simple_integral[symmetric] arg_cong[where f="SUPR UNIV"] ext)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1159
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1160
lemma positive_integral_max_0:
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1161
  "(\<integral>\<^isup>+x. max 0 (f x) \<partial>M) = integral\<^isup>P M f"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1162
  by (simp add: le_fun_def positive_integral_def)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1163
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1164
lemma (in measure_space) positive_integral_cong_pos:
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1165
  assumes "\<And>x. x \<in> space M \<Longrightarrow> f x \<le> 0 \<and> g x \<le> 0 \<or> f x = g x"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1166
  shows "integral\<^isup>P M f = integral\<^isup>P M g"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1167
proof -
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1168
  have "integral\<^isup>P M (\<lambda>x. max 0 (f x)) = integral\<^isup>P M (\<lambda>x. max 0 (g x))"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1169
  proof (intro positive_integral_cong)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1170
    fix x assume "x \<in> space M"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1171
    from assms[OF this] show "max 0 (f x) = max 0 (g x)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1172
      by (auto split: split_max)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1173
  qed
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1174
  then show ?thesis by (simp add: positive_integral_max_0)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1175
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1176
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1177
lemma (in measure_space) SUP_simple_integral_sequences:
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1178
  assumes f: "incseq f" "\<And>i x. 0 \<le> f i x" "\<And>i. simple_function M (f i)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1179
  and g: "incseq g" "\<And>i x. 0 \<le> g i x" "\<And>i. simple_function M (g i)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1180
  and eq: "AE x. (SUP i. f i x) = (SUP i. g i 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
  1181
  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
  1182
    (is "SUPR _ ?F = SUPR _ ?G")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1183
proof -
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1184
  have "(SUP i. integral\<^isup>S M (f i)) = (\<integral>\<^isup>+x. (SUP i. f i x) \<partial>M)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1185
    using f by (rule positive_integral_monotone_convergence_simple)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1186
  also have "\<dots> = (\<integral>\<^isup>+x. (SUP i. g i x) \<partial>M)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1187
    unfolding eq[THEN positive_integral_cong_AE] ..
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1188
  also have "\<dots> = (SUP i. ?G i)"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1189
    using g by (rule positive_integral_monotone_convergence_simple[symmetric])
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1190
  finally show ?thesis by simp
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1191
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1192
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1193
lemma (in measure_space) positive_integral_const[simp]:
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1194
  "0 \<le> c \<Longrightarrow> (\<integral>\<^isup>+ x. c \<partial>M) = c * \<mu> (space M)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1195
  by (subst positive_integral_eq_simple_integral) auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1196
41661
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41545
diff changeset
  1197
lemma (in measure_space) positive_integral_vimage:
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1198
  assumes T: "sigma_algebra M'" "T \<in> measure_preserving M M'"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1199
  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
  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
  1201
proof -
41831
91a2b435dd7a use measure_preserving in ..._vimage lemmas
hoelzl
parents: 41705
diff changeset
  1202
  interpret T: measure_space M' by (rule measure_space_vimage[OF T])
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1203
  from T.borel_measurable_implies_simple_function_sequence'[OF f]
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1204
  guess f' . note f' = this
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1205
  let "?f i x" = "f' i (T x)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1206
  have inc: "incseq ?f" using f' by (force simp: le_fun_def incseq_def)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1207
  have sup: "\<And>x. (SUP i. ?f i x) = max 0 (f (T x))"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1208
    using f'(4) .
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1209
  have sf: "\<And>i. simple_function M (\<lambda>x. f' i (T x))"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1210
    using simple_function_vimage[OF T(1) measure_preservingD2[OF T(2)] f'(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
  1211
  show "integral\<^isup>P M' f = (\<integral>\<^isup>+ x. f (T x) \<partial>M)"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1212
    using
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1213
      T.positive_integral_monotone_convergence_simple[OF f'(2,5,1)]
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1214
      positive_integral_monotone_convergence_simple[OF inc f'(5) sf]
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1215
    by (simp add: positive_integral_max_0 simple_integral_vimage[OF T f'(1)] f')
41661
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41545
diff changeset
  1216
qed
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41545
diff changeset
  1217
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1218
lemma (in measure_space) positive_integral_linear:
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1219
  assumes f: "f \<in> borel_measurable M" "\<And>x. 0 \<le> f x" and "0 \<le> a"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1220
  and g: "g \<in> borel_measurable M" "\<And>x. 0 \<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
  1221
  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
  1222
    (is "integral\<^isup>P M ?L = _")
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1223
proof -
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1224
  from borel_measurable_implies_simple_function_sequence'[OF f(1)] guess u .
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1225
  note u = positive_integral_monotone_convergence_simple[OF this(2,5,1)] this
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1226
  from borel_measurable_implies_simple_function_sequence'[OF g(1)] guess v .
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1227
  note v = positive_integral_monotone_convergence_simple[OF this(2,5,1)] this
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1228
  let "?L' i x" = "a * u i x + v i x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1229
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1230
  have "?L \<in> borel_measurable M" using assms by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1231
  from borel_measurable_implies_simple_function_sequence'[OF this] guess l .
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1232
  note l = positive_integral_monotone_convergence_simple[OF this(2,5,1)] this
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1233
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1234
  have inc: "incseq (\<lambda>i. a * integral\<^isup>S M (u i))" "incseq (\<lambda>i. integral\<^isup>S M (v i))"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1235
    using u v `0 \<le> a`
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1236
    by (auto simp: incseq_Suc_iff le_fun_def
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1237
             intro!: add_mono ereal_mult_left_mono simple_integral_mono)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1238
  have pos: "\<And>i. 0 \<le> integral\<^isup>S M (u i)" "\<And>i. 0 \<le> integral\<^isup>S M (v i)" "\<And>i. 0 \<le> a * integral\<^isup>S M (u i)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1239
    using u v `0 \<le> a` by (auto simp: simple_integral_positive)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1240
  { fix i from pos[of i] have "a * integral\<^isup>S M (u i) \<noteq> -\<infinity>" "integral\<^isup>S M (v i) \<noteq> -\<infinity>"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1241
      by (auto split: split_if_asm) }
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1242
  note not_MInf = this
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1243
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1244
  have l': "(SUP i. integral\<^isup>S M (l i)) = (SUP i. integral\<^isup>S M (?L' i))"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1245
  proof (rule SUP_simple_integral_sequences[OF l(3,6,2)])
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1246
    show "incseq ?L'" "\<And>i x. 0 \<le> ?L' i x" "\<And>i. simple_function M (?L' i)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1247
      using u v  `0 \<le> a` unfolding incseq_Suc_iff le_fun_def
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1248
      by (auto intro!: add_mono ereal_mult_left_mono ereal_add_nonneg_nonneg)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1249
    { fix x
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1250
      { fix i have "a * u i x \<noteq> -\<infinity>" "v i x \<noteq> -\<infinity>" "u i x \<noteq> -\<infinity>" using `0 \<le> a` u(6)[of i x] v(6)[of i x]
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1251
          by auto }
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1252
      then have "(SUP i. a * u i x + v i x) = a * (SUP i. u i x) + (SUP i. v i x)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1253
        using `0 \<le> a` u(3) v(3) u(6)[of _ x] v(6)[of _ x]
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1254
        by (subst SUPR_ereal_cmult[symmetric, OF u(6) `0 \<le> a`])
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1255
           (auto intro!: SUPR_ereal_add
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1256
                 simp: incseq_Suc_iff le_fun_def add_mono ereal_mult_left_mono ereal_add_nonneg_nonneg) }
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1257
    then show "AE x. (SUP i. l i x) = (SUP i. ?L' i x)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1258
      unfolding l(5) using `0 \<le> a` u(5) v(5) l(5) f(2) g(2)
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1259
      by (intro AE_I2) (auto split: split_max simp add: ereal_add_nonneg_nonneg)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1260
  qed
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1261
  also have "\<dots> = (SUP i. a * integral\<^isup>S M (u i) + integral\<^isup>S M (v i))"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1262
    using u(2, 6) v(2, 6) `0 \<le> a` by (auto intro!: arg_cong[where f="SUPR UNIV"] ext)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1263
  finally have "(\<integral>\<^isup>+ x. max 0 (a * f x + g x) \<partial>M) = a * (\<integral>\<^isup>+x. max 0 (f x) \<partial>M) + (\<integral>\<^isup>+x. max 0 (g x) \<partial>M)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1264
    unfolding l(5)[symmetric] u(5)[symmetric] v(5)[symmetric]
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1265
    unfolding l(1)[symmetric] u(1)[symmetric] v(1)[symmetric]
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1266
    apply (subst SUPR_ereal_cmult[symmetric, OF pos(1) `0 \<le> a`])
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1267
    apply (subst SUPR_ereal_add[symmetric, OF inc not_MInf]) .
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1268
  then show ?thesis by (simp add: positive_integral_max_0)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1269
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1270
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1271
lemma (in measure_space) positive_integral_cmult:
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1272
  assumes f: "f \<in> borel_measurable M" "AE x. 0 \<le> f x" "0 \<le> c"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1273
  shows "(\<integral>\<^isup>+ x. c * f x \<partial>M) = c * integral\<^isup>P M f"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1274
proof -
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1275
  have [simp]: "\<And>x. c * max 0 (f x) = max 0 (c * f x)" using `0 \<le> c`
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1276
    by (auto split: split_max simp: ereal_zero_le_0_iff)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1277
  have "(\<integral>\<^isup>+ x. c * f x \<partial>M) = (\<integral>\<^isup>+ x. c * max 0 (f x) \<partial>M)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1278
    by (simp add: positive_integral_max_0)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1279
  then show ?thesis
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1280
    using positive_integral_linear[OF _ _ `0 \<le> c`, of "\<lambda>x. max 0 (f x)" "\<lambda>x. 0"] f
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1281
    by (auto simp: positive_integral_max_0)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1282
qed
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1283
41096
843c40bbc379 integral over setprod
hoelzl
parents: 41095
diff changeset
  1284
lemma (in measure_space) positive_integral_multc:
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1285
  assumes "f \<in> borel_measurable M" "AE x. 0 \<le> f x" "0 \<le> c"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1286
  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
  1287
  unfolding mult_commute[of _ c] positive_integral_cmult[OF assms] by simp
843c40bbc379 integral over setprod
hoelzl
parents: 41095
diff changeset
  1288
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1289
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
  1290
  "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
  1291
  by (subst positive_integral_eq_simple_integral)
c3b977fee8a3 introduced integral syntax
hoelzl
parents: 41097
diff changeset
  1292
     (auto simp: simple_function_indicator simple_integral_indicator)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1293
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1294
lemma (in measure_space) positive_integral_cmult_indicator:
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1295
  "0 \<le> c \<Longrightarrow> 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
  1296
  by (subst positive_integral_eq_simple_integral)
c3b977fee8a3 introduced integral syntax
hoelzl
parents: 41097
diff changeset
  1297
     (auto simp: simple_function_indicator simple_integral_indicator)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1298
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1299
lemma (in measure_space) positive_integral_add:
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1300
  assumes f: "f \<in> borel_measurable M" "AE x. 0 \<le> f x"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1301
  and g: "g \<in> borel_measurable M" "AE x. 0 \<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
  1302
  shows "(\<integral>\<^isup>+ x. f x + g x \<partial>M) = integral\<^isup>P M f + integral\<^isup>P M g"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1303
proof -
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1304
  have ae: "AE x. max 0 (f x) + max 0 (g x) = max 0 (f x + g x)"
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1305
    using assms by (auto split: split_max simp: ereal_add_nonneg_nonneg)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1306
  have "(\<integral>\<^isup>+ x. f x + g x \<partial>M) = (\<integral>\<^isup>+ x. max 0 (f x + g x) \<partial>M)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1307
    by (simp add: positive_integral_max_0)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1308
  also have "\<dots> = (\<integral>\<^isup>+ x. max 0 (f x) + max 0 (g x) \<partial>M)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1309
    unfolding ae[THEN positive_integral_cong_AE] ..
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1310
  also have "\<dots> = (\<integral>\<^isup>+ x. max 0 (f x) \<partial>M) + (\<integral>\<^isup>+ x. max 0 (g x) \<partial>M)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1311
    using positive_integral_linear[of "\<lambda>x. max 0 (f x)" 1 "\<lambda>x. max 0 (g x)"] f g
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1312
    by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1313
  finally show ?thesis
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1314
    by (simp add: positive_integral_max_0)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1315
qed
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1316
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1317
lemma (in measure_space) positive_integral_setsum:
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1318
  assumes "\<And>i. i\<in>P \<Longrightarrow> f i \<in> borel_measurable M" "\<And>i. i\<in>P \<Longrightarrow> AE x. 0 \<le> f i 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
  1319
  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
  1320
proof cases
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1321
  assume f: "finite P"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1322
  from assms have "AE x. \<forall>i\<in>P. 0 \<le> f i x" unfolding AE_finite_all[OF f] by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1323
  from f this assms(1) show ?thesis
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1324
  proof induct
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1325
    case (insert i P)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1326
    then have "f i \<in> borel_measurable M" "AE x. 0 \<le> f i x"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1327
      "(\<lambda>x. \<Sum>i\<in>P. f i x) \<in> borel_measurable M" "AE x. 0 \<le> (\<Sum>i\<in>P. f i x)"
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1328
      by (auto intro!: borel_measurable_ereal_setsum setsum_nonneg)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1329
    from positive_integral_add[OF this]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1330
    show ?case using insert by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1331
  qed simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1332
qed simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1333
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1334
lemma (in measure_space) positive_integral_Markov_inequality:
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1335
  assumes u: "u \<in> borel_measurable M" "AE x. 0 \<le> u x" and "A \<in> sets M" and c: "0 \<le> c" "c \<noteq> \<infinity>"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1336
  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)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1337
    (is "\<mu> ?A \<le> _ * ?PI")
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1338
proof -
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1339
  have "?A \<in> sets M"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1340
    using `A \<in> sets M` u by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1341
  hence "\<mu> ?A = (\<integral>\<^isup>+ x. indicator ?A x \<partial>M)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1342
    using positive_integral_indicator by simp
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1343
  also have "\<dots> \<le> (\<integral>\<^isup>+ x. c * (u x * indicator A x) \<partial>M)" using u c
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1344
    by (auto intro!: positive_integral_mono_AE
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1345
      simp: indicator_def ereal_zero_le_0_iff)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1346
  also have "\<dots> = c * (\<integral>\<^isup>+ x. u x * indicator A x \<partial>M)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1347
    using assms
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1348
    by (auto intro!: positive_integral_cmult borel_measurable_indicator simp: ereal_zero_le_0_iff)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1349
  finally show ?thesis .
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1350
qed
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1351
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1352
lemma (in measure_space) positive_integral_noteq_infinite:
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1353
  assumes g: "g \<in> borel_measurable M" "AE x. 0 \<le> g x"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1354
  and "integral\<^isup>P M g \<noteq> \<infinity>"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1355
  shows "AE x. g x \<noteq> \<infinity>"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1356
proof (rule ccontr)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1357
  assume c: "\<not> (AE x. g x \<noteq> \<infinity>)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1358
  have "\<mu> {x\<in>space M. g x = \<infinity>} \<noteq> 0"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1359
    using c g by (simp add: AE_iff_null_set)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1360
  moreover have "0 \<le> \<mu> {x\<in>space M. g x = \<infinity>}" using g by (auto intro: measurable_sets)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1361
  ultimately have "0 < \<mu> {x\<in>space M. g x = \<infinity>}" by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1362
  then have "\<infinity> = \<infinity> * \<mu> {x\<in>space M. g x = \<infinity>}" by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1363
  also have "\<dots> \<le> (\<integral>\<^isup>+x. \<infinity> * indicator {x\<in>space M. g x = \<infinity>} x \<partial>M)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1364
    using g by (subst positive_integral_cmult_indicator) auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1365
  also have "\<dots> \<le> integral\<^isup>P M g"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1366
    using assms by (auto intro!: positive_integral_mono_AE simp: indicator_def)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1367
  finally show False using `integral\<^isup>P M g \<noteq> \<infinity>` by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1368
qed
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1369
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1370
lemma (in measure_space) positive_integral_diff:
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1371
  assumes f: "f \<in> borel_measurable M"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1372
  and g: "g \<in> borel_measurable M" "AE x. 0 \<le> g x"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1373
  and fin: "integral\<^isup>P M g \<noteq> \<infinity>"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1374
  and mono: "AE x. 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
  1375
  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
  1376
proof -
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1377
  have diff: "(\<lambda>x. f x - g x) \<in> borel_measurable M" "AE x. 0 \<le> f x - g x"
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1378
    using assms by (auto intro: ereal_diff_positive)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1379
  have pos_f: "AE x. 0 \<le> f x" using mono g by auto
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1380
  { fix a b :: ereal assume "0 \<le> a" "a \<noteq> \<infinity>" "0 \<le> b" "a \<le> b" then have "b - a + a = b"
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1381
      by (cases rule: ereal2_cases[of a b]) auto }
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1382
  note * = this
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1383
  then have "AE x. f x = f x - g x + g x"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1384
    using mono positive_integral_noteq_infinite[OF g fin] assms by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1385
  then have **: "integral\<^isup>P M f = (\<integral>\<^isup>+x. f x - g x \<partial>M) + integral\<^isup>P M g"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1386
    unfolding positive_integral_add[OF diff g, symmetric]
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1387
    by (rule positive_integral_cong_AE)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1388
  show ?thesis unfolding **
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1389
    using fin positive_integral_positive[of g]
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1390
    by (cases rule: ereal2_cases[of "\<integral>\<^isup>+ x. f x - g x \<partial>M" "integral\<^isup>P M g"]) auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1391
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1392
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1393
lemma (in measure_space) positive_integral_suminf:
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1394
  assumes f: "\<And>i. f i \<in> borel_measurable M" "\<And>i. AE x. 0 \<le> f i x"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1395
  shows "(\<integral>\<^isup>+ x. (\<Sum>i. f i x) \<partial>M) = (\<Sum>i. integral\<^isup>P M (f i))"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1396
proof -
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1397
  have all_pos: "AE x. \<forall>i. 0 \<le> f i x"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1398
    using assms by (auto simp: AE_all_countable)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1399
  have "(\<Sum>i. integral\<^isup>P M (f i)) = (SUP n. \<Sum>i<n. integral\<^isup>P M (f i))"
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1400
    using positive_integral_positive by (rule suminf_ereal_eq_SUPR)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1401
  also have "\<dots> = (SUP n. \<integral>\<^isup>+x. (\<Sum>i<n. f i x) \<partial>M)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1402
    unfolding positive_integral_setsum[OF f] ..
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1403
  also have "\<dots> = \<integral>\<^isup>+x. (SUP n. \<Sum>i<n. f i x) \<partial>M" using f all_pos
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1404
    by (intro positive_integral_monotone_convergence_SUP_AE[symmetric])
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1405
       (elim AE_mp, auto simp: setsum_nonneg simp del: setsum_lessThan_Suc intro!: AE_I2 setsum_mono3)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1406
  also have "\<dots> = \<integral>\<^isup>+x. (\<Sum>i. f i x) \<partial>M" using all_pos
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1407
    by (intro positive_integral_cong_AE) (auto simp: suminf_ereal_eq_SUPR)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1408
  finally show ?thesis by simp
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1409
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1410
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1411
text {* Fatou's lemma: convergence theorem on limes inferior *}
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1412
lemma (in measure_space) positive_integral_lim_INF:
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1413
  fixes u :: "nat \<Rightarrow> 'a \<Rightarrow> ereal"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1414
  assumes u: "\<And>i. u i \<in> borel_measurable M" "\<And>i. AE x. 0 \<le> u i x"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1415
  shows "(\<integral>\<^isup>+ x. liminf (\<lambda>n. u n x) \<partial>M) \<le> liminf (\<lambda>n. integral\<^isup>P M (u n))"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1416
proof -
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1417
  have pos: "AE x. \<forall>i. 0 \<le> u i x" using u by (auto simp: AE_all_countable)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1418
  have "(\<integral>\<^isup>+ x. liminf (\<lambda>n. u n x) \<partial>M) =
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1419
    (SUP n. \<integral>\<^isup>+ x. (INF i:{n..}. u i x) \<partial>M)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1420
    unfolding liminf_SUPR_INFI using pos u
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1421
    by (intro positive_integral_monotone_convergence_SUP_AE)
44937
22c0857b8aab removed further legacy rules from Complete_Lattices
hoelzl
parents: 44928
diff changeset
  1422
       (elim AE_mp, auto intro!: AE_I2 intro: INF_greatest INF_superset_mono)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1423
  also have "\<dots> \<le> liminf (\<lambda>n. integral\<^isup>P M (u n))"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1424
    unfolding liminf_SUPR_INFI
44928
7ef6505bde7f renamed Complete_Lattices lemmas, removed legacy names
hoelzl
parents: 44890
diff changeset
  1425
    by (auto intro!: SUP_mono exI INF_greatest positive_integral_mono INF_lower)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1426
  finally show ?thesis .
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1427
qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1428
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1429
lemma (in measure_space) measure_space_density:
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1430
  assumes u: "u \<in> borel_measurable M" "AE x. 0 \<le> 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
  1431
    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
  1432
  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
  1433
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
  1434
  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
  1435
  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
  1436
  proof
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1437
    have pos: "\<And>A. AE x. 0 \<le> u x * indicator A x"
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1438
      using u by (auto simp: ereal_zero_le_0_iff)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1439
    then show "positive M' (measure M')" unfolding M'
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1440
      using u(1) by (auto simp: positive_def intro!: positive_integral_positive)
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1441
    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
  1442
    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
  1443
      fix A :: "nat \<Rightarrow> 'a set" assume "range A \<subseteq> sets M'"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1444
      then have *: "\<And>i. (\<lambda>x. u x * indicator (A i) x) \<in> borel_measurable M"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1445
        using u by (auto intro: borel_measurable_indicator)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1446
      assume disj: "disjoint_family A"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1447
      have "(\<Sum>n. measure M' (A n)) = (\<integral>\<^isup>+ x. (\<Sum>n. u x * indicator (A n) x) \<partial>M)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1448
        unfolding M' using u(1) *
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1449
        by (simp add: positive_integral_suminf[OF _ pos, symmetric])
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1450
      also have "\<dots> = (\<integral>\<^isup>+ x. u x * (\<Sum>n. indicator (A n) x) \<partial>M)" using u
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1451
        by (intro positive_integral_cong_AE)
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1452
           (elim AE_mp, auto intro!: AE_I2 suminf_cmult_ereal)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1453
      also have "\<dots> = (\<integral>\<^isup>+ x. u x * indicator (\<Union>n. A n) x \<partial>M)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1454
        unfolding suminf_indicator[OF disj] ..
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1455
      finally show "(\<Sum>n. measure M' (A n)) = measure M' (\<Union>x. A x)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1456
        unfolding M' 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
  1457
    qed
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1458
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1459
qed
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1460
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1461
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
  1462
  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
  1463
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
  1464
  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
  1465
  proof (intro positive_integral_cong_AE AE_I)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1466
    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
  1467
      by (auto simp: indicator_def)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1468
    show "\<mu> N = 0" "N \<in> sets M"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1469
      using assms by auto
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1470
  qed
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1471
  then show ?thesis by simp
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1472
qed
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1473
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1474
lemma (in measure_space) positive_integral_translated_density:
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1475
  assumes f: "f \<in> borel_measurable M" "AE x. 0 \<le> f x"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1476
  assumes g: "g \<in> borel_measurable M" "AE x. 0 \<le> g x"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1477
    and M': "M' = (M\<lparr> measure := \<lambda>A. (\<integral>\<^isup>+ x. f x * indicator A x \<partial>M)\<rparr>)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1478
  shows "integral\<^isup>P M' g = (\<integral>\<^isup>+ x. f x * g x \<partial>M)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1479
proof -
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1480
  from measure_space_density[OF f M']
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1481
  interpret T: measure_space M' .
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1482
  have borel[simp]:
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1483
    "borel_measurable M' = borel_measurable M"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1484
    "simple_function M' = simple_function M"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1485
    unfolding measurable_def simple_function_def_raw by (auto simp: M')
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1486
  from borel_measurable_implies_simple_function_sequence'[OF g(1)] guess G . note G = this
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1487
  note G' = borel_measurable_simple_function[OF this(1)] simple_functionD[OF G(1)]
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1488
  note G'(2)[simp]
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1489
  { fix P have "AE x. P x \<Longrightarrow> AE x in M'. P x"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1490
      using positive_integral_null_set[of _ f]
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1491
      unfolding T.almost_everywhere_def almost_everywhere_def
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1492
      by (auto simp: M') }
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1493
  note ac = this
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1494
  from G(4) g(2) have G_M': "AE x in M'. (SUP i. G i x) = g x"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1495
    by (auto intro!: ac split: split_max)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1496
  { fix i
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1497
    let "?I y x" = "indicator (G i -` {y} \<inter> space M) x"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1498
    { fix x assume *: "x \<in> space M" "0 \<le> f x" "0 \<le> g x"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1499
      then have [simp]: "G i ` space M \<inter> {y. G i x = y \<and> x \<in> space M} = {G i x}" by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1500
      from * G' G have "(\<Sum>y\<in>G i`space M. y * (f x * ?I y x)) = f x * (\<Sum>y\<in>G i`space M. (y * ?I y x))"
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1501
        by (subst setsum_ereal_right_distrib) (auto simp: ac_simps)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1502
      also have "\<dots> = f x * G i x"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1503
        by (simp add: indicator_def if_distrib setsum_cases)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1504
      finally have "(\<Sum>y\<in>G i`space M. y * (f x * ?I y x)) = f x * G i x" . }
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1505
    note to_singleton = this
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1506
    have "integral\<^isup>P M' (G i) = integral\<^isup>S M' (G i)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1507
      using G T.positive_integral_eq_simple_integral by simp
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1508
    also have "\<dots> = (\<Sum>y\<in>G i`space M. y * (\<integral>\<^isup>+x. f x * ?I y x \<partial>M))"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1509
      unfolding simple_integral_def M' by simp
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1510
    also have "\<dots> = (\<Sum>y\<in>G i`space M. (\<integral>\<^isup>+x. y * (f x * ?I y x) \<partial>M))"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1511
      using f G' G by (auto intro!: setsum_cong positive_integral_cmult[symmetric])
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1512
    also have "\<dots> = (\<integral>\<^isup>+x. (\<Sum>y\<in>G i`space M. y * (f x * ?I y x)) \<partial>M)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1513
      using f G' G by (auto intro!: positive_integral_setsum[symmetric])
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1514
    finally have "integral\<^isup>P M' (G i) = (\<integral>\<^isup>+x. f x * G i x \<partial>M)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1515
      using f g G' to_singleton by (auto intro!: positive_integral_cong_AE) }
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1516
  note [simp] = this
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1517
  have "integral\<^isup>P M' g = (SUP i. integral\<^isup>P M' (G i))" using G'(1) G_M'(1) G
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1518
    using T.positive_integral_monotone_convergence_SUP[symmetric, OF `incseq G`]
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1519
    by (simp cong: T.positive_integral_cong_AE)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1520
  also have "\<dots> = (SUP i. (\<integral>\<^isup>+x. f x * G i x \<partial>M))" by simp
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1521
  also have "\<dots> = (\<integral>\<^isup>+x. (SUP i. f x * G i x) \<partial>M)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1522
    using f G' G(2)[THEN incseq_SucD] G
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1523
    by (intro positive_integral_monotone_convergence_SUP_AE[symmetric])
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1524
       (auto simp: ereal_mult_left_mono le_fun_def ereal_zero_le_0_iff)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1525
  also have "\<dots> = (\<integral>\<^isup>+x. f x * g x \<partial>M)" using f G' G g
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1526
    by (intro positive_integral_cong_AE)
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1527
       (auto simp add: SUPR_ereal_cmult split: split_max)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1528
  finally show "integral\<^isup>P M' g = (\<integral>\<^isup>+x. f x * g x \<partial>M)" .
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1529
qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1530
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1531
lemma (in measure_space) positive_integral_0_iff:
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1532
  assumes u: "u \<in> borel_measurable M" and pos: "AE x. 0 \<le> 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
  1533
  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
  1534
    (is "_ \<longleftrightarrow> \<mu> ?A = 0")
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1535
proof -
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1536
  have u_eq: "(\<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
  1537
    by (auto intro!: positive_integral_cong simp: indicator_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1538
  show ?thesis
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1539
  proof
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1540
    assume "\<mu> ?A = 0"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1541
    with positive_integral_null_set[of ?A u] u
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1542
    show "integral\<^isup>P M u = 0" by (simp add: u_eq)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1543
  next
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1544
    { fix r :: ereal and n :: nat assume gt_1: "1 \<le> real n * r"
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1545
      then have "0 < real n * r" by (cases r) (auto split: split_if_asm simp: one_ereal_def)
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1546
      then have "0 \<le> r" by (auto simp add: ereal_zero_less_0_iff) }
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1547
    note gt_1 = 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
  1548
    assume *: "integral\<^isup>P M u = 0"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1549
    let "?M n" = "{x \<in> space M. 1 \<le> real (n::nat) * u x}"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1550
    have "0 = (SUP n. \<mu> (?M n \<inter> ?A))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1551
    proof -
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1552
      { fix n :: nat
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1553
        from positive_integral_Markov_inequality[OF u pos, of ?A "ereal (real n)"]
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1554
        have "\<mu> (?M n \<inter> ?A) \<le> 0" unfolding u_eq * using u by simp
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1555
        moreover have "0 \<le> \<mu> (?M n \<inter> ?A)" using u by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1556
        ultimately have "\<mu> (?M n \<inter> ?A) = 0" by auto }
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1557
      thus ?thesis by simp
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1558
    qed
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1559
    also have "\<dots> = \<mu> (\<Union>n. ?M n \<inter> ?A)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1560
    proof (safe intro!: continuity_from_below)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1561
      fix n show "?M n \<inter> ?A \<in> sets M"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1562
        using u by (auto intro!: Int)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1563
    next
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1564
      show "incseq (\<lambda>n. {x \<in> space M. 1 \<le> real n * u x} \<inter> {x \<in> space M. u x \<noteq> 0})"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1565
      proof (safe intro!: incseq_SucI)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1566
        fix n :: nat and x
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1567
        assume *: "1 \<le> real n * u x"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1568
        also from gt_1[OF this] have "real n * u x \<le> real (Suc n) * u x"
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1569
          using `0 \<le> u x` by (auto intro!: ereal_mult_right_mono)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1570
        finally show "1 \<le> real (Suc n) * u x" by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1571
      qed
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1572
    qed
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1573
    also have "\<dots> = \<mu> {x\<in>space M. 0 < u x}"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1574
    proof (safe intro!: arg_cong[where f="\<mu>"] dest!: gt_1)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1575
      fix x assume "0 < u x" and [simp, intro]: "x \<in> space M"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1576
      show "x \<in> (\<Union>n. ?M n \<inter> ?A)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1577
      proof (cases "u x")
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1578
        case (real r) with `0 < u x` have "0 < r" by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1579
        obtain j :: nat where "1 / r \<le> real j" using real_arch_simple ..
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1580
        hence "1 / r * r \<le> real j * r" unfolding mult_le_cancel_right using `0 < r` by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1581
        hence "1 \<le> real j * r" using real `0 < r` by auto
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1582
        thus ?thesis using `0 < r` real by (auto simp: one_ereal_def)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1583
      qed (insert `0 < u x`, auto)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1584
    qed auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1585
    finally have "\<mu> {x\<in>space M. 0 < u x} = 0" by simp
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1586
    moreover
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1587
    from pos have "AE x. \<not> (u x < 0)" by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1588
    then have "\<mu> {x\<in>space M. u x < 0} = 0"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1589
      using AE_iff_null_set u by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1590
    moreover have "\<mu> {x\<in>space M. u x \<noteq> 0} = \<mu> {x\<in>space M. u x < 0} + \<mu> {x\<in>space M. 0 < u x}"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1591
      using u by (subst measure_additive) (auto intro!: arg_cong[where f=\<mu>])
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1592
    ultimately show "\<mu> ?A = 0" by simp
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1593
  qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1594
qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1595
41705
1100512e16d8 add auto support for AE_mp
hoelzl
parents: 41689
diff changeset
  1596
lemma (in measure_space) positive_integral_0_iff_AE:
1100512e16d8 add auto support for AE_mp
hoelzl
parents: 41689
diff changeset
  1597
  assumes u: "u \<in> borel_measurable M"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1598
  shows "integral\<^isup>P M u = 0 \<longleftrightarrow> (AE x. u x \<le> 0)"
41705
1100512e16d8 add auto support for AE_mp
hoelzl
parents: 41689
diff changeset
  1599
proof -
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1600
  have sets: "{x\<in>space M. max 0 (u x) \<noteq> 0} \<in> sets M"
41705
1100512e16d8 add auto support for AE_mp
hoelzl
parents: 41689
diff changeset
  1601
    using u by auto
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1602
  from positive_integral_0_iff[of "\<lambda>x. max 0 (u x)"]
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1603
  have "integral\<^isup>P M u = 0 \<longleftrightarrow> (AE x. max 0 (u x) = 0)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1604
    unfolding positive_integral_max_0
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1605
    using AE_iff_null_set[OF sets] u by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1606
  also have "\<dots> \<longleftrightarrow> (AE x. u x \<le> 0)" by (auto split: split_max)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1607
  finally show ?thesis .
41705
1100512e16d8 add auto support for AE_mp
hoelzl
parents: 41689
diff changeset
  1608
qed
1100512e16d8 add auto support for AE_mp
hoelzl
parents: 41689
diff changeset
  1609
42991
3fa22920bf86 integral strong monotone; finite subadditivity for measure
hoelzl
parents: 42950
diff changeset
  1610
lemma (in measure_space) positive_integral_const_If:
3fa22920bf86 integral strong monotone; finite subadditivity for measure
hoelzl
parents: 42950
diff changeset
  1611
  "(\<integral>\<^isup>+x. a \<partial>M) = (if 0 \<le> a then a * \<mu> (space M) else 0)"
3fa22920bf86 integral strong monotone; finite subadditivity for measure
hoelzl
parents: 42950
diff changeset
  1612
  by (auto intro!: positive_integral_0_iff_AE[THEN iffD2])
3fa22920bf86 integral strong monotone; finite subadditivity for measure
hoelzl
parents: 42950
diff changeset
  1613
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
  1614
lemma (in measure_space) positive_integral_restricted:
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1615
  assumes A: "A \<in> sets M"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1616
  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
  1617
    (is "integral\<^isup>P ?R f = integral\<^isup>P M ?f")
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
  1618
proof -
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1619
  interpret R: measure_space ?R
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1620
    by (rule restricted_measure_space) fact
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1621
  let "?I g x" = "g x * indicator A x :: ereal"
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
  1622
  show ?thesis
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1623
    unfolding positive_integral_def
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1624
    unfolding simple_function_restricted[OF A]
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1625
    unfolding AE_restricted[OF A]
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1626
  proof (safe intro!: SUPR_eq)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1627
    fix g assume g: "simple_function M (?I g)" and le: "g \<le> max 0 \<circ> f"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1628
    show "\<exists>j\<in>{g. simple_function M g \<and> g \<le> max 0 \<circ> ?I f}.
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1629
      integral\<^isup>S (restricted_space A) g \<le> integral\<^isup>S M j"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1630
    proof (safe intro!: bexI[of _ "?I g"])
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1631
      show "integral\<^isup>S (restricted_space A) g \<le> integral\<^isup>S M (?I g)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1632
        using g A by (simp add: simple_integral_restricted)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1633
      show "?I g \<le> max 0 \<circ> ?I f"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1634
        using le by (auto simp: le_fun_def max_def indicator_def split: split_if_asm)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1635
    qed fact
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
  1636
  next
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1637
    fix g assume g: "simple_function M g" and le: "g \<le> max 0 \<circ> ?I f"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1638
    show "\<exists>i\<in>{g. simple_function M (?I g) \<and> g \<le> max 0 \<circ> f}.
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1639
      integral\<^isup>S M g \<le> integral\<^isup>S (restricted_space A) i"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1640
    proof (safe intro!: bexI[of _ "?I g"])
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1641
      show "?I g \<le> max 0 \<circ> f"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1642
        using le by (auto simp: le_fun_def max_def indicator_def split: split_if_asm)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1643
      from le have "\<And>x. g x \<le> ?I (?I g) x"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1644
        by (auto simp: le_fun_def max_def indicator_def split: split_if_asm)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1645
      then show "integral\<^isup>S M g \<le> integral\<^isup>S (restricted_space A) (?I g)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1646
        using A g by (auto intro!: simple_integral_mono simp: simple_integral_restricted)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1647
      show "simple_function M (?I (?I g))" using g A by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1648
    qed
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
  1649
  qed
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
  1650
qed
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
  1651
41545
9c869baf1c66 tuned formalization of subalgebra
hoelzl
parents: 41544
diff changeset
  1652
lemma (in measure_space) positive_integral_subalgebra:
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1653
  assumes f: "f \<in> borel_measurable N" "AE x in N. 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
  1654
  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
  1655
  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
  1656
  shows "integral\<^isup>P N f = integral\<^isup>P M f"
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
  1657
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
  1658
  interpret N: measure_space N using measure_space_subalgebra[OF sa N] .
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1659
  from N.borel_measurable_implies_simple_function_sequence'[OF f(1)] guess fs . note fs = this
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1660
  note sf = simple_function_subalgebra[OF fs(1) N(1,2)]
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1661
  from N.positive_integral_monotone_convergence_simple[OF fs(2,5,1), 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
  1662
  have "integral\<^isup>P N f = (SUP i. \<Sum>x\<in>fs i ` space M. x * N.\<mu> (fs i -` {x} \<inter> space M))"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1663
    unfolding fs(4) positive_integral_max_0
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1664
    unfolding simple_integral_def `space N = space M` 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
  1665
  also have "\<dots> = (SUP i. \<Sum>x\<in>fs i ` space M. x * \<mu> (fs i -` {x} \<inter> space M))"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1666
    using N N.simple_functionD(2)[OF fs(1)] unfolding `space N = space M` by auto
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1667
  also have "\<dots> = integral\<^isup>P M f"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1668
    using positive_integral_monotone_convergence_simple[OF fs(2,5) sf, symmetric]
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1669
    unfolding fs(4) positive_integral_max_0
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1670
    unfolding simple_integral_def `space N = space M` 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
  1671
  finally show ?thesis .
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
  1672
qed
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
  1673
35692
f1315bbf1bc9 Moved theorems in Lebesgue to the right places
hoelzl
parents: 35582
diff changeset
  1674
section "Lebesgue Integral"
f1315bbf1bc9 Moved theorems in Lebesgue to the right places
hoelzl
parents: 35582
diff changeset
  1675
41689
3e39b0e730d6 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
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
  1677
  "integrable M f \<longleftrightarrow> f \<in> borel_measurable M \<and>
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1678
    (\<integral>\<^isup>+ x. ereal (f x) \<partial>M) \<noteq> \<infinity> \<and> (\<integral>\<^isup>+ x. ereal (- f x) \<partial>M) \<noteq> \<infinity>"
35692
f1315bbf1bc9 Moved theorems in Lebesgue to the right places
hoelzl
parents: 35582
diff changeset
  1679
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1680
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
  1681
  assumes "integrable M f"
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1682
  shows "f \<in> borel_measurable M" "(\<integral>\<^isup>+ x. ereal (f x) \<partial>M) \<noteq> \<infinity>" "(\<integral>\<^isup>+ x. ereal (- f x) \<partial>M) \<noteq> \<infinity>"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1683
  using assms unfolding integrable_def by auto
35692
f1315bbf1bc9 Moved theorems in Lebesgue to the right places
hoelzl
parents: 35582
diff changeset
  1684
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1685
definition lebesgue_integral_def:
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1686
  "integral\<^isup>L M f = real ((\<integral>\<^isup>+ x. ereal (f x) \<partial>M)) - real ((\<integral>\<^isup>+ x. ereal (- f x) \<partial>M))"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1687
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1688
syntax
45342
5c760e1692b3 proper syntactic category for abstraction syntax, to avoid low-level exception for malformed "\<integral> x y. f \<partial>M", for example;
wenzelm
parents: 44937
diff changeset
  1689
  "_lebesgue_integral" :: "pttrn \<Rightarrow> real \<Rightarrow> ('a, 'b) measure_space_scheme \<Rightarrow> real" ("\<integral> _. _ \<partial>_" [60,61] 110)
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1690
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1691
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
  1692
  "\<integral> x. f \<partial>M" == "CONST integral\<^isup>L M (%x. f)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1693
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1694
lemma (in measure_space) integrableE:
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1695
  assumes "integrable M f"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1696
  obtains r q where
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1697
    "(\<integral>\<^isup>+x. ereal (f x)\<partial>M) = ereal r"
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1698
    "(\<integral>\<^isup>+x. ereal (-f x)\<partial>M) = ereal q"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1699
    "f \<in> borel_measurable M" "integral\<^isup>L M f = r - q"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1700
  using assms unfolding integrable_def lebesgue_integral_def
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1701
  using positive_integral_positive[of "\<lambda>x. ereal (f x)"]
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1702
  using positive_integral_positive[of "\<lambda>x. ereal (-f x)"]
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1703
  by (cases rule: ereal2_cases[of "(\<integral>\<^isup>+x. ereal (-f x)\<partial>M)" "(\<integral>\<^isup>+x. ereal (f x)\<partial>M)"]) auto
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1704
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1705
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
  1706
  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
  1707
  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
  1708
  using assms by (simp cong: positive_integral_cong add: lebesgue_integral_def)
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1709
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1710
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
  1711
  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
  1712
  shows "integral\<^isup>L N f = integral\<^isup>L M f"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1713
  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
  1714
43339
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
  1715
lemma (in measure_space) integrable_cong_measure:
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
  1716
  assumes "\<And>A. A \<in> sets M \<Longrightarrow> measure N A = \<mu> A" "sets N = sets M" "space N = space M"
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
  1717
  shows "integrable N f \<longleftrightarrow> integrable M f"
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
  1718
  using assms
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
  1719
  by (simp add: positive_integral_cong_measure[OF assms] integrable_def measurable_def)
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
  1720
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1721
lemma (in measure_space) integral_cong_AE:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1722
  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
  1723
  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
  1724
proof -
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1725
  have *: "AE x. ereal (f x) = ereal (g x)"
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1726
    "AE x. ereal (- f x) = ereal (- g x)" using cong by auto
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1727
  show ?thesis
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1728
    unfolding *[THEN positive_integral_cong_AE] lebesgue_integral_def ..
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1729
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1730
43339
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
  1731
lemma (in measure_space) integrable_cong_AE:
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
  1732
  assumes borel: "f \<in> borel_measurable M" "g \<in> borel_measurable M"
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
  1733
  assumes "AE x. f x = g x"
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
  1734
  shows "integrable M f = integrable M g"
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
  1735
proof -
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1736
  have "(\<integral>\<^isup>+ x. ereal (f x) \<partial>M) = (\<integral>\<^isup>+ x. ereal (g x) \<partial>M)"
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1737
    "(\<integral>\<^isup>+ x. ereal (- f x) \<partial>M) = (\<integral>\<^isup>+ x. ereal (- g x) \<partial>M)"
43339
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
  1738
    using assms by (auto intro!: positive_integral_cong_AE)
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
  1739
  with assms show ?thesis
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
  1740
    by (auto simp: integrable_def)
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
  1741
qed
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
  1742
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1743
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
  1744
  "(\<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
  1745
  by (simp cong: positive_integral_cong measurable_cong add: integrable_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1746
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1747
lemma (in measure_space) integral_eq_positive_integral:
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1748
  assumes f: "\<And>x. 0 \<le> f x"
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1749
  shows "integral\<^isup>L M f = real (\<integral>\<^isup>+ x. ereal (f x) \<partial>M)"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1750
proof -
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1751
  { fix x have "max 0 (ereal (- f x)) = 0" using f[of x] by (simp split: split_max) }
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1752
  then have "0 = (\<integral>\<^isup>+ x. max 0 (ereal (- f x)) \<partial>M)" by simp
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1753
  also have "\<dots> = (\<integral>\<^isup>+ x. ereal (- f x) \<partial>M)" unfolding positive_integral_max_0 ..
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1754
  finally show ?thesis
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1755
    unfolding lebesgue_integral_def by simp
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1756
qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1757
41661
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41545
diff changeset
  1758
lemma (in measure_space) integral_vimage:
41831
91a2b435dd7a use measure_preserving in ..._vimage lemmas
hoelzl
parents: 41705
diff changeset
  1759
  assumes T: "sigma_algebra M'" "T \<in> measure_preserving M M'"
91a2b435dd7a use measure_preserving in ..._vimage lemmas
hoelzl
parents: 41705
diff changeset
  1760
  assumes f: "f \<in> borel_measurable M'"
91a2b435dd7a use measure_preserving in ..._vimage lemmas
hoelzl
parents: 41705
diff changeset
  1761
  shows "integral\<^isup>L M' f = (\<integral>x. f (T x) \<partial>M)"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1762
proof -
41831
91a2b435dd7a use measure_preserving in ..._vimage lemmas
hoelzl
parents: 41705
diff changeset
  1763
  interpret T: measure_space M' by (rule measure_space_vimage[OF T])
91a2b435dd7a use measure_preserving in ..._vimage lemmas
hoelzl
parents: 41705
diff changeset
  1764
  from measurable_comp[OF measure_preservingD2[OF T(2)], of f borel]
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1765
  have borel: "(\<lambda>x. ereal (f x)) \<in> borel_measurable M'" "(\<lambda>x. ereal (- f x)) \<in> borel_measurable M'"
41661
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41545
diff changeset
  1766
    and "(\<lambda>x. f (T x)) \<in> borel_measurable M"
41831
91a2b435dd7a use measure_preserving in ..._vimage lemmas
hoelzl
parents: 41705
diff changeset
  1767
    using f by (auto simp: comp_def)
91a2b435dd7a use measure_preserving in ..._vimage lemmas
hoelzl
parents: 41705
diff changeset
  1768
  then 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
  1769
    using f unfolding lebesgue_integral_def integrable_def
41831
91a2b435dd7a use measure_preserving in ..._vimage lemmas
hoelzl
parents: 41705
diff changeset
  1770
    by (auto simp: borel[THEN positive_integral_vimage[OF T]])
91a2b435dd7a use measure_preserving in ..._vimage lemmas
hoelzl
parents: 41705
diff changeset
  1771
qed
91a2b435dd7a use measure_preserving in ..._vimage lemmas
hoelzl
parents: 41705
diff changeset
  1772
91a2b435dd7a use measure_preserving in ..._vimage lemmas
hoelzl
parents: 41705
diff changeset
  1773
lemma (in measure_space) integrable_vimage:
91a2b435dd7a use measure_preserving in ..._vimage lemmas
hoelzl
parents: 41705
diff changeset
  1774
  assumes T: "sigma_algebra M'" "T \<in> measure_preserving M M'"
91a2b435dd7a use measure_preserving in ..._vimage lemmas
hoelzl
parents: 41705
diff changeset
  1775
  assumes f: "integrable M' f"
91a2b435dd7a use measure_preserving in ..._vimage lemmas
hoelzl
parents: 41705
diff changeset
  1776
  shows "integrable M (\<lambda>x. f (T x))"
91a2b435dd7a use measure_preserving in ..._vimage lemmas
hoelzl
parents: 41705
diff changeset
  1777
proof -
91a2b435dd7a use measure_preserving in ..._vimage lemmas
hoelzl
parents: 41705
diff changeset
  1778
  interpret T: measure_space M' by (rule measure_space_vimage[OF T])
91a2b435dd7a use measure_preserving in ..._vimage lemmas
hoelzl
parents: 41705
diff changeset
  1779
  from measurable_comp[OF measure_preservingD2[OF T(2)], of f borel]
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1780
  have borel: "(\<lambda>x. ereal (f x)) \<in> borel_measurable M'" "(\<lambda>x. ereal (- f x)) \<in> borel_measurable M'"
41831
91a2b435dd7a use measure_preserving in ..._vimage lemmas
hoelzl
parents: 41705
diff changeset
  1781
    and "(\<lambda>x. f (T x)) \<in> borel_measurable M"
91a2b435dd7a use measure_preserving in ..._vimage lemmas
hoelzl
parents: 41705
diff changeset
  1782
    using f by (auto simp: comp_def)
91a2b435dd7a use measure_preserving in ..._vimage lemmas
hoelzl
parents: 41705
diff changeset
  1783
  then show ?thesis
91a2b435dd7a use measure_preserving in ..._vimage lemmas
hoelzl
parents: 41705
diff changeset
  1784
    using f unfolding lebesgue_integral_def integrable_def
91a2b435dd7a use measure_preserving in ..._vimage lemmas
hoelzl
parents: 41705
diff changeset
  1785
    by (auto simp: borel[THEN positive_integral_vimage[OF T]])
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1786
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1787
43339
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
  1788
lemma (in measure_space) integral_translated_density:
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
  1789
  assumes f: "f \<in> borel_measurable M" "AE x. 0 \<le> f x"
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
  1790
    and g: "g \<in> borel_measurable M"
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
  1791
    and N: "space N = space M" "sets N = sets M"
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
  1792
    and density: "\<And>A. A \<in> sets M \<Longrightarrow> measure N A = (\<integral>\<^isup>+ x. f x * indicator A x \<partial>M)"
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
  1793
      (is "\<And>A. _ \<Longrightarrow> _ = ?d A")
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
  1794
  shows "integral\<^isup>L N g = (\<integral> x. f x * g x \<partial>M)" (is ?integral)
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
  1795
    and "integrable N g = integrable M (\<lambda>x. f x * g x)" (is ?integrable)
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
  1796
proof -
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
  1797
  from f have ms: "measure_space (M\<lparr>measure := ?d\<rparr>)"
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1798
    by (intro measure_space_density[where u="\<lambda>x. ereal (f x)"]) auto
43339
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
  1799
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1800
  from ms density N have "(\<integral>\<^isup>+ x. g x \<partial>N) =  (\<integral>\<^isup>+ x. max 0 (ereal (g x)) \<partial>M\<lparr>measure := ?d\<rparr>)"
43339
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
  1801
    unfolding positive_integral_max_0
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
  1802
    by (intro measure_space.positive_integral_cong_measure) auto
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1803
  also have "\<dots> = (\<integral>\<^isup>+ x. ereal (f x) * max 0 (ereal (g x)) \<partial>M)"
43339
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
  1804
    using f g by (intro positive_integral_translated_density) auto
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1805
  also have "\<dots> = (\<integral>\<^isup>+ x. max 0 (ereal (f x * g x)) \<partial>M)"
43339
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
  1806
    using f by (intro positive_integral_cong_AE)
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1807
               (auto simp: ereal_max_0 zero_le_mult_iff split: split_max)
43339
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
  1808
  finally have pos: "(\<integral>\<^isup>+ x. g x \<partial>N) = (\<integral>\<^isup>+ x. f x * g x \<partial>M)"
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
  1809
    by (simp add: positive_integral_max_0)
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
  1810
  
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1811
  from ms density N have "(\<integral>\<^isup>+ x. - (g x) \<partial>N) =  (\<integral>\<^isup>+ x. max 0 (ereal (- g x)) \<partial>M\<lparr>measure := ?d\<rparr>)"
43339
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
  1812
    unfolding positive_integral_max_0
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
  1813
    by (intro measure_space.positive_integral_cong_measure) auto
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1814
  also have "\<dots> = (\<integral>\<^isup>+ x. ereal (f x) * max 0 (ereal (- g x)) \<partial>M)"
43339
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
  1815
    using f g by (intro positive_integral_translated_density) auto
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1816
  also have "\<dots> = (\<integral>\<^isup>+ x. max 0 (ereal (- f x * g x)) \<partial>M)"
43339
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
  1817
    using f by (intro positive_integral_cong_AE)
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1818
               (auto simp: ereal_max_0 mult_le_0_iff split: split_max)
43339
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
  1819
  finally have neg: "(\<integral>\<^isup>+ x. - g x \<partial>N) = (\<integral>\<^isup>+ x. - (f x * g x) \<partial>M)"
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
  1820
    by (simp add: positive_integral_max_0)
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
  1821
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
  1822
  have g_N: "g \<in> borel_measurable N"
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
  1823
    using g N unfolding measurable_def by simp
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
  1824
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
  1825
  show ?integral ?integrable
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
  1826
    unfolding lebesgue_integral_def integrable_def
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
  1827
    using pos neg f g g_N by auto
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
  1828
qed
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
  1829
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1830
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
  1831
  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
  1832
  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
  1833
  using assms by (auto simp: integrable_def lebesgue_integral_def)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1834
42991
3fa22920bf86 integral strong monotone; finite subadditivity for measure
hoelzl
parents: 42950
diff changeset
  1835
lemma (in measure_space) integral_minus_iff[simp]:
3fa22920bf86 integral strong monotone; finite subadditivity for measure
hoelzl
parents: 42950
diff changeset
  1836
  "integrable M (\<lambda>x. - f x) \<longleftrightarrow> integrable M f"
3fa22920bf86 integral strong monotone; finite subadditivity for measure
hoelzl
parents: 42950
diff changeset
  1837
proof
3fa22920bf86 integral strong monotone; finite subadditivity for measure
hoelzl
parents: 42950
diff changeset
  1838
  assume "integrable M (\<lambda>x. - f x)"
3fa22920bf86 integral strong monotone; finite subadditivity for measure
hoelzl
parents: 42950
diff changeset
  1839
  then have "integrable M (\<lambda>x. - (- f x))"
3fa22920bf86 integral strong monotone; finite subadditivity for measure
hoelzl
parents: 42950
diff changeset
  1840
    by (rule integral_minus)
3fa22920bf86 integral strong monotone; finite subadditivity for measure
hoelzl
parents: 42950
diff changeset
  1841
  then show "integrable M f" by simp
3fa22920bf86 integral strong monotone; finite subadditivity for measure
hoelzl
parents: 42950
diff changeset
  1842
qed (rule integral_minus)
3fa22920bf86 integral strong monotone; finite subadditivity for measure
hoelzl
parents: 42950
diff changeset
  1843
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1844
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
  1845
  assumes integrable: "integrable M u" "integrable M v"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1846
  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
  1847
  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
  1848
proof -
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1849
  let "?f x" = "max 0 (ereal (f x))"
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1850
  let "?mf x" = "max 0 (ereal (- f x))"
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1851
  let "?u x" = "max 0 (ereal (u x))"
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1852
  let "?v x" = "max 0 (ereal (v x))"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1853
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1854
  from borel_measurable_diff[of u v] integrable
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1855
  have f_borel: "?f \<in> borel_measurable M" and
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1856
    mf_borel: "?mf \<in> borel_measurable M" and
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1857
    v_borel: "?v \<in> borel_measurable M" and
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1858
    u_borel: "?u \<in> borel_measurable M" and
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1859
    "f \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1860
    by (auto simp: f_def[symmetric] integrable_def)
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1861
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1862
  have "(\<integral>\<^isup>+ x. ereal (u x - v x) \<partial>M) \<le> integral\<^isup>P M ?u"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1863
    using pos by (auto intro!: positive_integral_mono simp: positive_integral_max_0)
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1864
  moreover have "(\<integral>\<^isup>+ x. ereal (v x - u x) \<partial>M) \<le> integral\<^isup>P M ?v"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1865
    using pos by (auto intro!: positive_integral_mono simp: positive_integral_max_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
  1866
  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
  1867
    using `integrable M u` `integrable M v` `f \<in> borel_measurable M`
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1868
    by (auto simp: integrable_def f_def positive_integral_max_0)
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1869
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1870
  have "\<And>x. ?u x + ?mf x = ?v x + ?f x"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1871
    unfolding f_def using pos by (simp split: split_max)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1872
  then have "(\<integral>\<^isup>+ x. ?u x + ?mf x \<partial>M) = (\<integral>\<^isup>+ x. ?v x + ?f x \<partial>M)" by simp
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1873
  then have "real (integral\<^isup>P M ?u + integral\<^isup>P M ?mf) =
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1874
      real (integral\<^isup>P M ?v + integral\<^isup>P M ?f)"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1875
    using positive_integral_add[OF u_borel _ mf_borel]
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1876
    using positive_integral_add[OF v_borel _ f_borel]
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1877
    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
  1878
  then show "integral\<^isup>L M f = integral\<^isup>L M u - integral\<^isup>L M v"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1879
    unfolding positive_integral_max_0
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1880
    unfolding pos[THEN integral_eq_positive_integral]
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1881
    using integrable f by (auto elim!: integrableE)
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1882
qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1883
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1884
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
  1885
  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
  1886
  shows "integrable M (\<lambda>t. a * f t + g t)"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1887
  and "(\<integral> t. a * f t + g t \<partial>M) = a * integral\<^isup>L M f + integral\<^isup>L M g" (is ?EQ)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1888
proof -
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1889
  let "?f x" = "max 0 (ereal (f x))"
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1890
  let "?g x" = "max 0 (ereal (g x))"
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1891
  let "?mf x" = "max 0 (ereal (- f x))"
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1892
  let "?mg x" = "max 0 (ereal (- g x))"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1893
  let "?p t" = "max 0 (a * f t) + max 0 (g t)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1894
  let "?n t" = "max 0 (- (a * f t)) + max 0 (- g t)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1895
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1896
  from assms have linear:
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1897
    "(\<integral>\<^isup>+ x. ereal a * ?f x + ?g x \<partial>M) = ereal a * integral\<^isup>P M ?f + integral\<^isup>P M ?g"
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1898
    "(\<integral>\<^isup>+ x. ereal a * ?mf x + ?mg x \<partial>M) = ereal a * integral\<^isup>P M ?mf + integral\<^isup>P M ?mg"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1899
    by (auto intro!: positive_integral_linear simp: integrable_def)
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1900
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1901
  have *: "(\<integral>\<^isup>+x. ereal (- ?p x) \<partial>M) = 0" "(\<integral>\<^isup>+x. ereal (- ?n x) \<partial>M) = 0"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1902
    using `0 \<le> a` assms by (auto simp: positive_integral_0_iff_AE integrable_def)
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1903
  have **: "\<And>x. ereal a * ?f x + ?g x = max 0 (ereal (?p x))"
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1904
           "\<And>x. ereal a * ?mf x + ?mg x = max 0 (ereal (?n x))"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1905
    using `0 \<le> a` by (auto split: split_max simp: zero_le_mult_iff mult_le_0_iff)
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1906
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1907
  have "integrable M ?p" "integrable M ?n"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1908
      "\<And>t. a * f t + g t = ?p t - ?n t" "\<And>t. 0 \<le> ?p t" "\<And>t. 0 \<le> ?n t"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1909
    using linear assms unfolding integrable_def ** *
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1910
    by (auto simp: positive_integral_max_0)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1911
  note diff = integral_of_positive_diff[OF this]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1912
41689
3e39b0e730d6 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
  show "integrable M (\<lambda>t. a * f t + g t)" by (rule diff)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1914
  from assms linear show ?EQ
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1915
    unfolding diff(2) ** positive_integral_max_0
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1916
    unfolding lebesgue_integral_def *
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1917
    by (auto elim!: integrableE simp: field_simps)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1918
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1919
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1920
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
  1921
  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
  1922
  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
  1923
  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
  1924
  using assms integral_linear[where a=1] by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1925
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1926
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
  1927
  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
  1928
  unfolding integrable_def lebesgue_integral_def
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1929
  by (auto simp add: borel_measurable_const)
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1930
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1931
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
  1932
  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
  1933
  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
  1934
  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
  1935
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
  1936
  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
  1937
  proof (cases rule: le_cases)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1938
    assume "0 \<le> a" show ?thesis
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1939
      using integral_linear[OF assms integral_zero(1) `0 \<le> a`]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1940
      by (simp add: integral_zero)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1941
  next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1942
    assume "a \<le> 0" hence "0 \<le> - a" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1943
    have *: "\<And>t. - a * t + 0 = (-a) * t" by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1944
    show ?thesis using integral_linear[OF assms integral_zero(1) `0 \<le> - a`]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1945
        integral_minus(1)[of "\<lambda>t. - a * f t"]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1946
      unfolding * integral_zero by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1947
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1948
  thus ?P ?I by auto
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1949
qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1950
41096
843c40bbc379 integral over setprod
hoelzl
parents: 41095
diff changeset
  1951
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
  1952
  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
  1953
  shows "(\<integral> x. f x * c \<partial>M) = integral\<^isup>L M f * c"
41096
843c40bbc379 integral over setprod
hoelzl
parents: 41095
diff changeset
  1954
  unfolding mult_commute[of _ c] integral_cmult[OF assms] ..
843c40bbc379 integral over setprod
hoelzl
parents: 41095
diff changeset
  1955
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1956
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
  1957
  assumes fg: "integrable M f" "integrable M g"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1958
  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
  1959
  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
  1960
proof -
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1961
  have "AE x. ereal (f x) \<le> ereal (g x)"
41705
1100512e16d8 add auto support for AE_mp
hoelzl
parents: 41689
diff changeset
  1962
    using mono by auto
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1963
  moreover have "AE x. ereal (- g x) \<le> ereal (- f x)"
41705
1100512e16d8 add auto support for AE_mp
hoelzl
parents: 41689
diff changeset
  1964
    using mono by auto
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1965
  ultimately show ?thesis using fg
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1966
    by (auto intro!: add_mono positive_integral_mono_AE real_of_ereal_positive_mono
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1967
             simp: positive_integral_positive lebesgue_integral_def diff_minus)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1968
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  1969
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1970
lemma (in measure_space) integral_mono:
41705
1100512e16d8 add auto support for AE_mp
hoelzl
parents: 41689
diff changeset
  1971
  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
  1972
  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
  1973
  using assms by (auto intro: integral_mono_AE)
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1974
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1975
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
  1976
  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
  1977
  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
  1978
  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
  1979
  using integral_add[OF f integral_minus(1)[OF g]]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1980
  unfolding diff_minus integral_minus(2)[OF g]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1981
  by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1982
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1983
lemma (in measure_space) integral_indicator[simp, intro]:
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1984
  assumes "A \<in> sets M" and "\<mu> A \<noteq> \<infinity>"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1985
  shows "integral\<^isup>L M (indicator A) = real (\<mu> A)" (is ?int)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1986
  and "integrable M (indicator A)" (is ?able)
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1987
proof -
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1988
  from `A \<in> sets M` have *:
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1989
    "\<And>x. ereal (indicator A x) = indicator A x"
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1990
    "(\<integral>\<^isup>+x. ereal (- indicator A x) \<partial>M) = 0"
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  1991
    by (auto split: split_indicator simp: positive_integral_0_iff_AE one_ereal_def)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1992
  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
  1993
    using assms unfolding lebesgue_integral_def integrable_def
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1994
    by (auto simp: * positive_integral_indicator borel_measurable_indicator)
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1995
qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1996
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  1997
lemma (in measure_space) integral_cmul_indicator:
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1998
  assumes "A \<in> sets M" and "c \<noteq> 0 \<Longrightarrow> \<mu> A \<noteq> \<infinity>"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1999
  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
  2000
  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
  2001
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2002
  show ?P
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2003
  proof (cases "c = 0")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2004
    case False with assms show ?thesis by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2005
  qed simp
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  2006
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2007
  show ?I
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2008
  proof (cases "c = 0")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2009
    case False with assms show ?thesis by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2010
  qed simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2011
qed
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  2012
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2013
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
  2014
  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
  2015
  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
  2016
    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
  2017
proof -
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2018
  have "?int S \<and> ?I S"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2019
  proof (cases "finite S")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2020
    assume "finite S"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2021
    from this assms show ?thesis by (induct S) simp_all
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2022
  qed simp
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  2023
  thus "?int S" and "?I S" by auto
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  2024
qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  2025
36624
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
  2026
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
  2027
  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
  2028
  shows "integrable M (\<lambda> x. \<bar>f x\<bar>)"
36624
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
  2029
proof -
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  2030
  from assms have *: "(\<integral>\<^isup>+x. ereal (- \<bar>f x\<bar>)\<partial>M) = 0"
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  2031
    "\<And>x. ereal \<bar>f x\<bar> = max 0 (ereal (f x)) + max 0 (ereal (- f x))"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  2032
    by (auto simp: integrable_def positive_integral_0_iff_AE split: split_max)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  2033
  with assms show ?thesis
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  2034
    by (simp add: positive_integral_add positive_integral_max_0 integrable_def)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2035
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2036
41545
9c869baf1c66 tuned formalization of subalgebra
hoelzl
parents: 41544
diff changeset
  2037
lemma (in measure_space) integral_subalgebra:
9c869baf1c66 tuned formalization of subalgebra
hoelzl
parents: 41544
diff changeset
  2038
  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
  2039
  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
  2040
  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
  2041
    and "integral\<^isup>L N f = integral\<^isup>L M f" (is ?I)
41545
9c869baf1c66 tuned formalization of subalgebra
hoelzl
parents: 41544
diff changeset
  2042
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
  2043
  interpret N: measure_space N using measure_space_subalgebra[OF sa N] .
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  2044
  have "(\<integral>\<^isup>+ x. max 0 (ereal (f x)) \<partial>N) = (\<integral>\<^isup>+ x. max 0 (ereal (f x)) \<partial>M)"
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  2045
       "(\<integral>\<^isup>+ x. max 0 (ereal (- f x)) \<partial>N) = (\<integral>\<^isup>+ x. max 0 (ereal (- f x)) \<partial>M)"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  2046
    using borel by (auto intro!: positive_integral_subalgebra N sa)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  2047
  moreover have "f \<in> borel_measurable M \<longleftrightarrow> f \<in> borel_measurable N"
41545
9c869baf1c66 tuned formalization of subalgebra
hoelzl
parents: 41544
diff changeset
  2048
    using assms unfolding measurable_def by auto
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  2049
  ultimately show ?P ?I
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  2050
    by (auto simp: integrable_def lebesgue_integral_def positive_integral_max_0)
41545
9c869baf1c66 tuned formalization of subalgebra
hoelzl
parents: 41544
diff changeset
  2051
qed
9c869baf1c66 tuned formalization of subalgebra
hoelzl
parents: 41544
diff changeset
  2052
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2053
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
  2054
  assumes "integrable M f"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2055
  and f: "\<And>x. x \<in> space M \<Longrightarrow> 0 \<le> f x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2056
    "\<And>x. x \<in> space M \<Longrightarrow> \<bar>g x\<bar> \<le> f x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2057
  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
  2058
  shows "integrable M g"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2059
proof -
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  2060
  have "(\<integral>\<^isup>+ x. ereal (g x) \<partial>M) \<le> (\<integral>\<^isup>+ x. ereal \<bar>g x\<bar> \<partial>M)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2061
    by (auto intro!: positive_integral_mono)
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  2062
  also have "\<dots> \<le> (\<integral>\<^isup>+ x. ereal (f x) \<partial>M)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2063
    using f by (auto intro!: positive_integral_mono)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  2064
  also have "\<dots> < \<infinity>"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  2065
    using `integrable M f` unfolding integrable_def by auto
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  2066
  finally have pos: "(\<integral>\<^isup>+ x. ereal (g x) \<partial>M) < \<infinity>" .
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2067
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  2068
  have "(\<integral>\<^isup>+ x. ereal (- g x) \<partial>M) \<le> (\<integral>\<^isup>+ x. ereal (\<bar>g x\<bar>) \<partial>M)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2069
    by (auto intro!: positive_integral_mono)
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  2070
  also have "\<dots> \<le> (\<integral>\<^isup>+ x. ereal (f x) \<partial>M)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2071
    using f by (auto intro!: positive_integral_mono)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  2072
  also have "\<dots> < \<infinity>"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  2073
    using `integrable M f` unfolding integrable_def by auto
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  2074
  finally have neg: "(\<integral>\<^isup>+ x. ereal (- g x) \<partial>M) < \<infinity>" .
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2075
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2076
  from neg pos borel show ?thesis
36624
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
  2077
    unfolding integrable_def by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2078
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2079
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2080
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
  2081
  "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
  2082
  by (auto intro!: integrable_bound[where g=f] integrable_abs)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2083
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2084
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
  2085
  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
  2086
  shows "integrable M (\<lambda> x. max (f x) (g x))"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2087
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
  2088
  show "integrable M (\<lambda>x. \<bar>f x\<bar> + \<bar>g x\<bar>)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2089
    using int by (simp add: integrable_abs)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2090
  show "(\<lambda>x. max (f x) (g x)) \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2091
    using int unfolding integrable_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2092
next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2093
  fix x assume "x \<in> space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2094
  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
  2095
    by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2096
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2097
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2098
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
  2099
  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
  2100
  shows "integrable M (\<lambda> x. min (f x) (g x))"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2101
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
  2102
  show "integrable M (\<lambda>x. \<bar>f x\<bar> + \<bar>g x\<bar>)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2103
    using int by (simp add: integrable_abs)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2104
  show "(\<lambda>x. min (f x) (g x)) \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2105
    using int unfolding integrable_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2106
next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2107
  fix x assume "x \<in> space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2108
  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
  2109
    by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2110
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2111
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2112
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
  2113
  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
  2114
  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
  2115
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
  2116
  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
  2117
  also have "\<dots> \<le> (\<integral>x. \<bar>f x\<bar> \<partial>M)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2118
      using assms integral_minus(2)[of f, symmetric]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2119
      by (auto intro!: integral_mono integrable_abs simp del: integral_minus)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2120
  finally show ?thesis .
36624
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
  2121
qed
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
  2122
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2123
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
  2124
  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
  2125
  shows "0 \<le> integral\<^isup>L M f"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2126
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
  2127
  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
  2128
  also have "\<dots> \<le> integral\<^isup>L M f"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2129
    using assms by (rule integral_mono[OF integral_zero(1)])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2130
  finally show ?thesis .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2131
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2132
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2133
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
  2134
  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
  2135
  and pos: "\<And>x i. 0 \<le> f i x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2136
  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
  2137
  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
  2138
  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
  2139
  and "integral\<^isup>L M u = x"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  2140
proof -
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2141
  { fix x have "0 \<le> u x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2142
      using mono pos[of 0 x] incseq_le[OF _ lim, of x 0]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2143
      by (simp add: mono_def incseq_def) }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2144
  note pos_u = this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2145
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  2146
  have SUP_F: "\<And>x. (SUP n. ereal (f n x)) = ereal (u x)"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  2147
    unfolding SUP_eq_LIMSEQ[OF mono] by (rule lim)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2148
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  2149
  have borel_f: "\<And>i. (\<lambda>x. ereal (f i x)) \<in> borel_measurable M"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2150
    using i unfolding integrable_def by auto
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  2151
  hence "(\<lambda>x. SUP i. ereal (f i x)) \<in> borel_measurable M"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  2152
    by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2153
  hence borel_u: "u \<in> borel_measurable M"
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  2154
    by (auto simp: borel_measurable_ereal_iff SUP_F)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2155
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  2156
  hence [simp]: "\<And>i. (\<integral>\<^isup>+x. ereal (- f i x) \<partial>M) = 0" "(\<integral>\<^isup>+x. ereal (- u x) \<partial>M) = 0"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  2157
    using i borel_u pos pos_u by (auto simp: positive_integral_0_iff_AE integrable_def)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  2158
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  2159
  have integral_eq: "\<And>n. (\<integral>\<^isup>+ x. ereal (f n x) \<partial>M) = ereal (integral\<^isup>L M (f n))"
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  2160
    using i positive_integral_positive by (auto simp: ereal_real lebesgue_integral_def integrable_def)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2161
41689
3e39b0e730d6 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
  have pos_integral: "\<And>n. 0 \<le> integral\<^isup>L M (f n)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2163
    using pos i by (auto simp: integral_positive)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2164
  hence "0 \<le> x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2165
    using LIMSEQ_le_const[OF ilim, of 0] by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2166
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  2167
  from mono pos i have pI: "(\<integral>\<^isup>+ x. ereal (u x) \<partial>M) = (SUP n. (\<integral>\<^isup>+ x. ereal (f n x) \<partial>M))"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  2168
    by (auto intro!: positive_integral_monotone_convergence_SUP
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  2169
      simp: integrable_def incseq_mono incseq_Suc_iff le_fun_def SUP_F[symmetric])
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  2170
  also have "\<dots> = ereal x" unfolding integral_eq
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2171
  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
  2172
    show "mono (\<lambda>n. integral\<^isup>L M (f n))"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2173
      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
  2174
    show "(\<lambda>n. integral\<^isup>L M (f n)) ----> x" using ilim .
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2175
  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
  2176
  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
  2177
    unfolding integrable_def lebesgue_integral_def by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2178
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2179
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2180
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
  2181
  assumes f: "\<And>i. integrable M (f i)" and "mono f"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2182
  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
  2183
  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
  2184
  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
  2185
  and "integral\<^isup>L M u = x"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2186
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
  2187
  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
  2188
      using f by (auto intro!: integral_diff)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2189
  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
  2190
      unfolding mono_def le_fun_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2191
  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
  2192
      unfolding mono_def le_fun_def by (auto simp: field_simps)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2193
  have 4: "\<And>x. (\<lambda>i. f i x - f 0 x) ----> u x - f 0 x"
44568
e6f291cb5810 discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
huffman
parents: 43941
diff changeset
  2194
    using lim by (auto intro!: tendsto_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
  2195
  have 5: "(\<lambda>i. (\<integral>x. f i x - f 0 x \<partial>M)) ----> x - integral\<^isup>L M (f 0)"
44568
e6f291cb5810 discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
huffman
parents: 43941
diff changeset
  2196
    using f ilim by (auto intro!: tendsto_diff simp: integral_diff)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2197
  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
  2198
  have "integrable M (\<lambda>x. (u x - f 0 x) + f 0 x)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2199
    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
  2200
  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
  2201
    by (auto simp: integral_diff)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2202
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2203
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2204
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
  2205
  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
  2206
  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
  2207
proof -
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  2208
  have *: "(\<integral>\<^isup>+x. ereal (- \<bar>f x\<bar>) \<partial>M) = 0"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  2209
    using assms by (auto simp: positive_integral_0_iff_AE integrable_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
  2210
  have "integrable M (\<lambda>x. \<bar>f x\<bar>)" using assms by (rule integrable_abs)
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  2211
  hence "(\<lambda>x. ereal (\<bar>f x\<bar>)) \<in> borel_measurable M"
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  2212
    "(\<integral>\<^isup>+ x. ereal \<bar>f x\<bar> \<partial>M) \<noteq> \<infinity>" unfolding integrable_def by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2213
  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
  2214
  show ?thesis unfolding lebesgue_integral_def *
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  2215
    using positive_integral_positive[of "\<lambda>x. ereal \<bar>f x\<bar>"]
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  2216
    by (auto simp add: real_of_ereal_eq_0)
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  2217
qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  2218
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  2219
lemma (in measure_space) positive_integral_PInf:
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  2220
  assumes f: "f \<in> borel_measurable M"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  2221
  and not_Inf: "integral\<^isup>P M f \<noteq> \<infinity>"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  2222
  shows "\<mu> (f -` {\<infinity>} \<inter> space M) = 0"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  2223
proof -
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  2224
  have "\<infinity> * \<mu> (f -` {\<infinity>} \<inter> space M) = (\<integral>\<^isup>+ x. \<infinity> * indicator (f -` {\<infinity>} \<inter> space M) x \<partial>M)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  2225
    using f by (subst positive_integral_cmult_indicator) (auto simp: measurable_sets)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  2226
  also have "\<dots> \<le> integral\<^isup>P M (\<lambda>x. max 0 (f x))"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  2227
    by (auto intro!: positive_integral_mono simp: indicator_def max_def)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  2228
  finally have "\<infinity> * \<mu> (f -` {\<infinity>} \<inter> space M) \<le> integral\<^isup>P M f"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  2229
    by (simp add: positive_integral_max_0)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  2230
  moreover have "0 \<le> \<mu> (f -` {\<infinity>} \<inter> space M)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  2231
    using f by (simp add: measurable_sets)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  2232
  ultimately show ?thesis
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  2233
    using assms by (auto split: split_if_asm)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  2234
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  2235
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  2236
lemma (in measure_space) positive_integral_PInf_AE:
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  2237
  assumes "f \<in> borel_measurable M" "integral\<^isup>P M f \<noteq> \<infinity>" shows "AE x. f x \<noteq> \<infinity>"
41026
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41023
diff changeset
  2238
proof (rule AE_I)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  2239
  show "\<mu> (f -` {\<infinity>} \<inter> space M) = 0"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  2240
    by (rule positive_integral_PInf[OF assms])
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  2241
  show "f -` {\<infinity>} \<inter> space M \<in> sets M"
41026
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41023
diff changeset
  2242
    using assms by (auto intro: borel_measurable_vimage)
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41023
diff changeset
  2243
qed auto
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41023
diff changeset
  2244
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  2245
lemma (in measure_space) simple_integral_PInf:
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  2246
  assumes "simple_function M f" "\<And>x. 0 \<le> f x"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  2247
  and "integral\<^isup>S M f \<noteq> \<infinity>"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  2248
  shows "\<mu> (f -` {\<infinity>} \<inter> space M) = 0"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  2249
proof (rule positive_integral_PInf)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  2250
  show "f \<in> borel_measurable M" using assms by (auto intro: borel_measurable_simple_function)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  2251
  show "integral\<^isup>P M f \<noteq> \<infinity>"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  2252
    using assms by (simp add: positive_integral_eq_simple_integral)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  2253
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 40786
diff changeset
  2254
41026
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41023
diff changeset
  2255
lemma (in measure_space) integral_real:
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  2256
  "AE x. \<bar>f x\<bar> \<noteq> \<infinity> \<Longrightarrow> (\<integral>x. real (f x) \<partial>M) = real (integral\<^isup>P M f) - real (integral\<^isup>P M (\<lambda>x. - f x))"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  2257
  using assms unfolding lebesgue_integral_def
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  2258
  by (subst (1 2) positive_integral_cong_AE) (auto simp add: ereal_real)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  2259
42991
3fa22920bf86 integral strong monotone; finite subadditivity for measure
hoelzl
parents: 42950
diff changeset
  2260
lemma (in finite_measure) lebesgue_integral_const[simp]:
3fa22920bf86 integral strong monotone; finite subadditivity for measure
hoelzl
parents: 42950
diff changeset
  2261
  shows "integrable M (\<lambda>x. a)"
3fa22920bf86 integral strong monotone; finite subadditivity for measure
hoelzl
parents: 42950
diff changeset
  2262
  and  "(\<integral>x. a \<partial>M) = a * \<mu>' (space M)"
3fa22920bf86 integral strong monotone; finite subadditivity for measure
hoelzl
parents: 42950
diff changeset
  2263
proof -
3fa22920bf86 integral strong monotone; finite subadditivity for measure
hoelzl
parents: 42950
diff changeset
  2264
  { fix a :: real assume "0 \<le> a"
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  2265
    then have "(\<integral>\<^isup>+ x. ereal a \<partial>M) = ereal a * \<mu> (space M)"
42991
3fa22920bf86 integral strong monotone; finite subadditivity for measure
hoelzl
parents: 42950
diff changeset
  2266
      by (subst positive_integral_const) auto
3fa22920bf86 integral strong monotone; finite subadditivity for measure
hoelzl
parents: 42950
diff changeset
  2267
    moreover
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  2268
    from `0 \<le> a` have "(\<integral>\<^isup>+ x. ereal (-a) \<partial>M) = 0"
42991
3fa22920bf86 integral strong monotone; finite subadditivity for measure
hoelzl
parents: 42950
diff changeset
  2269
      by (subst positive_integral_0_iff_AE) auto
3fa22920bf86 integral strong monotone; finite subadditivity for measure
hoelzl
parents: 42950
diff changeset
  2270
    ultimately have "integrable M (\<lambda>x. a)" by (auto simp: integrable_def) }
3fa22920bf86 integral strong monotone; finite subadditivity for measure
hoelzl
parents: 42950
diff changeset
  2271
  note * = this
3fa22920bf86 integral strong monotone; finite subadditivity for measure
hoelzl
parents: 42950
diff changeset
  2272
  show "integrable M (\<lambda>x. a)"
3fa22920bf86 integral strong monotone; finite subadditivity for measure
hoelzl
parents: 42950
diff changeset
  2273
  proof cases
3fa22920bf86 integral strong monotone; finite subadditivity for measure
hoelzl
parents: 42950
diff changeset
  2274
    assume "0 \<le> a" with * show ?thesis .
3fa22920bf86 integral strong monotone; finite subadditivity for measure
hoelzl
parents: 42950
diff changeset
  2275
  next
3fa22920bf86 integral strong monotone; finite subadditivity for measure
hoelzl
parents: 42950
diff changeset
  2276
    assume "\<not> 0 \<le> a"
3fa22920bf86 integral strong monotone; finite subadditivity for measure
hoelzl
parents: 42950
diff changeset
  2277
    then have "0 \<le> -a" by auto
3fa22920bf86 integral strong monotone; finite subadditivity for measure
hoelzl
parents: 42950
diff changeset
  2278
    from *[OF this] show ?thesis by simp
3fa22920bf86 integral strong monotone; finite subadditivity for measure
hoelzl
parents: 42950
diff changeset
  2279
  qed
3fa22920bf86 integral strong monotone; finite subadditivity for measure
hoelzl
parents: 42950
diff changeset
  2280
  show "(\<integral>x. a \<partial>M) = a * \<mu>' (space M)"
3fa22920bf86 integral strong monotone; finite subadditivity for measure
hoelzl
parents: 42950
diff changeset
  2281
    by (simp add: \<mu>'_def lebesgue_integral_def positive_integral_const_If)
3fa22920bf86 integral strong monotone; finite subadditivity for measure
hoelzl
parents: 42950
diff changeset
  2282
qed
3fa22920bf86 integral strong monotone; finite subadditivity for measure
hoelzl
parents: 42950
diff changeset
  2283
43339
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
  2284
lemma indicator_less[simp]:
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  2285
  "indicator A x \<le> (indicator B x::ereal) \<longleftrightarrow> (x \<in> A \<longrightarrow> x \<in> B)"
43339
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
  2286
  by (simp add: indicator_def not_le)
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
  2287
42991
3fa22920bf86 integral strong monotone; finite subadditivity for measure
hoelzl
parents: 42950
diff changeset
  2288
lemma (in finite_measure) integral_less_AE:
3fa22920bf86 integral strong monotone; finite subadditivity for measure
hoelzl
parents: 42950
diff changeset
  2289
  assumes int: "integrable M X" "integrable M Y"
43339
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
  2290
  assumes A: "\<mu> A \<noteq> 0" "A \<in> sets M" "AE x. x \<in> A \<longrightarrow> X x \<noteq> Y x"
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
  2291
  assumes gt: "AE x. X x \<le> Y x"
42991
3fa22920bf86 integral strong monotone; finite subadditivity for measure
hoelzl
parents: 42950
diff changeset
  2292
  shows "integral\<^isup>L M X < integral\<^isup>L M Y"
3fa22920bf86 integral strong monotone; finite subadditivity for measure
hoelzl
parents: 42950
diff changeset
  2293
proof -
3fa22920bf86 integral strong monotone; finite subadditivity for measure
hoelzl
parents: 42950
diff changeset
  2294
  have "integral\<^isup>L M X \<le> integral\<^isup>L M Y"
3fa22920bf86 integral strong monotone; finite subadditivity for measure
hoelzl
parents: 42950
diff changeset
  2295
    using gt int by (intro integral_mono_AE) auto
3fa22920bf86 integral strong monotone; finite subadditivity for measure
hoelzl
parents: 42950
diff changeset
  2296
  moreover
3fa22920bf86 integral strong monotone; finite subadditivity for measure
hoelzl
parents: 42950
diff changeset
  2297
  have "integral\<^isup>L M X \<noteq> integral\<^isup>L M Y"
3fa22920bf86 integral strong monotone; finite subadditivity for measure
hoelzl
parents: 42950
diff changeset
  2298
  proof
3fa22920bf86 integral strong monotone; finite subadditivity for measure
hoelzl
parents: 42950
diff changeset
  2299
    assume eq: "integral\<^isup>L M X = integral\<^isup>L M Y"
3fa22920bf86 integral strong monotone; finite subadditivity for measure
hoelzl
parents: 42950
diff changeset
  2300
    have "integral\<^isup>L M (\<lambda>x. \<bar>Y x - X x\<bar>) = integral\<^isup>L M (\<lambda>x. Y x - X x)"
3fa22920bf86 integral strong monotone; finite subadditivity for measure
hoelzl
parents: 42950
diff changeset
  2301
      using gt by (intro integral_cong_AE) auto
3fa22920bf86 integral strong monotone; finite subadditivity for measure
hoelzl
parents: 42950
diff changeset
  2302
    also have "\<dots> = 0"
3fa22920bf86 integral strong monotone; finite subadditivity for measure
hoelzl
parents: 42950
diff changeset
  2303
      using eq int by simp
43339
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
  2304
    finally have "\<mu> {x \<in> space M. Y x - X x \<noteq> 0} = 0"
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
  2305
      using int by (simp add: integral_0_iff)
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
  2306
    moreover
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
  2307
    have "(\<integral>\<^isup>+x. indicator A x \<partial>M) \<le> (\<integral>\<^isup>+x. indicator {x \<in> space M. Y x - X x \<noteq> 0} x \<partial>M)"
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
  2308
      using A by (intro positive_integral_mono_AE) auto
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
  2309
    then have "\<mu> A \<le> \<mu> {x \<in> space M. Y x - X x \<noteq> 0}"
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
  2310
      using int A by (simp add: integrable_def)
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
  2311
    moreover note `\<mu> A \<noteq> 0` positive_measure[OF `A \<in> sets M`]
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
  2312
    ultimately show False by auto
42991
3fa22920bf86 integral strong monotone; finite subadditivity for measure
hoelzl
parents: 42950
diff changeset
  2313
  qed
3fa22920bf86 integral strong monotone; finite subadditivity for measure
hoelzl
parents: 42950
diff changeset
  2314
  ultimately show ?thesis by auto
3fa22920bf86 integral strong monotone; finite subadditivity for measure
hoelzl
parents: 42950
diff changeset
  2315
qed
3fa22920bf86 integral strong monotone; finite subadditivity for measure
hoelzl
parents: 42950
diff changeset
  2316
43339
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
  2317
lemma (in finite_measure) integral_less_AE_space:
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
  2318
  assumes int: "integrable M X" "integrable M Y"
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
  2319
  assumes gt: "AE x. X x < Y x" "\<mu> (space M) \<noteq> 0"
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
  2320
  shows "integral\<^isup>L M X < integral\<^isup>L M Y"
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
  2321
  using gt by (intro integral_less_AE[OF int, where A="space M"]) auto
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
  2322
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2323
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
  2324
  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
  2325
  and w: "integrable M w"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2326
  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
  2327
  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
  2328
  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
  2329
  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
  2330
proof -
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2331
  { fix x j assume x: "x \<in> space M"
44568
e6f291cb5810 discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
huffman
parents: 43941
diff changeset
  2332
    from u'[OF x] have "(\<lambda>i. \<bar>u i x\<bar>) ----> \<bar>u' x\<bar>" by (rule tendsto_rabs)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2333
    from LIMSEQ_le_const2[OF this]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2334
    have "\<bar>u' x\<bar> \<le> w x" using bound[OF x] by auto }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2335
  note u'_bound = this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2336
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2337
  from u[unfolded integrable_def]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2338
  have u'_borel: "u' \<in> borel_measurable M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2339
    using u' by (blast intro: borel_measurable_LIMSEQ[of u])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2340
41705
1100512e16d8 add auto support for AE_mp
hoelzl
parents: 41689
diff changeset
  2341
  { fix x assume x: "x \<in> space M"
1100512e16d8 add auto support for AE_mp
hoelzl
parents: 41689
diff changeset
  2342
    then have "0 \<le> \<bar>u 0 x\<bar>" by auto
1100512e16d8 add auto support for AE_mp
hoelzl
parents: 41689
diff changeset
  2343
    also have "\<dots> \<le> w x" using bound[OF x] by auto
1100512e16d8 add auto support for AE_mp
hoelzl
parents: 41689
diff changeset
  2344
    finally have "0 \<le> w x" . }
1100512e16d8 add auto support for AE_mp
hoelzl
parents: 41689
diff changeset
  2345
  note w_pos = this
1100512e16d8 add auto support for AE_mp
hoelzl
parents: 41689
diff changeset
  2346
41689
3e39b0e730d6 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
  show "integrable M u'"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2348
  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
  2349
    show "integrable M w" by fact
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2350
    show "u' \<in> borel_measurable M" by fact
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2351
  next
41705
1100512e16d8 add auto support for AE_mp
hoelzl
parents: 41689
diff changeset
  2352
    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
  2353
    show "\<bar>u' x\<bar> \<le> w x" using u'_bound[OF x] .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2354
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2355
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2356
  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
  2357
  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
  2358
    using w u `integrable M u'`
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2359
    by (auto intro!: integral_add integral_diff integral_cmult integrable_abs)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2360
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2361
  { fix j x assume x: "x \<in> space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2362
    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
  2363
    also have "\<dots> \<le> w x + w x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2364
      by (rule add_mono[OF bound[OF x] u'_bound[OF x]])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2365
    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
  2366
  note diff_less_2w = this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2367
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  2368
  have PI_diff: "\<And>n. (\<integral>\<^isup>+ x. ereal (?diff n x) \<partial>M) =
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  2369
    (\<integral>\<^isup>+ x. ereal (2 * w x) \<partial>M) - (\<integral>\<^isup>+ x. ereal \<bar>u n x - u' x\<bar> \<partial>M)"
41705
1100512e16d8 add auto support for AE_mp
hoelzl
parents: 41689
diff changeset
  2370
    using diff w diff_less_2w w_pos
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2371
    by (subst positive_integral_diff[symmetric])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2372
       (auto simp: integrable_def intro!: positive_integral_cong)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2373
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2374
  have "integrable M (\<lambda>x. 2 * w x)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2375
    using w by (auto intro: integral_cmult)
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  2376
  hence I2w_fin: "(\<integral>\<^isup>+ x. ereal (2 * w x) \<partial>M) \<noteq> \<infinity>" and
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  2377
    borel_2w: "(\<lambda>x. ereal (2 * w x)) \<in> borel_measurable M"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2378
    unfolding integrable_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2379
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  2380
  have "limsup (\<lambda>n. \<integral>\<^isup>+ x. ereal \<bar>u n x - u' x\<bar> \<partial>M) = 0" (is "limsup ?f = 0")
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2381
  proof cases
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  2382
    assume eq_0: "(\<integral>\<^isup>+ x. max 0 (ereal (2 * w x)) \<partial>M) = 0" (is "?wx = 0")
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  2383
    { fix n
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  2384
      have "?f n \<le> ?wx" (is "integral\<^isup>P M ?f' \<le> _")
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  2385
        using diff_less_2w[of _ n] unfolding positive_integral_max_0
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  2386
        by (intro positive_integral_mono) auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  2387
      then have "?f n = 0"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  2388
        using positive_integral_positive[of ?f'] eq_0 by auto }
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  2389
    then show ?thesis by (simp add: Limsup_const)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2390
  next
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  2391
    assume neq_0: "(\<integral>\<^isup>+ x. max 0 (ereal (2 * w x)) \<partial>M) \<noteq> 0" (is "?wx \<noteq> 0")
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  2392
    have "0 = limsup (\<lambda>n. 0 :: ereal)" by (simp add: Limsup_const)
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  2393
    also have "\<dots> \<le> limsup (\<lambda>n. \<integral>\<^isup>+ x. ereal \<bar>u n x - u' x\<bar> \<partial>M)"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  2394
      by (intro limsup_mono positive_integral_positive)
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  2395
    finally have pos: "0 \<le> limsup (\<lambda>n. \<integral>\<^isup>+ x. ereal \<bar>u n x - u' x\<bar> \<partial>M)" .
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  2396
    have "?wx = (\<integral>\<^isup>+ x. liminf (\<lambda>n. max 0 (ereal (?diff n x))) \<partial>M)"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  2397
    proof (rule positive_integral_cong)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2398
      fix x assume x: "x \<in> space M"
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  2399
      show "max 0 (ereal (2 * w x)) = liminf (\<lambda>n. max 0 (ereal (?diff n x)))"
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  2400
        unfolding ereal_max_0
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  2401
      proof (rule lim_imp_Liminf[symmetric], unfold lim_ereal)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2402
        have "(\<lambda>i. ?diff i x) ----> 2 * w x - \<bar>u' x - u' x\<bar>"
44568
e6f291cb5810 discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
huffman
parents: 43941
diff changeset
  2403
          using u'[OF x] by (safe intro!: tendsto_intros)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  2404
        then show "(\<lambda>i. max 0 (?diff i x)) ----> max 0 (2 * w x)"
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  2405
          by (auto intro!: tendsto_real_max simp add: lim_ereal)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  2406
      qed (rule trivial_limit_sequentially)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2407
    qed
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  2408
    also have "\<dots> \<le> liminf (\<lambda>n. \<integral>\<^isup>+ x. max 0 (ereal (?diff n x)) \<partial>M)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2409
      using u'_borel w u unfolding integrable_def
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  2410
      by (intro positive_integral_lim_INF) (auto intro!: positive_integral_lim_INF)
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  2411
    also have "\<dots> = (\<integral>\<^isup>+ x. ereal (2 * w x) \<partial>M) -
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  2412
        limsup (\<lambda>n. \<integral>\<^isup>+ x. ereal \<bar>u n x - u' x\<bar> \<partial>M)"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  2413
      unfolding PI_diff positive_integral_max_0
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  2414
      using positive_integral_positive[of "\<lambda>x. ereal (2 * w x)"]
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  2415
      by (subst liminf_ereal_cminus) auto
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  2416
    finally show ?thesis
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  2417
      using neq_0 I2w_fin positive_integral_positive[of "\<lambda>x. ereal (2 * w x)"] pos
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  2418
      unfolding positive_integral_max_0
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  2419
      by (cases rule: ereal2_cases[of "\<integral>\<^isup>+ x. ereal (2 * w x) \<partial>M" "limsup (\<lambda>n. \<integral>\<^isup>+ x. ereal \<bar>u n x - u' x\<bar> \<partial>M)"])
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  2420
         auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2421
  qed
41705
1100512e16d8 add auto support for AE_mp
hoelzl
parents: 41689
diff changeset
  2422
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  2423
  have "liminf ?f \<le> limsup ?f"
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  2424
    by (intro ereal_Liminf_le_Limsup trivial_limit_sequentially)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  2425
  moreover
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  2426
  { have "0 = liminf (\<lambda>n. 0 :: ereal)" by (simp add: Liminf_const)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  2427
    also have "\<dots> \<le> liminf ?f"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  2428
      by (intro liminf_mono positive_integral_positive)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  2429
    finally have "0 \<le> liminf ?f" . }
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  2430
  ultimately have liminf_limsup_eq: "liminf ?f = ereal 0" "limsup ?f = ereal 0"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  2431
    using `limsup ?f = 0` by auto
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  2432
  have "\<And>n. (\<integral>\<^isup>+ x. ereal \<bar>u n x - u' x\<bar> \<partial>M) = ereal (\<integral>x. \<bar>u n x - u' x\<bar> \<partial>M)"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  2433
    using diff positive_integral_positive
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  2434
    by (subst integral_eq_positive_integral) (auto simp: ereal_real integrable_def)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  2435
  then show ?lim_diff
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  2436
    using ereal_Liminf_eq_Limsup[OF trivial_limit_sequentially liminf_limsup_eq]
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  2437
    by (simp add: lim_ereal)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2438
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2439
  show ?lim
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2440
  proof (rule LIMSEQ_I)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2441
    fix r :: real assume "0 < r"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2442
    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
  2443
    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
  2444
      using diff by (auto simp: integral_positive)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2445
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2446
    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
  2447
    proof (safe intro!: exI[of _ N])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2448
      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
  2449
      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
  2450
        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
  2451
      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
  2452
        by (rule_tac integral_triangle_inequality) (auto intro!: integral_diff)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2453
      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
  2454
      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
  2455
    qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2456
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2457
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2458
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2459
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
  2460
  assumes borel: "\<And>i. integrable M (f i)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2461
  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
  2462
  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
  2463
  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
  2464
  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
  2465
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2466
  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
  2467
    using summable unfolding summable_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2468
  from bchoice[OF this]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2469
  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
  2470
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2471
  let "?w y" = "if y \<in> space M then w y else 0"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2472
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2473
  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
  2474
    using sums unfolding summable_def ..
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2475
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2476
  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
  2477
    using borel by (auto intro!: integral_setsum)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2478
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2479
  { fix j x assume [simp]: "x \<in> space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2480
    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
  2481
    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
  2482
    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
  2483
  note 2 = this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2484
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2485
  have 3: "integrable M ?w"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2486
  proof (rule integral_monotone_convergence(1))
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2487
    let "?F n y" = "(\<Sum>i = 0..<n. \<bar>f i y\<bar>)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2488
    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
  2489
    have "\<And>n. integrable M (?F n)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2490
      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
  2491
    thus "\<And>n. integrable M (?w' n)" by (simp cong: integrable_cong)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2492
    show "mono ?w'"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2493
      by (auto simp: mono_def le_fun_def intro!: setsum_mono2)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2494
    { fix x show "(\<lambda>n. ?w' n x) ----> ?w x"
44568
e6f291cb5810 discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
huffman
parents: 43941
diff changeset
  2495
        using w by (cases "x \<in> space M") (simp_all add: tendsto_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
  2496
    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
  2497
      using borel by (simp add: integral_setsum integrable_abs cong: integral_cong)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2498
    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
  2499
    show "(\<lambda>i. integral\<^isup>L M (?w' i)) ----> x" unfolding * sums_def .
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2500
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2501
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2502
  from summable[THEN summable_rabs_cancel]
41705
1100512e16d8 add auto support for AE_mp
hoelzl
parents: 41689
diff changeset
  2503
  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
  2504
    by (auto intro: summable_sumr_LIMSEQ_suminf)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2505
41705
1100512e16d8 add auto support for AE_mp
hoelzl
parents: 41689
diff changeset
  2506
  note int = integral_dominated_convergence(1,3)[OF 1 2 3 4]
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2507
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2508
  from int show "integrable M ?S" by simp
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2509
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2510
  show ?integral unfolding sums_def integral_setsum(1)[symmetric, OF borel]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2511
    using int(2) by simp
36624
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
  2512
qed
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
  2513
35748
5f35613d9a65 Equality of integral and infinite sum.
hoelzl
parents: 35692
diff changeset
  2514
section "Lebesgue integration on countable spaces"
5f35613d9a65 Equality of integral and infinite sum.
hoelzl
parents: 35692
diff changeset
  2515
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2516
lemma (in measure_space) integral_on_countable:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2517
  assumes f: "f \<in> borel_measurable M"
35748
5f35613d9a65 Equality of integral and infinite sum.
hoelzl
parents: 35692
diff changeset
  2518
  and bij: "bij_betw enum S (f ` space M)"
5f35613d9a65 Equality of integral and infinite sum.
hoelzl
parents: 35692
diff changeset
  2519
  and enum_zero: "enum ` (-S) \<subseteq> {0}"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  2520
  and fin: "\<And>x. x \<noteq> 0 \<Longrightarrow> \<mu> (f -` {x} \<inter> space M) \<noteq> \<infinity>"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2521
  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
  2522
  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
  2523
  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
  2524
proof -
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2525
  let "?A r" = "f -` {enum r} \<inter> space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2526
  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
  2527
  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
  2528
    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
  2529
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2530
  { fix x assume "x \<in> space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2531
    hence "f x \<in> enum ` S" using bij unfolding bij_betw_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2532
    then obtain i where "i\<in>S" "enum i = f x" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2533
    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
  2534
    proof cases
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2535
      fix j assume "j = i"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2536
      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
  2537
    next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2538
      fix j assume "j \<noteq> i"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2539
      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
  2540
        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
  2541
    qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2542
    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
  2543
    have "(\<lambda>i. ?F i x) sums f x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2544
         "(\<lambda>i. \<bar>?F i x\<bar>) sums \<bar>f x\<bar>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2545
      by (auto intro!: sums_single simp: F F_abs) }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2546
  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
  2547
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2548
  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
  2549
    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
  2550
5f35613d9a65 Equality of integral and infinite sum.
hoelzl
parents: 35692
diff changeset
  2551
  { 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
  2552
    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
  2553
      by (auto simp: indicator_def intro!: integral_cong)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2554
    also have "\<dots> = \<bar>enum r\<bar> * real (\<mu> (?A r))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2555
      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
  2556
    finally have "(\<integral>x. \<bar>?F r x\<bar> \<partial>M) = \<bar>enum r * real (\<mu> (?A r))\<bar>"
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  2557
      using f by (subst (2) abs_mult_pos[symmetric]) (auto intro!: real_of_ereal_pos measurable_sets) }
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2558
  note int_abs_F = this
35748
5f35613d9a65 Equality of integral and infinite sum.
hoelzl
parents: 35692
diff changeset
  2559
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2560
  have 1: "\<And>i. integrable M (\<lambda>x. ?F i x)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2561
    using f fin by (simp add: borel_measurable_vimage integral_cmul_indicator)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2562
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2563
  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
  2564
    using F_abs_sums_f unfolding sums_iff by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2565
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2566
  from integral_sums(2)[OF 1 2, unfolded int_abs_F, OF _ abs_summable]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2567
  show ?sums unfolding enum_eq int_f by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2568
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2569
  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
  2570
  show "integrable M f" unfolding int_f by simp
35748
5f35613d9a65 Equality of integral and infinite sum.
hoelzl
parents: 35692
diff changeset
  2571
qed
5f35613d9a65 Equality of integral and infinite sum.
hoelzl
parents: 35692
diff changeset
  2572
35692
f1315bbf1bc9 Moved theorems in Lebesgue to the right places
hoelzl
parents: 35582
diff changeset
  2573
section "Lebesgue integration on finite space"
f1315bbf1bc9 Moved theorems in Lebesgue to the right places
hoelzl
parents: 35582
diff changeset
  2574
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2575
lemma (in measure_space) integral_on_finite:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2576
  assumes f: "f \<in> borel_measurable M" and finite: "finite (f`space M)"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  2577
  and fin: "\<And>x. x \<noteq> 0 \<Longrightarrow> \<mu> (f -` {x} \<inter> space M) \<noteq> \<infinity>"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2578
  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
  2579
  and "(\<integral>x. f x \<partial>M) =
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2580
    (\<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
  2581
proof -
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2582
  let "?A r" = "f -` {r} \<inter> space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2583
  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
  2584
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2585
  { fix x assume "x \<in> space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2586
    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
  2587
      using finite `x \<in> space M` by (simp add: setsum_cases)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2588
    also have "\<dots> = ?S x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2589
      by (auto intro!: setsum_cong)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2590
    finally have "f x = ?S x" . }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2591
  note f_eq = this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2592
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2593
  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
  2594
    by (auto intro!: integrable_cong integral_cong simp only: f_eq)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2595
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2596
  show "integrable M f" ?integral using fin f f_eq_S
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  2597
    by (simp_all add: integral_cmul_indicator borel_measurable_vimage)
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  2598
qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  2599
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  2600
lemma (in finite_measure_space) simple_function_finite[simp, intro]: "simple_function M f"
40875
9a9d33f6fb46 generalized simple_functionD
hoelzl
parents: 40873
diff changeset
  2601
  unfolding simple_function_def using finite_space by auto
35977
30d42bfd0174 Added finite measure space.
hoelzl
parents: 35833
diff changeset
  2602
38705
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 38656
diff changeset
  2603
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
  2604
  by (auto intro: borel_measurable_simple_function)
38705
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 38656
diff changeset
  2605
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 38656
diff changeset
  2606
lemma (in finite_measure_space) positive_integral_finite_eq_setsum:
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  2607
  assumes pos: "\<And>x. x \<in> space M \<Longrightarrow> 0 \<le> f x"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  2608
  shows "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
  2609
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
  2610
  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
  2611
    by (auto intro!: positive_integral_cong simp add: indicator_def if_distrib setsum_cases[OF finite_space])
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  2612
  show ?thesis unfolding * using borel_measurable_finite[of f] pos
40875
9a9d33f6fb46 generalized simple_functionD
hoelzl
parents: 40873
diff changeset
  2613
    by (simp add: positive_integral_setsum positive_integral_cmult_indicator)
38705
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 38656
diff changeset
  2614
qed
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 38656
diff changeset
  2615
35977
30d42bfd0174 Added finite measure space.
hoelzl
parents: 35833
diff changeset
  2616
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
  2617
  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
  2618
  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
  2619
proof -
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  2620
  have *:
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  2621
    "(\<integral>\<^isup>+ x. max 0 (ereal (f x)) \<partial>M) = (\<Sum>x \<in> space M. max 0 (ereal (f x)) * \<mu> {x})"
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  2622
    "(\<integral>\<^isup>+ x. max 0 (ereal (- f x)) \<partial>M) = (\<Sum>x \<in> space M. max 0 (ereal (- f x)) * \<mu> {x})"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  2623
    by (simp_all add: positive_integral_finite_eq_setsum)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  2624
  then show "integrable M f" using finite_space finite_measure
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  2625
    by (simp add: setsum_Pinfty integrable_def positive_integral_max_0
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  2626
             split: split_max)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  2627
  show ?I using finite_measure *
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  2628
    apply (simp add: positive_integral_max_0 lebesgue_integral_def)
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43339
diff changeset
  2629
    apply (subst (1 2) setsum_real_of_ereal[symmetric])
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  2630
    apply (simp_all split: split_max add: setsum_subtractf[symmetric])
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  2631
    apply (intro setsum_cong[OF refl])
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  2632
    apply (simp split: split_max)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  2633
    done
35977
30d42bfd0174 Added finite measure space.
hoelzl
parents: 35833
diff changeset
  2634
qed
30d42bfd0174 Added finite measure space.
hoelzl
parents: 35833
diff changeset
  2635
35748
5f35613d9a65 Equality of integral and infinite sum.
hoelzl
parents: 35692
diff changeset
  2636
end