src/HOL/Probability/Caratheodory.thy
author hoelzl
Thu, 02 Sep 2010 17:28:00 +0200
changeset 39096 111756225292
parent 38656 d5d342611edb
child 40859 de0b30e6c2d2
permissions -rw-r--r--
merged
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 {*Caratheodory Extension Theorem*}
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 Caratheodory
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
     4
  imports Sigma_Algebra Positive_Infinite_Real
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
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
     9
subsection {* Measure Spaces *}
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    10
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
    11
definition "positive f \<longleftrightarrow> f {} = (0::pinfreal)" -- "Positive is enforced by the type"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    12
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    13
definition
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    14
  additive  where
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
    15
  "additive M f \<longleftrightarrow>
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
    16
    (\<forall>x \<in> sets M. \<forall>y \<in> sets M. x \<inter> y = {}
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    17
    \<longrightarrow> f (x \<union> y) = f x + f y)"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    18
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    19
definition
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    20
  countably_additive  where
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
    21
  "countably_additive M f \<longleftrightarrow>
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
    22
    (\<forall>A. range A \<subseteq> sets M \<longrightarrow>
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    23
         disjoint_family A \<longrightarrow>
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
    24
         (\<Union>i. A i) \<in> sets M \<longrightarrow>
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
    25
         (\<Sum>\<^isub>\<infinity> n. f (A n)) = f (\<Union>i. A i))"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    26
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    27
definition
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    28
  increasing  where
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    29
  "increasing M f \<longleftrightarrow> (\<forall>x \<in> sets M. \<forall>y \<in> sets M. x \<subseteq> y \<longrightarrow> f x \<le> f y)"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    30
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    31
definition
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    32
  subadditive  where
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
    33
  "subadditive M f \<longleftrightarrow>
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
    34
    (\<forall>x \<in> sets M. \<forall>y \<in> sets M. x \<inter> y = {}
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    35
    \<longrightarrow> f (x \<union> y) \<le> f x + f y)"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    36
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    37
definition
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    38
  countably_subadditive  where
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
    39
  "countably_subadditive M f \<longleftrightarrow>
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
    40
    (\<forall>A. range A \<subseteq> sets M \<longrightarrow>
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    41
         disjoint_family A \<longrightarrow>
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
    42
         (\<Union>i. A i) \<in> sets M \<longrightarrow>
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
    43
         f (\<Union>i. A i) \<le> psuminf (\<lambda>n. f (A n)))"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    44
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    45
definition
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    46
  lambda_system where
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
    47
  "lambda_system M f =
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    48
    {l. l \<in> sets M & (\<forall>x \<in> sets M. f (l \<inter> x) + f ((space M - l) \<inter> x) = f x)}"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    49
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    50
definition
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    51
  outer_measure_space where
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
    52
  "outer_measure_space M f  \<longleftrightarrow>
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
    53
     positive f \<and> increasing M f \<and> countably_subadditive M f"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    54
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    55
definition
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    56
  measure_set where
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    57
  "measure_set M f X =
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
    58
     {r . \<exists>A. range A \<subseteq> sets M \<and> disjoint_family A \<and> X \<subseteq> (\<Union>i. A i) \<and> (\<Sum>\<^isub>\<infinity> i. f (A i)) = r}"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    59
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    60
locale measure_space = sigma_algebra +
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
    61
  fixes \<mu> :: "'a set \<Rightarrow> pinfreal"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
    62
  assumes empty_measure [simp]: "\<mu> {} = 0"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
    63
      and ca: "countably_additive M \<mu>"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    64
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    65
lemma increasingD:
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    66
     "increasing M f \<Longrightarrow> x \<subseteq> y \<Longrightarrow> x\<in>sets M \<Longrightarrow> y\<in>sets M \<Longrightarrow> f x \<le> f y"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    67
  by (auto simp add: increasing_def)
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    68
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    69
lemma subadditiveD:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
    70
     "subadditive M f \<Longrightarrow> x \<inter> y = {} \<Longrightarrow> x\<in>sets M \<Longrightarrow> y\<in>sets M
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    71
      \<Longrightarrow> f (x \<union> y) \<le> f x + f y"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    72
  by (auto simp add: subadditive_def)
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    73
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    74
lemma additiveD:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
    75
     "additive M f \<Longrightarrow> x \<inter> y = {} \<Longrightarrow> x\<in>sets M \<Longrightarrow> y\<in>sets M
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    76
      \<Longrightarrow> f (x \<union> y) = f x + f y"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    77
  by (auto simp add: additive_def)
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    78
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    79
lemma countably_additiveD:
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33536
diff changeset
    80
  "countably_additive M f \<Longrightarrow> range A \<subseteq> sets M \<Longrightarrow> disjoint_family A
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
    81
   \<Longrightarrow> (\<Union>i. A i) \<in> sets M \<Longrightarrow> (\<Sum>\<^isub>\<infinity> n. f (A n)) = f (\<Union>i. A i)"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33536
diff changeset
    82
  by (simp add: countably_additive_def)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    83
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
    84
section "Extend binary sets"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    85
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33536
diff changeset
    86
lemma LIMSEQ_binaryset:
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    87
  assumes f: "f {} = 0"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    88
  shows  "(\<lambda>n. \<Sum>i = 0..<n. f (binaryset A B i)) ----> f A + f B"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    89
proof -
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    90
  have "(\<lambda>n. \<Sum>i = 0..< Suc (Suc n). f (binaryset A B i)) = (\<lambda>n. f A + f B)"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33536
diff changeset
    91
    proof
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    92
      fix n
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    93
      show "(\<Sum>i\<Colon>nat = 0\<Colon>nat..<Suc (Suc n). f (binaryset A B i)) = f A + f B"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33536
diff changeset
    94
        by (induct n)  (auto simp add: binaryset_def f)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    95
    qed
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    96
  moreover
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33536
diff changeset
    97
  have "... ----> f A + f B" by (rule LIMSEQ_const)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    98
  ultimately
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33536
diff changeset
    99
  have "(\<lambda>n. \<Sum>i = 0..< Suc (Suc n). f (binaryset A B i)) ----> f A + f B"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   100
    by metis
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   101
  hence "(\<lambda>n. \<Sum>i = 0..< n+2. f (binaryset A B i)) ----> f A + f B"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   102
    by simp
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   103
  thus ?thesis by (rule LIMSEQ_offset [where k=2])
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   104
qed
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   105
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   106
lemma binaryset_sums:
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   107
  assumes f: "f {} = 0"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   108
  shows  "(\<lambda>n. f (binaryset A B n)) sums (f A + f B)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   109
    by (simp add: sums_def LIMSEQ_binaryset [where f=f, OF f])
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   110
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   111
lemma suminf_binaryset_eq:
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   112
     "f {} = 0 \<Longrightarrow> suminf (\<lambda>n. f (binaryset A B n)) = f A + f B"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   113
  by (metis binaryset_sums sums_unique)
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   114
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   115
lemma binaryset_psuminf:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   116
  assumes "f {} = 0"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   117
  shows "(\<Sum>\<^isub>\<infinity> n. f (binaryset A B n)) = f A + f B" (is "?suminf = ?sum")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   118
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   119
  have *: "{..<2} = {0, 1::nat}" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   120
  have "\<forall>n\<ge>2. f (binaryset A B n) = 0"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   121
    unfolding binaryset_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   122
    using assms by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   123
  hence "?suminf = (\<Sum>N<2. f (binaryset A B N))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   124
    by (rule psuminf_finite)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   125
  also have "... = ?sum" unfolding * binaryset_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   126
    by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   127
  finally show ?thesis .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   128
