src/HOL/Probability/Measure_Space.thy
author hoelzl
Wed, 25 Apr 2012 19:26:00 +0200
changeset 47761 dfe747e72fa8
parent 47694 05663f75964c
child 47762 d31085f07f60
permissions -rw-r--r--
moved lemmas to appropriate places
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
     1
(*  Title:      HOL/Probability/Measure_Space.thy
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
     2
    Author:     Lawrence C Paulson
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
     3
    Author:     Johannes Hölzl, TU München
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
     4
    Author:     Armin Heller, TU München
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
     5
*)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
     6
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
     7
header {* Measure spaces and their properties *}
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
     8
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
     9
theory Measure_Space
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    10
imports
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    11
  Sigma_Algebra
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    12
  "~~/src/HOL/Multivariate_Analysis/Extended_Real_Limits"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    13
begin
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    14
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    15
lemma suminf_cmult_indicator:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    16
  fixes f :: "nat \<Rightarrow> ereal"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    17
  assumes "disjoint_family A" "x \<in> A i" "\<And>i. 0 \<le> f i"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    18
  shows "(\<Sum>n. f n * indicator (A n) x) = f i"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    19
proof -
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    20
  have **: "\<And>n. f n * indicator (A n) x = (if n = i then f n else 0 :: ereal)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    21
    using `x \<in> A i` assms unfolding disjoint_family_on_def indicator_def by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    22
  then have "\<And>n. (\<Sum>j<n. f j * indicator (A j) x) = (if i < n then f i else 0 :: ereal)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    23
    by (auto simp: setsum_cases)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    24
  moreover have "(SUP n. if i < n then f i else 0) = (f i :: ereal)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    25
  proof (rule ereal_SUPI)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    26
    fix y :: ereal assume "\<And>n. n \<in> UNIV \<Longrightarrow> (if i < n then f i else 0) \<le> y"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    27
    from this[of "Suc i"] show "f i \<le> y" by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    28
  qed (insert assms, simp)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    29
  ultimately show ?thesis using assms
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    30
    by (subst suminf_ereal_eq_SUPR) (auto simp: indicator_def)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    31
qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    32
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    33
lemma suminf_indicator:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    34
  assumes "disjoint_family A"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    35
  shows "(\<Sum>n. indicator (A n) x :: ereal) = indicator (\<Union>i. A i) x"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    36
proof cases
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    37
  assume *: "x \<in> (\<Union>i. A i)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    38
  then obtain i where "x \<in> A i" by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    39
  from suminf_cmult_indicator[OF assms(1), OF `x \<in> A i`, of "\<lambda>k. 1"]
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    40
  show ?thesis using * by simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    41
qed simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    42
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    43
text {*
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    44
  The type for emeasure spaces is already defined in @{theory Sigma_Algebra}, as it is also used to
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    45
  represent sigma algebras (with an arbitrary emeasure).
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    46
*}
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    47
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    48
section "Extend binary sets"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    49
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    50
lemma LIMSEQ_binaryset:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    51
  assumes f: "f {} = 0"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    52
  shows  "(\<lambda>n. \<Sum>i<n. f (binaryset A B i)) ----> f A + f B"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    53
proof -
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    54
  have "(\<lambda>n. \<Sum>i < Suc (Suc n). f (binaryset A B i)) = (\<lambda>n. f A + f B)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    55
    proof
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    56
      fix n
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    57
      show "(\<Sum>i < Suc (Suc n). f (binaryset A B i)) = f A + f B"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    58
        by (induct n)  (auto simp add: binaryset_def f)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    59
    qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    60
  moreover
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    61
  have "... ----> f A + f B" by (rule tendsto_const)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    62
  ultimately
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    63
  have "(\<lambda>n. \<Sum>i< Suc (Suc n). f (binaryset A B i)) ----> f A + f B"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    64
    by metis
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    65
  hence "(\<lambda>n. \<Sum>i< n+2. f (binaryset A B i)) ----> f A + f B"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    66
    by simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    67
  thus ?thesis by (rule LIMSEQ_offset [where k=2])
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    68
qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    69
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    70
lemma binaryset_sums:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    71
  assumes f: "f {} = 0"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    72
  shows  "(\<lambda>n. f (binaryset A B n)) sums (f A + f B)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    73
    by (simp add: sums_def LIMSEQ_binaryset [where f=f, OF f] atLeast0LessThan)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    74
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    75
lemma suminf_binaryset_eq:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    76
  fixes f :: "'a set \<Rightarrow> 'b::{comm_monoid_add, t2_space}"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    77
  shows "f {} = 0 \<Longrightarrow> (\<Sum>n. f (binaryset A B n)) = f A + f B"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    78
  by (metis binaryset_sums sums_unique)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    79
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    80
section {* Properties of a premeasure @{term \<mu>} *}
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    81
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    82
text {*
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    83
  The definitions for @{const positive} and @{const countably_additive} should be here, by they are
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    84
  necessary to define @{typ "'a measure"} in @{theory Sigma_Algebra}.
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    85
*}
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    86
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    87
definition additive where
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    88
  "additive M \<mu> \<longleftrightarrow> (\<forall>x\<in>M. \<forall>y\<in>M. x \<inter> y = {} \<longrightarrow> \<mu> (x \<union> y) = \<mu> x + \<mu> y)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    89
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    90
definition increasing where
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    91
  "increasing M \<mu> \<longleftrightarrow> (\<forall>x\<in>M. \<forall>y\<in>M. x \<subseteq> y \<longrightarrow> \<mu> x \<le> \<mu> y)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    92
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    93
lemma positiveD_empty:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    94
  "positive M f \<Longrightarrow> f {} = 0"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    95
  by (auto simp add: positive_def)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    96
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    97
lemma additiveD:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    98
  "additive M f \<Longrightarrow> x \<inter> y = {} \<Longrightarrow> x \<in> M \<Longrightarrow> y \<in> M \<Longrightarrow> f (x \<union> y) = f x + f y"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    99
  by (auto simp add: additive_def)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   100
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   101
lemma increasingD:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   102
  "increasing M f \<Longrightarrow> x \<subseteq> y \<Longrightarrow> x\<in>M \<Longrightarrow> y\<in>M \<Longrightarrow> f x \<le> f y"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   103
  by (auto simp add: increasing_def)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   104
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   105
lemma countably_additiveI:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   106
  "(\<And>A. range A \<subseteq> M \<Longrightarrow> disjoint_family A \<Longrightarrow> (\<Union>i. A i) \<in> M \<Longrightarrow> (\<Sum>i. f (A i)) = f (\<Union>i. A i))
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   107
  \<Longrightarrow> countably_additive M f"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   108
  by (simp add: countably_additive_def)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   109
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   110
lemma (in ring_of_sets) disjointed_additive:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   111
  assumes f: "positive M f" "additive M f" and A: "range A \<subseteq> M" "incseq A"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   112
  shows "(\<Sum>i\<le>n. f (disjointed A i)) = f (A n)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   113
proof (induct n)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   114
  case (Suc n)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   115
  then have "(\<Sum>i\<le>Suc n. f (disjointed A i)) = f (A n) + f (disjointed A (Suc n))"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   116
    by simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   117
  also have "\<dots> = f (A n \<union> disjointed A (Suc n))"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   118
    using A by (subst f(2)[THEN additiveD]) (auto simp: disjointed_incseq)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   119
  also have "A n \<union> disjointed A (Suc n) = A (Suc n)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   120
    using `incseq A` by (auto dest: incseq_SucD simp: disjointed_incseq)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   121
  finally show ?case .
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   122
qed simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   123
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   124
lemma (in ring_of_sets) additive_sum:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   125
  fixes A:: "'i \<Rightarrow> 'a set"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   126
  assumes f: "positive M f" and ad: "additive M f" and "finite S"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   127
      and A: "A`S \<subseteq> M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   128
      and disj: "disjoint_family_on A S"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   129
  shows  "(\<Sum>i\<in>S. f (A i)) = f (\<Union>i\<in>S. A i)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   130
using `finite S` disj A proof induct
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   131
  case empty show ?case using f by (simp add: positive_def)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   132
next
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   133
  case (insert s S)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   134
  then have "A s \<inter> (\<Union>i\<in>S. A i) = {}"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   135
    by (auto simp add: disjoint_family_on_def neq_iff)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   136
  moreover
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   137
  have "A s \<in> M" using insert by blast
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   138
  moreover have "(\<Union>i\<in>S. A i) \<in> M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   139
    using insert `finite S` by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   140
  moreover
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   141
  ultimately have "f (A s \<union> (\<Union>i\<in>S. A i)) = f (A s) + f(\<Union>i\<in>S. A i)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   142
    using ad UNION_in_sets A by (auto simp add: additive_def)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   143
  with insert show ?case using ad disjoint_family_on_mono[of S "insert s S" A]
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   144
    by (auto simp add: additive_def subset_insertI)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   145
qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   146
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   147
lemma (in ring_of_sets) additive_increasing:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   148
  assumes posf: "positive M f" and addf: "additive M f"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   149
  shows "increasing M f"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   150
proof (auto simp add: increasing_def)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   151
  fix x y
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   152
  assume xy: "x \<in> M" "y \<in> M" "x \<subseteq> y"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   153
  then have "y - x \<in> M" by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   154
  then have "0 \<le> f (y-x)" using posf[unfolded positive_def] by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   155
  then have "f x + 0 \<le> f x + f (y-x)" by (intro add_left_mono) auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   156
  also have "... = f (x \<union> (y-x))" using addf
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   157
    by (auto simp add: additive_def) (metis Diff_disjoint Un_Diff_cancel Diff xy(1,2))
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   158
  also have "... = f y"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   159
    by (metis Un_Diff_cancel Un_absorb1 xy(3))
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   160
  finally show "f x \<le> f y" by simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   161
qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   162
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   163
lemma (in ring_of_sets) countably_additive_additive:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   164
  assumes posf: "positive M f" and ca: "countably_additive M f"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   165
  shows "additive M f"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   166
proof (auto simp add: additive_def)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   167
  fix x y
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   168
  assume x: "x \<in> M" and y: "y \<in> M" and "x \<inter> y = {}"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   169
  hence "disjoint_family (binaryset x y)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   170
    by (auto simp add: disjoint_family_on_def binaryset_def)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   171
  hence "range (binaryset x y) \<subseteq> M \<longrightarrow>
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   172
         (\<Union>i. binaryset x y i) \<in> M \<longrightarrow>
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   173
         f (\<Union>i. binaryset x y i) = (\<Sum> n. f (binaryset x y n))"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   174
    using ca
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   175
    by (simp add: countably_additive_def)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   176
  hence "{x,y,{}} \<subseteq> M \<longrightarrow> x \<union> y \<in> M \<longrightarrow>
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   177
         f (x \<union> y) = (\<Sum>n. f (binaryset x y n))"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   178
    by (simp add: range_binaryset_eq UN_binaryset_eq)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   179
  thus "f (x \<union> y) = f x + f y" using posf x y
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   180
    by (auto simp add: Un suminf_binaryset_eq positive_def)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   181
qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   182
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   183
lemma (in algebra) increasing_additive_bound:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   184
  fixes A:: "nat \<Rightarrow> 'a set" and  f :: "'a set \<Rightarrow> ereal"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   185
  assumes f: "positive M f" and ad: "additive M f"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   186
      and inc: "increasing M f"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   187
      and A: "range A \<subseteq> M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   188
      and disj: "disjoint_family A"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   189
  shows  "(\<Sum>i. f (A i)) \<le> f \<Omega>"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   190
