src/HOL/Probability/Lebesgue_Measure.thy
author hoelzl
Tue, 18 Jan 2011 21:37:23 +0100
changeset 41654 32fe42892983
parent 41546 2a12c23b7a34
child 41661 baf1964bc468
permissions -rw-r--r--
Gauge measure removed
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
     1
(*  Author: Robert Himmelmann, TU Muenchen *)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
     2
header {* Lebsegue measure *}
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
     3
theory Lebesgue_Measure
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
     4
  imports Product_Measure Complete_Measure
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
     5
begin
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
     6
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
     7
subsection {* Standard Cubes *}
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
     8
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
     9
definition cube :: "nat \<Rightarrow> 'a::ordered_euclidean_space set" where
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
    10
  "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
    11
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
    12
lemma cube_closed[intro]: "closed (cube n)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
    13
  unfolding cube_def by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
    14
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
    15
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
    16
  by (fastsimp simp: eucl_le[where 'a='a] cube_def)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    17
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
    18
lemma cube_subset_iff:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
    19
  "cube n \<subseteq> cube N \<longleftrightarrow> n \<le> N"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
    20
proof
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
    21
  assume subset: "cube n \<subseteq> (cube N::'a set)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
    22
  then have "((\<chi>\<chi> i. real n)::'a) \<in> cube N"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
    23
    using DIM_positive[where 'a='a]
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
    24
    by (fastsimp simp: cube_def eucl_le[where 'a='a])
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
    25
  then show "n \<le> N"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
    26
    by (fastsimp simp: cube_def eucl_le[where 'a='a])
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
    27
next
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
    28
  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
    29
qed
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    30
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    31
lemma ball_subset_cube:"ball (0::'a::ordered_euclidean_space) (real n) \<subseteq> cube n"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    32
  unfolding cube_def subset_eq mem_interval apply safe unfolding euclidean_lambda_beta'
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    33
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
    34
  thus "- real n \<le> x $$ i" "real n \<ge> x $$ i"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    35
    using component_le_norm[of x i] by(auto simp: dist_norm)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    36
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    37
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    38
lemma mem_big_cube: obtains n where "x \<in> cube n"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    39
proof- from real_arch_lt[of "norm x"] guess n ..
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    40
  thus ?thesis apply-apply(rule that[where n=n])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    41
    apply(rule ball_subset_cube[unfolded subset_eq,rule_format])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    42
    by (auto simp add:dist_norm)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    43
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    44
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
    45
definition lebesgue :: "'a::ordered_euclidean_space algebra" where
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
    46
  "lebesgue = \<lparr> space = UNIV, sets = {A. \<forall>n. (indicator A :: 'a \<Rightarrow> real) integrable_on cube n} \<rparr>"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
    47
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
    48
lemma space_lebesgue[simp]: "space lebesgue = UNIV"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
    49
  unfolding lebesgue_def by simp
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
    50
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
    51
lemma lebesgueD: "A \<in> sets lebesgue \<Longrightarrow> (indicator A :: _ \<Rightarrow> real) integrable_on cube n"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
    52
  unfolding lebesgue_def by simp
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
    53
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
    54
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
    55
  unfolding lebesgue_def by simp
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
    56
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
    57
lemma absolutely_integrable_on_indicator[simp]:
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
    58
  fixes A :: "'a::ordered_euclidean_space set"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
    59
  shows "((indicator A :: _ \<Rightarrow> real) absolutely_integrable_on X) \<longleftrightarrow>
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
    60
    (indicator A :: _ \<Rightarrow> real) integrable_on X"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
    61
  unfolding absolutely_integrable_on_def by simp
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
    62
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
    63
lemma LIMSEQ_indicator_UN:
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
    64
  "(\<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
    65
proof cases
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
    66
  assume "\<exists>i. x \<in> A i" then guess i .. note i = this
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
    67
  then have *: "\<And>n. (indicator (\<Union> i<n + Suc i. A i) x :: real) = 1"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
    68
    "(indicator (\<Union> i. A i) x :: real) = 1" by (auto simp: indicator_def)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
    69
  show ?thesis
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
    70
    apply (rule LIMSEQ_offset[of _ "Suc i"]) unfolding * by auto
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
    71
qed (auto simp: indicator_def)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    72
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
    73
lemma indicator_add:
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
    74
  "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
    75
  unfolding indicator_def by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    76
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
    77
interpretation lebesgue: sigma_algebra lebesgue
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
    78
proof (intro sigma_algebra_iff2[THEN iffD2] conjI allI ballI impI lebesgueI)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
    79
  fix A n assume A: "A \<in> sets lebesgue"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
    80
  have "indicator (space lebesgue - A) = (\<lambda>x. 1 - indicator A x :: real)"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
    81
    by (auto simp: fun_eq_iff indicator_def)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
    82
  then show "(indicator (space lebesgue - A) :: _ \<Rightarrow> real) integrable_on cube n"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
    83
    using A by (auto intro!: integrable_sub dest: lebesgueD simp: cube_def)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
    84
next
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
    85
  fix n show "(indicator {} :: _\<Rightarrow>real) integrable_on cube n"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
    86
    by (auto simp: cube_def indicator_def_raw)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
    87
next
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
    88
  fix A :: "nat \<Rightarrow> 'a set" and n ::nat assume "range A \<subseteq> sets lebesgue"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
    89
  then have A: "\<And>i. (indicator (A i) :: _ \<Rightarrow> real) integrable_on cube n"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
    90
    by (auto dest: lebesgueD)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
    91
  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
    92
  proof (intro dominated_convergence[where g="?g"] ballI)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
    93
    fix k show "(indicator (\<Union>i<k. A i) :: _ \<Rightarrow> real) integrable_on cube n"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
    94
    proof (induct k)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
    95
      case (Suc k)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
    96
      have *: "(\<Union> i<Suc k. A i) = (\<Union> i<k. A i) \<union> A k"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
    97
        unfolding lessThan_Suc UN_insert by auto
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
    98
      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
    99
          indicator (\<Union> i<Suc k. A i)" (is "(\<lambda>x. max (?f x) (?g x)) = _")
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   100
        by (auto simp: fun_eq_iff * indicator_def)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   101
      show ?case
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   102
        using absolutely_integrable_max[of ?f "cube n" ?g] A Suc by (simp add: *)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   103
    qed auto
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   104
  qed (auto intro: LIMSEQ_indicator_UN simp: cube_def)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   105
qed simp
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   106
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   107
definition "lmeasure A = (SUP n. Real (integral (cube n) (indicator A)))"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   108
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   109
interpretation lebesgue: measure_space lebesgue lmeasure
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   110
proof
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   111
  have *: "indicator {} = (\<lambda>x. 0 :: real)" by (simp add: fun_eq_iff)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   112
  show "lmeasure {} = 0" by (simp add: integral_0 * lmeasure_def)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   113
