src/HOL/Probability/Caratheodory.thy
author haftmann
Tue, 09 Aug 2011 20:24:48 +0200
changeset 44106 0e018cbcc0de
parent 43920 cedb5cb948fd
child 44568 e6f291cb5810
permissions -rw-r--r--
tuned proofs
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
42067
66c8281349ec standardized headers
hoelzl
parents: 42066
diff changeset
     1
(*  Title:      HOL/Probability/Caratheodory.thy
66c8281349ec standardized headers
hoelzl
parents: 42066
diff changeset
     2
    Author:     Lawrence C Paulson
66c8281349ec standardized headers
hoelzl
parents: 42066
diff changeset
     3
    Author:     Johannes Hölzl, TU München
66c8281349ec standardized headers
hoelzl
parents: 42066
diff changeset
     4
*)
66c8281349ec standardized headers
hoelzl
parents: 42066
diff changeset
     5
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
     6
header {*Caratheodory Extension Theorem*}
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
     7
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
     8
theory Caratheodory
44106
0e018cbcc0de tuned proofs
haftmann
parents: 43920
diff changeset
     9
imports Sigma_Algebra "~~/src/HOL/Multivariate_Analysis/Extended_Real_Limits"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    10
begin
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    11
42950
6e5c2a3c69da move lemmas to Extended_Reals and Extended_Real_Limits
hoelzl
parents: 42145
diff changeset
    12
lemma sums_def2:
6e5c2a3c69da move lemmas to Extended_Reals and Extended_Real_Limits
hoelzl
parents: 42145
diff changeset
    13
  "f sums x \<longleftrightarrow> (\<lambda>n. (\<Sum>i\<le>n. f i)) ----> x"
6e5c2a3c69da move lemmas to Extended_Reals and Extended_Real_Limits
hoelzl
parents: 42145
diff changeset
    14
  unfolding sums_def
6e5c2a3c69da move lemmas to Extended_Reals and Extended_Real_Limits
hoelzl
parents: 42145
diff changeset
    15
  apply (subst LIMSEQ_Suc_iff[symmetric])
6e5c2a3c69da move lemmas to Extended_Reals and Extended_Real_Limits
hoelzl
parents: 42145
diff changeset
    16
  unfolding atLeastLessThanSuc_atLeastAtMost atLeast0AtMost ..
6e5c2a3c69da move lemmas to Extended_Reals and Extended_Real_Limits
hoelzl
parents: 42145
diff changeset
    17
42067
66c8281349ec standardized headers
hoelzl
parents: 42066
diff changeset
    18
text {*
66c8281349ec standardized headers
hoelzl
parents: 42066
diff changeset
    19
  Originally from the Hurd/Coble measure theory development, translated by Lawrence Paulson.
66c8281349ec standardized headers
hoelzl
parents: 42066
diff changeset
    20
*}
66c8281349ec standardized headers
hoelzl
parents: 42066
diff changeset
    21
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
    22
lemma suminf_ereal_2dimen:
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
    23
  fixes f:: "nat \<times> nat \<Rightarrow> ereal"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
    24
  assumes pos: "\<And>p. 0 \<le> f p"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
    25
  assumes "\<And>m. g m = (\<Sum>n. f (m,n))"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
    26
  shows "(\<Sum>i. f (prod_decode i)) = suminf g"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
    27
proof -
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
    28
  have g_def: "g = (\<lambda>m. (\<Sum>n. f (m,n)))"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
    29
    using assms by (simp add: fun_eq_iff)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
    30
  have reindex: "\<And>B. (\<Sum>x\<in>B. f (prod_decode x)) = setsum f (prod_decode ` B)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
    31
    by (simp add: setsum_reindex[OF inj_prod_decode] comp_def)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
    32
  { fix n
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
    33
    let ?M = "\<lambda>f. Suc (Max (f ` prod_decode ` {..<n}))"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
    34
    { fix a b x assume "x < n" and [symmetric]: "(a, b) = prod_decode x"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
    35
      then have "a < ?M fst" "b < ?M snd"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
    36
        by (auto intro!: Max_ge le_imp_less_Suc image_eqI) }
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
    37
    then have "setsum f (prod_decode ` {..<n}) \<le> setsum f ({..<?M fst} \<times> {..<?M snd})"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
    38
      by (auto intro!: setsum_mono3 simp: pos)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
    39
    then have "\<exists>a b. setsum f (prod_decode ` {..<n}) \<le> setsum f ({..<a} \<times> {..<b})" by auto }
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
    40
  moreover
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
    41
  { fix a b
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
    42
    let ?M = "prod_decode ` {..<Suc (Max (prod_encode ` ({..<a} \<times> {..<b})))}"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
    43
    { fix a' b' assume "a' < a" "b' < b" then have "(a', b') \<in> ?M"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
    44
        by (auto intro!: Max_ge le_imp_less_Suc image_eqI[where x="prod_encode (a', b')"]) }
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
    45
    then have "setsum f ({..<a} \<times> {..<b}) \<le> setsum f ?M"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
    46
      by (auto intro!: setsum_mono3 simp: pos) }
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
    47
  ultimately
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
    48
  show ?thesis unfolding g_def using pos
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
    49
    by (auto intro!: SUPR_eq  simp: setsum_cartesian_product reindex le_SUPI2
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
    50
                     setsum_nonneg suminf_ereal_eq_SUPR SUPR_pair
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
    51
                     SUPR_ereal_setsum[symmetric] incseq_setsumI setsum_nonneg)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
    52
qed
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
    53
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    54
subsection {* Measure Spaces *}
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    55
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
    56
record 'a measure_space = "'a algebra" +
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
    57
  measure :: "'a set \<Rightarrow> ereal"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    58
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
    59
definition positive where "positive M f \<longleftrightarrow> f {} = (0::ereal) \<and> (\<forall>A\<in>sets M. 0 \<le> f A)"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    60
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
    61
definition additive where "additive M f \<longleftrightarrow>
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
    62
  (\<forall>x \<in> sets M. \<forall>y \<in> sets M. x \<inter> y = {} \<longrightarrow> f (x \<union> y) = f x + f y)"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    63
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
    64
definition countably_additive :: "('a, 'b) algebra_scheme \<Rightarrow> ('a set \<Rightarrow> ereal) \<Rightarrow> bool" where
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
    65
  "countably_additive M f \<longleftrightarrow> (\<forall>A. range A \<subseteq> sets M \<longrightarrow> disjoint_family A \<longrightarrow> (\<Union>i. A i) \<in> sets M \<longrightarrow>
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
    66
    (\<Sum>i. f (A i)) = f (\<Union>i. A i))"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    67
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
    68
definition increasing where "increasing M f \<longleftrightarrow>
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
    69
  (\<forall>x \<in> sets M. \<forall>y \<in> sets M. x \<subseteq> y \<longrightarrow> f x \<le> f y)"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    70
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
    71
definition subadditive where "subadditive M f \<longleftrightarrow>
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
    72
  (\<forall>x \<in> sets M. \<forall>y \<in> sets M. x \<inter> y = {} \<longrightarrow> f (x \<union> y) \<le> f x + f y)"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    73
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
    74
definition countably_subadditive where "countably_subadditive M f \<longleftrightarrow>
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
    75
  (\<forall>A. range A \<subseteq> sets M \<longrightarrow> disjoint_family A \<longrightarrow> (\<Union>i. A i) \<in> sets M \<longrightarrow>
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
    76
    (f (\<Union>i. A i) \<le> (\<Sum>i. f (A i))))"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
    77
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
    78
definition lambda_system where "lambda_system M f = {l \<in> sets M.
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
    79
  \<forall>x \<in> sets M. f (l \<inter> x) + f ((space M - l) \<inter> x) = f x}"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    80
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
    81
definition outer_measure_space where "outer_measure_space M f \<longleftrightarrow>
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
    82
  positive M f \<and> increasing M f \<and> countably_subadditive M f"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
    83
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
    84
definition measure_set where "measure_set M f X = {r.
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
    85
  \<exists>A. range A \<subseteq> sets M \<and> disjoint_family A \<and> X \<subseteq> (\<Union>i. A i) \<and> (\<Sum>i. f (A i)) = r}"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    86
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
    87
locale measure_space = sigma_algebra M for M :: "('a, 'b) measure_space_scheme" +
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
    88
  assumes measure_positive: "positive M (measure M)"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
    89
      and ca: "countably_additive M (measure M)"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    90
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
    91
abbreviation (in measure_space) "\<mu> \<equiv> measure M"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    92
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
    93