proof (safe intro!: suminf_bound)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   191
  fix N
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   192
  note disj_N = disjoint_family_on_mono[OF _ disj, of "{..<N}"]
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   193
  have "(\<Sum>i<N. f (A i)) = f (\<Union>i\<in>{..<N}. A i)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   194
    using A by (intro additive_sum [OF f ad _ _]) (auto simp: disj_N)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   195
  also have "... \<le> f \<Omega>" using space_closed A
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   196
    by (intro increasingD[OF inc] finite_UN) auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   197
  finally show "(\<Sum>i<N. f (A i)) \<le> f \<Omega>" by simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   198
qed (insert f A, auto simp: positive_def)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   199
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   200
lemma (in ring_of_sets) countably_additiveI_finite:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   201
  assumes "finite \<Omega>" "positive M \<mu>" "additive M \<mu>"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   202
  shows "countably_additive M \<mu>"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   203
proof (rule countably_additiveI)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   204
  fix F :: "nat \<Rightarrow> 'a set" assume F: "range F \<subseteq> M" "(\<Union>i. F i) \<in> M" and disj: "disjoint_family F"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   205
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   206
  have "\<forall>i\<in>{i. F i \<noteq> {}}. \<exists>x. x \<in> F i" by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   207
  from bchoice[OF this] obtain f where f: "\<And>i. F i \<noteq> {} \<Longrightarrow> f i \<in> F i" by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   208
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   209
  have inj_f: "inj_on f {i. F i \<noteq> {}}"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   210
  proof (rule inj_onI, simp)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   211
    fix i j a b assume *: "f i = f j" "F i \<noteq> {}" "F j \<noteq> {}"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   212
    then have "f i \<in> F i" "f j \<in> F j" using f by force+
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   213
    with disj * show "i = j" by (auto simp: disjoint_family_on_def)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   214
  qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   215
  have "finite (\<Union>i. F i)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   216
    by (metis F(2) assms(1) infinite_super sets_into_space)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   217
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   218
  have F_subset: "{i. \<mu> (F i) \<noteq> 0} \<subseteq> {i. F i \<noteq> {}}"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   219
    by (auto simp: positiveD_empty[OF `positive M \<mu>`])
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   220
  moreover have fin_not_empty: "finite {i. F i \<noteq> {}}"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   221
  proof (rule finite_imageD)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   222
    from f have "f`{i. F i \<noteq> {}} \<subseteq> (\<Union>i. F i)" by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   223
    then show "finite (f`{i. F i \<noteq> {}})"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   224
      by (rule finite_subset) fact
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   225
  qed fact
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   226
  ultimately have fin_not_0: "finite {i. \<mu> (F i) \<noteq> 0}"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   227
    by (rule finite_subset)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   228
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   229
  have disj_not_empty: "disjoint_family_on F {i. F i \<noteq> {}}"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   230
    using disj by (auto simp: disjoint_family_on_def)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   231
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   232
  from fin_not_0 have "(\<Sum>i. \<mu> (F i)) = (\<Sum>i | \<mu> (F i) \<noteq> 0. \<mu> (F i))"
47761
dfe747e72fa8 moved lemmas to appropriate places
hoelzl
parents: 47694
diff changeset
   233
    by (rule suminf_finite) auto
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   234
  also have "\<dots> = (\<Sum>i | F i \<noteq> {}. \<mu> (F i))"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   235
    using fin_not_empty F_subset by (rule setsum_mono_zero_left) auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   236
  also have "\<dots> = \<mu> (\<Union>i\<in>{i. F i \<noteq> {}}. F i)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   237
    using `positive M \<mu>` `additive M \<mu>` fin_not_empty disj_not_empty F by (intro additive_sum) auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   238
  also have "\<dots> = \<mu> (\<Union>i. F i)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   239
    by (rule arg_cong[where f=\<mu>]) auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   240
  finally show "(\<Sum>i. \<mu> (F i)) = \<mu> (\<Union>i. F i)" .
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   241
qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   242
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   243
section {* Properties of @{const emeasure} *}
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   244
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   245
lemma emeasure_positive: "positive (sets M) (emeasure M)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   246
  by (cases M) (auto simp: sets_def emeasure_def Abs_measure_inverse measure_space_def)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   247
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   248
lemma emeasure_empty[simp, intro]: "emeasure M {} = 0"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   249
  using emeasure_positive[of M] by (simp add: positive_def)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   250
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   251
lemma emeasure_nonneg[intro!]: "0 \<le> emeasure M A"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   252
  using emeasure_notin_sets[of A M] emeasure_positive[of M]
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   253
  by (cases "A \<in> sets M") (auto simp: positive_def)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   254
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   255
lemma emeasure_not_MInf[simp]: "emeasure M A \<noteq> - \<infinity>"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   256
  using emeasure_nonneg[of M A] by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   257
  
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   258
lemma emeasure_countably_additive: "countably_additive (sets M) (emeasure M)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   259
  by (cases M) (auto simp: sets_def emeasure_def Abs_measure_inverse measure_space_def)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   260
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   261
lemma suminf_emeasure:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   262
  "range A \<subseteq> sets M \<Longrightarrow> disjoint_family A \<Longrightarrow> (\<Sum>i. emeasure M (A i)) = emeasure M (\<Union>i. A i)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   263
  using countable_UN[of A UNIV M] emeasure_countably_additive[of M]
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   264
  by (simp add: countably_additive_def)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   265
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   266
lemma emeasure_additive: "additive (sets M) (emeasure M)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   267
  by (metis countably_additive_additive emeasure_positive emeasure_countably_additive)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   268
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   269
lemma plus_emeasure:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   270
  "a \<in> sets M \<Longrightarrow> b \<in> sets M \<Longrightarrow> a \<inter> b = {} \<Longrightarrow> emeasure M a + emeasure M b = emeasure M (a \<union> b)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   271
  using additiveD[OF emeasure_additive] ..
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   272
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   273
lemma setsum_emeasure:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   274
  "F`I \<subseteq> sets M \<Longrightarrow> disjoint_family_on F I \<Longrightarrow> finite I \<Longrightarrow>
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   275
    (\<Sum>i\<in>I. emeasure M (F i)) = emeasure M (\<Union>i\<in>I. F i)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   276
  by (metis additive_sum emeasure_positive emeasure_additive)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   277
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   278
lemma emeasure_mono:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   279
  "a \<subseteq> b \<Longrightarrow> b \<in> sets M \<Longrightarrow> emeasure M a \<le> emeasure M b"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   280
  by (metis additive_increasing emeasure_additive emeasure_nonneg emeasure_notin_sets
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   281
            emeasure_positive increasingD)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   282
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   283
lemma emeasure_space:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   284
  "emeasure M A \<le> emeasure M (space M)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   285
  by (metis emeasure_mono emeasure_nonneg emeasure_notin_sets sets_into_space top)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   286
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   287
lemma emeasure_compl:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   288
  assumes s: "s \<in> sets M" and fin: "emeasure M s \<noteq> \<infinity>"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   289
  shows "emeasure M (space M - s) = emeasure M (space M) - emeasure M s"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   290
proof -
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   291
  from s have "0 \<le> emeasure M s" by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   292
  have "emeasure M (space M) = emeasure M (s \<union> (space M - s))" using s
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   293
    by (metis Un_Diff_cancel Un_absorb1 s sets_into_space)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   294
  also have "... = emeasure M s + emeasure M (space M - s)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   295
    by (rule plus_emeasure[symmetric]) (auto simp add: s)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   296
  finally have "emeasure M (space M) = emeasure M s + emeasure M (space M - s)" .
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   297
  then show ?thesis
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   298
    using fin `0 \<le> emeasure M s`
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   299
    unfolding ereal_eq_minus_iff by (auto simp: ac_simps)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   300
qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   301
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   302
lemma emeasure_Diff:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   303
  assumes finite: "emeasure M B \<noteq> \<infinity>"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   304
  and measurable: "A \<in> sets M" "B \<in> sets M" "B \<subseteq> A"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   305
  shows "emeasure M (A - B) = emeasure M A - emeasure M B"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   306
proof -
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   307
  have "0 \<le> emeasure M B" using assms by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   308
  have "(A - B) \<union> B = A" using `B \<subseteq> A` by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   309
  then have "emeasure M A = emeasure M ((A - B) \<union> B)" by simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   310
  also have "\<dots> = emeasure M (A - B) + emeasure M B"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   311
    using measurable by (subst plus_emeasure[symmetric]) auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   312
  finally show "emeasure M (A - B) = emeasure M A - emeasure M B"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   313
    unfolding ereal_eq_minus_iff
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   314
    using finite `0 \<le> emeasure M B` by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   315
qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   316
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   317
lemma emeasure_countable_increasing:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   318
  assumes A: "range A \<subseteq> sets M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   319
      and A0: "A 0 = {}"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   320
      and ASuc: "\<And>n. A n \<subseteq> A (Suc n)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   321
  shows "(SUP n. emeasure M (A n)) = emeasure M (\<Union>i. A i)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   322
proof -
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   323
  { fix n
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   324
    have "emeasure M (A n) = (\<Sum>i<n. emeasure M (A (Suc i) - A i))"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   325
      proof (induct n)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   326
        case 0 thus ?case by (auto simp add: A0)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   327
      next
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   328
        case (Suc m)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   329
        have "A (Suc m) = A m \<union> (A (Suc m) - A m)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   330
          by (metis ASuc Un_Diff_cancel Un_absorb1)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   331
        hence "emeasure M (A (Suc m)) =
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   332
               emeasure M (A m) + emeasure M (A (Suc m) - A m)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   333
          by (subst plus_emeasure)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   334
             (auto simp add: emeasure_additive range_subsetD [OF A])
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   335
        with Suc show ?case
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   336
          by simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   337
      qed }
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   338
  note Meq = this
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   339
  have Aeq: "(\<Union>i. A (Suc i) - A i) = (\<Union>i. A i)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   340
    proof (rule UN_finite2_eq [where k=1], simp)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   341
      fix i
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   342
      show "(\<Union>i\<in>{0..<i}. A (Suc i) - A i) = (\<Union>i\<in>{0..<Suc i}. A i)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   343
        proof (induct i)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   344
          case 0 thus ?case by (simp add: A0)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   345
        next
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   346
          case (Suc i)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   347
          thus ?case
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   348
            by (auto simp add: atLeastLessThanSuc intro: subsetD [OF ASuc])
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   349
        qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   350
    qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   351
  have A1: "\<And>i. A (Suc i) - A i \<in> sets M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   352
    by (metis A Diff range_subsetD)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   353
  have A2: "(\<Union>i. A (Suc i) - A i) \<in> sets M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   354
    by (blast intro: range_subsetD [OF A])
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   355
  have "(SUP n. \<Sum>i<n. emeasure M (A (Suc i) - A i)) = (\<Sum>i. emeasure M (A (Suc i) - A i))"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   356
    using A by (auto intro!: suminf_ereal_eq_SUPR[symmetric])
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   357
  also have "\<dots> = emeasure M (\<Union>i. A (Suc i) - A i)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   358
    by (rule suminf_emeasure)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   359
       (auto simp add: disjoint_family_Suc ASuc A1 A2)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   360
  also have "... =  emeasure M (\<Union>i. A i)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   361
    by (simp add: Aeq)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   362
  finally have "(SUP n. \<Sum>i<n. emeasure M (A (Suc i) - A i)) = emeasure M (\<Union>i. A i)" .
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   363
  then show ?thesis by (auto simp add: Meq)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   364
qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   365
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   366
lemma SUP_emeasure_incseq:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   367
  assumes A: "range A \<subseteq> sets M" and "incseq A"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   368
  shows "(SUP n. emeasure M (A n)) = emeasure M (\<Union>i. A i)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   369
