src/HOL/Probability/Measure.thy
author blanchet
Thu, 02 Sep 2010 22:50:16 +0200
changeset 39110 a74bd9bfa880
parent 38656 d5d342611edb
child 39089 df379a447753
permissions -rw-r--r--
show the number of facts for each prover in "verbose" mode
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
     1
header {*Measures*}
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
     2
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
     3
theory Measure
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
     4
  imports Caratheodory
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
     5
begin
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
     6
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
     7
text{*From the Hurd/Coble measure theory development, translated by Lawrence Paulson.*}
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
     8
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
     9
section {* Equations for the measure function @{text \<mu>} *}
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    10
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    11
lemma (in measure_space) measure_countably_additive:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    12
  assumes "range A \<subseteq> sets M" "disjoint_family A"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    13
  shows "psuminf (\<lambda>i. \<mu> (A i)) = \<mu> (\<Union>i. A i)"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    14
proof -
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    15
  have "(\<Union> i. A i) \<in> sets M" using assms(1) by (rule countable_UN)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    16
  with ca assms show ?thesis by (simp add: countably_additive_def)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    17
qed
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    18
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    19
lemma (in measure_space) measure_space_cong:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    20
  assumes "\<And>A. A \<in> sets M \<Longrightarrow> \<nu> A = \<mu> A"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    21
  shows "measure_space M \<nu>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    22
proof
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    23
  show "\<nu> {} = 0" using assms by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    24
  show "countably_additive M \<nu>" unfolding countably_additive_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    25
  proof safe
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    26
    fix A :: "nat \<Rightarrow> 'a set" assume A: "range A \<subseteq> sets M" "disjoint_family A"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    27
    then have "\<And>i. A i \<in> sets M" "(UNION UNIV A) \<in> sets M" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    28
    from this[THEN assms] measure_countably_additive[OF A]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    29
    show "(\<Sum>\<^isub>\<infinity>n. \<nu> (A n)) = \<nu> (UNION UNIV A)" by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    30
  qed
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    31
qed
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    32
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    33
lemma (in measure_space) additive: "additive M \<mu>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    34
proof (rule algebra.countably_additive_additive [OF _ _ ca])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    35
  show "algebra M" by default
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    36
  show "positive \<mu>" by (simp add: positive_def)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    37
qed
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    38
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    39
lemma (in measure_space) measure_additive:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    40
     "a \<in> sets M \<Longrightarrow> b \<in> sets M \<Longrightarrow> a \<inter> b = {}
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    41
      \<Longrightarrow> \<mu> a + \<mu> b = \<mu> (a \<union> b)"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    42
  by (metis additiveD additive)
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    43
36624
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
    44
lemma (in measure_space) measure_mono:
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
    45
  assumes "a \<subseteq> b" "a \<in> sets M" "b \<in> sets M"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    46
  shows "\<mu> a \<le> \<mu> b"
36624
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
    47
proof -
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
    48
  have "b = a \<union> (b - a)" using assms by auto
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
    49
  moreover have "{} = a \<inter> (b - a)" by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    50
  ultimately have "\<mu> b = \<mu> a + \<mu> (b - a)"
36624
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
    51
    using measure_additive[of a "b - a"] local.Diff[of b a] assms by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    52
  moreover have "\<mu> (b - a) \<ge> 0" using assms by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    53
  ultimately show "\<mu> a \<le> \<mu> b" by auto
36624
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
    54
qed
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
    55
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    56
lemma (in measure_space) measure_compl:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    57
  assumes s: "s \<in> sets M" and fin: "\<mu> s \<noteq> \<omega>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    58
  shows "\<mu> (space M - s) = \<mu> (space M) - \<mu> s"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    59
proof -
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    60
  have s_less_space: "\<mu> s \<le> \<mu> (space M)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    61
    using s by (auto intro!: measure_mono sets_into_space)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    62
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    63
  have "\<mu> (space M) = \<mu> (s \<union> (space M - s))" using s
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    64
    by (metis Un_Diff_cancel Un_absorb1 s sets_into_space)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    65
  also have "... = \<mu> s + \<mu> (space M - s)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    66
    by (rule additiveD [OF additive]) (auto simp add: s)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    67
  finally have "\<mu> (space M) = \<mu> s + \<mu> (space M - s)" .
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    68
  thus ?thesis
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    69
    unfolding minus_pinfreal_eq2[OF s_less_space fin]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    70
    by (simp add: ac_simps)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    71
qed
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    72
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    73
lemma (in measure_space) measure_Diff:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    74
  assumes finite: "\<mu> B \<noteq> \<omega>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    75
  and measurable: "A \<in> sets M" "B \<in> sets M" "B \<subseteq> A"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    76
  shows "\<mu> (A - B) = \<mu> A - \<mu> B"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    77
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    78
  have *: "(A - B) \<union> B = A" using `B \<subseteq> A` by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    79
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    80
  have "\<mu> ((A - B) \<union> B) = \<mu> (A - B) + \<mu> B"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    81
    using measurable by (rule_tac measure_additive[symmetric]) auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    82
  thus ?thesis unfolding * using `\<mu> B \<noteq> \<omega>`
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    83
    by (simp add: pinfreal_cancel_plus_minus)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    84
qed
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    85
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    86
lemma (in measure_space) measure_countable_increasing:
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    87
  assumes A: "range A \<subseteq> sets M"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    88
      and A0: "A 0 = {}"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    89
      and ASuc: "\<And>n.  A n \<subseteq> A (Suc n)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    90
  shows "(SUP n. \<mu> (A n)) = \<mu> (\<Union>i. A i)"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    91
proof -
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    92
  {
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    93
    fix n
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    94
    have "\<mu> (A n) =
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    95
          setsum (\<mu> \<circ> (\<lambda>i. A (Suc i) - A i)) {..<n}"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    96
      proof (induct n)
37032
58a0757031dd speed up some proofs and fix some warnings
huffman
parents: 36670
diff changeset
    97
        case 0 thus ?case by (auto simp add: A0)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    98
      next
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
    99
        case (Suc m)
33536
fd28b7399f2b eliminated hard tabulators;
wenzelm
parents: 33273
diff changeset
   100
        have "A (Suc m) = A m \<union> (A (Suc m) - A m)"
fd28b7399f2b eliminated hard tabulators;
wenzelm
parents: 33273
diff changeset
   101
          by (metis ASuc Un_Diff_cancel Un_absorb1)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   102
        hence "\<mu> (A (Suc m)) =
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   103
               \<mu> (A m) + \<mu> (A (Suc m) - A m)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   104
          by (subst measure_additive)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   105
             (auto simp add: measure_additive range_subsetD [OF A])
33536
fd28b7399f2b eliminated hard tabulators;
wenzelm
parents: 33273
diff changeset
   106
        with Suc show ?case
fd28b7399f2b eliminated hard tabulators;
wenzelm
parents: 33273
diff changeset
   107
          by simp
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   108
      qed }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   109
  note Meq = this
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   110
  have Aeq: "(\<Union>i. A (Suc i) - A i) = (\<Union>i. A i)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   111
    proof (rule UN_finite2_eq [where k=1], simp)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   112
      fix i
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   113
      show "(\<Union>i\<in>{0..<i}. A (Suc i) - A i) = (\<Union>i\<in>{0..<Suc i}. A i)"
33536
fd28b7399f2b eliminated hard tabulators;
wenzelm
parents: 33273
diff changeset
   114
        proof (induct i)
fd28b7399f2b eliminated hard tabulators;
wenzelm
parents: 33273
diff changeset
   115
          case 0 thus ?case by (simp add: A0)
