src/HOL/Probability/Lebesgue_Measure.thy
author hoelzl
Wed, 10 Oct 2012 12:12:35 +0200
changeset 49801 f3471f09bb86
parent 49784 5e5b2da42a69
child 50003 8c213922ed49
permissions -rw-r--r--
add induction for real Borel measurable functions
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
42067
66c8281349ec standardized headers
hoelzl
parents: 41981
diff changeset
     1
(*  Title:      HOL/Probability/Lebesgue_Measure.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:     Robert Himmelmann, TU München
66c8281349ec standardized headers
hoelzl
parents: 41981
diff changeset
     4
*)
66c8281349ec standardized headers
hoelzl
parents: 41981
diff changeset
     5
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
     6
header {* Lebsegue measure *}
42067
66c8281349ec standardized headers
hoelzl
parents: 41981
diff changeset
     7
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
     8
theory Lebesgue_Measure
42146
5b52c6a9c627 split Product_Measure into Binary_Product_Measure and Finite_Product_Measure
hoelzl
parents: 42067
diff changeset
     9
  imports Finite_Product_Measure
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    10
begin
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    11
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
    12
lemma borel_measurable_indicator':
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
    13
  "A \<in> sets borel \<Longrightarrow> f \<in> borel_measurable M \<Longrightarrow> (\<lambda>x. indicator A (f x)) \<in> borel_measurable M"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
    14
  using measurable_comp[OF _ borel_measurable_indicator, of f M borel A] by (auto simp add: comp_def)
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
    15
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
    16
lemma borel_measurable_sets:
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
    17
  assumes "f \<in> measurable borel M" "A \<in> sets M"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
    18
  shows "f -` A \<in> sets borel"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
    19
  using measurable_sets[OF assms] by simp
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
    20
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    21
subsection {* Standard Cubes *}
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    22
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
    23
definition cube :: "nat \<Rightarrow> 'a::ordered_euclidean_space set" where
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
    24
  "cube n \<equiv> {\<chi>\<chi> i. - real n .. \<chi>\<chi> i. real n}"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
    25
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
    26
lemma borel_cube[intro]: "cube n \<in> sets borel"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
    27
  unfolding cube_def by auto
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
    28
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
    29
lemma cube_closed[intro]: "closed (cube n)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
    30
  unfolding cube_def by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
    31
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
    32
lemma cube_subset[intro]: "n \<le> N \<Longrightarrow> cube n \<subseteq> cube N"
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 44666
diff changeset
    33
  by (fastforce simp: eucl_le[where 'a='a] cube_def)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    34
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
    35
lemma cube_subset_iff:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
    36
  "cube n \<subseteq> cube N \<longleftrightarrow> n \<le> N"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
    37
proof
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
    38
  assume subset: "cube n \<subseteq> (cube N::'a set)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
    39
  then have "((\<chi>\<chi> i. real n)::'a) \<in> cube N"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
    40
    using DIM_positive[where 'a='a]
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 44666
diff changeset
    41
    by (fastforce simp: cube_def eucl_le[where 'a='a])
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
    42
  then show "n \<le> N"
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 44666
diff changeset
    43
    by (fastforce simp: cube_def eucl_le[where 'a='a])
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
    44
next
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
    45
  assume "n \<le> N" then show "cube n \<subseteq> (cube N::'a set)" by (rule cube_subset)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
    46
qed
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    47
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    48
lemma ball_subset_cube:"ball (0::'a::ordered_euclidean_space) (real n) \<subseteq> cube n"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    49
  unfolding cube_def subset_eq mem_interval apply safe unfolding euclidean_lambda_beta'
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    50
proof- fix x::'a and i assume as:"x\<in>ball 0 (real n)" "i<DIM('a)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    51
  thus "- real n \<le> x $$ i" "real n \<ge> x $$ i"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    52
    using component_le_norm[of x i] by(auto simp: dist_norm)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    53
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    54
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    55
lemma mem_big_cube: obtains n where "x \<in> cube n"
44666
8670a39d4420 remove more duplicate lemmas
huffman
parents: 43920
diff changeset
    56
proof- from reals_Archimedean2[of "norm x"] guess n ..
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    57
  thus ?thesis apply-apply(rule that[where n=n])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    58
    apply(rule ball_subset_cube[unfolded subset_eq,rule_format])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    59
    by (auto simp add:dist_norm)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    60
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    61
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
    62
lemma cube_subset_Suc[intro]: "cube n \<subseteq> cube (Suc n)"
46905
6b1c0a80a57a prefer abs_def over def_raw;
wenzelm
parents: 46731
diff changeset
    63
  unfolding cube_def subset_eq apply safe unfolding mem_interval apply auto done
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
    64
47757
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
    65
subsection {* Lebesgue measure *}
41706
a207a858d1f6 prefer p2e before e2p; use measure_unique_Int_stable_vimage;
hoelzl
parents: 41704
diff changeset
    66
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
    67
definition lebesgue :: "'a::ordered_euclidean_space measure" where
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
    68
  "lebesgue = measure_of UNIV {A. \<forall>n. (indicator A :: 'a \<Rightarrow> real) integrable_on cube n}
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
    69
    (\<lambda>A. SUP n. ereal (integral (cube n) (indicator A)))"
41661
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41654
diff changeset
    70
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
    71
lemma space_lebesgue[simp]: "space lebesgue = UNIV"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
    72
  unfolding lebesgue_def by simp
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
    73
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
    74
lemma lebesgueI: "(\<And>n. (indicator A :: _ \<Rightarrow> real) integrable_on cube n) \<Longrightarrow> A \<in> sets lebesgue"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
    75
  unfolding lebesgue_def by simp
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
    76
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
    77
lemma absolutely_integrable_on_indicator[simp]:
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
    78
  fixes A :: "'a::ordered_euclidean_space set"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
    79
  shows "((indicator A :: _ \<Rightarrow> real) absolutely_integrable_on X) \<longleftrightarrow>
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
    80
    (indicator A :: _ \<Rightarrow> real) integrable_on X"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
    81
  unfolding absolutely_integrable_on_def by simp
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
    82
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
    83
lemma LIMSEQ_indicator_UN:
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
    84
  "(\<lambda>k. indicator (\<Union> i<k. A i) x) ----> (indicator (\<Union>i. A i) x :: real)"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
    85
proof cases
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
    86
  assume "\<exists>i. x \<in> A i" then guess i .. note i = this
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
    87
  then have *: "\<And>n. (indicator (\<Union> i<n + Suc i. A i) x :: real) = 1"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
    88
    "(indicator (\<Union> i. A i) x :: real) = 1" by (auto simp: indicator_def)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
    89
  show ?thesis
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
    90
    apply (rule LIMSEQ_offset[of _ "Suc i"]) unfolding * by auto
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
    91
qed (auto simp: indicator_def)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    92
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
    93
lemma indicator_add:
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
    94
  "A \<inter> B = {} \<Longrightarrow> (indicator A x::_::monoid_add) + indicator B x = indicator (A \<union> B) x"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
    95
  unfolding indicator_def by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    96
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
    97
lemma sigma_algebra_lebesgue:
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
    98
  defines "leb \<equiv> {A. \<forall>n. (indicator A :: 'a::ordered_euclidean_space \<Rightarrow> real) integrable_on cube n}"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
    99
  shows "sigma_algebra UNIV leb"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   100
proof (safe intro!: sigma_algebra_iff2[THEN iffD2])
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   101
  fix A assume A: "A \<in> leb"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   102
  moreover have "indicator (UNIV - A) = (\<lambda>x. 1 - indicator A x :: real)"
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   103
    by (auto simp: fun_eq_iff indicator_def)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   104
  ultimately show "UNIV - A \<in> leb"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   105
    using A by (auto intro!: integrable_sub simp: cube_def leb_def)
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   106
next
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   107
  fix n show "{} \<in> leb"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   108
    by (auto simp: cube_def indicator_def[abs_def] leb_def)
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   109
next
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   110
  fix A :: "nat \<Rightarrow> _" assume A: "range A \<subseteq> leb"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   111
  have "\<forall>n. (indicator (\<Union>i. A i) :: _ \<Rightarrow> real) integrable_on cube n" (is "\<forall>n. ?g integrable_on _")
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   112
  proof (intro dominated_convergence[where g="?g"] ballI allI)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   113
    fix k n show "(indicator (\<Union>i<k. A i) :: _ \<Rightarrow> real) integrable_on cube n"
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   114
    proof (induct k)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   115
      case (Suc k)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   116
      have *: "(\<Union> i<Suc k. A i) = (\<Union> i<k. A i) \<union> A k"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   117
        unfolding lessThan_Suc UN_insert by auto
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   118
      have *: "(\<lambda>x. max (indicator (\<Union> i<k. A i) x) (indicator (A k) x) :: real) =
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   119
          indicator (\<Union> i<Suc k. A i)" (is "(\<lambda>x. max (?f x) (?g x)) = _")
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   120
        by (auto simp: fun_eq_iff * indicator_def)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   121
      show ?case
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   122
        using absolutely_integrable_max[of ?f "cube n" ?g] A Suc
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   123
        by (simp add: * leb_def subset_eq)
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   124
    qed auto
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   125
  qed (auto intro: LIMSEQ_indicator_UN simp: cube_def)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   126
  then show "(\<Union>i. A i) \<in> leb" by (auto simp: leb_def)
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   127
qed simp
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   128
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   129
lemma sets_lebesgue: "sets lebesgue = {A. \<forall>n. (indicator A :: _ \<Rightarrow> real) integrable_on cube n}"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   130
  unfolding lebesgue_def sigma_algebra.sets_measure_of_eq[OF sigma_algebra_lebesgue] ..
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   131
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   132
lemma lebesgueD: "A \<in> sets lebesgue \<Longrightarrow> (indicator A :: _ \<Rightarrow> real) integrable_on cube n"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   133
  unfolding sets_lebesgue by simp
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   134
47757
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   135
lemma emeasure_lebesgue:
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   136
  assumes "A \<in> sets lebesgue"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   137
  shows "emeasure lebesgue A = (SUP n. ereal (integral (cube n) (indicator A)))"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   138
    (is "_ = ?\<mu> A")
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   139
proof (rule emeasure_measure_of[OF lebesgue_def])
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   140
  have *: "indicator {} = (\<lambda>x. 0 :: real)" by (simp add: fun_eq_iff)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   141
  show "positive (sets lebesgue) ?\<mu>"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   142
  proof (unfold positive_def, intro conjI ballI)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   143
    show "?\<mu> {} = 0" by (simp add: integral_0 *)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   144
    fix A :: "'a set" assume "A \<in> sets lebesgue" then show "0 \<le> ?\<mu> A"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   145
      by (auto intro!: SUP_upper2 Integration.integral_nonneg simp: sets_lebesgue)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   146
  qed
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   147
next
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   148
  show "countably_additive (sets lebesgue) ?\<mu>"
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   149
  proof (intro countably_additive_def[THEN iffD2] allI impI)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   150
    fix A :: "nat \<Rightarrow> 'a set" assume rA: "range A \<subseteq> sets lebesgue" "disjoint_family A"
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   151
    then have A[simp, intro]: "\<And>i n. (indicator (A i) :: _ \<Rightarrow> real) integrable_on cube n"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   152
      by (auto dest: lebesgueD)
46731
5302e932d1e5 avoid undeclared variables in let bindings;
wenzelm
parents: 44928
diff changeset
   153
    let ?m = "\<lambda>n i. integral (cube n) (indicator (A i) :: _\<Rightarrow>real)"