next
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   114
  show "countably_additive lebesgue lmeasure"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   115
  proof (intro countably_additive_def[THEN iffD2] allI impI)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   116
    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
   117
    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
   118
      by (auto dest: lebesgueD)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   119
    let "?m n i" = "integral (cube n) (indicator (A i) :: _\<Rightarrow>real)"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   120
    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
   121
    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
   122
    assume "(\<Union>i. A i) \<in> sets lebesgue"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   123
    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
   124
      by (auto dest: lebesgueD)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   125
    show "(\<Sum>\<^isub>\<infinity>n. lmeasure (A n)) = lmeasure (\<Union>i. A i)" unfolding lmeasure_def
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   126
    proof (subst psuminf_SUP_eq)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   127
      fix n i show "Real (?m n i) \<le> Real (?m (Suc n) i)"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   128
        using cube_subset[of n "Suc n"] by (auto intro!: integral_subset_le)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   129
    next
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   130
      show "(SUP n. \<Sum>\<^isub>\<infinity>i. Real (?m n i)) = (SUP n. Real (?M n UNIV))"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   131
        unfolding psuminf_def
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   132
      proof (subst setsum_Real, (intro arg_cong[where f="SUPR UNIV"] ext ballI nn SUP_eq_LIMSEQ[THEN iffD2])+)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   133
        fix n :: nat show "mono (\<lambda>m. \<Sum>x<m. ?m n x)"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   134
        proof (intro mono_iff_le_Suc[THEN iffD2] allI)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   135
          fix m show "(\<Sum>x<m. ?m n x) \<le> (\<Sum>x<Suc m. ?m n x)"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   136
            using nn[of n m] by auto
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   137
        qed
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   138
        show "0 \<le> ?M n UNIV"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   139
          using UN_A by (auto intro!: integral_nonneg)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   140
        fix m show "0 \<le> (\<Sum>x<m. ?m n x)" by (auto intro!: setsum_nonneg)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   141
      next
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   142
        fix n
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   143
        have "\<And>m. (UNION {..<m} A) \<in> sets lebesgue" using rA by auto
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   144
        from lebesgueD[OF this]
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   145
        have "(\<lambda>m. ?M n {..< m}) ----> ?M n UNIV"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   146
          (is "(\<lambda>m. integral _ (?A m)) ----> ?I")
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   147
          by (intro dominated_convergence(2)[where f="?A" and h="\<lambda>x. 1::real"])
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   148
             (auto intro: LIMSEQ_indicator_UN simp: cube_def)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   149
        moreover
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   150
        { fix m have *: "(\<Sum>x<m. ?m n x) = ?M n {..< m}"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   151
          proof (induct m)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   152
            case (Suc m)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   153
            have "(\<Union>i<m. A i) \<in> sets lebesgue" using rA by auto
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   154
            then have "(indicator (\<Union>i<m. A i) :: _\<Rightarrow>real) integrable_on (cube n)"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   155
              by (auto dest!: lebesgueD)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   156
            moreover
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   157
            have "(\<Union>i<m. A i) \<inter> A m = {}"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   158
              using rA(2)[unfolded disjoint_family_on_def, THEN bspec, of m]
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   159
              by auto
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   160
            then have "\<And>x. indicator (\<Union>i<Suc m. A i) x =
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   161
              indicator (\<Union>i<m. A i) x + (indicator (A m) x :: real)"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   162
              by (auto simp: indicator_add lessThan_Suc ac_simps)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   163
            ultimately show ?case
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   164
              using Suc A by (simp add: integral_add[symmetric])
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   165
          qed auto }
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   166
        ultimately show "(\<lambda>m. \<Sum>x<m. ?m n x) ----> ?M n UNIV"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   167
          by simp
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   168
      qed
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   169
    qed
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   170
  qed
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   171
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   172
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   173
lemma has_integral_interval_cube:
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   174
  fixes a b :: "'a::ordered_euclidean_space"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   175
  shows "(indicator {a .. b} has_integral
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   176
    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
   177
    (is "(?I has_integral content ?R) (cube n)")
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   178
proof -
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   179
  let "{?N .. ?P}" = ?R
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   180
  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
   181
    by (auto simp: indicator_def cube_def fun_eq_iff eucl_le[where 'a='a])
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   182
  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
   183
    unfolding has_integral_restrict_univ[where s="cube n", symmetric] by simp
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   184
  also have "\<dots> \<longleftrightarrow> ((\<lambda>x. 1) has_integral content ?R) ?R"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   185
    unfolding indicator_def_raw has_integral_restrict_univ ..
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   186
  finally show ?thesis
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   187
    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
   188
qed
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   189
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   190
lemma lebesgueI_borel[intro, simp]:
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   191
  fixes s::"'a::ordered_euclidean_space set"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   192
  assumes "s \<in> sets borel" shows "s \<in> sets lebesgue"
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   193
proof -
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   194
  let ?S = "range (\<lambda>(a, b). {a .. (b :: 'a\<Colon>ordered_euclidean_space)})"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   195
  have *:"?S \<subseteq> sets lebesgue"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   196
  proof (safe intro!: lebesgueI)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   197
    fix n :: nat and a b :: 'a
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   198
    let ?N = "\<chi>\<chi> i. max (- real n) (a $$ i)"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   199
    let ?P = "\<chi>\<chi> i. min (real n) (b $$ i)"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   200
    show "(indicator {a..b} :: 'a\<Rightarrow>real) integrable_on cube n"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   201
      unfolding integrable_on_def
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   202
      using has_integral_interval_cube[of a b] by auto
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   203
  qed
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   204
  have "s \<in> sigma_sets UNIV ?S" using assms
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   205
    unfolding borel_eq_atLeastAtMost by (simp add: sigma_def)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   206
  thus ?thesis
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   207
    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
   208
    by (auto simp: sigma_def)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   209
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   210
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   211
lemma lebesgueI_negligible[dest]: fixes s::"'a::ordered_euclidean_space set"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   212
  assumes "negligible s" shows "s \<in> sets lebesgue"
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   213
  using assms by (force simp: cube_def integrable_on_def negligible_def intro!: lebesgueI)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   214
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   215
lemma lmeasure_eq_0:
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   216
  fixes S :: "'a::ordered_euclidean_space set" assumes "negligible S" shows "lmeasure S = 0"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   217
proof -
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   218
  have "\<And>n. integral (cube n) (indicator S :: 'a\<Rightarrow>real) = 0"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   219
    unfolding integral_def using assms
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   220
    by (intro some1_equality ex_ex1I has_integral_unique)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   221
       (auto simp: cube_def negligible_def intro: )
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   222
  then show ?thesis unfolding lmeasure_def by auto
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   223
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   224
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   225
lemma lmeasure_iff_LIMSEQ:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   226
  assumes "A \<in> sets lebesgue" "0 \<le> m"
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   227
  shows "lmeasure A = Real m \<longleftrightarrow> (\<lambda>n. integral (cube n) (indicator A :: _ \<Rightarrow> real)) ----> m"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   228
  unfolding lmeasure_def
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   229
proof (intro SUP_eq_LIMSEQ)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   230
  show "mono (\<lambda>n. integral (cube n) (indicator A::_=>real))"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   231
    using cube_subset assms by (intro monoI integral_subset_le) (auto dest!: lebesgueD)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   232
  fix n show "0 \<le> integral (cube n) (indicator A::_=>real)"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   233
    using assms by (auto dest!: lebesgueD intro!: integral_nonneg)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   234
qed fact
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   235
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   236
lemma has_integral_indicator_UNIV:
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   237
  fixes s A :: "'a::ordered_euclidean_space set" and x :: real
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   238
  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
   239
proof -
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   240
  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
   241
    by (auto simp: fun_eq_iff indicator_def)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   242
  then show ?thesis
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   243
    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
   244
qed
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   245
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   246
lemma
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   247
  fixes s a :: "'a::ordered_euclidean_space set"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   248
  shows integral_indicator_UNIV:
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   249
    "integral UNIV (indicator (s \<inter> A) :: 'a\<Rightarrow>real) = integral A (indicator s :: _\<Rightarrow>real)"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   250
  and integrable_indicator_UNIV:
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   251
    "(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
   252
  unfolding integral_def integrable_on_def has_integral_indicator_UNIV by auto
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   253
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   254
lemma lmeasure_finite_has_integral:
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   255
  fixes s :: "'a::ordered_euclidean_space set"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   256
  assumes "s \<in> sets lebesgue" "lmeasure s = Real m" "0 \<le> m"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   257
  shows "(indicator s has_integral m) UNIV"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   258
proof -
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   259
  let ?I = "indicator :: 'a set \<Rightarrow> 'a \<Rightarrow> real"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   260
  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
   261
  proof (intro monotone_convergence_increasing allI ballI)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   262
    have LIMSEQ: "(\<lambda>n. integral (cube n) (?I s)) ----> m"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   263
      using assms(2) unfolding lmeasure_iff_LIMSEQ[OF assms(1, 3)] .
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   264
    { fix n have "integral (cube n) (?I s) \<le> m"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   265
        using cube_subset assms
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   266
        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
   267
           (auto dest!: lebesgueD) }
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   268
    moreover
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   269
    { fix n have "0 \<le> integral (cube n) (?I s)"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   270
      using assms by (auto dest!: lebesgueD intro!: integral_nonneg) }
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   271
    ultimately
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   272
    show "bounded {integral UNIV (?I (s \<inter> cube k)) |k. True}"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   273
      unfolding bounded_def
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   274
      apply (rule_tac exI[of _ 0])
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   275
      apply (rule_tac exI[of _ m])
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   276
      by (auto simp: dist_real_def integral_indicator_UNIV)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   277
    fix k show "?I (s \<inter> cube k) integrable_on UNIV"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   278
      unfolding integrable_indicator_UNIV using assms by (auto dest!: lebesgueD)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   279
    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
   280
      using cube_subset[of k "Suc k"] by (auto simp: indicator_def)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   281
  next
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   282
    fix x :: 'a
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   283
    from mem_big_cube obtain k where k: "x \<in> cube k" .
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   284
    { fix n have "?I (s \<inter> cube (n + k)) x = ?I s x"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   285
      using k cube_subset[of k "n + k"] by (auto simp: indicator_def) }
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   286
    note * = this
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   287
    show "(\<lambda>k. ?I (s \<inter> cube k) x) ----> ?I s x"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   288
      by (rule LIMSEQ_offset[where k=k]) (auto simp: *)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   289
  qed
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   290
  note ** = conjunctD2[OF this]
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   291
  have m: "m = integral UNIV (?I s)"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   292
    apply (intro LIMSEQ_unique[OF _ **(2)])
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   293
    using assms(2) unfolding lmeasure_iff_LIMSEQ[OF assms(1,3)] integral_indicator_UNIV .
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   294
  show ?thesis
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   295
    unfolding m by (intro integrable_integral **)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   296
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   297
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   298
lemma lmeasure_finite_integrable: assumes "s \<in> sets lebesgue" "lmeasure s \<noteq> \<omega>"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   299
  shows "(indicator s :: _ \<Rightarrow> real) integrable_on UNIV"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   300
