src/HOL/Probability/Caratheodory.thy
author wenzelm
Wed, 04 Nov 2015 23:27:00 +0100
changeset 61578 6623c81cb15a
parent 61427 3c69ea85f8dd
child 61808 fc1556774cfe
permissions -rw-r--r--
avoid ligatures;
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
58876
1888e3cb8048 modernized header;
wenzelm
parents: 57447
diff changeset
     6
section {*Caratheodory Extension Theorem*}
33271
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
61273
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
    44
                     suminf_ereal_eq_SUP SUP_pair
56212
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
61273
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
    50
definition subadditive where
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
    51
  "subadditive M f \<longleftrightarrow> (\<forall>x\<in>M. \<forall>y\<in>M. x \<inter> y = {} \<longrightarrow> f (x \<union> y) \<le> f x + f 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
    52
61273
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
    53
definition countably_subadditive where
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
    54
  "countably_subadditive M f \<longleftrightarrow>
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
    55
    (\<forall>A. range A \<subseteq> M \<longrightarrow> disjoint_family A \<longrightarrow> (\<Union>i. A i) \<in> M \<longrightarrow> (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
61273
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
    57
definition outer_measure_space where
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
    58
  "outer_measure_space M f \<longleftrightarrow> positive M f \<and> increasing M f \<and> countably_subadditive M f"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    59
61273
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
    60
lemma subadditiveD: "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
    61
  by (auto simp add: subadditive_def)
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    62
56994
8d5e5ec1cac3 fixed document generation for HOL-Probability
hoelzl
parents: 56212
diff changeset
    63
subsubsection {* Lambda Systems *}
8d5e5ec1cac3 fixed document generation for HOL-Probability
hoelzl
parents: 56212
diff changeset
    64
61273
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
    65
definition lambda_system where
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
    66
  "lambda_system \<Omega> M f = {l \<in> M. \<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
    67
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    68
lemma (in algebra) lambda_system_eq:
61273
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
    69
  "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
    70
proof -
61273
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
    71
  have [simp]: "\<And>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
    72
    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
    73
  show ?thesis
37032
58a0757031dd speed up some proofs and fix some warnings
huffman
parents: 36649
diff changeset
    74
    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
    75
qed
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    76
61273
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
    77
lemma (in algebra) lambda_system_empty: "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
    78
  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
    79
61273
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
    80
lemma lambda_system_sets: "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
    81
  by (simp add: lambda_system_def)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    82
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    83
lemma (in algebra) lambda_system_Compl:
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
    84
  fixes f:: "'a set \<Rightarrow> ereal"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
    85
  assumes x: "x \<in> lambda_system \<Omega> M f"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
    86
  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
    87
proof -
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
    88
  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
    89
    by (metis sets_into_space lambda_system_sets x)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
    90
  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
    91
    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
    92
  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
    93
    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
    94
qed
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    95
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
    96
lemma (in algebra) lambda_system_Int:
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
    97
  fixes f:: "'a set \<Rightarrow> ereal"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
    98
  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
    99
  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
   100
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
   101
  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
   102
  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
   103
    fix u
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   104
    assume x: "x \<in> M" and y: "y \<in> M" and u: "u \<in> M"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   105
       and fx: "\<forall>z\<in>M. f (z \<inter> x) + f (z - x) = f z"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   106
       and fy: "\<forall>z\<in>M. f (z \<inter> y) + f (z - y) = f z"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   107
    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
   108
      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
   109
    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
   110
    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
   111
    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
   112
    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
   113
    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
   114
    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
   115
      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
   116
    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
   117
          = (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
   118
      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
   119
    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
   120
      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
   121
    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
   122
      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
   123
      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
   124
    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
   125
      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
   126
    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
   127
  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
   128
qed
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   129
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   130
lemma (in algebra) lambda_system_Un:
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
   131
  fixes f:: "'a set \<Rightarrow> ereal"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   132
  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
   133
  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
   134
proof -
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   135
  have "(\<Omega> - x) \<inter> (\<Omega> - y) \<in> M"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   136
    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
   137
  moreover
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   138
  have "x \<union> y = \<Omega> - ((\<Omega> - x) \<inter> (\<Omega> - y))"
46731
5302e932d1e5 avoid undeclared variables in let bindings;
wenzelm
parents: 44928
diff changeset
   139
    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
   140
  ultimately show ?thesis
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   141
    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
   142
qed
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   143
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   144
lemma (in algebra) lambda_system_algebra:
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   145
  "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
   146
  apply (auto simp add: algebra_iff_Un)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   147
  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
   148
  apply (metis lambda_system_empty)
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   149
  apply (metis lambda_system_Compl)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   150
  apply (metis lambda_system_Un)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   151
  done
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   152
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   153
lemma (in algebra) lambda_system_strong_additive:
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   154
  assumes z: "z \<in> M" and disj: "x \<inter> y = {}"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   155
      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
   156
  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
   157
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
   158
  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
   159
  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
   160
  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
   161
  moreover
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   162
  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
   163
    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
   164
  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
   165
    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
   166
qed
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   167
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   168
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
   169
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
   170
  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
   171
  assume disj: "x \<inter> y = {}"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   172
     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
   173
  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
   174
  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
   175
    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
   176
    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
   177
qed
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   178
42145
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   179
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
   180
  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
   181
  shows  "subadditive M f"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33536
diff changeset
   182
proof (auto simp add: subadditive_def)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   183
  fix x y
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   184
  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
   185
  hence "disjoint_family (binaryset x y)"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents: 33536
diff changeset
   186
    by (auto simp add: disjoint_family_on_def binaryset_def)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   187
  hence "range (binaryset x y) \<subseteq> M \<longrightarrow>
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   188
         (\<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
   189
         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
   190
    using cs by (auto simp add: countably_subadditive_def)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   191
  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
   192
         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
   193
    by (simp add: range_binaryset_eq UN_binaryset_eq)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   194
  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
   195
    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
   196
qed
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   197
61273
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   198
lemma lambda_system_increasing: "increasing M f \<Longrightarrow> increasing (lambda_system \<Omega> M f) f"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   199
  by (simp add: increasing_def lambda_system_def)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   200
61273
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   201
lemma lambda_system_positive: "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
   202
  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
   203
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   204
lemma (in algebra) lambda_system_strong_sum:
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
   205
  fixes A:: "nat \<Rightarrow> 'a set" and f :: "'a set \<Rightarrow> ereal"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   206
  assumes f: "positive M f" and a: "a \<in> M"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   207
      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
   208
      and disj: "disjoint_family A"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   209
  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
   210
proof (induct n)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   211
  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
   212
next
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   213
  case (Suc n)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   214
  have 2: "A n \<inter> UNION {0..<n} A = {}" using disj
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   215
    by (force simp add: disjoint_family_on_def neq_iff)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   216
  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
   217
    by blast
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   218
  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
   219
    using f by (rule lambda_system_algebra)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   220
  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
   221
    using A l.UNION_in_sets by simp
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   222
  from Suc.hyps show ?case
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   223
    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
   224
qed
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   225
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   226
lemma (in sigma_algebra) lambda_system_caratheodory:
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   227
  assumes oms: "outer_measure_space M f"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   228
      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
   229
      and disj: "disjoint_family A"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   230
  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
   231
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
   232
  have pos: "positive M f" and inc: "increasing M f"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   233
   and csa: "countably_subadditive M f"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   234
    by (metis oms outer_measure_space_def)+
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   235
  have sa: "subadditive M f"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   236
    by (metis countably_subadditive_subadditive csa pos)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   237
  have A': "\<And>S. A`S \<subseteq> (lambda_system \<Omega> M f)" using A
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   238
    by auto
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   239
  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
   240
    using pos by (rule lambda_system_algebra)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   241
  have A'': "range A \<subseteq> M"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   242
     by (metis A image_subset_iff lambda_system_sets)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   243
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   244
  have U_in: "(\<Union>i. A i) \<in> M"
37032
58a0757031dd speed up some proofs and fix some warnings
huffman
parents: 36649
diff changeset
   245
    by (metis A'' countable_UN)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   246
  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
   247
  proof (rule antisym)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   248
    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
   249
      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
   250
    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
   251
    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
   252
    show "(\<Sum>i. f (A i)) \<le> f (\<Union>i. A i)"
61273
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   253
      using ls.additive_sum [OF lambda_system_positive[OF pos] lambda_system_additive _ A' dis] A''
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   254
      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
   255
  qed
61273
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   256
  have "f (a \<inter> (\<Union>i. A i)) + f (a - (\<Union>i. A i)) = f a"
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   257
    if a [iff]: "a \<in> M" for a
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   258
  proof (rule antisym)
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   259
    have "range (\<lambda>i. a \<inter> A i) \<subseteq> M" using A''
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   260
      by blast
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   261
    moreover
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   262
    have "disjoint_family (\<lambda>i. a \<inter> A i)" using disj
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   263
      by (auto simp add: disjoint_family_on_def)
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   264
    moreover
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   265
    have "a \<inter> (\<Union>i. A i) \<in> M"
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   266
      by (metis Int U_in a)
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   267
    ultimately
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   268
    have "f (a \<inter> (\<Union>i. A i)) \<le> (\<Sum>i. f (a \<inter> A i))"
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   269
      using csa[unfolded countably_subadditive_def, rule_format, of "(\<lambda>i. a \<inter> A i)"]
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   270
      by (simp add: o_def)
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   271
    hence "f (a \<inter> (\<Union>i. A i)) + f (a - (\<Union>i. A i)) \<le> (\<Sum>i. f (a \<inter> A i)) + f (a - (\<Union>i. A i))"
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   272
      by (rule add_right_mono)
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   273
    also have "\<dots> \<le> f a"
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   274
    proof (intro suminf_bound_add allI)
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   275
      fix n
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   276
      have UNION_in: "(\<Union>i\<in>{0..<n}. A i) \<in> M"
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   277
        by (metis A'' UNION_in_sets)
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   278
      have le_fa: "f (UNION {0..<n} A \<inter> a) \<le> f a" using A''
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   279
        by (blast intro: increasingD [OF inc] A'' UNION_in_sets)
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   280
      have ls: "(\<Union>i\<in>{0..<n}. A i) \<in> lambda_system \<Omega> M f"
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   281
        using ls.UNION_in_sets by (simp add: A)
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   282
      hence eq_fa: "f a = f (a \<inter> (\<Union>i\<in>{0..<n}. A i)) + f (a - (\<Union>i\<in>{0..<n}. A i))"
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   283
        by (simp add: lambda_system_eq UNION_in)
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   284
      have "f (a - (\<Union>i. A i)) \<le> f (a - (\<Union>i\<in>{0..<n}. A i))"
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   285
        by (blast intro: increasingD [OF inc] UNION_in U_in)
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   286
      thus "(\<Sum>i<n. f (a \<inter> A i)) + f (a - (\<Union>i. A i)) \<le> f a"
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   287
        by (simp add: lambda_system_strong_sum pos A disj eq_fa add_left_mono atLeast0LessThan[symmetric])
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   288
    next
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   289
      have "\<And>i. a \<inter> A i \<in> M" using A'' by auto
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   290
      then show "\<And>i. 0 \<le> f (a \<inter> A i)" using pos[unfolded positive_def] by auto
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   291
      have "\<And>i. a - (\<Union>i. A i) \<in> M" using A'' by auto
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   292
      then have "\<And>i. 0 \<le> f (a - (\<Union>i. A i))" using pos[unfolded positive_def] by auto
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   293
      then show "f (a - (\<Union>i. A i)) \<noteq> -\<infinity>" by auto
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   294
    qed
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   295
    finally show "f (a \<inter> (\<Union>i. A i)) + f (a - (\<Union>i. A i)) \<le> f a" .
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   296
  next
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   297
    have "f a \<le> f (a \<inter> (\<Union>i. A i) \<union> (a - (\<Union>i. A i)))"
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   298
      by (blast intro:  increasingD [OF inc] U_in)
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   299
    also have "... \<le>  f (a \<inter> (\<Union>i. A i)) + f (a - (\<Union>i. A i))"
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   300
      by (blast intro: subadditiveD [OF sa] U_in)
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   301
    finally show "f a \<le> f (a \<inter> (\<Union>i. A i)) + f (a - (\<Union>i. A i))" .
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   302
  qed
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   303
  thus  ?thesis
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   304
    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
   305
qed
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   306
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   307
lemma (in sigma_algebra) caratheodory_lemma:
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   308
  assumes oms: "outer_measure_space M f"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   309
  defines "L \<equiv> lambda_system \<Omega> M f"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   310
  shows "measure_space \<Omega> L f"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   311
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
   312
  have pos: "positive M f"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   313
    by (metis oms outer_measure_space_def)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   314
  have alg: "algebra \<Omega> L"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   315
    using lambda_system_algebra [of f, OF pos]
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   316
    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
   317
  then
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   318
  have "sigma_algebra \<Omega> L"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   319
    using lambda_system_caratheodory [OF oms]
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   320
    by (simp add: sigma_algebra_disjoint_iff L_def)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   321
  moreover
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   322
  have "countably_additive L f" "positive L f"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   323
    using pos lambda_system_caratheodory [OF oms]
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   324
    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
   325
  ultimately
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   326
  show ?thesis
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   327
    using pos by (simp add: measure_space_def)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   328
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   329
61273
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   330
definition outer_measure :: "'a set set \<Rightarrow> ('a set \<Rightarrow> ereal) \<Rightarrow> 'a set \<Rightarrow> ereal" where
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   331
   "outer_measure M f X =
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   332
     (INF A:{A. range A \<subseteq> M \<and> disjoint_family A \<and> X \<subseteq> (\<Union>i. A i)}. \<Sum>i. f (A i))"
39096
hoelzl
parents: 38656
diff changeset
   333
61273
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   334
lemma (in ring_of_sets) outer_measure_agrees:
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   335
  assumes posf: "positive M f" and ca: "countably_additive M f" and s: "s \<in> M"
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   336
  shows "outer_measure M f s = f s"
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   337
  unfolding outer_measure_def
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   338
proof (safe intro!: antisym INF_greatest)
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   339
  fix A :: "nat \<Rightarrow> 'a set" assume A: "range A \<subseteq> M" and dA: "disjoint_family A" and sA: "s \<subseteq> (\<Union>x. A x)"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   340
  have inc: "increasing M f"
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   341
    by (metis additive_increasing ca countably_additive_additive posf)
61273
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   342
  have "f s = f (\<Union>i. A i \<inter> s)"
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   343
    using sA by (auto simp: Int_absorb1)
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   344
  also have "\<dots> = (\<Sum>i. f (A i \<inter> s))"
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   345
    using sA dA A s
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   346
    by (intro ca[unfolded countably_additive_def, rule_format, symmetric])
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   347
       (auto simp: Int_absorb1 disjoint_family_on_def)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   348
  also have "... \<le> (\<Sum>i. f (A i))"
61273
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   349
    using A s by (intro suminf_le_pos increasingD[OF inc] positiveD2[OF posf]) auto
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   350
  finally show "f s \<le> (\<Sum>i. f (A i))" .
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   351
next
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   352
  have "(\<Sum>i. f (if i = 0 then s else {})) \<le> f s"
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   353
    using positiveD1[OF posf] by (subst suminf_finite[of "{0}"]) auto
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   354
  with s show "(INF A:{A. range A \<subseteq> M \<and> disjoint_family A \<and> s \<subseteq> UNION UNIV A}. \<Sum>i. f (A i)) \<le> f s"
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   355
    by (intro INF_lower2[of "\<lambda>i. if i = 0 then s else {}"])
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   356
       (auto simp: disjoint_family_on_def)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   357
qed
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41689
diff changeset
   358
61273
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   359
lemma outer_measure_nonneg: "positive M f \<Longrightarrow> 0 \<le> outer_measure M f X"
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   360
  by (auto intro!: INF_greatest suminf_0_le intro: positiveD2 simp: outer_measure_def)
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   361
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   362
lemma outer_measure_empty:
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   363
  assumes posf: "positive M f" and "{} \<in> M"
61273
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   364
  shows "outer_measure M f {} = 0"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   365
proof (rule antisym)
61273
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   366
  show "outer_measure M f {} \<le> 0"
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   367
    using assms by (auto intro!: INF_lower2[of "\<lambda>_. {}"] simp: outer_measure_def disjoint_family_on_def positive_def)
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   368
qed (intro outer_measure_nonneg posf)
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   369
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   370
lemma (in ring_of_sets) positive_outer_measure:
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   371
  assumes "positive M f" shows "positive (Pow \<Omega>) (outer_measure M f)"
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   372
  unfolding positive_def by (auto simp: assms outer_measure_empty outer_measure_nonneg)
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   373
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   374
lemma (in ring_of_sets) increasing_outer_measure: "increasing (Pow \<Omega>) (outer_measure M f)"
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   375
  by (force simp: increasing_def outer_measure_def intro!: INF_greatest intro: INF_lower)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   376
61273
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   377
lemma (in ring_of_sets) outer_measure_le:
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   378
  assumes pos: "positive M f" and inc: "increasing M f" and A: "range A \<subseteq> M" and X: "X \<subseteq> (\<Union>i. A i)"
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   379
  shows "outer_measure M f X \<le> (\<Sum>i. f (A i))"
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   380
  unfolding outer_measure_def
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   381
proof (safe intro!: INF_lower2[of "disjointed A"] del: subsetI)
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   382
  show dA: "range (disjointed A) \<subseteq> M"
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   383
    by (auto intro!: A range_disjointed_sets)
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   384
  have "\<forall>n. f (disjointed A n) \<le> f (A n)"
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   385
    by (metis increasingD [OF inc] UNIV_I dA image_subset_iff disjointed_subset A)
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   386
  moreover have "\<forall>i. 0 \<le> f (disjointed A i)"
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   387
    using pos dA unfolding positive_def by auto
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   388
  ultimately show "(\<Sum>i. f (disjointed A i)) \<le> (\<Sum>i. f (A i))"
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   389
    by (blast intro!: suminf_le_pos)
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   390
qed (auto simp: X UN_disjointed_eq disjoint_family_disjointed)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   391
61273
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   392
lemma (in ring_of_sets) outer_measure_close:
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   393
  assumes posf: "positive M f" and "0 < e" and "outer_measure M f X \<noteq> \<infinity>"
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   394
  shows "\<exists>A. range A \<subseteq> M \<and> disjoint_family A \<and> X \<subseteq> (\<Union>i. A i) \<and> (\<Sum>i. f (A i)) \<le> outer_measure M f X + e"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   395
proof -
61273
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   396
  from `outer_measure M f X \<noteq> \<infinity>` have fin: "\<bar>outer_measure M f X\<bar> \<noteq> \<infinity>"
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   397
    using outer_measure_nonneg[OF posf, of X] by auto
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   398
  show ?thesis
61273
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   399
    using Inf_ereal_close[OF fin[unfolded outer_measure_def INF_def], OF \<open>0 < e\<close>]
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   400
    unfolding INF_def[symmetric] outer_measure_def[symmetric] by (auto intro: less_imp_le)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   401
qed
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   402
61273
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   403
lemma (in ring_of_sets) countably_subadditive_outer_measure:
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
   404
  assumes posf: "positive M f" and inc: "increasing M f"
61273
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   405
  shows "countably_subadditive (Pow \<Omega>) (outer_measure M f)"
42066
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   406
proof (simp add: countably_subadditive_def, safe)
61273
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   407
  fix A :: "nat \<Rightarrow> _" assume A: "range A \<subseteq> Pow (\<Omega>)" and sb: "(\<Union>i. A i) \<subseteq> \<Omega>"
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   408
  let ?O = "outer_measure M f"
42066
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   409
61273
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   410
  { fix e :: ereal assume e: "0 < e" and "\<forall>i. ?O (A i) \<noteq> \<infinity>"
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   411
    hence "\<exists>B. \<forall>n. range (B n) \<subseteq> M \<and> disjoint_family (B n) \<and> A n \<subseteq> (\<Union>i. B n i) \<and>
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   412
        (\<Sum>i. f (B n i)) \<le> ?O (A n) + e * (1/2)^(Suc n)"
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   413
      using e sb by (auto intro!: choice outer_measure_close [of f, OF posf] simp: ereal_zero_less_0_iff one_ereal_def)
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   414
    then obtain B
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   415
      where B: "\<And>n. range (B n) \<subseteq> M"
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   416
      and sbB: "\<And>n. A n \<subseteq> (\<Union>i. B n i)"
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   417
      and Ble: "\<And>n. (\<Sum>i. f (B n i)) \<le> ?O (A n) + e * (1/2)^(Suc n)"
42066
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   418
      by auto blast
61273
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   419
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61273
diff changeset
   420
    def C \<equiv> "case_prod B o prod_decode"
61273
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   421
    from B have B_in_M: "\<And>i j. B i j \<in> M"
61032
b57df8eecad6 standardized some occurences of ancient "split" alias
haftmann
parents: 60585
diff changeset
   422
      by (rule range_subsetD)
61273
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   423
    then have C: "range C \<subseteq> M"
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   424
      by (auto simp add: C_def split_def)
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   425
    have A_C: "(\<Union>i. A i) \<subseteq> (\<Union>i. C i)"
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   426
      using sbB by (auto simp add: C_def subset_eq) (metis prod.case prod_encode_inverse)
42066
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   427
61273
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   428
    have "?O (\<Union>i. A i) \<le> ?O (\<Union>i. C i)"  
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   429
      using A_C A C by (intro increasing_outer_measure[THEN increasingD]) (auto dest!: sets_into_space)
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   430
    also have "\<dots> \<le> (\<Sum>i. f (C i))"
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   431
      using C by (intro outer_measure_le[OF posf inc]) auto
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   432
    also have "\<dots> = (\<Sum>n. \<Sum>i. f (B n i))"
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   433
      using B_in_M unfolding C_def comp_def by (intro suminf_ereal_2dimen positiveD2[OF posf]) auto
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   434
    also have "\<dots> \<le> (\<Sum>n. ?O (A n) + e*(1/2) ^ Suc n)"
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   435
      using B_in_M by (intro suminf_le_pos[OF Ble] suminf_0_le posf[THEN positiveD2]) auto
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   436
    also have "... = (\<Sum>n. ?O (A n)) + (\<Sum>n. e*(1/2) ^ Suc n)"
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   437
      using e by (subst suminf_add_ereal) (auto simp add: ereal_zero_le_0_iff outer_measure_nonneg posf)
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   438
    also have "(\<Sum>n. e*(1/2) ^ Suc n) = e"
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   439
      using suminf_half_series_ereal e by (simp add: ereal_zero_le_0_iff suminf_cmult_ereal)
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   440
    finally have "?O (\<Union>i. A i) \<le> (\<Sum>n. ?O (A n)) + e" . }
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   441
  note * = this
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   442
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   443
  show "?O (\<Union>i. A i) \<le> (\<Sum>n. ?O (A n))"
42066
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   444
  proof cases
61273
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   445
    assume "\<forall>i. ?O (A i) \<noteq> \<infinity>" with * show ?thesis
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42950
diff changeset
   446
      by (intro ereal_le_epsilon) auto
61273
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   447
  qed (metis suminf_PInfty[OF outer_measure_nonneg, OF posf] ereal_less_eq(1))
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   448
qed
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   449
61273
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   450
lemma (in ring_of_sets) outer_measure_space_outer_measure:
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   451
  "positive M f \<Longrightarrow> increasing M f \<Longrightarrow> outer_measure_space (Pow \<Omega>) (outer_measure M f)"
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   452
  by (simp add: outer_measure_space_def
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   453
    positive_outer_measure increasing_outer_measure countably_subadditive_outer_measure)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   454
42066
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   455
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
   456
  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
   457
      and add: "additive M f"
61273
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   458
  shows "M \<subseteq> lambda_system \<Omega> (Pow \<Omega>) (outer_measure M f)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   459
proof (auto dest: sets_into_space
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   460
            simp add: algebra.lambda_system_eq [OF algebra_Pow])
61273
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   461
  fix x s assume x: "x \<in> M" and s: "s \<subseteq> \<Omega>"
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   462
  have [simp]: "\<And>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
   463
    by blast
61273
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   464
  have "outer_measure M f (s \<inter> x) + outer_measure M f (s - x) \<le> outer_measure M f s"
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   465
    unfolding outer_measure_def[of M f s]
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   466
  proof (safe intro!: INF_greatest)
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   467
    fix A :: "nat \<Rightarrow> 'a set" assume A: "disjoint_family A" "range A \<subseteq> M" "s \<subseteq> (\<Union>i. A i)"
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   468
    have "outer_measure M f (s \<inter> x) \<le> (\<Sum>i. f (A i \<inter> x))"
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   469
      unfolding outer_measure_def
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   470
    proof (safe intro!: INF_lower2[of "\<lambda>i. A i \<inter> x"])
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   471
      from A(1) show "disjoint_family (\<lambda>i. A i \<inter> x)"
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   472
        by (rule disjoint_family_on_bisimulation) auto
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   473
    qed (insert x A, auto)
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   474
    moreover
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   475
    have "outer_measure M f (s - x) \<le> (\<Sum>i. f (A i - x))"
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   476
      unfolding outer_measure_def
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   477
    proof (safe intro!: INF_lower2[of "\<lambda>i. A i - x"])
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   478
      from A(1) show "disjoint_family (\<lambda>i. A i - x)"
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   479
        by (rule disjoint_family_on_bisimulation) auto
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   480
    qed (insert x A, auto)
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   481
    ultimately have "outer_measure M f (s \<inter> x) + outer_measure M f (s - x) \<le>
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   482
        (\<Sum>i. f (A i \<inter> x)) + (\<Sum>i. f (A i - x))" by (rule add_mono)
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   483
    also have "\<dots> = (\<Sum>i. f (A i \<inter> x) + f (A i - x))"
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   484
      using A(2) x posf by (subst suminf_add_ereal) (auto simp: positive_def)
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   485
    also have "\<dots> = (\<Sum>i. f (A i))"
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   486
      using A x
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   487
      by (subst add[THEN additiveD, symmetric])
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   488
         (auto intro!: arg_cong[where f=suminf] arg_cong[where f=f])
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   489
    finally show "outer_measure M f (s \<inter> x) + outer_measure M f (s - x) \<le> (\<Sum>i. f (A i))" .
42066
6db76c88907a generalized Caratheodory from algebra to ring_of_sets
hoelzl
parents: 42065
diff changeset
   490
  qed
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   491
  moreover
61273
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   492
  have "outer_measure M f s \<le> outer_measure M f (s \<inter> x) + outer_measure M f (s - x)"
42145
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   493
  proof -
61273
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   494
    have "outer_measure M f s = outer_measure M f ((s \<inter> x) \<union> (s - x))"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   495
      by (metis Un_Diff_Int Un_commute)
61273
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   496
    also have "... \<le> outer_measure M f (s \<inter> x) + outer_measure M f (s - x)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   497
      apply (rule subadditiveD)
42145
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   498
      apply (rule ring_of_sets.countably_subadditive_subadditive [OF ring_of_sets_Pow])
61273
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   499
      apply (simp add: positive_def outer_measure_empty[OF posf] outer_measure_nonneg[OF posf])
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   500
      apply (rule countably_subadditive_outer_measure)
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
   501
      using s by (auto intro!: posf inc)
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   502
    finally show ?thesis .
42145
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   503
  qed
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 37591
diff changeset
   504
  ultimately
61273
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   505
  show "outer_measure M f (s \<inter> x) + outer_measure M f (s - x) = outer_measure M f s"
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   506
    by (rule order_antisym)
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   507
qed
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   508
61273
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   509
lemma measure_down: "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
   510
  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
   511
56994
8d5e5ec1cac3 fixed document generation for HOL-Probability
hoelzl
parents: 56212
diff changeset
   512
subsection {* Caratheodory's theorem *}
8d5e5ec1cac3 fixed document generation for HOL-Probability
hoelzl
parents: 56212
diff changeset
   513
47762
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   514
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
   515
  assumes posf: "positive M f" and ca: "countably_additive M f"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   516
  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
   517
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
   518
  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
   519
    by (metis additive_increasing ca countably_additive_additive posf)
61273
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   520
  let ?O = "outer_measure M f"
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   521
  def ls \<equiv> "lambda_system \<Omega> (Pow \<Omega>) ?O"
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   522
  have mls: "measure_space \<Omega> ls ?O"
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
   523
    using sigma_algebra.caratheodory_lemma
61273
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   524
            [OF sigma_algebra_Pow outer_measure_space_outer_measure [OF posf inc]]
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
   525
    by (simp add: ls_def)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   526
  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
   527
    by (simp add: measure_space_def)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   528
  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
   529
    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
   530
       (metis ca posf inc countably_additive_additive algebra_subset_lambda_system)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   531
  hence sgs_sb: "sigma_sets (\<Omega>) (M) \<subseteq> ls"
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   532
    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
   533
    by simp
61273
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   534
  have "measure_space \<Omega> (sigma_sets \<Omega> M) ?O"
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
   535
    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
   536
       (simp_all add: sgs_sb space_closed)
61273
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   537
  thus ?thesis using outer_measure_agrees [OF posf ca]
542a4d9bac96 Caratheodory: cleanup and modernisation
hoelzl
parents: 61032
diff changeset
   538
    by (intro exI[of _ ?O]) auto
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
   539
qed
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   540
42145
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   541
lemma (in ring_of_sets) caratheodory_empty_continuous:
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   542
  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
   543
  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
   544
  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
   545
proof (intro caratheodory' empty_continuous_imp_countably_additive f)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46731
diff changeset
   546
  show "\<forall>A\<in>M. f A \<noteq> \<infinity>" using fin by auto
42145
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   547
qed (rule cont)
8448713d48b7 proved caratheodory_empty_continuous
hoelzl
parents: 42067
diff changeset
   548
56994
8d5e5ec1cac3 fixed document generation for HOL-Probability
hoelzl
parents: 56212
diff changeset
   549
subsection {* Volumes *}
47762
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   550
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   551
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
   552
  "volume M f \<longleftrightarrow>
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   553
  (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
   554
  (\<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
   555
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   556
lemma volumeI:
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   557
  assumes "f {} = 0"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   558
  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
   559
  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
   560
  shows "volume M f"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   561
  using assms by (auto simp: volume_def)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   562
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   563
lemma volume_positive:
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   564
  "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
   565
  by (auto simp: volume_def)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   566
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   567
lemma volume_empty:
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   568
  "volume M f \<Longrightarrow> f {} = 0"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   569
  by (auto simp: volume_def)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   570
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   571
lemma volume_finite_additive:
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   572
  assumes "volume M f" 
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   573
  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
   574
  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
   575
proof -
52141
eff000cab70f weaker precendence of syntax for big intersection and union on sets
haftmann
parents: 51329
diff changeset
   576
  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
   577
    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
   578
  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
   579
    unfolding volume_def by blast
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   580
  also have "\<dots> = (\<Sum>i\<in>I. f (A i))"
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56994
diff changeset
   581
  proof (subst setsum.reindex_nontrivial)
47762
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   582
    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
   583
    with `disjoint_family_on A I` have "A i = {}"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   584
      by (auto simp: disjoint_family_on_def)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   585
    then show "f (A i) = 0"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   586
      using volume_empty[OF `volume M f`] by simp
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   587
  qed (auto intro: `finite I`)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   588
  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
   589
    by simp
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   590
qed
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   591
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   592
lemma (in ring_of_sets) volume_additiveI:
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   593
  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
   594
  assumes [simp]: "\<mu> {} = 0"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   595
  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
   596
  shows "volume M \<mu>"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   597
proof (unfold volume_def, safe)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   598
  fix C assume "finite C" "C \<subseteq> M" "disjoint C"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   599
  then show "\<mu> (\<Union>C) = setsum \<mu> C"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   600
  proof (induct C)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   601
    case (insert c C)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   602
    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
   603
      by (auto intro!: add simp: disjoint_def)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   604
    with insert show ?case
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   605
      by (simp add: disjoint_def)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   606
  qed simp
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   607
qed fact+
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   608
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   609
lemma (in semiring_of_sets) extend_volume:
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   610
  assumes "volume M \<mu>"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   611
  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
   612
proof -
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   613
  let ?R = generated_ring
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   614
  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
   615
    by (auto simp: generated_ring_def)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   616
  from bchoice[OF this] guess \<mu>' .. note \<mu>'_spec = this
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   617
  
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   618
  { 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
   619
    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
   620
    assume "\<Union>C = \<Union>D"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   621
    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
   622
    proof (intro setsum.cong refl)
47762
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   623
      fix d assume "d \<in> D"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   624
      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
   625
        using `d \<in> D` `\<Union>C = \<Union>D` by auto
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   626
      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
   627
      proof (rule volume_finite_additive)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   628
        { 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
   629
            using C D `d \<in> D` by auto }
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   630
        show "(\<Union>a\<in>C. a \<inter> d) \<in> M"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   631
          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
   632
        show "disjoint_family_on (\<lambda>a. a \<inter> d) C"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   633
          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
   634
      qed fact+
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   635
      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
   636
    qed }
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   637
  note split_sum = this
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   638
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   639
  { 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
   640
    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
   641
    assume "\<Union>C = \<Union>D"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   642
    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
   643
    have "(\<Sum>d\<in>D. \<mu> d) = (\<Sum>c\<in>C. \<mu> c)"
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56994
diff changeset
   644
      by (simp, subst setsum.commute, simp add: ac_simps) }
47762
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   645
  note sum_eq = this
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   646
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   647
  { 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
   648
    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
   649
    with \<mu>'_spec[THEN bspec, of "\<Union>C"]
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   650
    obtain D where
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   651
      D: "D \<subseteq> M" "finite D" "disjoint D" "\<Union>C = \<Union>D" and "\<mu>' (\<Union>C) = (\<Sum>d\<in>D. \<mu> d)"
61427
3c69ea85f8dd Fixed nonterminating "blast" proof
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
   652
      by auto
47762
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   653
    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
   654
  note \<mu>' = this
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   655
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   656
  show ?thesis
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   657
  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
   658
    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
   659
      by (simp add: disjoint_def)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   660
  next
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   661
    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
   662
    with \<mu>'[of Ca] `volume M \<mu>`[THEN volume_positive]
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   663
    show "0 \<le> \<mu>' a"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   664
      by (auto intro!: setsum_nonneg)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   665
  next
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   666
    show "\<mu>' {} = 0" using \<mu>'[of "{}"] by auto
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   667
  next
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   668
    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
   669
    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
   670
    assume "a \<inter> b = {}"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   671
    with Ca Cb have "Ca \<inter> Cb \<subseteq> {{}}" by auto
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   672
    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
   673
60585
48fdff264eb2 tuned whitespace;
wenzelm
parents: 58876
diff changeset
   674
    from `a \<inter> b = {}` have "\<mu>' (\<Union>(Ca \<union> Cb)) = (\<Sum>c\<in>Ca \<union> Cb. \<mu> c)"
47762
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   675
      using Ca Cb by (intro \<mu>') (auto intro!: disjoint_union)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   676
    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
   677
      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
   678
    also have "\<dots> = (\<Sum>c\<in>Ca. \<mu> c) + (\<Sum>c\<in>Cb. \<mu> c)"
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56994
diff changeset
   679
      using Ca Cb by (simp add: setsum.union_inter)
47762
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   680
    also have "\<dots> = \<mu>' a + \<mu>' b"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   681
      using Ca Cb by (simp add: \<mu>')
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   682
    finally show "\<mu>' (a \<union> b) = \<mu>' a + \<mu>' b"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   683
      using Ca Cb by simp
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   684
  qed
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   685
qed
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   686
56994
8d5e5ec1cac3 fixed document generation for HOL-Probability
hoelzl
parents: 56212
diff changeset
   687
subsubsection {* Caratheodory on semirings *}
47762
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
theorem (in semiring_of_sets) caratheodory:
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   690
  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
   691
  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
   692
proof -
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   693
  have "volume M \<mu>"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   694
  proof (rule volumeI)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   695
    { 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
   696
        using pos unfolding positive_def by auto }
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   697
    note p = this
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   698
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   699
    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
   700
    have "\<exists>F'. bij_betw F' {..<card C} C"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   701
      by (rule finite_same_card_bij[OF _ `finite C`]) auto
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   702
    then guess F' .. note F' = this
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   703
    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
   704
      by (auto simp: bij_betw_def)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   705
    { 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
   706
      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
   707
        unfolding inj_on_def by auto
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   708
      with `disjoint C`[THEN disjointD]
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   709
      have "F' i \<inter> F' j = {}"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   710
        by auto }
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   711
    note F'_disj = this
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   712
    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
   713
    then have "disjoint_family F"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   714
      using F'_disj by (auto simp: disjoint_family_on_def)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   715
    moreover from F' have "(\<Union>i. F i) = \<Union>C"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   716
      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
   717
    moreover have sets_F: "\<And>i. F i \<in> M"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   718
      using F' sets_C by (auto simp: F_def)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   719
    moreover note sets_C
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   720
    ultimately have "\<mu> (\<Union>C) = (\<Sum>i. \<mu> (F i))"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   721
      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
   722
    also have "\<dots> = (\<Sum>i<card C. \<mu> (F' i))"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   723
    proof -
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   724
      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
   725
        by (rule sums_If_finite_set) auto
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   726
      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
   727
        using pos by (auto simp: positive_def F_def)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   728
      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
   729
        by (simp add: sums_iff)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   730
    qed
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   731
    also have "\<dots> = (\<Sum>c\<in>C. \<mu> c)"
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56994
diff changeset
   732
      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
   733
    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
   734
  next
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   735
    show "\<mu> {} = 0"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   736
      using `positive M \<mu>` by (rule positiveD1)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   737
  qed
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   738
  from extend_volume[OF this] obtain \<mu>_r where
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   739
    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
   740
    by auto
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
  interpret G: ring_of_sets \<Omega> generated_ring
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   743
    by (rule generating_ring)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   744
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   745
  have pos: "positive generated_ring \<mu>_r"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   746
    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
   747
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   748
  have "countably_additive generated_ring \<mu>_r"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   749
  proof (rule countably_additiveI)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   750
    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
   751
      and Un_A: "(\<Union>i. A' i) \<in> generated_ring"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   752
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   753
    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
   754
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   755
    { fix c assume "c \<in> C'"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   756
      moreover def A \<equiv> "\<lambda>i. A' i \<inter> c"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   757
      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
   758
        and Un_A: "(\<Union>i. A i) \<in> generated_ring"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   759
        using A' C'
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   760
        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
   761
      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
   762
        by (auto simp: A_def)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   763
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   764
      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
   765
        (is "\<forall>i. ?P i")
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   766
      proof
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   767
        fix i
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   768
        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
   769
        from generated_ringE[OF this] guess C . note C = this
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   770
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   771
        have "\<exists>F'. bij_betw F' {..<card C} C"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   772
          by (rule finite_same_card_bij[OF _ `finite C`]) auto
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   773
        then guess F .. note F = this
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   774
        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
   775
        then have f: "bij_betw f {..< card C} C"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   776
          by (intro bij_betw_cong[THEN iffD1, OF _ F]) auto
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   777
        with C have "\<forall>j. f j \<in> M"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   778
          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
   779
        moreover
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   780
        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
   781
          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
   782
        then have "disjoint_family f"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   783
          by (auto simp: disjoint_family_on_def f_def)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   784
        moreover
60585
48fdff264eb2 tuned whitespace;
wenzelm
parents: 58876
diff changeset
   785
        have Ai_eq: "A i = (\<Union>x<card C. f x)"
47762
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   786
          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
   787
        then have "\<Union>range f = A i"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   788
          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
   789
        moreover 
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   790
        { 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
   791
            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
   792
          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
   793
            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
   794
          also have "\<dots> = \<mu>_r (A i)"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   795
            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
   796
            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
   797
               (auto simp: Pi_iff Ai_eq intro: generated_ringI_Basic)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   798
          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
   799
        ultimately show "?P i"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   800
          by blast
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   801
      qed
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   802
      from choice[OF this] guess f .. note f = this
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61273
diff changeset
   803
      then have UN_f_eq: "(\<Union>i. case_prod f (prod_decode i)) = (\<Union>i. A i)"
47762
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   804
        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
   805
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61273
diff changeset
   806
      have d: "disjoint_family (\<lambda>i. case_prod f (prod_decode i))"
47762
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   807
        unfolding disjoint_family_on_def
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   808
      proof (intro ballI impI)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   809
        fix m n :: nat assume "m \<noteq> n"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   810
        then have neq: "prod_decode m \<noteq> prod_decode n"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   811
          using inj_prod_decode[of UNIV] by (auto simp: inj_on_def)
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61273
diff changeset
   812
        show "case_prod f (prod_decode m) \<inter> case_prod f (prod_decode n) = {}"
47762
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   813
        proof cases
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   814
          assume "fst (prod_decode m) = fst (prod_decode n)"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   815
          then show ?thesis
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   816
            using neq f by (fastforce simp: disjoint_family_on_def)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   817
        next
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   818
          assume neq: "fst (prod_decode m) \<noteq> fst (prod_decode n)"
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61273
diff changeset
   819
          have "case_prod f (prod_decode m) \<subseteq> A (fst (prod_decode m))"
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61273
diff changeset
   820
            "case_prod f (prod_decode n) \<subseteq> A (fst (prod_decode n))"
47762
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   821
            using f[THEN spec, of "fst (prod_decode m)"]
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   822
            using f[THEN spec, of "fst (prod_decode n)"]
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   823
            by (auto simp: set_eq_iff)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   824
          with f A neq show ?thesis
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   825
            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
   826
        qed
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   827
      qed
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61273
diff changeset
   828
      from f have "(\<Sum>n. \<mu>_r (A n)) = (\<Sum>n. \<mu>_r (case_prod f (prod_decode n)))"
47762
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   829
        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
   830
         (auto split: prod.split)
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61273
diff changeset
   831
      also have "\<dots> = (\<Sum>n. \<mu> (case_prod f (prod_decode n)))"
47762
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   832
        using f V(2) by (auto intro!: arg_cong[where f=suminf] split: prod.split)
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61273
diff changeset
   833
      also have "\<dots> = \<mu> (\<Union>i. case_prod f (prod_decode i))"
47762
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   834
        using f `c \<in> C'` C'
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   835
        by (intro ca[unfolded countably_additive_def, rule_format])
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   836
           (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
   837
      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
   838
        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
   839
    note eq = this
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   840
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   841
    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
   842
      using C' A'
47762
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   843
      by (subst volume_finite_additive[symmetric, OF V(1)])
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 55642
diff changeset
   844
         (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
   845
               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
   846
               intro: generated_ringI_Basic)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   847
    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
   848
      using C' A'
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   849
      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
   850
         (auto intro: generated_ringI_Basic)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   851
    also have "\<dots> = (\<Sum>c\<in>C'. \<mu>_r c)"
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56994
diff changeset
   852
      using eq V C' by (auto intro!: setsum.cong)
47762
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   853
    also have "\<dots> = \<mu>_r (\<Union>C')"
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   854
      using C' Un_A
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   855
      by (subst volume_finite_additive[symmetric, OF V(1)])
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 55642
diff changeset
   856
         (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
   857
               intro: generated_ringI_Basic)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   858
    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
   859
      using C' by simp
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   860
  qed
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   861
  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
   862
  guess \<mu>' ..
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   863
  with V show ?thesis
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   864
    unfolding sigma_sets_generated_ring_eq
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   865
    by (intro exI[of _ \<mu>']) (auto intro: generated_ringI_Basic)
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   866
qed
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   867
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
   868
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
   869
  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
   870
  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
   871
  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
   872
  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
   873
  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
   874
  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
   875
  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
   876
  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
   877
    (\<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
   878
  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
   879
proof -
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
   880
  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
   881
    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
   882
  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
   883
    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
   884
  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
   885
    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
   886
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
   887
  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
   888
  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
   889
    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
   890
      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
   891
    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
   892
    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
   893
      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
   894
      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
   895
        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
   896
    qed
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
   897
  qed
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
   898
  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
   899
    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
   900
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
   901
  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
   902
  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
   903
    { 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
   904
      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
   905
    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
   906
      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
   907
    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
   908
      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
   909
  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
   910
qed
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
   911
  
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
   912
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
   913
  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
   914
  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
   915
  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
   916
  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
   917
  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
   918
  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
   919
  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
   920
  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
   921
    (\<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
   922
    (\<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
   923
  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
   924
proof -
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
   925
  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
   926
  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
   927
    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
   928
      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
   929
  next
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
   930
    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
   931
      "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
   932
      "(\<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
   933
    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
   934
      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
   935
      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
   936
  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
   937
  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
   938
qed
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
   939
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents:
diff changeset
   940
end