proof -
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   370
  have *: "(SUP n. emeasure M (nat_case {} A (Suc n))) = (SUP n. emeasure M (nat_case {} A n))"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   371
    using A by (auto intro!: SUPR_eq exI split: nat.split)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   372
  have ueq: "(\<Union>i. nat_case {} A i) = (\<Union>i. A i)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   373
    by (auto simp add: split: nat.splits)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   374
  have meq: "\<And>n. emeasure M (A n) = (emeasure M \<circ> nat_case {} A) (Suc n)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   375
    by simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   376
  have "(SUP n. emeasure M (nat_case {} A n)) = emeasure M (\<Union>i. nat_case {} A i)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   377
    using range_subsetD[OF A] incseq_SucD[OF `incseq A`]
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   378
    by (force split: nat.splits intro!: emeasure_countable_increasing)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   379
  also have "emeasure M (\<Union>i. nat_case {} A i) = emeasure M (\<Union>i. A i)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   380
    by (simp add: ueq)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   381
  finally have "(SUP n. emeasure M (nat_case {} A n)) = emeasure M (\<Union>i. A i)" .
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   382
  thus ?thesis unfolding meq * comp_def .
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   383
qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   384
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   385
lemma incseq_emeasure:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   386
  assumes "range B \<subseteq> sets M" "incseq B"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   387
  shows "incseq (\<lambda>i. emeasure M (B i))"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   388
  using assms by (auto simp: incseq_def intro!: emeasure_mono)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   389
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   390
lemma Lim_emeasure_incseq:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   391
  assumes A: "range A \<subseteq> sets M" "incseq A"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   392
  shows "(\<lambda>i. (emeasure M (A i))) ----> emeasure M (\<Union>i. A i)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   393
  using LIMSEQ_ereal_SUPR[OF incseq_emeasure, OF A]
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   394
    SUP_emeasure_incseq[OF A] by simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   395
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   396
lemma decseq_emeasure:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   397
  assumes "range B \<subseteq> sets M" "decseq B"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   398
  shows "decseq (\<lambda>i. emeasure M (B i))"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   399
  using assms by (auto simp: decseq_def intro!: emeasure_mono)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   400
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   401
lemma INF_emeasure_decseq:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   402
  assumes A: "range A \<subseteq> sets M" and "decseq A"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   403
  and finite: "\<And>i. emeasure M (A i) \<noteq> \<infinity>"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   404
  shows "(INF n. emeasure M (A n)) = emeasure M (\<Inter>i. A i)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   405
proof -
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   406
  have le_MI: "emeasure M (\<Inter>i. A i) \<le> emeasure M (A 0)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   407
    using A by (auto intro!: emeasure_mono)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   408
  hence *: "emeasure M (\<Inter>i. A i) \<noteq> \<infinity>" using finite[of 0] by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   409
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   410
  have A0: "0 \<le> emeasure M (A 0)" using A by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   411
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   412
  have "emeasure M (A 0) - (INF n. emeasure M (A n)) = emeasure M (A 0) + (SUP n. - emeasure M (A n))"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   413
    by (simp add: ereal_SUPR_uminus minus_ereal_def)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   414
  also have "\<dots> = (SUP n. emeasure M (A 0) - emeasure M (A n))"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   415
    unfolding minus_ereal_def using A0 assms
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   416
    by (subst SUPR_ereal_add) (auto simp add: decseq_emeasure)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   417
  also have "\<dots> = (SUP n. emeasure M (A 0 - A n))"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   418
    using A finite `decseq A`[unfolded decseq_def] by (subst emeasure_Diff) auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   419
  also have "\<dots> = emeasure M (\<Union>i. A 0 - A i)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   420
  proof (rule SUP_emeasure_incseq)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   421
    show "range (\<lambda>n. A 0 - A n) \<subseteq> sets M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   422
      using A by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   423
    show "incseq (\<lambda>n. A 0 - A n)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   424
      using `decseq A` by (auto simp add: incseq_def decseq_def)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   425
  qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   426
  also have "\<dots> = emeasure M (A 0) - emeasure M (\<Inter>i. A i)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   427
    using A finite * by (simp, subst emeasure_Diff) auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   428
  finally show ?thesis
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   429
    unfolding ereal_minus_eq_minus_iff using finite A0 by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   430
qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   431
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   432
lemma Lim_emeasure_decseq:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   433
  assumes A: "range A \<subseteq> sets M" "decseq A" and fin: "\<And>i. emeasure M (A i) \<noteq> \<infinity>"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   434
  shows "(\<lambda>i. emeasure M (A i)) ----> emeasure M (\<Inter>i. A i)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   435
  using LIMSEQ_ereal_INFI[OF decseq_emeasure, OF A]
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   436
  using INF_emeasure_decseq[OF A fin] by simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   437
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   438
lemma emeasure_subadditive:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   439
  assumes measurable: "A \<in> sets M" "B \<in> sets M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   440
  shows "emeasure M (A \<union> B) \<le> emeasure M A + emeasure M B"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   441
proof -
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   442
  from plus_emeasure[of A M "B - A"]
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   443
  have "emeasure M (A \<union> B) = emeasure M A + emeasure M (B - A)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   444
    using assms by (simp add: Diff)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   445
  also have "\<dots> \<le> emeasure M A + emeasure M B"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   446
    using assms by (auto intro!: add_left_mono emeasure_mono)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   447
  finally show ?thesis .
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   448
qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   449
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   450
lemma emeasure_subadditive_finite:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   451
  assumes "finite I" "A ` I \<subseteq> sets M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   452
  shows "emeasure M (\<Union>i\<in>I. A i) \<le> (\<Sum>i\<in>I. emeasure M (A i))"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   453
using assms proof induct
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   454
  case (insert i I)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   455
  then have "emeasure M (\<Union>i\<in>insert i I. A i) = emeasure M (A i \<union> (\<Union>i\<in>I. A i))"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   456
    by simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   457
  also have "\<dots> \<le> emeasure M (A i) + emeasure M (\<Union>i\<in>I. A i)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   458
    using insert by (intro emeasure_subadditive finite_UN) auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   459
  also have "\<dots> \<le> emeasure M (A i) + (\<Sum>i\<in>I. emeasure M (A i))"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   460
    using insert by (intro add_mono) auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   461
  also have "\<dots> = (\<Sum>i\<in>insert i I. emeasure M (A i))"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   462
    using insert by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   463
  finally show ?case .
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   464
qed simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   465
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   466
lemma emeasure_subadditive_countably:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   467
  assumes "range f \<subseteq> sets M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   468
  shows "emeasure M (\<Union>i. f i) \<le> (\<Sum>i. emeasure M (f i))"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   469
proof -
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   470
  have "emeasure M (\<Union>i. f i) = emeasure M (\<Union>i. disjointed f i)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   471
    unfolding UN_disjointed_eq ..
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   472
  also have "\<dots> = (\<Sum>i. emeasure M (disjointed f i))"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   473
    using range_disjointed_sets[OF assms] suminf_emeasure[of "disjointed f"]
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   474
    by (simp add:  disjoint_family_disjointed comp_def)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   475
  also have "\<dots> \<le> (\<Sum>i. emeasure M (f i))"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   476
    using range_disjointed_sets[OF assms] assms
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   477
    by (auto intro!: suminf_le_pos emeasure_mono disjointed_subset)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   478
  finally show ?thesis .
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   479
qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   480
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   481
lemma emeasure_insert:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   482
  assumes sets: "{x} \<in> sets M" "A \<in> sets M" and "x \<notin> A"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   483
  shows "emeasure M (insert x A) = emeasure M {x} + emeasure M A"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   484
proof -
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   485
  have "{x} \<inter> A = {}" using `x \<notin> A` by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   486
  from plus_emeasure[OF sets this] show ?thesis by simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   487
qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   488
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   489
lemma emeasure_eq_setsum_singleton:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   490
  assumes "finite S" "\<And>x. x \<in> S \<Longrightarrow> {x} \<in> sets M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   491
  shows "emeasure M S = (\<Sum>x\<in>S. emeasure M {x})"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   492
  using setsum_emeasure[of "\<lambda>x. {x}" S M] assms
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   493
  by (auto simp: disjoint_family_on_def subset_eq)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   494
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   495
lemma setsum_emeasure_cover:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   496
  assumes "finite S" and "A \<in> sets M" and br_in_M: "B ` S \<subseteq> sets M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   497
  assumes A: "A \<subseteq> (\<Union>i\<in>S. B i)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   498
  assumes disj: "disjoint_family_on B S"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   499
  shows "emeasure M A = (\<Sum>i\<in>S. emeasure M (A \<inter> (B i)))"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   500
proof -
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   501
  have "(\<Sum>i\<in>S. emeasure M (A \<inter> (B i))) = emeasure M (\<Union>i\<in>S. A \<inter> (B i))"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   502
  proof (rule setsum_emeasure)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   503
    show "disjoint_family_on (\<lambda>i. A \<inter> B i) S"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   504
      using `disjoint_family_on B S`
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   505
      unfolding disjoint_family_on_def by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   506
  qed (insert assms, auto)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   507
  also have "(\<Union>i\<in>S. A \<inter> (B i)) = A"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   508
    using A by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   509
  finally show ?thesis by simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   510
qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   511
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   512
lemma emeasure_eq_0:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   513
  "N \<in> sets M \<Longrightarrow> emeasure M N = 0 \<Longrightarrow> K \<subseteq> N \<Longrightarrow> emeasure M K = 0"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   514
  by (metis emeasure_mono emeasure_nonneg order_eq_iff)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   515
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   516
lemma emeasure_UN_eq_0:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   517
  assumes "\<And>i::nat. emeasure M (N i) = 0" and "range N \<subseteq> sets M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   518
  shows "emeasure M (\<Union> i. N i) = 0"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   519
proof -
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   520
  have "0 \<le> emeasure M (\<Union> i. N i)" using assms by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   521
  moreover have "emeasure M (\<Union> i. N i) \<le> 0"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   522
    using emeasure_subadditive_countably[OF assms(2)] assms(1) by simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   523
  ultimately show ?thesis by simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   524
qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   525
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   526
lemma measure_eqI_finite:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   527
  assumes [simp]: "sets M = Pow A" "sets N = Pow A" and "finite A"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   528
  assumes eq: "\<And>a. a \<in> A \<Longrightarrow> emeasure M {a} = emeasure N {a}"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   529
  shows "M = N"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   530
proof (rule measure_eqI)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   531
  fix X assume "X \<in> sets M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   532
  then have X: "X \<subseteq> A" by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   533
  then have "emeasure M X = (\<Sum>a\<in>X. emeasure M {a})"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   534
    using `finite A` by (subst emeasure_eq_setsum_singleton) (auto dest: finite_subset)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   535
  also have "\<dots> = (\<Sum>a\<in>X. emeasure N {a})"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   536
    using X eq by (auto intro!: setsum_cong)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   537
  also have "\<dots> = emeasure N X"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   538
    using X `finite A` by (subst emeasure_eq_setsum_singleton) (auto dest: finite_subset)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   539
  finally show "emeasure M X = emeasure N X" .
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   540
qed simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   541
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   542
lemma measure_eqI_generator_eq:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   543
  fixes M N :: "'a measure" and E :: "'a set set" and A :: "nat \<Rightarrow> 'a set"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   544
  assumes "Int_stable E" "E \<subseteq> Pow \<Omega>"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   545
  and eq: "\<And>X. X \<in> E \<Longrightarrow> emeasure M X = emeasure N X"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   546
  and M: "sets M = sigma_sets \<Omega> E"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   547
  and N: "sets N = sigma_sets \<Omega> E"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   548
  and A: "range A \<subseteq> E" "incseq A" "(\<Union>i. A i) = \<Omega>" "\<And>i. emeasure M (A i) \<noteq> \<infinity>"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   549
  shows "M = N"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   550