qed
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   129
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   130
subsection {* Lambda Systems *}
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   131
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   132
lemma (in algebra) lambda_system_eq:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   133
    "lambda_system M f =
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   134
        {l. l \<in> sets M & (\<forall>x \<in> sets M. f (x \<inter> l) + f (x - l) = f x)}"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   135
proof -
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   136
  have [simp]: "!!l x. l \<in> sets M \<Longrightarrow> x \<in> sets M \<Longrightarrow> (space M - l) \<inter> x = x - l"
37032
58a0757031dd speed up some proofs and fix some warnings
huffman
parents: 36649
diff changeset
   137
    by (metis Int_Diff Int_absorb1 Int_commute sets_into_space)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   138
  show ?thesis
37032
58a0757031dd speed up some proofs and fix some warnings
huffman
parents: 36649
diff changeset
   139
    by (auto simp add: lambda_system_def) (metis Int_commute)+
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   140
qed
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   141
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   142
lemma (in algebra) lambda_system_empty:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   143
  "positive f \<Longrightarrow> {} \<in> lambda_system M f"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   144
  by (auto simp add: positive_def lambda_system_eq)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   145
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   146
lemma lambda_system_sets:
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   147
    "x \<in> lambda_system M f \<Longrightarrow> x \<in> sets M"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   148
  by (simp add:  lambda_system_def)
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   149
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   150
lemma (in algebra) lambda_system_Compl:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   151
  fixes f:: "'a set \<Rightarrow> pinfreal"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   152
  assumes x: "x \<in> lambda_system M f"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   153
  shows "space M - x \<in> lambda_system M f"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   154
  proof -
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   155
    have "x \<subseteq> space M"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   156
      by (metis sets_into_space lambda_system_sets x)
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   157
    hence "space M - (space M - x) = x"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   158
      by (metis double_diff equalityE)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   159
    with x show ?thesis
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   160
      by (force simp add: lambda_system_def ac_simps)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   161
  qed
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   162
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   163
lemma (in algebra) lambda_system_Int:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   164
  fixes f:: "'a set \<Rightarrow> pinfreal"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   165
  assumes xl: "x \<in> lambda_system M f" and yl: "y \<in> lambda_system M f"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   166
  shows "x \<inter> y \<in> lambda_system M f"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   167
  proof -
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   168
    from xl yl show ?thesis
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   169
      proof (auto simp add: positive_def lambda_system_eq Int)
33536
fd28b7399f2b eliminated hard tabulators;
wenzelm
parents: 33271
diff changeset
   170
        fix u
fd28b7399f2b eliminated hard tabulators;
wenzelm
parents: 33271
diff changeset
   171
        assume x: "x \<in> sets M" and y: "y \<in> sets M" and u: "u \<in> sets M"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   172
           and fx: "\<forall>z\<in>sets M. f (z \<inter> x) + f (z - x) = f z"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   173
           and fy: "\<forall>z\<in>sets M. f (z \<inter> y) + f (z - y) = f z"
33536
fd28b7399f2b eliminated hard tabulators;
wenzelm
parents: 33271
diff changeset
   174
        have "u - x \<inter> y \<in> sets M"
fd28b7399f2b eliminated hard tabulators;
wenzelm
parents: 33271
diff changeset
   175
          by (metis Diff Diff_Int Un u x y)
fd28b7399f2b eliminated hard tabulators;
wenzelm
parents: 33271
diff changeset
   176
        moreover
fd28b7399f2b eliminated hard tabulators;
wenzelm
parents: 33271
diff changeset
   177
        have "(u - (x \<inter> y)) \<inter> y = u \<inter> y - x" by blast
fd28b7399f2b eliminated hard tabulators;
wenzelm
parents: 33271
diff changeset
   178
        moreover
fd28b7399f2b eliminated hard tabulators;
wenzelm
parents: 33271
diff changeset
   179
        have "u - x \<inter> y - y = u - y" by blast
fd28b7399f2b eliminated hard tabulators;
wenzelm
parents: 33271
diff changeset
   180
        ultimately
fd28b7399f2b eliminated hard tabulators;
wenzelm
parents: 33271
diff changeset
   181
        have ey: "f (u - x \<inter> y) = f (u \<inter> y - x) + f (u - y)" using fy
fd28b7399f2b eliminated hard tabulators;
wenzelm
parents: 33271
diff changeset
   182
          by force
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   183
        have "f (u \<inter> (x \<inter> y)) + f (u - x \<inter> y)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   184
              = (f (u \<inter> (x \<inter> y)) + f (u \<inter> y - x)) + f (u - y)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   185
          by (simp add: ey ac_simps)
33536
fd28b7399f2b eliminated hard tabulators;
wenzelm
parents: 33271
diff changeset
   186
        also have "... =  (f ((u \<inter> y) \<inter> x) + f (u \<inter> y - x)) + f (u - y)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   187
          by (simp add: Int_ac)
33536
fd28b7399f2b eliminated hard tabulators;
wenzelm
parents: 33271
diff changeset
   188
        also have "... = f (u \<inter> y) + f (u - y)"
fd28b7399f2b eliminated hard tabulators;
wenzelm
parents: 33271
diff changeset
   189
          using fx [THEN bspec, of "u \<inter> y"] Int y u
fd28b7399f2b eliminated hard tabulators;
wenzelm
parents: 33271
diff changeset
   190
          by force
fd28b7399f2b eliminated hard tabulators;
wenzelm
parents: 33271
diff changeset
   191
        also have "... = f u"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   192
          by (metis fy u)
33536
fd28b7399f2b eliminated hard tabulators;
wenzelm
parents: 33271
diff changeset
   193
        finally show "f (u \<inter> (x \<inter> y)) + f (u - x \<inter> y) = f u" .
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   194
      qed
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   195
  qed
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   196
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   197
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   198
lemma (in algebra) lambda_system_Un:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   199
  fixes f:: "'a set \<Rightarrow> pinfreal"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   200
  assumes xl: "x \<in> lambda_system M f" and yl: "y \<in> lambda_system M f"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   201
  shows "x \<union> y \<in> lambda_system M f"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   202