lemma (in measure_space)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
    94
  shows empty_measure[simp, intro]: "\<mu> {} = 0"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
    95
  and positive_measure[simp, intro!]: "\<And>A. A \<in> sets M \<Longrightarrow> 0 \<le> \<mu> A"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
    96
  using measure_positive unfolding positive_def by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
    97
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    98
lemma increasingD:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
    99
  "increasing M f \<Longrightarrow> x \<subseteq> y \<Longrightarrow> x\<in>sets M \<Longrightarrow> y\<in>sets M \<Longrightarrow> f x \<le> f y"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   100
  by (auto simp add: increasing_def)
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   101
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   102
lemma subadditiveD:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   103
  "subadditive M f \<Longrightarrow> x \<inter> y = {} \<Longrightarrow> x \<in> sets M \<Longrightarrow> y \<in> sets M
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   104
    \<Longrightarrow> f (x \<union> y) \<le> f x + f y"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   105
  by (auto simp add: subadditive_def)
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   106
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   107
lemma additiveD:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   108
  "additive M f \<Longrightarrow> x \<inter> y = {} \<Longrightarrow> x \<in> sets M \<Longrightarrow> y \<in> sets M
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   109
    \<Longrightarrow> f (x \<union> y) = f x + f y"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   110
  by (auto simp add: additive_def)
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   111
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   112
lemma countably_additiveI:
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   113
  assumes "\<And>A x. range A \<subseteq> sets M \<Longrightarrow> disjoint_family A \<Longrightarrow> (\<Union>i. A i) \<in> sets M
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   114
    \<Longrightarrow> (\<Sum>i. f (A i)) = f (\<Union>i. A i)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   115
  shows "countably_additive M f"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   116
  using assms by (simp add: countably_additive_def)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   117
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   118
section "Extend binary sets"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   119
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33536
diff changeset
   120
lemma LIMSEQ_binaryset:
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   121
  assumes f: "f {} = 0"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   122
  shows  "(\<lambda>n. \<Sum>i<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
   123
proof -
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   124
  have "(\<lambda>n. \<Sum>i < 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
   125
    proof
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   126
      fix n
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   127
      show "(\<Sum>i < 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
   128
        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
   129
    qed
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   130
  moreover
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33536
diff changeset
   131
  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
   132
  ultimately
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   133
  have "(\<lambda>n. \<Sum>i< 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
   134
    by metis
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   135
  hence "(\<lambda>n. \<Sum>i< n+2. 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
   136
    by simp
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   137
  thus ?thesis by (rule LIMSEQ_offset [where k=2])
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   138
qed
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   139
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   140
lemma binaryset_sums:
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   141
  assumes f: "f {} = 0"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   142
  shows  "(\<lambda>n. f (binaryset A B n)) sums (f A + f B)"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   143
    by (simp add: sums_def LIMSEQ_binaryset [where f=f, OF f] atLeast0LessThan)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   144
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   145
lemma suminf_binaryset_eq:
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   146
  fixes f :: "'a set \<Rightarrow> 'b::{comm_monoid_add, t2_space}"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   147
  shows "f {} = 0 \<Longrightarrow> (\<Sum>n. f (binaryset A B n)) = f A + f B"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   148
  by (metis binaryset_sums sums_unique)
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
subsection {* Lambda Systems *}
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   151
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   152
lemma (in algebra) lambda_system_eq:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   153
  shows "lambda_system M f = {l \<in> sets M.
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   154
    \<forall>x \<in> sets M. f (x \<inter> l) + f (x - l) = f x}"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   155
proof -
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   156
  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
   157
    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
   158
  show ?thesis
37032
58a0757031dd speed up some proofs and fix some warnings
huffman
parents: 36649
diff changeset
   159
    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
   160
qed
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   161
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   162
lemma (in algebra) lambda_system_empty:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   163
  "positive M f \<Longrightarrow> {} \<in> lambda_system M f"
42066
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   164
  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
   165
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   166
lemma lambda_system_sets:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   167
  "x \<in> lambda_system M f \<Longrightarrow> x \<in> sets M"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   168
  by (simp add: lambda_system_def)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   169
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   170
lemma (in algebra) lambda_system_Compl:
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
   171
  fixes f:: "'a set \<Rightarrow> ereal"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   172
  assumes x: "x \<in> lambda_system M f"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   173
  shows "space M - x \<in> lambda_system M f"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   174
proof -
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   175
  have "x \<subseteq> space M"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   176
    by (metis sets_into_space lambda_system_sets x)
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   177
  hence "space M - (space M - x) = x"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   178
    by (metis double_diff equalityE)
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   179
  with x show ?thesis
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   180
    by (force simp add: lambda_system_def ac_simps)
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   181
qed
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   182
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   183
lemma (in algebra) lambda_system_Int:
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
   184
  fixes f:: "'a set \<Rightarrow> ereal"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   185
  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
   186
  shows "x \<inter> y \<in> lambda_system M f"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   187
proof -
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   188
  from xl yl show ?thesis
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   189
  proof (auto simp add: positive_def lambda_system_eq Int)
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   190
    fix u
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   191
    assume x: "x \<in> sets M" and y: "y \<in> sets M" and u: "u \<in> sets M"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   192
       and fx: "\<forall>z\<in>sets M. f (z \<inter> x) + f (z - x) = f z"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   193
       and fy: "\<forall>z\<in>sets M. f (z \<inter> y) + f (z - y) = f z"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   194
    have "u - x \<inter> y \<in> sets M"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   195
      by (metis Diff Diff_Int Un u x y)
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   196
    moreover
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   197
    have "(u - (x \<inter> y)) \<inter> y = u \<inter> y - x" by blast
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   198
    moreover
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   199
    have "u - x \<inter> y - y = u - y" by blast
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   200
    ultimately
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   201
    have ey: "f (u - x \<inter> y) = f (u \<inter> y - x) + f (u - y)" using fy
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   202
      by force
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   203
    have "f (u \<inter> (x \<inter> y)) + f (u - x \<inter> y)
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   204
          = (f (u \<inter> (x \<inter> y)) + f (u \<inter> y - x)) + f (u - y)"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   205
      by (simp add: ey ac_simps)
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   206
    also have "... =  (f ((u \<inter> y) \<inter> x) + f (u \<inter> y - x)) + f (u - y)"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   207
      by (simp add: Int_ac)
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   208
    also have "... = f (u \<inter> y) + f (u - y)"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   209
      using fx [THEN bspec, of "u \<inter> y"] Int y u
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   210
      by force
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   211
    also have "... = f u"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   212
      by (metis fy u)
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   213
    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
   214
  qed
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   215
qed
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   216
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   217
lemma (in algebra) lambda_system_Un:
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
   218
  fixes f:: "'a set \<Rightarrow> ereal"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   219
  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
   220
  shows "x \<union> y \<in> lambda_system M f"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   221
proof -
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   222
  have "(space M - x) \<inter> (space M - y) \<in> sets M"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   223
    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
   224
  moreover
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   225
  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
   226
    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
   227
  ultimately show ?thesis
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   228
    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
   229
qed
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   230
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   231
lemma (in algebra) lambda_system_algebra:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   232
  "positive M f \<Longrightarrow> algebra (M\<lparr>sets := lambda_system M f\<rparr>)"
42065
2b98b4c2e2f1 add ring_of_sets and subset_class as basis for algebra
hoelzl
parents: 41981
diff changeset
   233
  apply (auto simp add: algebra_iff_Un)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   234
  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
   235
  apply (metis lambda_system_empty)
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   236
  apply (metis lambda_system_Compl)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   237
  apply (metis lambda_system_Un)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   238
  done
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   239
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   240
lemma (in algebra) lambda_system_strong_additive:
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   241
  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
   242
      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
   243
  shows "f (z \<inter> (x \<union> y)) = f (z \<inter> x) + f (z \<inter> y)"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   244
proof -
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   245
  have "z \<inter> x = (z \<inter> (x \<union> y)) \<inter> x" using disj by blast
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   246
  moreover
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   247
  have "z \<inter> y = (z \<inter> (x \<union> y)) - x" using disj by blast
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   248
  moreover
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   249
  have "(z \<inter> (x \<union> y)) \<in> sets M"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   250
    by (metis Int Un lambda_system_sets xl yl z)
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   251
  ultimately show ?thesis using xl yl
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   252
    by (simp add: lambda_system_eq)
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   253
qed
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   254
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   255
lemma (in algebra) lambda_system_additive:
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   256
     "additive (M (|sets := lambda_system M f|)) f"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   257
proof (auto simp add: additive_def)
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   258
  fix x and y
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   259
  assume disj: "x \<inter> y = {}"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   260
     and xl: "x \<in> lambda_system M f" and yl: "y \<in> lambda_system M f"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   261
  hence  "x \<in> sets M" "y \<in> sets M" by (blast intro: lambda_system_sets)+
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   262
  thus "f (x \<union> y) = f x + f y"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   263
    using lambda_system_strong_additive [OF top disj xl yl]
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   264
    by (simp add: Un)
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   265
qed
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   266
42145
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   267
lemma (in ring_of_sets) disjointed_additive:
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   268
  assumes f: "positive M f" "additive M f" and A: "range A \<subseteq> sets M" "incseq A"
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   269
  shows "(\<Sum>i\<le>n. f (disjointed A i)) = f (A n)"
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   270
proof (induct n)
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   271
  case (Suc n)
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   272
  then have "(\<Sum>i\<le>Suc n. f (disjointed A i)) = f (A n) + f (disjointed A (Suc n))"
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   273
    by simp
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   274
  also have "\<dots> = f (A n \<union> disjointed A (Suc n))"
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   275
    using A by (subst f(2)[THEN additiveD]) (auto simp: disjointed_incseq)
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   276
  also have "A n \<union> disjointed A (Suc n) = A (Suc n)"
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   277
    using `incseq A` by (auto dest: incseq_SucD simp: disjointed_incseq)
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   278
  finally show ?case .
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   279
qed simp
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   280
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   281
lemma (in ring_of_sets) countably_subadditive_subadditive:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   282
  assumes f: "positive M f" and cs: "countably_subadditive M f"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   283
  shows  "subadditive M f"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33536