5302e932d1e5 avoid undeclared variables in let bindings;
wenzelm
parents: 44928
diff changeset
   154
    let ?M = "\<lambda>n I. integral (cube n) (indicator (\<Union>i\<in>I. A i) :: _\<Rightarrow>real)"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   155
    have nn[simp, intro]: "\<And>i n. 0 \<le> ?m n i" by (auto intro!: Integration.integral_nonneg)
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   156
    assume "(\<Union>i. A i) \<in> sets lebesgue"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   157
    then have UN_A[simp, intro]: "\<And>i n. (indicator (\<Union>i. A i) :: _ \<Rightarrow> real) integrable_on cube n"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   158
      by (auto simp: sets_lebesgue)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   159
    show "(\<Sum>n. ?\<mu> (A n)) = ?\<mu> (\<Union>i. A i)"
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   160
    proof (subst suminf_SUP_eq, safe intro!: incseq_SucI) 
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
   161
      fix i n show "ereal (?m n i) \<le> ereal (?m (Suc n) i)"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   162
        using cube_subset[of n "Suc n"] by (auto intro!: integral_subset_le incseq_SucI)
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   163
    next
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
   164
      fix i n show "0 \<le> ereal (?m n i)"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   165
        using rA unfolding lebesgue_def
44928
7ef6505bde7f renamed Complete_Lattices lemmas, removed legacy names
hoelzl
parents: 44890
diff changeset
   166
        by (auto intro!: SUP_upper2 integral_nonneg)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   167
    next
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
   168
      show "(SUP n. \<Sum>i. ereal (?m n i)) = (SUP n. ereal (?M n UNIV))"
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
   169
      proof (intro arg_cong[where f="SUPR UNIV"] ext sums_unique[symmetric] sums_ereal[THEN iffD2] sums_def[THEN iffD2])
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   170
        fix n
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   171
        have "\<And>m. (UNION {..<m} A) \<in> sets lebesgue" using rA by auto
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   172
        from lebesgueD[OF this]
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   173
        have "(\<lambda>m. ?M n {..< m}) ----> ?M n UNIV"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   174
          (is "(\<lambda>m. integral _ (?A m)) ----> ?I")
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   175
          by (intro dominated_convergence(2)[where f="?A" and h="\<lambda>x. 1::real"])
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   176
             (auto intro: LIMSEQ_indicator_UN simp: cube_def)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   177
        moreover
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   178
        { fix m have *: "(\<Sum>x<m. ?m n x) = ?M n {..< m}"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   179
          proof (induct m)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   180
            case (Suc m)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   181
            have "(\<Union>i<m. A i) \<in> sets lebesgue" using rA by auto
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   182
            then have "(indicator (\<Union>i<m. A i) :: _\<Rightarrow>real) integrable_on (cube n)"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   183
              by (auto dest!: lebesgueD)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   184
            moreover
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   185
            have "(\<Union>i<m. A i) \<inter> A m = {}"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   186
              using rA(2)[unfolded disjoint_family_on_def, THEN bspec, of m]
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   187
              by auto
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   188
            then have "\<And>x. indicator (\<Union>i<Suc m. A i) x =
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   189
              indicator (\<Union>i<m. A i) x + (indicator (A m) x :: real)"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   190
              by (auto simp: indicator_add lessThan_Suc ac_simps)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   191
            ultimately show ?case
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   192
              using Suc A by (simp add: Integration.integral_add[symmetric])
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   193
          qed auto }
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   194
        ultimately show "(\<lambda>m. \<Sum>x = 0..<m. ?m n x) ----> ?M n UNIV"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   195
          by (simp add: atLeast0LessThan)
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   196
      qed
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   197
    qed
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   198
  qed
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   199
qed (auto, fact)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   200
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   201
lemma has_integral_interval_cube:
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   202
  fixes a b :: "'a::ordered_euclidean_space"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   203
  shows "(indicator {a .. b} has_integral
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   204
    content ({\<chi>\<chi> i. max (- real n) (a $$ i) .. \<chi>\<chi> i. min (real n) (b $$ i)} :: 'a set)) (cube n)"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   205
    (is "(?I has_integral content ?R) (cube n)")
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   206
proof -
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   207
  let "{?N .. ?P}" = ?R
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   208
  have [simp]: "(\<lambda>x. if x \<in> cube n then ?I x else 0) = indicator ?R"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   209
    by (auto simp: indicator_def cube_def fun_eq_iff eucl_le[where 'a='a])
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   210
  have "(?I has_integral content ?R) (cube n) \<longleftrightarrow> (indicator ?R has_integral content ?R) UNIV"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   211
    unfolding has_integral_restrict_univ[where s="cube n", symmetric] by simp
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   212
  also have "\<dots> \<longleftrightarrow> ((\<lambda>x. 1) has_integral content ?R) ?R"
46905
6b1c0a80a57a prefer abs_def over def_raw;
wenzelm
parents: 46731
diff changeset
   213
    unfolding indicator_def [abs_def] has_integral_restrict_univ ..
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   214
  finally show ?thesis
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   215
    using has_integral_const[of "1::real" "?N" "?P"] by simp
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   216
qed
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   217
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   218
lemma lebesgueI_borel[intro, simp]:
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   219
  fixes s::"'a::ordered_euclidean_space set"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   220
  assumes "s \<in> sets borel" shows "s \<in> sets lebesgue"
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   221
proof -
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   222
  have "s \<in> sigma_sets (space lebesgue) (range (\<lambda>(a, b). {a .. (b :: 'a\<Colon>ordered_euclidean_space)}))"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   223
    using assms by (simp add: borel_eq_atLeastAtMost)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   224
  also have "\<dots> \<subseteq> sets lebesgue"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   225
  proof (safe intro!: sigma_sets_subset lebesgueI)
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   226
    fix n :: nat and a b :: 'a
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   227
    let ?N = "\<chi>\<chi> i. max (- real n) (a $$ i)"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   228
    let ?P = "\<chi>\<chi> i. min (real n) (b $$ i)"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   229
    show "(indicator {a..b} :: 'a\<Rightarrow>real) integrable_on cube n"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   230
      unfolding integrable_on_def
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   231
      using has_integral_interval_cube[of a b] by auto
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   232
  qed
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   233
  finally show ?thesis .
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   234
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   235
47757
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   236
lemma borel_measurable_lebesgueI:
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   237
  "f \<in> borel_measurable borel \<Longrightarrow> f \<in> borel_measurable lebesgue"
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   238
  unfolding measurable_def by simp
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   239
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   240
lemma lebesgueI_negligible[dest]: fixes s::"'a::ordered_euclidean_space set"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   241
  assumes "negligible s" shows "s \<in> sets lebesgue"
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   242
  using assms by (force simp: cube_def integrable_on_def negligible_def intro!: lebesgueI)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   243
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   244
lemma lmeasure_eq_0:
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   245
  fixes S :: "'a::ordered_euclidean_space set"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   246
  assumes "negligible S" shows "emeasure lebesgue S = 0"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   247
proof -
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   248
  have "\<And>n. integral (cube n) (indicator S :: 'a\<Rightarrow>real) = 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
   249
    unfolding lebesgue_integral_def using 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
   250
    by (intro integral_unique some1_equality ex_ex1I)
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   251
       (auto simp: cube_def negligible_def)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   252
  then show ?thesis
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   253
    using assms by (simp add: emeasure_lebesgue lebesgueI_negligible)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   254
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   255
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   256
lemma lmeasure_iff_LIMSEQ:
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   257
  assumes A: "A \<in> sets lebesgue" and "0 \<le> m"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   258
  shows "emeasure lebesgue A = ereal m \<longleftrightarrow> (\<lambda>n. integral (cube n) (indicator A :: _ \<Rightarrow> real)) ----> m"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   259
proof (subst emeasure_lebesgue[OF A], intro SUP_eq_LIMSEQ)
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   260
  show "mono (\<lambda>n. integral (cube n) (indicator A::_=>real))"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   261
    using cube_subset assms by (intro monoI integral_subset_le) (auto dest!: lebesgueD)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   262
qed
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   263
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   264
lemma has_integral_indicator_UNIV:
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   265
  fixes s A :: "'a::ordered_euclidean_space set" and x :: real
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   266
  shows "((indicator (s \<inter> A) :: 'a\<Rightarrow>real) has_integral x) UNIV = ((indicator s :: _\<Rightarrow>real) has_integral x) A"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   267
proof -
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   268
  have "(\<lambda>x. if x \<in> A then indicator s x else 0) = (indicator (s \<inter> A) :: _\<Rightarrow>real)"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   269
    by (auto simp: fun_eq_iff indicator_def)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   270
  then show ?thesis
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   271
    unfolding has_integral_restrict_univ[where s=A, symmetric] by simp
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   272
qed
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   273
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   274
lemma
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   275
  fixes s a :: "'a::ordered_euclidean_space set"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   276
  shows integral_indicator_UNIV:
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   277
    "integral UNIV (indicator (s \<inter> A) :: 'a\<Rightarrow>real) = integral A (indicator s :: _\<Rightarrow>real)"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   278
  and integrable_indicator_UNIV:
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   279
    "(indicator (s \<inter> A) :: 'a\<Rightarrow>real) integrable_on UNIV \<longleftrightarrow> (indicator s :: 'a\<Rightarrow>real) integrable_on A"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   280
  unfolding integral_def integrable_on_def has_integral_indicator_UNIV by auto
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   281
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   282
lemma lmeasure_finite_has_integral:
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   283
  fixes s :: "'a::ordered_euclidean_space set"
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   284
  assumes "s \<in> sets lebesgue" "emeasure lebesgue s = ereal m"
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   285
  shows "(indicator s has_integral m) UNIV"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   286
proof -
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   287
  let ?I = "indicator :: 'a set \<Rightarrow> 'a \<Rightarrow> real"
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   288
  have "0 \<le> m"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   289
    using emeasure_nonneg[of lebesgue s] `emeasure lebesgue s = ereal m` by simp
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   290
  have **: "(?I s) integrable_on UNIV \<and> (\<lambda>k. integral UNIV (?I (s \<inter> cube k))) ----> integral UNIV (?I s)"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   291
  proof (intro monotone_convergence_increasing allI ballI)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   292
    have LIMSEQ: "(\<lambda>n. integral (cube n) (?I s)) ----> m"
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   293
      using assms(2) unfolding lmeasure_iff_LIMSEQ[OF assms(1) `0 \<le> m`] .
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   294
    { fix n have "integral (cube n) (?I s) \<le> m"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   295
        using cube_subset assms
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   296
        by (intro incseq_le[where L=m] LIMSEQ incseq_def[THEN iffD2] integral_subset_le allI impI)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   297
           (auto dest!: lebesgueD) }
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   298
    moreover
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   299
    { fix n have "0 \<le> integral (cube n) (?I s)"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   300
      using assms by (auto dest!: lebesgueD intro!: Integration.integral_nonneg) }
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   301
    ultimately
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   302
    show "bounded {integral UNIV (?I (s \<inter> cube k)) |k. True}"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   303
      unfolding bounded_def
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   304
      apply (rule_tac exI[of _ 0])
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   305
      apply (rule_tac exI[of _ m])
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   306
      by (auto simp: dist_real_def integral_indicator_UNIV)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   307
    fix k show "?I (s \<inter> cube k) integrable_on UNIV"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   308
      unfolding integrable_indicator_UNIV using assms by (auto dest!: lebesgueD)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   309
    fix x show "?I (s \<inter> cube k) x \<le> ?I (s \<inter> cube (Suc k)) x"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   310
      using cube_subset[of k "Suc k"] by (auto simp: indicator_def)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   311
  next
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   312
    fix x :: 'a
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   313
    from mem_big_cube obtain k where k: "x \<in> cube k" .
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   314
    { fix n have "?I (s \<inter> cube (n + k)) x = ?I s x"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   315
      using k cube_subset[of k "n + k"] by (auto simp: indicator_def) }
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   316
    note * = this
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   317
    show "(\<lambda>k. ?I (s \<inter> cube k) x) ----> ?I s x"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   318
      by (rule LIMSEQ_offset[where k=k]) (auto simp: *)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   319
  qed
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   320
  note ** = conjunctD2[OF this]
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   321
  have m: "m = integral UNIV (?I s)"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   322
    apply (intro LIMSEQ_unique[OF _ **(2)])
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   323
    using assms(2) unfolding lmeasure_iff_LIMSEQ[OF assms(1) `0 \<le> m`] integral_indicator_UNIV .
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   324
  show ?thesis
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   325
    unfolding m by (intro integrable_integral **)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   326
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   327
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   328
lemma lmeasure_finite_integrable: assumes s: "s \<in> sets lebesgue" and "emeasure lebesgue s \<noteq> \<infinity>"
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   329
  shows "(indicator s :: _ \<Rightarrow> real) integrable_on UNIV"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   330
proof (cases "emeasure lebesgue s")
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   331
  case (real m)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   332
  with lmeasure_finite_has_integral[OF `s \<in> sets lebesgue` this] emeasure_nonneg[of lebesgue s]
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   333
  show ?thesis unfolding integrable_on_def by auto
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   334
qed (insert assms emeasure_nonneg[of lebesgue s], auto)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   335
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   336
lemma has_integral_lebesgue: assumes "((indicator s :: _\<Rightarrow>real) has_integral m) UNIV"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   337
  shows "s \<in> sets lebesgue"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   338
proof (intro lebesgueI)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   339
  let ?I = "indicator :: 'a set \<Rightarrow> 'a \<Rightarrow> real"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   340
  fix n show "(?I s) integrable_on cube n" unfolding cube_def
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   341
  proof (intro integrable_on_subinterval)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   342
    show "(?I s) integrable_on UNIV"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   343
      unfolding integrable_on_def using assms by auto
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   344
  qed auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   345
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   346
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   347
lemma has_integral_lmeasure: assumes "((indicator s :: _\<Rightarrow>real) has_integral m) UNIV"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   348
  shows "emeasure lebesgue s = ereal m"
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   349
proof (intro lmeasure_iff_LIMSEQ[THEN iffD2])
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   350
  let ?I = "indicator :: 'a set \<Rightarrow> 'a \<Rightarrow> real"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   351
  show "s \<in> sets lebesgue" using has_integral_lebesgue[OF assms] .
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   352
  show "0 \<le> m" using assms by (rule has_integral_nonneg) auto
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   353
  have "(\<lambda>n. integral UNIV (?I (s \<inter> cube n))) ----> integral UNIV (?I s)"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   354
  proof (intro dominated_convergence(2) ballI)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   355
    show "(?I s) integrable_on UNIV" unfolding integrable_on_def using assms by auto
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   356
    fix n show "?I (s \<inter> cube n) integrable_on UNIV"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   357
      unfolding integrable_indicator_UNIV using `s \<in> sets lebesgue` by (auto dest: lebesgueD)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   358
    fix x show "norm (?I (s \<inter> cube n) x) \<le> ?I s x" by (auto simp: indicator_def)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   359
  next
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   360
    fix x :: 'a
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   361
    from mem_big_cube obtain k where k: "x \<in> cube k" .
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   362
    { fix n have "?I (s \<inter> cube (n + k)) x = ?I s x"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   363
      using k cube_subset[of k "n + k"] by (auto simp: indicator_def) }
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   364
    note * = this
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   365
    show "(\<lambda>k. ?I (s \<inter> cube k) x) ----> ?I s x"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   366
      by (rule LIMSEQ_offset[where k=k]) (auto simp: *)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   367
  qed
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   368
  then show "(\<lambda>n. integral (cube n) (?I s)) ----> m"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   369
    unfolding integral_unique[OF assms] integral_indicator_UNIV by simp
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   370
qed
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   371
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   372
lemma has_integral_iff_lmeasure:
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   373
  "(indicator A has_integral m) UNIV \<longleftrightarrow> (A \<in> sets lebesgue \<and> emeasure lebesgue A = ereal m)"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   374
proof
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   375
  assume "(indicator A has_integral m) UNIV"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   376
  with has_integral_lmeasure[OF this] has_integral_lebesgue[OF this]
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   377
  show "A \<in> sets lebesgue \<and> emeasure lebesgue A = ereal m"
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   378
    by (auto intro: has_integral_nonneg)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   379
next
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   380
  assume "A \<in> sets lebesgue \<and> emeasure lebesgue A = ereal m"
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   381
  then show "(indicator A has_integral m) UNIV" by (intro lmeasure_finite_has_integral) auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   382
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   383
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   384
lemma lmeasure_eq_integral: assumes "(indicator s::_\<Rightarrow>real) integrable_on UNIV"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   385
  shows "emeasure lebesgue s = ereal (integral UNIV (indicator s))"
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   386
  using assms unfolding integrable_on_def
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   387
proof safe
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   388
  fix y :: real assume "(indicator s has_integral y) UNIV"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   389
  from this[unfolded has_integral_iff_lmeasure] integral_unique[OF this]
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   390
  show "emeasure lebesgue s = ereal (integral UNIV (indicator s))" by simp
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   391
qed
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   392
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   393
lemma lebesgue_simple_function_indicator:
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
   394
  fixes f::"'a::ordered_euclidean_space \<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
   395
  assumes f:"simple_function lebesgue f"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   396
  shows "f = (\<lambda>x. (\<Sum>y \<in> f ` UNIV. y * indicator (f -` {y}) x))"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   397
  by (rule, subst simple_function_indicator_representation[OF f]) auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   398
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   399
lemma integral_eq_lmeasure:
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   400
  "(indicator s::_\<Rightarrow>real) integrable_on UNIV \<Longrightarrow> integral UNIV (indicator s) = real (emeasure lebesgue s)"
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   401
  by (subst lmeasure_eq_integral) (auto intro!: integral_nonneg)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   402
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   403
lemma lmeasure_finite: assumes "(indicator s::_\<Rightarrow>real) integrable_on UNIV" shows "emeasure lebesgue s \<noteq> \<infinity>"
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   404
  using lmeasure_eq_integral[OF assms] by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   405
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   406
lemma negligible_iff_lebesgue_null_sets:
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   407
  "negligible A \<longleftrightarrow> A \<in> null_sets lebesgue"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   408
proof
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   409
  assume "negligible A"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   410
  from this[THEN lebesgueI_negligible] this[THEN lmeasure_eq_0]
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   411
  show "A \<in> null_sets lebesgue" by auto
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   412
next
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   413
  assume A: "A \<in> null_sets lebesgue"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   414
  then have *:"((indicator A) has_integral (0::real)) UNIV" using lmeasure_finite_has_integral[of A]
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   415
    by (auto simp: null_sets_def)
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   416
  show "negligible A" unfolding negligible_def
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   417
  proof (intro allI)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   418
    fix a b :: 'a
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   419
    have integrable: "(indicator A :: _\<Rightarrow>real) integrable_on {a..b}"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   420
      by (intro integrable_on_subinterval has_integral_integrable) (auto intro: *)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   421
    then have "integral {a..b} (indicator A) \<le> (integral UNIV (indicator A) :: real)"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   422
      using * by (auto intro!: integral_subset_le)
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   423
    moreover have "(0::real) \<le> integral {a..b} (indicator A)"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   424
      using integrable by (auto intro!: integral_nonneg)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   425
    ultimately have "integral {a..b} (indicator A) = (0::real)"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   426
      using integral_unique[OF *] by auto
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   427
    then show "(indicator A has_integral (0::real)) {a..b}"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   428
      using integrable_integral[OF integrable] by simp
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   429
  qed
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   430
qed
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   431
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   432
lemma integral_const[simp]:
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   433
  fixes a b :: "'a::ordered_euclidean_space"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   434
  shows "integral {a .. b} (\<lambda>x. c) = content {a .. b} *\<^sub>R c"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   435
  by (rule integral_unique) (rule has_integral_const)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   436
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   437
lemma lmeasure_UNIV[intro]: "emeasure lebesgue (UNIV::'a::ordered_euclidean_space set) = \<infinity>"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   438
proof (simp add: emeasure_lebesgue, intro SUP_PInfty bexI)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   439
  fix n :: nat
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   440
  have "indicator UNIV = (\<lambda>x::'a. 1 :: real)" by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   441
  moreover
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   442
  { have "real n \<le> (2 * real n) ^ DIM('a)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   443
    proof (cases n)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   444
      case 0 then show ?thesis by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   445
    next
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   446
      case (Suc n')
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   447
      have "real n \<le> (2 * real n)^1" by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   448
      also have "(2 * real n)^1 \<le> (2 * real n) ^ DIM('a)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   449
        using Suc DIM_positive[where 'a='a] by (intro power_increasing) (auto simp: real_of_nat_Suc)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   450
      finally show ?thesis .
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   451
    qed }
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
   452
  ultimately show "ereal (real n) \<le> ereal (integral (cube n) (indicator UNIV::'a\<Rightarrow>real))"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   453
    using integral_const DIM_positive[where 'a='a]
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   454
    by (auto simp: cube_def content_closed_interval_cases setprod_constant)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   455
qed simp
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   456
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   457
lemma lmeasure_complete: "A \<subseteq> B \<Longrightarrow> B \<in> null_sets lebesgue \<Longrightarrow> A \<in> null_sets lebesgue"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   458
  unfolding negligible_iff_lebesgue_null_sets[symmetric] by (auto simp: negligible_subset)
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   459
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   460
lemma
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   461
  fixes a b ::"'a::ordered_euclidean_space"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   462
  shows lmeasure_atLeastAtMost[simp]: "emeasure lebesgue {a..b} = ereal (content {a..b})"
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   463
proof -
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   464
  have "(indicator (UNIV \<inter> {a..b})::_\<Rightarrow>real) integrable_on UNIV"
46905
6b1c0a80a57a prefer abs_def over def_raw;
wenzelm
parents: 46731
diff changeset
   465
    unfolding integrable_indicator_UNIV by (simp add: integrable_const indicator_def [abs_def])
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   466
  from lmeasure_eq_integral[OF this] show ?thesis unfolding integral_indicator_UNIV
46905
6b1c0a80a57a prefer abs_def over def_raw;
wenzelm
parents: 46731
diff changeset
   467
    by (simp add: indicator_def [abs_def])
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   468
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   469
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   470
lemma atLeastAtMost_singleton_euclidean[simp]:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   471
  fixes a :: "'a::ordered_euclidean_space" shows "{a .. a} = {a}"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   472
  by (force simp: eucl_le[where 'a='a] euclidean_eq[where 'a='a])
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   473
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   474
lemma content_singleton[simp]: "content {a} = 0"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   475
proof -
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   476
  have "content {a .. a} = 0"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   477
    by (subst content_closed_interval) auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   478
  then show ?thesis by simp
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   479
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   480
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   481
lemma lmeasure_singleton[simp]:
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   482
  fixes a :: "'a::ordered_euclidean_space" shows "emeasure lebesgue {a} = 0"
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   483
  using lmeasure_atLeastAtMost[of a a] by simp
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   484
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   485
lemma AE_lebesgue_singleton:
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   486
  fixes a :: "'a::ordered_euclidean_space" shows "AE x in lebesgue. x \<noteq> a"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   487
  by (rule AE_I[where N="{a}"]) auto
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   488
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   489
declare content_real[simp]
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   490
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   491
lemma
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   492
  fixes a b :: real
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   493
  shows lmeasure_real_greaterThanAtMost[simp]:
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   494
    "emeasure lebesgue {a <.. b} = ereal (if a \<le> b then b - a else 0)"
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   495
proof -
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   496
  have "emeasure lebesgue {a <.. b} = emeasure lebesgue {a .. b}"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   497
    using AE_lebesgue_singleton[of a]
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   498
    by (intro emeasure_eq_AE) auto
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   499
  then show ?thesis by auto
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   500
qed
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   501
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   502
lemma
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   503
  fixes a b :: real
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   504
  shows lmeasure_real_atLeastLessThan[simp]:
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   505
    "emeasure lebesgue {a ..< b} = ereal (if a \<le> b then b - a else 0)"
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   506
proof -
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   507
  have "emeasure lebesgue {a ..< b} = emeasure lebesgue {a .. b}"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   508
    using AE_lebesgue_singleton[of b]
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   509
    by (intro emeasure_eq_AE) auto
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   510
  then show ?thesis by auto
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   511
qed
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   512
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   513
lemma
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   514
  fixes a b :: real
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   515
  shows lmeasure_real_greaterThanLessThan[simp]:
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   516
    "emeasure lebesgue {a <..< b} = ereal (if a \<le> b then b - a else 0)"
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   517
proof -
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   518
  have "emeasure lebesgue {a <..< b} = emeasure lebesgue {a .. b}"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   519
    using AE_lebesgue_singleton[of a] AE_lebesgue_singleton[of b]
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   520
    by (intro emeasure_eq_AE) auto
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   521
  then show ?thesis by auto
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   522
qed
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   523
41706
a207a858d1f6 prefer p2e before e2p; use measure_unique_Int_stable_vimage;
hoelzl
parents: 41704
diff changeset
   524
subsection {* Lebesgue-Borel measure *}
a207a858d1f6 prefer p2e before e2p; use measure_unique_Int_stable_vimage;
hoelzl
parents: 41704
diff changeset
   525
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   526
definition "lborel = measure_of UNIV (sets borel) (emeasure lebesgue)"
41689
3e39b0e730d6 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
3e39b0e730d6 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
lemma
3e39b0e730d6 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
  shows space_lborel[simp]: "space lborel = UNIV"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   530
  and sets_lborel[simp]: "sets lborel = sets borel"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   531
  and measurable_lborel1[simp]: "measurable lborel = measurable borel"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   532
  and measurable_lborel2[simp]: "measurable A lborel = measurable A borel"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   533
  using sigma_sets_eq[of borel]
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   534
  by (auto simp add: lborel_def measurable_def[abs_def])
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   535
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   536
lemma emeasure_lborel[simp]: "A \<in> sets borel \<Longrightarrow> emeasure lborel A = emeasure lebesgue A"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   537
  by (rule emeasure_measure_of[OF lborel_def])
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   538
     (auto simp: positive_def emeasure_nonneg countably_additive_def intro!: suminf_emeasure)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   539
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   540
interpretation lborel: sigma_finite_measure lborel
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   541
proof (default, intro conjI exI[of _ "\<lambda>n. cube n"])
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   542
  show "range cube \<subseteq> sets lborel" by (auto intro: borel_closed)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   543
  { fix x :: 'a have "\<exists>n. x\<in>cube n" using mem_big_cube by auto }
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   544
  then show "(\<Union>i. cube i) = (space lborel :: 'a set)" using mem_big_cube by auto
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   545
  show "\<forall>i. emeasure lborel (cube i) \<noteq> \<infinity>" by (simp add: cube_def)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   546
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
   547
3e39b0e730d6 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
interpretation lebesgue: sigma_finite_measure lebesgue
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   549
proof
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   550
  from lborel.sigma_finite guess A :: "nat \<Rightarrow> 'a set" ..
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   551
  then show "\<exists>A::nat \<Rightarrow> 'a set. range A \<subseteq> sets lebesgue \<and> (\<Union>i. A i) = space lebesgue \<and> (\<forall>i. emeasure lebesgue (A i) \<noteq> \<infinity>)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   552
    by (intro exI[of _ A]) (auto simp: subset_eq)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   553
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   554
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   555
lemma Int_stable_atLeastAtMost:
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   556
  fixes x::"'a::ordered_euclidean_space"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   557
  shows "Int_stable (range (\<lambda>(a, b::'a). {a..b}))"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   558
  by (auto simp: inter_interval Int_stable_def)
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   559
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   560
lemma lborel_eqI:
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   561
  fixes M :: "'a::ordered_euclidean_space measure"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   562
  assumes emeasure_eq: "\<And>a b. emeasure M {a .. b} = content {a .. b}"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   563
  assumes sets_eq: "sets M = sets borel"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   564
  shows "lborel = M"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   565
proof (rule measure_eqI_generator_eq[OF Int_stable_atLeastAtMost])
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   566
  let ?P = "\<Pi>\<^isub>M i\<in>{..<DIM('a::ordered_euclidean_space)}. lborel"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   567
  let ?E = "range (\<lambda>(a, b). {a..b} :: 'a set)"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   568
  show "?E \<subseteq> Pow UNIV" "sets lborel = sigma_sets UNIV ?E" "sets M = sigma_sets UNIV ?E"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   569
    by (simp_all add: borel_eq_atLeastAtMost sets_eq)
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   570
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   571
  show "range cube \<subseteq> ?E" unfolding cube_def [abs_def] by auto
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   572
  { fix x :: 'a have "\<exists>n. x \<in> cube n" using mem_big_cube[of x] by fastforce }
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   573
  then show "(\<Union>i. cube i :: 'a set) = UNIV" by auto
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   574
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   575
  { fix i show "emeasure lborel (cube i) \<noteq> \<infinity>" unfolding cube_def by auto }
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   576
  { fix X assume "X \<in> ?E" then show "emeasure lborel X = emeasure M X"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   577
      by (auto simp: emeasure_eq) }
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   578
qed
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   579
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   580
lemma lebesgue_real_affine:
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   581
  fixes c :: real assumes "c \<noteq> 0"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   582
  shows "lborel = density (distr lborel borel (\<lambda>x. t + c * x)) (\<lambda>_. \<bar>c\<bar>)" (is "_ = ?D")
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   583
proof (rule lborel_eqI)
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   584
  fix a b show "emeasure ?D {a..b} = content {a .. b}"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   585
  proof cases
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   586
    assume "0 < c"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   587
    then have "(\<lambda>x. t + c * x) -` {a..b} = {(a - t) / c .. (b - t) / c}"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   588
      by (auto simp: field_simps)
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   589
    with `0 < c` show ?thesis
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   590
      by (cases "a \<le> b")
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   591
         (auto simp: field_simps emeasure_density positive_integral_distr positive_integral_cmult
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   592
                     borel_measurable_indicator' emeasure_distr)
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   593
  next
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   594
    assume "\<not> 0 < c" with `c \<noteq> 0` have "c < 0" by auto
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   595
    then have *: "(\<lambda>x. t + c * x) -` {a..b} = {(b - t) / c .. (a - t) / c}"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   596
      by (auto simp: field_simps)
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   597
    with `c < 0` show ?thesis
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   598
      by (cases "a \<le> b")
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   599
         (auto simp: field_simps emeasure_density positive_integral_distr
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   600
                     positive_integral_cmult borel_measurable_indicator' emeasure_distr)
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   601
  qed
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   602
qed simp
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   603
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   604
lemma lebesgue_integral_real_affine:
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   605
  fixes c :: real assumes c: "c \<noteq> 0" and f: "f \<in> borel_measurable borel"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   606
  shows "(\<integral> x. f x \<partial> lborel) = \<bar>c\<bar> * (\<integral> x. f (t + c * x) \<partial>lborel)"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   607
  by (subst lebesgue_real_affine[OF c, of t])
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   608
     (simp add: f integral_density integral_distr lebesgue_integral_cmult)
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   609
41706
a207a858d1f6 prefer p2e before e2p; use measure_unique_Int_stable_vimage;
hoelzl
parents: 41704
diff changeset
   610
subsection {* Lebesgue integrable implies Gauge integrable *}
a207a858d1f6 prefer p2e before e2p; use measure_unique_Int_stable_vimage;
hoelzl
parents: 41704
diff changeset
   611
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   612
lemma has_integral_cmult_real:
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   613
  fixes c :: real
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   614
  assumes "c \<noteq> 0 \<Longrightarrow> (f has_integral x) A"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   615
  shows "((\<lambda>x. c * f x) has_integral c * x) A"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   616
proof cases
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   617
  assume "c \<noteq> 0"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   618
  from has_integral_cmul[OF assms[OF this], of c] show ?thesis
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   619
    unfolding real_scaleR_def .
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   620
qed simp
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   621
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   622
lemma simple_function_has_integral:
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
   623
  fixes f::"'a::ordered_euclidean_space \<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
   624
  assumes f:"simple_function lebesgue f"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   625
  and f':"range f \<subseteq> {0..<\<infinity>}"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   626
  and om:"\<And>x. x \<in> range f \<Longrightarrow> emeasure lebesgue (f -` {x} \<inter> UNIV) = \<infinity> \<Longrightarrow> 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
   627
  shows "((\<lambda>x. real (f x)) has_integral (real (integral\<^isup>S lebesgue f))) UNIV"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   628
  unfolding simple_integral_def space_lebesgue
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   629
proof (subst lebesgue_simple_function_indicator)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   630
  let ?M = "\<lambda>x. emeasure lebesgue (f -` {x} \<inter> UNIV)"
46731
5302e932d1e5 avoid undeclared variables in let bindings;
wenzelm
parents: 44928
diff changeset
   631
  let ?F = "\<lambda>x. indicator (f -` {x})"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   632
  { fix x y assume "y \<in> range f"
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
   633
    from subsetD[OF f' this] have "y * ?F y x = ereal (real y * ?F y x)"
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
   634
      by (cases rule: ereal2_cases[of y "?F y x"])
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
   635
         (auto simp: indicator_def one_ereal_def split: split_if_asm) }
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   636
  moreover
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   637
  { fix x assume x: "x\<in>range f"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   638
    have "x * ?M x = real x * real (?M x)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   639
    proof cases
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   640
      assume "x \<noteq> 0" with om[OF x] have "?M x \<noteq> \<infinity>" by auto
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   641
      with subsetD[OF f' x] f[THEN simple_functionD(2)] show ?thesis
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
   642
        by (cases rule: ereal2_cases[of x "?M x"]) auto
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   643
    qed simp }
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   644
  ultimately
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   645
  have "((\<lambda>x. real (\<Sum>y\<in>range f. y * ?F y x)) has_integral real (\<Sum>x\<in>range f. x * ?M x)) UNIV \<longleftrightarrow>
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   646
    ((\<lambda>x. \<Sum>y\<in>range f. real y * ?F y x) has_integral (\<Sum>x\<in>range f. real x * real (?M x))) UNIV"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   647
    by simp
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   648
  also have \<dots>
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   649
  proof (intro has_integral_setsum has_integral_cmult_real lmeasure_finite_has_integral
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   650
               real_of_ereal_pos emeasure_nonneg ballI)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   651
    show *: "finite (range f)" "\<And>y. f -` {y} \<in> sets lebesgue"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   652
      using simple_functionD[OF f] by auto
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   653
    fix y assume "real y \<noteq> 0" "y \<in> range f"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   654
    with * om[OF this(2)] show "emeasure lebesgue (f -` {y}) = ereal (real (?M y))"
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
   655
      by (auto simp: ereal_real)
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   656
  qed
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   657
  finally show "((\<lambda>x. real (\<Sum>y\<in>range f. y * ?F y x)) has_integral real (\<Sum>x\<in>range f. x * ?M x)) UNIV" .
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   658
qed fact
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   659
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   660
lemma bounded_realI: assumes "\<forall>x\<in>s. abs (x::real) \<le> B" shows "bounded s"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   661
  unfolding bounded_def dist_real_def apply(rule_tac x=0 in exI)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   662
  using assms by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   663
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   664
lemma simple_function_has_integral':
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
   665
  fixes f::"'a::ordered_euclidean_space \<Rightarrow> ereal"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   666
  assumes f: "simple_function lebesgue 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
   667
  and i: "integral\<^isup>S lebesgue f \<noteq> \<infinity>"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   668
  shows "((\<lambda>x. real (f x)) has_integral (real (integral\<^isup>S lebesgue f))) UNIV"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   669
proof -
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   670
  let ?f = "\<lambda>x. if x \<in> f -` {\<infinity>} then 0 else f x"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   671
  note f(1)[THEN simple_functionD(2)]
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   672
  then have [simp, intro]: "\<And>X. f -` X \<in> sets lebesgue" by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   673
  have f': "simple_function lebesgue ?f"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   674
    using f by (intro simple_function_If_set) auto
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   675
  have rng: "range ?f \<subseteq> {0..<\<infinity>}" using f by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   676
  have "AE x in lebesgue. f x = ?f x"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   677
    using simple_integral_PInf[OF f i]
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   678
    by (intro AE_I[where N="f -` {\<infinity>} \<inter> space lebesgue"]) auto
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   679
  from f(1) f' this have eq: "integral\<^isup>S lebesgue f = integral\<^isup>S lebesgue ?f"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   680
    by (rule simple_integral_cong_AE)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   681
  have real_eq: "\<And>x. real (f x) = real (?f x)" by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   682
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   683
  show ?thesis
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   684
    unfolding eq real_eq
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   685
  proof (rule simple_function_has_integral[OF f' rng])
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   686
    fix x assume x: "x \<in> range ?f" and inf: "emeasure lebesgue (?f -` {x} \<inter> UNIV) = \<infinity>"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   687
    have "x * emeasure lebesgue (?f -` {x} \<inter> UNIV) = (\<integral>\<^isup>S y. x * indicator (?f -` {x}) y \<partial>lebesgue)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   688
      using f'[THEN simple_functionD(2)]
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   689
      by (simp add: simple_integral_cmult_indicator)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   690
    also have "\<dots> \<le> integral\<^isup>S lebesgue f"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   691
      using f'[THEN simple_functionD(2)] f
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   692
      by (intro simple_integral_mono simple_function_mult simple_function_indicator)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   693
         (auto split: split_indicator)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   694
    finally show "x = 0" unfolding inf using i subsetD[OF rng x] by (auto split: split_if_asm)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   695
  qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   696
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   697
49801
f3471f09bb86 add induction for real Borel measurable functions
hoelzl
parents: 49784
diff changeset
   698
lemma borel_measurable_real_induct[consumes 2, case_names cong set mult add seq, induct set: borel_measurable]:
f3471f09bb86 add induction for real Borel measurable functions
hoelzl
parents: 49784
diff changeset
   699
  fixes u :: "'a \<Rightarrow> real"
f3471f09bb86 add induction for real Borel measurable functions
hoelzl
parents: 49784
diff changeset
   700
  assumes u: "u \<in> borel_measurable M" "\<And>x. 0 \<le> u x"
f3471f09bb86 add induction for real Borel measurable functions
hoelzl
parents: 49784
diff changeset
   701
  assumes cong: "\<And>f g. f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> (\<And>x. x \<in> space M \<Longrightarrow> f x = g x) \<Longrightarrow> P g \<Longrightarrow> P f"
f3471f09bb86 add induction for real Borel measurable functions
hoelzl
parents: 49784
diff changeset
   702
  assumes set: "\<And>A. A \<in> sets M \<Longrightarrow> P (indicator A)"
f3471f09bb86 add induction for real Borel measurable functions
hoelzl
parents: 49784
diff changeset
   703
  assumes mult: "\<And>u c. 0 \<le> c \<Longrightarrow> u \<in> borel_measurable M \<Longrightarrow> (\<And>x. 0 \<le> u x) \<Longrightarrow> P u \<Longrightarrow> P (\<lambda>x. c * u x)"
f3471f09bb86 add induction for real Borel measurable functions
hoelzl
parents: 49784
diff changeset
   704
  assumes add: "\<And>u v. u \<in> borel_measurable M \<Longrightarrow> (\<And>x. 0 \<le> u x) \<Longrightarrow> P u \<Longrightarrow> v \<in> borel_measurable M \<Longrightarrow> (\<And>x. 0 \<le> v x) \<Longrightarrow> P v \<Longrightarrow> P (\<lambda>x. v x + u x)"
f3471f09bb86 add induction for real Borel measurable functions
hoelzl
parents: 49784
diff changeset
   705
  assumes seq: "\<And>U u. (\<And>i. U i \<in> borel_measurable M) \<Longrightarrow> (\<And>i x. 0 \<le> U i x) \<Longrightarrow> (\<And>x. (\<lambda>i. U i x) ----> u x) \<Longrightarrow> (\<And>i. P (U i)) \<Longrightarrow> incseq U \<Longrightarrow> P u"
f3471f09bb86 add induction for real Borel measurable functions
hoelzl
parents: 49784
diff changeset
   706
  shows "P u"
f3471f09bb86 add induction for real Borel measurable functions
hoelzl
parents: 49784
diff changeset
   707
proof -
f3471f09bb86 add induction for real Borel measurable functions
hoelzl
parents: 49784
diff changeset
   708
  have "(\<lambda>x. ereal (u x)) \<in> borel_measurable M"
f3471f09bb86 add induction for real Borel measurable functions
hoelzl
parents: 49784
diff changeset
   709
    using u by auto
f3471f09bb86 add induction for real Borel measurable functions
hoelzl
parents: 49784
diff changeset
   710
  then obtain U where U: "\<And>i. simple_function M (U i)" "incseq U" "\<And>i. \<infinity> \<notin> range (U i)" and
f3471f09bb86 add induction for real Borel measurable functions
hoelzl
parents: 49784
diff changeset
   711
    "\<And>x. (SUP i. U i x) = max 0 (ereal (u x))" and nn: "\<And>i x. 0 \<le> U i x"
f3471f09bb86 add induction for real Borel measurable functions
hoelzl
parents: 49784
diff changeset
   712
    using borel_measurable_implies_simple_function_sequence'[of u M] by auto
f3471f09bb86 add induction for real Borel measurable functions
hoelzl
parents: 49784
diff changeset
   713
  then have u_eq: "\<And>x. ereal (u x) = (SUP i. U i x)"
f3471f09bb86 add induction for real Borel measurable functions
hoelzl
parents: 49784
diff changeset
   714
    using u by (auto simp: max_def)
f3471f09bb86 add induction for real Borel measurable functions
hoelzl
parents: 49784
diff changeset
   715
f3471f09bb86 add induction for real Borel measurable functions
hoelzl
parents: 49784
diff changeset
   716
  have [simp]: "\<And>i x. U i x \<noteq> \<infinity>" using U by (auto simp: image_def eq_commute)
f3471f09bb86 add induction for real Borel measurable functions
hoelzl
parents: 49784
diff changeset
   717
f3471f09bb86 add induction for real Borel measurable functions
hoelzl
parents: 49784
diff changeset
   718
  { fix i x have [simp]: "U i x \<noteq> -\<infinity>" using nn[of i x] by auto }
f3471f09bb86 add induction for real Borel measurable functions
hoelzl
parents: 49784
diff changeset
   719
  note this[simp]
f3471f09bb86 add induction for real Borel measurable functions
hoelzl
parents: 49784
diff changeset
   720
f3471f09bb86 add induction for real Borel measurable functions
hoelzl
parents: 49784
diff changeset
   721
  show "P u"
f3471f09bb86 add induction for real Borel measurable functions
hoelzl
parents: 49784
diff changeset
   722
  proof (rule seq)
f3471f09bb86 add induction for real Borel measurable functions
hoelzl
parents: 49784
diff changeset
   723
    show "\<And>i. (\<lambda>x. real (U i x)) \<in> borel_measurable M"
f3471f09bb86 add induction for real Borel measurable functions
hoelzl
parents: 49784
diff changeset
   724
      using U by (auto intro: borel_measurable_simple_function)
f3471f09bb86 add induction for real Borel measurable functions
hoelzl
parents: 49784
diff changeset
   725
    show "\<And>i x. 0 \<le> real (U i x)"
f3471f09bb86 add induction for real Borel measurable functions
hoelzl
parents: 49784
diff changeset
   726
      using nn by (auto simp: real_of_ereal_pos)
f3471f09bb86 add induction for real Borel measurable functions
hoelzl
parents: 49784
diff changeset
   727
    show "incseq (\<lambda>i x. real (U i x))"
f3471f09bb86 add induction for real Borel measurable functions
hoelzl
parents: 49784
diff changeset
   728
      using U(2) by (auto simp: incseq_def image_iff le_fun_def intro!: real_of_ereal_positive_mono nn)
f3471f09bb86 add induction for real Borel measurable functions
hoelzl
parents: 49784
diff changeset
   729
    then show "\<And>x. (\<lambda>i. real (U i x)) ----> u x"
f3471f09bb86 add induction for real Borel measurable functions
hoelzl
parents: 49784
diff changeset
   730
      by (intro SUP_eq_LIMSEQ[THEN iffD1])
f3471f09bb86 add induction for real Borel measurable functions
hoelzl
parents: 49784
diff changeset
   731
         (auto simp: incseq_mono incseq_def le_fun_def u_eq ereal_real
f3471f09bb86 add induction for real Borel measurable functions
hoelzl
parents: 49784
diff changeset
   732
               intro!: arg_cong2[where f=SUPR] ext)
f3471f09bb86 add induction for real Borel measurable functions
hoelzl
parents: 49784
diff changeset
   733
    show "\<And>i. P (\<lambda>x. real (U i x))"
f3471f09bb86 add induction for real Borel measurable functions
hoelzl
parents: 49784
diff changeset
   734
    proof (rule cong)
f3471f09bb86 add induction for real Borel measurable functions
hoelzl
parents: 49784
diff changeset
   735
      fix x i assume x: "x \<in> space M"
f3471f09bb86 add induction for real Borel measurable functions
hoelzl
parents: 49784
diff changeset
   736
      have [simp]: "\<And>A x. real (indicator A x :: ereal) = indicator A x"
f3471f09bb86 add induction for real Borel measurable functions
hoelzl
parents: 49784
diff changeset
   737
        by (auto simp: indicator_def one_ereal_def)
f3471f09bb86 add induction for real Borel measurable functions
hoelzl
parents: 49784
diff changeset
   738
      { fix y assume "y \<in> U i ` space M"
f3471f09bb86 add induction for real Borel measurable functions
hoelzl
parents: 49784
diff changeset
   739
        then have "0 \<le> y" "y \<noteq> \<infinity>" using nn by auto
f3471f09bb86 add induction for real Borel measurable functions
hoelzl
parents: 49784
diff changeset
   740
        then have "\<bar>y * indicator (U i -` {y} \<inter> space M) x\<bar> \<noteq> \<infinity>"
f3471f09bb86 add induction for real Borel measurable functions
hoelzl
parents: 49784
diff changeset
   741
          by (auto simp: indicator_def) }
f3471f09bb86 add induction for real Borel measurable functions
hoelzl
parents: 49784
diff changeset
   742
      then show "real (U i x) = (\<Sum>y \<in> U i ` space M. real y * indicator (U i -` {y} \<inter> space M) x)"
f3471f09bb86 add induction for real Borel measurable functions
hoelzl
parents: 49784
diff changeset
   743
        unfolding simple_function_indicator_representation[OF U(1) x]
f3471f09bb86 add induction for real Borel measurable functions
hoelzl
parents: 49784
diff changeset
   744
        by (subst setsum_real_of_ereal[symmetric]) auto
f3471f09bb86 add induction for real Borel measurable functions
hoelzl
parents: 49784
diff changeset
   745
    next
f3471f09bb86 add induction for real Borel measurable functions
hoelzl
parents: 49784
diff changeset
   746
      fix i
f3471f09bb86 add induction for real Borel measurable functions
hoelzl
parents: 49784
diff changeset
   747
      have "finite (U i ` space M)" "\<And>x. x \<in> U i ` space M \<Longrightarrow> 0 \<le> x""\<And>x. x \<in> U i ` space M \<Longrightarrow> x \<noteq> \<infinity>"
f3471f09bb86 add induction for real Borel measurable functions
hoelzl
parents: 49784
diff changeset
   748
        using U(1) nn by (auto simp: simple_function_def)
f3471f09bb86 add induction for real Borel measurable functions
hoelzl
parents: 49784
diff changeset
   749
      then show "P (\<lambda>x. \<Sum>y \<in> U i ` space M. real y * indicator (U i -` {y} \<inter> space M) x)"
f3471f09bb86 add induction for real Borel measurable functions
hoelzl
parents: 49784
diff changeset
   750
      proof induct
f3471f09bb86 add induction for real Borel measurable functions
hoelzl
parents: 49784
diff changeset
   751
        case empty then show ?case
f3471f09bb86 add induction for real Borel measurable functions
hoelzl
parents: 49784
diff changeset
   752
          using set[of "{}"] by (simp add: indicator_def[abs_def])
f3471f09bb86 add induction for real Borel measurable functions
hoelzl
parents: 49784
diff changeset
   753
      qed (auto intro!: add mult set simple_functionD U setsum_nonneg borel_measurable_setsum mult_nonneg_nonneg real_of_ereal_pos)
f3471f09bb86 add induction for real Borel measurable functions
hoelzl
parents: 49784
diff changeset
   754
    qed (auto intro: borel_measurable_simple_function U simple_functionD intro!: borel_measurable_setsum borel_measurable_times)
f3471f09bb86 add induction for real Borel measurable functions
hoelzl
parents: 49784
diff changeset
   755
  qed
f3471f09bb86 add induction for real Borel measurable functions
hoelzl
parents: 49784
diff changeset
   756
qed
f3471f09bb86 add induction for real Borel measurable functions
hoelzl
parents: 49784
diff changeset
   757
f3471f09bb86 add induction for real Borel measurable functions
hoelzl
parents: 49784
diff changeset
   758
lemma ereal_indicator: "ereal (indicator A x) = indicator A x"
f3471f09bb86 add induction for real Borel measurable functions
hoelzl
parents: 49784
diff changeset
   759
  by (auto simp: indicator_def one_ereal_def)
f3471f09bb86 add induction for real Borel measurable functions
hoelzl
parents: 49784
diff changeset
   760
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   761
lemma positive_integral_has_integral:
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
   762
  fixes f :: "'a::ordered_euclidean_space \<Rightarrow> ereal"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   763
  assumes f: "f \<in> borel_measurable lebesgue" "range f \<subseteq> {0..<\<infinity>}" "integral\<^isup>P lebesgue f \<noteq> \<infinity>"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   764
  shows "((\<lambda>x. real (f x)) has_integral (real (integral\<^isup>P lebesgue f))) UNIV"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   765
proof -
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   766
  from borel_measurable_implies_simple_function_sequence'[OF f(1)]
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   767
  guess u . note u = this
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   768
  have SUP_eq: "\<And>x. (SUP i. u i x) = f x"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   769
    using u(4) f(2)[THEN subsetD] by (auto split: split_max)
46731
5302e932d1e5 avoid undeclared variables in let bindings;
wenzelm
parents: 44928
diff changeset
   770
  let ?u = "\<lambda>i x. real (u i x)"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   771
  note u_eq = positive_integral_eq_simple_integral[OF u(1,5), symmetric]
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   772
  { fix i
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   773
    note u_eq
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   774
    also have "integral\<^isup>P lebesgue (u i) \<le> (\<integral>\<^isup>+x. max 0 (f x) \<partial>lebesgue)"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   775
      by (intro positive_integral_mono) (auto intro: SUP_upper simp: u(4)[symmetric])
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   776
    finally have "integral\<^isup>S lebesgue (u i) \<noteq> \<infinity>"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   777
      unfolding positive_integral_max_0 using f by auto }
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   778
  note u_fin = this
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   779
  then have u_int: "\<And>i. (?u i has_integral real (integral\<^isup>S lebesgue (u i))) UNIV"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   780
    by (rule simple_function_has_integral'[OF u(1,5)])
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
   781
  have "\<forall>x. \<exists>r\<ge>0. f x = ereal r"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   782
  proof
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   783
    fix x from f(2) have "0 \<le> f x" "f x \<noteq> \<infinity>" by (auto simp: subset_eq)
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
   784
    then show "\<exists>r\<ge>0. f x = ereal r" by (cases "f x") auto
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   785
  qed
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
   786
  from choice[OF this] obtain f' where f': "f = (\<lambda>x. ereal (f' x))" "\<And>x. 0 \<le> f' x" by auto
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   787
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
   788
  have "\<forall>i. \<exists>r. \<forall>x. 0 \<le> r x \<and> u i x = ereal (r x)"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   789
  proof
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
   790
    fix i show "\<exists>r. \<forall>x. 0 \<le> r x \<and> u i x = ereal (r x)"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   791
    proof (intro choice allI)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   792
      fix i x have "u i x \<noteq> \<infinity>" using u(3)[of i] by (auto simp: image_iff) metis
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
   793
      then show "\<exists>r\<ge>0. u i x = ereal r" using u(5)[of i x] by (cases "u i x") auto
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   794
    qed
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   795
  qed
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   796
  from choice[OF this] obtain u' where
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
   797
      u': "u = (\<lambda>i x. ereal (u' i x))" "\<And>i x. 0 \<le> u' i x" by (auto simp: fun_eq_iff)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   798
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   799
  have convergent: "f' integrable_on UNIV \<and>
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   800
    (\<lambda>k. integral UNIV (u' k)) ----> integral UNIV f'"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   801
  proof (intro monotone_convergence_increasing allI ballI)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   802
    show int: "\<And>k. (u' k) integrable_on UNIV"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   803
      using u_int unfolding integrable_on_def u' by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   804
    show "\<And>k x. u' k x \<le> u' (Suc k) x" using u(2,3,5)
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
   805
      by (auto simp: incseq_Suc_iff le_fun_def image_iff u' intro!: real_of_ereal_positive_mono)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   806
    show "\<And>x. (\<lambda>k. u' k x) ----> f' x"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   807
      using SUP_eq u(2)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   808
      by (intro SUP_eq_LIMSEQ[THEN iffD1]) (auto simp: u' f' incseq_mono incseq_Suc_iff le_fun_def)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   809
    show "bounded {integral UNIV (u' k)|k. True}"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   810
    proof (safe intro!: bounded_realI)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   811
      fix k
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   812
      have "\<bar>integral UNIV (u' k)\<bar> = integral UNIV (u' k)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   813
        by (intro abs_of_nonneg integral_nonneg int ballI u')
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   814
      also have "\<dots> = real (integral\<^isup>S lebesgue (u k))"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   815
        using u_int[THEN integral_unique] by (simp add: u')
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   816
      also have "\<dots> = real (integral\<^isup>P lebesgue (u k))"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   817
        using positive_integral_eq_simple_integral[OF u(1,5)] by simp
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   818
      also have "\<dots> \<le> real (integral\<^isup>P lebesgue f)" using f
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   819
        by (auto intro!: real_of_ereal_positive_mono positive_integral_positive
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   820
             positive_integral_mono SUP_upper simp: SUP_eq[symmetric])
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   821
      finally show "\<bar>integral UNIV (u' k)\<bar> \<le> real (integral\<^isup>P lebesgue f)" .
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   822
    qed
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   823
  qed
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   824
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
   825
  have "integral\<^isup>P lebesgue f = ereal (integral UNIV f')"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   826
  proof (rule tendsto_unique[OF trivial_limit_sequentially])
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   827
    have "(\<lambda>i. integral\<^isup>S lebesgue (u i)) ----> (SUP i. integral\<^isup>P lebesgue (u i))"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   828
      unfolding u_eq by (intro LIMSEQ_ereal_SUPR incseq_positive_integral u)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   829
    also note positive_integral_monotone_convergence_SUP
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   830
      [OF u(2)  borel_measurable_simple_function[OF u(1)] u(5), symmetric]
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   831
    finally show "(\<lambda>k. integral\<^isup>S lebesgue (u k)) ----> integral\<^isup>P lebesgue f"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   832
      unfolding SUP_eq .
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   833
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   834
    { fix k
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   835
      have "0 \<le> integral\<^isup>S lebesgue (u k)"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   836
        using u by (auto intro!: simple_integral_positive)
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
   837
      then have "integral\<^isup>S lebesgue (u k) = ereal (real (integral\<^isup>S lebesgue (u k)))"
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
   838
        using u_fin by (auto simp: ereal_real) }
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   839
    note * = this
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
   840
    show "(\<lambda>k. integral\<^isup>S lebesgue (u k)) ----> ereal (integral UNIV f')"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   841
      using convergent using u_int[THEN integral_unique, symmetric]
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   842
      by (subst *) (simp add: u')
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   843
  qed
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   844
  then show ?thesis using convergent by (simp add: f' integrable_integral)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   845
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   846
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   847
lemma lebesgue_integral_has_integral:
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   848
  fixes f :: "'a::ordered_euclidean_space \<Rightarrow> real"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   849
  assumes f: "integrable lebesgue f"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   850
  shows "(f has_integral (integral\<^isup>L lebesgue f)) UNIV"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   851
proof -
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
   852
  let ?n = "\<lambda>x. real (ereal (max 0 (- f x)))" and ?p = "\<lambda>x. real (ereal (max 0 (f x)))"
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
   853
  have *: "f = (\<lambda>x. ?p x - ?n x)" by (auto simp del: ereal_max)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   854
  { fix f :: "'a \<Rightarrow> real" have "(\<integral>\<^isup>+ x. ereal (f x) \<partial>lebesgue) = (\<integral>\<^isup>+ x. ereal (max 0 (f x)) \<partial>lebesgue)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   855
      by (intro positive_integral_cong_pos) (auto split: split_max) }
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   856
  note eq = this
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   857
  show ?thesis
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   858
    unfolding lebesgue_integral_def
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   859
    apply (subst *)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   860
    apply (rule has_integral_sub)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   861
    unfolding eq[of f] eq[of "\<lambda>x. - f x"]
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   862
    apply (safe intro!: positive_integral_has_integral)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   863
    using integrableD[OF f]
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
   864
    by (auto simp: zero_ereal_def[symmetric] positive_integral_max_0  split: split_max
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   865
             intro!: measurable_If)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   866
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   867
47757
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   868
lemma lebesgue_simple_integral_eq_borel:
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   869
  assumes f: "f \<in> borel_measurable borel"
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   870
  shows "integral\<^isup>S lebesgue f = integral\<^isup>S lborel f"
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   871
  using f[THEN measurable_sets]
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   872
  by (auto intro!: setsum_cong arg_cong2[where f="op *"] emeasure_lborel[symmetric]
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   873
           simp: simple_integral_def)
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   874
41546
2a12c23b7a34 integral on lebesgue measure is extension of integral on borel measure
hoelzl
parents: 41097
diff changeset
   875
lemma lebesgue_positive_integral_eq_borel:
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   876
  assumes f: "f \<in> borel_measurable borel"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   877
  shows "integral\<^isup>P lebesgue f = integral\<^isup>P lborel f"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   878
proof -
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   879
  from f have "integral\<^isup>P lebesgue (\<lambda>x. max 0 (f x)) = integral\<^isup>P lborel (\<lambda>x. max 0 (f x))"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   880
    by (auto intro!: positive_integral_subalgebra[symmetric])
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   881
  then show ?thesis unfolding positive_integral_max_0 .
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   882
qed
41546
2a12c23b7a34 integral on lebesgue measure is extension of integral on borel measure
hoelzl
parents: 41097
diff changeset
   883
2a12c23b7a34 integral on lebesgue measure is extension of integral on borel measure
hoelzl
parents: 41097
diff changeset
   884
lemma lebesgue_integral_eq_borel:
2a12c23b7a34 integral on lebesgue measure is extension of integral on borel measure
hoelzl
parents: 41097
diff changeset
   885
  assumes "f \<in> borel_measurable borel"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   886
  shows "integrable lebesgue f \<longleftrightarrow> integrable lborel 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
   887
    and "integral\<^isup>L lebesgue f = integral\<^isup>L lborel f" (is ?I)
41546
2a12c23b7a34 integral on lebesgue measure is extension of integral on borel measure
hoelzl
parents: 41097
diff changeset
   888
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
   889
  have "sets lborel \<subseteq> sets lebesgue" by auto
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   890
  from integral_subalgebra[of f lborel, OF _ this _ _] assms
41546
2a12c23b7a34 integral on lebesgue measure is extension of integral on borel measure
hoelzl
parents: 41097
diff changeset
   891
  show ?P ?I by auto
2a12c23b7a34 integral on lebesgue measure is extension of integral on borel measure
hoelzl
parents: 41097
diff changeset
   892
qed
2a12c23b7a34 integral on lebesgue measure is extension of integral on borel measure
hoelzl
parents: 41097
diff changeset
   893
2a12c23b7a34 integral on lebesgue measure is extension of integral on borel measure
hoelzl
parents: 41097
diff changeset
   894
lemma borel_integral_has_integral:
2a12c23b7a34 integral on lebesgue measure is extension of integral on borel measure
hoelzl
parents: 41097
diff changeset
   895
  fixes f::"'a::ordered_euclidean_space => real"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   896
  assumes f:"integrable lborel 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
   897
  shows "(f has_integral (integral\<^isup>L lborel f)) UNIV"
41546
2a12c23b7a34 integral on lebesgue measure is extension of integral on borel measure
hoelzl
parents: 41097
diff changeset
   898
proof -
2a12c23b7a34 integral on lebesgue measure is extension of integral on borel measure
hoelzl
parents: 41097
diff changeset
   899
  have borel: "f \<in> borel_measurable borel"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   900
    using f unfolding integrable_def by auto
41546
2a12c23b7a34 integral on lebesgue measure is extension of integral on borel measure
hoelzl
parents: 41097
diff changeset
   901
  from f show ?thesis
2a12c23b7a34 integral on lebesgue measure is extension of integral on borel measure
hoelzl
parents: 41097
diff changeset
   902
    using lebesgue_integral_has_integral[of f]
2a12c23b7a34 integral on lebesgue measure is extension of integral on borel measure
hoelzl
parents: 41097
diff changeset
   903
    unfolding lebesgue_integral_eq_borel[OF borel] by simp
2a12c23b7a34 integral on lebesgue measure is extension of integral on borel measure
hoelzl
parents: 41097
diff changeset
   904
qed
2a12c23b7a34 integral on lebesgue measure is extension of integral on borel measure
hoelzl
parents: 41097
diff changeset
   905
47757
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   906
lemma integrable_on_cmult_iff:
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   907
  fixes c :: real assumes "c \<noteq> 0"
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   908
  shows "(\<lambda>x. c * f x) integrable_on s \<longleftrightarrow> f integrable_on s"
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   909
  using integrable_cmul[of "\<lambda>x. c * f x" s "1 / c"] integrable_cmul[of f s c] `c \<noteq> 0`
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   910
  by auto
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   911
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   912
lemma positive_integral_lebesgue_has_integral:
47757
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   913
  fixes f :: "'a::ordered_euclidean_space \<Rightarrow> real"
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   914
  assumes f_borel: "f \<in> borel_measurable lebesgue" and nonneg: "\<And>x. 0 \<le> f x"
47757
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   915
  assumes I: "(f has_integral I) UNIV"
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   916
  shows "(\<integral>\<^isup>+x. f x \<partial>lebesgue) = I"
47757
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   917
proof -
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   918
  from f_borel have "(\<lambda>x. ereal (f x)) \<in> borel_measurable lebesgue" by auto
47757
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   919
  from borel_measurable_implies_simple_function_sequence'[OF this] guess F . note F = this
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   920
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   921
  have "(\<integral>\<^isup>+ x. ereal (f x) \<partial>lebesgue) = (SUP i. integral\<^isup>S lebesgue (F i))"
47757
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   922
    using F
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   923
    by (subst positive_integral_monotone_convergence_simple)
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   924
       (simp_all add: positive_integral_max_0 simple_function_def)
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   925
  also have "\<dots> \<le> ereal I"
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   926
  proof (rule SUP_least)
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   927
    fix i :: nat
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   928
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   929
    { fix z
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   930
      from F(4)[of z] have "F i z \<le> ereal (f z)"
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   931
        by (metis SUP_upper UNIV_I ereal_max_0 min_max.sup_absorb2 nonneg)
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   932
      with F(5)[of i z] have "real (F i z) \<le> f z"
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   933
        by (cases "F i z") simp_all }
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   934
    note F_bound = this
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   935
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   936
    { fix x :: ereal assume x: "x \<noteq> 0" "x \<in> range (F i)"
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   937
      with F(3,5)[of i] have [simp]: "real x \<noteq> 0"
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   938
        by (metis image_iff order_eq_iff real_of_ereal_le_0)
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   939
      let ?s = "(\<lambda>n z. real x * indicator (F i -` {x} \<inter> cube n) z) :: nat \<Rightarrow> 'a \<Rightarrow> real"
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   940
      have "(\<lambda>z::'a. real x * indicator (F i -` {x}) z) integrable_on UNIV"
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   941
      proof (rule dominated_convergence(1))
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   942
        fix n :: nat
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   943
        have "(\<lambda>z. indicator (F i -` {x} \<inter> cube n) z :: real) integrable_on cube n"
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   944
          using x F(1)[of i]
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   945
          by (intro lebesgueD) (auto simp: simple_function_def)
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   946
        then have cube: "?s n integrable_on cube n"
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   947
          by (simp add: integrable_on_cmult_iff)
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   948
        show "?s n integrable_on UNIV"
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   949
          by (rule integrable_on_superset[OF _ _ cube]) auto
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   950
      next
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   951
        show "f integrable_on UNIV"
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   952
          unfolding integrable_on_def using I by auto
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   953
      next
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   954
        fix n from F_bound show "\<forall>x\<in>UNIV. norm (?s n x) \<le> f x"
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   955
          using nonneg F(5) by (auto split: split_indicator)
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   956
      next
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   957
        show "\<forall>z\<in>UNIV. (\<lambda>n. ?s n z) ----> real x * indicator (F i -` {x}) z"
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   958
        proof
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   959
          fix z :: 'a
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   960
          from mem_big_cube[of z] guess j .
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   961
          then have x: "eventually (\<lambda>n. ?s n z = real x * indicator (F i -` {x}) z) sequentially"
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   962
            by (auto intro!: eventually_sequentiallyI[where c=j] dest!: cube_subset split: split_indicator)
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   963
          then show "(\<lambda>n. ?s n z) ----> real x * indicator (F i -` {x}) z"
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   964
            by (rule Lim_eventually)
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   965
        qed
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   966
      qed
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   967
      then have "(indicator (F i -` {x}) :: 'a \<Rightarrow> real) integrable_on UNIV"
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   968
        by (simp add: integrable_on_cmult_iff) }
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   969
    note F_finite = lmeasure_finite[OF this]
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   970
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   971
    have "((\<lambda>x. real (F i x)) has_integral real (integral\<^isup>S lebesgue (F i))) UNIV"
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   972
    proof (rule simple_function_has_integral[of "F i"])
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   973
      show "simple_function lebesgue (F i)"
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   974
        using F(1) by (simp add: simple_function_def)
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   975
      show "range (F i) \<subseteq> {0..<\<infinity>}"
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   976
        using F(3,5)[of i] by (auto simp: image_iff) metis
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   977
    next
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   978
      fix x assume "x \<in> range (F i)" "emeasure lebesgue (F i -` {x} \<inter> UNIV) = \<infinity>"
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   979
      with F_finite[of x] show "x = 0" by auto
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   980
    qed
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   981
    from this I have "real (integral\<^isup>S lebesgue (F i)) \<le> I"
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   982
      by (rule has_integral_le) (intro ballI F_bound)
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   983
    moreover
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   984
    { fix x assume x: "x \<in> range (F i)"
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   985
      with F(3,5)[of i] have "x = 0 \<or> (0 < x \<and> x \<noteq> \<infinity>)"
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   986
        by (auto  simp: image_iff le_less) metis
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   987
      with F_finite[OF _ x] x have "x * emeasure lebesgue (F i -` {x} \<inter> UNIV) \<noteq> \<infinity>"
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   988
        by auto }
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   989
    then have "integral\<^isup>S lebesgue (F i) \<noteq> \<infinity>"
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   990
      unfolding simple_integral_def setsum_Pinfty space_lebesgue by blast
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   991
    moreover have "0 \<le> integral\<^isup>S lebesgue (F i)"
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   992
      using F(1,5) by (intro simple_integral_positive) (auto simp: simple_function_def)
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   993
    ultimately show "integral\<^isup>S lebesgue (F i) \<le> ereal I"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   994
      by (cases "integral\<^isup>S lebesgue (F i)") auto
47757
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   995
  qed
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   996
  also have "\<dots> < \<infinity>" by simp
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   997
  finally have finite: "(\<integral>\<^isup>+ x. ereal (f x) \<partial>lebesgue) \<noteq> \<infinity>" by simp
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   998
  have borel: "(\<lambda>x. ereal (f x)) \<in> borel_measurable lebesgue"
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   999
    using f_borel by (auto intro: borel_measurable_lebesgueI)
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
  1000
  from positive_integral_has_integral[OF borel _ finite]
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
  1001
  have "(f has_integral real (\<integral>\<^isup>+ x. ereal (f x) \<partial>lebesgue)) UNIV"
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
  1002
    using nonneg by (simp add: subset_eq)
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
  1003
  with I have "I = real (\<integral>\<^isup>+ x. ereal (f x) \<partial>lebesgue)"
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
  1004
    by (rule has_integral_unique)
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
  1005
  with finite positive_integral_positive[of _ "\<lambda>x. ereal (f x)"] show ?thesis
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1006
    by (cases "\<integral>\<^isup>+ x. ereal (f x) \<partial>lebesgue") auto
47757
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
  1007
qed
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
  1008
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1009
lemma has_integral_iff_positive_integral_lebesgue:
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1010
  fixes f :: "'a::ordered_euclidean_space \<Rightarrow> real"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1011
  assumes f: "f \<in> borel_measurable lebesgue" "\<And>x. 0 \<le> f x"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1012
  shows "(f has_integral I) UNIV \<longleftrightarrow> integral\<^isup>P lebesgue f = I"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1013
  using f positive_integral_lebesgue_has_integral[of f I] positive_integral_has_integral[of f]
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1014
  by (auto simp: subset_eq)
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1015
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1016
lemma has_integral_iff_positive_integral_lborel:
47757
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
  1017
  fixes f :: "'a::ordered_euclidean_space \<Rightarrow> real"
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
  1018
  assumes f: "f \<in> borel_measurable borel" "\<And>x. 0 \<le> f x"
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
  1019
  shows "(f has_integral I) UNIV \<longleftrightarrow> integral\<^isup>P lborel f = I"
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1020
  using assms
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1021
  by (subst has_integral_iff_positive_integral_lebesgue)
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1022
     (auto simp: borel_measurable_lebesgueI lebesgue_positive_integral_eq_borel)
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1023
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1024
subsection {* Equivalence between product spaces and euclidean spaces *}
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1025
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1026
definition e2p :: "'a::ordered_euclidean_space \<Rightarrow> (nat \<Rightarrow> real)" where
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1027
  "e2p x = (\<lambda>i\<in>{..<DIM('a)}. x$$i)"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1028
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1029
definition p2e :: "(nat \<Rightarrow> real) \<Rightarrow> 'a::ordered_euclidean_space" where
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1030
  "p2e x = (\<chi>\<chi> i. x i)"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1031
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1032
lemma e2p_p2e[simp]:
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1033
  "x \<in> extensional {..<DIM('a)} \<Longrightarrow> e2p (p2e x::'a::ordered_euclidean_space) = x"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1034
  by (auto simp: fun_eq_iff extensional_def p2e_def e2p_def)
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1035
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1036
lemma p2e_e2p[simp]:
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1037
  "p2e (e2p x) = (x::'a::ordered_euclidean_space)"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1038
  by (auto simp: euclidean_eq[where 'a='a] p2e_def e2p_def)
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1039
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1040
interpretation lborel_product: product_sigma_finite "\<lambda>x. lborel::real measure"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1041
  by default
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1042
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1043
interpretation lborel_space: finite_product_sigma_finite "\<lambda>x. lborel::real measure" "{..<n}" for n :: nat
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1044
  by default auto
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1045
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1046
lemma bchoice_iff: "(\<forall>x\<in>A. \<exists>y. P x y) \<longleftrightarrow> (\<exists>f. \<forall>x\<in>A. P x (f x))"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1047
  by metis
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1048
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1049
lemma sets_product_borel:
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1050
  assumes I: "finite I"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1051
  shows "sets (\<Pi>\<^isub>M i\<in>I. lborel) = sigma_sets (\<Pi>\<^isub>E i\<in>I. UNIV) { \<Pi>\<^isub>E i\<in>I. {..< x i :: real} | x. True}" (is "_ = ?G")
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1052
proof (subst sigma_prod_algebra_sigma_eq[where S="\<lambda>_ i::nat. {..<real i}" and E="\<lambda>_. range lessThan", OF I])
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1053
  show "sigma_sets (space (Pi\<^isub>M I (\<lambda>i. lborel))) {Pi\<^isub>E I F |F. \<forall>i\<in>I. F i \<in> range lessThan} = ?G"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1054
    by (intro arg_cong2[where f=sigma_sets]) (auto simp: space_PiM image_iff bchoice_iff)
49779
1484b4b82855 remove incseq assumption from sigma_prod_algebra_sigma_eq
hoelzl
parents: 49777
diff changeset
  1055
qed (auto simp: borel_eq_lessThan reals_Archimedean2)
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1056
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1057
lemma measurable_e2p:
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1058
  "e2p \<in> measurable (borel::'a::ordered_euclidean_space measure) (\<Pi>\<^isub>M i\<in>{..<DIM('a)}. (lborel :: real measure))"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1059
proof (rule measurable_sigma_sets[OF sets_product_borel])
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1060
  fix A :: "(nat \<Rightarrow> real) set" assume "A \<in> {\<Pi>\<^isub>E i\<in>{..<DIM('a)}. {..<x i} |x. True} "
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1061
  then obtain x where  "A = (\<Pi>\<^isub>E i\<in>{..<DIM('a)}. {..<x i})" by auto
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1062
  then have "e2p -` A = {..< (\<chi>\<chi> i. x i) :: 'a}"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1063
    using DIM_positive by (auto simp add: Pi_iff set_eq_iff e2p_def
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1064
      euclidean_eq[where 'a='a] eucl_less[where 'a='a])
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1065
  then show "e2p -` A \<inter> space (borel::'a measure) \<in> sets borel" by simp
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1066
qed (auto simp: e2p_def)
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1067
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1068
lemma measurable_p2e:
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1069
  "p2e \<in> measurable (\<Pi>\<^isub>M i\<in>{..<DIM('a)}. (lborel :: real measure))
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1070
    (borel :: 'a::ordered_euclidean_space measure)"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1071
  (is "p2e \<in> measurable ?P _")
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1072
proof (safe intro!: borel_measurable_iff_halfspace_le[THEN iffD2])
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1073
  fix x i
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1074
  let ?A = "{w \<in> space ?P. (p2e w :: 'a) $$ i \<le> x}"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1075
  assume "i < DIM('a)"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1076
  then have "?A = (\<Pi>\<^isub>E j\<in>{..<DIM('a)}. if i = j then {.. x} else UNIV)"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1077
    using DIM_positive by (auto simp: space_PiM p2e_def split: split_if_asm)
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1078
  then show "?A \<in> sets ?P"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1079
    by auto
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1080
qed
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1081
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1082
lemma lborel_eq_lborel_space:
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1083
  "(lborel :: 'a measure) = distr (\<Pi>\<^isub>M i\<in>{..<DIM('a::ordered_euclidean_space)}. lborel) borel p2e"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1084
  (is "?B = ?D")
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1085
proof (rule lborel_eqI)
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1086
  show "sets ?D = sets borel" by simp
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1087
  let ?P = "(\<Pi>\<^isub>M i\<in>{..<DIM('a::ordered_euclidean_space)}. lborel)"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1088
  fix a b :: 'a
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1089
  have *: "p2e -` {a .. b} \<inter> space ?P = (\<Pi>\<^isub>E i\<in>{..<DIM('a)}. {a $$ i .. b $$ i})"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1090
    by (auto simp: Pi_iff eucl_le[where 'a='a] p2e_def space_PiM)
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1091
  have "emeasure ?P (p2e -` {a..b} \<inter> space ?P) = content {a..b}"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1092
  proof cases
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1093
    assume "{a..b} \<noteq> {}"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1094
    then have "a \<le> b"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1095
      by (simp add: interval_ne_empty eucl_le[where 'a='a])
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1096
    then have "emeasure lborel {a..b} = (\<Prod>x<DIM('a). emeasure lborel {a $$ x .. b $$ x})"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1097
      by (auto simp: content_closed_interval eucl_le[where 'a='a]
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1098
               intro!: setprod_ereal[symmetric])
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1099
    also have "\<dots> = emeasure ?P (p2e -` {a..b} \<inter> space ?P)"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1100
      unfolding * by (subst lborel_space.measure_times) auto
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1101
    finally show ?thesis by simp
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1102
  qed simp
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1103
  then show "emeasure ?D {a .. b} = content {a .. b}"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1104
    by (simp add: emeasure_distr measurable_p2e)
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1105
qed
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1106
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1107
lemma borel_fubini_positiv_integral:
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1108
  fixes f :: "'a::ordered_euclidean_space \<Rightarrow> ereal"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1109
  assumes f: "f \<in> borel_measurable borel"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1110
  shows "integral\<^isup>P lborel f = \<integral>\<^isup>+x. f (p2e x) \<partial>(\<Pi>\<^isub>M i\<in>{..<DIM('a)}. lborel)"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1111
  by (subst lborel_eq_lborel_space) (simp add: positive_integral_distr measurable_p2e f)
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1112
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1113
lemma borel_fubini_integrable:
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1114
  fixes f :: "'a::ordered_euclidean_space \<Rightarrow> real"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1115
  shows "integrable lborel f \<longleftrightarrow>
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1116
    integrable (\<Pi>\<^isub>M i\<in>{..<DIM('a)}. lborel) (\<lambda>x. f (p2e x))"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1117
    (is "_ \<longleftrightarrow> integrable ?B ?f")
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1118
proof
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1119
  assume "integrable lborel f"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1120
  moreover then have f: "f \<in> borel_measurable borel"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1121
    by auto
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1122
  moreover with measurable_p2e
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1123
  have "f \<circ> p2e \<in> borel_measurable ?B"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1124
    by (rule measurable_comp)
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1125
  ultimately show "integrable ?B ?f"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1126
    by (simp add: comp_def borel_fubini_positiv_integral integrable_def)
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1127
next
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1128
  assume "integrable ?B ?f"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1129
  moreover
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1130
  then have "?f \<circ> e2p \<in> borel_measurable (borel::'a measure)"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1131
    by (auto intro!: measurable_e2p)
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1132
  then have "f \<in> borel_measurable borel"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1133
    by (simp cong: measurable_cong)
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1134
  ultimately show "integrable lborel f"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1135
    by (simp add: borel_fubini_positiv_integral integrable_def)
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1136
qed
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1137
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1138
lemma borel_fubini:
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1139
  fixes f :: "'a::ordered_euclidean_space \<Rightarrow> real"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1140
  assumes f: "f \<in> borel_measurable borel"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1141
  shows "integral\<^isup>L lborel f = \<integral>x. f (p2e x) \<partial>((\<Pi>\<^isub>M i\<in>{..<DIM('a)}. lborel))"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1142
  using f by (simp add: borel_fubini_positiv_integral lebesgue_integral_def)
47757
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
  1143
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1144
end