fd28b7399f2b eliminated hard tabulators;
wenzelm
parents: 33273
diff changeset
   116
        next
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   117
          case (Suc i)
33536
fd28b7399f2b eliminated hard tabulators;
wenzelm
parents: 33273
diff changeset
   118
          thus ?case
fd28b7399f2b eliminated hard tabulators;
wenzelm
parents: 33273
diff changeset
   119
            by (auto simp add: atLeastLessThanSuc intro: subsetD [OF ASuc])
fd28b7399f2b eliminated hard tabulators;
wenzelm
parents: 33273
diff changeset
   120
        qed
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   121
    qed
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   122
  have A1: "\<And>i. A (Suc i) - A i \<in> sets M"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   123
    by (metis A Diff range_subsetD)
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   124
  have A2: "(\<Union>i. A (Suc i) - A i) \<in> sets M"
37032
58a0757031dd speed up some proofs and fix some warnings
huffman
parents: 36670
diff changeset
   125
    by (blast intro: range_subsetD [OF A])
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   126
  have "psuminf ( (\<lambda>i. \<mu> (A (Suc i) - A i))) = \<mu> (\<Union>i. A (Suc i) - A i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   127
    by (rule measure_countably_additive)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   128
       (auto simp add: disjoint_family_Suc ASuc A1 A2)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   129
  also have "... =  \<mu> (\<Union>i. A i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   130
    by (simp add: Aeq)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   131
  finally have "psuminf (\<lambda>i. \<mu> (A (Suc i) - A i)) = \<mu> (\<Union>i. A i)" .
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   132
  thus ?thesis
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   133
    by (auto simp add: Meq psuminf_def)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   134
qed
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   135
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   136
lemma (in measure_space) continuity_from_below:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   137
  assumes A: "range A \<subseteq> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   138
      and ASuc: "!!n.  A n \<subseteq> A (Suc n)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   139
  shows "(SUP n. \<mu> (A n)) = \<mu> (\<Union>i. A i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   140
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   141
  have *: "(SUP n. \<mu> (nat_case {} A (Suc n))) = (SUP n. \<mu> (nat_case {} A n))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   142
    apply (rule Sup_mono_offset_Suc)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   143
    apply (rule measure_mono)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   144
    using assms by (auto split: nat.split)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   145
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   146
  have ueq: "(\<Union>i. nat_case {} A i) = (\<Union>i. A i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   147
    by (auto simp add: split: nat.splits)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   148
  have meq: "\<And>n. \<mu> (A n) = (\<mu> \<circ> nat_case {} A) (Suc n)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   149
    by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   150
  have "(SUP n. \<mu> (nat_case {} A n)) = \<mu> (\<Union>i. nat_case {} A i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   151
    by (rule measure_countable_increasing)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   152
       (auto simp add: range_subsetD [OF A] subsetD [OF ASuc] split: nat.splits)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   153
  also have "\<mu> (\<Union>i. nat_case {} A i) = \<mu> (\<Union>i. A i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   154
    by (simp add: ueq)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   155
  finally have "(SUP n. \<mu> (nat_case {} A n)) = \<mu> (\<Union>i. A i)" .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   156
  thus ?thesis unfolding meq * comp_def .
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   157
qed
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   158
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   159
lemma (in measure_space) measure_up:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   160
  assumes "P \<in> sets M" "\<And>i. B i \<in> sets M" "B \<up> P"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   161
  shows "(\<lambda>i. \<mu> (B i)) \<up> \<mu> P"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   162
  using assms unfolding isoton_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   163
  by (auto intro!: measure_mono continuity_from_below)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   164
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   165
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   166
lemma (in measure_space) continuity_from_above:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   167
  assumes A: "range A \<subseteq> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   168
  and mono_Suc: "\<And>n.  A (Suc n) \<subseteq> A n"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   169
  and finite: "\<And>i. \<mu> (A i) \<noteq> \<omega>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   170
  shows "(INF n. \<mu> (A n)) = \<mu> (\<Inter>i. A i)"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   171
proof -
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   172
  { fix n have "A n \<subseteq> A 0" using mono_Suc by (induct n) auto }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   173
  note mono = this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   174
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   175
  have le_MI: "\<mu> (\<Inter>i. A i) \<le> \<mu> (A 0)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   176
    using A by (auto intro!: measure_mono)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   177
  hence *: "\<mu> (\<Inter>i. A i) \<noteq> \<omega>" using finite[of 0] by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   178
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   179
  have le_IM: "(INF n. \<mu> (A n)) \<le> \<mu> (A 0)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   180
    by (rule INF_leI) simp
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   181
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   182
  have "\<mu> (A 0) - (INF n. \<mu> (A n)) = (SUP n. \<mu> (A 0 - A n))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   183
    unfolding pinfreal_SUP_minus[symmetric]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   184
    using mono A finite by (subst measure_Diff) auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   185
  also have "\<dots> = \<mu> (\<Union>i. A 0 - A i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   186
  proof (rule continuity_from_below)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   187
    show "range (\<lambda>n. A 0 - A n) \<subseteq> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   188
      using A by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   189
    show "\<And>n. A 0 - A n \<subseteq> A 0 - A (Suc n)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   190
      using mono_Suc by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   191
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   192
  also have "\<dots> = \<mu> (A 0) - \<mu> (\<Inter>i. A i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   193
    using mono A finite * by (simp, subst measure_Diff) auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   194
  finally show ?thesis
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   195
    by (rule pinfreal_diff_eq_diff_imp_eq[OF finite[of 0] le_IM le_MI])
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   196
qed
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   197
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   198
lemma (in measure_space) measure_down:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   199
  assumes "P \<in> sets M" "\<And>i. B i \<in> sets M" "B \<down> P"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   200
  and finite: "\<And>i. \<mu> (B i) \<noteq> \<omega>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   201
  shows "(\<lambda>i. \<mu> (B i)) \<down> \<mu> P"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   202
  using assms unfolding antiton_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   203
  by (auto intro!: measure_mono continuity_from_above)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   204
lemma (in measure_space) measure_insert:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   205
  assumes sets: "{x} \<in> sets M" "A \<in> sets M" and "x \<notin> A"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   206
  shows "\<mu> (insert x A) = \<mu> {x} + \<mu> A"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   207
proof -
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   208
  have "{x} \<inter> A = {}" using `x \<notin> A` by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   209
  from measure_additive[OF sets this] show ?thesis by simp
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   210
qed
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   211
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   212
lemma (in measure_space) measure_finite_singleton:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   213
  assumes fin: "finite S"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   214
  and ssets: "\<And>x. x \<in> S \<Longrightarrow> {x} \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   215
  shows "\<mu> S = (\<Sum>x\<in>S. \<mu> {x})"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   216
using assms proof induct
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   217
  case (insert x S)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   218
  have *: "\<mu> S = (\<Sum>x\<in>S. \<mu> {x})" "{x} \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   219
    using insert.prems by (blast intro!: insert.hyps(3))+
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   220
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   221
  have "(\<Union>x\<in>S. {x}) \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   222
    using  insert.prems `finite S` by (blast intro!: finite_UN)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   223
  hence "S \<in> sets M" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   224
  from measure_insert[OF `{x} \<in> sets M` this `x \<notin> S`]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   225
  show ?case using `x \<notin> S` `finite S` * by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   226
qed simp
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33657
diff changeset
   227
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33657
diff changeset
   228
lemma (in measure_space) measure_finitely_additive':
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   229
  assumes "f \<in> ({..< n :: nat} \<rightarrow> sets M)"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33657