proof -
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   551
  let ?D = "\<lambda>F. {D. D \<in> sigma_sets \<Omega> E \<and> emeasure M (F \<inter> D) = emeasure N (F \<inter> D)}"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   552
  interpret S: sigma_algebra \<Omega> "sigma_sets \<Omega> E" by (rule sigma_algebra_sigma_sets) fact
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   553
  { fix F assume "F \<in> E" and "emeasure M F \<noteq> \<infinity>"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   554
    then have [intro]: "F \<in> sigma_sets \<Omega> E" by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   555
    have "emeasure N F \<noteq> \<infinity>" using `emeasure M F \<noteq> \<infinity>` `F \<in> E` eq by simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   556
    interpret D: dynkin_system \<Omega> "?D F"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   557
    proof (rule dynkin_systemI, simp_all)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   558
      fix A assume "A \<in> sigma_sets \<Omega> E \<and> emeasure M (F \<inter> A) = emeasure N (F \<inter> A)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   559
      then show "A \<subseteq> \<Omega>" using S.sets_into_space by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   560
    next
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   561
      have "F \<inter> \<Omega> = F" using `F \<in> E` `E \<subseteq> Pow \<Omega>` by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   562
      then show "emeasure M (F \<inter> \<Omega>) = emeasure N (F \<inter> \<Omega>)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   563
        using `F \<in> E` eq by (auto intro: sigma_sets_top)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   564
    next
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   565
      fix A assume *: "A \<in> sigma_sets \<Omega> E \<and> emeasure M (F \<inter> A) = emeasure N (F \<inter> A)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   566
      then have **: "F \<inter> (\<Omega> - A) = F - (F \<inter> A)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   567
        and [intro]: "F \<inter> A \<in> sigma_sets \<Omega> E"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   568
        using `F \<in> E` S.sets_into_space by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   569
      have "emeasure N (F \<inter> A) \<le> emeasure N F" by (auto intro!: emeasure_mono simp: M N)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   570
      then have "emeasure N (F \<inter> A) \<noteq> \<infinity>" using `emeasure N F \<noteq> \<infinity>` by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   571
      have "emeasure M (F \<inter> A) \<le> emeasure M F" by (auto intro!: emeasure_mono simp: M N)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   572
      then have "emeasure M (F \<inter> A) \<noteq> \<infinity>" using `emeasure M F \<noteq> \<infinity>` by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   573
      then have "emeasure M (F \<inter> (\<Omega> - A)) = emeasure M F - emeasure M (F \<inter> A)" unfolding **
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   574
        using `F \<inter> A \<in> sigma_sets \<Omega> E` by (auto intro!: emeasure_Diff simp: M N)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   575
      also have "\<dots> = emeasure N F - emeasure N (F \<inter> A)" using eq `F \<in> E` * by simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   576
      also have "\<dots> = emeasure N (F \<inter> (\<Omega> - A))" unfolding **
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   577
        using `F \<inter> A \<in> sigma_sets \<Omega> E` `emeasure N (F \<inter> A) \<noteq> \<infinity>`
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   578
        by (auto intro!: emeasure_Diff[symmetric] simp: M N)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   579
      finally show "\<Omega> - A \<in> sigma_sets \<Omega> E \<and> emeasure M (F \<inter> (\<Omega> - A)) = emeasure N (F \<inter> (\<Omega> - A))"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   580
        using * by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   581
    next
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   582
      fix A :: "nat \<Rightarrow> 'a set"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   583
      assume "disjoint_family A" "range A \<subseteq> {X \<in> sigma_sets \<Omega> E. emeasure M (F \<inter> X) = emeasure N (F \<inter> X)}"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   584
      then have A: "range (\<lambda>i. F \<inter> A i) \<subseteq> sigma_sets \<Omega> E" "F \<inter> (\<Union>x. A x) = (\<Union>x. F \<inter> A x)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   585
        "disjoint_family (\<lambda>i. F \<inter> A i)" "\<And>i. emeasure M (F \<inter> A i) = emeasure N (F \<inter> A i)" "range A \<subseteq> sigma_sets \<Omega> E"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   586
        by (auto simp: disjoint_family_on_def subset_eq)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   587
      then show "(\<Union>x. A x) \<in> sigma_sets \<Omega> E \<and> emeasure M (F \<inter> (\<Union>x. A x)) = emeasure N (F \<inter> (\<Union>x. A x))"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   588
        by (auto simp: M N suminf_emeasure[symmetric] simp del: UN_simps)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   589
    qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   590
    have *: "sigma_sets \<Omega> E = ?D F"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   591
      using `F \<in> E` `Int_stable E`
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   592
      by (intro D.dynkin_lemma) (auto simp add: Int_stable_def eq)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   593
    have "\<And>D. D \<in> sigma_sets \<Omega> E \<Longrightarrow> emeasure M (F \<inter> D) = emeasure N (F \<inter> D)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   594
      by (subst (asm) *) auto }
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   595
  note * = this
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   596
  show "M = N"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   597
  proof (rule measure_eqI)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   598
    show "sets M = sets N"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   599
      using M N by simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   600
    fix X assume "X \<in> sets M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   601
    then have "X \<in> sigma_sets \<Omega> E"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   602
      using M by simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   603
    let ?A = "\<lambda>i. A i \<inter> X"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   604
    have "range ?A \<subseteq> sigma_sets \<Omega> E" "incseq ?A"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   605
      using A(1,2) `X \<in> sigma_sets \<Omega> E` by (auto simp: incseq_def)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   606
    moreover
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   607
    { fix i have "emeasure M (?A i) = emeasure N (?A i)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   608
        using *[of "A i" X] `X \<in> sigma_sets \<Omega> E` A finite by auto }
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   609
    ultimately show "emeasure M X = emeasure N X"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   610
      using SUP_emeasure_incseq[of ?A M] SUP_emeasure_incseq[of ?A N] A(3) `X \<in> sigma_sets \<Omega> E`
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   611
      by (auto simp: M N SUP_emeasure_incseq)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   612
  qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   613
qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   614
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   615
lemma measure_of_of_measure: "measure_of (space M) (sets M) (emeasure M) = M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   616
proof (intro measure_eqI emeasure_measure_of_sigma)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   617
  show "sigma_algebra (space M) (sets M)" ..
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   618
  show "positive (sets M) (emeasure M)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   619
    by (simp add: positive_def emeasure_nonneg)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   620
  show "countably_additive (sets M) (emeasure M)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   621
    by (simp add: emeasure_countably_additive)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   622
qed simp_all
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   623
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   624
section "@{text \<mu>}-null sets"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   625
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   626
definition null_sets :: "'a measure \<Rightarrow> 'a set set" where
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   627
  "null_sets M = {N\<in>sets M. emeasure M N = 0}"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   628
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   629
lemma null_setsD1[dest]: "A \<in> null_sets M \<Longrightarrow> emeasure M A = 0"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   630
  by (simp add: null_sets_def)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   631
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   632
lemma null_setsD2[dest]: "A \<in> null_sets M \<Longrightarrow> A \<in> sets M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   633
  unfolding null_sets_def by simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   634
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   635
lemma null_setsI[intro]: "emeasure M A = 0 \<Longrightarrow> A \<in> sets M \<Longrightarrow> A \<in> null_sets M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   636
  unfolding null_sets_def by simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   637
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   638
interpretation null_sets: ring_of_sets "space M" "null_sets M" for M
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   639
proof
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   640
  show "null_sets M \<subseteq> Pow (space M)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   641
    using sets_into_space by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   642
  show "{} \<in> null_sets M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   643
    by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   644
  fix A B assume sets: "A \<in> null_sets M" "B \<in> null_sets M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   645
  then have "A \<in> sets M" "B \<in> sets M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   646
    by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   647
  moreover then have "emeasure M (A \<union> B) \<le> emeasure M A + emeasure M B"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   648
    "emeasure M (A - B) \<le> emeasure M A"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   649
    by (auto intro!: emeasure_subadditive emeasure_mono)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   650
  moreover have "emeasure M B = 0" "emeasure M A = 0"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   651
    using sets by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   652
  ultimately show "A - B \<in> null_sets M" "A \<union> B \<in> null_sets M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   653
    by (auto intro!: antisym)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   654
qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   655
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   656
lemma UN_from_nat: "(\<Union>i. N i) = (\<Union>i. N (Countable.from_nat i))"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   657
proof -
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   658
  have "(\<Union>i. N i) = (\<Union>i. (N \<circ> Countable.from_nat) i)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   659
    unfolding SUP_def image_compose
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   660
    unfolding surj_from_nat ..
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   661
  then show ?thesis by simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   662
qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   663
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   664
lemma null_sets_UN[intro]:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   665
  assumes "\<And>i::'i::countable. N i \<in> null_sets M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   666
  shows "(\<Union>i. N i) \<in> null_sets M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   667
proof (intro conjI CollectI null_setsI)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   668
  show "(\<Union>i. N i) \<in> sets M" using assms by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   669
  have "0 \<le> emeasure M (\<Union>i. N i)" by (rule emeasure_nonneg)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   670
  moreover have "emeasure M (\<Union>i. N i) \<le> (\<Sum>n. emeasure M (N (Countable.from_nat n)))"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   671
    unfolding UN_from_nat[of N]
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   672
    using assms by (intro emeasure_subadditive_countably) auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   673
  ultimately show "emeasure M (\<Union>i. N i) = 0"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   674
    using assms by (auto simp: null_setsD1)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   675
qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   676
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   677
lemma null_set_Int1:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   678
  assumes "B \<in> null_sets M" "A \<in> sets M" shows "A \<inter> B \<in> null_sets M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   679
proof (intro CollectI conjI null_setsI)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   680
  show "emeasure M (A \<inter> B) = 0" using assms
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   681
    by (intro emeasure_eq_0[of B _ "A \<inter> B"]) auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   682
qed (insert assms, auto)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   683
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   684
lemma null_set_Int2:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   685
  assumes "B \<in> null_sets M" "A \<in> sets M" shows "B \<inter> A \<in> null_sets M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   686
  using assms by (subst Int_commute) (rule null_set_Int1)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   687
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   688
lemma emeasure_Diff_null_set:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   689
  assumes "B \<in> null_sets M" "A \<in> sets M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   690
  shows "emeasure M (A - B) = emeasure M A"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   691
proof -
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   692
  have *: "A - B = (A - (A \<inter> B))" by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   693
  have "A \<inter> B \<in> null_sets M" using assms by (rule null_set_Int1)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   694
  then show ?thesis
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   695
    unfolding * using assms
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   696
    by (subst emeasure_Diff) auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   697
qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   698
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   699
lemma null_set_Diff:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   700
  assumes "B \<in> null_sets M" "A \<in> sets M" shows "B - A \<in> null_sets M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   701
proof (intro CollectI conjI null_setsI)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   702
  show "emeasure M (B - A) = 0" using assms by (intro emeasure_eq_0[of B _ "B - A"]) auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   703
qed (insert assms, auto)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   704
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   705
lemma emeasure_Un_null_set:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   706
  assumes "A \<in> sets M" "B \<in> null_sets M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   707
  shows "emeasure M (A \<union> B) = emeasure M A"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   708