proof (cases "lmeasure s")
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   301
  case (preal m) from lmeasure_finite_has_integral[OF `s \<in> sets lebesgue` this]
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   302
  show ?thesis unfolding integrable_on_def by auto
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   303
qed (insert assms, auto)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   304
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   305
lemma has_integral_lebesgue: assumes "((indicator s :: _\<Rightarrow>real) has_integral m) UNIV"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   306
  shows "s \<in> sets lebesgue"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   307
proof (intro lebesgueI)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   308
  let ?I = "indicator :: 'a set \<Rightarrow> 'a \<Rightarrow> real"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   309
  fix n show "(?I s) integrable_on cube n" unfolding cube_def
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   310
  proof (intro integrable_on_subinterval)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   311
    show "(?I s) integrable_on UNIV"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   312
      unfolding integrable_on_def using assms by auto
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   313
  qed auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   314
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   315
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   316
lemma has_integral_lmeasure: assumes "((indicator s :: _\<Rightarrow>real) has_integral m) UNIV"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   317
  shows "lmeasure s = Real m"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   318
proof (intro lmeasure_iff_LIMSEQ[THEN iffD2])
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   319
  let ?I = "indicator :: 'a set \<Rightarrow> 'a \<Rightarrow> real"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   320
  show "s \<in> sets lebesgue" using has_integral_lebesgue[OF assms] .
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   321
  show "0 \<le> m" using assms by (rule has_integral_nonneg) auto
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   322
  have "(\<lambda>n. integral UNIV (?I (s \<inter> cube n))) ----> integral UNIV (?I s)"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   323
  proof (intro dominated_convergence(2) ballI)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   324
    show "(?I s) integrable_on UNIV" unfolding integrable_on_def using assms by auto
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   325
    fix n show "?I (s \<inter> cube n) integrable_on UNIV"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   326
      unfolding integrable_indicator_UNIV using `s \<in> sets lebesgue` by (auto dest: lebesgueD)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   327
    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
   328
  next
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   329
    fix x :: 'a
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   330
    from mem_big_cube obtain k where k: "x \<in> cube k" .
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   331
    { fix n have "?I (s \<inter> cube (n + k)) x = ?I s x"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   332
      using k cube_subset[of k "n + k"] by (auto simp: indicator_def) }
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   333
    note * = this
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   334
    show "(\<lambda>k. ?I (s \<inter> cube k) x) ----> ?I s x"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   335
      by (rule LIMSEQ_offset[where k=k]) (auto simp: *)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   336
  qed
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   337
  then show "(\<lambda>n. integral (cube n) (?I s)) ----> m"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   338
    unfolding integral_unique[OF assms] integral_indicator_UNIV by simp
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   339
qed
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   340
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   341
lemma has_integral_iff_lmeasure:
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   342
  "(indicator A has_integral m) UNIV \<longleftrightarrow> (A \<in> sets lebesgue \<and> 0 \<le> m \<and> lmeasure A = Real m)"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   343
proof
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   344
  assume "(indicator A has_integral m) UNIV"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   345
  with has_integral_lmeasure[OF this] has_integral_lebesgue[OF this]
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   346
  show "A \<in> sets lebesgue \<and> 0 \<le> m \<and> lmeasure A = Real m"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   347
    by (auto intro: has_integral_nonneg)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   348
next
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   349
  assume "A \<in> sets lebesgue \<and> 0 \<le> m \<and> lmeasure A = Real m"
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   350
  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
   351
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   352
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   353
lemma lmeasure_eq_integral: assumes "(indicator s::_\<Rightarrow>real) integrable_on UNIV"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   354
  shows "lmeasure s = Real (integral UNIV (indicator s))"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   355
  using assms unfolding integrable_on_def
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   356
proof safe
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   357
  fix y :: real assume "(indicator s has_integral y) UNIV"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   358
  from this[unfolded has_integral_iff_lmeasure] integral_unique[OF this]
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   359
  show "lmeasure s = Real (integral UNIV (indicator s))" by simp
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   360
qed
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   361
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   362
lemma lebesgue_simple_function_indicator:
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   363
  fixes f::"'a::ordered_euclidean_space \<Rightarrow> pextreal"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   364
  assumes f:"lebesgue.simple_function f"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   365
  shows "f = (\<lambda>x. (\<Sum>y \<in> f ` UNIV. y * indicator (f -` {y}) x))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   366
  apply(rule,subst lebesgue.simple_function_indicator_representation[OF f]) by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   367
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   368
lemma integral_eq_lmeasure:
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   369
  "(indicator s::_\<Rightarrow>real) integrable_on UNIV \<Longrightarrow> integral UNIV (indicator s) = real (lmeasure s)"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   370
  by (subst lmeasure_eq_integral) (auto intro!: integral_nonneg)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   371
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   372
lemma lmeasure_finite: assumes "(indicator s::_\<Rightarrow>real) integrable_on UNIV" shows "lmeasure s \<noteq> \<omega>"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   373
  using lmeasure_eq_integral[OF assms] by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   374
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   375
lemma negligible_iff_lebesgue_null_sets:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   376
  "negligible A \<longleftrightarrow> A \<in> lebesgue.null_sets"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   377
proof
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   378
  assume "negligible A"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   379
  from this[THEN lebesgueI_negligible] this[THEN lmeasure_eq_0]
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   380
  show "A \<in> lebesgue.null_sets" by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   381
next
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   382
  assume A: "A \<in> lebesgue.null_sets"
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   383
  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
   384
  show "negligible A" unfolding negligible_def
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   385
  proof (intro allI)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   386
    fix a b :: 'a
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   387
    have integrable: "(indicator A :: _\<Rightarrow>real) integrable_on {a..b}"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   388
      by (intro integrable_on_subinterval has_integral_integrable) (auto intro: *)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   389
    then have "integral {a..b} (indicator A) \<le> (integral UNIV (indicator A) :: real)"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   390
      using * by (auto intro!: integral_subset_le has_integral_integrable)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   391
    moreover have "(0::real) \<le> integral {a..b} (indicator A)"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   392
      using integrable by (auto intro!: integral_nonneg)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   393
    ultimately have "integral {a..b} (indicator A) = (0::real)"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   394
      using integral_unique[OF *] by auto
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   395
    then show "(indicator A has_integral (0::real)) {a..b}"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   396
      using integrable_integral[OF integrable] by simp
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   397
  qed
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   398
qed
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   399
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   400
lemma integral_const[simp]:
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   401
  fixes a b :: "'a::ordered_euclidean_space"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   402
  shows "integral {a .. b} (\<lambda>x. c) = content {a .. b} *\<^sub>R c"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   403
  by (rule integral_unique) (rule has_integral_const)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   404
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   405
lemma lmeasure_UNIV[intro]: "lmeasure (UNIV::'a::ordered_euclidean_space set) = \<omega>"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   406
  unfolding lmeasure_def SUP_\<omega>
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   407
proof (intro allI impI)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   408
  fix x assume "x < \<omega>"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   409
  then obtain r where r: "x = Real r" "0 \<le> r" by (cases x) auto
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   410
  then obtain n where n: "r < of_nat n" using ex_less_of_nat by auto
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   411
  show "\<exists>i\<in>UNIV. x < Real (integral (cube i) (indicator UNIV::'a\<Rightarrow>real))"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   412
  proof (intro bexI[of _ n])
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   413
    have [simp]: "indicator UNIV = (\<lambda>x. 1)" by (auto simp: fun_eq_iff)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   414
    { fix m :: nat assume "0 < m" then have "real n \<le> (\<Prod>x<m. 2 * real n)"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   415
      proof (induct m)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   416
        case (Suc m)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   417
        show ?case
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   418
        proof cases
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   419
          assume "m = 0" then show ?thesis by (simp add: lessThan_Suc)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   420
        next
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   421
          assume "m \<noteq> 0" then have "real n \<le> (\<Prod>x<m. 2 * real n)" using Suc by auto
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   422
          then show ?thesis
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   423
            by (auto simp: lessThan_Suc field_simps mult_le_cancel_left1)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   424
        qed
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   425
      qed auto } note this[OF DIM_positive[where 'a='a], simp]
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   426
    then have [simp]: "0 \<le> (\<Prod>x<DIM('a). 2 * real n)" using real_of_nat_ge_zero by arith
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   427
    have "x < Real (of_nat n)" using n r by auto
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   428
    also have "Real (of_nat n) \<le> Real (integral (cube n) (indicator UNIV::'a\<Rightarrow>real))"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   429
      by (auto simp: real_eq_of_nat[symmetric] cube_def content_closed_interval_cases)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   430
    finally show "x < Real (integral (cube n) (indicator UNIV::'a\<Rightarrow>real))" .
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   431
  qed auto
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   432
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   433
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   434
lemma
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   435
  fixes a b ::"'a::ordered_euclidean_space"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   436
  shows lmeasure_atLeastAtMost[simp]: "lmeasure {a..b} = Real (content {a..b})"
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   437
proof -
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   438
  have "(indicator (UNIV \<inter> {a..b})::_\<Rightarrow>real) integrable_on UNIV"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   439
    unfolding integrable_indicator_UNIV by (simp add: integrable_const indicator_def_raw)
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   440
  from lmeasure_eq_integral[OF this] show ?thesis unfolding integral_indicator_UNIV
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   441
    by (simp add: indicator_def_raw)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   442
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   443
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   444
lemma atLeastAtMost_singleton_euclidean[simp]:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   445
  fixes a :: "'a::ordered_euclidean_space" shows "{a .. a} = {a}"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   446
  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
   447
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   448
lemma content_singleton[simp]: "content {a} = 0"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   449
proof -
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   450
  have "content {a .. a} = 0"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   451
    by (subst content_closed_interval) auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   452
  then show ?thesis by simp
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   453
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   454
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   455
lemma lmeasure_singleton[simp]:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   456
  fixes a :: "'a::ordered_euclidean_space" shows "lmeasure {a} = 0"
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   457
  using lmeasure_atLeastAtMost[of a a] by simp
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   458
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   459
declare content_real[simp]
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
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   462
  fixes a b :: real
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   463
  shows lmeasure_real_greaterThanAtMost[simp]:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   464
    "lmeasure {a <.. b} = Real (if a \<le> b then b - a else 0)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   465
