src/HOL/Probability/Lebesgue_Measure.thy
author hoelzl
Fri, 30 May 2014 15:56:30 +0200
changeset 57137 f174712d0a84
parent 56996 891e992e510f
child 57138 7b3146180291
permissions -rw-r--r--
better support for restrict_space
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
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
     9
  imports Finite_Product_Measure Bochner_Integration
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    10
begin
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    11
50104
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50003
diff changeset
    12
lemma absolutely_integrable_on_indicator[simp]:
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50003
diff changeset
    13
  fixes A :: "'a::ordered_euclidean_space set"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50003
diff changeset
    14
  shows "((indicator A :: _ \<Rightarrow> real) absolutely_integrable_on X) \<longleftrightarrow>
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50003
diff changeset
    15
    (indicator A :: _ \<Rightarrow> real) integrable_on X"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50003
diff changeset
    16
  unfolding absolutely_integrable_on_def by simp
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
    17
50104
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50003
diff changeset
    18
lemma has_integral_indicator_UNIV:
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50003
diff changeset
    19
  fixes s A :: "'a::ordered_euclidean_space set" and x :: real
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50003
diff changeset
    20
  shows "((indicator (s \<inter> A) :: 'a\<Rightarrow>real) has_integral x) UNIV = ((indicator s :: _\<Rightarrow>real) has_integral x) A"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50003
diff changeset
    21
proof -
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50003
diff changeset
    22
  have "(\<lambda>x. if x \<in> A then indicator s x else 0) = (indicator (s \<inter> A) :: _\<Rightarrow>real)"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50003
diff changeset
    23
    by (auto simp: fun_eq_iff indicator_def)
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50003
diff changeset
    24
  then show ?thesis
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50003
diff changeset
    25
    unfolding has_integral_restrict_univ[where s=A, symmetric] by simp
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50003
diff changeset
    26
qed
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50003
diff changeset
    27
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50003
diff changeset
    28
lemma
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50003
diff changeset
    29
  fixes s a :: "'a::ordered_euclidean_space set"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50003
diff changeset
    30
  shows integral_indicator_UNIV:
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50003
diff changeset
    31
    "integral UNIV (indicator (s \<inter> A) :: 'a\<Rightarrow>real) = integral A (indicator s :: _\<Rightarrow>real)"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50003
diff changeset
    32
  and integrable_indicator_UNIV:
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50003
diff changeset
    33
    "(indicator (s \<inter> A) :: 'a\<Rightarrow>real) integrable_on UNIV \<longleftrightarrow> (indicator s :: 'a\<Rightarrow>real) integrable_on A"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50003
diff changeset
    34
  unfolding integral_def integrable_on_def has_integral_indicator_UNIV by auto
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
    35
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    36
subsection {* Standard Cubes *}
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    37
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
    38
definition cube :: "nat \<Rightarrow> 'a::ordered_euclidean_space set" where
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50418
diff changeset
    39
  "cube n \<equiv> {\<Sum>i\<in>Basis. - n *\<^sub>R i .. \<Sum>i\<in>Basis. n *\<^sub>R i}"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
    40
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
    41
lemma borel_cube[intro]: "cube n \<in> sets borel"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
    42
  unfolding cube_def by auto
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
    43
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
    44
lemma cube_closed[intro]: "closed (cube n)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
    45
  unfolding cube_def by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
    46
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
    47
lemma cube_subset[intro]: "n \<le> N \<Longrightarrow> cube n \<subseteq> cube N"
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50418
diff changeset
    48
  by (fastforce simp: eucl_le[where 'a='a] cube_def setsum_negf)
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50418
diff changeset
    49
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50418
diff changeset
    50
lemma cube_subset_iff: "cube n \<subseteq> cube N \<longleftrightarrow> n \<le> N"
56188
0268784f60da use cbox to relax class constraints
immler
parents: 56181
diff changeset
    51
  unfolding cube_def subset_box by (simp add: setsum_negf ex_in_conv eucl_le[where 'a='a])
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    52
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50418
diff changeset
    53
lemma ball_subset_cube: "ball (0::'a::ordered_euclidean_space) (real n) \<subseteq> cube n"
56188
0268784f60da use cbox to relax class constraints
immler
parents: 56181
diff changeset
    54
  apply (simp add: cube_def subset_eq mem_box setsum_negf eucl_le[where 'a='a])
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50418
diff changeset
    55
proof safe
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50418
diff changeset
    56
  fix x i :: 'a assume x: "x \<in> ball 0 (real n)" and i: "i \<in> Basis" 
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50418
diff changeset
    57
  thus "- real n \<le> x \<bullet> i" "real n \<ge> x \<bullet> i"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50418
diff changeset
    58
    using Basis_le_norm[OF i, of x] by(auto simp: dist_norm)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    59
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    60
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    61
lemma mem_big_cube: obtains n where "x \<in> cube n"
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50418
diff changeset
    62
proof -
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50418
diff changeset
    63
  from reals_Archimedean2[of "norm x"] guess n ..
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50418
diff changeset
    64
  with ball_subset_cube[unfolded subset_eq, of n]
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50418
diff changeset
    65
  show ?thesis
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50418
diff changeset
    66
    by (intro that[where n=n]) (auto simp add: dist_norm)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    67
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    68
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
    69
lemma cube_subset_Suc[intro]: "cube n \<subseteq> cube (Suc n)"
56188
0268784f60da use cbox to relax class constraints
immler
parents: 56181
diff changeset
    70
  unfolding cube_def cbox_interval[symmetric] subset_box by (simp add: setsum_negf)
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
    71
50104
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50003
diff changeset
    72
lemma has_integral_interval_cube:
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50003
diff changeset
    73
  fixes a b :: "'a::ordered_euclidean_space"
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50418
diff changeset
    74
  shows "(indicator {a .. b} has_integral content ({a .. b} \<inter> cube n)) (cube n)"
50104
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50003
diff changeset
    75
    (is "(?I has_integral content ?R) (cube n)")
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50003
diff changeset
    76
proof -
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50003
diff changeset
    77
  have [simp]: "(\<lambda>x. if x \<in> cube n then ?I x else 0) = indicator ?R"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50003
diff changeset
    78
    by (auto simp: indicator_def cube_def fun_eq_iff eucl_le[where 'a='a])
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50003
diff changeset
    79
  have "(?I has_integral content ?R) (cube n) \<longleftrightarrow> (indicator ?R has_integral content ?R) UNIV"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50003
diff changeset
    80
    unfolding has_integral_restrict_univ[where s="cube n", symmetric] by simp
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50418
diff changeset
    81
  also have "\<dots> \<longleftrightarrow> ((\<lambda>x. 1::real) has_integral content ?R *\<^sub>R 1) ?R"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50418
diff changeset
    82
    unfolding indicator_def [abs_def] has_integral_restrict_univ real_scaleR_def mult_1_right ..
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50418
diff changeset
    83
  also have "((\<lambda>x. 1) has_integral content ?R *\<^sub>R 1) ?R"
56188
0268784f60da use cbox to relax class constraints
immler
parents: 56181
diff changeset
    84
    unfolding cube_def inter_interval cbox_interval[symmetric] by (rule has_integral_const)
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50418
diff changeset
    85
  finally show ?thesis .
50104
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50003
diff changeset
    86
qed
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50003
diff changeset
    87
47757
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
    88
subsection {* Lebesgue measure *}
41706
a207a858d1f6 prefer p2e before e2p; use measure_unique_Int_stable_vimage;
hoelzl
parents: 41704
diff changeset
    89
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
    90
definition lebesgue :: "'a::ordered_euclidean_space measure" where
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
    91
  "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
    92
    (\<lambda>A. SUP n. ereal (integral (cube n) (indicator A)))"
41661
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41654
diff changeset
    93
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
    94
lemma space_lebesgue[simp]: "space lebesgue = UNIV"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
    95
  unfolding lebesgue_def by simp
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
    96
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
    97
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
    98
  unfolding lebesgue_def by simp
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
    99
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   100
lemma sigma_algebra_lebesgue:
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   101
  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
   102
  shows "sigma_algebra UNIV leb"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   103