proof -
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   203
  have "(space M - x) \<inter> (space M - y) \<in> sets M"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   204
    by (metis Diff_Un Un compl_sets lambda_system_sets xl yl)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   205
  moreover
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   206
  have "x \<union> y = space M - ((space M - x) \<inter> (space M - y))"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   207
    by auto  (metis subsetD lambda_system_sets sets_into_space xl yl)+
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   208
  ultimately show ?thesis
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   209
    by (metis lambda_system_Compl lambda_system_Int xl yl)
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
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   212
lemma (in algebra) lambda_system_algebra:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   213
  "positive f \<Longrightarrow> algebra (M (|sets := lambda_system M f|))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   214
  apply (auto simp add: algebra_def)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   215
  apply (metis lambda_system_sets set_mp sets_into_space)
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   216
  apply (metis lambda_system_empty)
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   217
  apply (metis lambda_system_Compl)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   218
  apply (metis lambda_system_Un)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   219
  done
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   220
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   221
lemma (in algebra) lambda_system_strong_additive:
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   222
  assumes z: "z \<in> sets M" and disj: "x \<inter> y = {}"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   223
      and xl: "x \<in> lambda_system M f" and yl: "y \<in> lambda_system M f"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   224
  shows "f (z \<inter> (x \<union> y)) = f (z \<inter> x) + f (z \<inter> y)"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   225
  proof -
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   226
    have "z \<inter> x = (z \<inter> (x \<union> y)) \<inter> x" using disj by blast
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   227
    moreover
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   228
    have "z \<inter> y = (z \<inter> (x \<union> y)) - x" using disj by blast
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   229
    moreover
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   230
    have "(z \<inter> (x \<union> y)) \<in> sets M"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   231
      by (metis Int Un lambda_system_sets xl yl z)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   232
    ultimately show ?thesis using xl yl
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   233
      by (simp add: lambda_system_eq)
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   234
  qed
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   235
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   236
lemma (in algebra) lambda_system_additive:
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   237
     "additive (M (|sets := lambda_system M f|)) f"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   238
  proof (auto simp add: additive_def)
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   239
    fix x and y
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   240
    assume disj: "x \<inter> y = {}"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   241
       and xl: "x \<in> lambda_system M f" and yl: "y \<in> lambda_system M f"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   242
    hence  "x \<in> sets M" "y \<in> sets M" by (blast intro: lambda_system_sets)+
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   243
    thus "f (x \<union> y) = f x + f y"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   244
      using lambda_system_strong_additive [OF top disj xl yl]
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   245
      by (simp add: Un)
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   246
  qed
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   247
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   248
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   249
lemma (in algebra) countably_subadditive_subadditive:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   250
  assumes f: "positive f" and cs: "countably_subadditive M f"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   251
  shows  "subadditive M f"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33536
diff changeset
   252
proof (auto simp add: subadditive_def)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   253
  fix x y
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   254
  assume x: "x \<in> sets M" and y: "y \<in> sets M" and "x \<inter> y = {}"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   255
  hence "disjoint_family (binaryset x y)"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33536
diff changeset
   256
    by (auto simp add: disjoint_family_on_def binaryset_def)
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33536
diff changeset
   257
  hence "range (binaryset x y) \<subseteq> sets M \<longrightarrow>
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33536
diff changeset
   258
         (\<Union>i. binaryset x y i) \<in> sets M \<longrightarrow>
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   259
         f (\<Union>i. binaryset x y i) \<le> (\<Sum>\<^isub>\<infinity> n. f (binaryset x y n))"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33536
diff changeset
   260
    using cs by (simp add: countably_subadditive_def)
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33536
diff changeset
   261
  hence "{x,y,{}} \<subseteq> sets M \<longrightarrow> x \<union> y \<in> sets M \<longrightarrow>
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   262
         f (x \<union> y) \<le> (\<Sum>\<^isub>\<infinity> n. f (binaryset x y n))"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   263
    by (simp add: range_binaryset_eq UN_binaryset_eq)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   264
  thus "f (x \<union> y) \<le>  f x + f y" using f x y
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   265
    by (auto simp add: Un o_def binaryset_psuminf positive_def)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   266
qed
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   267
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   268
lemma (in algebra) additive_sum:
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   269
  fixes A:: "nat \<Rightarrow> 'a set"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   270
  assumes f: "positive f" and ad: "additive M f"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   271
      and A: "range A \<subseteq> sets M"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   272
      and disj: "disjoint_family A"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   273
  shows  "setsum (f \<circ> A) {0..<n} = f (\<Union>i\<in>{0..<n}. A i)"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   274
proof (induct n)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   275
  case 0 show ?case using f by (simp add: positive_def)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   276
next
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   277
  case (Suc n)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   278
  have "A n \<inter> (\<Union>i\<in>{0..<n}. A i) = {}" using disj
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33536
diff changeset
   279
    by (auto simp add: disjoint_family_on_def neq_iff) blast
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   280
  moreover
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   281
  have "A n \<in> sets M" using A by blast
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   282
  moreover have "(\<Union>i\<in>{0..<n}. A i) \<in> sets M"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   283
    by (metis A UNION_in_sets atLeast0LessThan)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   284
  moreover
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   285
  ultimately have "f (A n \<union> (\<Union>i\<in>{0..<n}. A i)) = f (A n) + f(\<Union>i\<in>{0..<n}. A i)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   286
    using ad UNION_in_sets A by (auto simp add: additive_def)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   287
  with Suc.hyps show ?case using ad
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   288
    by (auto simp add: atLeastLessThanSuc additive_def)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   289
qed
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   290
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   291
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   292
lemma countably_subadditiveD:
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   293
  "countably_subadditive M f \<Longrightarrow> range A \<subseteq> sets M \<Longrightarrow> disjoint_family A \<Longrightarrow>
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   294
   (\<Union>i. A i) \<in> sets M \<Longrightarrow> f (\<Union>i. A i) \<le> psuminf (f o A)"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   295
  by (auto simp add: countably_subadditive_def o_def)
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   296
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   297
lemma (in algebra) increasing_additive_bound:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   298
  fixes A:: "nat \<Rightarrow> 'a set" and  f :: "'a set \<Rightarrow> pinfreal"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   299
  assumes f: "positive f" and ad: "additive M f"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   300
      and inc: "increasing M f"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   301
      and A: "range A \<subseteq> sets M"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   302
      and disj: "disjoint_family A"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   303
  shows  "psuminf (f \<circ> A) \<le> f (space M)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   304
proof (safe intro!: psuminf_bound)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   305
  fix N
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   306
  have "setsum (f \<circ> A) {0..<N} = f (\<Union>i\<in>{0..<N}. A i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   307
    by (rule additive_sum [OF f ad A disj])
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   308
  also have "... \<le> f (space M)" using space_closed A
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   309
    by (blast intro: increasingD [OF inc] UNION_in_sets top)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   310
  finally show "setsum (f \<circ> A) {..<N} \<le> f (space M)" by (simp add: atLeast0LessThan)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   311
