src/HOL/Probability/Caratheodory.thy
author hoelzl
Mon, 30 Jun 2014 15:45:21 +0200
changeset 57447 87429bdecad5
parent 57446 06e195515deb
child 58876 1888e3cb8048
permissions -rw-r--r--
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
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
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
     9
  imports Measure_Space
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
42067
66c8281349ec standardized headers
hoelzl
parents: 42066
diff changeset
    12
text {*
66c8281349ec standardized headers
hoelzl
parents: 42066
diff changeset
    13
  Originally from the Hurd/Coble measure theory development, translated by Lawrence Paulson.
66c8281349ec standardized headers
hoelzl
parents: 42066
diff changeset
    14
*}
66c8281349ec standardized headers
hoelzl
parents: 42066
diff changeset
    15
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
    16
lemma suminf_ereal_2dimen:
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
    17
  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
    18
  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
    19
  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
    20
  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
    21
proof -
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
    22
  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
    23
    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
    24
  have reindex: "\<And>B. (\<Sum>x\<in>B. f (prod_decode x)) = setsum f (prod_decode ` B)"
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56994
diff changeset
    25
    by (simp add: setsum.reindex[OF inj_prod_decode] comp_def)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
    26
  { fix n
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
    27
    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
    28
    { 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
    29
      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
    30
        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
    31
    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
    32
      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
    33
    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
    34
  moreover
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
    35
  { fix a b
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
    36
    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
    37
    { 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
    38
        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
    39
    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
    40
      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
    41
  ultimately
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
    42
  show ?thesis unfolding g_def using pos
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56994
diff changeset
    43
    by (auto intro!: SUP_eq  simp: setsum.cartesian_product reindex SUP_upper2
56212
3253aaf73a01 consolidated theorem names containing INFI and SUPR: have INF and SUP instead uniformly
haftmann
parents: 56166
diff changeset
    44
                     setsum_nonneg suminf_ereal_eq_SUP SUP_pair
3253aaf73a01 consolidated theorem names containing INFI and SUPR: have INF and SUP instead uniformly
haftmann
parents: 56166
diff changeset
    45
                     SUP_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
    46
qed
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
    47
56994
8d5e5ec1cac3 fixed document generation for HOL-Probability
hoelzl
parents: 56212
diff changeset
    48
subsection {* Characterizations of Measures *}
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    49
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
    50
definition subadditive where "subadditive M f \<longleftrightarrow>
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
    51
  (\<forall>x\<in>M. \<forall>y\<in>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
    52
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
    53
definition countably_subadditive where "countably_subadditive M f \<longleftrightarrow>
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
    54
  (\<forall>A. range A \<subseteq> M \<longrightarrow> disjoint_family A \<longrightarrow> (\<Union>i. A i) \<in> M \<longrightarrow>
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
    55
    (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
    56
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
    57
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
    58
  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
    59
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
    60
definition measure_set where "measure_set M f X = {r.
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
    61
  \<exists>A. range A \<subseteq> 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
    62
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    63
lemma subadditiveD:
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
    64
  "subadditive M f \<Longrightarrow> x \<inter> y = {} \<Longrightarrow> x \<in> M \<Longrightarrow> y \<in> M \<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
    65
  by (auto simp add: subadditive_def)
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    66
56994
8d5e5ec1cac3 fixed document generation for HOL-Probability
hoelzl
parents: 56212
diff changeset
    67
subsubsection {* Lambda Systems *}
8d5e5ec1cac3 fixed document generation for HOL-Probability
hoelzl
parents: 56212
diff changeset
    68
8d5e5ec1cac3 fixed document generation for HOL-Probability
hoelzl
parents: 56212
diff changeset
    69
definition lambda_system where "lambda_system \<Omega> M f = {l \<in> M.
8d5e5ec1cac3 fixed document generation for HOL-Probability
hoelzl
parents: 56212
diff changeset
    70
  \<forall>x \<in> M. f (l \<inter> x) + f ((\<Omega> - l) \<inter> x) = f x}"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    71
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    72
lemma (in algebra) lambda_system_eq:
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
    73
  shows "lambda_system \<Omega> M f = {l \<in> M. \<forall>x \<in> 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
    74
proof -
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
    75
  have [simp]: "!!l x. l \<in> M \<Longrightarrow> x \<in> M \<Longrightarrow> (\<Omega> - l) \<inter> x = x - l"
37032
58a0757031dd speed up some proofs and fix some warnings
huffman
parents: 36649
diff changeset
    76
    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
    77
  show ?thesis
37032
58a0757031dd speed up some proofs and fix some warnings
huffman
parents: 36649
diff changeset
    78
    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
    79
qed
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    80
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    81
lemma (in algebra) lambda_system_empty:
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
    82
  "positive M f \<Longrightarrow> {} \<in> lambda_system \<Omega> M f"
42066
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
    83
  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
    84
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    85
lemma lambda_system_sets:
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
    86
  "x \<in> lambda_system \<Omega> M f \<Longrightarrow> x \<in> 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
    87
  by (simp add: lambda_system_def)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    88
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    89
lemma (in algebra) lambda_system_Compl:
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
    90
  fixes f:: "'a set \<Rightarrow> ereal"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
    91
  assumes x: "x \<in> lambda_system \<Omega> M f"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
    92
  shows "\<Omega> - x \<in> lambda_system \<Omega> 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
    93
proof -
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
    94
  have "x \<subseteq> \<Omega>"
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
    95
    by (metis sets_into_space lambda_system_sets x)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
    96
  hence "\<Omega> - (\<Omega> - x) = x"
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
    97
    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
    98
  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
    99
    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
   100
qed
33271
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 (in algebra) lambda_system_Int:
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
   103
  fixes f:: "'a set \<Rightarrow> ereal"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   104
  assumes xl: "x \<in> lambda_system \<Omega> M f" and yl: "y \<in> lambda_system \<Omega> M f"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   105
  shows "x \<inter> y \<in> lambda_system \<Omega> 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
   106
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
   107
  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
   108
  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
   109
    fix u
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   110
    assume x: "x \<in> M" and y: "y \<in> M" and u: "u \<in> M"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   111
       and fx: "\<forall>z\<in>M. f (z \<inter> x) + f (z - x) = f z"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   112
       and fy: "\<forall>z\<in>M. f (z \<inter> y) + f (z - y) = f z"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   113
    have "u - x \<inter> y \<in> 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
   114
      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
   115
    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
   116
    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
   117
    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
   118
    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
   119
    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
   120
    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
   121
      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
   122
    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
   123
          = (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
   124
      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
   125
    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
   126
      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
   127
    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
   128
      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
   129
      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
   130
    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
   131
      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
   132
    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
   133
  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
   134
qed
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   135
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   136
lemma (in algebra) lambda_system_Un:
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
   137
  fixes f:: "'a set \<Rightarrow> ereal"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   138
  assumes xl: "x \<in> lambda_system \<Omega> M f" and yl: "y \<in> lambda_system \<Omega> M f"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   139
  shows "x \<union> y \<in> lambda_system \<Omega> M f"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   140
proof -
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   141
  have "(\<Omega> - x) \<inter> (\<Omega> - y) \<in> M"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   142
    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
   143
  moreover
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   144
  have "x \<union> y = \<Omega> - ((\<Omega> - x) \<inter> (\<Omega> - y))"
46731
5302e932d1e5 avoid undeclared variables in let bindings;
wenzelm
parents: 44928
diff changeset
   145
    by auto (metis subsetD lambda_system_sets sets_into_space xl yl)+
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   146
  ultimately show ?thesis
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   147
    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
   148
qed
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   149
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   150
lemma (in algebra) lambda_system_algebra:
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   151
  "positive M f \<Longrightarrow> algebra \<Omega> (lambda_system \<Omega> M f)"
42065
2b98b4c2e2f1 add ring_of_sets and subset_class as basis for algebra
hoelzl
parents: 41981
diff changeset
   152
  apply (auto simp add: algebra_iff_Un)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   153
  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
   154
  apply (metis lambda_system_empty)
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   155
  apply (metis lambda_system_Compl)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   156
  apply (metis lambda_system_Un)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   157
  done
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   158
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   159
lemma (in algebra) lambda_system_strong_additive:
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   160
  assumes z: "z \<in> M" and disj: "x \<inter> y = {}"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   161
      and xl: "x \<in> lambda_system \<Omega> M f" and yl: "y \<in> lambda_system \<Omega> M f"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   162
  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
   163
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
   164
  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
   165
  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
   166
  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
   167
  moreover
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   168
  have "(z \<inter> (x \<union> y)) \<in> 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
   169
    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
   170
  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
   171
    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
   172
qed
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   173
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   174
lemma (in algebra) lambda_system_additive: "additive (lambda_system \<Omega> 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
   175
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
   176
  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
   177
  assume disj: "x \<inter> y = {}"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   178
     and xl: "x \<in> lambda_system \<Omega> M f" and yl: "y \<in> lambda_system \<Omega> M f"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   179
  hence  "x \<in> M" "y \<in> M" by (blast intro: 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
   180
  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
   181
    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
   182
    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
   183
qed
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   184
42145
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   185
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
   186
  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
   187
  shows  "subadditive M f"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33536
diff changeset
   188
proof (auto simp add: subadditive_def)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   189
  fix x y
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   190
  assume x: "x \<in> M" and y: "y \<in> M" and "x \<inter> y = {}"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   191
  hence "disjoint_family (binaryset x y)"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33536
diff changeset
   192
    by (auto simp add: disjoint_family_on_def binaryset_def)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   193
  hence "range (binaryset x y) \<subseteq> M \<longrightarrow>
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   194
         (\<Union>i. binaryset x y i) \<in> M \<longrightarrow>
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   195
         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
   196
    using cs by (auto simp add: countably_subadditive_def)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   197
  hence "{x,y,{}} \<subseteq> M \<longrightarrow> x \<union> y \<in> M \<longrightarrow>
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   198
         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
   199
    by (simp add: range_binaryset_eq UN_binaryset_eq)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   200
  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
   201
    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
   202
qed
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   203
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   204
lemma lambda_system_increasing:
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   205
 "increasing M f \<Longrightarrow> increasing (lambda_system \<Omega> M f) f"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   206
  by (simp add: increasing_def lambda_system_def)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   207
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
   208
lemma lambda_system_positive:
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   209
  "positive M f \<Longrightarrow> positive (lambda_system \<Omega> 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
   210
  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
   211
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   212
lemma (in algebra) lambda_system_strong_sum:
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
   213
  fixes A:: "nat \<Rightarrow> 'a set" and f :: "'a set \<Rightarrow> ereal"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   214
  assumes f: "positive M f" and a: "a \<in> M"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   215
      and A: "range A \<subseteq> lambda_system \<Omega> M f"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   216
      and disj: "disjoint_family A"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   217
  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
   218
proof (induct n)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   219
  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
   220
next
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   221
  case (Suc n)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   222
  have 2: "A n \<inter> UNION {0..<n} A = {}" using disj
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   223
    by (force simp add: disjoint_family_on_def neq_iff)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   224
  have 3: "A n \<in> lambda_system \<Omega> M f" using A
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   225
    by blast
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   226
  interpret l: algebra \<Omega> "lambda_system \<Omega> M f"
42065
2b98b4c2e2f1 add ring_of_sets and subset_class as basis for algebra
hoelzl
parents: 41981
diff changeset
   227
    using f by (rule lambda_system_algebra)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   228
  have 4: "UNION {0..<n} A \<in> lambda_system \<Omega> M f"
42065
2b98b4c2e2f1 add ring_of_sets and subset_class as basis for algebra
hoelzl
parents: 41981
diff changeset
   229
    using A l.UNION_in_sets by simp
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   230
  from Suc.hyps show ?case
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   231
    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
   232
qed
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   233
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   234
lemma (in sigma_algebra) lambda_system_caratheodory:
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   235
  assumes oms: "outer_measure_space M f"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   236
      and A: "range A \<subseteq> lambda_system \<Omega> M f"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   237
      and disj: "disjoint_family A"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   238
  shows  "(\<Union>i. A i) \<in> lambda_system \<Omega> 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
   239
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
   240
  have pos: "positive M f" and inc: "increasing M f"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   241
   and csa: "countably_subadditive M f"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   242
    by (metis oms outer_measure_space_def)+
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   243
  have sa: "subadditive M f"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   244
    by (metis countably_subadditive_subadditive csa pos)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   245
  have A': "\<And>S. A`S \<subseteq> (lambda_system \<Omega> M f)" using A
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   246
    by auto
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   247
  interpret ls: algebra \<Omega> "lambda_system \<Omega> M f"
42065
2b98b4c2e2f1 add ring_of_sets and subset_class as basis for algebra
hoelzl
parents: 41981
diff changeset
   248
    using pos by (rule lambda_system_algebra)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   249
  have A'': "range A \<subseteq> M"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   250
     by (metis A image_subset_iff lambda_system_sets)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   251
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   252
  have U_in: "(\<Union>i. A i) \<in> M"
37032
58a0757031dd speed up some proofs and fix some warnings
huffman
parents: 36649
diff changeset
   253
    by (metis A'' countable_UN)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   254
  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
   255
  proof (rule antisym)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   256
    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
   257
      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
   258
    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
   259
    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
   260
    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
   261
      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
   262
      using A''
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   263
      by (intro suminf_bound[OF _ *]) (auto intro!: increasingD[OF inc] 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
   264
  qed
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   265
  {
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   266
    fix a
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   267
    assume a [iff]: "a \<in> M"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   268
    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
   269
    proof -
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   270
      show ?thesis
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   271
      proof (rule antisym)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   272
        have "range (\<lambda>i. a \<inter> A i) \<subseteq> M" using A''
33536
fd28b7399f2b eliminated hard tabulators;
wenzelm
parents: 33271
diff changeset
   273
          by blast
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   274
        moreover
33536
fd28b7399f2b eliminated hard tabulators;
wenzelm
parents: 33271
diff changeset
   275
        have "disjoint_family (\<lambda>i. a \<inter> A i)" using disj
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   276
          by (auto simp add: disjoint_family_on_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   277
        moreover
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   278
        have "a \<inter> (\<Union>i. A i) \<in> M"
33536
fd28b7399f2b eliminated hard tabulators;
wenzelm
parents: 33271
diff changeset
   279
          by (metis Int U_in a)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   280
        ultimately
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   281
        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
   282
          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
   283
          by (simp add: o_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   284
        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
   285
            (\<Sum>i. f (a \<inter> A i)) + f (a - (\<Union>i. A i))"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   286
          by (rule add_right_mono)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   287
        moreover
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   288
        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
   289
          proof (intro suminf_bound_add allI)
33536
fd28b7399f2b eliminated hard tabulators;
wenzelm
parents: 33271
diff changeset
   290
            fix n
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   291
            have UNION_in: "(\<Union>i\<in>{0..<n}. A i) \<in> M"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   292
              by (metis A'' UNION_in_sets)
33536
fd28b7399f2b eliminated hard tabulators;
wenzelm
parents: 33271
diff changeset
   293
            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
   294
              by (blast intro: increasingD [OF inc] A'' UNION_in_sets)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   295
            have ls: "(\<Union>i\<in>{0..<n}. A i) \<in> lambda_system \<Omega> M f"
42065
2b98b4c2e2f1 add ring_of_sets and subset_class as basis for algebra
hoelzl
parents: 41981
diff changeset
   296
              using ls.UNION_in_sets by (simp add: A)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   297
            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
   298
              by (simp add: lambda_system_eq UNION_in)
33536
fd28b7399f2b eliminated hard tabulators;
wenzelm
parents: 33271
diff changeset
   299
            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
   300
              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
   301
            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
   302
              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
   303
          next
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   304
            have "\<And>i. a \<inter> A i \<in> M" using A'' by auto
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   305
            then show "\<And>i. 0 \<le> f (a \<inter> A i)" using pos[unfolded positive_def] by auto
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   306
            have "\<And>i. a - (\<Union>i. A i) \<in> M" using A'' by auto
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   307
            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
   308
            then show "f (a - (\<Union>i. A i)) \<noteq> -\<infinity>" by auto
33536
fd28b7399f2b eliminated hard tabulators;
wenzelm
parents: 33271
diff changeset
   309
          qed
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   310
        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
   311
          by (rule order_trans)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   312
      next
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   313
        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
   314
          by (blast intro:  increasingD [OF inc] U_in)
33536
fd28b7399f2b eliminated hard tabulators;
wenzelm
parents: 33271
diff changeset
   315
        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
   316
          by (blast intro: subadditiveD [OF sa] U_in)
33536
fd28b7399f2b eliminated hard tabulators;
wenzelm
parents: 33271
diff changeset
   317
        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
   318
        qed
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   319
     qed
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   320
  }
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   321
  thus  ?thesis
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   322
    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
   323
qed
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   324
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   325
lemma (in sigma_algebra) caratheodory_lemma:
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   326
  assumes oms: "outer_measure_space M f"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   327
  defines "L \<equiv> lambda_system \<Omega> M f"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   328
  shows "measure_space \<Omega> L f"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   329
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
   330
  have pos: "positive M f"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   331
    by (metis oms outer_measure_space_def)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   332
  have alg: "algebra \<Omega> L"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   333
    using lambda_system_algebra [of f, OF pos]
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   334
    by (simp add: algebra_iff_Un L_def)
42065
2b98b4c2e2f1 add ring_of_sets and subset_class as basis for algebra
hoelzl
parents: 41981
diff changeset
   335
  then
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   336
  have "sigma_algebra \<Omega> L"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   337
    using lambda_system_caratheodory [OF oms]
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   338
    by (simp add: sigma_algebra_disjoint_iff L_def)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   339
  moreover
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   340
  have "countably_additive L f" "positive L f"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   341
    using pos lambda_system_caratheodory [OF oms]
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   342
    by (auto simp add: lambda_system_sets L_def countably_additive_def positive_def)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   343
  ultimately
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   344
  show ?thesis
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   345
    using pos by (simp add: measure_space_def)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   346
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   347
39096
hoelzl
parents: 38656
diff changeset
   348
lemma inf_measure_nonempty:
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   349
  assumes f: "positive M f" and b: "b \<in> M" and a: "a \<subseteq> b" "{} \<in> M"
39096
hoelzl
parents: 38656
diff changeset
   350
  shows "f b \<in> measure_set M f a"
hoelzl
parents: 38656
diff changeset
   351
proof -
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   352
  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
   353
  have "(\<Sum>i. f (?A i)) = (\<Sum>i<1::nat. f (?A i))"
47761
dfe747e72fa8 moved lemmas to appropriate places
hoelzl
parents: 47694
diff changeset
   354
    by (rule suminf_finite) (simp_all add: f[unfolded positive_def])
39096
hoelzl
parents: 38656
diff changeset
   355
  also have "... = f b"
hoelzl
parents: 38656
diff changeset
   356
    by simp
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   357
  finally show ?thesis using assms
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   358
    by (auto intro!: exI [of _ ?A]
39096
hoelzl
parents: 38656
diff changeset
   359
             simp: measure_set_def disjoint_family_on_def split_if_mem2 comp_def)
hoelzl
parents: 38656
diff changeset
   360
qed
hoelzl
parents: 38656
diff changeset
   361
42066
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   362
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
   363
  assumes posf: "positive M f" and ca: "countably_additive M f"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   364
      and s: "s \<in> M"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   365
  shows "Inf (measure_set M f s) = f s"
51329
4a3c453f99a1 split dense into inner_dense_order and no_top/no_bot
hoelzl
parents: 49773
diff changeset
   366
proof (intro Inf_eqI)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   367
  fix z
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   368
  assume z: "z \<in> measure_set M f s"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   369
  from this obtain A where
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   370
    A: "range A \<subseteq> 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
   371
    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
   372
    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
   373
  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
   374
  have inc: "increasing M f"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   375
    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
   376
  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
   377
    proof (rule ca[unfolded countably_additive_def, rule_format])
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   378
      show "range (\<lambda>n. A n \<inter> s) \<subseteq> M" using A s
33536
fd28b7399f2b eliminated hard tabulators;
wenzelm
parents: 33271
diff changeset
   379
        by blast
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   380
      show "disjoint_family (\<lambda>n. A n \<inter> s)" using disj
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33536
diff changeset
   381
        by (auto simp add: disjoint_family_on_def)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   382
      show "(\<Union>i. A i \<inter> s) \<in> M" using A s
33536
fd28b7399f2b eliminated hard tabulators;
wenzelm
parents: 33271
diff changeset
   383
        by (metis UN_extend_simps(4) s seq)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   384
    qed
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   385
  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
   386
    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
   387
  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
   388
    proof (rule suminf_le_pos)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   389
      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
   390
        by (force intro: increasingD [OF inc])
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   391
      fix N have "A N \<inter> s \<in> M"  using A s by auto
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   392
      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
   393
    qed
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   394
  also have "... = z" by (rule si)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   395
  finally show "f s \<le> z" .
51329
4a3c453f99a1 split dense into inner_dense_order and no_top/no_bot
hoelzl
parents: 49773
diff changeset
   396
qed (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
   397
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   398
lemma measure_set_pos:
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   399
  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
   400
  shows "0 \<le> r"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   401
proof -
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   402
  obtain A where "range A \<subseteq> M" and r: "r = (\<Sum>i. f (A i))"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   403
    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
   404
  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
   405
    by (intro suminf_0_le) auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   406
qed
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   407
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   408
lemma inf_measure_pos:
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   409
  assumes posf: "positive M f"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   410
  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
   411
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
   412
  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
   413
    by (rule measure_set_pos)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   414
qed
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   415
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
   416
lemma inf_measure_empty:
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   417
  assumes posf: "positive M f" and "{} \<in> M"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   418
  shows "Inf (measure_set M f {}) = 0"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   419
proof (rule antisym)
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   420
  show "Inf (measure_set M f {}) \<le> 0"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   421
    by (metis complete_lattice_class.Inf_lower `{} \<in> 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
   422
              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
   423
qed (rule inf_measure_pos[OF posf])
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   424
42066
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   425
lemma (in ring_of_sets) inf_measure_positive:
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   426
  assumes p: "positive M f" and "{} \<in> M"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   427
  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
   428
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
   429
  show "Inf (measure_set M f {}) = 0" using inf_measure_empty[OF assms] by auto
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   430
  fix A assume "A \<in> M"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   431
qed (rule inf_measure_pos[OF p])
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   432
42066
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   433
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
   434
  assumes posf: "positive M f"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   435
  shows "increasing (Pow \<Omega>) (\<lambda>x. Inf (measure_set M f x))"
44918
6a80fbc4e72c tune simpset for Complete_Lattices
noschinl
parents: 44568
diff changeset
   436
apply (clarsimp simp add: increasing_def)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   437
apply (rule complete_lattice_class.Inf_greatest)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   438
apply (rule complete_lattice_class.Inf_lower)
37032
58a0757031dd speed up some proofs and fix some warnings
huffman
parents: 36649
diff changeset
   439
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
   440
done
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   441
42066
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   442
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
   443
  assumes posf: "positive M f" and inc: "increasing M f"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   444
      and x: "x \<in> {r . \<exists>A. range A \<subseteq> 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
   445
  shows "Inf (measure_set M f s) \<le> x"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   446
proof -
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   447
  obtain A where A: "range A \<subseteq> 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
   448
             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
   449
    using x by auto
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   450
  have dA: "range (disjointed A) \<subseteq> M"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   451
    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
   452
  have "\<forall>n. f (disjointed A n) \<le> f (A n)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   453
    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
   454
  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
   455
    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
   456
  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
   457
    by (blast intro!: suminf_le_pos)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   458
  hence ley: "(\<Sum>i. f (disjointed A i)) \<le> x"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   459
    by (metis xeq)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   460
  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
   461
    apply (auto simp add: measure_set_def)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   462
    apply (rule_tac x="disjointed A" in exI)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   463
    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
   464
    done
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   465
  show ?thesis
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   466
    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
   467
qed
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   468
42066
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   469
lemma (in ring_of_sets) inf_measure_close:
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
   470
  fixes e :: ereal
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   471
  assumes posf: "positive M f" and e: "0 < e" and ss: "s \<subseteq> (\<Omega>)" and "Inf (measure_set M f s) \<noteq> \<infinity>"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   472
  shows "\<exists>A. range A \<subseteq> 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
   473
               (\<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
   474
proof -
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   475
  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
   476
    using inf_measure_pos[OF posf, of s] by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   477
  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
   478
    using Inf_ereal_close[OF fin e] by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   479
  thus ?thesis
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   480
    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
   481
qed
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   482
42066
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   483
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
   484
  assumes posf: "positive M f" and inc: "increasing M f"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   485
  shows "countably_subadditive (Pow \<Omega>) (\<lambda>x. Inf (measure_set M f x))"
42066
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   486
proof (simp add: countably_subadditive_def, safe)
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   487
  fix A :: "nat \<Rightarrow> 'a set"
46731
5302e932d1e5 avoid undeclared variables in let bindings;
wenzelm
parents: 44928
diff changeset
   488
  let ?outer = "\<lambda>B. Inf (measure_set M f B)"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   489
  assume A: "range A \<subseteq> Pow (\<Omega>)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   490
     and disj: "disjoint_family A"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   491
     and sb: "(\<Union>i. A i) \<subseteq> \<Omega>"
42066
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   492
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
   493
  { fix e :: ereal assume e: "0 < e" and "\<forall>i. ?outer (A i) \<noteq> \<infinity>"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   494
    hence "\<exists>BB. \<forall>n. range (BB n) \<subseteq> M \<and> disjoint_family (BB n) \<and>
42066
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   495
        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
   496
      apply (safe intro!: choice inf_measure_close [of f, OF posf])
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
   497
      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
   498
    then obtain BB
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   499
      where BB: "\<And>n. (range (BB n) \<subseteq> M)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   500
      and disjBB: "\<And>n. disjoint_family (BB n)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   501
      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
   502
      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
   503
      by auto blast
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   504
    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
   505
    proof -
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   506
      have sum_eq_1: "(\<Sum>n. e*(1/2) ^ Suc n) = e"
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
   507
        using suminf_half_series_ereal e
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
   508
        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
   509
      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
   510
      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
   511
      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
   512
        by (rule suminf_le_pos[OF BBle])
42066
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   513
      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
   514
        using sum_eq_1 inf_measure_pos[OF posf] e
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
   515
        by (subst suminf_add_ereal) (auto simp add: ereal_zero_le_0_iff)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   516
      finally show ?thesis .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   517
    qed
42066
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   518
    def C \<equiv> "(split BB) o prod_decode"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   519
    have C: "!!n. C n \<in> M"
42066
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   520
      apply (rule_tac p="prod_decode n" in PairE)
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   521
      apply (simp add: C_def)
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   522
      apply (metis BB subsetD rangeI)
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   523
      done
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   524
    have sbC: "(\<Union>i. A i) \<subseteq> (\<Union>i. C i)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   525
    proof (auto simp add: C_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   526
      fix x i
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   527
      assume x: "x \<in> A i"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   528
      with sbBB [of i] obtain j where "x \<in> BB i j"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   529
        by blast
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   530
      thus "\<exists>i. x \<in> split BB (prod_decode i)"
55642
63beb38e9258 adapted to renaming of datatype 'cases' and 'recs' to 'case' and 'rec'
blanchet
parents: 52141
diff changeset
   531
        by (metis prod_encode_inverse prod.case)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   532
    qed
42066
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   533
    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
   534
      by (rule ext)  (auto simp add: C_def)
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   535
    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
   536
      using BB posf[unfolded positive_def]
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
   537
      by (force intro!: suminf_ereal_2dimen simp: o_def)
42066
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   538
    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
   539
    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
   540
      apply (rule inf_measure_le [OF posf(1) inc], auto)
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   541
      apply (rule_tac x="C" in exI)
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   542
      apply (auto simp add: C sbC Csums)
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   543
      done
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   544
    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
   545
      by blast
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   546
    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
   547
  note for_finite_Inf = this
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   548
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   549
  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
   550
  proof cases
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   551
    assume "\<forall>i. ?outer (A i) \<noteq> \<infinity>"
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   552
    with for_finite_Inf show ?thesis
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
   553
      by (intro ereal_le_epsilon) auto
42066
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   554
  next
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   555
    assume "\<not> (\<forall>i. ?outer (A i) \<noteq> \<infinity>)"
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   556
    then have "\<exists>i. ?outer (A i) = \<infinity>"
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   557
      by auto
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   558
    then have "(\<Sum>n. ?outer (A n)) = \<infinity>"
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   559
      using suminf_PInfty[OF inf_measure_pos, OF posf]
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   560
      by metis
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   561
    then show ?thesis by simp
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   562
  qed
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   563
qed
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   564
42066
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   565
lemma (in ring_of_sets) inf_measure_outer:
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   566
  "\<lbrakk> positive M f ; increasing M f \<rbrakk> \<Longrightarrow>
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   567
    outer_measure_space (Pow \<Omega>) (\<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
   568
  using inf_measure_pos[of M f]
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   569
  by (simp add: outer_measure_space_def inf_measure_empty
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   570
                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
   571
42066
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   572
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
   573
  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
   574
      and add: "additive M f"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   575
  shows "M \<subseteq> lambda_system \<Omega> (Pow \<Omega>) (\<lambda>x. Inf (measure_set M f x))"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   576
proof (auto dest: sets_into_space
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   577
            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
   578
  fix x s
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   579
  assume x: "x \<in> M"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   580
     and s: "s \<subseteq> \<Omega>"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   581
  have [simp]: "!!x. x \<in> M \<Longrightarrow> s \<inter> (\<Omega> - x) = s-x" using s
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   582
    by blast
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   583
  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
   584
        \<le> Inf (measure_set M f s)"
42066
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   585
  proof cases
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   586
    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
   587
  next
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   588
    assume fin: "Inf (measure_set M f s) \<noteq> \<infinity>"
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   589
    then have "measure_set M f s \<noteq> {}"
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
   590
      by (auto simp: top_ereal_def)
42066
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   591
    show ?thesis
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   592
    proof (rule complete_lattice_class.Inf_greatest)
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   593
      fix r assume "r \<in> measure_set M f s"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   594
      then obtain A where A: "disjoint_family A" "range A \<subseteq> M" "s \<subseteq> (\<Union>i. A i)"
42066
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   595
        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
   596
      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
   597
        unfolding measure_set_def
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   598
      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
   599
        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
   600
          by (rule disjoint_family_on_bisimulation) auto
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   601
      qed (insert x A, auto)
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   602
      moreover
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   603
      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
   604
        unfolding measure_set_def
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   605
      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
   606
        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
   607
          by (rule disjoint_family_on_bisimulation) auto
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   608
      qed (insert x A, auto)
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   609
      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
   610
          (\<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
   611
      also have "\<dots> = (\<Sum>i. f (A i \<inter> x) + f (A i - x))"
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
   612
        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
   613
      also have "\<dots> = (\<Sum>i. f (A i))"
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   614
        using A x
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   615
        by (subst add[THEN additiveD, symmetric])
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   616
           (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
   617
      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
   618
        using r by simp
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   619
    qed
42066
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   620
  qed
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   621
  moreover
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   622
  have "Inf (measure_set M f s)
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   623
       \<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
   624
  proof -
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   625
    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
   626
      by (metis Un_Diff_Int Un_commute)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   627
    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
   628
      apply (rule subadditiveD)
42145
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   629
      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
   630
      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
   631
      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
   632
      using s by (auto intro!: posf inc)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   633
    finally show ?thesis .
42145
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   634
  qed
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   635
  ultimately
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   636
  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
   637
        = Inf (measure_set M f s)"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   638
    by (rule order_antisym)
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   639
qed
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   640
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   641
lemma measure_down:
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   642
  "measure_space \<Omega> N \<mu> \<Longrightarrow> sigma_algebra \<Omega> M \<Longrightarrow> M \<subseteq> N \<Longrightarrow> measure_space \<Omega> M \<mu>"
57446
06e195515deb some lemmas about the indicator function; removed lemma sums_def2
hoelzl
parents: 57418
diff changeset
   643
  by (auto simp add: measure_space_def positive_def countably_additive_def subset_eq)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   644
56994
8d5e5ec1cac3 fixed document generation for HOL-Probability
hoelzl
parents: 56212
diff changeset
   645
subsection {* Caratheodory's theorem *}
8d5e5ec1cac3 fixed document generation for HOL-Probability
hoelzl
parents: 56212
diff changeset
   646
47762
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   647
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
   648
  assumes posf: "positive M f" and ca: "countably_additive M f"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   649
  shows "\<exists>\<mu> :: 'a set \<Rightarrow> ereal. (\<forall>s \<in> M. \<mu> s = f s) \<and> measure_space \<Omega> (sigma_sets \<Omega> M) \<mu>"
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
   650
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
   651
  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
   652
    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
   653
  let ?infm = "(\<lambda>x. Inf (measure_set M f x))"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   654
  def ls \<equiv> "lambda_system \<Omega> (Pow \<Omega>) ?infm"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   655
  have mls: "measure_space \<Omega> ls ?infm"
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
   656
    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
   657
            [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
   658
    by (simp add: ls_def)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   659
  hence sls: "sigma_algebra \<Omega> ls"
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
   660
    by (simp add: measure_space_def)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   661
  have "M \<subseteq> ls"
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
   662
    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
   663
       (metis ca posf inc countably_additive_additive algebra_subset_lambda_system)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   664
  hence sgs_sb: "sigma_sets (\<Omega>) (M) \<subseteq> ls"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   665
    using sigma_algebra.sigma_sets_subset [OF sls, of "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
   666
    by simp
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   667
  have "measure_space \<Omega> (sigma_sets \<Omega> M) ?infm"
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
   668
    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
   669
       (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
   670
  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
   671
    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
   672
qed
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   673
42145
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   674
lemma (in ring_of_sets) caratheodory_empty_continuous:
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   675
  assumes f: "positive M f" "additive M f" and fin: "\<And>A. A \<in> M \<Longrightarrow> f A \<noteq> \<infinity>"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   676
  assumes cont: "\<And>A. range A \<subseteq> M \<Longrightarrow> decseq A \<Longrightarrow> (\<Inter>i. A i) = {} \<Longrightarrow> (\<lambda>i. f (A i)) ----> 0"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   677
  shows "\<exists>\<mu> :: 'a set \<Rightarrow> ereal. (\<forall>s \<in> M. \<mu> s = f s) \<and> measure_space \<Omega> (sigma_sets \<Omega> M) \<mu>"
47762
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   678
proof (intro caratheodory' empty_continuous_imp_countably_additive f)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   679
  show "\<forall>A\<in>M. f A \<noteq> \<infinity>" using fin by auto
42145
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   680
qed (rule cont)
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   681
56994
8d5e5ec1cac3 fixed document generation for HOL-Probability
hoelzl
parents: 56212
diff changeset
   682
subsection {* Volumes *}
47762
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   683
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   684
definition volume :: "'a set set \<Rightarrow> ('a set \<Rightarrow> ereal) \<Rightarrow> bool" where
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   685
  "volume M f \<longleftrightarrow>
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   686
  (f {} = 0) \<and> (\<forall>a\<in>M. 0 \<le> f a) \<and>
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   687
  (\<forall>C\<subseteq>M. disjoint C \<longrightarrow> finite C \<longrightarrow> \<Union>C \<in> M \<longrightarrow> f (\<Union>C) = (\<Sum>c\<in>C. f c))"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   688
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   689
lemma volumeI:
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   690
  assumes "f {} = 0"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   691
  assumes "\<And>a. a \<in> M \<Longrightarrow> 0 \<le> f a"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   692
  assumes "\<And>C. C \<subseteq> M \<Longrightarrow> disjoint C \<Longrightarrow> finite C \<Longrightarrow> \<Union>C \<in> M \<Longrightarrow> f (\<Union>C) = (\<Sum>c\<in>C. f c)"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   693
  shows "volume M f"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   694
  using assms by (auto simp: volume_def)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   695
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   696
lemma volume_positive:
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   697
  "volume M f \<Longrightarrow> a \<in> M \<Longrightarrow> 0 \<le> f a"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   698
  by (auto simp: volume_def)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   699
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   700
lemma volume_empty:
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   701
  "volume M f \<Longrightarrow> f {} = 0"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   702
  by (auto simp: volume_def)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   703
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   704
lemma volume_finite_additive:
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   705
  assumes "volume M f" 
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   706
  assumes A: "\<And>i. i \<in> I \<Longrightarrow> A i \<in> M" "disjoint_family_on A I" "finite I" "UNION I A \<in> M"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   707
  shows "f (UNION I A) = (\<Sum>i\<in>I. f (A i))"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   708
proof -
52141
eff000cab70f weaker precendence of syntax for big intersection and union on sets
haftmann
parents: 51329
diff changeset
   709
  have "A`I \<subseteq> M" "disjoint (A`I)" "finite (A`I)" "\<Union>(A`I) \<in> M"
47762
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   710
    using A unfolding SUP_def by (auto simp: disjoint_family_on_disjoint_image)
52141
eff000cab70f weaker precendence of syntax for big intersection and union on sets
haftmann
parents: 51329
diff changeset
   711
  with `volume M f` have "f (\<Union>(A`I)) = (\<Sum>a\<in>A`I. f a)"
47762
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   712
    unfolding volume_def by blast
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   713
  also have "\<dots> = (\<Sum>i\<in>I. f (A i))"
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56994
diff changeset
   714
  proof (subst setsum.reindex_nontrivial)
47762
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   715
    fix i j assume "i \<in> I" "j \<in> I" "i \<noteq> j" "A i = A j"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   716
    with `disjoint_family_on A I` have "A i = {}"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   717
      by (auto simp: disjoint_family_on_def)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   718
    then show "f (A i) = 0"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   719
      using volume_empty[OF `volume M f`] by simp
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   720
  qed (auto intro: `finite I`)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   721
  finally show "f (UNION I A) = (\<Sum>i\<in>I. f (A i))"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   722
    by simp
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   723
qed
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   724
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   725
lemma (in ring_of_sets) volume_additiveI:
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   726
  assumes pos: "\<And>a. a \<in> M \<Longrightarrow> 0 \<le> \<mu> a" 
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   727
  assumes [simp]: "\<mu> {} = 0"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   728
  assumes add: "\<And>a b. a \<in> M \<Longrightarrow> b \<in> M \<Longrightarrow> a \<inter> b = {} \<Longrightarrow> \<mu> (a \<union> b) = \<mu> a + \<mu> b"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   729
  shows "volume M \<mu>"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   730
proof (unfold volume_def, safe)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   731
  fix C assume "finite C" "C \<subseteq> M" "disjoint C"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   732
  then show "\<mu> (\<Union>C) = setsum \<mu> C"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   733
  proof (induct C)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   734
    case (insert c C)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   735
    from insert(1,2,4,5) have "\<mu> (\<Union>insert c C) = \<mu> c + \<mu> (\<Union>C)"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   736
      by (auto intro!: add simp: disjoint_def)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   737
    with insert show ?case
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   738
      by (simp add: disjoint_def)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   739
  qed simp
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   740
qed fact+
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   741
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   742
lemma (in semiring_of_sets) extend_volume:
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   743
  assumes "volume M \<mu>"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   744
  shows "\<exists>\<mu>'. volume generated_ring \<mu>' \<and> (\<forall>a\<in>M. \<mu>' a = \<mu> a)"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   745
proof -
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   746
  let ?R = generated_ring
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   747
  have "\<forall>a\<in>?R. \<exists>m. \<exists>C\<subseteq>M. a = \<Union>C \<and> finite C \<and> disjoint C \<and> m = (\<Sum>c\<in>C. \<mu> c)"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   748
    by (auto simp: generated_ring_def)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   749
  from bchoice[OF this] guess \<mu>' .. note \<mu>'_spec = this
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   750
  
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   751
  { fix C assume C: "C \<subseteq> M" "finite C" "disjoint C"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   752
    fix D assume D: "D \<subseteq> M" "finite D" "disjoint D"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   753
    assume "\<Union>C = \<Union>D"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   754
    have "(\<Sum>d\<in>D. \<mu> d) = (\<Sum>d\<in>D. \<Sum>c\<in>C. \<mu> (c \<inter> d))"
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56994
diff changeset
   755
    proof (intro setsum.cong refl)
47762
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   756
      fix d assume "d \<in> D"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   757
      have Un_eq_d: "(\<Union>c\<in>C. c \<inter> d) = d"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   758
        using `d \<in> D` `\<Union>C = \<Union>D` by auto
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   759
      moreover have "\<mu> (\<Union>c\<in>C. c \<inter> d) = (\<Sum>c\<in>C. \<mu> (c \<inter> d))"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   760
      proof (rule volume_finite_additive)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   761
        { fix c assume "c \<in> C" then show "c \<inter> d \<in> M"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   762
            using C D `d \<in> D` by auto }
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   763
        show "(\<Union>a\<in>C. a \<inter> d) \<in> M"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   764
          unfolding Un_eq_d using `d \<in> D` D by auto
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   765
        show "disjoint_family_on (\<lambda>a. a \<inter> d) C"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   766
          using `disjoint C` by (auto simp: disjoint_family_on_def disjoint_def)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   767
      qed fact+
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   768
      ultimately show "\<mu> d = (\<Sum>c\<in>C. \<mu> (c \<inter> d))" by simp
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   769
    qed }
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   770
  note split_sum = this
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   771
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   772
  { fix C assume C: "C \<subseteq> M" "finite C" "disjoint C"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   773
    fix D assume D: "D \<subseteq> M" "finite D" "disjoint D"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   774
    assume "\<Union>C = \<Union>D"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   775
    with split_sum[OF C D] split_sum[OF D C]
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   776
    have "(\<Sum>d\<in>D. \<mu> d) = (\<Sum>c\<in>C. \<mu> c)"
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56994
diff changeset
   777
      by (simp, subst setsum.commute, simp add: ac_simps) }
47762
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   778
  note sum_eq = this
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   779
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   780
  { fix C assume C: "C \<subseteq> M" "finite C" "disjoint C"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   781
    then have "\<Union>C \<in> ?R" by (auto simp: generated_ring_def)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   782
    with \<mu>'_spec[THEN bspec, of "\<Union>C"]
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   783
    obtain D where
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   784
      D: "D \<subseteq> M" "finite D" "disjoint D" "\<Union>C = \<Union>D" and "\<mu>' (\<Union>C) = (\<Sum>d\<in>D. \<mu> d)"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   785
      by blast
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   786
    with sum_eq[OF C D] have "\<mu>' (\<Union>C) = (\<Sum>c\<in>C. \<mu> c)" by simp }
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   787
  note \<mu>' = this
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   788
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   789
  show ?thesis
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   790
  proof (intro exI conjI ring_of_sets.volume_additiveI[OF generating_ring] ballI)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   791
    fix a assume "a \<in> M" with \<mu>'[of "{a}"] show "\<mu>' a = \<mu> a"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   792
      by (simp add: disjoint_def)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   793
  next
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   794
    fix a assume "a \<in> ?R" then guess Ca .. note Ca = this
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   795
    with \<mu>'[of Ca] `volume M \<mu>`[THEN volume_positive]
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   796
    show "0 \<le> \<mu>' a"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   797
      by (auto intro!: setsum_nonneg)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   798
  next
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   799
    show "\<mu>' {} = 0" using \<mu>'[of "{}"] by auto
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   800
  next
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   801
    fix a assume "a \<in> ?R" then guess Ca .. note Ca = this
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   802
    fix b assume "b \<in> ?R" then guess Cb .. note Cb = this
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   803
    assume "a \<inter> b = {}"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   804
    with Ca Cb have "Ca \<inter> Cb \<subseteq> {{}}" by auto
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   805
    then have C_Int_cases: "Ca \<inter> Cb = {{}} \<or> Ca \<inter> Cb = {}" by auto
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   806
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   807
    from `a \<inter> b = {}` have "\<mu>' (\<Union> (Ca \<union> Cb)) = (\<Sum>c\<in>Ca \<union> Cb. \<mu> c)"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   808
      using Ca Cb by (intro \<mu>') (auto intro!: disjoint_union)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   809
    also have "\<dots> = (\<Sum>c\<in>Ca \<union> Cb. \<mu> c) + (\<Sum>c\<in>Ca \<inter> Cb. \<mu> c)"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   810
      using C_Int_cases volume_empty[OF `volume M \<mu>`] by (elim disjE) simp_all
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   811
    also have "\<dots> = (\<Sum>c\<in>Ca. \<mu> c) + (\<Sum>c\<in>Cb. \<mu> c)"
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56994
diff changeset
   812
      using Ca Cb by (simp add: setsum.union_inter)
47762
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   813
    also have "\<dots> = \<mu>' a + \<mu>' b"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   814
      using Ca Cb by (simp add: \<mu>')
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   815
    finally show "\<mu>' (a \<union> b) = \<mu>' a + \<mu>' b"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   816
      using Ca Cb by simp
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   817
  qed
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   818
qed
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   819
56994
8d5e5ec1cac3 fixed document generation for HOL-Probability
hoelzl
parents: 56212
diff changeset
   820
subsubsection {* Caratheodory on semirings *}
47762
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   821
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   822
theorem (in semiring_of_sets) caratheodory:
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   823
  assumes pos: "positive M \<mu>" and ca: "countably_additive M \<mu>"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   824
  shows "\<exists>\<mu>' :: 'a set \<Rightarrow> ereal. (\<forall>s \<in> M. \<mu>' s = \<mu> s) \<and> measure_space \<Omega> (sigma_sets \<Omega> M) \<mu>'"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   825
proof -
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   826
  have "volume M \<mu>"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   827
  proof (rule volumeI)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   828
    { fix a assume "a \<in> M" then show "0 \<le> \<mu> a"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   829
        using pos unfolding positive_def by auto }
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   830
    note p = this
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   831
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   832
    fix C assume sets_C: "C \<subseteq> M" "\<Union>C \<in> M" and "disjoint C" "finite C"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   833
    have "\<exists>F'. bij_betw F' {..<card C} C"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   834
      by (rule finite_same_card_bij[OF _ `finite C`]) auto
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   835
    then guess F' .. note F' = this
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   836
    then have F': "C = F' ` {..< card C}" "inj_on F' {..< card C}"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   837
      by (auto simp: bij_betw_def)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   838
    { fix i j assume *: "i < card C" "j < card C" "i \<noteq> j"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   839
      with F' have "F' i \<in> C" "F' j \<in> C" "F' i \<noteq> F' j"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   840
        unfolding inj_on_def by auto
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   841
      with `disjoint C`[THEN disjointD]
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   842
      have "F' i \<inter> F' j = {}"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   843
        by auto }
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   844
    note F'_disj = this
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   845
    def F \<equiv> "\<lambda>i. if i < card C then F' i else {}"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   846
    then have "disjoint_family F"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   847
      using F'_disj by (auto simp: disjoint_family_on_def)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   848
    moreover from F' have "(\<Union>i. F i) = \<Union>C"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   849
      by (auto simp: F_def set_eq_iff split: split_if_asm)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   850
    moreover have sets_F: "\<And>i. F i \<in> M"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   851
      using F' sets_C by (auto simp: F_def)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   852
    moreover note sets_C
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   853
    ultimately have "\<mu> (\<Union>C) = (\<Sum>i. \<mu> (F i))"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   854
      using ca[unfolded countably_additive_def, THEN spec, of F] by auto
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   855
    also have "\<dots> = (\<Sum>i<card C. \<mu> (F' i))"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   856
    proof -
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   857
      have "(\<lambda>i. if i \<in> {..< card C} then \<mu> (F' i) else 0) sums (\<Sum>i<card C. \<mu> (F' i))"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   858
        by (rule sums_If_finite_set) auto
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   859
      also have "(\<lambda>i. if i \<in> {..< card C} then \<mu> (F' i) else 0) = (\<lambda>i. \<mu> (F i))"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   860
        using pos by (auto simp: positive_def F_def)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   861
      finally show "(\<Sum>i. \<mu> (F i)) = (\<Sum>i<card C. \<mu> (F' i))"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   862
        by (simp add: sums_iff)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   863
    qed
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   864
    also have "\<dots> = (\<Sum>c\<in>C. \<mu> c)"
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56994
diff changeset
   865
      using F'(2) by (subst (2) F') (simp add: setsum.reindex)
47762
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   866
    finally show "\<mu> (\<Union>C) = (\<Sum>c\<in>C. \<mu> c)" .
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   867
  next
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   868
    show "\<mu> {} = 0"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   869
      using `positive M \<mu>` by (rule positiveD1)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   870
  qed
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   871
  from extend_volume[OF this] obtain \<mu>_r where
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   872
    V: "volume generated_ring \<mu>_r" "\<And>a. a \<in> M \<Longrightarrow> \<mu> a = \<mu>_r a"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   873
    by auto
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   874
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   875
  interpret G: ring_of_sets \<Omega> generated_ring
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   876
    by (rule generating_ring)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   877
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   878
  have pos: "positive generated_ring \<mu>_r"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   879
    using V unfolding positive_def by (auto simp: positive_def intro!: volume_positive volume_empty)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   880
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   881
  have "countably_additive generated_ring \<mu>_r"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   882
  proof (rule countably_additiveI)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   883
    fix A' :: "nat \<Rightarrow> 'a set" assume A': "range A' \<subseteq> generated_ring" "disjoint_family A'"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   884
      and Un_A: "(\<Union>i. A' i) \<in> generated_ring"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   885
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   886
    from generated_ringE[OF Un_A] guess C' . note C' = this
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   887
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   888
    { fix c assume "c \<in> C'"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   889
      moreover def A \<equiv> "\<lambda>i. A' i \<inter> c"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   890
      ultimately have A: "range A \<subseteq> generated_ring" "disjoint_family A"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   891
        and Un_A: "(\<Union>i. A i) \<in> generated_ring"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   892
        using A' C'
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   893
        by (auto intro!: G.Int G.finite_Union intro: generated_ringI_Basic simp: disjoint_family_on_def)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   894
      from A C' `c \<in> C'` have UN_eq: "(\<Union>i. A i) = c"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   895
        by (auto simp: A_def)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   896
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   897
      have "\<forall>i::nat. \<exists>f::nat \<Rightarrow> 'a set. \<mu>_r (A i) = (\<Sum>j. \<mu>_r (f j)) \<and> disjoint_family f \<and> \<Union>range f = A i \<and> (\<forall>j. f j \<in> M)"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   898
        (is "\<forall>i. ?P i")
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   899
      proof
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   900
        fix i
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   901
        from A have Ai: "A i \<in> generated_ring" by auto
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   902
        from generated_ringE[OF this] guess C . note C = this
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   903
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   904
        have "\<exists>F'. bij_betw F' {..<card C} C"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   905
          by (rule finite_same_card_bij[OF _ `finite C`]) auto
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   906
        then guess F .. note F = this
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   907
        def f \<equiv> "\<lambda>i. if i < card C then F i else {}"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   908
        then have f: "bij_betw f {..< card C} C"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   909
          by (intro bij_betw_cong[THEN iffD1, OF _ F]) auto
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   910
        with C have "\<forall>j. f j \<in> M"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   911
          by (auto simp: Pi_iff f_def dest!: bij_betw_imp_funcset)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   912
        moreover
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   913
        from f C have d_f: "disjoint_family_on f {..<card C}"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   914
          by (intro disjoint_image_disjoint_family_on) (auto simp: bij_betw_def)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   915
        then have "disjoint_family f"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   916
          by (auto simp: disjoint_family_on_def f_def)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   917
        moreover
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   918
        have Ai_eq: "A i = (\<Union> x<card C. f x)"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   919
          using f C Ai unfolding bij_betw_def by (simp add: Union_image_eq[symmetric])
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   920
        then have "\<Union>range f = A i"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   921
          using f C Ai unfolding bij_betw_def by (auto simp: f_def)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   922
        moreover 
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   923
        { have "(\<Sum>j. \<mu>_r (f j)) = (\<Sum>j. if j \<in> {..< card C} then \<mu>_r (f j) else 0)"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   924
            using volume_empty[OF V(1)] by (auto intro!: arg_cong[where f=suminf] simp: f_def)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   925
          also have "\<dots> = (\<Sum>j<card C. \<mu>_r (f j))"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   926
            by (rule sums_If_finite_set[THEN sums_unique, symmetric]) simp
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   927
          also have "\<dots> = \<mu>_r (A i)"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   928
            using C f[THEN bij_betw_imp_funcset] unfolding Ai_eq
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   929
            by (intro volume_finite_additive[OF V(1) _ d_f, symmetric])
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   930
               (auto simp: Pi_iff Ai_eq intro: generated_ringI_Basic)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   931
          finally have "\<mu>_r (A i) = (\<Sum>j. \<mu>_r (f j))" .. }
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   932
        ultimately show "?P i"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   933
          by blast
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   934
      qed
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   935
      from choice[OF this] guess f .. note f = this
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   936
      then have UN_f_eq: "(\<Union>i. split f (prod_decode i)) = (\<Union>i. A i)"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   937
        unfolding UN_extend_simps surj_prod_decode by (auto simp: set_eq_iff)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   938
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   939
      have d: "disjoint_family (\<lambda>i. split f (prod_decode i))"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   940
        unfolding disjoint_family_on_def
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   941
      proof (intro ballI impI)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   942
        fix m n :: nat assume "m \<noteq> n"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   943
        then have neq: "prod_decode m \<noteq> prod_decode n"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   944
          using inj_prod_decode[of UNIV] by (auto simp: inj_on_def)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   945
        show "split f (prod_decode m) \<inter> split f (prod_decode n) = {}"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   946
        proof cases
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   947
          assume "fst (prod_decode m) = fst (prod_decode n)"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   948
          then show ?thesis
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   949
            using neq f by (fastforce simp: disjoint_family_on_def)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   950
        next
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   951
          assume neq: "fst (prod_decode m) \<noteq> fst (prod_decode n)"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   952
          have "split f (prod_decode m) \<subseteq> A (fst (prod_decode m))"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   953
            "split f (prod_decode n) \<subseteq> A (fst (prod_decode n))"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   954
            using f[THEN spec, of "fst (prod_decode m)"]
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   955
            using f[THEN spec, of "fst (prod_decode n)"]
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   956
            by (auto simp: set_eq_iff)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   957
          with f A neq show ?thesis
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   958
            by (fastforce simp: disjoint_family_on_def subset_eq set_eq_iff)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   959
        qed
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   960
      qed
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   961
      from f have "(\<Sum>n. \<mu>_r (A n)) = (\<Sum>n. \<mu>_r (split f (prod_decode n)))"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   962
        by (intro suminf_ereal_2dimen[symmetric] positiveD2[OF pos] generated_ringI_Basic)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   963
         (auto split: prod.split)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   964
      also have "\<dots> = (\<Sum>n. \<mu> (split f (prod_decode n)))"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   965
        using f V(2) by (auto intro!: arg_cong[where f=suminf] split: prod.split)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   966
      also have "\<dots> = \<mu> (\<Union>i. split f (prod_decode i))"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   967
        using f `c \<in> C'` C'
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   968
        by (intro ca[unfolded countably_additive_def, rule_format])
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   969
           (auto split: prod.split simp: UN_f_eq d UN_eq)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   970
      finally have "(\<Sum>n. \<mu>_r (A' n \<inter> c)) = \<mu> c"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   971
        using UN_f_eq UN_eq by (simp add: A_def) }
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   972
    note eq = this
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   973
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   974
    have "(\<Sum>n. \<mu>_r (A' n)) = (\<Sum>n. \<Sum>c\<in>C'. \<mu>_r (A' n \<inter> c))"
49394
52e636ace94e removing find_theorems commands that were left in the developments accidently
bulwahn
parents: 47762
diff changeset
   975
      using C' A'
47762
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   976
      by (subst volume_finite_additive[symmetric, OF V(1)])
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 55642
diff changeset
   977
         (auto simp: disjoint_def disjoint_family_on_def Union_image_eq[symmetric] simp del: Sup_image_eq Union_image_eq
47762
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   978
               intro!: G.Int G.finite_Union arg_cong[where f="\<lambda>X. suminf (\<lambda>i. \<mu>_r (X i))"] ext
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   979
               intro: generated_ringI_Basic)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   980
    also have "\<dots> = (\<Sum>c\<in>C'. \<Sum>n. \<mu>_r (A' n \<inter> c))"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   981
      using C' A'
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   982
      by (intro suminf_setsum_ereal positiveD2[OF pos] G.Int G.finite_Union)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   983
         (auto intro: generated_ringI_Basic)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   984
    also have "\<dots> = (\<Sum>c\<in>C'. \<mu>_r c)"
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56994
diff changeset
   985
      using eq V C' by (auto intro!: setsum.cong)
47762
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   986
    also have "\<dots> = \<mu>_r (\<Union>C')"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   987
      using C' Un_A
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   988
      by (subst volume_finite_additive[symmetric, OF V(1)])
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 55642
diff changeset
   989
         (auto simp: disjoint_family_on_def disjoint_def Union_image_eq[symmetric] simp del: Sup_image_eq Union_image_eq 
47762
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   990
               intro: generated_ringI_Basic)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   991
    finally show "(\<Sum>n. \<mu>_r (A' n)) = \<mu>_r (\<Union>i. A' i)"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   992
      using C' by simp
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   993
  qed
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   994
  from G.caratheodory'[OF `positive generated_ring \<mu>_r` `countably_additive generated_ring \<mu>_r`]
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   995
  guess \<mu>' ..
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   996
  with V show ?thesis
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   997
    unfolding sigma_sets_generated_ring_eq
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   998
    by (intro exI[of _ \<mu>']) (auto intro: generated_ringI_Basic)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   999
qed
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
  1000
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1001
lemma extend_measure_caratheodory:
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1002
  fixes G :: "'i \<Rightarrow> 'a set"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1003
  assumes M: "M = extend_measure \<Omega> I G \<mu>"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1004
  assumes "i \<in> I"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1005
  assumes "semiring_of_sets \<Omega> (G ` I)"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1006
  assumes empty: "\<And>i. i \<in> I \<Longrightarrow> G i = {} \<Longrightarrow> \<mu> i = 0"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1007
  assumes inj: "\<And>i j. i \<in> I \<Longrightarrow> j \<in> I \<Longrightarrow> G i = G j \<Longrightarrow> \<mu> i = \<mu> j"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1008
  assumes nonneg: "\<And>i. i \<in> I \<Longrightarrow> 0 \<le> \<mu> i"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1009
  assumes add: "\<And>A::nat \<Rightarrow> 'i. \<And>j. A \<in> UNIV \<rightarrow> I \<Longrightarrow> j \<in> I \<Longrightarrow> disjoint_family (G \<circ> A) \<Longrightarrow>
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1010
    (\<Union>i. G (A i)) = G j \<Longrightarrow> (\<Sum>n. \<mu> (A n)) = \<mu> j"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1011
  shows "emeasure M (G i) = \<mu> i"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1012
proof -
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1013
  interpret semiring_of_sets \<Omega> "G ` I"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1014
    by fact
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1015
  have "\<forall>g\<in>G`I. \<exists>i\<in>I. g = G i"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1016
    by auto
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1017
  then obtain sel where sel: "\<And>g. g \<in> G ` I \<Longrightarrow> sel g \<in> I" "\<And>g. g \<in> G ` I \<Longrightarrow> G (sel g) = g"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1018
    by metis
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1019
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1020
  have "\<exists>\<mu>'. (\<forall>s\<in>G ` I. \<mu>' s = \<mu> (sel s)) \<and> measure_space \<Omega> (sigma_sets \<Omega> (G ` I)) \<mu>'"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1021
  proof (rule caratheodory)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1022
    show "positive (G ` I) (\<lambda>s. \<mu> (sel s))"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1023
      by (auto simp: positive_def intro!: empty sel nonneg)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1024
    show "countably_additive (G ` I) (\<lambda>s. \<mu> (sel s))"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1025
    proof (rule countably_additiveI)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1026
      fix A :: "nat \<Rightarrow> 'a set" assume "range A \<subseteq> G ` I" "disjoint_family A" "(\<Union>i. A i) \<in> G ` I"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1027
      then show "(\<Sum>i. \<mu> (sel (A i))) = \<mu> (sel (\<Union>i. A i))"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1028
        by (intro add) (auto simp: sel image_subset_iff_funcset comp_def Pi_iff intro!: sel)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1029
    qed
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1030
  qed
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1031
  then obtain \<mu>' where \<mu>': "\<forall>s\<in>G ` I. \<mu>' s = \<mu> (sel s)" "measure_space \<Omega> (sigma_sets \<Omega> (G ` I)) \<mu>'"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1032
    by metis
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1033
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1034
  show ?thesis
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1035
  proof (rule emeasure_extend_measure[OF M])
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1036
    { fix i assume "i \<in> I" then show "\<mu>' (G i) = \<mu> i"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1037
      using \<mu>' by (auto intro!: inj sel) }
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1038
    show "G ` I \<subseteq> Pow \<Omega>"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1039
      by fact
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1040
    then show "positive (sets M) \<mu>'" "countably_additive (sets M) \<mu>'"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1041
      using \<mu>' by (simp_all add: M sets_extend_measure measure_space_def)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1042
  qed fact
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1043
qed
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1044
  
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1045
lemma extend_measure_caratheodory_pair:
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1046
  fixes G :: "'i \<Rightarrow> 'j \<Rightarrow> 'a set"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1047
  assumes M: "M = extend_measure \<Omega> {(a, b). P a b} (\<lambda>(a, b). G a b) (\<lambda>(a, b). \<mu> a b)"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1048
  assumes "P i j"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1049
  assumes semiring: "semiring_of_sets \<Omega> {G a b | a b. P a b}"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1050
  assumes empty: "\<And>i j. P i j \<Longrightarrow> G i j = {} \<Longrightarrow> \<mu> i j = 0"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1051
  assumes inj: "\<And>i j k l. P i j \<Longrightarrow> P k l \<Longrightarrow> G i j = G k l \<Longrightarrow> \<mu> i j = \<mu> k l"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1052
  assumes nonneg: "\<And>i j. P i j \<Longrightarrow> 0 \<le> \<mu> i j"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1053
  assumes add: "\<And>A::nat \<Rightarrow> 'i. \<And>B::nat \<Rightarrow> 'j. \<And>j k.
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1054
    (\<And>n. P (A n) (B n)) \<Longrightarrow> P j k \<Longrightarrow> disjoint_family (\<lambda>n. G (A n) (B n)) \<Longrightarrow>
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1055
    (\<Union>i. G (A i) (B i)) = G j k \<Longrightarrow> (\<Sum>n. \<mu> (A n) (B n)) = \<mu> j k"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1056
  shows "emeasure M (G i j) = \<mu> i j"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1057
proof -
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1058
  have "emeasure M ((\<lambda>(a, b). G a b) (i, j)) = (\<lambda>(a, b). \<mu> a b) (i, j)"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1059
  proof (rule extend_measure_caratheodory[OF M])
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1060
    show "semiring_of_sets \<Omega> ((\<lambda>(a, b). G a b) ` {(a, b). P a b})"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1061
      using semiring by (simp add: image_def conj_commute)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1062
  next
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1063
    fix A :: "nat \<Rightarrow> ('i \<times> 'j)" and j assume "A \<in> UNIV \<rightarrow> {(a, b). P a b}" "j \<in> {(a, b). P a b}"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1064
      "disjoint_family ((\<lambda>(a, b). G a b) \<circ> A)"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1065
      "(\<Union>i. case A i of (a, b) \<Rightarrow> G a b) = (case j of (a, b) \<Rightarrow> G a b)"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1066
    then show "(\<Sum>n. case A n of (a, b) \<Rightarrow> \<mu> a b) = (case j of (a, b) \<Rightarrow> \<mu> a b)"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1067
      using add[of "\<lambda>i. fst (A i)" "\<lambda>i. snd (A i)" "fst j" "snd j"]
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1068
      by (simp add: split_beta' comp_def Pi_iff)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1069
  qed (auto split: prod.splits intro: assms)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1070
  then show ?thesis by simp
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1071
qed
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1072
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1073
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
  1074
end