proof (safe intro!: sigma_algebra_iff2[THEN iffD2])
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   104
  fix A assume A: "A \<in> leb"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   105
  moreover have "indicator (UNIV - A) = (\<lambda>x. 1 - indicator A x :: real)"
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   106
    by (auto simp: fun_eq_iff indicator_def)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   107
  ultimately show "UNIV - A \<in> leb"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   108
    using A by (auto intro!: integrable_sub simp: cube_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 n show "{} \<in> leb"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   111
    by (auto simp: cube_def indicator_def[abs_def] leb_def)
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   112
next
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   113
  fix A :: "nat \<Rightarrow> _" assume A: "range A \<subseteq> leb"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   114
  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
   115
  proof (intro dominated_convergence[where g="?g"] ballI allI)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   116
    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
   117
    proof (induct k)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   118
      case (Suc k)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   119
      have *: "(\<Union> i<Suc k. A i) = (\<Union> i<k. A i) \<union> A k"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   120
        unfolding lessThan_Suc UN_insert by auto
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   121
      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
   122
          indicator (\<Union> i<Suc k. A i)" (is "(\<lambda>x. max (?f x) (?g x)) = _")
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   123
        by (auto simp: fun_eq_iff * indicator_def)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   124
      show ?case
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   125
        using absolutely_integrable_max[of ?f "cube n" ?g] A Suc
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   126
        by (simp add: * leb_def subset_eq)
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   127
    qed auto
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   128
  qed (auto intro: LIMSEQ_indicator_UN simp: cube_def)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   129
  then show "(\<Union>i. A i) \<in> leb" by (auto simp: leb_def)
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   130
qed simp
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   131
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   132
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
   133
  unfolding lebesgue_def sigma_algebra.sets_measure_of_eq[OF sigma_algebra_lebesgue] ..
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   134
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   135
lemma lebesgueD: "A \<in> sets lebesgue \<Longrightarrow> (indicator A :: _ \<Rightarrow> real) integrable_on cube n"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   136
  unfolding sets_lebesgue by simp
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   137
47757
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   138
lemma emeasure_lebesgue:
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   139
  assumes "A \<in> sets lebesgue"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   140
  shows "emeasure lebesgue A = (SUP n. ereal (integral (cube n) (indicator A)))"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   141
    (is "_ = ?\<mu> A")
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   142
proof (rule emeasure_measure_of[OF lebesgue_def])
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   143
  have *: "indicator {} = (\<lambda>x. 0 :: real)" by (simp add: fun_eq_iff)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   144
  show "positive (sets lebesgue) ?\<mu>"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   145
  proof (unfold positive_def, intro conjI ballI)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   146
    show "?\<mu> {} = 0" by (simp add: integral_0 *)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   147
    fix A :: "'a set" assume "A \<in> sets lebesgue" then show "0 \<le> ?\<mu> A"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   148
      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
   149
  qed
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   150
next
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   151
  show "countably_additive (sets lebesgue) ?\<mu>"
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   152
  proof (intro countably_additive_def[THEN iffD2] allI impI)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   153
    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
   154
    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
   155
      by (auto dest: lebesgueD)
46731
5302e932d1e5 avoid undeclared variables in let bindings;
wenzelm
parents: 44928
diff changeset
   156
    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
   157
    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
   158
    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
   159
    assume "(\<Union>i. A i) \<in> sets lebesgue"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   160
    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
   161
      by (auto simp: sets_lebesgue)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   162
    show "(\<Sum>n. ?\<mu> (A n)) = ?\<mu> (\<Union>i. A i)"
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   163
    proof (subst suminf_SUP_eq, safe intro!: incseq_SucI) 
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
   164
      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
   165
        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
   166
    next
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
   167
      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
   168
        using rA unfolding lebesgue_def
44928
7ef6505bde7f renamed Complete_Lattices lemmas, removed legacy names
hoelzl
parents: 44890
diff changeset
   169
        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
   170
    next
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
   171
      show "(SUP n. \<Sum>i. ereal (?m n i)) = (SUP n. ereal (?M n UNIV))"
56218
1c3f1f2431f9 elongated INFI and SUPR, to reduced risk of confusing theorems names in the future while still being consistent with INTER and UNION
haftmann
parents: 56193
diff changeset
   172
      proof (intro arg_cong[where f="SUPREMUM UNIV"] ext sums_unique[symmetric] sums_ereal[THEN iffD2] sums_def[THEN iffD2])
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   173
        fix n
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   174
        have "\<And>m. (UNION {..<m} A) \<in> sets lebesgue" using rA by auto
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   175
        from lebesgueD[OF this]
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   176
        have "(\<lambda>m. ?M n {..< m}) ----> ?M n UNIV"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   177
          (is "(\<lambda>m. integral _ (?A m)) ----> ?I")
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   178
          by (intro dominated_convergence(2)[where f="?A" and h="\<lambda>x. 1::real"])
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   179
             (auto intro: LIMSEQ_indicator_UN simp: cube_def)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   180
        moreover
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   181
        { fix m have *: "(\<Sum>x<m. ?m n x) = ?M n {..< m}"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   182
          proof (induct m)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   183
            case (Suc m)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   184
            have "(\<Union>i<m. A i) \<in> sets lebesgue" using rA by auto
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   185
            then have "(indicator (\<Union>i<m. A i) :: _\<Rightarrow>real) integrable_on (cube n)"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   186
              by (auto dest!: lebesgueD)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   187
            moreover
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   188
            have "(\<Union>i<m. A i) \<inter> A m = {}"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   189
              using rA(2)[unfolded disjoint_family_on_def, THEN bspec, of m]
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   190
              by auto
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   191
            then have "\<And>x. indicator (\<Union>i<Suc m. A i) x =
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   192
              indicator (\<Union>i<m. A i) x + (indicator (A m) x :: real)"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   193
              by (auto simp: indicator_add lessThan_Suc ac_simps)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   194
            ultimately show ?case
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   195
              using Suc A by (simp add: Integration.integral_add[symmetric])
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   196
          qed auto }
56193
c726ecfb22b6 cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
hoelzl
parents: 56188
diff changeset
   197
        ultimately show "(\<lambda>m. \<Sum>x<m. ?m n x) ----> ?M n UNIV"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   198
          by (simp add: atLeast0LessThan)
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   199
      qed
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   200
    qed
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   201
  qed
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   202
qed (auto, fact)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   203
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   204
lemma lebesgueI_borel[intro, simp]:
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   205
  fixes s::"'a::ordered_euclidean_space set"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   206
  assumes "s \<in> sets borel" shows "s \<in> sets lebesgue"
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   207
proof -
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   208
  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
   209
    using assms by (simp add: borel_eq_atLeastAtMost)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   210
  also have "\<dots> \<subseteq> sets lebesgue"
50244
de72bbe42190 qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents: 50123
diff changeset
   211
  proof (safe intro!: sets.sigma_sets_subset lebesgueI)
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   212
    fix n :: nat and a b :: 'a
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   213
    show "(indicator {a..b} :: 'a\<Rightarrow>real) integrable_on cube n"
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50418
diff changeset
   214
      unfolding integrable_on_def using has_integral_interval_cube[of a b] by auto
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   215
  qed
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   216
  finally show ?thesis .
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   217
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   218
47757
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   219
lemma borel_measurable_lebesgueI:
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   220
  "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
   221
  unfolding measurable_def by simp
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   222
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   223
lemma lebesgueI_negligible[dest]: fixes s::"'a::ordered_euclidean_space set"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   224
  assumes "negligible s" shows "s \<in> sets lebesgue"
56188
0268784f60da use cbox to relax class constraints
immler
parents: 56181
diff changeset
   225
  using assms by (force simp: cbox_interval[symmetric] cube_def integrable_on_def negligible_def intro!: lebesgueI)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   226
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   227
lemma lmeasure_eq_0:
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   228
  fixes S :: "'a::ordered_euclidean_space set"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   229
  assumes "negligible S" shows "emeasure lebesgue S = 0"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   230
proof -
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   231
  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
   232
    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
   233
    by (intro integral_unique some1_equality ex_ex1I)
56188
0268784f60da use cbox to relax class constraints
immler
parents: 56181
diff changeset
   234
       (auto simp: cube_def negligible_def cbox_interval[symmetric])
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   235
  then show ?thesis
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   236
    using assms by (simp add: emeasure_lebesgue lebesgueI_negligible)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   237
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   238
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   239
lemma lmeasure_iff_LIMSEQ:
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   240
  assumes A: "A \<in> sets lebesgue" and "0 \<le> m"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   241
  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
   242
proof (subst emeasure_lebesgue[OF A], intro SUP_eq_LIMSEQ)
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   243
  show "mono (\<lambda>n. integral (cube n) (indicator A::_=>real))"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   244
    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
   245
qed
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   246
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   247
lemma lmeasure_finite_has_integral:
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   248
  fixes s :: "'a::ordered_euclidean_space set"
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   249
  assumes "s \<in> sets lebesgue" "emeasure lebesgue s = ereal m"
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   250
  shows "(indicator s has_integral m) UNIV"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   251
proof -
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   252
  let ?I = "indicator :: 'a set \<Rightarrow> 'a \<Rightarrow> real"
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   253
  have "0 \<le> m"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   254
    using emeasure_nonneg[of lebesgue s] `emeasure lebesgue s = ereal m` by simp
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   255
  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
   256
  proof (intro monotone_convergence_increasing allI ballI)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   257
    have LIMSEQ: "(\<lambda>n. integral (cube n) (?I s)) ----> m"
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   258
      using assms(2) unfolding lmeasure_iff_LIMSEQ[OF assms(1) `0 \<le> m`] .
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   259
    { fix n have "integral (cube n) (?I s) \<le> m"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   260
        using cube_subset assms
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   261
        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
   262
           (auto dest!: lebesgueD) }
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   263
    moreover
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   264
    { fix n have "0 \<le> integral (cube n) (?I s)"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   265
      using assms by (auto dest!: lebesgueD intro!: Integration.integral_nonneg) }
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   266
    ultimately
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   267
    show "bounded {integral UNIV (?I (s \<inter> cube k)) |k. True}"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   268
      unfolding bounded_def
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   269
      apply (rule_tac exI[of _ 0])
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   270
      apply (rule_tac exI[of _ m])
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   271
      by (auto simp: dist_real_def integral_indicator_UNIV)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   272
    fix k show "?I (s \<inter> cube k) integrable_on UNIV"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   273
      unfolding integrable_indicator_UNIV using assms by (auto dest!: lebesgueD)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   274
    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
   275
      using cube_subset[of k "Suc k"] by (auto simp: indicator_def)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   276
  next
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   277
    fix x :: 'a
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   278
    from mem_big_cube obtain k where k: "x \<in> cube k" .
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   279
    { fix n have "?I (s \<inter> cube (n + k)) x = ?I s x"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   280
      using k cube_subset[of k "n + k"] by (auto simp: indicator_def) }
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   281
    note * = this
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   282
    show "(\<lambda>k. ?I (s \<inter> cube k) x) ----> ?I s x"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   283
      by (rule LIMSEQ_offset[where k=k]) (auto simp: *)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   284
  qed
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   285
  note ** = conjunctD2[OF this]
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   286
  have m: "m = integral UNIV (?I s)"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   287
    apply (intro LIMSEQ_unique[OF _ **(2)])
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   288
    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
   289
  show ?thesis
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   290
    unfolding m by (intro integrable_integral **)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   291
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   292
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   293
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
   294
  shows "(indicator s :: _ \<Rightarrow> real) integrable_on UNIV"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   295
proof (cases "emeasure lebesgue s")
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   296
  case (real m)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   297
  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
   298
  show ?thesis unfolding integrable_on_def by auto
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   299
qed (insert assms emeasure_nonneg[of lebesgue s], auto)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   300
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   301
lemma has_integral_lebesgue: assumes "((indicator s :: _\<Rightarrow>real) has_integral m) UNIV"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   302
  shows "s \<in> sets lebesgue"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   303
proof (intro lebesgueI)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   304
  let ?I = "indicator :: 'a set \<Rightarrow> 'a \<Rightarrow> real"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   305
  fix n show "(?I s) integrable_on cube n" unfolding cube_def
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   306
  proof (intro integrable_on_subinterval)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   307
    show "(?I s) integrable_on UNIV"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   308
      unfolding integrable_on_def using assms by auto
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   309
  qed auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   310
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   311
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   312
lemma has_integral_lmeasure: assumes "((indicator s :: _\<Rightarrow>real) has_integral m) UNIV"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   313
  shows "emeasure lebesgue s = ereal m"
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   314
proof (intro lmeasure_iff_LIMSEQ[THEN iffD2])
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   315
  let ?I = "indicator :: 'a set \<Rightarrow> 'a \<Rightarrow> real"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   316
  show "s \<in> sets lebesgue" using has_integral_lebesgue[OF assms] .
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   317
  show "0 \<le> m" using assms by (rule has_integral_nonneg) auto
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   318
  have "(\<lambda>n. integral UNIV (?I (s \<inter> cube n))) ----> integral UNIV (?I s)"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   319
  proof (intro dominated_convergence(2) ballI)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   320
    show "(?I s) integrable_on UNIV" unfolding integrable_on_def using assms by auto
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   321
    fix n show "?I (s \<inter> cube n) integrable_on UNIV"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   322
      unfolding integrable_indicator_UNIV using `s \<in> sets lebesgue` by (auto dest: lebesgueD)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   323
    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
   324
  next
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   325
    fix x :: 'a
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   326
    from mem_big_cube obtain k where k: "x \<in> cube k" .
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   327
    { fix n have "?I (s \<inter> cube (n + k)) x = ?I s x"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   328
      using k cube_subset[of k "n + k"] by (auto simp: indicator_def) }
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   329
    note * = this
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   330
    show "(\<lambda>k. ?I (s \<inter> cube k) x) ----> ?I s x"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   331
      by (rule LIMSEQ_offset[where k=k]) (auto simp: *)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   332
  qed
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   333
  then show "(\<lambda>n. integral (cube n) (?I s)) ----> m"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   334
    unfolding integral_unique[OF assms] integral_indicator_UNIV by simp
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   335
qed
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   336
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   337
lemma has_integral_iff_lmeasure:
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   338
  "(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
   339
proof
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   340
  assume "(indicator A has_integral m) UNIV"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   341
  with has_integral_lmeasure[OF this] has_integral_lebesgue[OF this]
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   342
  show "A \<in> sets lebesgue \<and> emeasure lebesgue A = ereal m"
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   343
    by (auto intro: has_integral_nonneg)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   344
next
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   345
  assume "A \<in> sets lebesgue \<and> emeasure lebesgue A = ereal m"
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   346
  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
   347
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   348
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   349
lemma lmeasure_eq_integral: assumes "(indicator s::_\<Rightarrow>real) integrable_on UNIV"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   350
  shows "emeasure lebesgue s = ereal (integral UNIV (indicator s))"
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   351
  using assms unfolding integrable_on_def
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   352
proof safe
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   353
  fix y :: real assume "(indicator s has_integral y) UNIV"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   354
  from this[unfolded has_integral_iff_lmeasure] integral_unique[OF this]
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   355
  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
   356
qed
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   357
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   358
lemma lebesgue_simple_function_indicator:
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
   359
  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
   360
  assumes f:"simple_function lebesgue f"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   361
  shows "f = (\<lambda>x. (\<Sum>y \<in> f ` UNIV. y * indicator (f -` {y}) x))"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   362
  by (rule, subst simple_function_indicator_representation[OF f]) auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   363
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   364
lemma integral_eq_lmeasure:
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   365
  "(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
   366
  by (subst lmeasure_eq_integral) (auto intro!: integral_nonneg)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   367
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   368
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
   369
  using lmeasure_eq_integral[OF assms] by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   370
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   371
lemma negligible_iff_lebesgue_null_sets:
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   372
  "negligible A \<longleftrightarrow> A \<in> null_sets lebesgue"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   373
proof
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   374
  assume "negligible A"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   375
  from this[THEN lebesgueI_negligible] this[THEN lmeasure_eq_0]
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   376
  show "A \<in> null_sets lebesgue" by auto
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   377
next
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   378
  assume A: "A \<in> null_sets lebesgue"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   379
  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
   380
    by (auto simp: null_sets_def)
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   381
  show "negligible A" unfolding negligible_def
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   382
  proof (intro allI)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   383
    fix a b :: 'a
56188
0268784f60da use cbox to relax class constraints
immler
parents: 56181
diff changeset
   384
    have integrable: "(indicator A :: _\<Rightarrow>real) integrable_on cbox a b"
0268784f60da use cbox to relax class constraints
immler
parents: 56181
diff changeset
   385
      by (intro integrable_on_subcbox has_integral_integrable) (auto intro: *)
0268784f60da use cbox to relax class constraints
immler
parents: 56181
diff changeset
   386
    then have "integral (cbox a b) (indicator A) \<le> (integral UNIV (indicator A) :: real)"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   387
      using * by (auto intro!: integral_subset_le)
56188
0268784f60da use cbox to relax class constraints
immler
parents: 56181
diff changeset
   388
    moreover have "(0::real) \<le> integral (cbox a b) (indicator A)"
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   389
      using integrable by (auto intro!: integral_nonneg)
56188
0268784f60da use cbox to relax class constraints
immler
parents: 56181
diff changeset
   390
    ultimately have "integral (cbox a b) (indicator A) = (0::real)"
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   391
      using integral_unique[OF *] by auto
56188
0268784f60da use cbox to relax class constraints
immler
parents: 56181
diff changeset
   392
    then show "(indicator A has_integral (0::real)) (cbox a b)"
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   393
      using integrable_integral[OF integrable] by simp
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   394
  qed
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   395
qed
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   396
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   397
lemma lmeasure_UNIV[intro]: "emeasure lebesgue (UNIV::'a::ordered_euclidean_space set) = \<infinity>"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   398
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
   399
  fix n :: nat
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   400
  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
   401
  moreover
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   402
  { 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
   403
    proof (cases n)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   404
      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
   405
    next
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   406
      case (Suc n')
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   407
      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
   408
      also have "(2 * real n)^1 \<le> (2 * real n) ^ DIM('a)"
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50418
diff changeset
   409
        using Suc DIM_positive[where 'a='a] 
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50418
diff changeset
   410
        by (intro power_increasing) (auto simp: real_of_nat_Suc simp del: DIM_positive)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   411
      finally show ?thesis .
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   412
    qed }
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
   413
  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
   414
    using integral_const DIM_positive[where 'a='a]
56188
0268784f60da use cbox to relax class constraints
immler
parents: 56181
diff changeset
   415
    by (auto simp: cube_def content_cbox_cases setprod_constant setsum_negf cbox_interval[symmetric])
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   416
qed simp
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   417
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   418
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
   419
  unfolding negligible_iff_lebesgue_null_sets[symmetric] by (auto simp: negligible_subset)
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   420
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   421
lemma
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   422
  fixes a b ::"'a::ordered_euclidean_space"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   423
  shows lmeasure_atLeastAtMost[simp]: "emeasure lebesgue {a..b} = ereal (content {a..b})"
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   424
proof -
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   425
  have "(indicator (UNIV \<inter> {a..b})::_\<Rightarrow>real) integrable_on UNIV"
56188
0268784f60da use cbox to relax class constraints
immler
parents: 56181
diff changeset
   426
    unfolding integrable_indicator_UNIV by (simp add: integrable_const indicator_def [abs_def] cbox_interval[symmetric])
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   427
  from lmeasure_eq_integral[OF this] show ?thesis unfolding integral_indicator_UNIV
56188
0268784f60da use cbox to relax class constraints
immler
parents: 56181
diff changeset
   428
    by (simp add: indicator_def [abs_def] cbox_interval[symmetric])
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   429
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   430
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   431
lemma lmeasure_singleton[simp]:
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   432
  fixes a :: "'a::ordered_euclidean_space" shows "emeasure lebesgue {a} = 0"
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   433
  using lmeasure_atLeastAtMost[of a a] by simp
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   434
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   435
lemma AE_lebesgue_singleton:
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   436
  fixes a :: "'a::ordered_euclidean_space" shows "AE x in lebesgue. x \<noteq> a"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   437
  by (rule AE_I[where N="{a}"]) auto
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   438
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   439
declare content_real[simp]
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   440
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   441
lemma
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   442
  fixes a b :: real
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   443
  shows lmeasure_real_greaterThanAtMost[simp]:
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   444
    "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
   445
proof -
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   446
  have "emeasure lebesgue {a <.. b} = emeasure lebesgue {a .. b}"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   447
    using AE_lebesgue_singleton[of a]
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   448
    by (intro emeasure_eq_AE) auto
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   449
  then show ?thesis by auto
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   450
qed
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   451
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   452
lemma
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   453
  fixes a b :: real
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   454
  shows lmeasure_real_atLeastLessThan[simp]:
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   455
    "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
   456
proof -
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   457
  have "emeasure lebesgue {a ..< b} = emeasure lebesgue {a .. b}"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   458
    using AE_lebesgue_singleton[of b]
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   459
    by (intro emeasure_eq_AE) auto
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   460
  then show ?thesis by auto
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   461
qed
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   462
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   463
lemma
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   464
  fixes a b :: real
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   465
  shows lmeasure_real_greaterThanLessThan[simp]:
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   466
    "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
   467
proof -
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   468
  have "emeasure lebesgue {a <..< b} = emeasure lebesgue {a .. b}"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   469
    using AE_lebesgue_singleton[of a] AE_lebesgue_singleton[of b]
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   470
    by (intro emeasure_eq_AE) auto
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   471
  then show ?thesis by auto
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   472
qed
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   473
41706
a207a858d1f6 prefer p2e before e2p; use measure_unique_Int_stable_vimage;
hoelzl
parents: 41704
diff changeset
   474
subsection {* Lebesgue-Borel measure *}
a207a858d1f6 prefer p2e before e2p; use measure_unique_Int_stable_vimage;
hoelzl
parents: 41704
diff changeset
   475
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   476
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
   477
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   478
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
   479
  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
   480
  and sets_lborel[simp]: "sets lborel = sets borel"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   481
  and measurable_lborel1[simp]: "measurable lborel = measurable borel"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   482
  and measurable_lborel2[simp]: "measurable A lborel = measurable A borel"
50244
de72bbe42190 qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents: 50123
diff changeset
   483
  using sets.sigma_sets_eq[of borel]
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   484
  by (auto simp add: lborel_def measurable_def[abs_def])
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   485
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   486
(* TODO: switch these rules! *)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   487
lemma emeasure_lborel[simp]: "A \<in> sets borel \<Longrightarrow> emeasure lborel A = emeasure lebesgue A"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   488
  by (rule emeasure_measure_of[OF lborel_def])
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   489
     (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
   490
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   491
lemma measure_lborel[simp]: "A \<in> sets borel \<Longrightarrow> measure lborel A = measure lebesgue A"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   492
  unfolding measure_def by simp
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   493
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   494
interpretation lborel: sigma_finite_measure lborel
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   495
proof (default, intro conjI exI[of _ "\<lambda>n. cube n"])
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   496
  show "range cube \<subseteq> sets lborel" by (auto intro: borel_closed)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   497
  { 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
   498
  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
   499
  show "\<forall>i. emeasure lborel (cube i) \<noteq> \<infinity>" by (simp add: cube_def)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   500
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
   501
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   502
interpretation lebesgue: sigma_finite_measure lebesgue
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   503
proof
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   504
  from lborel.sigma_finite guess A :: "nat \<Rightarrow> 'a set" ..
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   505
  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
   506
    by (intro exI[of _ A]) (auto simp: subset_eq)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   507
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   508
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   509
lemma Int_stable_atLeastAtMost:
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   510
  fixes x::"'a::ordered_euclidean_space"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   511
  shows "Int_stable (range (\<lambda>(a, b::'a). {a..b}))"
56188
0268784f60da use cbox to relax class constraints
immler
parents: 56181
diff changeset
   512
  by (auto simp: inter_interval Int_stable_def cbox_interval[symmetric])
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   513
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   514
lemma lborel_eqI:
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   515
  fixes M :: "'a::ordered_euclidean_space measure"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   516
  assumes emeasure_eq: "\<And>a b. emeasure M {a .. b} = content {a .. b}"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   517
  assumes sets_eq: "sets M = sets borel"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   518
  shows "lborel = M"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   519
proof (rule measure_eqI_generator_eq[OF Int_stable_atLeastAtMost])
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51478
diff changeset
   520
  let ?P = "\<Pi>\<^sub>M i\<in>{..<DIM('a::ordered_euclidean_space)}. lborel"
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   521
  let ?E = "range (\<lambda>(a, b). {a..b} :: 'a set)"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   522
  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
   523
    by (simp_all add: borel_eq_atLeastAtMost sets_eq)
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   524
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   525
  show "range cube \<subseteq> ?E" unfolding cube_def [abs_def] by auto
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   526
  { 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
   527
  then show "(\<Union>i. cube i :: 'a set) = UNIV" by auto
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   528
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   529
  { fix i show "emeasure lborel (cube i) \<noteq> \<infinity>" unfolding cube_def by auto }
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   530
  { fix X assume "X \<in> ?E" then show "emeasure lborel X = emeasure M X"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   531
      by (auto simp: emeasure_eq) }
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   532
qed
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   533
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   534
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   535
(* GENEREALIZE to euclidean_spaces *)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   536
lemma lborel_real_affine:
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   537
  fixes c :: real assumes "c \<noteq> 0"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   538
  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
   539
proof (rule lborel_eqI)
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   540
  fix a b show "emeasure ?D {a..b} = content {a .. b}"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   541
  proof cases
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   542
    assume "0 < c"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   543
    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
   544
      by (auto simp: field_simps)
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   545
    with `0 < c` show ?thesis
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   546
      by (cases "a \<le> b")
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56993
diff changeset
   547
         (auto simp: field_simps emeasure_density nn_integral_distr nn_integral_cmult
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   548
                     borel_measurable_indicator' emeasure_distr)
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   549
  next
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   550
    assume "\<not> 0 < c" with `c \<noteq> 0` have "c < 0" by auto
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   551
    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
   552
      by (auto simp: field_simps)
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   553
    with `c < 0` show ?thesis
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   554
      by (cases "a \<le> b")
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56993
diff changeset
   555
         (auto simp: field_simps emeasure_density nn_integral_distr
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56993
diff changeset
   556
                     nn_integral_cmult borel_measurable_indicator' emeasure_distr)
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   557
  qed
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   558
qed simp
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   559
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56993
diff changeset
   560
lemma nn_integral_real_affine:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   561
  fixes c :: real assumes [measurable]: "f \<in> borel_measurable borel" and c: "c \<noteq> 0"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   562
  shows "(\<integral>\<^sup>+x. f x \<partial>lborel) = \<bar>c\<bar> * (\<integral>\<^sup>+x. f (t + c * x) \<partial>lborel)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   563
  by (subst lborel_real_affine[OF c, of t])
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56993
diff changeset
   564
     (simp add: nn_integral_density nn_integral_distr nn_integral_cmult)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   565
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   566
lemma lborel_integrable_real_affine:
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   567
  fixes f :: "real \<Rightarrow> _ :: {banach, second_countable_topology}"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   568
  assumes f: "integrable lborel f"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   569
  shows "c \<noteq> 0 \<Longrightarrow> integrable lborel (\<lambda>x. f (t + c * x))"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   570
  using f f[THEN borel_measurable_integrable] unfolding integrable_iff_bounded
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56993
diff changeset
   571
  by (subst (asm) nn_integral_real_affine[where c=c and t=t]) auto
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   572
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   573
lemma lborel_integrable_real_affine_iff:
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   574
  fixes f :: "real \<Rightarrow> 'a :: {banach, second_countable_topology}"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   575
  shows "c \<noteq> 0 \<Longrightarrow> integrable lborel (\<lambda>x. f (t + c * x)) \<longleftrightarrow> integrable lborel f"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   576
  using
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   577
    lborel_integrable_real_affine[of f c t]
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   578
    lborel_integrable_real_affine[of "\<lambda>x. f (t + c * x)" "1/c" "-t/c"]
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   579
  by (auto simp add: field_simps)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   580
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   581
lemma lborel_integral_real_affine:
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   582
  fixes f :: "real \<Rightarrow> 'a :: {banach, second_countable_topology}" and c :: real
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   583
  assumes c: "c \<noteq> 0" and f[measurable]: "integrable lborel f"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   584
  shows "(\<integral>x. f x \<partial> lborel) = \<bar>c\<bar> *\<^sub>R (\<integral>x. f (t + c * x) \<partial>lborel)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   585
  using c f f[THEN borel_measurable_integrable] f[THEN lborel_integrable_real_affine, of c t]
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   586
  by (subst lborel_real_affine[OF c, of t]) (simp add: integral_density integral_distr)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   587
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   588
lemma divideR_right: 
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   589
  fixes x y :: "'a::real_normed_vector"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   590
  shows "r \<noteq> 0 \<Longrightarrow> y = x /\<^sub>R r \<longleftrightarrow> r *\<^sub>R y = x"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   591
  using scaleR_cancel_left[of r y "x /\<^sub>R r"] by simp
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   592
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   593
lemma integrable_on_cmult_iff2:
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   594
  fixes c :: real
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   595
  shows "(\<lambda>x. c * f x) integrable_on s \<longleftrightarrow> c = 0 \<or> f integrable_on s"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   596
  using integrable_cmul[of "\<lambda>x. c * f x" s "1 / c"] integrable_cmul[of f s c]
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   597
  by (cases "c = 0") auto
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   598
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   599
lemma lborel_has_bochner_integral_real_affine_iff:
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   600
  fixes x :: "'a :: {banach, second_countable_topology}"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   601
  shows "c \<noteq> 0 \<Longrightarrow>
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   602
    has_bochner_integral lborel f x \<longleftrightarrow>
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   603
    has_bochner_integral lborel (\<lambda>x. f (t + c * x)) (x /\<^sub>R \<bar>c\<bar>)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   604
  unfolding has_bochner_integral_iff lborel_integrable_real_affine_iff
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   605
  by (simp_all add: lborel_integral_real_affine[symmetric] divideR_right cong: conj_cong)
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   606
41706
a207a858d1f6 prefer p2e before e2p; use measure_unique_Int_stable_vimage;
hoelzl
parents: 41704
diff changeset
   607
subsection {* Lebesgue integrable implies Gauge integrable *}
a207a858d1f6 prefer p2e before e2p; use measure_unique_Int_stable_vimage;
hoelzl
parents: 41704
diff changeset
   608
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   609
lemma has_integral_scaleR_left: 
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   610
  "(f has_integral y) s \<Longrightarrow> ((\<lambda>x. f x *\<^sub>R c) has_integral (y *\<^sub>R c)) s"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   611
  using has_integral_linear[OF _ bounded_linear_scaleR_left] by (simp add: comp_def)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   612
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   613
lemma has_integral_mult_left:
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   614
  fixes c :: "_ :: {real_normed_algebra}"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   615
  shows "(f has_integral y) s \<Longrightarrow> ((\<lambda>x. f x * c) has_integral (y * c)) s"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   616
  using has_integral_linear[OF _ bounded_linear_mult_left] by (simp add: comp_def)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   617
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   618
(* GENERALIZE Integration.dominated_convergence, then generalize the following theorems *)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   619
lemma has_integral_dominated_convergence:
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   620
  fixes f :: "nat \<Rightarrow> 'n::ordered_euclidean_space \<Rightarrow> real"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   621
  assumes "\<And>k. (f k has_integral y k) s" "h integrable_on s"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   622
    "\<And>k. \<forall>x\<in>s. norm (f k x) \<le> h x" "\<forall>x\<in>s. (\<lambda>k. f k x) ----> g x"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   623
    and x: "y ----> x"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   624
  shows "(g has_integral x) s"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   625
proof -
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   626
  have int_f: "\<And>k. (f k) integrable_on s"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   627
    using assms by (auto simp: integrable_on_def)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   628
  have "(g has_integral (integral s g)) s"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   629
    by (intro integrable_integral dominated_convergence[OF int_f assms(2)]) fact+
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   630
  moreover have "integral s g = x"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   631
  proof (rule LIMSEQ_unique)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   632
    show "(\<lambda>i. integral s (f i)) ----> x"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   633
      using integral_unique[OF assms(1)] x by simp
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   634
    show "(\<lambda>i. integral s (f i)) ----> integral s g"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   635
      by (intro dominated_convergence[OF int_f assms(2)]) fact+
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   636
  qed
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   637
  ultimately show ?thesis
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   638
    by simp
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   639
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   640
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56993
diff changeset
   641
lemma nn_integral_has_integral:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   642
  fixes f::"'a::ordered_euclidean_space \<Rightarrow> real"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   643
  assumes f: "f \<in> borel_measurable lebesgue" "\<And>x. 0 \<le> f x" "(\<integral>\<^sup>+x. f x \<partial>lebesgue) = ereal r"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   644
  shows "(f has_integral r) UNIV"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   645
using f proof (induct arbitrary: r rule: borel_measurable_induct_real)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   646
  case (set A) then show ?case
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   647
    by (auto simp add: ereal_indicator has_integral_iff_lmeasure)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   648
next
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   649
  case (mult g c)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   650
  then have "ereal c * (\<integral>\<^sup>+ x. g x \<partial>lebesgue) = ereal r"
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56993
diff changeset
   651
    by (subst nn_integral_cmult[symmetric]) auto
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   652
  then obtain r' where "(c = 0 \<and> r = 0) \<or> ((\<integral>\<^sup>+ x. ereal (g x) \<partial>lebesgue) = ereal r' \<and> r = c * r')"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   653
    by (cases "\<integral>\<^sup>+ x. ereal (g x) \<partial>lebesgue") (auto split: split_if_asm)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   654
  with mult show ?case
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   655
    by (auto intro!: has_integral_cmult_real)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   656
next
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   657
  case (add g h)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   658
  moreover
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   659
  then have "(\<integral>\<^sup>+ x. h x + g x \<partial>lebesgue) = (\<integral>\<^sup>+ x. h x \<partial>lebesgue) + (\<integral>\<^sup>+ x. g x \<partial>lebesgue)"
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56993
diff changeset
   660
    unfolding plus_ereal.simps[symmetric] by (subst nn_integral_add) auto
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   661
  with add obtain a b where "(\<integral>\<^sup>+ x. h x \<partial>lebesgue) = ereal a" "(\<integral>\<^sup>+ x. g x \<partial>lebesgue) = ereal b" "r = a + b"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   662
    by (cases "\<integral>\<^sup>+ x. h x \<partial>lebesgue" "\<integral>\<^sup>+ x. g x \<partial>lebesgue" rule: ereal2_cases) auto
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   663
  ultimately show ?case
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   664
    by (auto intro!: has_integral_add)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   665
next
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   666
  case (seq U)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   667
  note seq(1)[measurable] and f[measurable]
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   668
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   669
  { fix i x 
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   670
    have "U i x \<le> f x"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   671
      using seq(5)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   672
      apply (rule LIMSEQ_le_const)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   673
      using seq(4)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   674
      apply (auto intro!: exI[of _ i] simp: incseq_def le_fun_def)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   675
      done }
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   676
  note U_le_f = this
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   677
  
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   678
  { fix i
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   679
    have "(\<integral>\<^sup>+x. ereal (U i x) \<partial>lebesgue) \<le> (\<integral>\<^sup>+x. ereal (f x) \<partial>lebesgue)"
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56993
diff changeset
   680
      using U_le_f by (intro nn_integral_mono) simp
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   681
    then obtain p where "(\<integral>\<^sup>+x. U i x \<partial>lebesgue) = ereal p" "p \<le> r"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   682
      using seq(6) by (cases "\<integral>\<^sup>+x. U i x \<partial>lebesgue") auto
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   683
    moreover then have "0 \<le> p"
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56993
diff changeset
   684
      by (metis ereal_less_eq(5) nn_integral_nonneg)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   685
    moreover note seq
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   686
    ultimately have "\<exists>p. (\<integral>\<^sup>+x. U i x \<partial>lebesgue) = ereal p \<and> 0 \<le> p \<and> p \<le> r \<and> (U i has_integral p) UNIV"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   687
      by auto }
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   688
  then obtain p where p: "\<And>i. (\<integral>\<^sup>+x. ereal (U i x) \<partial>lebesgue) = ereal (p i)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   689
    and bnd: "\<And>i. p i \<le> r" "\<And>i. 0 \<le> p i"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   690
    and U_int: "\<And>i.(U i has_integral (p i)) UNIV" by metis
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   691
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   692
  have int_eq: "\<And>i. integral UNIV (U i) = p i" using U_int by (rule integral_unique)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   693
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   694
  have *: "f integrable_on UNIV \<and> (\<lambda>k. integral UNIV (U k)) ----> integral UNIV f"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   695
  proof (rule monotone_convergence_increasing)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   696
    show "\<forall>k. U k integrable_on UNIV" using U_int by auto
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   697
    show "\<forall>k. \<forall>x\<in>UNIV. U k x \<le> U (Suc k) x" using `incseq U` by (auto simp: incseq_def le_fun_def)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   698
    then show "bounded {integral UNIV (U k) |k. True}"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   699
      using bnd int_eq by (auto simp: bounded_real intro!: exI[of _ r])
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   700
    show "\<forall>x\<in>UNIV. (\<lambda>k. U k x) ----> f x"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   701
      using seq by auto
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   702
  qed
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   703
  moreover have "(\<lambda>i. (\<integral>\<^sup>+x. U i x \<partial>lebesgue)) ----> (\<integral>\<^sup>+x. f x \<partial>lebesgue)"
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56993
diff changeset
   704
    using seq U_le_f by (intro nn_integral_dominated_convergence[where w=f]) auto
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   705
  ultimately have "integral UNIV f = r"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   706
    by (auto simp add: int_eq p seq intro: LIMSEQ_unique)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   707
  with * show ?case
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   708
    by (simp add: has_integral_integral)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   709
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   710
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   711
lemma has_integral_integrable_lebesgue_nonneg:
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   712
  fixes f::"'a::ordered_euclidean_space \<Rightarrow> real"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   713
  assumes f: "integrable lebesgue f" "\<And>x. 0 \<le> f x"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   714
  shows "(f has_integral integral\<^sup>L lebesgue f) UNIV"
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56993
diff changeset
   715
proof (rule nn_integral_has_integral)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   716
  show "(\<integral>\<^sup>+ x. ereal (f x) \<partial>lebesgue) = ereal (integral\<^sup>L lebesgue f)"
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56993
diff changeset
   717
    using f by (intro nn_integral_eq_integral) auto
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   718
qed (insert f, auto)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   719
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   720
lemma has_integral_lebesgue_integral_lebesgue:
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   721
  fixes f::"'a::ordered_euclidean_space \<Rightarrow> real"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   722
  assumes f: "integrable lebesgue f"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51478
diff changeset
   723
  shows "(f has_integral (integral\<^sup>L lebesgue f)) UNIV"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   724
using f proof induct
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   725
  case (base A c) then show ?case
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   726
    by (auto intro!: has_integral_mult_left simp: has_integral_iff_lmeasure)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   727
       (simp add: emeasure_eq_ereal_measure)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   728
next
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   729
  case (add f g) then show ?case
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   730
    by (auto intro!: has_integral_add)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   731
next
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   732
  case (lim f s)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   733
  show ?case
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   734
  proof (rule has_integral_dominated_convergence)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   735
    show "\<And>i. (s i has_integral integral\<^sup>L lebesgue (s i)) UNIV" by fact
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   736
    show "(\<lambda>x. norm (2 * f x)) integrable_on UNIV"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   737
      using lim by (intro has_integral_integrable[OF has_integral_integrable_lebesgue_nonneg]) auto
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   738
    show "\<And>k. \<forall>x\<in>UNIV. norm (s k x) \<le> norm (2 * f x)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   739
      using lim by (auto simp add: abs_mult)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   740
    show "\<forall>x\<in>UNIV. (\<lambda>k. s k x) ----> f x"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   741
      using lim by auto
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   742
    show "(\<lambda>k. integral\<^sup>L lebesgue (s k)) ----> integral\<^sup>L lebesgue f"
57137
f174712d0a84 better support for restrict_space
hoelzl
parents: 56996
diff changeset
   743
      using lim by (intro integral_dominated_convergence[where w="\<lambda>x. 2 * norm (f x)"]) auto
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   744
  qed
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   745
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   746
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56993
diff changeset
   747
lemma lebesgue_nn_integral_eq_borel:
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   748
  assumes f: "f \<in> borel_measurable borel"
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56993
diff changeset
   749
  shows "integral\<^sup>N lebesgue f = integral\<^sup>N lborel f"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   750
proof -
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56993
diff changeset
   751
  from f have "integral\<^sup>N lebesgue (\<lambda>x. max 0 (f x)) = integral\<^sup>N lborel (\<lambda>x. max 0 (f x))"
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56993
diff changeset
   752
    by (auto intro!: nn_integral_subalgebra[symmetric])
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56993
diff changeset
   753
  then show ?thesis unfolding nn_integral_max_0 .
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   754
qed
41546
2a12c23b7a34 integral on lebesgue measure is extension of integral on borel measure
hoelzl
parents: 41097
diff changeset
   755
2a12c23b7a34 integral on lebesgue measure is extension of integral on borel measure
hoelzl
parents: 41097
diff changeset
   756
lemma lebesgue_integral_eq_borel:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   757
  fixes f :: "_ \<Rightarrow> _::{banach, second_countable_topology}"
41546
2a12c23b7a34 integral on lebesgue measure is extension of integral on borel measure
hoelzl
parents: 41097
diff changeset
   758
  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
   759
  shows "integrable lebesgue f \<longleftrightarrow> integrable lborel f" (is ?P)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51478
diff changeset
   760
    and "integral\<^sup>L lebesgue f = integral\<^sup>L lborel f" (is ?I)
41546
2a12c23b7a34 integral on lebesgue measure is extension of integral on borel measure
hoelzl
parents: 41097
diff changeset
   761
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
   762
  have "sets lborel \<subseteq> sets lebesgue" by auto
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   763
  from integral_subalgebra[of f lborel, OF _ this _ _]
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   764
       integrable_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
   765
  show ?P ?I by auto
2a12c23b7a34 integral on lebesgue measure is extension of integral on borel measure
hoelzl
parents: 41097
diff changeset
   766
qed
2a12c23b7a34 integral on lebesgue measure is extension of integral on borel measure
hoelzl
parents: 41097
diff changeset
   767
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   768
lemma has_integral_lebesgue_integral:
41546
2a12c23b7a34 integral on lebesgue measure is extension of integral on borel measure
hoelzl
parents: 41097
diff changeset
   769
  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
   770
  assumes f:"integrable lborel f"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51478
diff changeset
   771
  shows "(f has_integral (integral\<^sup>L lborel f)) UNIV"
41546
2a12c23b7a34 integral on lebesgue measure is extension of integral on borel measure
hoelzl
parents: 41097
diff changeset
   772
proof -
2a12c23b7a34 integral on lebesgue measure is extension of integral on borel measure
hoelzl
parents: 41097
diff changeset
   773
  have borel: "f \<in> borel_measurable borel"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   774
    using f unfolding integrable_iff_bounded by auto
41546
2a12c23b7a34 integral on lebesgue measure is extension of integral on borel measure
hoelzl
parents: 41097
diff changeset
   775
  from f show ?thesis
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   776
    using has_integral_lebesgue_integral_lebesgue[of f]
41546
2a12c23b7a34 integral on lebesgue measure is extension of integral on borel measure
hoelzl
parents: 41097
diff changeset
   777
    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
   778
qed
2a12c23b7a34 integral on lebesgue measure is extension of integral on borel measure
hoelzl
parents: 41097
diff changeset
   779
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56993
diff changeset
   780
lemma nn_integral_bound_simple_function:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   781
  assumes bnd: "\<And>x. x \<in> space M \<Longrightarrow> 0 \<le> f x" "\<And>x. x \<in> space M \<Longrightarrow> f x < \<infinity>"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   782
  assumes f[measurable]: "simple_function M f"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   783
  assumes supp: "emeasure M {x\<in>space M. f x \<noteq> 0} < \<infinity>"
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56993
diff changeset
   784
  shows "nn_integral M f < \<infinity>"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   785
proof cases
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   786
  assume "space M = {}"
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56993
diff changeset
   787
  then have "nn_integral M f = (\<integral>\<^sup>+x. 0 \<partial>M)"
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56993
diff changeset
   788
    by (intro nn_integral_cong) auto
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   789
  then show ?thesis by simp
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   790
next
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   791
  assume "space M \<noteq> {}"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   792
  with simple_functionD(1)[OF f] bnd have bnd: "0 \<le> Max (f`space M) \<and> Max (f`space M) < \<infinity>"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   793
    by (subst Max_less_iff) (auto simp: Max_ge_iff)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   794
  
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56993
diff changeset
   795
  have "nn_integral M f \<le> (\<integral>\<^sup>+x. Max (f`space M) * indicator {x\<in>space M. f x \<noteq> 0} x \<partial>M)"
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56993
diff changeset
   796
  proof (rule nn_integral_mono)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   797
    fix x assume "x \<in> space M"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   798
    with f show "f x \<le> Max (f ` space M) * indicator {x \<in> space M. f x \<noteq> 0} x"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   799
      by (auto split: split_indicator intro!: Max_ge simple_functionD)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   800
  qed
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   801
  also have "\<dots> < \<infinity>"
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56993
diff changeset
   802
    using bnd supp by (subst nn_integral_cmult) auto
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   803
  finally show ?thesis .
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   804
qed
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   805
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   806
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   807
lemma
47757
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   808
  fixes f :: "'a::ordered_euclidean_space \<Rightarrow> real"
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   809
  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
   810
  assumes I: "(f has_integral I) UNIV"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   811
  shows integrable_has_integral_lebesgue_nonneg: "integrable lebesgue f"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   812
    and integral_has_integral_lebesgue_nonneg: "integral\<^sup>L lebesgue f = I"
47757
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   813
proof -
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   814
  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
   815
  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
   816
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56993
diff changeset
   817
  have "(\<integral>\<^sup>+ x. ereal (f x) \<partial>lebesgue) = (SUP i. integral\<^sup>N lebesgue (F i))"
47757
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   818
    using F
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56993
diff changeset
   819
    by (subst nn_integral_monotone_convergence_SUP[symmetric])
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56993
diff changeset
   820
       (simp_all add: nn_integral_max_0 borel_measurable_simple_function)
47757
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   821
  also have "\<dots> \<le> ereal I"
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   822
  proof (rule SUP_least)
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   823
    fix i :: nat
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   824
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   825
    { fix z
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   826
      from F(4)[of z] have "F i z \<le> ereal (f z)"
54863
82acc20ded73 prefer more canonical names for lemmas on min/max
haftmann
parents: 54775
diff changeset
   827
        by (metis SUP_upper UNIV_I ereal_max_0 max.absorb2 nonneg)
47757
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   828
      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
   829
        by (cases "F i z") simp_all }
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   830
    note F_bound = this
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   831
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   832
    { 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
   833
      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
   834
        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
   835
      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
   836
      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
   837
      proof (rule dominated_convergence(1))
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   838
        fix n :: nat
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   839
        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
   840
          using x F(1)[of i]
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   841
          by (intro lebesgueD) (auto simp: simple_function_def)
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   842
        then have cube: "?s n integrable_on cube n"
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   843
          by (simp add: integrable_on_cmult_iff)
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   844
        show "?s n integrable_on UNIV"
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   845
          by (rule integrable_on_superset[OF _ _ cube]) auto
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   846
      next
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   847
        show "f integrable_on UNIV"
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   848
          unfolding integrable_on_def using I by auto
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   849
      next
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   850
        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
   851
          using nonneg F(5) by (auto split: split_indicator)
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   852
      next
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   853
        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
   854
        proof
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   855
          fix z :: 'a
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   856
          from mem_big_cube[of z] guess j .
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   857
          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
   858
            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
   859
          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
   860
            by (rule Lim_eventually)
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   861
        qed
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   862
      qed
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   863
      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
   864
        by (simp add: integrable_on_cmult_iff) }
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   865
    note F_finite = lmeasure_finite[OF this]
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   866
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   867
    have F_eq: "\<And>x. F i x = ereal (norm (real (F i x)))"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   868
      using F(3,5) by (auto simp: fun_eq_iff ereal_real image_iff eq_commute)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   869
    have F_eq2: "\<And>x. F i x = ereal (real (F i x))"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   870
      using F(3,5) by (auto simp: fun_eq_iff ereal_real image_iff eq_commute)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   871
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   872
    have int: "integrable lebesgue (\<lambda>x. real (F i x))"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   873
      unfolding integrable_iff_bounded
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   874
    proof
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   875
      have "(\<integral>\<^sup>+x. F i x \<partial>lebesgue) < \<infinity>"
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56993
diff changeset
   876
      proof (rule nn_integral_bound_simple_function)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   877
        fix x::'a assume "x \<in> space lebesgue" then show "0 \<le> F i x" "F i x < \<infinity>"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   878
          using F by (auto simp: image_iff eq_commute)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   879
      next
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   880
        have eq: "{x \<in> space lebesgue. F i x \<noteq> 0} = (\<Union>x\<in>F i ` space lebesgue - {0}. F i -` {x} \<inter> space lebesgue)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   881
          by auto
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   882
        show "emeasure lebesgue {x \<in> space lebesgue. F i x \<noteq> 0} < \<infinity>"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   883
          unfolding eq using simple_functionD[OF F(1)]
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   884
          by (subst setsum_emeasure[symmetric])
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   885
             (auto simp: disjoint_family_on_def setsum_Pinfty F_finite)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   886
      qed fact
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   887
      with F_eq show "(\<integral>\<^sup>+x. norm (real (F i x)) \<partial>lebesgue) < \<infinity>" by simp
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   888
    qed (insert F(1), auto intro!: borel_measurable_real_of_ereal dest: borel_measurable_simple_function)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   889
    then have "((\<lambda>x. real (F i x)) has_integral integral\<^sup>L lebesgue (\<lambda>x. real (F i x))) UNIV"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   890
      by (rule has_integral_lebesgue_integral_lebesgue)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   891
    from this I have "integral\<^sup>L lebesgue (\<lambda>x. real (F i x)) \<le> I"
47757
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   892
      by (rule has_integral_le) (intro ballI F_bound)
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56993
diff changeset
   893
    moreover have "integral\<^sup>N lebesgue (F i) = integral\<^sup>L lebesgue (\<lambda>x. real (F i x))"
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56993
diff changeset
   894
      using int F by (subst nn_integral_eq_integral[symmetric])  (auto simp: F_eq2[symmetric] real_of_ereal_pos)
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56993
diff changeset
   895
    ultimately show "integral\<^sup>N lebesgue (F i) \<le> ereal I"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   896
      by simp
47757
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   897
  qed
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   898
  finally show "integrable lebesgue f"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   899
    using f_borel by (auto simp: integrable_iff_bounded nonneg)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   900
  from has_integral_lebesgue_integral_lebesgue[OF this] I
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   901
  show "integral\<^sup>L lebesgue f = I"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   902
    by (metis has_integral_unique)
47757
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   903
qed
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
   904
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   905
lemma has_integral_iff_has_bochner_integral_lebesgue_nonneg:
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   906
  fixes f :: "'a::ordered_euclidean_space \<Rightarrow> real"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   907
  shows "f \<in> borel_measurable lebesgue \<Longrightarrow> (\<And>x. 0 \<le> f x) \<Longrightarrow>
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   908
    (f has_integral I) UNIV \<longleftrightarrow> has_bochner_integral lebesgue f I"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   909
  by (metis has_bochner_integral_iff has_integral_unique has_integral_lebesgue_integral_lebesgue
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   910
            integrable_has_integral_lebesgue_nonneg)
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   911
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   912
lemma
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"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   914
  assumes "f \<in> borel_measurable borel" "\<And>x. 0 \<le> f x" "(f has_integral I) UNIV"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   915
  shows integrable_has_integral_nonneg: "integrable lborel f"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   916
    and integral_has_integral_nonneg: "integral\<^sup>L lborel f = I"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   917
  by (metis assms borel_measurable_lebesgueI integrable_has_integral_lebesgue_nonneg lebesgue_integral_eq_borel(1))
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
   918
     (metis assms borel_measurable_lebesgueI has_integral_lebesgue_integral has_integral_unique integrable_has_integral_lebesgue_nonneg lebesgue_integral_eq_borel(1))
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   919
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   920
subsection {* Equivalence between product spaces and euclidean spaces *}
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   921
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50418
diff changeset
   922
definition e2p :: "'a::ordered_euclidean_space \<Rightarrow> ('a \<Rightarrow> real)" where
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50418
diff changeset
   923
  "e2p x = (\<lambda>i\<in>Basis. x \<bullet> i)"
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   924
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50418
diff changeset
   925
definition p2e :: "('a \<Rightarrow> real) \<Rightarrow> 'a::ordered_euclidean_space" where
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50418
diff changeset
   926
  "p2e x = (\<Sum>i\<in>Basis. x i *\<^sub>R i)"
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   927
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   928
lemma e2p_p2e[simp]:
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50418
diff changeset
   929
  "x \<in> extensional Basis \<Longrightarrow> e2p (p2e x::'a::ordered_euclidean_space) = x"
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   930
  by (auto simp: fun_eq_iff extensional_def p2e_def e2p_def)
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   931
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   932
lemma p2e_e2p[simp]:
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   933
  "p2e (e2p x) = (x::'a::ordered_euclidean_space)"
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50418
diff changeset
   934
  by (auto simp: euclidean_eq_iff[where 'a='a] p2e_def e2p_def)
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   935
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   936
interpretation lborel_product: product_sigma_finite "\<lambda>x. lborel::real measure"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   937
  by default
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   938
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50418
diff changeset
   939
interpretation lborel_space: finite_product_sigma_finite "\<lambda>x. lborel::real measure" "Basis"
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   940
  by default auto
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   941
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   942
lemma sets_product_borel:
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   943
  assumes I: "finite I"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51478
diff changeset
   944
  shows "sets (\<Pi>\<^sub>M i\<in>I. lborel) = sigma_sets (\<Pi>\<^sub>E i\<in>I. UNIV) { \<Pi>\<^sub>E i\<in>I. {..< x i :: real} | x. True}" (is "_ = ?G")
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   945
proof (subst sigma_prod_algebra_sigma_eq[where S="\<lambda>_ i::nat. {..<real i}" and E="\<lambda>_. range lessThan", OF I])
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51478
diff changeset
   946
  show "sigma_sets (space (Pi\<^sub>M I (\<lambda>i. lborel))) {Pi\<^sub>E I F |F. \<forall>i\<in>I. F i \<in> range lessThan} = ?G"
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   947
    by (intro arg_cong2[where f=sigma_sets]) (auto simp: space_PiM image_iff bchoice_iff)
54775
2d3df8633dad prefer box over greaterThanLessThan on euclidean_space
immler
parents: 54257
diff changeset
   948
qed (auto simp: borel_eq_lessThan eucl_lessThan reals_Archimedean2)
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   949
50003
8c213922ed49 use measurability prover
hoelzl
parents: 49801
diff changeset
   950
lemma measurable_e2p[measurable]:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51478
diff changeset
   951
  "e2p \<in> measurable (borel::'a::ordered_euclidean_space measure) (\<Pi>\<^sub>M (i::'a)\<in>Basis. (lborel :: real measure))"
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   952
proof (rule measurable_sigma_sets[OF sets_product_borel])
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51478
diff changeset
   953
  fix A :: "('a \<Rightarrow> real) set" assume "A \<in> {\<Pi>\<^sub>E (i::'a)\<in>Basis. {..<x i} |x. True} "
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51478
diff changeset
   954
  then obtain x where  "A = (\<Pi>\<^sub>E (i::'a)\<in>Basis. {..<x i})" by auto
54775
2d3df8633dad prefer box over greaterThanLessThan on euclidean_space
immler
parents: 54257
diff changeset
   955
  then have "e2p -` A = {y :: 'a. eucl_less y (\<Sum>i\<in>Basis. x i *\<^sub>R i)}"
2d3df8633dad prefer box over greaterThanLessThan on euclidean_space
immler
parents: 54257
diff changeset
   956
    using DIM_positive by (auto simp add: set_eq_iff e2p_def eucl_less_def)
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   957
  then show "e2p -` A \<inter> space (borel::'a measure) \<in> sets borel" by simp
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   958
qed (auto simp: e2p_def)
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   959
50003
8c213922ed49 use measurability prover
hoelzl
parents: 49801
diff changeset
   960
(* FIXME: conversion in measurable prover *)
50385
837e38a42d8c Remove looping rule from measurability prover
hoelzl
parents: 50244
diff changeset
   961
lemma lborelD_Collect[measurable (raw)]: "{x\<in>space borel. P x} \<in> sets borel \<Longrightarrow> {x\<in>space lborel. P x} \<in> sets lborel" by simp
837e38a42d8c Remove looping rule from measurability prover
hoelzl
parents: 50244
diff changeset
   962
lemma lborelD[measurable (raw)]: "A \<in> sets borel \<Longrightarrow> A \<in> sets lborel" by simp
50003
8c213922ed49 use measurability prover
hoelzl
parents: 49801
diff changeset
   963
8c213922ed49 use measurability prover
hoelzl
parents: 49801
diff changeset
   964
lemma measurable_p2e[measurable]:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51478
diff changeset
   965
  "p2e \<in> measurable (\<Pi>\<^sub>M (i::'a)\<in>Basis. (lborel :: real measure))
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   966
    (borel :: 'a::ordered_euclidean_space measure)"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   967
  (is "p2e \<in> measurable ?P _")
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   968
proof (safe intro!: borel_measurable_iff_halfspace_le[THEN iffD2])
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50418
diff changeset
   969
  fix x and i :: 'a
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50418
diff changeset
   970
  let ?A = "{w \<in> space ?P. (p2e w :: 'a) \<bullet> i \<le> x}"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50418
diff changeset
   971
  assume "i \<in> Basis"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51478
diff changeset
   972
  then have "?A = (\<Pi>\<^sub>E j\<in>Basis. if i = j then {.. x} else UNIV)"
50123
69b35a75caf3 merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents: 50105
diff changeset
   973
    using DIM_positive by (auto simp: space_PiM p2e_def PiE_def split: split_if_asm)
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   974
  then show "?A \<in> sets ?P"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   975
    by auto
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   976
qed
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   977
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   978
lemma lborel_eq_lborel_space:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51478
diff changeset
   979
  "(lborel :: 'a measure) = distr (\<Pi>\<^sub>M (i::'a::ordered_euclidean_space)\<in>Basis. lborel) borel p2e"
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   980
  (is "?B = ?D")
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   981
proof (rule lborel_eqI)
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   982
  show "sets ?D = sets borel" by simp
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51478
diff changeset
   983
  let ?P = "(\<Pi>\<^sub>M (i::'a)\<in>Basis. lborel)"
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   984
  fix a b :: 'a
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51478
diff changeset
   985
  have *: "p2e -` {a .. b} \<inter> space ?P = (\<Pi>\<^sub>E i\<in>Basis. {a \<bullet> i .. b \<bullet> i})"
50123
69b35a75caf3 merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents: 50105
diff changeset
   986
    by (auto simp: eucl_le[where 'a='a] p2e_def space_PiM PiE_def Pi_iff)
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   987
  have "emeasure ?P (p2e -` {a..b} \<inter> space ?P) = content {a..b}"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   988
  proof cases
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   989
    assume "{a..b} \<noteq> {}"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   990
    then have "a \<le> b"
56188
0268784f60da use cbox to relax class constraints
immler
parents: 56181
diff changeset
   991
      by (simp add: eucl_le[where 'a='a])
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50418
diff changeset
   992
    then have "emeasure lborel {a..b} = (\<Prod>x\<in>Basis. emeasure lborel {a \<bullet> x .. b \<bullet> x})"
56188
0268784f60da use cbox to relax class constraints
immler
parents: 56181
diff changeset
   993
      by (auto simp: eucl_le[where 'a='a] content_closed_interval
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   994
               intro!: setprod_ereal[symmetric])
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   995
    also have "\<dots> = emeasure ?P (p2e -` {a..b} \<inter> space ?P)"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   996
      unfolding * by (subst lborel_space.measure_times) auto
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   997
    finally show ?thesis by simp
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   998
  qed simp
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
   999
  then show "emeasure ?D {a .. b} = content {a .. b}"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1000
    by (simp add: emeasure_distr measurable_p2e)
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1001
qed
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1002
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1003
lemma borel_fubini_positiv_integral:
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1004
  fixes f :: "'a::ordered_euclidean_space \<Rightarrow> ereal"
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1005
  assumes f: "f \<in> borel_measurable borel"
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56993
diff changeset
  1006
  shows "integral\<^sup>N lborel f = \<integral>\<^sup>+x. f (p2e x) \<partial>(\<Pi>\<^sub>M (i::'a)\<in>Basis. lborel)"
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56993
diff changeset
  1007
  by (subst lborel_eq_lborel_space) (simp add: nn_integral_distr measurable_p2e f)
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1008
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1009
lemma borel_fubini_integrable:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
  1010
  fixes f :: "'a::ordered_euclidean_space \<Rightarrow> _::{banach, second_countable_topology}"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51478
diff changeset
  1011
  shows "integrable lborel f \<longleftrightarrow> integrable (\<Pi>\<^sub>M (i::'a)\<in>Basis. lborel) (\<lambda>x. f (p2e x))"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
  1012
  unfolding integrable_iff_bounded
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
  1013
proof (intro conj_cong[symmetric])
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
  1014
  show "((\<lambda>x. f (p2e x)) \<in> borel_measurable (Pi\<^sub>M Basis (\<lambda>i. lborel))) = (f \<in> borel_measurable lborel)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
  1015
  proof
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
  1016
    assume "((\<lambda>x. f (p2e x)) \<in> borel_measurable (Pi\<^sub>M Basis (\<lambda>i. lborel)))"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
  1017
    then have "(\<lambda>x. f (p2e (e2p x))) \<in> borel_measurable borel"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
  1018
      by measurable
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
  1019
    then show "f \<in> borel_measurable lborel"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
  1020
      by simp
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
  1021
  qed simp
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
  1022
qed (simp add: borel_fubini_positiv_integral)
49777
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1023
6ac97ab9b295 tuned Lebesgue measure proofs
hoelzl
parents: 47757
diff changeset
  1024
lemma borel_fubini:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
  1025
  fixes f :: "'a::ordered_euclidean_space \<Rightarrow> _::{banach, second_countable_topology}"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
  1026
  shows "f \<in> borel_measurable borel \<Longrightarrow>
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
  1027
    integral\<^sup>L lborel f = \<integral>x. f (p2e x) \<partial>((\<Pi>\<^sub>M (i::'a)\<in>Basis. lborel))"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
  1028
  by (subst lborel_eq_lborel_space) (simp add: integral_distr)
47757
5e6fe71e2390 equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents: 47694
diff changeset
  1029
50418
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1030
lemma integrable_on_borel_integrable:
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1031
  fixes f :: "'a::ordered_euclidean_space \<Rightarrow> real"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
  1032
  shows "f \<in> borel_measurable borel \<Longrightarrow> (\<And>x. 0 \<le> f x) \<Longrightarrow> f integrable_on UNIV \<Longrightarrow> integrable lborel f"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
  1033
  by (metis borel_measurable_lebesgueI integrable_has_integral_nonneg integrable_on_def
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
  1034
            lebesgue_integral_eq_borel(1))
50418
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1035
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1036
subsection {* Fundamental Theorem of Calculus for the Lebesgue integral *}
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1037
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1038
lemma borel_integrable_atLeastAtMost:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
  1039
  fixes f :: "real \<Rightarrow> real"
50418
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1040
  assumes f: "\<And>x. a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow> isCont f x"
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1041
  shows "integrable lborel (\<lambda>x. f x * indicator {a .. b} x)" (is "integrable _ ?f")
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1042
proof cases
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1043
  assume "a \<le> b"
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1044
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1045
  from isCont_Lb_Ub[OF `a \<le> b`, of f] f
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1046
  obtain M L where
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1047
    bounds: "\<And>x. a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow> f x \<le> M" "\<And>x. a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow> L \<le> f x"
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1048
    by metis
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1049
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1050
  show ?thesis
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1051
  proof (rule integrable_bound)
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1052
    show "integrable lborel (\<lambda>x. max \<bar>M\<bar> \<bar>L\<bar> * indicator {a..b} x)"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
  1053
      by (intro integrable_mult_right integrable_real_indicator) simp_all
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
  1054
    show "AE x in lborel. norm (?f x) \<le> norm (max \<bar>M\<bar> \<bar>L\<bar> * indicator {a..b} x)"
50418
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1055
    proof (rule AE_I2)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
  1056
      fix x show "norm (?f x) \<le> norm (max \<bar>M\<bar> \<bar>L\<bar> * indicator {a..b} x)"
50418
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1057
        using bounds[of x] by (auto split: split_indicator)
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1058
    qed
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1059
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1060
    let ?g = "\<lambda>x. if x = a then f a else if x = b then f b else if x \<in> {a <..< b} then f x else 0"
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1061
    from f have "continuous_on {a <..< b} f"
51478
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51000
diff changeset
  1062
      by (subst continuous_on_eq_continuous_at) auto
50418
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1063
    then have "?g \<in> borel_measurable borel"
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1064
      using borel_measurable_continuous_on_open[of "{a <..< b }" f "\<lambda>x. x" borel 0]
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1065
      by (auto intro!: measurable_If[where P="\<lambda>x. x = a"] measurable_If[where P="\<lambda>x. x = b"])
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1066
    also have "?g = ?f"
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50418
diff changeset
  1067
      using `a \<le> b` by (intro ext) (auto split: split_indicator)
50418
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1068
    finally show "?f \<in> borel_measurable lborel"
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1069
      by simp
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1070
  qed
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1071
qed simp
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1072
56181
2aa0b19e74f3 unify syntax for has_derivative and differentiable
hoelzl
parents: 56020
diff changeset
  1073
lemma has_field_derivative_subset:
2aa0b19e74f3 unify syntax for has_derivative and differentiable
hoelzl
parents: 56020
diff changeset
  1074
  "(f has_field_derivative y) (at x within s) \<Longrightarrow> t \<subseteq> s \<Longrightarrow> (f has_field_derivative y) (at x within t)"
2aa0b19e74f3 unify syntax for has_derivative and differentiable
hoelzl
parents: 56020
diff changeset
  1075
  unfolding has_field_derivative_def by (rule has_derivative_subset)
2aa0b19e74f3 unify syntax for has_derivative and differentiable
hoelzl
parents: 56020
diff changeset
  1076
50418
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1077
lemma integral_FTC_atLeastAtMost:
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1078
  fixes a b :: real
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1079
  assumes "a \<le> b"
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1080
    and F: "\<And>x. a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow> DERIV F x :> f x"
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1081
    and f: "\<And>x. a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow> isCont f x"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51478
diff changeset
  1082
  shows "integral\<^sup>L lborel (\<lambda>x. f x * indicator {a .. b} x) = F b - F a"
50418
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1083
proof -
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1084
  let ?f = "\<lambda>x. f x * indicator {a .. b} x"
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1085
  have "(?f has_integral (\<integral>x. ?f x \<partial>lborel)) UNIV"
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1086
    using borel_integrable_atLeastAtMost[OF f]
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
  1087
    by (rule has_integral_lebesgue_integral)
50418
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1088
  moreover
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1089
  have "(f has_integral F b - F a) {a .. b}"
56181
2aa0b19e74f3 unify syntax for has_derivative and differentiable
hoelzl
parents: 56020
diff changeset
  1090
    by (intro fundamental_theorem_of_calculus)
2aa0b19e74f3 unify syntax for has_derivative and differentiable
hoelzl
parents: 56020
diff changeset
  1091
       (auto simp: has_field_derivative_iff_has_vector_derivative[symmetric]
2aa0b19e74f3 unify syntax for has_derivative and differentiable
hoelzl
parents: 56020
diff changeset
  1092
             intro: has_field_derivative_subset[OF F] assms(1))
50418
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1093
  then have "(?f has_integral F b - F a) {a .. b}"
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1094
    by (subst has_integral_eq_eq[where g=f]) auto
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1095
  then have "(?f has_integral F b - F a) UNIV"
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1096
    by (intro has_integral_on_superset[where t=UNIV and s="{a..b}"]) auto
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51478
diff changeset
  1097
  ultimately show "integral\<^sup>L lborel ?f = F b - F a"
50418
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1098
    by (rule has_integral_unique)
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1099
qed
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1100
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1101
text {*
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1102
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1103
For the positive integral we replace continuity with Borel-measurability. 
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1104
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1105
*}
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1106
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
  1107
lemma
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
  1108
  fixes f :: "real \<Rightarrow> real"
50418
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1109
  assumes f_borel: "f \<in> borel_measurable borel"
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1110
  assumes f: "\<And>x. x \<in> {a..b} \<Longrightarrow> DERIV F x :> f x" "\<And>x. x \<in> {a..b} \<Longrightarrow> 0 \<le> f x" and "a \<le> b"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
  1111
  shows integral_FTC_Icc_nonneg: "(\<integral>x. f x * indicator {a .. b} x \<partial>lborel) = F b - F a" (is ?eq)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
  1112
    and integrable_FTC_Icc_nonneg: "integrable lborel (\<lambda>x. f x * indicator {a .. b} x)" (is ?int)
50418
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1113
proof -
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1114
  have i: "(f has_integral F b - F a) {a..b}"
56181
2aa0b19e74f3 unify syntax for has_derivative and differentiable
hoelzl
parents: 56020
diff changeset
  1115
    by (intro fundamental_theorem_of_calculus)
2aa0b19e74f3 unify syntax for has_derivative and differentiable
hoelzl
parents: 56020
diff changeset
  1116
       (auto simp: has_field_derivative_iff_has_vector_derivative[symmetric]
2aa0b19e74f3 unify syntax for has_derivative and differentiable
hoelzl
parents: 56020
diff changeset
  1117
             intro: has_field_derivative_subset[OF f(1)] `a \<le> b`)
50418
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1118
  have i: "((\<lambda>x. f x * indicator {a..b} x) has_integral F b - F a) {a..b}"
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1119
    by (rule has_integral_eq[OF _ i]) auto
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1120
  have i: "((\<lambda>x. f x * indicator {a..b} x) has_integral F b - F a) UNIV"
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1121
    by (rule has_integral_on_superset[OF _ _ i]) auto
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
  1122
  from i f f_borel show ?eq
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
  1123
    by (intro integral_has_integral_nonneg) (auto split: split_indicator)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
  1124
  from i f f_borel show ?int
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
  1125
    by (intro integrable_has_integral_nonneg) (auto split: split_indicator)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
  1126
qed
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
  1127
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56993
diff changeset
  1128
lemma nn_integral_FTC_atLeastAtMost:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
  1129
  assumes "f \<in> borel_measurable borel" "\<And>x. x \<in> {a..b} \<Longrightarrow> DERIV F x :> f x" "\<And>x. x \<in> {a..b} \<Longrightarrow> 0 \<le> f x" "a \<le> b"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
  1130
  shows "(\<integral>\<^sup>+x. f x * indicator {a .. b} x \<partial>lborel) = F b - F a"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
  1131
proof -
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
  1132
  have "integrable lborel (\<lambda>x. f x * indicator {a .. b} x)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
  1133
    by (rule integrable_FTC_Icc_nonneg) fact+
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
  1134
  then have "(\<integral>\<^sup>+x. f x * indicator {a .. b} x \<partial>lborel) = (\<integral>x. f x * indicator {a .. b} x \<partial>lborel)"
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56993
diff changeset
  1135
    using assms by (intro nn_integral_eq_integral) (auto simp: indicator_def)
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
  1136
  also have "(\<integral>x. f x * indicator {a .. b} x \<partial>lborel) = F b - F a"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
  1137
    by (rule integral_FTC_Icc_nonneg) fact+
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
  1138
  finally show ?thesis
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56218
diff changeset
  1139
    unfolding ereal_indicator[symmetric] by simp
50418
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1140
qed
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1141
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56993
diff changeset
  1142
lemma nn_integral_FTC_atLeast:
50418
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1143
  fixes f :: "real \<Rightarrow> real"
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1144
  assumes f_borel: "f \<in> borel_measurable borel"
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1145
  assumes f: "\<And>x. a \<le> x \<Longrightarrow> DERIV F x :> f x" 
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1146
  assumes nonneg: "\<And>x. a \<le> x \<Longrightarrow> 0 \<le> f x"
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1147
  assumes lim: "(F ---> T) at_top"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51478
diff changeset
  1148
  shows "(\<integral>\<^sup>+x. ereal (f x) * indicator {a ..} x \<partial>lborel) = T - F a"
50418
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1149
proof -
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1150
  let ?f = "\<lambda>(i::nat) (x::real). ereal (f x) * indicator {a..a + real i} x"
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1151
  let ?fR = "\<lambda>x. ereal (f x) * indicator {a ..} x"
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1152
  have "\<And>x. (SUP i::nat. ?f i x) = ?fR x"
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1153
  proof (rule SUP_Lim_ereal)
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1154
    show "\<And>x. incseq (\<lambda>i. ?f i x)"
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1155
      using nonneg by (auto simp: incseq_def le_fun_def split: split_indicator)
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1156
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1157
    fix x
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1158
    from reals_Archimedean2[of "x - a"] guess n ..
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1159
    then have "eventually (\<lambda>n. ?f n x = ?fR x) sequentially"
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1160
      by (auto intro!: eventually_sequentiallyI[where c=n] split: split_indicator)
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1161
    then show "(\<lambda>n. ?f n x) ----> ?fR x"
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1162
      by (rule Lim_eventually)
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1163
  qed
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56993
diff changeset
  1164
  then have "integral\<^sup>N lborel ?fR = (\<integral>\<^sup>+ x. (SUP i::nat. ?f i x) \<partial>lborel)"
50418
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1165
    by simp
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51478
diff changeset
  1166
  also have "\<dots> = (SUP i::nat. (\<integral>\<^sup>+ x. ?f i x \<partial>lborel))"
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56993
diff changeset
  1167
  proof (rule nn_integral_monotone_convergence_SUP)
50418
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1168
    show "incseq ?f"
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1169
      using nonneg by (auto simp: incseq_def le_fun_def split: split_indicator)
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1170
    show "\<And>i. (?f i) \<in> borel_measurable lborel"
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1171
      using f_borel by auto
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1172
    show "\<And>i x. 0 \<le> ?f i x"
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1173
      using nonneg by (auto split: split_indicator)
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1174
  qed
54257
5c7a3b6b05a9 generalize SUP and INF to the syntactic type classes Sup and Inf
hoelzl
parents: 53374
diff changeset
  1175
  also have "\<dots> = (SUP i::nat. ereal (F (a + real i) - F a))"
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56993
diff changeset
  1176
    by (subst nn_integral_FTC_atLeastAtMost[OF f_borel f nonneg]) auto
50418
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1177
  also have "\<dots> = T - F a"
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1178
  proof (rule SUP_Lim_ereal)
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1179
    show "incseq (\<lambda>n. ereal (F (a + real n) - F a))"
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1180
    proof (simp add: incseq_def, safe)
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1181
      fix m n :: nat assume "m \<le> n"
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1182
      with f nonneg show "F (a + real m) \<le> F (a + real n)"
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1183
        by (intro DERIV_nonneg_imp_nondecreasing[where f=F])
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1184
           (simp, metis add_increasing2 order_refl order_trans real_of_nat_ge_zero)
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1185
    qed 
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1186
    have "(\<lambda>x. F (a + real x)) ----> T"
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1187
      apply (rule filterlim_compose[OF lim filterlim_tendsto_add_at_top])
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1188
      apply (rule LIMSEQ_const_iff[THEN iffD2, OF refl])
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1189
      apply (rule filterlim_real_sequentially)
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1190
      done
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1191
    then show "(\<lambda>n. ereal (F (a + real n) - F a)) ----> ereal (T - F a)"
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1192
      unfolding lim_ereal
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1193
      by (intro tendsto_diff) auto
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1194
  qed
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1195
  finally show ?thesis .
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1196
qed
bd68cf816dd3 fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents: 50385
diff changeset
  1197
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1198
end