diff changeset
   230
  assumes "\<And> a b. \<lbrakk>a < n ; b < n ; a \<noteq> b\<rbrakk> \<Longrightarrow> f a \<inter> f b = {}"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   231
  assumes "s = \<Union> (f ` {..< n})"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   232
  shows "(\<Sum>i<n. (\<mu> \<circ> f) i) = \<mu> s"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33657
diff changeset
   233
proof -
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33657
diff changeset
   234
  def f' == "\<lambda> i. (if i < n then f i else {})"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   235
  have rf: "range f' \<subseteq> sets M" unfolding f'_def
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33657
diff changeset
   236
    using assms empty_sets by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   237
  have df: "disjoint_family f'" unfolding f'_def disjoint_family_on_def
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33657
diff changeset
   238
    using assms by simp
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   239
  have "\<Union> range f' = (\<Union> i \<in> {..< n}. f i)"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33657
diff changeset
   240
    unfolding f'_def by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   241
  then have "\<mu> s = \<mu> (\<Union> range f')"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33657
diff changeset
   242
    using assms by simp
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   243
  then have part1: "(\<Sum>\<^isub>\<infinity> i. \<mu> (f' i)) = \<mu> s"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33657
diff changeset
   244
    using df rf ca[unfolded countably_additive_def, rule_format, of f']
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33657
diff changeset
   245
    by auto
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33657
diff changeset
   246
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   247
  have "(\<Sum>\<^isub>\<infinity> i. \<mu> (f' i)) = (\<Sum> i< n. \<mu> (f' i))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   248
    by (rule psuminf_finite) (simp add: f'_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   249
  also have "\<dots> = (\<Sum>i<n. \<mu> (f i))"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33657
diff changeset
   250
    unfolding f'_def by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   251
  finally have part2: "(\<Sum>\<^isub>\<infinity> i. \<mu> (f' i)) = (\<Sum>i<n. \<mu> (f i))" by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   252
  show ?thesis using part1 part2 by auto
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33657
diff changeset
   253
qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33657
diff changeset
   254
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33657
diff changeset
   255
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33657
diff changeset
   256
lemma (in measure_space) measure_finitely_additive:
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33657
diff changeset
   257
  assumes "finite r"
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33657
diff changeset
   258
  assumes "r \<subseteq> sets M"
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33657
diff changeset
   259
  assumes d: "\<And> a b. \<lbrakk>a \<in> r ; b \<in> r ; a \<noteq> b\<rbrakk> \<Longrightarrow> a \<inter> b = {}"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   260
  shows "(\<Sum> i \<in> r. \<mu> i) = \<mu> (\<Union> r)"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33657
diff changeset
   261
using assms
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33657
diff changeset
   262
proof -
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33657
diff changeset
   263
  (* counting the elements in r is enough *)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   264
  let ?R = "{..<card r}"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33657
diff changeset
   265
  obtain f where f: "f ` ?R = r" "inj_on f ?R"
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33657
diff changeset
   266
    using ex_bij_betw_nat_finite[unfolded bij_betw_def, OF `finite r`]
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   267
    unfolding atLeast0LessThan by auto
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33657
diff changeset
   268
  hence f_into_sets: "f \<in> ?R \<rightarrow> sets M" using assms by auto
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33657
diff changeset
   269
  have disj: "\<And> a b. \<lbrakk>a \<in> ?R ; b \<in> ?R ; a \<noteq> b\<rbrakk> \<Longrightarrow> f a \<inter> f b = {}"
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33657
diff changeset
   270
  proof -
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33657
diff changeset
   271
    fix a b assume asm: "a \<in> ?R" "b \<in> ?R" "a \<noteq> b"
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33657
diff changeset
   272
    hence neq: "f a \<noteq> f b" using f[unfolded inj_on_def, rule_format] by blast
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33657
diff changeset
   273
    from asm have "f a \<in> r" "f b \<in> r" using f by auto
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33657
diff changeset
   274
    thus "f a \<inter> f b = {}" using d[of "f a" "f b"] f using neq by auto
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33657
diff changeset
   275
  qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33657
diff changeset
   276
  have "(\<Union> r) = (\<Union> i \<in> ?R. f i)"
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33657
diff changeset
   277
    using f by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   278
  hence "\<mu> (\<Union> r)= \<mu> (\<Union> i \<in> ?R. f i)" by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   279
  also have "\<dots> = (\<Sum> i \<in> ?R. \<mu> (f i))"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33657
diff changeset
   280
    using measure_finitely_additive'[OF f_into_sets disj] by simp
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   281
  also have "\<dots> = (\<Sum> a \<in> r. \<mu> a)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   282
    using f[rule_format] setsum_reindex[of f ?R "\<lambda> a. \<mu> a"] by auto
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33657
diff changeset
   283
  finally show ?thesis by simp
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33657
diff changeset
   284
qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33657
diff changeset
   285
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33657
diff changeset
   286
lemma (in measure_space) measure_finitely_additive'':
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33657
diff changeset
   287
  assumes "finite s"
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33657
diff changeset
   288
  assumes "\<And> i. i \<in> s \<Longrightarrow> a i \<in> sets M"
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33657
diff changeset
   289
  assumes d: "disjoint_family_on a s"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   290
  shows "(\<Sum> i \<in> s. \<mu> (a i)) = \<mu> (\<Union> i \<in> s. a i)"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33657
diff changeset
   291
using assms
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33657
diff changeset
   292
proof -
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33657
diff changeset
   293
  (* counting the elements in s is enough *)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   294
  let ?R = "{..< card s}"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33657
diff changeset
   295
  obtain f where f: "f ` ?R = s" "inj_on f ?R"
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33657
diff changeset
   296
    using ex_bij_betw_nat_finite[unfolded bij_betw_def, OF `finite s`]
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   297
    unfolding atLeast0LessThan by auto
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33657
diff changeset
   298
  hence f_into_sets: "a \<circ> f \<in> ?R \<rightarrow> sets M" using assms unfolding o_def by auto
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33657
diff changeset
   299
  have disj: "\<And> i j. \<lbrakk>i \<in> ?R ; j \<in> ?R ; i \<noteq> j\<rbrakk> \<Longrightarrow> (a \<circ> f) i \<inter> (a \<circ> f) j = {}"
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33657
diff changeset
   300
  proof -
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33657
diff changeset
   301
    fix i j assume asm: "i \<in> ?R" "j \<in> ?R" "i \<noteq> j"
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33657
diff changeset
   302
    hence neq: "f i \<noteq> f j" using f[unfolded inj_on_def, rule_format] by blast
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33657
diff changeset
   303
    from asm have "f i \<in> s" "f j \<in> s" using f by auto
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33657
diff changeset
   304
    thus "(a \<circ> f) i \<inter> (a \<circ> f) j = {}"
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33657
diff changeset
   305
      using d f neq unfolding disjoint_family_on_def by auto
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33657
diff changeset
   306
  qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33657
diff changeset
   307
  have "(\<Union> i \<in> s. a i) = (\<Union> i \<in> f ` ?R. a i)" using f by auto
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33657
diff changeset
   308
  hence "(\<Union> i \<in> s. a i) = (\<Union> i \<in> ?R. a (f i))" by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   309
  hence "\<mu> (\<Union> i \<in> s. a i) = (\<Sum> i \<in> ?R. \<mu> (a (f i)))"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33657
diff changeset
   310
    using measure_finitely_additive'[OF f_into_sets disj] by simp
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   311
  also have "\<dots> = (\<Sum> i \<in> s. \<mu> (a i))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   312
    using f[rule_format] setsum_reindex[of f ?R "\<lambda> i. \<mu> (a i)"] by auto
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33657
diff changeset
   313
  finally show ?thesis by simp
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33657
diff changeset
   314
qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33657
diff changeset
   315
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   316
lemma (in sigma_algebra) finite_additivity_sufficient:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   317
  assumes fin: "finite (space M)" and pos: "positive \<mu>" and add: "additive M \<mu>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   318
  shows "measure_space M \<mu>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   319
proof
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   320
  show [simp]: "\<mu> {} = 0" using pos by (simp add: positive_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   321
  show "countably_additive M \<mu>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   322
    proof (auto simp add: countably_additive_def)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   323
      fix A :: "nat \<Rightarrow> 'a set"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   324
      assume A: "range A \<subseteq> sets M"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   325
         and disj: "disjoint_family A"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   326
         and UnA: "(\<Union>i. A i) \<in> sets M"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   327
      def I \<equiv> "{i. A i \<noteq> {}}"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   328
      have "Union (A ` I) \<subseteq> space M" using A
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   329
        by auto (metis range_subsetD subsetD sets_into_space)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   330
      hence "finite (A ` I)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   331
        by (metis finite_UnionD finite_subset fin)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   332
      moreover have "inj_on A I" using disj
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   333
        by (auto simp add: I_def disjoint_family_on_def inj_on_def)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   334
      ultimately have finI: "finite I"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   335
        by (metis finite_imageD)
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   336
      hence "\<exists>N. \<forall>m\<ge>N. A m = {}"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   337
        proof (cases "I = {}")
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   338
          case True thus ?thesis by (simp add: I_def)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   339
        next
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   340
          case False
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   341
          hence "\<forall>i\<in>I. i < Suc(Max I)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   342
            by (simp add: Max_less_iff [symmetric] finI)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   343
          hence "\<forall>m \<ge> Suc(Max I). A m = {}"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   344
            by (simp add: I_def) (metis less_le_not_le)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   345
          thus ?thesis
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   346
            by blast
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   347
        qed
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   348
      then obtain N where N: "\<forall>m\<ge>N. A m = {}" by blast
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   349
      then have "\<forall>m\<ge>N. \<mu> (A m) = 0" by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   350
      then have "(\<Sum>\<^isub>\<infinity> n. \<mu> (A n)) = setsum (\<lambda>m. \<mu> (A m)) {..<N}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   351
        by (simp add: psuminf_finite)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   352
      also have "... = \<mu> (\<Union>i<N. A i)"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   353
        proof (induct N)
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   354
          case 0 thus ?case by simp
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   355
        next
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   356
          case (Suc n)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   357
          have "\<mu> (A n \<union> (\<Union> x<n. A x)) = \<mu> (A n) + \<mu> (\<Union> i<n. A i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   358
            proof (rule Caratheodory.additiveD [OF add])
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   359
              show "A n \<inter> (\<Union> x<n. A x) = {}" using disj
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33657
diff changeset
   360
                by (auto simp add: disjoint_family_on_def nat_less_le) blast
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   361
              show "A n \<in> sets M" using A
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   362
                by force
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   363
              show "(\<Union>i<n. A i) \<in> sets M"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   364
                proof (induct n)
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   365
                  case 0 thus ?case by simp
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   366
                next
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   367
                  case (Suc n) thus ?case using A
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   368
                    by (simp add: lessThan_Suc Un range_subsetD)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   369
                qed
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   370
            qed
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   371
          thus ?case using Suc
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   372
            by (simp add: lessThan_Suc)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   373
        qed
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   374
      also have "... = \<mu> (\<Union>i. A i)"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   375
        proof -
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   376
          have "(\<Union> i<N. A i) = (\<Union>i. A i)" using N
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   377
            by auto (metis Int_absorb N disjoint_iff_not_equal lessThan_iff not_leE)
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   378
          thus ?thesis by simp
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   379
        qed
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   380
      finally show "(\<Sum>\<^isub>\<infinity> n. \<mu> (A n)) = \<mu> (\<Union>i. A i)" .
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   381
    qed
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   382
qed
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   383
35692
f1315bbf1bc9 Moved theorems in Lebesgue to the right places
hoelzl
parents: 35582
diff changeset
   384
lemma (in measure_space) measure_setsum_split:
f1315bbf1bc9 Moved theorems in Lebesgue to the right places
hoelzl
parents: 35582
diff changeset
   385
  assumes "finite r" and "a \<in> sets M" and br_in_M: "b ` r \<subseteq> sets M"
f1315bbf1bc9 Moved theorems in Lebesgue to the right places
hoelzl
parents: 35582
diff changeset
   386
  assumes "(\<Union>i \<in> r. b i) = space M"
f1315bbf1bc9 Moved theorems in Lebesgue to the right places
hoelzl
parents: 35582
diff changeset
   387
  assumes "disjoint_family_on b r"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   388
  shows "\<mu> a = (\<Sum> i \<in> r. \<mu> (a \<inter> (b i)))"
35692
f1315bbf1bc9 Moved theorems in Lebesgue to the right places
hoelzl
parents: 35582
diff changeset
   389
proof -
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   390
  have *: "\<mu> a = \<mu> (\<Union>i \<in> r. a \<inter> b i)"
35692
f1315bbf1bc9 Moved theorems in Lebesgue to the right places
hoelzl
parents: 35582
diff changeset
   391
    using assms by auto
f1315bbf1bc9 Moved theorems in Lebesgue to the right places
hoelzl
parents: 35582
diff changeset
   392
  show ?thesis unfolding *
f1315bbf1bc9 Moved theorems in Lebesgue to the right places
hoelzl
parents: 35582
diff changeset
   393
  proof (rule measure_finitely_additive''[symmetric])
f1315bbf1bc9 Moved theorems in Lebesgue to the right places
hoelzl
parents: 35582
diff changeset
   394
    show "finite r" using `finite r` by auto
f1315bbf1bc9 Moved theorems in Lebesgue to the right places
hoelzl
parents: 35582
diff changeset
   395
    { fix i assume "i \<in> r"
f1315bbf1bc9 Moved theorems in Lebesgue to the right places
hoelzl
parents: 35582
diff changeset
   396
      hence "b i \<in> sets M" using br_in_M by auto
f1315bbf1bc9 Moved theorems in Lebesgue to the right places
hoelzl
parents: 35582
diff changeset
   397
      thus "a \<inter> b i \<in> sets M" using `a \<in> sets M` by auto
f1315bbf1bc9 Moved theorems in Lebesgue to the right places
hoelzl
parents: 35582
diff changeset
   398
    }
f1315bbf1bc9 Moved theorems in Lebesgue to the right places
hoelzl
parents: 35582
diff changeset
   399
    show "disjoint_family_on (\<lambda>i. a \<inter> b i) r"
f1315bbf1bc9 Moved theorems in Lebesgue to the right places
hoelzl
parents: 35582
diff changeset
   400
      using `disjoint_family_on b r`
f1315bbf1bc9 Moved theorems in Lebesgue to the right places
hoelzl
parents: 35582
diff changeset
   401
      unfolding disjoint_family_on_def by auto
f1315bbf1bc9 Moved theorems in Lebesgue to the right places
hoelzl
parents: 35582
diff changeset
   402
  qed
f1315bbf1bc9 Moved theorems in Lebesgue to the right places
hoelzl
parents: 35582
diff changeset
   403
qed
f1315bbf1bc9 Moved theorems in Lebesgue to the right places
hoelzl
parents: 35582
diff changeset
   404
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   405
lemma (in measure_space) measure_subadditive:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   406
  assumes measurable: "A \<in> sets M" "B \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   407
  shows "\<mu> (A \<union> B) \<le> \<mu> A + \<mu> B"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   408
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   409
  from measure_additive[of A "B - A"]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   410
  have "\<mu> (A \<union> B) = \<mu> A + \<mu> (B - A)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   411
    using assms by (simp add: Diff)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   412
  also have "\<dots> \<le> \<mu> A + \<mu> B"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   413
    using assms by (auto intro!: add_left_mono measure_mono)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   414
  finally show ?thesis .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   415
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   416
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   417
lemma (in measure_space) measurable_countably_subadditive:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   418
  assumes "range f \<subseteq> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   419
  shows "\<mu> (\<Union>i. f i) \<le> (\<Sum>\<^isub>\<infinity> i. \<mu> (f i))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   420
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   421
  have "\<mu> (\<Union>i. f i) = \<mu> (\<Union>i. disjointed f i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   422
    unfolding UN_disjointed_eq ..
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   423
  also have "\<dots> = (\<Sum>\<^isub>\<infinity> i. \<mu> (disjointed f i))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   424
    using range_disjointed_sets[OF assms] measure_countably_additive
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   425
    by (simp add:  disjoint_family_disjointed comp_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   426
  also have "\<dots> \<le> (\<Sum>\<^isub>\<infinity> i. \<mu> (f i))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   427
  proof (rule psuminf_le, rule measure_mono)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   428
    fix i show "disjointed f i \<subseteq> f i" by (rule disjointed_subset)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   429
    show "f i \<in> sets M" "disjointed f i \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   430
      using assms range_disjointed_sets[OF assms] by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   431
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   432
  finally show ?thesis .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   433
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   434
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   435
lemma (in measure_space) restricted_measure_space:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   436
  assumes "S \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   437
  shows "measure_space (M\<lparr> space := S, sets := (\<lambda>A. S \<inter> A) ` sets M \<rparr>) \<mu>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   438
    (is "measure_space ?r \<mu>")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   439
  unfolding measure_space_def measure_space_axioms_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   440
proof safe
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   441
  show "sigma_algebra ?r" using restricted_sigma_algebra[OF assms] .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   442
  show "\<mu> {} = 0" by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   443
  show "countably_additive ?r \<mu>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   444
    unfolding countably_additive_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   445
  proof safe
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   446
    fix A :: "nat \<Rightarrow> 'a set"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   447
    assume *: "range A \<subseteq> sets ?r" and **: "disjoint_family A"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   448
    from restriction_in_sets[OF assms *[simplified]] **
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   449
    show "(\<Sum>\<^isub>\<infinity> n. \<mu> (A n)) = \<mu> (\<Union>i. A i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   450
      using measure_countably_additive by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   451
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   452
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   453
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   454
section "@{text \<sigma>}-finite Measures"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   455
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   456
locale sigma_finite_measure = measure_space +
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   457
  assumes sigma_finite: "\<exists>A::nat \<Rightarrow> 'a set. range A \<subseteq> sets M \<and> (\<Union>i. A i) = space M \<and> (\<forall>i. \<mu> (A i) \<noteq> \<omega>)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   458
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   459
lemma (in sigma_finite_measure) restricted_sigma_finite_measure:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   460
  assumes "S \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   461
  shows "sigma_finite_measure (M\<lparr> space := S, sets := (\<lambda>A. S \<inter> A) ` sets M \<rparr>) \<mu>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   462
    (is "sigma_finite_measure ?r _")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   463
  unfolding sigma_finite_measure_def sigma_finite_measure_axioms_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   464
proof safe
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   465
  show "measure_space ?r \<mu>" using restricted_measure_space[OF assms] .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   466
next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   467
  obtain A :: "nat \<Rightarrow> 'a set" where
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   468
      "range A \<subseteq> sets M" "(\<Union>i. A i) = space M" "\<And>i. \<mu> (A i) \<noteq> \<omega>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   469
    using sigma_finite by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   470
  show "\<exists>A::nat \<Rightarrow> 'a set. range A \<subseteq> sets ?r \<and> (\<Union>i. A i) = space ?r \<and> (\<forall>i. \<mu> (A i) \<noteq> \<omega>)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   471
  proof (safe intro!: exI[of _ "\<lambda>i. A i \<inter> S"] del: notI)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   472
    fix i
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   473
    show "A i \<inter> S \<in> sets ?r"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   474
      using `range A \<subseteq> sets M` `S \<in> sets M` by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   475
  next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   476
    fix x i assume "x \<in> S" thus "x \<in> space ?r" by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   477
  next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   478
    fix x assume "x \<in> space ?r" thus "x \<in> (\<Union>i. A i \<inter> S)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   479
      using `(\<Union>i. A i) = space M` `S \<in> sets M` by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   480
  next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   481
    fix i
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   482
    have "\<mu> (A i \<inter> S) \<le> \<mu> (A i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   483
      using `range A \<subseteq> sets M` `S \<in> sets M` by (auto intro!: measure_mono)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   484
    also have "\<dots> < \<omega>" using `\<mu> (A i) \<noteq> \<omega>` by (auto simp: pinfreal_less_\<omega>)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   485
    finally show "\<mu> (A i \<inter> S) \<noteq> \<omega>" by (auto simp: pinfreal_less_\<omega>)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   486
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   487
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   488
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   489
section "Real measure values"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   490
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   491
lemma (in measure_space) real_measure_Union:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   492
  assumes finite: "\<mu> A \<noteq> \<omega>" "\<mu> B \<noteq> \<omega>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   493
  and measurable: "A \<in> sets M" "B \<in> sets M" "A \<inter> B = {}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   494
  shows "real (\<mu> (A \<union> B)) = real (\<mu> A) + real (\<mu> B)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   495
  unfolding measure_additive[symmetric, OF measurable]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   496
  using finite by (auto simp: real_of_pinfreal_add)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   497
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   498
lemma (in measure_space) real_measure_finite_Union:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   499
  assumes measurable:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   500
    "finite S" "\<And>i. i \<in> S \<Longrightarrow> A i \<in> sets M" "disjoint_family_on A S"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   501
  assumes finite: "\<And>i. i \<in> S \<Longrightarrow> \<mu> (A i) \<noteq> \<omega>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   502
  shows "real (\<mu> (\<Union>i\<in>S. A i)) = (\<Sum>i\<in>S. real (\<mu> (A i)))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   503
  using real_of_pinfreal_setsum[of S, OF finite]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   504
        measure_finitely_additive''[symmetric, OF measurable]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   505
  by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   506
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   507
lemma (in measure_space) real_measure_Diff:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   508
  assumes finite: "\<mu> A \<noteq> \<omega>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   509
  and measurable: "A \<in> sets M" "B \<in> sets M" "B \<subseteq> A"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   510
  shows "real (\<mu> (A - B)) = real (\<mu> A) - real (\<mu> B)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   511
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   512
  have "\<mu> (A - B) \<le> \<mu> A"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   513
    "\<mu> B \<le> \<mu> A"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   514
    using measurable by (auto intro!: measure_mono)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   515
  hence "real (\<mu> ((A - B) \<union> B)) = real (\<mu> (A - B)) + real (\<mu> B)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   516
    using measurable finite by (rule_tac real_measure_Union) auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   517
  thus ?thesis using `B \<subseteq> A` by (auto simp: Un_absorb2)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   518
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   519
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   520
lemma (in measure_space) real_measure_UNION:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   521
  assumes measurable: "range A \<subseteq> sets M" "disjoint_family A"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   522
  assumes finite: "\<mu> (\<Union>i. A i) \<noteq> \<omega>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   523
  shows "(\<lambda>i. real (\<mu> (A i))) sums (real (\<mu> (\<Union>i. A i)))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   524
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   525
  have *: "(\<Sum>\<^isub>\<infinity> i. \<mu> (A i)) = \<mu> (\<Union>i. A i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   526
    using measure_countably_additive[OF measurable] by (simp add: comp_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   527
  hence "(\<Sum>\<^isub>\<infinity> i. \<mu> (A i)) \<noteq> \<omega>" using finite by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   528
  from psuminf_imp_suminf[OF this]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   529
  show ?thesis using * by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   530
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   531
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   532
lemma (in measure_space) real_measure_subadditive:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   533
  assumes measurable: "A \<in> sets M" "B \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   534
  and fin: "\<mu> A \<noteq> \<omega>" "\<mu> B \<noteq> \<omega>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   535
  shows "real (\<mu> (A \<union> B)) \<le> real (\<mu> A) + real (\<mu> B)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   536
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   537
  have "real (\<mu> (A \<union> B)) \<le> real (\<mu> A + \<mu> B)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   538
    using measure_subadditive[OF measurable] fin by (auto intro!: real_of_pinfreal_mono)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   539
  also have "\<dots> = real (\<mu> A) + real (\<mu> B)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   540
    using fin by (simp add: real_of_pinfreal_add)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   541
  finally show ?thesis .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   542
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   543
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   544
lemma (in measure_space) real_measurable_countably_subadditive:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   545
  assumes "range f \<subseteq> sets M" and "(\<Sum>\<^isub>\<infinity> i. \<mu> (f i)) \<noteq> \<omega>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   546
  shows "real (\<mu> (\<Union>i. f i)) \<le> (\<Sum> i. real (\<mu> (f i)))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   547
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   548
  have "real (\<mu> (\<Union>i. f i)) \<le> real (\<Sum>\<^isub>\<infinity> i. \<mu> (f i))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   549
    using assms by (auto intro!: real_of_pinfreal_mono measurable_countably_subadditive)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   550
  also have "\<dots> = (\<Sum> i. real (\<mu> (f i)))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   551
    using assms by (auto intro!: sums_unique psuminf_imp_suminf)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   552
  finally show ?thesis .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   553
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   554
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   555
lemma (in measure_space) real_measure_setsum_singleton:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   556
  assumes S: "finite S" "\<And>x. x \<in> S \<Longrightarrow> {x} \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   557
  and fin: "\<And>x. x \<in> S \<Longrightarrow> \<mu> {x} \<noteq> \<omega>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   558
  shows "real (\<mu> S) = (\<Sum>x\<in>S. real (\<mu> {x}))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   559
  using measure_finite_singleton[OF S] fin by (simp add: real_of_pinfreal_setsum)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   560
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   561
lemma (in measure_space) real_continuity_from_below:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   562
  assumes A: "range A \<subseteq> sets M" "\<And>i. A i \<subseteq> A (Suc i)" and "\<mu> (\<Union>i. A i) \<noteq> \<omega>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   563
  shows "(\<lambda>i. real (\<mu> (A i))) ----> real (\<mu> (\<Union>i. A i))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   564
proof (rule SUP_eq_LIMSEQ[THEN iffD1])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   565
  { fix n have "\<mu> (A n) \<le> \<mu> (\<Union>i. A i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   566
      using A by (auto intro!: measure_mono)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   567
    hence "\<mu> (A n) \<noteq> \<omega>" using assms by auto }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   568
  note this[simp]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   569
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   570
  show "mono (\<lambda>i. real (\<mu> (A i)))" unfolding mono_iff_le_Suc using A
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   571
    by (auto intro!: real_of_pinfreal_mono measure_mono)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   572
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   573
  show "(SUP n. Real (real (\<mu> (A n)))) = Real (real (\<mu> (\<Union>i. A i)))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   574
    using continuity_from_below[OF A(1), OF A(2)]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   575
    using assms by (simp add: Real_real)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   576
qed simp_all
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   577
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   578
lemma (in measure_space) real_continuity_from_above:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   579
  assumes A: "range A \<subseteq> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   580
  and mono_Suc: "\<And>n.  A (Suc n) \<subseteq> A n"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   581
  and finite: "\<And>i. \<mu> (A i) \<noteq> \<omega>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   582
  shows "(\<lambda>n. real (\<mu> (A n))) ----> real (\<mu> (\<Inter>i. A i))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   583
proof (rule INF_eq_LIMSEQ[THEN iffD1])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   584
  { fix n have "\<mu> (\<Inter>i. A i) \<le> \<mu> (A n)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   585
      using A by (auto intro!: measure_mono)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   586
    hence "\<mu> (\<Inter>i. A i) \<noteq> \<omega>" using assms by auto }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   587
  note this[simp]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   588
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   589
  show "mono (\<lambda>i. - real (\<mu> (A i)))" unfolding mono_iff_le_Suc using assms
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   590
    by (auto intro!: real_of_pinfreal_mono measure_mono)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   591
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   592
  show "(INF n. Real (real (\<mu> (A n)))) =
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   593
    Real (real (\<mu> (\<Inter>i. A i)))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   594
    using continuity_from_above[OF A, OF mono_Suc finite]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   595
    using assms by (simp add: Real_real)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   596
qed simp_all
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   597
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   598
locale finite_measure = measure_space +
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   599
  assumes finite_measure_of_space: "\<mu> (space M) \<noteq> \<omega>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   600
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   601
sublocale finite_measure < sigma_finite_measure
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   602
proof
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   603
  show "\<exists>A. range A \<subseteq> sets M \<and> (\<Union>i. A i) = space M \<and> (\<forall>i. \<mu> (A i) \<noteq> \<omega>)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   604
    using finite_measure_of_space by (auto intro!: exI[of _ "\<lambda>x. space M"])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   605
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   606
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   607
lemma (in finite_measure) finite_measure:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   608
  assumes "A \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   609
  shows "\<mu> A \<noteq> \<omega>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   610
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   611
  from `A \<in> sets M` have "A \<subseteq> space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   612
    using sets_into_space by blast
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   613
  hence "\<mu> A \<le> \<mu> (space M)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   614
    using assms top by (rule measure_mono)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   615
  also have "\<dots> < \<omega>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   616
    using finite_measure_of_space unfolding pinfreal_less_\<omega> .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   617
  finally show ?thesis unfolding pinfreal_less_\<omega> .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   618
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   619
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   620
lemma (in finite_measure) restricted_finite_measure:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   621
  assumes "S \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   622
  shows "finite_measure (M\<lparr> space := S, sets := (\<lambda>A. S \<inter> A) ` sets M \<rparr>) \<mu>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   623
    (is "finite_measure ?r _")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   624
  unfolding finite_measure_def finite_measure_axioms_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   625
proof (safe del: notI)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   626
  show "measure_space ?r \<mu>" using restricted_measure_space[OF assms] .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   627
next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   628
  show "\<mu> (space ?r) \<noteq> \<omega>" using finite_measure[OF `S \<in> sets M`] by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   629
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   630
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   631
lemma (in finite_measure) real_finite_measure_Diff:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   632
  assumes "A \<in> sets M" "B \<in> sets M" "B \<subseteq> A"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   633
  shows "real (\<mu> (A - B)) = real (\<mu> A) - real (\<mu> B)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   634
  using finite_measure[OF `A \<in> sets M`] by (rule real_measure_Diff[OF _ assms])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   635
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   636
lemma (in finite_measure) real_finite_measure_Union:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   637
  assumes sets: "A \<in> sets M" "B \<in> sets M" and "A \<inter> B = {}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   638
  shows "real (\<mu> (A \<union> B)) = real (\<mu> A) + real (\<mu> B)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   639
  using sets by (auto intro!: real_measure_Union[OF _ _ assms] finite_measure)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   640
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   641
lemma (in finite_measure) real_finite_measure_finite_Union:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   642
  assumes "finite S" and dis: "disjoint_family_on A S"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   643
  and *: "\<And>i. i \<in> S \<Longrightarrow> A i \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   644
  shows "real (\<mu> (\<Union>i\<in>S. A i)) = (\<Sum>i\<in>S. real (\<mu> (A i)))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   645
proof (rule real_measure_finite_Union[OF `finite S` _ dis])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   646
  fix i assume "i \<in> S" from *[OF this] show "A i \<in> sets M" .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   647
  from finite_measure[OF this] show "\<mu> (A i) \<noteq> \<omega>" .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   648
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   649
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   650
lemma (in finite_measure) real_finite_measure_UNION:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   651
  assumes measurable: "range A \<subseteq> sets M" "disjoint_family A"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   652
  shows "(\<lambda>i. real (\<mu> (A i))) sums (real (\<mu> (\<Union>i. A i)))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   653
proof (rule real_measure_UNION[OF assms])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   654
  have "(\<Union>i. A i) \<in> sets M" using measurable(1) by (rule countable_UN)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   655
  thus "\<mu> (\<Union>i. A i) \<noteq> \<omega>" by (rule finite_measure)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   656
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   657
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   658
lemma (in finite_measure) real_measure_mono:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   659
  "A \<in> sets M \<Longrightarrow> B \<in> sets M \<Longrightarrow> A \<subseteq> B \<Longrightarrow> real (\<mu> A) \<le> real (\<mu> B)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   660
  by (auto intro!: measure_mono real_of_pinfreal_mono finite_measure)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   661
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   662
lemma (in finite_measure) real_finite_measure_subadditive:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   663
  assumes measurable: "A \<in> sets M" "B \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   664
  shows "real (\<mu> (A \<union> B)) \<le> real (\<mu> A) + real (\<mu> B)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   665
  using measurable measurable[THEN finite_measure] by (rule real_measure_subadditive)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   666
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   667
lemma (in finite_measure) real_finite_measurable_countably_subadditive:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   668
  assumes "range f \<subseteq> sets M" and "summable (\<lambda>i. real (\<mu> (f i)))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   669
  shows "real (\<mu> (\<Union>i. f i)) \<le> (\<Sum> i. real (\<mu> (f i)))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   670
proof (rule real_measurable_countably_subadditive[OF assms(1)])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   671
  have *: "\<And>i. f i \<in> sets M" using assms by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   672
  have "(\<lambda>i. real (\<mu> (f i))) sums (\<Sum>i. real (\<mu> (f i)))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   673
    using assms(2) by (rule summable_sums)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   674
  from suminf_imp_psuminf[OF this]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   675
  have "(\<Sum>\<^isub>\<infinity>i. \<mu> (f i)) = Real (\<Sum>i. real (\<mu> (f i)))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   676
    using * by (simp add: Real_real finite_measure)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   677
  thus "(\<Sum>\<^isub>\<infinity>i. \<mu> (f i)) \<noteq> \<omega>" by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   678
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   679
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   680
lemma (in finite_measure) real_finite_measure_finite_singelton:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   681
  assumes "finite S" and *: "\<And>x. x \<in> S \<Longrightarrow> {x} \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   682
  shows "real (\<mu> S) = (\<Sum>x\<in>S. real (\<mu> {x}))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   683
proof (rule real_measure_setsum_singleton[OF `finite S`])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   684
  fix x assume "x \<in> S" thus "{x} \<in> sets M" by (rule *)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   685
  with finite_measure show "\<mu> {x} \<noteq> \<omega>" .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   686
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   687
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   688
lemma (in finite_measure) real_finite_continuity_from_below:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   689
  assumes "range A \<subseteq> sets M" "\<And>i. A i \<subseteq> A (Suc i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   690
  shows "(\<lambda>i. real (\<mu> (A i))) ----> real (\<mu> (\<Union>i. A i))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   691
  using real_continuity_from_below[OF assms(1), OF assms(2) finite_measure] assms by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   692
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   693
lemma (in finite_measure) real_finite_continuity_from_above:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   694
  assumes A: "range A \<subseteq> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   695
  and mono_Suc: "\<And>n.  A (Suc n) \<subseteq> A n"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   696
  shows "(\<lambda>n. real (\<mu> (A n))) ----> real (\<mu> (\<Inter>i. A i))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   697
  using real_continuity_from_above[OF A, OF mono_Suc finite_measure] A
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   698
  by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   699
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   700
lemma (in finite_measure) real_finite_measure_finite_Union':
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   701
  assumes "finite S" "A`S \<subseteq> sets M" "disjoint_family_on A S"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   702
  shows "real (\<mu> (\<Union>i\<in>S. A i)) = (\<Sum>i\<in>S. real (\<mu> (A i)))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   703
  using assms real_finite_measure_finite_Union[of S A] by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   704
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   705
lemma (in finite_measure) finite_measure_compl:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   706
  assumes s: "s \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   707
  shows "\<mu> (space M - s) = \<mu> (space M) - \<mu> s"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   708
  using measure_compl[OF s, OF finite_measure, OF s] .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   709
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   710
section {* Measure preserving *}
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   711
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   712
definition "measure_preserving A \<mu> B \<nu> =
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   713
    {f \<in> measurable A B. (\<forall>y \<in> sets B. \<mu> (f -` y \<inter> space A) = \<nu> y)}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   714
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   715
lemma (in finite_measure) measure_preserving_lift:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   716
  fixes f :: "'a \<Rightarrow> 'a2" and A :: "('a2, 'b2) algebra_scheme"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   717
  assumes "algebra A"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   718
  assumes "finite_measure (sigma (space A) (sets A)) \<nu>" (is "finite_measure ?sA _")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   719
  assumes fmp: "f \<in> measure_preserving M \<mu> A \<nu>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   720
  shows "f \<in> measure_preserving M \<mu> (sigma (space A) (sets A)) \<nu>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   721
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   722
  interpret sA: finite_measure ?sA \<nu> by fact
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   723
  interpret A: algebra A by fact
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   724
  show ?thesis using fmp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   725
    proof (clarsimp simp add: measure_preserving_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   726
      assume fm: "f \<in> measurable M A"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   727
         and meq: "\<forall>y\<in>sets A. \<mu> (f -` y \<inter> space M) = \<nu> y"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   728
      have f12: "f \<in> measurable M ?sA"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   729
        using measurable_subset[OF A.sets_into_space] fm by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   730
      hence ffn: "f \<in> space M \<rightarrow> space A"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   731
        by (simp add: measurable_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   732
      have "space M \<subseteq> f -` (space A)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   733
        by auto (metis PiE ffn)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   734
      hence fveq [simp]: "(f -` (space A)) \<inter> space M = space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   735
        by blast
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   736
      {
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   737
        fix y
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   738
        assume y: "y \<in> sets ?sA"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   739
        have "sets ?sA = sigma_sets (space A) (sets A)" (is "_ = ?A") by (auto simp: sigma_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   740
        also have "\<dots> \<subseteq> {s . \<mu> ((f -` s) \<inter> space M) = \<nu> s}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   741
          proof (rule A.sigma_property_disjoint, auto)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   742
            fix x assume "x \<in> sets A" then show "\<mu> (f -` x \<inter> space M) = \<nu> x" by (simp add: meq)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   743
          next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   744
            fix s
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   745
            assume eq: "\<mu> (f -` s \<inter> space M) = \<nu> s" and s: "s \<in> ?A"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   746
            then have s': "s \<in> sets ?sA" by (simp add: sigma_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   747
            show "\<mu> (f -` (space A - s) \<inter> space M) = \<nu> (space A - s)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   748
              using sA.finite_measure_compl[OF s']
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   749
              using measurable_sets[OF f12 s'] meq[THEN bspec, OF A.top]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   750
              by (simp add: vimage_Diff Diff_Int_distrib2 finite_measure_compl eq)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   751
          next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   752
            fix S
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   753
            assume S0: "S 0 = {}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   754
               and SSuc: "\<And>n.  S n \<subseteq> S (Suc n)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   755
               and rS1: "range S \<subseteq> {s. \<mu> (f -` s \<inter> space M) = \<nu> s}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   756
               and "range S \<subseteq> ?A"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   757
            hence rS2: "range S \<subseteq> sets ?sA" by (simp add: sigma_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   758
            have eq1: "\<And>i. \<mu> (f -` S i \<inter> space M) = \<nu> (S i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   759
              using rS1 by blast
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   760
            have *: "(\<lambda>n. \<nu> (S n)) = (\<lambda>n. \<mu> (f -` S n \<inter> space M))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   761
              by (simp add: eq1)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   762
            have "(SUP n. ... n) = \<mu> (\<Union>i. f -` S i \<inter> space M)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   763
              proof (rule measure_countable_increasing)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   764
                show "range (\<lambda>i. f -` S i \<inter> space M) \<subseteq> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   765
                  using f12 rS2 by (auto simp add: measurable_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   766
                show "f -` S 0 \<inter> space M = {}" using S0
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   767
                  by blast
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   768
                show "\<And>n. f -` S n \<inter> space M \<subseteq> f -` S (Suc n) \<inter> space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   769
                  using SSuc by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   770
              qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   771
            also have "\<mu> (\<Union>i. f -` S i \<inter> space M) = \<mu> (f -` (\<Union>i. S i) \<inter> space M)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   772
              by (simp add: vimage_UN)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   773
            finally have "(SUP n. \<nu> (S n)) = \<mu> (f -` (\<Union>i. S i) \<inter> space M)" unfolding * .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   774
            moreover
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   775
            have "(SUP n. \<nu> (S n)) = \<nu> (\<Union>i. S i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   776
              by (rule sA.measure_countable_increasing[OF rS2, OF S0 SSuc])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   777
            ultimately
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   778
            show "\<mu> (f -` (\<Union>i. S i) \<inter> space M) = \<nu> (\<Union>i. S i)" by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   779
          next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   780
            fix S :: "nat => 'a2 set"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   781
              assume dS: "disjoint_family S"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   782
                 and rS1: "range S \<subseteq> {s. \<mu> (f -` s \<inter> space M) = \<nu> s}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   783
                 and "range S \<subseteq> ?A"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   784
              hence rS2: "range S \<subseteq> sets ?sA" by (simp add: sigma_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   785
              have "\<And>i. \<mu> (f -` S i \<inter> space M) = \<nu> (S i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   786
                using rS1 by blast
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   787
              hence *: "(\<lambda>i. \<nu> (S i)) = (\<lambda>n. \<mu> (f -` S n \<inter> space M))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   788
                by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   789
              have "psuminf ... = \<mu> (\<Union>i. f -` S i \<inter> space M)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   790
                proof (rule measure_countably_additive)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   791
                  show "range (\<lambda>i. f -` S i \<inter> space M) \<subseteq> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   792
                    using f12 rS2 by (auto simp add: measurable_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   793
                  show "disjoint_family (\<lambda>i. f -` S i \<inter> space M)" using dS
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   794
                    by (auto simp add: disjoint_family_on_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   795
                qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   796
              hence "(\<Sum>\<^isub>\<infinity> i. \<nu> (S i)) = \<mu> (\<Union>i. f -` S i \<inter> space M)" unfolding * .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   797
              with sA.measure_countably_additive [OF rS2 dS]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   798
              have "\<mu> (\<Union>i. f -` S i \<inter> space M) = \<nu> (\<Union>i. S i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   799
                by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   800
              moreover have "\<mu> (f -` (\<Union>i. S i) \<inter> space M) = \<mu> (\<Union>i. f -` S i \<inter> space M)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   801
                by (simp add: vimage_UN)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   802
              ultimately show "\<mu> (f -` (\<Union>i. S i) \<inter> space M) = \<nu> (\<Union>i. S i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   803
                by metis
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   804
          qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   805
        finally have "sets ?sA \<subseteq> {s . \<mu> ((f -` s) \<inter> space M) = \<nu> s}" .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   806
        hence "\<mu> (f -` y \<inter> space M) = \<nu> y" using y by force
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   807
      }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   808
      thus "f \<in> measurable M ?sA \<and> (\<forall>y\<in>sets ?sA. \<mu> (f -` y \<inter> space M) = \<nu> y)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   809
        by (blast intro: f12)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   810
    qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   811
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   812
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   813
section "Finite spaces"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   814
36624
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   815
locale finite_measure_space = measure_space +
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   816
  assumes finite_space: "finite (space M)"
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   817
  and sets_eq_Pow: "sets M = Pow (space M)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   818
  and finite_single_measure: "\<And>x. x \<in> space M \<Longrightarrow> \<mu> {x} \<noteq> \<omega>"
36624
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   819
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   820
lemma (in finite_measure_space) sum_over_space: "(\<Sum>x\<in>space M. \<mu> {x}) = \<mu> (space M)"
36624
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   821
  using measure_finitely_additive''[of "space M" "\<lambda>i. {i}"]
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   822
  by (simp add: sets_eq_Pow disjoint_family_on_def finite_space)
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   823
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   824
sublocale finite_measure_space < finite_measure
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   825
proof
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   826
  show "\<mu> (space M) \<noteq> \<omega>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   827
    unfolding sum_over_space[symmetric] setsum_\<omega>
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   828
    using finite_space finite_single_measure by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   829
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   830
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   831
lemma psuminf_cmult_indicator:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   832
  assumes "disjoint_family A" "x \<in> A i"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   833
  shows "(\<Sum>\<^isub>\<infinity> n. f n * indicator (A n) x) = f i"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   834
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   835
  have **: "\<And>n. f n * indicator (A n) x = (if n = i then f n else 0 :: pinfreal)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   836
    using `x \<in> A i` assms unfolding disjoint_family_on_def indicator_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   837
  then have "\<And>n. (\<Sum>j<n. f j * indicator (A j) x) = (if i < n then f i else 0 :: pinfreal)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   838
    by (auto simp: setsum_cases)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   839
  moreover have "(SUP n. if i < n then f i else 0) = (f i :: pinfreal)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   840
  proof (rule pinfreal_SUPI)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   841
    fix y :: pinfreal assume "\<And>n. n \<in> UNIV \<Longrightarrow> (if i < n then f i else 0) \<le> y"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   842
    from this[of "Suc i"] show "f i \<le> y" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   843
  qed simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   844
  ultimately show ?thesis using `x \<in> A i` unfolding psuminf_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   845
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   846
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   847
lemma psuminf_indicator:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   848
  assumes "disjoint_family A"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   849
  shows "(\<Sum>\<^isub>\<infinity> n. indicator (A n) x) = indicator (\<Union>i. A i) x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   850
proof cases
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   851
  assume *: "x \<in> (\<Union>i. A i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   852
  then obtain i where "x \<in> A i" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   853
  from psuminf_cmult_indicator[OF assms, OF this, of "\<lambda>i. 1"]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   854
  show ?thesis using * by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   855
qed simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37032
diff changeset
   856
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33657
diff changeset
   857
end