proof cases
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   466
  assume "a < b"
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   467
  then have "lmeasure {a <.. b} = lmeasure {a .. b} - lmeasure {a}"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   468
    by (subst lebesgue.measure_Diff[symmetric])
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   469
       (auto intro!: arg_cong[where f=lmeasure])
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   470
  then show ?thesis by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   471
qed auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   472
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   473
lemma
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   474
  fixes a b :: real
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   475
  shows lmeasure_real_atLeastLessThan[simp]:
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   476
    "lmeasure {a ..< b} = Real (if a \<le> b then b - a else 0)"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   477
proof cases
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   478
  assume "a < b"
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   479
  then have "lmeasure {a ..< b} = lmeasure {a .. b} - lmeasure {b}"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   480
    by (subst lebesgue.measure_Diff[symmetric])
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   481
       (auto intro!: arg_cong[where f=lmeasure])
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   482
  then show ?thesis by auto
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   483
qed auto
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   484
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   485
lemma
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   486
  fixes a b :: real
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   487
  shows lmeasure_real_greaterThanLessThan[simp]:
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   488
    "lmeasure {a <..< b} = Real (if a \<le> b then b - a else 0)"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   489
proof cases
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   490
  assume "a < b"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   491
  then have "lmeasure {a <..< b} = lmeasure {a <.. b} - lmeasure {b}"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   492
    by (subst lebesgue.measure_Diff[symmetric])
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   493
       (auto intro!: arg_cong[where f=lmeasure])
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   494
  then show ?thesis by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   495
qed auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   496
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   497
interpretation borel: measure_space borel lmeasure
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   498
proof
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   499
  show "countably_additive borel lmeasure"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   500
    using lebesgue.ca unfolding countably_additive_def
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   501
    apply safe apply (erule_tac x=A in allE) by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   502