diff changeset
   284
proof (auto simp add: subadditive_def)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   285
  fix x y
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   286
  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
   287
  hence "disjoint_family (binaryset x y)"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33536
diff changeset
   288
    by (auto simp add: disjoint_family_on_def binaryset_def)
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33536
diff changeset
   289
  hence "range (binaryset x y) \<subseteq> sets M \<longrightarrow>
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33536
diff changeset
   290
         (\<Union>i. binaryset x y i) \<in> sets M \<longrightarrow>
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   291
         f (\<Union>i. binaryset x y i) \<le> (\<Sum> n. f (binaryset x y n))"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   292
    using cs by (auto simp add: countably_subadditive_def)
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33536
diff changeset
   293
  hence "{x,y,{}} \<subseteq> sets M \<longrightarrow> x \<union> y \<in> sets M \<longrightarrow>
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   294
         f (x \<union> y) \<le> (\<Sum> n. f (binaryset x y n))"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   295
    by (simp add: range_binaryset_eq UN_binaryset_eq)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   296
  thus "f (x \<union> y) \<le>  f x + f y" using f x y
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   297
    by (auto simp add: Un o_def suminf_binaryset_eq positive_def)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   298
qed
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   299
42145
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   300
lemma (in ring_of_sets) additive_sum:
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   301
  fixes A:: "nat \<Rightarrow> 'a set"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   302
  assumes f: "positive M f" and ad: "additive M f" and "finite S"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   303
      and A: "range A \<subseteq> sets M"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   304
      and disj: "disjoint_family_on A S"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   305
  shows  "(\<Sum>i\<in>S. f (A i)) = f (\<Union>i\<in>S. A i)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   306
using `finite S` disj proof induct
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   307
  case empty 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
   308
next
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   309
  case (insert s S)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   310
  then have "A s \<inter> (\<Union>i\<in>S. A i) = {}"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   311
    by (auto simp add: disjoint_family_on_def neq_iff)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   312
  moreover
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   313
  have "A s \<in> sets M" using A by blast
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   314
  moreover have "(\<Union>i\<in>S. A i) \<in> sets M"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   315
    using A `finite S` by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   316
  moreover
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   317
  ultimately have "f (A s \<union> (\<Union>i\<in>S. A i)) = f (A s) + f(\<Union>i\<in>S. A i)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   318
    using ad UNION_in_sets A by (auto simp add: additive_def)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   319
  with insert show ?case using ad disjoint_family_on_mono[of S "insert s S" A]
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   320
    by (auto simp add: additive_def subset_insertI)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   321
qed
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   322
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   323
lemma (in algebra) increasing_additive_bound:
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
   324
  fixes A:: "nat \<Rightarrow> 'a set" and  f :: "'a set \<Rightarrow> ereal"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   325
  assumes f: "positive M f" and ad: "additive M f"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   326
      and inc: "increasing M f"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   327
      and A: "range A \<subseteq> sets M"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   328
      and disj: "disjoint_family A"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   329
  shows  "(\<Sum>i. f (A i)) \<le> f (space M)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   330
proof (safe intro!: suminf_bound)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   331
  fix N
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   332
  note disj_N = disjoint_family_on_mono[OF _ disj, of "{..<N}"]
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   333
  have "(\<Sum>i<N. f (A i)) = f (\<Union>i\<in>{..<N}. A i)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   334
    by (rule additive_sum [OF f ad _ A]) (auto simp: disj_N)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   335
  also have "... \<le> f (space M)" using space_closed A
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   336
    by (intro increasingD[OF inc] finite_UN) auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   337
  finally show "(\<Sum>i<N. f (A i)) \<le> f (space M)" by simp
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   338
qed (insert f A, auto simp: positive_def)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   339
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   340
lemma lambda_system_increasing:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   341
 "increasing M f \<Longrightarrow> increasing (M (|sets := lambda_system M f|)) f"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   342
  by (simp add: increasing_def lambda_system_def)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   343
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   344
lemma lambda_system_positive:
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   345
  "positive M f \<Longrightarrow> positive (M (|sets := lambda_system M f|)) f"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   346
  by (simp add: positive_def lambda_system_def)
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   347
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   348
lemma (in algebra) lambda_system_strong_sum:
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
   349
  fixes A:: "nat \<Rightarrow> 'a set" and f :: "'a set \<Rightarrow> ereal"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   350
  assumes f: "positive M f" and a: "a \<in> sets M"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   351
      and A: "range A \<subseteq> lambda_system M f"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   352
      and disj: "disjoint_family A"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   353
  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
   354
proof (induct n)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   355
  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
   356
next
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   357
  case (Suc n)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   358
  have 2: "A n \<inter> UNION {0..<n} A = {}" using disj
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   359
    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
   360
  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
   361
    by blast
42065
2b98b4c2e2f1 add ring_of_sets and subset_class as basis for algebra
hoelzl
parents: 41981
diff changeset
   362
  interpret l: algebra "M\<lparr>sets := lambda_system M f\<rparr>"
2b98b4c2e2f1 add ring_of_sets and subset_class as basis for algebra
hoelzl
parents: 41981
diff changeset
   363
    using f by (rule lambda_system_algebra)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   364
  have 4: "UNION {0..<n} A \<in> lambda_system M f"
42065
2b98b4c2e2f1 add ring_of_sets and subset_class as basis for algebra
hoelzl
parents: 41981
diff changeset
   365
    using A l.UNION_in_sets by simp
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   366
  from Suc.hyps show ?case
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   367
    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
   368
qed
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   369
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   370
lemma (in sigma_algebra) lambda_system_caratheodory:
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   371
  assumes oms: "outer_measure_space M f"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   372
      and A: "range A \<subseteq> lambda_system M f"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   373
      and disj: "disjoint_family A"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   374
  shows  "(\<Union>i. A i) \<in> lambda_system M f \<and> (\<Sum>i. f (A i)) = f (\<Union>i. A i)"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   375
proof -
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   376
  have pos: "positive M f" and inc: "increasing M f"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   377
   and csa: "countably_subadditive M f"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   378
    by (metis oms outer_measure_space_def)+
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   379
  have sa: "subadditive M f"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   380
    by (metis countably_subadditive_subadditive csa pos)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   381
  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
   382
    by simp
42065
2b98b4c2e2f1 add ring_of_sets and subset_class as basis for algebra
hoelzl
parents: 41981
diff changeset
   383
  interpret ls: algebra "M\<lparr>sets := lambda_system M f\<rparr>"
2b98b4c2e2f1 add ring_of_sets and subset_class as basis for algebra
hoelzl
parents: 41981
diff changeset
   384
    using pos by (rule lambda_system_algebra)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   385
  have A'': "range A \<subseteq> sets M"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   386
     by (metis A image_subset_iff lambda_system_sets)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   387
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   388
  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
   389
    by (metis A'' countable_UN)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   390
  have U_eq: "f (\<Union>i. A i) = (\<Sum>i. f (A i))"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   391
  proof (rule antisym)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   392
    show "f (\<Union>i. A i) \<le> (\<Sum>i. f (A i))"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   393
      using csa[unfolded countably_subadditive_def] A'' disj U_in by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   394
    have *: "\<And>i. 0 \<le> f (A i)" using pos A'' unfolding positive_def by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   395
    have dis: "\<And>N. disjoint_family_on A {..<N}" by (intro disjoint_family_on_mono[OF _ disj]) auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   396
    show "(\<Sum>i. f (A i)) \<le> f (\<Union>i. A i)"