proof -
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   709
  have *: "A \<union> B = A \<union> (B - A)" by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   710
  have "B - A \<in> null_sets M" using assms(2,1) by (rule null_set_Diff)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   711
  then show ?thesis
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   712
    unfolding * using assms
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   713
    by (subst plus_emeasure[symmetric]) auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   714
qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   715
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   716
section "Formalize almost everywhere"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   717
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   718
definition ae_filter :: "'a measure \<Rightarrow> 'a filter" where
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   719
  "ae_filter M = Abs_filter (\<lambda>P. \<exists>N\<in>null_sets M. {x \<in> space M. \<not> P x} \<subseteq> N)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   720
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   721
abbreviation
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   722
  almost_everywhere :: "'a measure \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> bool" where
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   723
  "almost_everywhere M P \<equiv> eventually P (ae_filter M)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   724
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   725
syntax
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   726
  "_almost_everywhere" :: "pttrn \<Rightarrow> 'a \<Rightarrow> bool \<Rightarrow> bool" ("AE _ in _. _" [0,0,10] 10)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   727
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   728
translations
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   729
  "AE x in M. P" == "CONST almost_everywhere M (%x. P)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   730
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   731
lemma eventually_ae_filter:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   732
  fixes M P
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   733
  defines [simp]: "F \<equiv> \<lambda>P. \<exists>N\<in>null_sets M. {x \<in> space M. \<not> P x} \<subseteq> N" 
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   734
  shows "eventually P (ae_filter M) \<longleftrightarrow> F P"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   735
  unfolding ae_filter_def F_def[symmetric]
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   736
proof (rule eventually_Abs_filter)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   737
  show "is_filter F"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   738
  proof
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   739
    fix P Q assume "F P" "F Q"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   740
    then obtain N L where N: "N \<in> null_sets M" "{x \<in> space M. \<not> P x} \<subseteq> N"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   741
      and L: "L \<in> null_sets M" "{x \<in> space M. \<not> Q x} \<subseteq> L"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   742
      by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   743
    then have "L \<union> N \<in> null_sets M" "{x \<in> space M. \<not> (P x \<and> Q x)} \<subseteq> L \<union> N" by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   744
    then show "F (\<lambda>x. P x \<and> Q x)" by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   745
  next
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   746
    fix P Q assume "F P"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   747
    then obtain N where N: "N \<in> null_sets M" "{x \<in> space M. \<not> P x} \<subseteq> N" by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   748
    moreover assume "\<forall>x. P x \<longrightarrow> Q x"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   749
    ultimately have "N \<in> null_sets M" "{x \<in> space M. \<not> Q x} \<subseteq> N" by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   750
    then show "F Q" by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   751
  qed auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   752
qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   753
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   754
lemma AE_I':
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   755
  "N \<in> null_sets M \<Longrightarrow> {x\<in>space M. \<not> P x} \<subseteq> N \<Longrightarrow> (AE x in M. P x)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   756
  unfolding eventually_ae_filter by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   757
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   758
lemma AE_iff_null:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   759
  assumes "{x\<in>space M. \<not> P x} \<in> sets M" (is "?P \<in> sets M")
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   760
  shows "(AE x in M. P x) \<longleftrightarrow> {x\<in>space M. \<not> P x} \<in> null_sets M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   761
proof
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   762
  assume "AE x in M. P x" then obtain N where N: "N \<in> sets M" "?P \<subseteq> N" "emeasure M N = 0"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   763
    unfolding eventually_ae_filter by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   764
  have "0 \<le> emeasure M ?P" by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   765
  moreover have "emeasure M ?P \<le> emeasure M N"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   766
    using assms N(1,2) by (auto intro: emeasure_mono)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   767
  ultimately have "emeasure M ?P = 0" unfolding `emeasure M N = 0` by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   768
  then show "?P \<in> null_sets M" using assms by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   769