qed auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   503
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   504
interpretation borel: sigma_finite_measure borel lmeasure
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   505
proof (default, intro conjI exI[of _ "\<lambda>n. cube n"])
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   506
  show "range cube \<subseteq> sets borel" by (auto intro: borel_closed)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   507
  { fix x have "\<exists>n. x\<in>cube n" using mem_big_cube by auto }
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   508
  thus "(\<Union>i. cube i) = space borel" by auto
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   509
  show "\<forall>i. lmeasure (cube i) \<noteq> \<omega>" unfolding cube_def by auto
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   510
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   511
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   512
interpretation lebesgue: sigma_finite_measure lebesgue lmeasure
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   513
proof
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   514
  from borel.sigma_finite guess A ..
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   515
  moreover then have "range A \<subseteq> sets lebesgue" using lebesgueI_borel by blast
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   516
  ultimately show "\<exists>A::nat \<Rightarrow> 'b set. range A \<subseteq> sets lebesgue \<and> (\<Union>i. A i) = space lebesgue \<and> (\<forall>i. lmeasure (A i) \<noteq> \<omega>)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   517
    by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   518
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   519
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   520
lemma simple_function_has_integral:
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   521
  fixes f::"'a::ordered_euclidean_space \<Rightarrow> pextreal"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   522
  assumes f:"lebesgue.simple_function f"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   523
  and f':"\<forall>x. f x \<noteq> \<omega>"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   524
  and om:"\<forall>x\<in>range f. lmeasure (f -` {x} \<inter> UNIV) = \<omega> \<longrightarrow> x = 0"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   525
  shows "((\<lambda>x. real (f x)) has_integral (real (lebesgue.simple_integral f))) UNIV"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   526
  unfolding lebesgue.simple_integral_def
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   527
  apply(subst lebesgue_simple_function_indicator[OF f])
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   528
proof -
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   529
  case goal1
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   530
  have *:"\<And>x. \<forall>y\<in>range f. y * indicator (f -` {y}) x \<noteq> \<omega>"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   531
    "\<forall>x\<in>range f. x * lmeasure (f -` {x} \<inter> UNIV) \<noteq> \<omega>"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   532
    using f' om unfolding indicator_def by auto
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   533
  show ?case unfolding space_lebesgue real_of_pextreal_setsum'[OF *(1),THEN sym]
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   534
    unfolding real_of_pextreal_setsum'[OF *(2),THEN sym]
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   535
    unfolding real_of_pextreal_setsum space_lebesgue
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   536
    apply(rule has_integral_setsum)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   537
  proof safe show "finite (range f)" using f by (auto dest: lebesgue.simple_functionD)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   538
    fix y::'a show "((\<lambda>x. real (f y * indicator (f -` {f y}) x)) has_integral
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   539
      real (f y * lmeasure (f -` {f y} \<inter> UNIV))) UNIV"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   540
    proof(cases "f y = 0") case False
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   541
      have mea:"(indicator (f -` {f y}) ::_\<Rightarrow>real) integrable_on UNIV"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   542
        apply(rule lmeasure_finite_integrable)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   543
        using assms unfolding lebesgue.simple_function_def using False by auto
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   544
      have *:"\<And>x. real (indicator (f -` {f y}) x::pextreal) = (indicator (f -` {f y}) x)"
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   545
        by (auto simp: indicator_def)
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   546
      show ?thesis unfolding real_of_pextreal_mult[THEN sym]
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   547
        apply(rule has_integral_cmul[where 'b=real, unfolded real_scaleR_def])
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   548
        unfolding Int_UNIV_right lmeasure_eq_integral[OF mea,THEN sym]
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   549
        unfolding integral_eq_lmeasure[OF mea, symmetric] *
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   550
        apply(rule integrable_integral) using mea .
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   551
    qed auto
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   552
  qed
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   553
qed
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   554
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   555
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
   556
  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
   557
  using assms by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   558
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   559
lemma simple_function_has_integral':
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   560
  fixes f::"'a::ordered_euclidean_space \<Rightarrow> pextreal"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   561
  assumes f:"lebesgue.simple_function f"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   562
  and i: "lebesgue.simple_integral f \<noteq> \<omega>"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   563
  shows "((\<lambda>x. real (f x)) has_integral (real (lebesgue.simple_integral f))) UNIV"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   564
proof- let ?f = "\<lambda>x. if f x = \<omega> then 0 else f x"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   565
  { fix x have "real (f x) = real (?f x)" by (cases "f x") auto } note * = this
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   566
  have **:"{x. f x \<noteq> ?f x} = f -` {\<omega>}" by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   567
  have **:"lmeasure {x\<in>space lebesgue. f x \<noteq> ?f x} = 0"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   568
    using lebesgue.simple_integral_omega[OF assms] by(auto simp add:**)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   569
  show ?thesis apply(subst lebesgue.simple_integral_cong'[OF f _ **])
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   570
    apply(rule lebesgue.simple_function_compose1[OF f])
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   571
    unfolding * defer apply(rule simple_function_has_integral)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   572
  proof-
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   573
    show "lebesgue.simple_function ?f"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   574
      using lebesgue.simple_function_compose1[OF f] .
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   575
    show "\<forall>x. ?f x \<noteq> \<omega>" by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   576
    show "\<forall>x\<in>range ?f. lmeasure (?f -` {x} \<inter> UNIV) = \<omega> \<longrightarrow> x = 0"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   577
    proof (safe, simp, safe, rule ccontr)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   578
      fix y assume "f y \<noteq> \<omega>" "f y \<noteq> 0"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   579
      hence "(\<lambda>x. if f x = \<omega> then 0 else f x) -` {if f y = \<omega> then 0 else f y} = f -` {f y}"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   580
        by (auto split: split_if_asm)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   581
      moreover assume "lmeasure ((\<lambda>x. if f x = \<omega> then 0 else f x) -` {if f y = \<omega> then 0 else f y}) = \<omega>"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   582
      ultimately have "lmeasure (f -` {f y}) = \<omega>" by simp
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   583
      moreover
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   584
      have "f y * lmeasure (f -` {f y}) \<noteq> \<omega>" using i f
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   585
        unfolding lebesgue.simple_integral_def setsum_\<omega> lebesgue.simple_function_def
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   586
        by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   587
      ultimately have "f y = 0" by (auto split: split_if_asm)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   588
      then show False using `f y \<noteq> 0` by simp
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   589
    qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   590
  qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   591
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   592
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   593
lemma (in measure_space) positive_integral_monotone_convergence:
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   594
  fixes f::"nat \<Rightarrow> 'a \<Rightarrow> pextreal"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   595
  assumes i: "\<And>i. f i \<in> borel_measurable M" and mono: "\<And>x. mono (\<lambda>n. f n x)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   596
  and lim: "\<And>x. (\<lambda>i. f i x) ----> u x"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   597
  shows "u \<in> borel_measurable M"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   598
  and "(\<lambda>i. positive_integral (f i)) ----> positive_integral u" (is ?ilim)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   599
proof -
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   600
  from positive_integral_isoton[unfolded isoton_fun_expand isoton_iff_Lim_mono, of f u]
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   601
  show ?ilim using mono lim i by auto
41097
a1abfa4e2b44 use SUPR_ and INFI_apply instead of SUPR_, INFI_fun_expand
hoelzl
parents: 41095
diff changeset
   602
  have "\<And>x. (SUP i. f i x) = u x" using mono lim SUP_Lim_pextreal
a1abfa4e2b44 use SUPR_ and INFI_apply instead of SUPR_, INFI_fun_expand
hoelzl
parents: 41095
diff changeset
   603
    unfolding fun_eq_iff mono_def by auto
a1abfa4e2b44 use SUPR_ and INFI_apply instead of SUPR_, INFI_fun_expand
hoelzl
parents: 41095
diff changeset
   604
  moreover have "(\<lambda>x. SUP i. f i x) \<in> borel_measurable M"
a1abfa4e2b44 use SUPR_ and INFI_apply instead of SUPR_, INFI_fun_expand
hoelzl
parents: 41095
diff changeset
   605
    using i by auto
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   606
  ultimately show "u \<in> borel_measurable M" by simp
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   607
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   608
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   609
lemma positive_integral_has_integral:
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   610
  fixes f::"'a::ordered_euclidean_space => pextreal"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   611
  assumes f:"f \<in> borel_measurable lebesgue"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   612
  and int_om:"lebesgue.positive_integral f \<noteq> \<omega>"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   613
  and f_om:"\<forall>x. f x \<noteq> \<omega>" (* TODO: remove this *)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   614
  shows "((\<lambda>x. real (f x)) has_integral (real (lebesgue.positive_integral f))) UNIV"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   615
proof- let ?i = "lebesgue.positive_integral f"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   616
  from lebesgue.borel_measurable_implies_simple_function_sequence[OF f]
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   617
  guess u .. note conjunctD2[OF this,rule_format] note u = conjunctD2[OF this(1)] this(2)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   618
  let ?u = "\<lambda>i x. real (u i x)" and ?f = "\<lambda>x. real (f x)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   619
  have u_simple:"\<And>k. lebesgue.simple_integral (u k) = lebesgue.positive_integral (u k)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   620
    apply(subst lebesgue.positive_integral_eq_simple_integral[THEN sym,OF u(1)]) ..
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   621
  have int_u_le:"\<And>k. lebesgue.simple_integral (u k) \<le> lebesgue.positive_integral f"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   622
    unfolding u_simple apply(rule lebesgue.positive_integral_mono)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   623
    using isoton_Sup[OF u(3)] unfolding le_fun_def by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   624
  have u_int_om:"\<And>i. lebesgue.simple_integral (u i) \<noteq> \<omega>"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   625
  proof- case goal1 thus ?case using int_u_le[of i] int_om by auto qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   626
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   627
  note u_int = simple_function_has_integral'[OF u(1) this]
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   628
  have "(\<lambda>x. real (f x)) integrable_on UNIV \<and>
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   629
    (\<lambda>k. Integration.integral UNIV (\<lambda>x. real (u k x))) ----> Integration.integral UNIV (\<lambda>x. real (f x))"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   630
    apply(rule monotone_convergence_increasing) apply(rule,rule,rule u_int)
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   631
  proof safe case goal1 show ?case apply(rule real_of_pextreal_mono) using u(2,3) by auto
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   632
  next case goal2 show ?case using u(3) apply(subst lim_Real[THEN sym])
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   633
      prefer 3 apply(subst Real_real') defer apply(subst Real_real')
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   634
      using isotone_Lim[OF u(3)[unfolded isoton_fun_expand, THEN spec]] using f_om u by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   635
  next case goal3
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   636
    show ?case apply(rule bounded_realI[where B="real (lebesgue.positive_integral f)"])
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   637
      apply safe apply(subst abs_of_nonneg) apply(rule integral_nonneg,rule) apply(rule u_int)
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   638
      unfolding integral_unique[OF u_int] defer apply(rule real_of_pextreal_mono[OF _ int_u_le])
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   639
      using u int_om by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   640
  qed note int = conjunctD2[OF this]
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   641
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   642
  have "(\<lambda>i. lebesgue.simple_integral (u i)) ----> ?i" unfolding u_simple
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   643
    apply(rule lebesgue.positive_integral_monotone_convergence(2))
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   644
    apply(rule lebesgue.borel_measurable_simple_function[OF u(1)])
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   645
    using isotone_Lim[OF u(3)[unfolded isoton_fun_expand, THEN spec]] by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   646
  hence "(\<lambda>i. real (lebesgue.simple_integral (u i))) ----> real ?i" apply-
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   647
    apply(subst lim_Real[THEN sym]) prefer 3
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   648
    apply(subst Real_real') defer apply(subst Real_real')
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   649
    using u f_om int_om u_int_om by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   650
  note * = LIMSEQ_unique[OF this int(2)[unfolded integral_unique[OF u_int]]]
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   651
  show ?thesis unfolding * by(rule integrable_integral[OF int(1)])
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   652
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   653
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   654
lemma lebesgue_integral_has_integral:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   655
  fixes f::"'a::ordered_euclidean_space => real"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   656
  assumes f:"lebesgue.integrable f"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   657
  shows "(f has_integral (lebesgue.integral f)) UNIV"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   658
proof- let ?n = "\<lambda>x. - min (f x) 0" and ?p = "\<lambda>x. max (f x) 0"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   659
  have *:"f = (\<lambda>x. ?p x - ?n x)" apply rule by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   660
  note f = lebesgue.integrableD[OF f]
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   661
  show ?thesis unfolding lebesgue.integral_def apply(subst *)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   662
  proof(rule has_integral_sub) case goal1
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   663
    have *:"\<forall>x. Real (f x) \<noteq> \<omega>" by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   664
    note lebesgue.borel_measurable_Real[OF f(1)]
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   665
    from positive_integral_has_integral[OF this f(2) *]
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   666
    show ?case unfolding real_Real_max .
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   667
  next case goal2
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   668
    have *:"\<forall>x. Real (- f x) \<noteq> \<omega>" by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   669
    note lebesgue.borel_measurable_uminus[OF f(1)]
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   670
    note lebesgue.borel_measurable_Real[OF this]
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   671
    from positive_integral_has_integral[OF this f(3) *]
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   672
    show ?case unfolding real_Real_max minus_min_eq_max by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   673
  qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   674
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   675
41546
2a12c23b7a34 integral on lebesgue measure is extension of integral on borel measure
hoelzl
parents: 41097
diff changeset
   676
lemma lebesgue_positive_integral_eq_borel:
2a12c23b7a34 integral on lebesgue measure is extension of integral on borel measure
hoelzl
parents: 41097
diff changeset
   677
  "f \<in> borel_measurable borel \<Longrightarrow> lebesgue.positive_integral f = borel.positive_integral f "
2a12c23b7a34 integral on lebesgue measure is extension of integral on borel measure
hoelzl
parents: 41097
diff changeset
   678
  by (auto intro!: lebesgue.positive_integral_subalgebra[symmetric]) default
2a12c23b7a34 integral on lebesgue measure is extension of integral on borel measure
hoelzl
parents: 41097
diff changeset
   679
2a12c23b7a34 integral on lebesgue measure is extension of integral on borel measure
hoelzl
parents: 41097
diff changeset
   680
lemma lebesgue_integral_eq_borel:
2a12c23b7a34 integral on lebesgue measure is extension of integral on borel measure
hoelzl
parents: 41097
diff changeset
   681
  assumes "f \<in> borel_measurable borel"
2a12c23b7a34 integral on lebesgue measure is extension of integral on borel measure
hoelzl
parents: 41097
diff changeset
   682
  shows "lebesgue.integrable f = borel.integrable f" (is ?P)
2a12c23b7a34 integral on lebesgue measure is extension of integral on borel measure
hoelzl
parents: 41097
diff changeset
   683
    and "lebesgue.integral f = borel.integral f" (is ?I)
2a12c23b7a34 integral on lebesgue measure is extension of integral on borel measure
hoelzl
parents: 41097
diff changeset
   684
proof -
2a12c23b7a34 integral on lebesgue measure is extension of integral on borel measure
hoelzl
parents: 41097
diff changeset
   685
  have *: "sigma_algebra borel" by default
2a12c23b7a34 integral on lebesgue measure is extension of integral on borel measure
hoelzl
parents: 41097
diff changeset
   686
  have "sets borel \<subseteq> sets lebesgue" by auto
2a12c23b7a34 integral on lebesgue measure is extension of integral on borel measure
hoelzl
parents: 41097
diff changeset
   687
  from lebesgue.integral_subalgebra[OF assms this _ *]
2a12c23b7a34 integral on lebesgue measure is extension of integral on borel measure
hoelzl
parents: 41097
diff changeset
   688
  show ?P ?I by auto
2a12c23b7a34 integral on lebesgue measure is extension of integral on borel measure
hoelzl
parents: 41097
diff changeset
   689
qed
2a12c23b7a34 integral on lebesgue measure is extension of integral on borel measure
hoelzl
parents: 41097
diff changeset
   690
2a12c23b7a34 integral on lebesgue measure is extension of integral on borel measure
hoelzl
parents: 41097
diff changeset
   691
lemma borel_integral_has_integral:
2a12c23b7a34 integral on lebesgue measure is extension of integral on borel measure
hoelzl
parents: 41097
diff changeset
   692
  fixes f::"'a::ordered_euclidean_space => real"
2a12c23b7a34 integral on lebesgue measure is extension of integral on borel measure
hoelzl
parents: 41097
diff changeset
   693
  assumes f:"borel.integrable f"
2a12c23b7a34 integral on lebesgue measure is extension of integral on borel measure
hoelzl
parents: 41097
diff changeset
   694
  shows "(f has_integral (borel.integral f)) UNIV"
2a12c23b7a34 integral on lebesgue measure is extension of integral on borel measure
hoelzl
parents: 41097
diff changeset
   695
proof -
2a12c23b7a34 integral on lebesgue measure is extension of integral on borel measure
hoelzl
parents: 41097
diff changeset
   696
  have borel: "f \<in> borel_measurable borel"
2a12c23b7a34 integral on lebesgue measure is extension of integral on borel measure
hoelzl
parents: 41097
diff changeset
   697
    using f unfolding borel.integrable_def by auto
2a12c23b7a34 integral on lebesgue measure is extension of integral on borel measure
hoelzl
parents: 41097
diff changeset
   698
  from f show ?thesis
2a12c23b7a34 integral on lebesgue measure is extension of integral on borel measure
hoelzl
parents: 41097
diff changeset
   699
    using lebesgue_integral_has_integral[of f]
2a12c23b7a34 integral on lebesgue measure is extension of integral on borel measure
hoelzl
parents: 41097
diff changeset
   700
    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
   701
qed
2a12c23b7a34 integral on lebesgue measure is extension of integral on borel measure
hoelzl
parents: 41097
diff changeset
   702
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   703
lemma continuous_on_imp_borel_measurable:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   704
  fixes f::"'a::ordered_euclidean_space \<Rightarrow> 'b::ordered_euclidean_space"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   705
  assumes "continuous_on UNIV f"
41546
2a12c23b7a34 integral on lebesgue measure is extension of integral on borel measure
hoelzl
parents: 41097
diff changeset
   706
  shows "f \<in> borel_measurable borel"
2a12c23b7a34 integral on lebesgue measure is extension of integral on borel measure
hoelzl
parents: 41097
diff changeset
   707
  apply(rule borel.borel_measurableI)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   708
  using continuous_open_preimage[OF assms] unfolding vimage_def by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   709
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   710
lemma (in measure_space) integral_monotone_convergence_pos':
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   711
  assumes i: "\<And>i. integrable (f i)" and mono: "\<And>x. mono (\<lambda>n. f n x)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   712
  and pos: "\<And>x i. 0 \<le> f i x"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   713
  and lim: "\<And>x. (\<lambda>i. f i x) ----> u x"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   714
  and ilim: "(\<lambda>i. integral (f i)) ----> x"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   715
  shows "integrable u \<and> integral u = x"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   716
  using integral_monotone_convergence_pos[OF assms] by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   717
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   718
definition e2p :: "'a::ordered_euclidean_space \<Rightarrow> (nat \<Rightarrow> real)" where
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   719
  "e2p x = (\<lambda>i\<in>{..<DIM('a)}. x$$i)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   720
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   721
definition p2e :: "(nat \<Rightarrow> real) \<Rightarrow> 'a::ordered_euclidean_space" where
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   722
  "p2e x = (\<chi>\<chi> i. x i)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   723
41095
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   724
lemma e2p_p2e[simp]:
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   725
  "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
   726
  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
   727
41095
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   728
lemma p2e_e2p[simp]:
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   729
  "p2e (e2p x) = (x::'a::ordered_euclidean_space)"
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   730
  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
   731
41095
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   732
lemma bij_inv_p2e_e2p:
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   733
  shows "bij_inv ({..<DIM('a)} \<rightarrow>\<^isub>E UNIV) (UNIV :: 'a::ordered_euclidean_space set)
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   734
     p2e e2p" (is "bij_inv ?P ?U _ _")
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   735
proof (rule bij_invI)
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   736
  show "p2e \<in> ?P \<rightarrow> ?U" "e2p \<in> ?U \<rightarrow> ?P" by (auto simp: e2p_def)
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   737
qed auto
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   738
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   739
interpretation borel_product: product_sigma_finite "\<lambda>x. borel::real algebra" "\<lambda>x. lmeasure"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   740
  by default
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   741
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   742
lemma cube_subset_Suc[intro]: "cube n \<subseteq> cube (Suc n)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   743
  unfolding cube_def_raw subset_eq apply safe unfolding mem_interval by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   744
41095
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   745
lemma Pi_iff: "f \<in> Pi I X \<longleftrightarrow> (\<forall>i\<in>I. f i \<in> X i)"
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   746
  unfolding Pi_def by auto
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   747
41095
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   748
lemma measurable_e2p_on_generator:
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   749
  "e2p \<in> measurable \<lparr> space = UNIV::'a set, sets = range lessThan \<rparr>
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   750
  (product_algebra
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   751
    (\<lambda>x. \<lparr> space = UNIV::real set, sets = range lessThan \<rparr>)
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   752
    {..<DIM('a::ordered_euclidean_space)})"
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   753
  (is "e2p \<in> measurable ?E ?P")
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   754
proof (unfold measurable_def, intro CollectI conjI ballI)
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   755
  show "e2p \<in> space ?E \<rightarrow> space ?P" by (auto simp: e2p_def)
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   756
  fix A assume "A \<in> sets ?P"
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   757
  then obtain E where A: "A = (\<Pi>\<^isub>E i\<in>{..<DIM('a)}. E i)"
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   758
    and "E \<in> {..<DIM('a)} \<rightarrow> (range lessThan)"
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   759
    by (auto elim!: product_algebraE)
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   760
  then have "\<forall>i\<in>{..<DIM('a)}. \<exists>xs. E i = {..< xs}" by auto
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   761
  from this[THEN bchoice] guess xs ..
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   762
  then have A_eq: "A = (\<Pi>\<^isub>E i\<in>{..<DIM('a)}. {..< xs i})"
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   763
    using A by auto
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   764
  have "e2p -` A = {..< (\<chi>\<chi> i. xs i) :: 'a}"
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   765
    using DIM_positive by (auto simp add: Pi_iff set_eq_iff e2p_def A_eq
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   766
      euclidean_eq[where 'a='a] eucl_less[where 'a='a])
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   767
  then show "e2p -` A \<inter> space ?E \<in> sets ?E" by simp
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   768
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   769
41095
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   770
lemma measurable_p2e_on_generator:
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   771
  "p2e \<in> measurable
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   772
    (product_algebra
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   773
      (\<lambda>x. \<lparr> space = UNIV::real set, sets = range lessThan \<rparr>)
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   774
      {..<DIM('a::ordered_euclidean_space)})
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   775
    \<lparr> space = UNIV::'a set, sets = range lessThan \<rparr>"
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   776
  (is "p2e \<in> measurable ?P ?E")
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   777
proof (unfold measurable_def, intro CollectI conjI ballI)
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   778
  show "p2e \<in> space ?P \<rightarrow> space ?E" by simp
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   779
  fix A assume "A \<in> sets ?E"
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   780
  then obtain x where "A = {..<x}" by auto
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   781
  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
   782
    using DIM_positive
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   783
    by (auto simp: Pi_iff set_eq_iff p2e_def
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   784
                   euclidean_eq[where 'a='a] eucl_less[where 'a='a])
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   785
  then show "p2e -` A \<inter> space ?P \<in> sets ?P" by auto
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   786
qed
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   787
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   788
lemma borel_vimage_algebra_eq:
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   789
  defines "F \<equiv> product_algebra (\<lambda>x. \<lparr> space = (UNIV::real set), sets = range lessThan \<rparr>) {..<DIM('a::ordered_euclidean_space)}"
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   790
  shows "sigma_algebra.vimage_algebra (borel::'a::ordered_euclidean_space algebra) (space (sigma F)) p2e = sigma F"
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   791
  unfolding borel_eq_lessThan
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   792
proof (intro vimage_algebra_sigma)
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   793
  let ?E = "\<lparr>space = (UNIV::'a set), sets = range lessThan\<rparr>"
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   794
  show "bij_inv (space (sigma F)) (space (sigma ?E)) p2e e2p"
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   795
    using bij_inv_p2e_e2p unfolding F_def by simp
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   796
  show "sets F \<subseteq> Pow (space F)" "sets ?E \<subseteq> Pow (space ?E)" unfolding F_def
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   797
    by (intro product_algebra_sets_into_space) auto
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   798
  show "p2e \<in> measurable F ?E"
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   799
    "e2p \<in> measurable ?E F"
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   800
    unfolding F_def using measurable_p2e_on_generator measurable_e2p_on_generator by auto
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   801
qed
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   802
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   803
lemma product_borel_eq_vimage:
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   804
  "sigma (product_algebra (\<lambda>x. borel) {..<DIM('a::ordered_euclidean_space)}) =
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   805
  sigma_algebra.vimage_algebra borel (extensional {..<DIM('a)})
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   806
  (p2e:: _ \<Rightarrow> 'a::ordered_euclidean_space)"
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   807
  unfolding borel_vimage_algebra_eq[simplified]
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   808
  unfolding borel_eq_lessThan
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   809
  apply(subst sigma_product_algebra_sigma_eq[where S="\<lambda>i. \<lambda>n. lessThan (real n)"])
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   810
  unfolding lessThan_iff
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   811
proof- fix i assume i:"i<DIM('a)"
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   812
  show "(\<lambda>n. {..<real n}) \<up> space \<lparr>space = UNIV, sets = range lessThan\<rparr>"
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   813
    by(auto intro!:real_arch_lt isotoneI)
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   814
qed auto
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   815
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   816
lemma e2p_Int:"e2p ` A \<inter> e2p ` B = e2p ` (A \<inter> B)" (is "?L = ?R")
41095
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   817
  apply(rule image_Int[THEN sym])
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   818
  using bij_inv_p2e_e2p[THEN bij_inv_bij_betw(2)]
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   819
  unfolding bij_betw_def by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   820
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   821
lemma Int_stable_cuboids: fixes x::"'a::ordered_euclidean_space"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   822
  shows "Int_stable \<lparr>space = UNIV, sets = range (\<lambda>(a, b::'a). e2p ` {a..b})\<rparr>"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   823
  unfolding Int_stable_def algebra.select_convs
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   824
proof safe fix a b x y::'a
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   825
  have *:"e2p ` {a..b} \<inter> e2p ` {x..y} =
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   826
    (\<lambda>(a, b). e2p ` {a..b}) (\<chi>\<chi> i. max (a $$ i) (x $$ i), \<chi>\<chi> i. min (b $$ i) (y $$ i)::'a)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   827
    unfolding e2p_Int inter_interval by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   828
  show "e2p ` {a..b} \<inter> e2p ` {x..y} \<in> range (\<lambda>(a, b). e2p ` {a..b::'a})" unfolding *
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   829
    apply(rule range_eqI) ..
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   830
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   831
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   832
lemma Int_stable_cuboids': fixes x::"'a::ordered_euclidean_space"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   833
  shows "Int_stable \<lparr>space = UNIV, sets = range (\<lambda>(a, b::'a). {a..b})\<rparr>"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   834
  unfolding Int_stable_def algebra.select_convs
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   835
  apply safe unfolding inter_interval by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   836
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   837
lemma inj_on_disjoint_family_on: assumes "disjoint_family_on A S" "inj f"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   838
  shows "disjoint_family_on (\<lambda>x. f ` A x) S"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   839
  unfolding disjoint_family_on_def
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   840
proof(rule,rule,rule)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   841
  fix x1 x2 assume x:"x1 \<in> S" "x2 \<in> S" "x1 \<noteq> x2"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   842
  show "f ` A x1 \<inter> f ` A x2 = {}"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   843
  proof(rule ccontr) case goal1
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   844
    then obtain z where z:"z \<in> f ` A x1 \<inter> f ` A x2" by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   845
    then obtain z1 z2 where z12:"z1 \<in> A x1" "z2 \<in> A x2" "f z1 = z" "f z2 = z" by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   846
    hence "z1 = z2" using assms(2) unfolding inj_on_def by blast
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   847
    hence "x1 = x2" using z12(1-2) using assms[unfolded disjoint_family_on_def] using x by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   848
    thus False using x(3) by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   849
  qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   850
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   851
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   852
declare restrict_extensional[intro]
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   853
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   854
lemma e2p_extensional[intro]:"e2p (y::'a::ordered_euclidean_space) \<in> extensional {..<DIM('a)}"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   855
  unfolding e2p_def by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   856
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   857
lemma e2p_image_vimage: fixes A::"'a::ordered_euclidean_space set"
41095
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   858
  shows "e2p ` A = p2e -` A \<inter> extensional {..<DIM('a)}"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   859
proof(rule set_eqI,rule)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   860
  fix x assume "x \<in> e2p ` A" then guess y unfolding image_iff .. note y=this
41095
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   861
  show "x \<in> p2e -` A \<inter> extensional {..<DIM('a)}"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   862
    apply safe apply(rule vimageI[OF _ y(1)]) unfolding y p2e_e2p by auto
41095
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   863
next fix x assume "x \<in> p2e -` A \<inter> extensional {..<DIM('a)}"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   864
  thus "x \<in> e2p ` A" unfolding image_iff apply(rule_tac x="p2e x" in bexI) apply(subst e2p_p2e) by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   865
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   866
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   867
lemma lmeasure_measure_eq_borel_prod:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   868
  fixes A :: "('a::ordered_euclidean_space) set"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   869
  assumes "A \<in> sets borel"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   870
  shows "lmeasure A = borel_product.product_measure {..<DIM('a)} (e2p ` A :: (nat \<Rightarrow> real) set)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   871
proof (rule measure_unique_Int_stable[where X=A and A=cube])
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   872
  interpret fprod: finite_product_sigma_finite "\<lambda>x. borel" "\<lambda>x. lmeasure" "{..<DIM('a)}" by default auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   873
  show "Int_stable \<lparr> space = UNIV :: 'a set, sets = range (\<lambda>(a,b). {a..b}) \<rparr>"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   874
    (is "Int_stable ?E" ) using Int_stable_cuboids' .
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   875
  show "borel = sigma ?E" using borel_eq_atLeastAtMost .
41654
32fe42892983 Gauge measure removed
hoelzl
parents: 41546
diff changeset
   876
  show "\<And>i. lmeasure (cube i) \<noteq> \<omega>" unfolding cube_def by auto
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   877
  show "\<And>X. X \<in> sets ?E \<Longrightarrow>
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   878
    lmeasure X = borel_product.product_measure {..<DIM('a)} (e2p ` X :: (nat \<Rightarrow> real) set)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   879
  proof- case goal1 then obtain a b where X:"X = {a..b}" by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   880
    { presume *:"X \<noteq> {} \<Longrightarrow> ?case"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   881
      show ?case apply(cases,rule *,assumption) by auto }
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   882
    def XX \<equiv> "\<lambda>i. {a $$ i .. b $$ i}" assume  "X \<noteq> {}"  note X' = this[unfolded X interval_ne_empty]
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   883
    have *:"Pi\<^isub>E {..<DIM('a)} XX = e2p ` X" apply(rule set_eqI)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   884
    proof fix x assume "x \<in> Pi\<^isub>E {..<DIM('a)} XX"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   885
      thus "x \<in> e2p ` X" unfolding image_iff apply(rule_tac x="\<chi>\<chi> i. x i" in bexI)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   886
        unfolding Pi_def extensional_def e2p_def restrict_def X mem_interval XX_def by rule auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   887
    next fix x assume "x \<in> e2p ` X" then guess y unfolding image_iff .. note y = this
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   888
      show "x \<in> Pi\<^isub>E {..<DIM('a)} XX" unfolding y using y(1)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   889
        unfolding Pi_def extensional_def e2p_def restrict_def X mem_interval XX_def by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   890
    qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   891
    have "lmeasure X = (\<Prod>x<DIM('a). Real (b $$ x - a $$ x))"  using X' apply- unfolding X
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   892
      unfolding lmeasure_atLeastAtMost content_closed_interval apply(subst Real_setprod) by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   893
    also have "... = (\<Prod>i<DIM('a). lmeasure (XX i))" apply(rule setprod_cong2)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   894
      unfolding XX_def lmeasure_atLeastAtMost apply(subst content_real) using X' by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   895
    also have "... = borel_product.product_measure {..<DIM('a)} (e2p ` X)" unfolding *[THEN sym]
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   896
      apply(rule fprod.measure_times[THEN sym]) unfolding XX_def by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   897
    finally show ?case .
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   898
  qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   899
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   900
  show "range cube \<subseteq> sets \<lparr>space = UNIV, sets = range (\<lambda>(a, b). {a..b})\<rparr>"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   901
    unfolding cube_def_raw by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   902
  have "\<And>x. \<exists>xa. x \<in> cube xa" apply(rule_tac x=x in mem_big_cube) by fastsimp
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   903
  thus "cube \<up> space \<lparr>space = UNIV, sets = range (\<lambda>(a, b). {a..b})\<rparr>"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   904
    apply-apply(rule isotoneI) apply(rule cube_subset_Suc) by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   905
  show "A \<in> sets borel " by fact
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   906
  show "measure_space borel lmeasure" by default
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   907
  show "measure_space borel
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   908
     (\<lambda>a::'a set. finite_product_sigma_finite.measure (\<lambda>x. borel) (\<lambda>x. lmeasure) {..<DIM('a)} (e2p ` a))"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   909
    apply default unfolding countably_additive_def
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   910
  proof safe fix A::"nat \<Rightarrow> 'a set" assume A:"range A \<subseteq> sets borel" "disjoint_family A"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   911
      "(\<Union>i. A i) \<in> sets borel"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   912
    note fprod.ca[unfolded countably_additive_def,rule_format]
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   913
    note ca = this[of "\<lambda> n. e2p ` (A n)"]
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   914
    show "(\<Sum>\<^isub>\<infinity>n. finite_product_sigma_finite.measure
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   915
        (\<lambda>x. borel) (\<lambda>x. lmeasure) {..<DIM('a)} (e2p ` A n)) =
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   916
           finite_product_sigma_finite.measure (\<lambda>x. borel)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   917
            (\<lambda>x. lmeasure) {..<DIM('a)} (e2p ` (\<Union>i. A i))" unfolding image_UN
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   918
    proof(rule ca) show "range (\<lambda>n. e2p ` A n) \<subseteq> sets
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   919
       (sigma (product_algebra (\<lambda>x. borel) {..<DIM('a)}))"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   920
        unfolding product_borel_eq_vimage
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   921
      proof case goal1
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   922
        then guess y unfolding image_iff .. note y=this(2)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   923
        show ?case unfolding borel.in_vimage_algebra y apply-
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   924
          apply(rule_tac x="A y" in bexI,rule e2p_image_vimage)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   925
          using A(1) by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   926
      qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   927
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   928
      show "disjoint_family (\<lambda>n. e2p ` A n)" apply(rule inj_on_disjoint_family_on)
41095
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   929
        using bij_inv_p2e_e2p[THEN bij_inv_bij_betw(2)] using A(2) unfolding bij_betw_def by auto
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   930
      show "(\<Union>n. e2p ` A n) \<in> sets (sigma (product_algebra (\<lambda>x. borel) {..<DIM('a)}))"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   931
        unfolding product_borel_eq_vimage borel.in_vimage_algebra
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   932
      proof(rule bexI[OF _ A(3)],rule set_eqI,rule)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   933
        fix x assume x:"x \<in> (\<Union>n. e2p ` A n)" hence "p2e x \<in> (\<Union>i. A i)" by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   934
        moreover have "x \<in> extensional {..<DIM('a)}"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   935
          using x unfolding extensional_def e2p_def_raw by auto
41095
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   936
        ultimately show "x \<in> p2e -` (\<Union>i. A i) \<inter> extensional {..<DIM('a)}" by auto
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   937
      next fix x assume x:"x \<in> p2e -` (\<Union>i. A i) \<inter> extensional {..<DIM('a)}"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   938
        hence "p2e x \<in> (\<Union>i. A i)" by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   939
        hence "\<exists>n. x \<in> e2p ` A n" apply safe apply(rule_tac x=i in exI)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   940
          unfolding image_iff apply(rule_tac x="p2e x" in bexI)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   941
          apply(subst e2p_p2e) using x by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   942
        thus "x \<in> (\<Union>n. e2p ` A n)" by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   943
      qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   944
    qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   945
  qed auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   946
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   947
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   948
lemma e2p_p2e'[simp]: fixes x::"'a::ordered_euclidean_space"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   949
  assumes "A \<subseteq> extensional {..<DIM('a)}"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   950
  shows "e2p ` (p2e ` A ::'a set) = A"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   951
  apply(rule set_eqI) unfolding image_iff Bex_def apply safe defer
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   952
  apply(rule_tac x="p2e x" in exI,safe) using assms by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   953
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   954
lemma range_p2e:"range (p2e::_\<Rightarrow>'a::ordered_euclidean_space) = UNIV"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   955
  apply safe defer unfolding image_iff apply(rule_tac x="\<lambda>i. x $$ i" in bexI)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   956
  unfolding p2e_def by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   957
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   958
lemma p2e_inv_extensional:"(A::'a::ordered_euclidean_space set)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   959
  = p2e ` (p2e -` A \<inter> extensional {..<DIM('a)})"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   960
  unfolding p2e_def_raw apply safe unfolding image_iff
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   961
proof- fix x assume "x\<in>A"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   962
  let ?y = "\<lambda>i. if i<DIM('a) then x$$i else undefined"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   963
  have *:"Chi ?y = x" apply(subst euclidean_eq) by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   964
  show "\<exists>xa\<in>Chi -` A \<inter> extensional {..<DIM('a)}. x = Chi xa" apply(rule_tac x="?y" in bexI)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   965
    apply(subst euclidean_eq) unfolding extensional_def using `x\<in>A` by(auto simp: *)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   966
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   967
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   968
lemma borel_fubini_positiv_integral:
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   969
  fixes f :: "'a::ordered_euclidean_space \<Rightarrow> pextreal"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   970
  assumes f: "f \<in> borel_measurable borel"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   971
  shows "borel.positive_integral f =
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   972
          borel_product.product_positive_integral {..<DIM('a)} (f \<circ> p2e)"
41095
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   973
proof- def U \<equiv> "extensional {..<DIM('a)} :: (nat \<Rightarrow> real) set"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   974
  interpret fprod: finite_product_sigma_finite "\<lambda>x. borel" "\<lambda>x. lmeasure" "{..<DIM('a)}" by default auto
41095
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   975
  have *:"sigma_algebra.vimage_algebra borel U (p2e:: _ \<Rightarrow> 'a)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   976
    = sigma (product_algebra (\<lambda>x. borel) {..<DIM('a)})"
41095
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   977
    unfolding U_def product_borel_eq_vimage[symmetric] ..
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   978
  show ?thesis
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   979
    unfolding borel.positive_integral_vimage[unfolded space_borel, OF bij_inv_p2e_e2p[THEN bij_inv_bij_betw(1)]]
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   980
    apply(subst fprod.positive_integral_cong_measure[THEN sym, of "\<lambda>A. lmeasure (p2e ` A)"])
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   981
    unfolding U_def[symmetric] *[THEN sym] o_def
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   982
  proof- fix A assume A:"A \<in> sets (sigma_algebra.vimage_algebra borel U (p2e ::_ \<Rightarrow> 'a))"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   983
    hence *:"A \<subseteq> extensional {..<DIM('a)}" unfolding U_def by auto
41095
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   984
    from A guess B unfolding borel.in_vimage_algebra U_def ..
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   985
    then have "(p2e ` A::'a set) \<in> sets borel"
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   986
      by (simp add: p2e_inv_extensional[of B, symmetric])
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   987
    from lmeasure_measure_eq_borel_prod[OF this] show "lmeasure (p2e ` A::'a set) =
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   988
      finite_product_sigma_finite.measure (\<lambda>x. borel) (\<lambda>x. lmeasure) {..<DIM('a)} A"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   989
      unfolding e2p_p2e'[OF *] .
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   990
  qed auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   991
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   992
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   993
lemma borel_fubini:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   994
  fixes f :: "'a::ordered_euclidean_space \<Rightarrow> real"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   995
  assumes f: "f \<in> borel_measurable borel"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   996
  shows "borel.integral f = borel_product.product_integral {..<DIM('a)} (f \<circ> p2e)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   997
proof- interpret fprod: finite_product_sigma_finite "\<lambda>x. borel" "\<lambda>x. lmeasure" "{..<DIM('a)}" by default auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   998
  have 1:"(\<lambda>x. Real (f x)) \<in> borel_measurable borel" using f by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
   999
  have 2:"(\<lambda>x. Real (- f x)) \<in> borel_measurable borel" using f by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
  1000
  show ?thesis unfolding fprod.integral_def borel.integral_def
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
  1001
    unfolding borel_fubini_positiv_integral[OF 1] borel_fubini_positiv_integral[OF 2]
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 38656
diff changeset
  1002
    unfolding o_def ..
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1003
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1004
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1005
end