42065
2b98b4c2e2f1 add ring_of_sets and subset_class as basis for algebra
hoelzl
parents: 41981
diff changeset
   397
      using ls.additive_sum [OF lambda_system_positive[OF pos] lambda_system_additive _ A' dis]
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   398
      using A''
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   399
      by (intro suminf_bound[OF _ *]) (auto intro!: increasingD[OF inc] allI countable_UN)
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   400
  qed
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   401
  {
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   402
    fix a
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   403
    assume a [iff]: "a \<in> sets M"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   404
    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
   405
    proof -
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   406
      show ?thesis
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   407
      proof (rule antisym)
33536
fd28b7399f2b eliminated hard tabulators;
wenzelm
parents: 33271
diff changeset
   408
        have "range (\<lambda>i. a \<inter> A i) \<subseteq> sets M" using A''
fd28b7399f2b eliminated hard tabulators;
wenzelm
parents: 33271
diff changeset
   409
          by blast
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   410
        moreover
33536
fd28b7399f2b eliminated hard tabulators;
wenzelm
parents: 33271
diff changeset
   411
        have "disjoint_family (\<lambda>i. a \<inter> A i)" using disj
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   412
          by (auto simp add: disjoint_family_on_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   413
        moreover
33536
fd28b7399f2b eliminated hard tabulators;
wenzelm
parents: 33271
diff changeset
   414
        have "a \<inter> (\<Union>i. A i) \<in> sets M"
fd28b7399f2b eliminated hard tabulators;
wenzelm
parents: 33271
diff changeset
   415
          by (metis Int U_in a)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   416
        ultimately
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   417
        have "f (a \<inter> (\<Union>i. A i)) \<le> (\<Sum>i. f (a \<inter> A i))"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   418
          using csa[unfolded countably_subadditive_def, rule_format, of "(\<lambda>i. a \<inter> A i)"]
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   419
          by (simp add: o_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   420
        hence "f (a \<inter> (\<Union>i. A i)) + f (a - (\<Union>i. A i)) \<le>
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   421
            (\<Sum>i. f (a \<inter> A i)) + f (a - (\<Union>i. A i))"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   422
          by (rule add_right_mono)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   423
        moreover
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   424
        have "(\<Sum>i. f (a \<inter> A i)) + f (a - (\<Union>i. A i)) \<le> f a"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   425
          proof (intro suminf_bound_add allI)
33536
fd28b7399f2b eliminated hard tabulators;
wenzelm
parents: 33271
diff changeset
   426
            fix n
fd28b7399f2b eliminated hard tabulators;
wenzelm
parents: 33271
diff changeset
   427
            have UNION_in: "(\<Union>i\<in>{0..<n}. A i) \<in> sets M"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   428
              by (metis A'' UNION_in_sets)
33536
fd28b7399f2b eliminated hard tabulators;
wenzelm
parents: 33271
diff changeset
   429
            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
   430
              by (blast intro: increasingD [OF inc] A'' UNION_in_sets)
33536
fd28b7399f2b eliminated hard tabulators;
wenzelm
parents: 33271
diff changeset
   431
            have ls: "(\<Union>i\<in>{0..<n}. A i) \<in> lambda_system M f"
42065
2b98b4c2e2f1 add ring_of_sets and subset_class as basis for algebra
hoelzl
parents: 41981
diff changeset
   432
              using ls.UNION_in_sets by (simp add: A)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   433
            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
   434
              by (simp add: lambda_system_eq UNION_in)
33536
fd28b7399f2b eliminated hard tabulators;
wenzelm
parents: 33271
diff changeset
   435
            have "f (a - (\<Union>i. A i)) \<le> f (a - (\<Union>i\<in>{0..<n}. A i))"
44106
0e018cbcc0de tuned proofs
haftmann
parents: 43920
diff changeset
   436
              by (blast intro: increasingD [OF inc] UNION_in U_in)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   437
            thus "(\<Sum>i<n. f (a \<inter> A i)) + f (a - (\<Union>i. A i)) \<le> f a"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   438
              by (simp add: lambda_system_strong_sum pos A disj eq_fa add_left_mono atLeast0LessThan[symmetric])
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   439
          next
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   440
            have "\<And>i. a \<inter> A i \<in> sets M" using A'' by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   441
            then show "\<And>i. 0 \<le> f (a \<inter> A i)" using pos[unfolded positive_def] by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   442
            have "\<And>i. a - (\<Union>i. A i) \<in> sets M" using A'' by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   443
            then have "\<And>i. 0 \<le> f (a - (\<Union>i. A i))" using pos[unfolded positive_def] by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   444
            then show "f (a - (\<Union>i. A i)) \<noteq> -\<infinity>" by auto
33536
fd28b7399f2b eliminated hard tabulators;
wenzelm
parents: 33271
diff changeset
   445
          qed
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   446
        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
   447
          by (rule order_trans)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   448
      next
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   449
        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
   450
          by (blast intro:  increasingD [OF inc] U_in)
33536
fd28b7399f2b eliminated hard tabulators;
wenzelm
parents: 33271
diff changeset
   451
        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
   452
          by (blast intro: subadditiveD [OF sa] U_in)
33536
fd28b7399f2b eliminated hard tabulators;
wenzelm
parents: 33271
diff changeset
   453
        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
   454
        qed
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   455
     qed
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   456
  }
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   457
  thus  ?thesis
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   458
    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
   459
qed
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   460
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   461
lemma (in sigma_algebra) caratheodory_lemma:
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   462
  assumes oms: "outer_measure_space M f"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   463
  shows "measure_space \<lparr> space = space M, sets = lambda_system M f, measure = f \<rparr>"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   464
    (is "measure_space ?M")
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   465
proof -
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   466
  have pos: "positive M f"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   467
    by (metis oms outer_measure_space_def)
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   468
  have alg: "algebra ?M"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   469
    using lambda_system_algebra [of f, OF pos]
42065
2b98b4c2e2f1 add ring_of_sets and subset_class as basis for algebra
hoelzl
parents: 41981
diff changeset
   470
    by (simp add: algebra_iff_Un)
2b98b4c2e2f1 add ring_of_sets and subset_class as basis for algebra
hoelzl
parents: 41981
diff changeset
   471
  then
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   472
  have "sigma_algebra ?M"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   473
    using lambda_system_caratheodory [OF oms]
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   474
    by (simp add: sigma_algebra_disjoint_iff)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   475
  moreover
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   476
  have "measure_space_axioms ?M"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   477
    using pos lambda_system_caratheodory [OF oms]
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   478
    by (simp add: measure_space_axioms_def positive_def lambda_system_sets
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   479
                  countably_additive_def o_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   480
  ultimately
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   481
  show ?thesis
42065
2b98b4c2e2f1 add ring_of_sets and subset_class as basis for algebra
hoelzl
parents: 41981
diff changeset
   482
    by (simp add: measure_space_def)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   483
qed
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   484
42066
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   485
lemma (in ring_of_sets) additive_increasing:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   486
  assumes posf: "positive M f" and addf: "additive M f"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   487
  shows "increasing M f"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   488
proof (auto simp add: increasing_def)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   489
  fix x y
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   490
  assume xy: "x \<in> sets M" "y \<in> sets M" "x \<subseteq> y"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   491
  then have "y - x \<in> sets M" by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   492
  then have "0 \<le> f (y-x)" using posf[unfolded positive_def] by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   493
  then have "f x + 0 \<le> f x + f (y-x)" by (intro add_left_mono) auto
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   494
  also have "... = f (x \<union> (y-x))" using addf
37032
58a0757031dd speed up some proofs and fix some warnings
huffman
parents: 36649
diff changeset
   495
    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
   496
  also have "... = f y"
37032
58a0757031dd speed up some proofs and fix some warnings
huffman
parents: 36649
diff changeset
   497
    by (metis Un_Diff_cancel Un_absorb1 xy(3))
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   498
  finally show "f x \<le> f y" by simp
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   499
qed
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   500
42066
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   501
lemma (in ring_of_sets) countably_additive_additive:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   502
  assumes posf: "positive M f" and ca: "countably_additive M f"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   503
  shows "additive M f"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   504
proof (auto simp add: additive_def)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   505
  fix x y
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   506
  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
   507
  hence "disjoint_family (binaryset x y)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   508
    by (auto simp add: disjoint_family_on_def binaryset_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   509
  hence "range (binaryset x y) \<subseteq> sets M \<longrightarrow>
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   510
         (\<Union>i. binaryset x y i) \<in> sets M \<longrightarrow>
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   511
         f (\<Union>i. binaryset x y i) = (\<Sum> n. f (binaryset x y n))"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   512
    using ca
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   513
    by (simp add: countably_additive_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   514
  hence "{x,y,{}} \<subseteq> sets M \<longrightarrow> x \<union> y \<in> sets M \<longrightarrow>
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   515
         f (x \<union> y) = (\<Sum>n. f (binaryset x y n))"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   516
    by (simp add: range_binaryset_eq UN_binaryset_eq)
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   517
  thus "f (x \<union> y) = f x + f y" using posf x y
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   518
    by (auto simp add: Un suminf_binaryset_eq positive_def)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   519
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   520
39096
hoelzl
parents: 38656
diff changeset
   521
lemma inf_measure_nonempty:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   522
  assumes f: "positive M f" and b: "b \<in> sets M" and a: "a \<subseteq> b" "{} \<in> sets M"
39096
hoelzl
parents: 38656
diff changeset
   523
  shows "f b \<in> measure_set M f a"
hoelzl
parents: 38656
diff changeset
   524
proof -
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   525
  let ?A = "\<lambda>i::nat. (if i = 0 then b else {})"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   526
  have "(\<Sum>i. f (?A i)) = (\<Sum>i<1::nat. f (?A i))"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   527
    by (rule suminf_finite) (simp add: f[unfolded positive_def])
39096
hoelzl
parents: 38656
diff changeset
   528
  also have "... = f b"
hoelzl
parents: 38656
diff changeset
   529
    by simp
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   530
  finally show ?thesis using assms
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   531
    by (auto intro!: exI [of _ ?A]
39096
hoelzl
parents: 38656
diff changeset
   532
             simp: measure_set_def disjoint_family_on_def split_if_mem2 comp_def)
hoelzl
parents: 38656
diff changeset
   533
qed
hoelzl
parents: 38656
diff changeset
   534
42066
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   535
lemma (in ring_of_sets) inf_measure_agrees:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   536
  assumes posf: "positive M f" and ca: "countably_additive M f"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   537
      and s: "s \<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 s) = f s"
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
   539
  unfolding Inf_ereal_def
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   540
proof (safe intro!: Greatest_equality)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   541
  fix z
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   542
  assume z: "z \<in> measure_set M f s"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   543
  from this obtain A where
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   544
    A: "range A \<subseteq> sets M" and disj: "disjoint_family A"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   545
    and "s \<subseteq> (\<Union>x. A x)" and si: "(\<Sum>i. f (A i)) = z"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   546
    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
   547
  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
   548
  have inc: "increasing M f"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   549
    by (metis additive_increasing ca countably_additive_additive posf)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   550
  have sums: "(\<Sum>i. f (A i \<inter> s)) = f (\<Union>i. A i \<inter> s)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   551
    proof (rule ca[unfolded countably_additive_def, rule_format])
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   552
      show "range (\<lambda>n. A n \<inter> s) \<subseteq> sets M" using A s
33536
fd28b7399f2b eliminated hard tabulators;
wenzelm
parents: 33271
diff changeset
   553
        by blast
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   554
      show "disjoint_family (\<lambda>n. A n \<inter> s)" using disj
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33536
diff changeset
   555
        by (auto simp add: disjoint_family_on_def)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   556
      show "(\<Union>i. A i \<inter> s) \<in> sets M" using A s
33536
fd28b7399f2b eliminated hard tabulators;
wenzelm
parents: 33271
diff changeset
   557
        by (metis UN_extend_simps(4) s seq)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   558
    qed
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   559
  hence "f s = (\<Sum>i. f (A i \<inter> s))"
37032
58a0757031dd speed up some proofs and fix some warnings
huffman
parents: 36649
diff changeset
   560
    using seq [symmetric] by (simp add: sums_iff)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   561
  also have "... \<le> (\<Sum>i. f (A i))"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   562
    proof (rule suminf_le_pos)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   563
      fix n show "f (A n \<inter> s) \<le> f (A n)" using A s
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   564
        by (force intro: increasingD [OF inc])
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   565
      fix N have "A N \<inter> s \<in> sets M"  using A s by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   566
      then show "0 \<le> f (A N \<inter> s)" using posf unfolding positive_def by auto
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   567
    qed
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   568
  also have "... = z" by (rule si)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   569
  finally show "f s \<le> z" .
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   570
next
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   571
  fix y
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   572
  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
   573
  thus "y \<le> f s"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   574
    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
   575
qed
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   576
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   577
lemma measure_set_pos:
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   578
  assumes posf: "positive M f" "r \<in> measure_set M f X"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   579
  shows "0 \<le> r"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   580
proof -
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   581
  obtain A where "range A \<subseteq> sets M" and r: "r = (\<Sum>i. f (A i))"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   582
    using `r \<in> measure_set M f X` unfolding measure_set_def by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   583
  then show "0 \<le> r" using posf unfolding r positive_def
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   584
    by (intro suminf_0_le) auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   585
qed
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   586
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   587
lemma inf_measure_pos:
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   588
  assumes posf: "positive M f"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   589
  shows "0 \<le> Inf (measure_set M f X)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   590
proof (rule complete_lattice_class.Inf_greatest)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   591
  fix r assume "r \<in> measure_set M f X" with posf show "0 \<le> r"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   592
    by (rule measure_set_pos)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   593
qed
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   594
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   595
lemma inf_measure_empty:
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   596
  assumes posf: "positive M f" and "{} \<in> sets M"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   597
  shows "Inf (measure_set M f {}) = 0"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   598
proof (rule antisym)
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   599
  show "Inf (measure_set M f {}) \<le> 0"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   600
    by (metis complete_lattice_class.Inf_lower `{} \<in> sets M`
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   601
              inf_measure_nonempty[OF posf] subset_refl posf[unfolded positive_def])
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   602
qed (rule inf_measure_pos[OF posf])
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   603
42066
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   604
lemma (in ring_of_sets) inf_measure_positive:
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   605
  assumes p: "positive M f" and "{} \<in> sets M"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   606
  shows "positive M (\<lambda>x. Inf (measure_set M f x))"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   607
proof (unfold positive_def, intro conjI ballI)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   608
  show "Inf (measure_set M f {}) = 0" using inf_measure_empty[OF assms] by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   609
  fix A assume "A \<in> sets M"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   610
qed (rule inf_measure_pos[OF p])
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   611
42066
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   612
lemma (in ring_of_sets) inf_measure_increasing:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   613
  assumes posf: "positive M f"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   614
  shows "increasing \<lparr> space = space M, sets = Pow (space M) \<rparr>
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   615
                    (\<lambda>x. Inf (measure_set M f x))"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   616
apply (auto simp add: increasing_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   617
apply (rule complete_lattice_class.Inf_greatest)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   618
apply (rule complete_lattice_class.Inf_lower)
37032
58a0757031dd speed up some proofs and fix some warnings
huffman
parents: 36649
diff changeset
   619
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
   620
done
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   621
42066
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   622
lemma (in ring_of_sets) inf_measure_le:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   623
  assumes posf: "positive M f" and inc: "increasing M f"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   624
      and x: "x \<in> {r . \<exists>A. range A \<subseteq> sets M \<and> s \<subseteq> (\<Union>i. A i) \<and> (\<Sum>i. f (A i)) = r}"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   625
  shows "Inf (measure_set M f s) \<le> x"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   626
proof -
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   627
  obtain A where A: "range A \<subseteq> sets M" and ss: "s \<subseteq> (\<Union>i. A i)"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   628
             and xeq: "(\<Sum>i. f (A i)) = x"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   629
    using x by auto
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   630
  have dA: "range (disjointed A) \<subseteq> sets M"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   631
    by (metis A range_disjointed_sets)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   632
  have "\<forall>n. f (disjointed A n) \<le> f (A n)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   633
    by (metis increasingD [OF inc] UNIV_I dA image_subset_iff disjointed_subset A comp_def)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   634
  moreover have "\<forall>i. 0 \<le> f (disjointed A i)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   635
    using posf dA unfolding positive_def by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   636
  ultimately have sda: "(\<Sum>i. f (disjointed A i)) \<le> (\<Sum>i. f (A i))"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   637
    by (blast intro!: suminf_le_pos)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   638
  hence ley: "(\<Sum>i. f (disjointed A i)) \<le> x"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   639
    by (metis xeq)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   640
  hence y: "(\<Sum>i. f (disjointed A i)) \<in> measure_set M f s"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   641
    apply (auto simp add: measure_set_def)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   642
    apply (rule_tac x="disjointed A" in exI)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   643
    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
   644
    done
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   645
  show ?thesis
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   646
    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
   647
qed
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   648
42066
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   649
lemma (in ring_of_sets) inf_measure_close:
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
   650
  fixes e :: ereal
42066
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   651
  assumes posf: "positive M f" and e: "0 < e" and ss: "s \<subseteq> (space M)" and "Inf (measure_set M f s) \<noteq> \<infinity>"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   652
  shows "\<exists>A. range A \<subseteq> sets M \<and> disjoint_family A \<and> s \<subseteq> (\<Union>i. A i) \<and>
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   653
               (\<Sum>i. f (A i)) \<le> Inf (measure_set M f s) + e"
42066
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   654
proof -
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   655
  from `Inf (measure_set M f s) \<noteq> \<infinity>` have fin: "\<bar>Inf (measure_set M f s)\<bar> \<noteq> \<infinity>"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   656
    using inf_measure_pos[OF posf, of s] by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   657
  obtain l where "l \<in> measure_set M f s" "l \<le> Inf (measure_set M f s) + e"
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
   658
    using Inf_ereal_close[OF fin e] by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   659
  thus ?thesis
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   660
    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
   661
qed
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   662
42066
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   663
lemma (in ring_of_sets) inf_measure_countably_subadditive:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   664
  assumes posf: "positive M f" and inc: "increasing M f"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   665
  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
   666
                  (\<lambda>x. Inf (measure_set M f x))"
42066
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   667
proof (simp add: countably_subadditive_def, safe)
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   668
  fix A :: "nat \<Rightarrow> 'a set"
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   669
  let "?outer B" = "Inf (measure_set M f B)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   670
  assume A: "range A \<subseteq> Pow (space M)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   671
     and disj: "disjoint_family A"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   672
     and sb: "(\<Union>i. A i) \<subseteq> space M"
42066
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   673
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
   674
  { fix e :: ereal assume e: "0 < e" and "\<forall>i. ?outer (A i) \<noteq> \<infinity>"
42066
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   675
    hence "\<exists>BB. \<forall>n. range (BB n) \<subseteq> sets M \<and> disjoint_family (BB n) \<and>
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   676
        A n \<subseteq> (\<Union>i. BB n i) \<and> (\<Sum>i. f (BB n i)) \<le> ?outer (A n) + e * (1/2)^(Suc n)"
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   677
      apply (safe intro!: choice inf_measure_close [of f, OF posf])
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
   678
      using e sb by (auto simp: ereal_zero_less_0_iff one_ereal_def)
42066
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   679
    then obtain BB
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   680
      where BB: "\<And>n. (range (BB n) \<subseteq> sets M)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   681
      and disjBB: "\<And>n. disjoint_family (BB n)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   682
      and sbBB: "\<And>n. A n \<subseteq> (\<Union>i. BB n i)"
42066
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   683
      and BBle: "\<And>n. (\<Sum>i. f (BB n i)) \<le> ?outer (A n) + e * (1/2)^(Suc n)"
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   684
      by auto blast
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   685
    have sll: "(\<Sum>n. \<Sum>i. (f (BB n i))) \<le> (\<Sum>n. ?outer (A n)) + e"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   686
    proof -
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   687
      have sum_eq_1: "(\<Sum>n. e*(1/2) ^ Suc n) = e"
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
   688
        using suminf_half_series_ereal e
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
   689
        by (simp add: ereal_zero_le_0_iff zero_le_divide_ereal suminf_cmult_ereal)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   690
      have "\<And>n i. 0 \<le> f (BB n i)" using posf[unfolded positive_def] BB by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   691
      then have "\<And>n. 0 \<le> (\<Sum>i. f (BB n i))" by (rule suminf_0_le)
42066
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   692
      then have "(\<Sum>n. \<Sum>i. (f (BB n i))) \<le> (\<Sum>n. ?outer (A n) + e*(1/2) ^ Suc n)"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   693
        by (rule suminf_le_pos[OF BBle])
42066
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   694
      also have "... = (\<Sum>n. ?outer (A n)) + e"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   695
        using sum_eq_1 inf_measure_pos[OF posf] e
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
   696
        by (subst suminf_add_ereal) (auto simp add: ereal_zero_le_0_iff)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   697
      finally show ?thesis .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   698
    qed
42066
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   699
    def C \<equiv> "(split BB) o prod_decode"
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   700
    have C: "!!n. C n \<in> sets M"
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   701
      apply (rule_tac p="prod_decode n" in PairE)
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   702
      apply (simp add: C_def)
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   703
      apply (metis BB subsetD rangeI)
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   704
      done
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   705
    have sbC: "(\<Union>i. A i) \<subseteq> (\<Union>i. C i)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   706
    proof (auto simp add: C_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   707
      fix x i
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   708
      assume x: "x \<in> A i"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   709
      with sbBB [of i] obtain j where "x \<in> BB i j"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   710
        by blast
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   711
      thus "\<exists>i. x \<in> split BB (prod_decode i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   712
        by (metis prod_encode_inverse prod.cases)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   713
    qed
42066
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   714
    have "(f \<circ> C) = (f \<circ> (\<lambda>(x, y). BB x y)) \<circ> prod_decode"
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   715
      by (rule ext)  (auto simp add: C_def)
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   716
    moreover have "suminf ... = (\<Sum>n. \<Sum>i. f (BB n i))" using BBle
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   717
      using BB posf[unfolded positive_def]
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
   718
      by (force intro!: suminf_ereal_2dimen simp: o_def)
42066
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   719
    ultimately have Csums: "(\<Sum>i. f (C i)) = (\<Sum>n. \<Sum>i. f (BB n i))" by (simp add: o_def)
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   720
    have "?outer (\<Union>i. A i) \<le> (\<Sum>n. \<Sum>i. f (BB n i))"
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   721
      apply (rule inf_measure_le [OF posf(1) inc], auto)
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   722
      apply (rule_tac x="C" in exI)
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   723
      apply (auto simp add: C sbC Csums)
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   724
      done
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   725
    also have "... \<le> (\<Sum>n. ?outer (A n)) + e" using sll
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   726
      by blast
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   727
    finally have "?outer (\<Union>i. A i) \<le> (\<Sum>n. ?outer (A n)) + e" . }
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   728
  note for_finite_Inf = this
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   729
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   730
  show "?outer (\<Union>i. A i) \<le> (\<Sum>n. ?outer (A n))"
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   731
  proof cases
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   732
    assume "\<forall>i. ?outer (A i) \<noteq> \<infinity>"
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   733
    with for_finite_Inf show ?thesis
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
   734
      by (intro ereal_le_epsilon) auto
42066
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   735
  next
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   736
    assume "\<not> (\<forall>i. ?outer (A i) \<noteq> \<infinity>)"
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   737
    then have "\<exists>i. ?outer (A i) = \<infinity>"
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   738
      by auto
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   739
    then have "(\<Sum>n. ?outer (A n)) = \<infinity>"
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   740
      using suminf_PInfty[OF inf_measure_pos, OF posf]
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   741
      by metis
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   742
    then show ?thesis by simp
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   743
  qed
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   744
qed
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   745
42066
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   746
lemma (in ring_of_sets) inf_measure_outer:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   747
  "\<lbrakk> positive M f ; increasing M f \<rbrakk>
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   748
   \<Longrightarrow> outer_measure_space \<lparr> space = space M, sets = Pow (space M) \<rparr>
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   749
                          (\<lambda>x. Inf (measure_set M f x))"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   750
  using inf_measure_pos[of M f]
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   751
  by (simp add: outer_measure_space_def inf_measure_empty
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   752
                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
   753
42066
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   754
lemma (in ring_of_sets) algebra_subset_lambda_system:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   755
  assumes posf: "positive M f" and inc: "increasing M f"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   756
      and add: "additive M f"
42066
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   757
  shows "sets M \<subseteq> lambda_system \<lparr> space = space M, sets = Pow (space M) \<rparr>
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   758
                                (\<lambda>x. Inf (measure_set M f x))"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   759
proof (auto dest: sets_into_space
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   760
            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
   761
  fix x s
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   762
  assume x: "x \<in> sets M"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   763
     and s: "s \<subseteq> space M"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   764
  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
   765
    by blast
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   766
  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
   767
        \<le> Inf (measure_set M f s)"
42066
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   768
  proof cases
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   769
    assume "Inf (measure_set M f s) = \<infinity>" then show ?thesis by simp
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   770
  next
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   771
    assume fin: "Inf (measure_set M f s) \<noteq> \<infinity>"
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   772
    then have "measure_set M f s \<noteq> {}"
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
   773
      by (auto simp: top_ereal_def)
42066
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   774
    show ?thesis
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   775
    proof (rule complete_lattice_class.Inf_greatest)
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   776
      fix r assume "r \<in> measure_set M f s"
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   777
      then obtain A where A: "disjoint_family A" "range A \<subseteq> sets M" "s \<subseteq> (\<Union>i. A i)"
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   778
        and r: "r = (\<Sum>i. f (A i))" unfolding measure_set_def by auto
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   779
      have "Inf (measure_set M f (s \<inter> x)) \<le> (\<Sum>i. f (A i \<inter> x))"
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   780
        unfolding measure_set_def
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   781
      proof (safe intro!: complete_lattice_class.Inf_lower exI[of _ "\<lambda>i. A i \<inter> x"])
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   782
        from A(1) show "disjoint_family (\<lambda>i. A i \<inter> x)"
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   783
          by (rule disjoint_family_on_bisimulation) auto
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   784
      qed (insert x A, auto)
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   785
      moreover
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   786
      have "Inf (measure_set M f (s - x)) \<le> (\<Sum>i. f (A i - x))"
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   787
        unfolding measure_set_def
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   788
      proof (safe intro!: complete_lattice_class.Inf_lower exI[of _ "\<lambda>i. A i - x"])
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   789
        from A(1) show "disjoint_family (\<lambda>i. A i - x)"
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   790
          by (rule disjoint_family_on_bisimulation) auto
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   791
      qed (insert x A, auto)
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   792
      ultimately have "Inf (measure_set M f (s \<inter> x)) + Inf (measure_set M f (s - x)) \<le>
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   793
          (\<Sum>i. f (A i \<inter> x)) + (\<Sum>i. f (A i - x))" by (rule add_mono)
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   794
      also have "\<dots> = (\<Sum>i. f (A i \<inter> x) + f (A i - x))"
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
   795
        using A(2) x posf by (subst suminf_add_ereal) (auto simp: positive_def)
42066
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   796
      also have "\<dots> = (\<Sum>i. f (A i))"
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   797
        using A x
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   798
        by (subst add[THEN additiveD, symmetric])
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   799
           (auto intro!: arg_cong[where f=suminf] arg_cong[where f=f])
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   800
      finally show "Inf (measure_set M f (s \<inter> x)) + Inf (measure_set M f (s - x)) \<le> r"
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   801
        using r by simp
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   802
    qed
42066
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   803
  qed
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   804
  moreover
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   805
  have "Inf (measure_set M f s)
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   806
       \<le> Inf (measure_set M f (s\<inter>x)) + Inf (measure_set M f (s-x))"
42145
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   807
  proof -
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   808
    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
   809
      by (metis Un_Diff_Int Un_commute)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   810
    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
   811
      apply (rule subadditiveD)
42145
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   812
      apply (rule ring_of_sets.countably_subadditive_subadditive [OF ring_of_sets_Pow])
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   813
      apply (simp add: positive_def inf_measure_empty[OF posf] inf_measure_pos[OF posf])
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   814
      apply (rule inf_measure_countably_subadditive)
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   815
      using s by (auto intro!: posf inc)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   816
    finally show ?thesis .
42145
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   817
  qed
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   818
  ultimately
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   819
  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
   820
        = Inf (measure_set M f s)"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   821
    by (rule order_antisym)
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   822
qed
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   823
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   824
lemma measure_down:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   825
  "measure_space N \<Longrightarrow> sigma_algebra M \<Longrightarrow> sets M \<subseteq> sets N \<Longrightarrow> measure N = measure M \<Longrightarrow> measure_space M"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   826
  by (simp add: measure_space_def measure_space_axioms_def positive_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   827
                countably_additive_def)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   828
     blast
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   829
42066
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   830
theorem (in ring_of_sets) caratheodory:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   831
  assumes posf: "positive M f" and ca: "countably_additive M f"
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
   832
  shows "\<exists>\<mu> :: 'a set \<Rightarrow> ereal. (\<forall>s \<in> sets M. \<mu> s = f s) \<and>
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   833
            measure_space \<lparr> space = space M, sets = sets (sigma M), measure = \<mu> \<rparr>"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   834
proof -
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   835
  have inc: "increasing M f"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   836
    by (metis additive_increasing ca countably_additive_additive posf)
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   837
  let ?infm = "(\<lambda>x. Inf (measure_set M f x))"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   838
  def ls \<equiv> "lambda_system (|space = space M, sets = Pow (space M)|) ?infm"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   839
  have mls: "measure_space \<lparr>space = space M, sets = ls, measure = ?infm\<rparr>"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   840
    using sigma_algebra.caratheodory_lemma
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   841
            [OF sigma_algebra_Pow  inf_measure_outer [OF posf inc]]
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   842
    by (simp add: ls_def)
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   843
  hence sls: "sigma_algebra (|space = space M, sets = ls, measure = ?infm|)"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   844
    by (simp add: measure_space_def)
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   845
  have "sets M \<subseteq> ls"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   846
    by (simp add: ls_def)
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   847
       (metis ca posf inc countably_additive_additive algebra_subset_lambda_system)
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   848
  hence sgs_sb: "sigma_sets (space M) (sets M) \<subseteq> ls"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   849
    using sigma_algebra.sigma_sets_subset [OF sls, of "sets M"]
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   850
    by simp
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   851
  have "measure_space \<lparr> space = space M, sets = sets (sigma M), measure = ?infm \<rparr>"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   852
    unfolding sigma_def
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   853
    by (rule measure_down [OF mls], rule sigma_algebra_sigma_sets)
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   854
       (simp_all add: sgs_sb space_closed)
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   855
  thus ?thesis using inf_measure_agrees [OF posf ca]
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   856
    by (intro exI[of _ ?infm]) auto
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41023
diff changeset
   857
qed
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   858
42145
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   859
subsubsection {*Alternative instances of caratheodory*}
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   860
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   861
lemma (in ring_of_sets) countably_additive_iff_continuous_from_below:
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   862
  assumes f: "positive M f" "additive M f"
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   863
  shows "countably_additive M f \<longleftrightarrow>
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   864
    (\<forall>A. range A \<subseteq> sets M \<longrightarrow> incseq A \<longrightarrow> (\<Union>i. A i) \<in> sets M \<longrightarrow> (\<lambda>i. f (A i)) ----> f (\<Union>i. A i))"
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   865
  unfolding countably_additive_def
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   866
proof safe
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   867
  assume count_sum: "\<forall>A. range A \<subseteq> sets M \<longrightarrow> disjoint_family A \<longrightarrow> UNION UNIV A \<in> sets M \<longrightarrow> (\<Sum>i. f (A i)) = f (UNION UNIV A)"
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   868
  fix A :: "nat \<Rightarrow> 'a set" assume A: "range A \<subseteq> sets M" "incseq A" "(\<Union>i. A i) \<in> sets M"
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   869
  then have dA: "range (disjointed A) \<subseteq> sets M" by (auto simp: range_disjointed_sets)
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   870
  with count_sum[THEN spec, of "disjointed A"] A(3)
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   871
  have f_UN: "(\<Sum>i. f (disjointed A i)) = f (\<Union>i. A i)"
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   872
    by (auto simp: UN_disjointed_eq disjoint_family_disjointed)
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   873
  moreover have "(\<lambda>n. (\<Sum>i=0..<n. f (disjointed A i))) ----> (\<Sum>i. f (disjointed A i))"
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   874
    using f(1)[unfolded positive_def] dA
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
   875
    by (auto intro!: summable_sumr_LIMSEQ_suminf summable_ereal_pos)
42145
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   876
  from LIMSEQ_Suc[OF this]
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   877
  have "(\<lambda>n. (\<Sum>i\<le>n. f (disjointed A i))) ----> (\<Sum>i. f (disjointed A i))"
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   878
    unfolding atLeastLessThanSuc_atLeastAtMost atLeast0AtMost .
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   879
  moreover have "\<And>n. (\<Sum>i\<le>n. f (disjointed A i)) = f (A n)"
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   880
    using disjointed_additive[OF f A(1,2)] .
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   881
  ultimately show "(\<lambda>i. f (A i)) ----> f (\<Union>i. A i)" by simp
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   882
next
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   883
  assume cont: "\<forall>A. range A \<subseteq> sets M \<longrightarrow> incseq A \<longrightarrow> (\<Union>i. A i) \<in> sets M \<longrightarrow> (\<lambda>i. f (A i)) ----> f (\<Union>i. A i)"
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   884
  fix A :: "nat \<Rightarrow> 'a set" assume A: "range A \<subseteq> sets M" "disjoint_family A" "(\<Union>i. A i) \<in> sets M"
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   885
  have *: "(\<Union>n. (\<Union>i\<le>n. A i)) = (\<Union>i. A i)" by auto
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   886
  have "(\<lambda>n. f (\<Union>i\<le>n. A i)) ----> f (\<Union>i. A i)"
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   887
  proof (unfold *[symmetric], intro cont[rule_format])
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   888
    show "range (\<lambda>i. \<Union> i\<le>i. A i) \<subseteq> sets M" "(\<Union>i. \<Union> i\<le>i. A i) \<in> sets M"
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   889
      using A * by auto
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   890
  qed (force intro!: incseq_SucI)
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   891
  moreover have "\<And>n. f (\<Union>i\<le>n. A i) = (\<Sum>i\<le>n. f (A i))"
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   892
    using A
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   893
    by (intro additive_sum[OF f, of _ A, symmetric])
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   894
       (auto intro: disjoint_family_on_mono[where B=UNIV])
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   895
  ultimately
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   896
  have "(\<lambda>i. f (A i)) sums f (\<Union>i. A i)"
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   897
    unfolding sums_def2 by simp
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   898
  from sums_unique[OF this]
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   899
  show "(\<Sum>i. f (A i)) = f (\<Union>i. A i)" by simp
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   900
qed
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   901
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   902
lemma (in ring_of_sets) continuous_from_above_iff_empty_continuous:
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   903
  assumes f: "positive M f" "additive M f"
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   904
  shows "(\<forall>A. range A \<subseteq> sets M \<longrightarrow> decseq A \<longrightarrow> (\<Inter>i. A i) \<in> sets M \<longrightarrow> (\<forall>i. f (A i) \<noteq> \<infinity>) \<longrightarrow> (\<lambda>i. f (A i)) ----> f (\<Inter>i. A i))
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   905
     \<longleftrightarrow> (\<forall>A. range A \<subseteq> sets M \<longrightarrow> decseq A \<longrightarrow> (\<Inter>i. A i) = {} \<longrightarrow> (\<forall>i. f (A i) \<noteq> \<infinity>) \<longrightarrow> (\<lambda>i. f (A i)) ----> 0)"
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   906
proof safe
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   907
  assume cont: "(\<forall>A. range A \<subseteq> sets M \<longrightarrow> decseq A \<longrightarrow> (\<Inter>i. A i) \<in> sets M \<longrightarrow> (\<forall>i. f (A i) \<noteq> \<infinity>) \<longrightarrow> (\<lambda>i. f (A i)) ----> f (\<Inter>i. A i))"
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   908
  fix A :: "nat \<Rightarrow> 'a set" assume A: "range A \<subseteq> sets M" "decseq A" "(\<Inter>i. A i) = {}" "\<forall>i. f (A i) \<noteq> \<infinity>"
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   909
  with cont[THEN spec, of A] show "(\<lambda>i. f (A i)) ----> 0"
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   910
    using `positive M f`[unfolded positive_def] by auto
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   911
next
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   912
  assume cont: "\<forall>A. range A \<subseteq> sets M \<longrightarrow> decseq A \<longrightarrow> (\<Inter>i. A i) = {} \<longrightarrow> (\<forall>i. f (A i) \<noteq> \<infinity>) \<longrightarrow> (\<lambda>i. f (A i)) ----> 0"
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   913
  fix A :: "nat \<Rightarrow> 'a set" assume A: "range A \<subseteq> sets M" "decseq A" "(\<Inter>i. A i) \<in> sets M" "\<forall>i. f (A i) \<noteq> \<infinity>"
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   914
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   915
  have f_mono: "\<And>a b. a \<in> sets M \<Longrightarrow> b \<in> sets M \<Longrightarrow> a \<subseteq> b \<Longrightarrow> f a \<le> f b"
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   916
    using additive_increasing[OF f] unfolding increasing_def by simp
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   917
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   918
  have decseq_fA: "decseq (\<lambda>i. f (A i))"
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   919
    using A by (auto simp: decseq_def intro!: f_mono)
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   920
  have decseq: "decseq (\<lambda>i. A i - (\<Inter>i. A i))"
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   921
    using A by (auto simp: decseq_def)
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   922
  then have decseq_f: "decseq (\<lambda>i. f (A i - (\<Inter>i. A i)))"
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   923
    using A unfolding decseq_def by (auto intro!: f_mono Diff)
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   924
  have "f (\<Inter>x. A x) \<le> f (A 0)"
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   925
    using A by (auto intro!: f_mono)
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   926
  then have f_Int_fin: "f (\<Inter>x. A x) \<noteq> \<infinity>"
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   927
    using A by auto
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   928
  { fix i
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   929
    have "f (A i - (\<Inter>i. A i)) \<le> f (A i)" using A by (auto intro!: f_mono)
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   930
    then have "f (A i - (\<Inter>i. A i)) \<noteq> \<infinity>"
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   931
      using A by auto }
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   932
  note f_fin = this
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   933
  have "(\<lambda>i. f (A i - (\<Inter>i. A i))) ----> 0"
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   934
  proof (intro cont[rule_format, OF _ decseq _ f_fin])
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   935
    show "range (\<lambda>i. A i - (\<Inter>i. A i)) \<subseteq> sets M" "(\<Inter>i. A i - (\<Inter>i. A i)) = {}"
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   936
      using A by auto
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   937
  qed
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
   938
  from INF_Lim_ereal[OF decseq_f this]
42145
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   939
  have "(INF n. f (A n - (\<Inter>i. A i))) = 0" .
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   940
  moreover have "(INF n. f (\<Inter>i. A i)) = f (\<Inter>i. A i)"
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   941
    by auto
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   942
  ultimately have "(INF n. f (A n - (\<Inter>i. A i)) + f (\<Inter>i. A i)) = 0 + f (\<Inter>i. A i)"
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   943
    using A(4) f_fin f_Int_fin
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
   944
    by (subst INFI_ereal_add) (auto simp: decseq_f)
42145
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   945
  moreover {
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   946
    fix n
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   947
    have "f (A n - (\<Inter>i. A i)) + f (\<Inter>i. A i) = f ((A n - (\<Inter>i. A i)) \<union> (\<Inter>i. A i))"
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   948
      using A by (subst f(2)[THEN additiveD]) auto
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   949
    also have "(A n - (\<Inter>i. A i)) \<union> (\<Inter>i. A i) = A n"
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   950
      by auto
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   951
    finally have "f (A n - (\<Inter>i. A i)) + f (\<Inter>i. A i) = f (A n)" . }
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   952
  ultimately have "(INF n. f (A n)) = f (\<Inter>i. A i)"
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   953
    by simp
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
   954
  with LIMSEQ_ereal_INFI[OF decseq_fA]
42145
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   955
  show "(\<lambda>i. f (A i)) ----> f (\<Inter>i. A i)" by simp
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   956
qed
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   957
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   958
lemma positiveD1: "positive M f \<Longrightarrow> f {} = 0" by (auto simp: positive_def)
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   959
lemma positiveD2: "positive M f \<Longrightarrow> A \<in> sets M \<Longrightarrow> 0 \<le> f A" by (auto simp: positive_def)
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   960
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   961
lemma (in ring_of_sets) empty_continuous_imp_continuous_from_below:
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   962
  assumes f: "positive M f" "additive M f" "\<forall>A\<in>sets M. f A \<noteq> \<infinity>"
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   963
  assumes cont: "\<forall>A. range A \<subseteq> sets M \<longrightarrow> decseq A \<longrightarrow> (\<Inter>i. A i) = {} \<longrightarrow> (\<lambda>i. f (A i)) ----> 0"
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   964
  assumes A: "range A \<subseteq> sets M" "incseq A" "(\<Union>i. A i) \<in> sets M"
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   965
  shows "(\<lambda>i. f (A i)) ----> f (\<Union>i. A i)"
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   966
proof -
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
   967
  have "\<forall>A\<in>sets M. \<exists>x. f A = ereal x"
42145
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   968
  proof
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
   969
    fix A assume "A \<in> sets M" with f show "\<exists>x. f A = ereal x"
42145
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   970
      unfolding positive_def by (cases "f A") auto
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   971
  qed
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   972
  from bchoice[OF this] guess f' .. note f' = this[rule_format]
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   973
  from A have "(\<lambda>i. f ((\<Union>i. A i) - A i)) ----> 0"
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   974
    by (intro cont[rule_format]) (auto simp: decseq_def incseq_def)
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   975
  moreover
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   976
  { fix i
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   977
    have "f ((\<Union>i. A i) - A i) + f (A i) = f ((\<Union>i. A i) - A i \<union> A i)"
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   978
      using A by (intro f(2)[THEN additiveD, symmetric]) auto
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   979
    also have "(\<Union>i. A i) - A i \<union> A i = (\<Union>i. A i)"
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   980
      by auto
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   981
    finally have "f' (\<Union>i. A i) - f' (A i) = f' ((\<Union>i. A i) - A i)"
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   982
      using A by (subst (asm) (1 2 3) f') auto
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
   983
    then have "f ((\<Union>i. A i) - A i) = ereal (f' (\<Union>i. A i) - f' (A i))"
42145
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   984
      using A f' by auto }
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   985
  ultimately have "(\<lambda>i. f' (\<Union>i. A i) - f' (A i)) ----> 0"
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
   986
    by (simp add: zero_ereal_def)
42145
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   987
  then have "(\<lambda>i. f' (A i)) ----> f' (\<Union>i. A i)"
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   988
    by (rule LIMSEQ_diff_approach_zero2[OF LIMSEQ_const])
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   989
  then show "(\<lambda>i. f (A i)) ----> f (\<Union>i. A i)"
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   990
    using A by (subst (1 2) f') auto
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   991
qed
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   992
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   993
lemma (in ring_of_sets) empty_continuous_imp_countably_additive:
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   994
  assumes f: "positive M f" "additive M f" and fin: "\<forall>A\<in>sets M. f A \<noteq> \<infinity>"
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   995
  assumes cont: "\<And>A. range A \<subseteq> sets M \<Longrightarrow> decseq A \<Longrightarrow> (\<Inter>i. A i) = {} \<Longrightarrow> (\<lambda>i. f (A i)) ----> 0"
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   996
  shows "countably_additive M f"
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   997
  using countably_additive_iff_continuous_from_below[OF f]
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   998
  using empty_continuous_imp_continuous_from_below[OF f fin] cont
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   999
  by blast
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
  1000
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
  1001
lemma (in ring_of_sets) caratheodory_empty_continuous:
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
  1002
  assumes f: "positive M f" "additive M f" and fin: "\<And>A. A \<in> sets M \<Longrightarrow> f A \<noteq> \<infinity>"
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
  1003
  assumes cont: "\<And>A. range A \<subseteq> sets M \<Longrightarrow> decseq A \<Longrightarrow> (\<Inter>i. A i) = {} \<Longrightarrow> (\<lambda>i. f (A i)) ----> 0"
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
  1004
  shows "\<exists>\<mu> :: 'a set \<Rightarrow> ereal. (\<forall>s \<in> sets M. \<mu> s = f s) \<and>
42145
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
  1005
            measure_space \<lparr> space = space M, sets = sets (sigma M), measure = \<mu> \<rparr>"
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
  1006
proof (intro caratheodory empty_continuous_imp_countably_additive f)
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
  1007
  show "\<forall>A\<in>sets M. f A \<noteq> \<infinity>" using fin by auto
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
  1008
qed (rule cont)
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
  1009
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
  1010
end