next
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   770
  assume "?P \<in> null_sets M" with assms show "AE x in M. P x" by (auto intro: AE_I')
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   771
qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   772
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   773
lemma AE_iff_null_sets:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   774
  "N \<in> sets M \<Longrightarrow> N \<in> null_sets M \<longleftrightarrow> (AE x in M. x \<notin> N)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   775
  using Int_absorb1[OF sets_into_space, of N M]
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   776
  by (subst AE_iff_null) (auto simp: Int_def[symmetric])
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   777
47761
dfe747e72fa8 moved lemmas to appropriate places
hoelzl
parents: 47694
diff changeset
   778
lemma AE_not_in:
dfe747e72fa8 moved lemmas to appropriate places
hoelzl
parents: 47694
diff changeset
   779
  "N \<in> null_sets M \<Longrightarrow> AE x in M. x \<notin> N"
dfe747e72fa8 moved lemmas to appropriate places
hoelzl
parents: 47694
diff changeset
   780
  by (metis AE_iff_null_sets null_setsD2)
dfe747e72fa8 moved lemmas to appropriate places
hoelzl
parents: 47694
diff changeset
   781
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   782
lemma AE_iff_measurable:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   783
  "N \<in> sets M \<Longrightarrow> {x\<in>space M. \<not> P x} = N \<Longrightarrow> (AE x in M. P x) \<longleftrightarrow> emeasure M N = 0"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   784
  using AE_iff_null[of _ P] by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   785
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   786
lemma AE_E[consumes 1]:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   787
  assumes "AE x in M. P x"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   788
  obtains N where "{x \<in> space M. \<not> P x} \<subseteq> N" "emeasure M N = 0" "N \<in> sets M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   789
  using assms unfolding eventually_ae_filter by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   790
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   791
lemma AE_E2:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   792
  assumes "AE x in M. P x" "{x\<in>space M. P x} \<in> sets M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   793
  shows "emeasure M {x\<in>space M. \<not> P x} = 0" (is "emeasure M ?P = 0")
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   794
proof -
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   795
  have "{x\<in>space M. \<not> P x} = space M - {x\<in>space M. P x}" by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   796
  with AE_iff_null[of M P] assms show ?thesis by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   797
qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   798
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   799
lemma AE_I:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   800
  assumes "{x \<in> space M. \<not> P x} \<subseteq> N" "emeasure M N = 0" "N \<in> sets M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   801
  shows "AE x in M. P x"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   802
  using assms unfolding eventually_ae_filter by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   803
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   804
lemma AE_mp[elim!]:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   805
  assumes AE_P: "AE x in M. P x" and AE_imp: "AE x in M. P x \<longrightarrow> Q x"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   806
  shows "AE x in M. Q x"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   807
proof -
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   808
  from AE_P obtain A where P: "{x\<in>space M. \<not> P x} \<subseteq> A"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   809
    and A: "A \<in> sets M" "emeasure M A = 0"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   810
    by (auto elim!: AE_E)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   811
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   812
  from AE_imp obtain B where imp: "{x\<in>space M. P x \<and> \<not> Q x} \<subseteq> B"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   813
    and B: "B \<in> sets M" "emeasure M B = 0"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   814
    by (auto elim!: AE_E)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   815
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   816
  show ?thesis
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   817
  proof (intro AE_I)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   818
    have "0 \<le> emeasure M (A \<union> B)" using A B by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   819
    moreover have "emeasure M (A \<union> B) \<le> 0"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   820
      using emeasure_subadditive[of A M B] A B by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   821
    ultimately show "A \<union> B \<in> sets M" "emeasure M (A \<union> B) = 0" using A B by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   822
    show "{x\<in>space M. \<not> Q x} \<subseteq> A \<union> B"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   823
      using P imp by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   824
  qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   825
qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   826
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   827
(* depricated replace by laws about eventually *)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   828
lemma
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   829
  shows AE_iffI: "AE x in M. P x \<Longrightarrow> AE x in M. P x \<longleftrightarrow> Q x \<Longrightarrow> AE x in M. Q x"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   830
    and AE_disjI1: "AE x in M. P x \<Longrightarrow> AE x in M. P x \<or> Q x"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   831
    and AE_disjI2: "AE x in M. Q x \<Longrightarrow> AE x in M. P x \<or> Q x"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   832
    and AE_conjI: "AE x in M. P x \<Longrightarrow> AE x in M. Q x \<Longrightarrow> AE x in M. P x \<and> Q x"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   833
    and AE_conj_iff[simp]: "(AE x in M. P x \<and> Q x) \<longleftrightarrow> (AE x in M. P x) \<and> (AE x in M. Q x)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   834
  by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   835
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   836
lemma AE_impI:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   837
  "(P \<Longrightarrow> AE x in M. Q x) \<Longrightarrow> AE x in M. P \<longrightarrow> Q x"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   838
  by (cases P) auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   839
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   840
lemma AE_measure:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   841
  assumes AE: "AE x in M. P x" and sets: "{x\<in>space M. P x} \<in> sets M" (is "?P \<in> sets M")
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   842
  shows "emeasure M {x\<in>space M. P x} = emeasure M (space M)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   843
proof -
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   844
  from AE_E[OF AE] guess N . note N = this
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   845
  with sets have "emeasure M (space M) \<le> emeasure M (?P \<union> N)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   846
    by (intro emeasure_mono) auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   847
  also have "\<dots> \<le> emeasure M ?P + emeasure M N"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   848
    using sets N by (intro emeasure_subadditive) auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   849
  also have "\<dots> = emeasure M ?P" using N by simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   850
  finally show "emeasure M ?P = emeasure M (space M)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   851
    using emeasure_space[of M "?P"] by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   852
qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   853
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   854
lemma AE_space: "AE x in M. x \<in> space M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   855
  by (rule AE_I[where N="{}"]) auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   856
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   857
lemma AE_I2[simp, intro]:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   858
  "(\<And>x. x \<in> space M \<Longrightarrow> P x) \<Longrightarrow> AE x in M. P x"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   859
  using AE_space by force
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   860
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   861
lemma AE_Ball_mp:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   862
  "\<forall>x\<in>space M. P x \<Longrightarrow> AE x in M. P x \<longrightarrow> Q x \<Longrightarrow> AE x in M. Q x"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   863
  by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   864
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   865
lemma AE_cong[cong]:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   866
  "(\<And>x. x \<in> space M \<Longrightarrow> P x \<longleftrightarrow> Q x) \<Longrightarrow> (AE x in M. P x) \<longleftrightarrow> (AE x in M. Q x)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   867
  by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   868
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   869
lemma AE_all_countable:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   870
  "(AE x in M. \<forall>i. P i x) \<longleftrightarrow> (\<forall>i::'i::countable. AE x in M. P i x)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   871
proof
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   872
  assume "\<forall>i. AE x in M. P i x"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   873
  from this[unfolded eventually_ae_filter Bex_def, THEN choice]
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   874
  obtain N where N: "\<And>i. N i \<in> null_sets M" "\<And>i. {x\<in>space M. \<not> P i x} \<subseteq> N i" by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   875
  have "{x\<in>space M. \<not> (\<forall>i. P i x)} \<subseteq> (\<Union>i. {x\<in>space M. \<not> P i x})" by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   876
  also have "\<dots> \<subseteq> (\<Union>i. N i)" using N by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   877
  finally have "{x\<in>space M. \<not> (\<forall>i. P i x)} \<subseteq> (\<Union>i. N i)" .
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   878
  moreover from N have "(\<Union>i. N i) \<in> null_sets M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   879
    by (intro null_sets_UN) auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   880
  ultimately show "AE x in M. \<forall>i. P i x"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   881
    unfolding eventually_ae_filter by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   882
qed auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   883
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   884
lemma AE_finite_all:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   885
  assumes f: "finite S" shows "(AE x in M. \<forall>i\<in>S. P i x) \<longleftrightarrow> (\<forall>i\<in>S. AE x in M. P i x)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   886
  using f by induct auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   887
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   888
lemma AE_finite_allI:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   889
  assumes "finite S"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   890
  shows "(\<And>s. s \<in> S \<Longrightarrow> AE x in M. Q s x) \<Longrightarrow> AE x in M. \<forall>s\<in>S. Q s x"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   891
  using AE_finite_all[OF `finite S`] by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   892
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   893
lemma emeasure_mono_AE:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   894
  assumes imp: "AE x in M. x \<in> A \<longrightarrow> x \<in> B"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   895
    and B: "B \<in> sets M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   896
  shows "emeasure M A \<le> emeasure M B"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   897
proof cases
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   898
  assume A: "A \<in> sets M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   899
  from imp obtain N where N: "{x\<in>space M. \<not> (x \<in> A \<longrightarrow> x \<in> B)} \<subseteq> N" "N \<in> null_sets M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   900
    by (auto simp: eventually_ae_filter)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   901
  have "emeasure M A = emeasure M (A - N)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   902
    using N A by (subst emeasure_Diff_null_set) auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   903
  also have "emeasure M (A - N) \<le> emeasure M (B - N)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   904
    using N A B sets_into_space by (auto intro!: emeasure_mono)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   905
  also have "emeasure M (B - N) = emeasure M B"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   906
    using N B by (subst emeasure_Diff_null_set) auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   907
  finally show ?thesis .
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   908
qed (simp add: emeasure_nonneg emeasure_notin_sets)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   909
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   910
lemma emeasure_eq_AE:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   911
  assumes iff: "AE x in M. x \<in> A \<longleftrightarrow> x \<in> B"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   912
  assumes A: "A \<in> sets M" and B: "B \<in> sets M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   913
  shows "emeasure M A = emeasure M B"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   914
  using assms by (safe intro!: antisym emeasure_mono_AE) auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   915
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   916
section {* @{text \<sigma>}-finite Measures *}
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   917
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   918
locale sigma_finite_measure =
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   919
  fixes M :: "'a measure"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   920
  assumes sigma_finite: "\<exists>A::nat \<Rightarrow> 'a set.
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   921
    range A \<subseteq> sets M \<and> (\<Union>i. A i) = space M \<and> (\<forall>i. emeasure M (A i) \<noteq> \<infinity>)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   922
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   923
lemma (in sigma_finite_measure) sigma_finite_disjoint:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   924
  obtains A :: "nat \<Rightarrow> 'a set"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   925
  where "range A \<subseteq> sets M" "(\<Union>i. A i) = space M" "\<And>i. emeasure M (A i) \<noteq> \<infinity>" "disjoint_family A"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   926
proof atomize_elim
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   927
  case goal1
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   928
  obtain A :: "nat \<Rightarrow> 'a set" where
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   929
    range: "range A \<subseteq> sets M" and
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   930
    space: "(\<Union>i. A i) = space M" and
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   931
    measure: "\<And>i. emeasure M (A i) \<noteq> \<infinity>"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   932
    using sigma_finite by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   933
  note range' = range_disjointed_sets[OF range] range
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   934
  { fix i
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   935
    have "emeasure M (disjointed A i) \<le> emeasure M (A i)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   936
      using range' disjointed_subset[of A i] by (auto intro!: emeasure_mono)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   937
    then have "emeasure M (disjointed A i) \<noteq> \<infinity>"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   938
      using measure[of i] by auto }
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   939
  with disjoint_family_disjointed UN_disjointed_eq[of A] space range'
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   940
  show ?case by (auto intro!: exI[of _ "disjointed A"])
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   941
qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   942
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   943
lemma (in sigma_finite_measure) sigma_finite_incseq:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   944
  obtains A :: "nat \<Rightarrow> 'a set"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   945
  where "range A \<subseteq> sets M" "(\<Union>i. A i) = space M" "\<And>i. emeasure M (A i) \<noteq> \<infinity>" "incseq A"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   946
proof atomize_elim
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   947
  case goal1
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   948
  obtain F :: "nat \<Rightarrow> 'a set" where
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   949
    F: "range F \<subseteq> sets M" "(\<Union>i. F i) = space M" "\<And>i. emeasure M (F i) \<noteq> \<infinity>"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   950
    using sigma_finite by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   951
  then show ?case
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   952
  proof (intro exI[of _ "\<lambda>n. \<Union>i\<le>n. F i"] conjI allI)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   953
    from F have "\<And>x. x \<in> space M \<Longrightarrow> \<exists>i. x \<in> F i" by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   954
    then show "(\<Union>n. \<Union> i\<le>n. F i) = space M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   955
      using F by fastforce
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   956
  next
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   957
    fix n
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   958
    have "emeasure M (\<Union> i\<le>n. F i) \<le> (\<Sum>i\<le>n. emeasure M (F i))" using F
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   959
      by (auto intro!: emeasure_subadditive_finite)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   960
    also have "\<dots> < \<infinity>"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   961
      using F by (auto simp: setsum_Pinfty)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   962
    finally show "emeasure M (\<Union> i\<le>n. F i) \<noteq> \<infinity>" by simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   963
  qed (force simp: incseq_def)+
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   964
qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   965
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   966
section {* Measure space induced by distribution of @{const measurable}-functions *}
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   967
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   968
definition distr :: "'a measure \<Rightarrow> 'b measure \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'b measure" where
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   969
  "distr M N f = measure_of (space N) (sets N) (\<lambda>A. emeasure M (f -` A \<inter> space M))"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   970
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   971
lemma
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   972
  shows sets_distr[simp]: "sets (distr M N f) = sets N"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   973
    and space_distr[simp]: "space (distr M N f) = space N"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   974
  by (auto simp: distr_def)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   975
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   976
lemma
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   977
  shows measurable_distr_eq1[simp]: "measurable (distr Mf Nf f) Mf' = measurable Nf Mf'"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   978
    and measurable_distr_eq2[simp]: "measurable Mg' (distr Mg Ng g) = measurable Mg' Ng"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   979
  by (auto simp: measurable_def)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   980
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   981
lemma emeasure_distr:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   982
  fixes f :: "'a \<Rightarrow> 'b"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   983
  assumes f: "f \<in> measurable M N" and A: "A \<in> sets N"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   984
  shows "emeasure (distr M N f) A = emeasure M (f -` A \<inter> space M)" (is "_ = ?\<mu> A")
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   985
  unfolding distr_def
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   986
proof (rule emeasure_measure_of_sigma)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   987
  show "positive (sets N) ?\<mu>"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   988
    by (auto simp: positive_def)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   989
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   990
  show "countably_additive (sets N) ?\<mu>"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   991
  proof (intro countably_additiveI)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   992
    fix A :: "nat \<Rightarrow> 'b set" assume "range A \<subseteq> sets N" "disjoint_family A"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   993
    then have A: "\<And>i. A i \<in> sets N" "(\<Union>i. A i) \<in> sets N" by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   994
    then have *: "range (\<lambda>i. f -` (A i) \<inter> space M) \<subseteq> sets M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   995
      using f by (auto simp: measurable_def)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   996
    moreover have "(\<Union>i. f -`  A i \<inter> space M) \<in> sets M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   997
      using * by blast
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   998
    moreover have **: "disjoint_family (\<lambda>i. f -` A i \<inter> space M)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   999
      using `disjoint_family A` by (auto simp: disjoint_family_on_def)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1000
    ultimately show "(\<Sum>i. ?\<mu> (A i)) = ?\<mu> (\<Union>i. A i)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1001
      using suminf_emeasure[OF _ **] A f
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1002
      by (auto simp: comp_def vimage_UN)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1003
  qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1004
  show "sigma_algebra (space N) (sets N)" ..
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1005
qed fact
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1006
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1007
lemma AE_distrD:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1008
  assumes f: "f \<in> measurable M M'"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1009
    and AE: "AE x in distr M M' f. P x"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1010
  shows "AE x in M. P (f x)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1011
proof -
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1012
  from AE[THEN AE_E] guess N .
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1013
  with f show ?thesis
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1014
    unfolding eventually_ae_filter
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1015
    by (intro bexI[of _ "f -` N \<inter> space M"])
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1016
       (auto simp: emeasure_distr measurable_def)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1017
qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1018
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1019
lemma null_sets_distr_iff:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1020
  "f \<in> measurable M N \<Longrightarrow> A \<in> null_sets (distr M N f) \<longleftrightarrow> f -` A \<inter> space M \<in> null_sets M \<and> A \<in> sets N"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1021
  by (auto simp add: null_sets_def emeasure_distr measurable_sets)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1022
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1023
lemma distr_distr:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1024
  assumes f: "g \<in> measurable N L" and g: "f \<in> measurable M N"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1025
  shows "distr (distr M N f) L g = distr M L (g \<circ> f)" (is "?L = ?R")
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1026
  using measurable_comp[OF g f] f g
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1027
  by (auto simp add: emeasure_distr measurable_sets measurable_space
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1028
           intro!: arg_cong[where f="emeasure M"] measure_eqI)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1029
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1030
section {* Real measure values *}
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1031
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1032
lemma measure_nonneg: "0 \<le> measure M A"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1033
  using emeasure_nonneg[of M A] unfolding measure_def by (auto intro: real_of_ereal_pos)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1034
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1035
lemma measure_empty[simp]: "measure M {} = 0"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1036
  unfolding measure_def by simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1037
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1038
lemma emeasure_eq_ereal_measure:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1039
  "emeasure M A \<noteq> \<infinity> \<Longrightarrow> emeasure M A = ereal (measure M A)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1040
  using emeasure_nonneg[of M A]
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1041
  by (cases "emeasure M A") (auto simp: measure_def)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1042
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1043
lemma measure_Union:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1044
  assumes finite: "emeasure M A \<noteq> \<infinity>" "emeasure M B \<noteq> \<infinity>"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1045
  and measurable: "A \<in> sets M" "B \<in> sets M" "A \<inter> B = {}"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1046
  shows "measure M (A \<union> B) = measure M A + measure M B"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1047
  unfolding measure_def
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1048
  using plus_emeasure[OF measurable, symmetric] finite
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1049
  by (simp add: emeasure_eq_ereal_measure)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1050
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1051
lemma measure_finite_Union:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1052
  assumes measurable: "A`S \<subseteq> sets M" "disjoint_family_on A S" "finite S"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1053
  assumes finite: "\<And>i. i \<in> S \<Longrightarrow> emeasure M (A i) \<noteq> \<infinity>"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1054
  shows "measure M (\<Union>i\<in>S. A i) = (\<Sum>i\<in>S. measure M (A i))"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1055
  unfolding measure_def
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1056
  using setsum_emeasure[OF measurable, symmetric] finite
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1057
  by (simp add: emeasure_eq_ereal_measure)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1058
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1059
lemma measure_Diff:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1060
  assumes finite: "emeasure M A \<noteq> \<infinity>"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1061
  and measurable: "A \<in> sets M" "B \<in> sets M" "B \<subseteq> A"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1062
  shows "measure M (A - B) = measure M A - measure M B"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1063
proof -
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1064
  have "emeasure M (A - B) \<le> emeasure M A" "emeasure M B \<le> emeasure M A"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1065
    using measurable by (auto intro!: emeasure_mono)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1066
  hence "measure M ((A - B) \<union> B) = measure M (A - B) + measure M B"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1067
    using measurable finite by (rule_tac measure_Union) auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1068
  thus ?thesis using `B \<subseteq> A` by (auto simp: Un_absorb2)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1069
qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1070
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1071
lemma measure_UNION:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1072
  assumes measurable: "range A \<subseteq> sets M" "disjoint_family A"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1073
  assumes finite: "emeasure M (\<Union>i. A i) \<noteq> \<infinity>"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1074
  shows "(\<lambda>i. measure M (A i)) sums (measure M (\<Union>i. A i))"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1075
proof -
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1076
  from summable_sums[OF summable_ereal_pos, of "\<lambda>i. emeasure M (A i)"]
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1077
       suminf_emeasure[OF measurable] emeasure_nonneg[of M]
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1078
  have "(\<lambda>i. emeasure M (A i)) sums (emeasure M (\<Union>i. A i))" by simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1079
  moreover
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1080
  { fix i
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1081
    have "emeasure M (A i) \<le> emeasure M (\<Union>i. A i)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1082
      using measurable by (auto intro!: emeasure_mono)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1083
    then have "emeasure M (A i) = ereal ((measure M (A i)))"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1084
      using finite by (intro emeasure_eq_ereal_measure) auto }
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1085
  ultimately show ?thesis using finite
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1086
    unfolding sums_ereal[symmetric] by (simp add: emeasure_eq_ereal_measure)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1087
qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1088
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1089
lemma measure_subadditive:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1090
  assumes measurable: "A \<in> sets M" "B \<in> sets M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1091
  and fin: "emeasure M A \<noteq> \<infinity>" "emeasure M B \<noteq> \<infinity>"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1092
  shows "(measure M (A \<union> B)) \<le> (measure M A) + (measure M B)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1093
proof -
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1094
  have "emeasure M (A \<union> B) \<noteq> \<infinity>"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1095
    using emeasure_subadditive[OF measurable] fin by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1096
  then show "(measure M (A \<union> B)) \<le> (measure M A) + (measure M B)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1097
    using emeasure_subadditive[OF measurable] fin
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1098
    by (auto simp: emeasure_eq_ereal_measure)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1099
qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1100
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1101
lemma measure_subadditive_finite:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1102
  assumes A: "finite I" "A`I \<subseteq> sets M" and fin: "\<And>i. i \<in> I \<Longrightarrow> emeasure M (A i) \<noteq> \<infinity>"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1103
  shows "measure M (\<Union>i\<in>I. A i) \<le> (\<Sum>i\<in>I. measure M (A i))"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1104
proof -
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1105
  { have "emeasure M (\<Union>i\<in>I. A i) \<le> (\<Sum>i\<in>I. emeasure M (A i))"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1106
      using emeasure_subadditive_finite[OF A] .
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1107
    also have "\<dots> < \<infinity>"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1108
      using fin by (simp add: setsum_Pinfty)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1109
    finally have "emeasure M (\<Union>i\<in>I. A i) \<noteq> \<infinity>" by simp }
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1110
  then show ?thesis
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1111
    using emeasure_subadditive_finite[OF A] fin
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1112
    unfolding measure_def by (simp add: emeasure_eq_ereal_measure suminf_ereal measure_nonneg)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1113
qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1114
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1115
lemma measure_subadditive_countably:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1116
  assumes A: "range A \<subseteq> sets M" and fin: "(\<Sum>i. emeasure M (A i)) \<noteq> \<infinity>"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1117
  shows "measure M (\<Union>i. A i) \<le> (\<Sum>i. measure M (A i))"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1118
proof -
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1119
  from emeasure_nonneg fin have "\<And>i. emeasure M (A i) \<noteq> \<infinity>" by (rule suminf_PInfty)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1120
  moreover
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1121
  { have "emeasure M (\<Union>i. A i) \<le> (\<Sum>i. emeasure M (A i))"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1122
      using emeasure_subadditive_countably[OF A] .
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1123
    also have "\<dots> < \<infinity>"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1124
      using fin by simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1125
    finally have "emeasure M (\<Union>i. A i) \<noteq> \<infinity>" by simp }
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1126
  ultimately  show ?thesis
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1127
    using emeasure_subadditive_countably[OF A] fin
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1128
    unfolding measure_def by (simp add: emeasure_eq_ereal_measure suminf_ereal measure_nonneg)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1129
qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1130
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1131
lemma measure_eq_setsum_singleton:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1132
  assumes S: "finite S" "\<And>x. x \<in> S \<Longrightarrow> {x} \<in> sets M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1133
  and fin: "\<And>x. x \<in> S \<Longrightarrow> emeasure M {x} \<noteq> \<infinity>"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1134
  shows "(measure M S) = (\<Sum>x\<in>S. (measure M {x}))"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1135
  unfolding measure_def
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1136
  using emeasure_eq_setsum_singleton[OF S] fin
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1137
  by simp (simp add: emeasure_eq_ereal_measure)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1138
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1139
lemma Lim_measure_incseq:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1140
  assumes A: "range A \<subseteq> sets M" "incseq A" and fin: "emeasure M (\<Union>i. A i) \<noteq> \<infinity>"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1141
  shows "(\<lambda>i. (measure M (A i))) ----> (measure M (\<Union>i. A i))"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1142
proof -
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1143
  have "ereal ((measure M (\<Union>i. A i))) = emeasure M (\<Union>i. A i)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1144
    using fin by (auto simp: emeasure_eq_ereal_measure)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1145
  then show ?thesis
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1146
    using Lim_emeasure_incseq[OF A]
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1147
    unfolding measure_def
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1148
    by (intro lim_real_of_ereal) simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1149
qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1150
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1151
lemma Lim_measure_decseq:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1152
  assumes A: "range A \<subseteq> sets M" "decseq A" and fin: "\<And>i. emeasure M (A i) \<noteq> \<infinity>"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1153
  shows "(\<lambda>n. measure M (A n)) ----> measure M (\<Inter>i. A i)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1154
proof -
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1155
  have "emeasure M (\<Inter>i. A i) \<le> emeasure M (A 0)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1156
    using A by (auto intro!: emeasure_mono)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1157
  also have "\<dots> < \<infinity>"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1158
    using fin[of 0] by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1159
  finally have "ereal ((measure M (\<Inter>i. A i))) = emeasure M (\<Inter>i. A i)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1160
    by (auto simp: emeasure_eq_ereal_measure)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1161
  then show ?thesis
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1162
    unfolding measure_def
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1163
    using Lim_emeasure_decseq[OF A fin]
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1164
    by (intro lim_real_of_ereal) simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1165
qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1166
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1167
section {* Measure spaces with @{term "emeasure M (space M) < \<infinity>"} *}
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1168
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1169
locale finite_measure = sigma_finite_measure M for M +
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1170
  assumes finite_emeasure_space: "emeasure M (space M) \<noteq> \<infinity>"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1171
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1172
lemma finite_measureI[Pure.intro!]:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1173
  assumes *: "emeasure M (space M) \<noteq> \<infinity>"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1174
  shows "finite_measure M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1175
proof
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1176
  show "\<exists>A. range A \<subseteq> sets M \<and> (\<Union>i. A i) = space M \<and> (\<forall>i. emeasure M (A i) \<noteq> \<infinity>)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1177
    using * by (auto intro!: exI[of _ "\<lambda>_. space M"])
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1178
qed fact
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1179
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1180
lemma (in finite_measure) emeasure_finite[simp, intro]: "emeasure M A \<noteq> \<infinity>"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1181
  using finite_emeasure_space emeasure_space[of M A] by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1182
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1183
lemma (in finite_measure) emeasure_eq_measure: "emeasure M A = ereal (measure M A)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1184
  unfolding measure_def by (simp add: emeasure_eq_ereal_measure)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1185
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1186
lemma (in finite_measure) emeasure_real: "\<exists>r. 0 \<le> r \<and> emeasure M A = ereal r"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1187
  using emeasure_finite[of A] emeasure_nonneg[of M A] by (cases "emeasure M A") auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1188
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1189
lemma (in finite_measure) bounded_measure: "measure M A \<le> measure M (space M)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1190
  using emeasure_space[of M A] emeasure_real[of A] emeasure_real[of "space M"] by (auto simp: measure_def)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1191
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1192
lemma (in finite_measure) finite_measure_Diff:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1193
  assumes sets: "A \<in> sets M" "B \<in> sets M" and "B \<subseteq> A"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1194
  shows "measure M (A - B) = measure M A - measure M B"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1195
  using measure_Diff[OF _ assms] by simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1196
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1197
lemma (in finite_measure) finite_measure_Union:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1198
  assumes sets: "A \<in> sets M" "B \<in> sets M" and "A \<inter> B = {}"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1199
  shows "measure M (A \<union> B) = measure M A + measure M B"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1200
  using measure_Union[OF _ _ assms] by simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1201
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1202
lemma (in finite_measure) finite_measure_finite_Union:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1203
  assumes measurable: "A`S \<subseteq> sets M" "disjoint_family_on A S" "finite S"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1204
  shows "measure M (\<Union>i\<in>S. A i) = (\<Sum>i\<in>S. measure M (A i))"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1205
  using measure_finite_Union[OF assms] by simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1206
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1207
lemma (in finite_measure) finite_measure_UNION:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1208
  assumes A: "range A \<subseteq> sets M" "disjoint_family A"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1209
  shows "(\<lambda>i. measure M (A i)) sums (measure M (\<Union>i. A i))"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1210
  using measure_UNION[OF A] by simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1211
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1212
lemma (in finite_measure) finite_measure_mono:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1213
  assumes "A \<subseteq> B" "B \<in> sets M" shows "measure M A \<le> measure M B"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1214
  using emeasure_mono[OF assms] emeasure_real[of A] emeasure_real[of B] by (auto simp: measure_def)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1215
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1216
lemma (in finite_measure) finite_measure_subadditive:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1217
  assumes m: "A \<in> sets M" "B \<in> sets M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1218
  shows "measure M (A \<union> B) \<le> measure M A + measure M B"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1219
  using measure_subadditive[OF m] by simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1220
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1221
lemma (in finite_measure) finite_measure_subadditive_finite:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1222
  assumes "finite I" "A`I \<subseteq> sets M" shows "measure M (\<Union>i\<in>I. A i) \<le> (\<Sum>i\<in>I. measure M (A i))"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1223
  using measure_subadditive_finite[OF assms] by simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1224
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1225
lemma (in finite_measure) finite_measure_subadditive_countably:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1226
  assumes A: "range A \<subseteq> sets M" and sum: "summable (\<lambda>i. measure M (A i))"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1227
  shows "measure M (\<Union>i. A i) \<le> (\<Sum>i. measure M (A i))"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1228
proof -
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1229
  from `summable (\<lambda>i. measure M (A i))`
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1230
  have "(\<lambda>i. ereal (measure M (A i))) sums ereal (\<Sum>i. measure M (A i))"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1231
    by (simp add: sums_ereal) (rule summable_sums)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1232
  from sums_unique[OF this, symmetric]
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1233
       measure_subadditive_countably[OF A]
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1234
  show ?thesis by (simp add: emeasure_eq_measure)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1235
qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1236
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1237
lemma (in finite_measure) finite_measure_eq_setsum_singleton:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1238
  assumes "finite S" and *: "\<And>x. x \<in> S \<Longrightarrow> {x} \<in> sets M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1239
  shows "measure M S = (\<Sum>x\<in>S. measure M {x})"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1240
  using measure_eq_setsum_singleton[OF assms] by simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1241
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1242
lemma (in finite_measure) finite_Lim_measure_incseq:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1243
  assumes A: "range A \<subseteq> sets M" "incseq A"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1244
  shows "(\<lambda>i. measure M (A i)) ----> measure M (\<Union>i. A i)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1245
  using Lim_measure_incseq[OF A] by simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1246
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1247
lemma (in finite_measure) finite_Lim_measure_decseq:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1248
  assumes A: "range A \<subseteq> sets M" "decseq A"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1249
  shows "(\<lambda>n. measure M (A n)) ----> measure M (\<Inter>i. A i)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1250
  using Lim_measure_decseq[OF A] by simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1251
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1252
lemma (in finite_measure) finite_measure_compl:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1253
  assumes S: "S \<in> sets M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1254
  shows "measure M (space M - S) = measure M (space M) - measure M S"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1255
  using measure_Diff[OF _ top S sets_into_space] S by simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1256
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1257
lemma (in finite_measure) finite_measure_mono_AE:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1258
  assumes imp: "AE x in M. x \<in> A \<longrightarrow> x \<in> B" and B: "B \<in> sets M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1259
  shows "measure M A \<le> measure M B"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1260
  using assms emeasure_mono_AE[OF imp B]
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1261
  by (simp add: emeasure_eq_measure)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1262
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1263
lemma (in finite_measure) finite_measure_eq_AE:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1264
  assumes iff: "AE x in M. x \<in> A \<longleftrightarrow> x \<in> B"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1265
  assumes A: "A \<in> sets M" and B: "B \<in> sets M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1266
  shows "measure M A = measure M B"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1267
  using assms emeasure_eq_AE[OF assms] by (simp add: emeasure_eq_measure)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1268
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1269
section {* Counting space *}
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1270
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1271
definition count_space :: "'a set \<Rightarrow> 'a measure" where
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1272
  "count_space \<Omega> = measure_of \<Omega> (Pow \<Omega>) (\<lambda>A. if finite A then ereal (card A) else \<infinity>)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1273
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1274
lemma 
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1275
  shows space_count_space[simp]: "space (count_space \<Omega>) = \<Omega>"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1276
    and sets_count_space[simp]: "sets (count_space \<Omega>) = Pow \<Omega>"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1277
  using sigma_sets_into_sp[of "Pow \<Omega>" \<Omega>]
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1278
  by (auto simp: count_space_def)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1279
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1280
lemma measurable_count_space_eq1[simp]:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1281
  "f \<in> measurable (count_space A) M \<longleftrightarrow> f \<in> A \<rightarrow> space M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1282
 unfolding measurable_def by simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1283
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1284
lemma measurable_count_space_eq2[simp]:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1285
  assumes "finite A"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1286
  shows "f \<in> measurable M (count_space A) \<longleftrightarrow> (f \<in> space M \<rightarrow> A \<and> (\<forall>a\<in>A. f -` {a} \<inter> space M \<in> sets M))"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1287
proof -
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1288
  { fix X assume "X \<subseteq> A" "f \<in> space M \<rightarrow> A"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1289
    with `finite A` have "f -` X \<inter> space M = (\<Union>a\<in>X. f -` {a} \<inter> space M)" "finite X"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1290
      by (auto dest: finite_subset)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1291
    moreover assume "\<forall>a\<in>A. f -` {a} \<inter> space M \<in> sets M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1292
    ultimately have "f -` X \<inter> space M \<in> sets M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1293
      using `X \<subseteq> A` by (auto intro!: finite_UN simp del: UN_simps) }
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1294
  then show ?thesis
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1295
    unfolding measurable_def by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1296
qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1297
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1298
lemma emeasure_count_space:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1299
  assumes "X \<subseteq> A" shows "emeasure (count_space A) X = (if finite X then ereal (card X) else \<infinity>)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1300
    (is "_ = ?M X")
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1301
  unfolding count_space_def
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1302
proof (rule emeasure_measure_of_sigma)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1303
  show "sigma_algebra A (Pow A)" by (rule sigma_algebra_Pow)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1304
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1305
  show "positive (Pow A) ?M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1306
    by (auto simp: positive_def)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1307
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1308
  show "countably_additive (Pow A) ?M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1309
  proof (unfold countably_additive_def, safe)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1310
      fix F :: "nat \<Rightarrow> 'a set" assume disj: "disjoint_family F"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1311
      show "(\<Sum>i. ?M (F i)) = ?M (\<Union>i. F i)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1312
      proof cases
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1313
        assume "\<forall>i. finite (F i)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1314
        then have finite_F: "\<And>i. finite (F i)" by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1315
        have "\<forall>i\<in>{i. F i \<noteq> {}}. \<exists>x. x \<in> F i" by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1316
        from bchoice[OF this] obtain f where f: "\<And>i. F i \<noteq> {} \<Longrightarrow> f i \<in> F i" by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1317
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1318
        have inj_f: "inj_on f {i. F i \<noteq> {}}"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1319
        proof (rule inj_onI, simp)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1320
          fix i j a b assume *: "f i = f j" "F i \<noteq> {}" "F j \<noteq> {}"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1321
          then have "f i \<in> F i" "f j \<in> F j" using f by force+
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1322
          with disj * show "i = j" by (auto simp: disjoint_family_on_def)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1323
        qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1324
        have fin_eq: "finite (\<Union>i. F i) \<longleftrightarrow> finite {i. F i \<noteq> {}}"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1325
        proof
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1326
          assume "finite (\<Union>i. F i)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1327
          show "finite {i. F i \<noteq> {}}"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1328
          proof (rule finite_imageD)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1329
            from f have "f`{i. F i \<noteq> {}} \<subseteq> (\<Union>i. F i)" by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1330
            then show "finite (f`{i. F i \<noteq> {}})"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1331
              by (rule finite_subset) fact
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1332
          qed fact
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1333
        next
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1334
          assume "finite {i. F i \<noteq> {}}"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1335
          with finite_F have "finite (\<Union>i\<in>{i. F i \<noteq> {}}. F i)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1336
            by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1337
          also have "(\<Union>i\<in>{i. F i \<noteq> {}}. F i) = (\<Union>i. F i)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1338
            by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1339
          finally show "finite (\<Union>i. F i)" .
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1340
        qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1341
        
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1342
        show ?thesis
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1343
        proof cases
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1344
          assume *: "finite (\<Union>i. F i)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1345
          with finite_F have "finite {i. ?M (F i) \<noteq> 0} "
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1346
            by (simp add: fin_eq)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1347
          then have "(\<Sum>i. ?M (F i)) = (\<Sum>i | ?M (F i) \<noteq> 0. ?M (F i))"
47761
dfe747e72fa8 moved lemmas to appropriate places
hoelzl
parents: 47694
diff changeset
  1348
            by (rule suminf_finite) auto
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1349
          also have "\<dots> = ereal (\<Sum>i | F i \<noteq> {}. card (F i))"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1350
            using finite_F by simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1351
          also have "\<dots> = ereal (card (\<Union>i \<in> {i. F i \<noteq> {}}. F i))"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1352
            using * finite_F disj
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1353
            by (subst card_UN_disjoint) (auto simp: disjoint_family_on_def fin_eq)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1354
          also have "\<dots> = ?M (\<Union>i. F i)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1355
            using * by (auto intro!: arg_cong[where f=card])
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1356
          finally show ?thesis .
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1357
        next
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1358
          assume inf: "infinite (\<Union>i. F i)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1359
          { fix i
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1360
            have "\<exists>N. i \<le> (\<Sum>i<N. card (F i))"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1361
            proof (induct i)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1362
              case (Suc j)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1363
              from Suc obtain N where N: "j \<le> (\<Sum>i<N. card (F i))" by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1364
              have "infinite ({i. F i \<noteq> {}} - {..< N})"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1365
                using inf by (auto simp: fin_eq)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1366
              then have "{i. F i \<noteq> {}} - {..< N} \<noteq> {}"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1367
                by (metis finite.emptyI)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1368
              then obtain i where i: "F i \<noteq> {}" "N \<le> i"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1369
                by (auto simp: not_less[symmetric])
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1370
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1371
              note N
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1372
              also have "(\<Sum>i<N. card (F i)) \<le> (\<Sum>i<i. card (F i))"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1373
                by (rule setsum_mono2) (auto simp: i)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1374
              also have "\<dots> < (\<Sum>i<i. card (F i)) + card (F i)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1375
                using finite_F `F i \<noteq> {}` by (simp add: card_gt_0_iff)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1376
              finally have "j < (\<Sum>i<Suc i. card (F i))"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1377
                by simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1378
              then show ?case unfolding Suc_le_eq by blast
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1379
            qed simp }
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1380
          with finite_F inf show ?thesis
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1381
            by (auto simp del: real_of_nat_setsum intro!: SUP_PInfty
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1382
                     simp add: suminf_ereal_eq_SUPR real_of_nat_setsum[symmetric])
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1383
        qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1384
      next
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1385
        assume "\<not> (\<forall>i. finite (F i))"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1386
        then obtain j where j: "infinite (F j)" by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1387
        then have "infinite (\<Union>i. F i)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1388
          using finite_subset[of "F j" "\<Union>i. F i"] by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1389
        moreover have "\<And>i. 0 \<le> ?M (F i)" by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1390
        ultimately show ?thesis
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1391
          using suminf_PInfty[of "\<lambda>i. ?M (F i)" j] j by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1392
      qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1393
  qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1394
  show "X \<in> Pow A" using `X \<subseteq> A` by simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1395
qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1396
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1397
lemma emeasure_count_space_finite[simp]:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1398
  "X \<subseteq> A \<Longrightarrow> finite X \<Longrightarrow> emeasure (count_space A) X = ereal (card X)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1399
  using emeasure_count_space[of X A] by simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1400
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1401
lemma emeasure_count_space_infinite[simp]:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1402
  "X \<subseteq> A \<Longrightarrow> infinite X \<Longrightarrow> emeasure (count_space A) X = \<infinity>"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1403
  using emeasure_count_space[of X A] by simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1404
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1405
lemma emeasure_count_space_eq_0:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1406
  "emeasure (count_space A) X = 0 \<longleftrightarrow> (X \<subseteq> A \<longrightarrow> X = {})"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1407
proof cases
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1408
  assume X: "X \<subseteq> A"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1409
  then show ?thesis
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1410
  proof (intro iffI impI)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1411
    assume "emeasure (count_space A) X = 0"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1412
    with X show "X = {}"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1413
      by (subst (asm) emeasure_count_space) (auto split: split_if_asm)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1414
  qed simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1415
qed (simp add: emeasure_notin_sets)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1416
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1417
lemma null_sets_count_space: "null_sets (count_space A) = { {} }"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1418
  unfolding null_sets_def by (auto simp add: emeasure_count_space_eq_0)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1419
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1420
lemma AE_count_space: "(AE x in count_space A. P x) \<longleftrightarrow> (\<forall>x\<in>A. P x)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1421
  unfolding eventually_ae_filter by (auto simp add: null_sets_count_space)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1422
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1423
lemma sigma_finite_measure_count_space:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1424
  fixes A :: "'a::countable set"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1425
  shows "sigma_finite_measure (count_space A)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1426
proof
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1427
  show "\<exists>F::nat \<Rightarrow> 'a set. range F \<subseteq> sets (count_space A) \<and> (\<Union>i. F i) = space (count_space A) \<and>
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1428
     (\<forall>i. emeasure (count_space A) (F i) \<noteq> \<infinity>)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1429
     using surj_from_nat by (intro exI[of _ "\<lambda>i. {from_nat i} \<inter> A"]) (auto simp del: surj_from_nat)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1430
qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1431
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1432
lemma finite_measure_count_space:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1433
  assumes [simp]: "finite A"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1434
  shows "finite_measure (count_space A)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1435
  by rule simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1436
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1437
lemma sigma_finite_measure_count_space_finite:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1438
  assumes A: "finite A" shows "sigma_finite_measure (count_space A)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1439
proof -
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1440
  interpret finite_measure "count_space A" using A by (rule finite_measure_count_space)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1441
  show "sigma_finite_measure (count_space A)" ..
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1442
qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1443
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1444
end
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1445