qed
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   312
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   313
lemma lambda_system_increasing:
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   314
   "increasing M f \<Longrightarrow> increasing (M (|sets := lambda_system M f|)) f"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   315
  by (simp add: increasing_def lambda_system_def)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   316
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   317
lemma (in algebra) lambda_system_strong_sum:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   318
  fixes A:: "nat \<Rightarrow> 'a set" and f :: "'a set \<Rightarrow> pinfreal"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   319
  assumes f: "positive f" and a: "a \<in> sets M"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   320
      and A: "range A \<subseteq> lambda_system M f"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   321
      and disj: "disjoint_family A"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   322
  shows  "(\<Sum>i = 0..<n. f (a \<inter>A i)) = f (a \<inter> (\<Union>i\<in>{0..<n}. A i))"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   323
proof (induct n)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   324
  case 0 show ?case using f by (simp add: positive_def)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   325
next
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   326
  case (Suc n)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   327
  have 2: "A n \<inter> UNION {0..<n} A = {}" using disj
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   328
    by (force simp add: disjoint_family_on_def neq_iff)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   329
  have 3: "A n \<in> lambda_system M f" using A
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   330
    by blast
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   331
  have 4: "UNION {0..<n} A \<in> lambda_system M f"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   332
    using A algebra.UNION_in_sets [OF local.lambda_system_algebra, of f, OF f]
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   333
    by simp
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   334
  from Suc.hyps show ?case
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   335
    by (simp add: atLeastLessThanSuc lambda_system_strong_additive [OF a 2 3 4])
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   336
qed
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   337
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   338
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   339
lemma (in sigma_algebra) lambda_system_caratheodory:
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   340
  assumes oms: "outer_measure_space M f"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   341
      and A: "range A \<subseteq> lambda_system M f"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   342
      and disj: "disjoint_family A"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   343
  shows  "(\<Union>i. A i) \<in> lambda_system M f \<and> psuminf (f \<circ> A) = f (\<Union>i. A i)"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   344
proof -
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   345
  have pos: "positive f" and inc: "increasing M f"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   346
   and csa: "countably_subadditive M f"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   347
    by (metis oms outer_measure_space_def)+
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   348
  have sa: "subadditive M f"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   349
    by (metis countably_subadditive_subadditive csa pos)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   350
  have A': "range A \<subseteq> sets (M(|sets := lambda_system M f|))" using A
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   351
    by simp
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   352
  have alg_ls: "algebra (M(|sets := lambda_system M f|))"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   353
    by (rule lambda_system_algebra) (rule pos)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   354
  have A'': "range A \<subseteq> sets M"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   355
     by (metis A image_subset_iff lambda_system_sets)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   356
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   357
  have U_in: "(\<Union>i. A i) \<in> sets M"
