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