37032
58a0757031dd speed up some proofs and fix some warnings
huffman
parents: 36649
diff changeset
   358
    by (metis A'' countable_UN)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   359
  have U_eq: "f (\<Union>i. A i) = psuminf (f o A)"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   360
    proof (rule antisym)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   361
      show "f (\<Union>i. A i) \<le> psuminf (f \<circ> A)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   362
        by (rule countably_subadditiveD [OF csa A'' disj U_in])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   363
      show "psuminf (f \<circ> A) \<le> f (\<Union>i. A i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   364
        by (rule psuminf_bound, unfold atLeast0LessThan[symmetric])
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   365
           (metis algebra.additive_sum [OF alg_ls] pos disj UN_Un Un_UNIV_right
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   366
                  lambda_system_additive subset_Un_eq increasingD [OF inc]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   367
                  A' A'' UNION_in_sets U_in)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   368
    qed
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   369
  {
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   370
    fix a
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   371
    assume a [iff]: "a \<in> sets M"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   372
    have "f (a \<inter> (\<Union>i. A i)) + f (a - (\<Union>i. A i)) = f a"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   373
    proof -
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   374
      show ?thesis
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   375
      proof (rule antisym)
33536
fd28b7399f2b eliminated hard tabulators;
wenzelm
parents: 33271
diff changeset
   376
        have "range (\<lambda>i. a \<inter> A i) \<subseteq> sets M" using A''
fd28b7399f2b eliminated hard tabulators;
wenzelm
parents: 33271
diff changeset
   377
          by blast
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   378
        moreover
33536
fd28b7399f2b eliminated hard tabulators;
wenzelm
parents: 33271
diff changeset
   379
        have "disjoint_family (\<lambda>i. a \<inter> A i)" using disj
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   380
          by (auto simp add: disjoint_family_on_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   381
        moreover
33536
fd28b7399f2b eliminated hard tabulators;
wenzelm
parents: 33271
diff changeset
   382
        have "a \<inter> (\<Union>i. A i) \<in> sets M"
fd28b7399f2b eliminated hard tabulators;
wenzelm
parents: 33271
diff changeset
   383
          by (metis Int U_in a)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   384
        ultimately
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   385
        have "f (a \<inter> (\<Union>i. A i)) \<le> psuminf (f \<circ> (\<lambda>i. a \<inter> i) \<circ> A)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   386
          using countably_subadditiveD [OF csa, of "(\<lambda>i. a \<inter> A i)"]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   387
          by (simp add: o_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   388
        hence "f (a \<inter> (\<Union>i. A i)) + f (a - (\<Union>i. A i)) \<le>
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   389
            psuminf (f \<circ> (\<lambda>i. a \<inter> i) \<circ> A) + f (a - (\<Union>i. A i))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   390
          by (rule add_right_mono)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   391
        moreover
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   392
        have "psuminf (f \<circ> (\<lambda>i. a \<inter> i) \<circ> A) + f (a - (\<Union>i. A i)) \<le> f a"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   393
          proof (safe intro!: psuminf_bound_add)
33536
fd28b7399f2b eliminated hard tabulators;
wenzelm
parents: 33271
diff changeset
   394
            fix n
fd28b7399f2b eliminated hard tabulators;
wenzelm
parents: 33271
diff changeset
   395
            have UNION_in: "(\<Union>i\<in>{0..<n}. A i) \<in> sets M"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   396
              by (metis A'' UNION_in_sets)
33536
fd28b7399f2b eliminated hard tabulators;
wenzelm
parents: 33271
diff changeset
   397
            have le_fa: "f (UNION {0..<n} A \<inter> a) \<le> f a" using A''
37032
58a0757031dd speed up some proofs and fix some warnings
huffman
parents: 36649
diff changeset
   398
              by (blast intro: increasingD [OF inc] A'' UNION_in_sets)
33536
fd28b7399f2b eliminated hard tabulators;
wenzelm
parents: 33271
diff changeset
   399
            have ls: "(\<Union>i\<in>{0..<n}. A i) \<in> lambda_system M f"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   400
              using algebra.UNION_in_sets [OF lambda_system_algebra [of f, OF pos]]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   401
              by (simp add: A)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   402
            hence eq_fa: "f a = f (a \<inter> (\<Union>i\<in>{0..<n}. A i)) + f (a - (\<Union>i\<in>{0..<n}. A i))"
37032
58a0757031dd speed up some proofs and fix some warnings
huffman
parents: 36649
diff changeset
   403
              by (simp add: lambda_system_eq UNION_in)
33536
fd28b7399f2b eliminated hard tabulators;
wenzelm
parents: 33271
diff changeset
   404
            have "f (a - (\<Union>i. A i)) \<le> f (a - (\<Union>i\<in>{0..<n}. A i))"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   405
              by (blast intro: increasingD [OF inc] UNION_eq_Union_image
37032
58a0757031dd speed up some proofs and fix some warnings
huffman
parents: 36649
diff changeset
   406
                               UNION_in U_in)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   407
            thus "setsum (f \<circ> (\<lambda>i. a \<inter> i) \<circ> A) {..<n} + f (a - (\<Union>i. A i)) \<le> f a"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   408
              by (simp add: lambda_system_strong_sum pos A disj eq_fa add_left_mono atLeast0LessThan[symmetric])
33536
fd28b7399f2b eliminated hard tabulators;
wenzelm
parents: 33271
diff changeset
   409
          qed
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   410
        ultimately show "f (a \<inter> (\<Union>i. A i)) + f (a - (\<Union>i. A i)) \<le> f a"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   411
          by (rule order_trans)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   412
      next
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   413
        have "f a \<le> f (a \<inter> (\<Union>i. A i) \<union> (a - (\<Union>i. A i)))"
37032
58a0757031dd speed up some proofs and fix some warnings
huffman
parents: 36649
diff changeset
   414
          by (blast intro:  increasingD [OF inc] U_in)
33536
fd28b7399f2b eliminated hard tabulators;
wenzelm
parents: 33271
diff changeset
   415
        also have "... \<le>  f (a \<inter> (\<Union>i. A i)) + f (a - (\<Union>i. A i))"
37032
58a0757031dd speed up some proofs and fix some warnings
huffman
parents: 36649
diff changeset
   416
          by (blast intro: subadditiveD [OF sa] U_in)
33536
fd28b7399f2b eliminated hard tabulators;
wenzelm
parents: 33271
diff changeset
   417
        finally show "f a \<le> f (a \<inter> (\<Union>i. A i)) + f (a - (\<Union>i. A i))" .
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   418
        qed
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   419
     qed
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   420
  }
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   421
  thus  ?thesis
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   422
    by (simp add: lambda_system_eq sums_iff U_eq U_in)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   423
qed
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   424
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   425
lemma (in sigma_algebra) caratheodory_lemma:
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   426
  assumes oms: "outer_measure_space M f"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   427
  shows "measure_space (|space = space M, sets = lambda_system M f|) f"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   428
proof -
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   429
  have pos: "positive f"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   430
    by (metis oms outer_measure_space_def)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   431
  have alg: "algebra (|space = space M, sets = lambda_system M f|)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   432
    using lambda_system_algebra [of f, OF pos]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   433
    by (simp add: algebra_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   434
  then moreover
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   435
  have "sigma_algebra (|space = space M, sets = lambda_system M f|)"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   436
    using lambda_system_caratheodory [OF oms]
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   437
    by (simp add: sigma_algebra_disjoint_iff)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   438
  moreover
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   439
  have "measure_space_axioms (|space = space M, sets = lambda_system M f|) f"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   440
    using pos lambda_system_caratheodory [OF oms]
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   441
    by (simp add: measure_space_axioms_def positive_def lambda_system_sets
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   442
                  countably_additive_def o_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   443
  ultimately
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   444
  show ?thesis
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   445
    by intro_locales (auto simp add: sigma_algebra_def)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   446
qed
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   447
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   448
lemma (in algebra) additive_increasing:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   449
  assumes posf: "positive f" and addf: "additive M f"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   450
  shows "increasing M f"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   451
proof (auto simp add: increasing_def)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   452
  fix x y
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   453
  assume xy: "x \<in> sets M" "y \<in> sets M" "x \<subseteq> y"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   454
  have "f x \<le> f x + f (y-x)" ..
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   455
  also have "... = f (x \<union> (y-x))" using addf
37032
58a0757031dd speed up some proofs and fix some warnings
huffman
parents: 36649
diff changeset
   456
    by (auto simp add: additive_def) (metis Diff_disjoint Un_Diff_cancel Diff xy(1,2))
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   457
  also have "... = f y"
37032
58a0757031dd speed up some proofs and fix some warnings
huffman
parents: 36649
diff changeset
   458
    by (metis Un_Diff_cancel Un_absorb1 xy(3))
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   459
  finally show "f x \<le> f y" .
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   460
qed
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   461
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   462
lemma (in algebra) countably_additive_additive:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   463
  assumes posf: "positive f" and ca: "countably_additive M f"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   464
  shows "additive M f"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   465
proof (auto simp add: additive_def)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   466
  fix x y
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   467
  assume x: "x \<in> sets M" and y: "y \<in> sets M" and "x \<inter> y = {}"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   468
  hence "disjoint_family (binaryset x y)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   469
    by (auto simp add: disjoint_family_on_def binaryset_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   470
  hence "range (binaryset x y) \<subseteq> sets M \<longrightarrow>
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   471
         (\<Union>i. binaryset x y i) \<in> sets M \<longrightarrow>
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   472
         f (\<Union>i. binaryset x y i) = (\<Sum>\<^isub>\<infinity> n. f (binaryset x y n))"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   473
    using ca
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   474
    by (simp add: countably_additive_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   475
  hence "{x,y,{}} \<subseteq> sets M \<longrightarrow> x \<union> y \<in> sets M \<longrightarrow>
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   476
         f (x \<union> y) = (\<Sum>\<^isub>\<infinity> n. f (binaryset x y n))"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   477
    by (simp add: range_binaryset_eq UN_binaryset_eq)
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   478
  thus "f (x \<union> y) = f x + f y" using posf x y
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   479
    by (auto simp add: Un binaryset_psuminf positive_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   480
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   481
39096
hoelzl
parents: 38656
diff changeset
   482
lemma inf_measure_nonempty:
hoelzl
parents: 38656
diff changeset
   483
  assumes f: "positive f" and b: "b \<in> sets M" and a: "a \<subseteq> b" "{} \<in> sets M"
hoelzl
parents: 38656
diff changeset
   484
  shows "f b \<in> measure_set M f a"
hoelzl
parents: 38656
diff changeset
   485
proof -
hoelzl
parents: 38656
diff changeset
   486
  have "psuminf (f \<circ> (\<lambda>i. {})(0 := b)) = setsum (f \<circ> (\<lambda>i. {})(0 := b)) {..<1::nat}"
hoelzl
parents: 38656
diff changeset
   487
    by (rule psuminf_finite) (simp add: f[unfolded positive_def])
hoelzl
parents: 38656
diff changeset
   488
  also have "... = f b"
hoelzl
parents: 38656
diff changeset
   489
    by simp
hoelzl
parents: 38656
diff changeset
   490
  finally have "psuminf (f \<circ> (\<lambda>i. {})(0 := b)) = f b" .
hoelzl
parents: 38656
diff changeset
   491
  thus ?thesis using assms
hoelzl
parents: 38656
diff changeset
   492
    by (auto intro!: exI [of _ "(\<lambda>i. {})(0 := b)"]
hoelzl
parents: 38656
diff changeset
   493
             simp: measure_set_def disjoint_family_on_def split_if_mem2 comp_def)
hoelzl
parents: 38656
diff changeset
   494
qed
hoelzl
parents: 38656
diff changeset
   495
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   496
lemma (in algebra) inf_measure_agrees:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   497
  assumes posf: "positive f" and ca: "countably_additive M f"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   498
      and s: "s \<in> sets M"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   499
  shows "Inf (measure_set M f s) = f s"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   500
  unfolding Inf_pinfreal_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   501
proof (safe intro!: Greatest_equality)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   502
  fix z
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   503
  assume z: "z \<in> measure_set M f s"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   504
  from this obtain A where
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   505
    A: "range A \<subseteq> sets M" and disj: "disjoint_family A"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   506
    and "s \<subseteq> (\<Union>x. A x)" and si: "psuminf (f \<circ> A) = z"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   507
    by (auto simp add: measure_set_def comp_def)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   508
  hence seq: "s = (\<Union>i. A i \<inter> s)" by blast
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   509
  have inc: "increasing M f"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   510
    by (metis additive_increasing ca countably_additive_additive posf)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   511
  have sums: "psuminf (\<lambda>i. f (A i \<inter> s)) = f (\<Union>i. A i \<inter> s)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   512
    proof (rule countably_additiveD [OF ca])
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   513
      show "range (\<lambda>n. A n \<inter> s) \<subseteq> sets M" using A s
33536
fd28b7399f2b eliminated hard tabulators;
wenzelm
parents: 33271
diff changeset
   514
        by blast
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   515
      show "disjoint_family (\<lambda>n. A n \<inter> s)" using disj
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33536
diff changeset
   516
        by (auto simp add: disjoint_family_on_def)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   517
      show "(\<Union>i. A i \<inter> s) \<in> sets M" using A s
33536
fd28b7399f2b eliminated hard tabulators;
wenzelm
parents: 33271
diff changeset
   518
        by (metis UN_extend_simps(4) s seq)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   519
    qed
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   520
  hence "f s = psuminf (\<lambda>i. f (A i \<inter> s))"
37032
58a0757031dd speed up some proofs and fix some warnings
huffman
parents: 36649
diff changeset
   521
    using seq [symmetric] by (simp add: sums_iff)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   522
  also have "... \<le> psuminf (f \<circ> A)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   523
    proof (rule psuminf_le)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   524
      fix n show "f (A n \<inter> s) \<le> (f \<circ> A) n" using A s
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   525
        by (force intro: increasingD [OF inc])
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   526
    qed
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   527
  also have "... = z" by (rule si)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   528
  finally show "f s \<le> z" .
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   529
next
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   530
  fix y
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   531
  assume y: "\<forall>u \<in> measure_set M f s. y \<le> u"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   532
  thus "y \<le> f s"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   533
    by (blast intro: inf_measure_nonempty [of f, OF posf s subset_refl])
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   534
qed
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   535
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   536
lemma (in algebra) inf_measure_empty:
39096
hoelzl
parents: 38656
diff changeset
   537
  assumes posf: "positive f"  "{} \<in> sets M"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   538
  shows "Inf (measure_set M f {}) = 0"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   539
proof (rule antisym)
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   540
  show "Inf (measure_set M f {}) \<le> 0"
39096
hoelzl
parents: 38656
diff changeset
   541
    by (metis complete_lattice_class.Inf_lower `{} \<in> sets M` inf_measure_nonempty[OF posf] subset_refl posf[unfolded positive_def])
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   542
qed simp
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   543
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   544
lemma (in algebra) inf_measure_positive:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   545
  "positive f \<Longrightarrow>
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   546
   positive (\<lambda>x. Inf (measure_set M f x))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   547
  by (simp add: positive_def inf_measure_empty) 
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   548
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   549
lemma (in algebra) inf_measure_increasing:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   550
  assumes posf: "positive f"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   551
  shows "increasing (| space = space M, sets = Pow (space M) |)
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   552
                    (\<lambda>x. Inf (measure_set M f x))"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   553
apply (auto simp add: increasing_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   554
apply (rule complete_lattice_class.Inf_greatest)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   555
apply (rule complete_lattice_class.Inf_lower)
37032
58a0757031dd speed up some proofs and fix some warnings
huffman
parents: 36649
diff changeset
   556
apply (clarsimp simp add: measure_set_def, rule_tac x=A in exI, blast)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   557
done
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   558
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   559
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   560
lemma (in algebra) inf_measure_le:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   561
  assumes posf: "positive f" and inc: "increasing M f"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   562
      and x: "x \<in> {r . \<exists>A. range A \<subseteq> sets M \<and> s \<subseteq> (\<Union>i. A i) \<and> psuminf (f \<circ> A) = r}"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   563
  shows "Inf (measure_set M f s) \<le> x"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   564
proof -
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   565
  from x
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   566
  obtain A where A: "range A \<subseteq> sets M" and ss: "s \<subseteq> (\<Union>i. A i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   567
             and xeq: "psuminf (f \<circ> A) = x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   568
    by auto
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   569
  have dA: "range (disjointed A) \<subseteq> sets M"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   570
    by (metis A range_disjointed_sets)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   571
  have "\<forall>n.(f o disjointed A) n \<le> (f \<circ> A) n" unfolding comp_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   572
    by (metis increasingD [OF inc] UNIV_I dA image_subset_iff disjointed_subset A comp_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   573
  hence sda: "psuminf (f o disjointed A) \<le> psuminf (f \<circ> A)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   574
    by (blast intro: psuminf_le)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   575
  hence ley: "psuminf (f o disjointed A) \<le> x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   576
    by (metis xeq)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   577
  hence y: "psuminf (f o disjointed A) \<in> measure_set M f s"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   578
    apply (auto simp add: measure_set_def)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   579
    apply (rule_tac x="disjointed A" in exI)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   580
    apply (simp add: disjoint_family_disjointed UN_disjointed_eq ss dA comp_def)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   581
    done
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   582
  show ?thesis
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   583
    by (blast intro: y order_trans [OF _ ley] posf complete_lattice_class.Inf_lower)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   584
qed
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   585
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   586
lemma (in algebra) inf_measure_close:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   587
  assumes posf: "positive f" and e: "0 < e" and ss: "s \<subseteq> (space M)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   588
  shows "\<exists>A. range A \<subseteq> sets M \<and> disjoint_family A \<and> s \<subseteq> (\<Union>i. A i) \<and>
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   589
               psuminf (f \<circ> A) \<le> Inf (measure_set M f s) + e"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   590
proof (cases "Inf (measure_set M f s) = \<omega>")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   591
  case False
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   592
  obtain l where "l \<in> measure_set M f s" "l \<le> Inf (measure_set M f s) + e"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   593
    using Inf_close[OF False e] by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   594
  thus ?thesis
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   595
    by (auto intro!: exI[of _ l] simp: measure_set_def comp_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   596
next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   597
  case True
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   598
  have "measure_set M f s \<noteq> {}"
39096
hoelzl
parents: 38656
diff changeset
   599
    by (metis emptyE ss inf_measure_nonempty [of f, OF posf top _ empty_sets])
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   600
  then obtain l where "l \<in> measure_set M f s" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   601
  moreover from True have "l \<le> Inf (measure_set M f s) + e" by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   602
  ultimately show ?thesis
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   603
    by (auto intro!: exI[of _ l] simp: measure_set_def comp_def)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   604
qed
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   605
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   606
lemma (in algebra) inf_measure_countably_subadditive:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   607
  assumes posf: "positive f" and inc: "increasing M f"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   608
  shows "countably_subadditive (| space = space M, sets = Pow (space M) |)
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   609
                  (\<lambda>x. Inf (measure_set M f x))"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   610
  unfolding countably_subadditive_def o_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   611
proof (safe, simp, rule pinfreal_le_epsilon)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   612
  fix A :: "nat \<Rightarrow> 'a set" and e :: pinfreal
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   613
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   614
  let "?outer n" = "Inf (measure_set M f (A n))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   615
  assume A: "range A \<subseteq> Pow (space M)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   616
     and disj: "disjoint_family A"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   617
     and sb: "(\<Union>i. A i) \<subseteq> space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   618
     and e: "0 < e"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   619
  hence "\<exists>BB. \<forall>n. range (BB n) \<subseteq> sets M \<and> disjoint_family (BB n) \<and>
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   620
                   A n \<subseteq> (\<Union>i. BB n i) \<and>
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   621
                   psuminf (f o BB n) \<le> ?outer n + e * (1/2)^(Suc n)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   622
    apply (safe intro!: choice inf_measure_close [of f, OF posf _])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   623
    using e sb by (cases e, auto simp add: not_le mult_pos_pos)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   624
  then obtain BB
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   625
    where BB: "\<And>n. (range (BB n) \<subseteq> sets M)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   626
      and disjBB: "\<And>n. disjoint_family (BB n)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   627
      and sbBB: "\<And>n. A n \<subseteq> (\<Union>i. BB n i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   628
      and BBle: "\<And>n. psuminf (f o BB n) \<le> ?outer n + e * (1/2)^(Suc n)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   629
    by auto blast
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   630
  have sll: "(\<Sum>\<^isub>\<infinity> n. psuminf (f o BB n)) \<le> psuminf ?outer + e"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   631
    proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   632
      have "(\<Sum>\<^isub>\<infinity> n. psuminf (f o BB n)) \<le> (\<Sum>\<^isub>\<infinity> n. ?outer n + e*(1/2) ^ Suc n)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   633
        by (rule psuminf_le[OF BBle])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   634
      also have "... = psuminf ?outer + e"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   635
        using psuminf_half_series by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   636
      finally show ?thesis .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   637
    qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   638
  def C \<equiv> "(split BB) o prod_decode"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   639
  have C: "!!n. C n \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   640
    apply (rule_tac p="prod_decode n" in PairE)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   641
    apply (simp add: C_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   642
    apply (metis BB subsetD rangeI)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   643
    done
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   644
  have sbC: "(\<Union>i. A i) \<subseteq> (\<Union>i. C i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   645
    proof (auto simp add: C_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   646
      fix x i
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   647
      assume x: "x \<in> A i"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   648
      with sbBB [of i] obtain j where "x \<in> BB i j"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   649
        by blast
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   650
      thus "\<exists>i. x \<in> split BB (prod_decode i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   651
        by (metis prod_encode_inverse prod.cases)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   652
    qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   653
  have "(f \<circ> C) = (f \<circ> (\<lambda>(x, y). BB x y)) \<circ> prod_decode"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   654
    by (rule ext)  (auto simp add: C_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   655
  moreover have "psuminf ... = (\<Sum>\<^isub>\<infinity> n. psuminf (f o BB n))" using BBle
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   656
    by (force intro!: psuminf_2dimen simp: o_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   657
  ultimately have Csums: "psuminf (f \<circ> C) = (\<Sum>\<^isub>\<infinity> n. psuminf (f o BB n))" by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   658
  have "Inf (measure_set M f (\<Union>i. A i)) \<le> (\<Sum>\<^isub>\<infinity> n. psuminf (f o BB n))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   659
    apply (rule inf_measure_le [OF posf(1) inc], auto)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   660
    apply (rule_tac x="C" in exI)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   661
    apply (auto simp add: C sbC Csums)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   662
    done
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   663
  also have "... \<le> (\<Sum>\<^isub>\<infinity>n. Inf (measure_set M f (A n))) + e" using sll
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   664
    by blast
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   665
  finally show "Inf (measure_set M f (\<Union>i. A i)) \<le> psuminf ?outer + e" .
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   666
qed
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   667
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   668
lemma (in algebra) inf_measure_outer:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   669
  "\<lbrakk> positive f ; increasing M f \<rbrakk>
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   670
   \<Longrightarrow> outer_measure_space (| space = space M, sets = Pow (space M) |)
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   671
                          (\<lambda>x. Inf (measure_set M f x))"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   672
  by (simp add: outer_measure_space_def inf_measure_empty
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   673
                inf_measure_increasing inf_measure_countably_subadditive positive_def)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   674
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   675
(*MOVE UP*)
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   676
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   677
lemma (in algebra) algebra_subset_lambda_system:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   678
  assumes posf: "positive f" and inc: "increasing M f"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   679
      and add: "additive M f"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   680
  shows "sets M \<subseteq> lambda_system (| space = space M, sets = Pow (space M) |)
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   681
                                (\<lambda>x. Inf (measure_set M f x))"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   682
proof (auto dest: sets_into_space
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   683
            simp add: algebra.lambda_system_eq [OF algebra_Pow])
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   684
  fix x s
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   685
  assume x: "x \<in> sets M"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   686
     and s: "s \<subseteq> space M"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   687
  have [simp]: "!!x. x \<in> sets M \<Longrightarrow> s \<inter> (space M - x) = s-x" using s
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   688
    by blast
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   689
  have "Inf (measure_set M f (s\<inter>x)) + Inf (measure_set M f (s-x))
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   690
        \<le> Inf (measure_set M f s)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   691
    proof (rule pinfreal_le_epsilon)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   692
      fix e :: pinfreal
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   693
      assume e: "0 < e"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   694
      from inf_measure_close [of f, OF posf e s]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   695
      obtain A where A: "range A \<subseteq> sets M" and disj: "disjoint_family A"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   696
                 and sUN: "s \<subseteq> (\<Union>i. A i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   697
                 and l: "psuminf (f \<circ> A) \<le> Inf (measure_set M f s) + e"
33536
fd28b7399f2b eliminated hard tabulators;
wenzelm
parents: 33271
diff changeset
   698
        by auto
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   699
      have [simp]: "!!x. x \<in> sets M \<Longrightarrow>
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   700
                      (f o (\<lambda>z. z \<inter> (space M - x)) o A) = (f o (\<lambda>z. z - x) o A)"
33536
fd28b7399f2b eliminated hard tabulators;
wenzelm
parents: 33271
diff changeset
   701
        by (rule ext, simp, metis A Int_Diff Int_space_eq2 range_subsetD)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   702
      have  [simp]: "!!n. f (A n \<inter> x) + f (A n - x) = f (A n)"
33536
fd28b7399f2b eliminated hard tabulators;
wenzelm
parents: 33271
diff changeset
   703
        by (subst additiveD [OF add, symmetric])
fd28b7399f2b eliminated hard tabulators;
wenzelm
parents: 33271
diff changeset
   704
           (auto simp add: x range_subsetD [OF A] Int_Diff_Un Int_Diff_disjoint)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   705
      { fix u
33536
fd28b7399f2b eliminated hard tabulators;
wenzelm
parents: 33271
diff changeset
   706
        assume u: "u \<in> sets M"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   707
        have [simp]: "\<And>n. f (A n \<inter> u) \<le> f (A n)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   708
          by (simp add: increasingD [OF inc] u Int range_subsetD [OF A])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   709
        have 2: "Inf (measure_set M f (s \<inter> u)) \<le> psuminf (f \<circ> (\<lambda>z. z \<inter> u) \<circ> A)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   710
          proof (rule complete_lattice_class.Inf_lower)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   711
            show "psuminf (f \<circ> (\<lambda>z. z \<inter> u) \<circ> A) \<in> measure_set M f (s \<inter> u)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   712
              apply (simp add: measure_set_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   713
              apply (rule_tac x="(\<lambda>z. z \<inter> u) o A" in exI)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   714
              apply (auto simp add: disjoint_family_subset [OF disj] o_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   715
              apply (blast intro: u range_subsetD [OF A])
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   716
              apply (blast dest: subsetD [OF sUN])
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   717
              done
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   718
          qed
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   719
      } note lesum = this
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   720
      have inf1: "Inf (measure_set M f (s\<inter>x)) \<le> psuminf (f o (\<lambda>z. z\<inter>x) o A)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   721
        and inf2: "Inf (measure_set M f (s \<inter> (space M - x)))
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   722
                   \<le> psuminf (f o (\<lambda>z. z \<inter> (space M - x)) o A)"
33536
fd28b7399f2b eliminated hard tabulators;
wenzelm
parents: 33271
diff changeset
   723
        by (metis Diff lesum top x)+
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   724
      hence "Inf (measure_set M f (s\<inter>x)) + Inf (measure_set M f (s-x))
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   725
           \<le> psuminf (f o (\<lambda>s. s\<inter>x) o A) + psuminf (f o (\<lambda>s. s-x) o A)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   726
        by (simp add: x add_mono)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   727
      also have "... \<le> psuminf (f o A)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   728
        by (simp add: x psuminf_add[symmetric] o_def)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   729
      also have "... \<le> Inf (measure_set M f s) + e"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   730
        by (rule l)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   731
      finally show "Inf (measure_set M f (s\<inter>x)) + Inf (measure_set M f (s-x))
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   732
        \<le> Inf (measure_set M f s) + e" .
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   733
    qed
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   734
  moreover
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   735
  have "Inf (measure_set M f s)
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   736
       \<le> Inf (measure_set M f (s\<inter>x)) + Inf (measure_set M f (s-x))"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   737
    proof -
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   738
    have "Inf (measure_set M f s) = Inf (measure_set M f ((s\<inter>x) \<union> (s-x)))"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   739
      by (metis Un_Diff_Int Un_commute)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   740
    also have "... \<le> Inf (measure_set M f (s\<inter>x)) + Inf (measure_set M f (s-x))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   741
      apply (rule subadditiveD)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   742
      apply (iprover intro: algebra.countably_subadditive_subadditive algebra_Pow
33536
fd28b7399f2b eliminated hard tabulators;
wenzelm
parents: 33271
diff changeset
   743
               inf_measure_positive inf_measure_countably_subadditive posf inc)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   744
      apply (auto simp add: subsetD [OF s])
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   745
      done
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   746
    finally show ?thesis .
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   747
    qed
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   748
  ultimately
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   749
  show "Inf (measure_set M f (s\<inter>x)) + Inf (measure_set M f (s-x))
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   750
        = Inf (measure_set M f s)"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   751
    by (rule order_antisym)
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   752
qed
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   753
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   754
lemma measure_down:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   755
     "measure_space N \<mu> \<Longrightarrow> sigma_algebra M \<Longrightarrow> sets M \<subseteq> sets N \<Longrightarrow>
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   756
      (\<nu> = \<mu>) \<Longrightarrow> measure_space M \<nu>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   757
  by (simp add: measure_space_def measure_space_axioms_def positive_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   758
                countably_additive_def)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   759
     blast
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   760
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   761
theorem (in algebra) caratheodory:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   762
  assumes posf: "positive f" and ca: "countably_additive M f"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   763
  shows "\<exists>\<mu> :: 'a set \<Rightarrow> pinfreal. (\<forall>s \<in> sets M. \<mu> s = f s) \<and> measure_space (sigma (space M) (sets M)) \<mu>"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   764
  proof -
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   765
    have inc: "increasing M f"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   766
      by (metis additive_increasing ca countably_additive_additive posf)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   767
    let ?infm = "(\<lambda>x. Inf (measure_set M f x))"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   768
    def ls \<equiv> "lambda_system (|space = space M, sets = Pow (space M)|) ?infm"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   769
    have mls: "measure_space \<lparr>space = space M, sets = ls\<rparr> ?infm"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   770
      using sigma_algebra.caratheodory_lemma
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   771
              [OF sigma_algebra_Pow  inf_measure_outer [OF posf inc]]
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   772
      by (simp add: ls_def)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   773
    hence sls: "sigma_algebra (|space = space M, sets = ls|)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   774
      by (simp add: measure_space_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   775
    have "sets M \<subseteq> ls"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   776
      by (simp add: ls_def)
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   777
         (metis ca posf inc countably_additive_additive algebra_subset_lambda_system)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   778
    hence sgs_sb: "sigma_sets (space M) (sets M) \<subseteq> ls"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   779
      using sigma_algebra.sigma_sets_subset [OF sls, of "sets M"]
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   780
      by simp
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   781
    have "measure_space (sigma (space M) (sets M)) ?infm"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   782
      unfolding sigma_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   783
      by (rule measure_down [OF mls], rule sigma_algebra_sigma_sets)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   784
         (simp_all add: sgs_sb space_closed)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   785
    thus ?thesis using inf_measure_agrees [OF posf ca] by (auto intro!: exI[of _ ?infm])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   786
  qed
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   787
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   788
end