src/HOL/Probability/Measure_Space.thy
author Andreas Lochbihler
Wed, 11 Nov 2015 10:13:40 +0100
changeset 61633 64e6d712af16
parent 61609 77b453bd616f
child 61634 48e2de1b1df5
permissions -rw-r--r--
add lemmas
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
     1
(*  Title:      HOL/Probability/Measure_Space.thy
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
     2
    Author:     Lawrence C Paulson
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
     3
    Author:     Johannes Hölzl, TU München
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
     4
    Author:     Armin Heller, TU München
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
     5
*)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
     6
58876
1888e3cb8048 modernized header;
wenzelm
parents: 58606
diff changeset
     7
section {* Measure spaces and their properties *}
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
     8
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
     9
theory Measure_Space
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    10
imports
59593
304ee0a475d1 fix import path
hoelzl
parents: 59425
diff changeset
    11
  Measurable "~~/src/HOL/Multivariate_Analysis/Multivariate_Analysis"
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    12
begin
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    13
50104
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
    14
subsection "Relate extended reals and the indicator function"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
    15
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    16
lemma suminf_cmult_indicator:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    17
  fixes f :: "nat \<Rightarrow> ereal"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    18
  assumes "disjoint_family A" "x \<in> A i" "\<And>i. 0 \<le> f i"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    19
  shows "(\<Sum>n. f n * indicator (A n) x) = f i"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    20
proof -
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    21
  have **: "\<And>n. f n * indicator (A n) x = (if n = i then f n else 0 :: ereal)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    22
    using `x \<in> A i` assms unfolding disjoint_family_on_def indicator_def by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    23
  then have "\<And>n. (\<Sum>j<n. f j * indicator (A j) x) = (if i < n then f i else 0 :: ereal)"
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57276
diff changeset
    24
    by (auto simp: setsum.If_cases)
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    25
  moreover have "(SUP n. if i < n then f i else 0) = (f i :: ereal)"
51000
c9adb50f74ad use order topology for extended reals
hoelzl
parents: 50419
diff changeset
    26
  proof (rule SUP_eqI)
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    27
    fix y :: ereal assume "\<And>n. n \<in> UNIV \<Longrightarrow> (if i < n then f i else 0) \<le> y"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    28
    from this[of "Suc i"] show "f i \<le> y" by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    29
  qed (insert assms, simp)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    30
  ultimately show ?thesis using assms
56212
3253aaf73a01 consolidated theorem names containing INFI and SUPR: have INF and SUP instead uniformly
haftmann
parents: 56193
diff changeset
    31
    by (subst suminf_ereal_eq_SUP) (auto simp: indicator_def)
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    32
qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    33
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    34
lemma suminf_indicator:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    35
  assumes "disjoint_family A"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    36
  shows "(\<Sum>n. indicator (A n) x :: ereal) = indicator (\<Union>i. A i) x"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    37
proof cases
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    38
  assume *: "x \<in> (\<Union>i. A i)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    39
  then obtain i where "x \<in> A i" by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    40
  from suminf_cmult_indicator[OF assms(1), OF `x \<in> A i`, of "\<lambda>k. 1"]
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    41
  show ?thesis using * by simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    42
qed simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    43
60727
53697011b03a move disjoint sets to their own theory
hoelzl
parents: 60715
diff changeset
    44
lemma setsum_indicator_disjoint_family:
53697011b03a move disjoint sets to their own theory
hoelzl
parents: 60715
diff changeset
    45
  fixes f :: "'d \<Rightarrow> 'e::semiring_1"
53697011b03a move disjoint sets to their own theory
hoelzl
parents: 60715
diff changeset
    46
  assumes d: "disjoint_family_on A P" and "x \<in> A j" and "finite P" and "j \<in> P"
53697011b03a move disjoint sets to their own theory
hoelzl
parents: 60715
diff changeset
    47
  shows "(\<Sum>i\<in>P. f i * indicator (A i) x) = f j"
53697011b03a move disjoint sets to their own theory
hoelzl
parents: 60715
diff changeset
    48
proof -
53697011b03a move disjoint sets to their own theory
hoelzl
parents: 60715
diff changeset
    49
  have "P \<inter> {i. x \<in> A i} = {j}"
53697011b03a move disjoint sets to their own theory
hoelzl
parents: 60715
diff changeset
    50
    using d `x \<in> A j` `j \<in> P` unfolding disjoint_family_on_def
53697011b03a move disjoint sets to their own theory
hoelzl
parents: 60715
diff changeset
    51
    by auto
53697011b03a move disjoint sets to their own theory
hoelzl
parents: 60715
diff changeset
    52
  thus ?thesis
53697011b03a move disjoint sets to their own theory
hoelzl
parents: 60715
diff changeset
    53
    unfolding indicator_def
53697011b03a move disjoint sets to their own theory
hoelzl
parents: 60715
diff changeset
    54
    by (simp add: if_distrib setsum.If_cases[OF `finite P`])
53697011b03a move disjoint sets to their own theory
hoelzl
parents: 60715
diff changeset
    55
qed
53697011b03a move disjoint sets to their own theory
hoelzl
parents: 60715
diff changeset
    56
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    57
text {*
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    58
  The type for emeasure spaces is already defined in @{theory Sigma_Algebra}, as it is also used to
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    59
  represent sigma algebras (with an arbitrary emeasure).
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    60
*}
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    61
56994
8d5e5ec1cac3 fixed document generation for HOL-Probability
hoelzl
parents: 56212
diff changeset
    62
subsection "Extend binary sets"
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    63
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    64
lemma LIMSEQ_binaryset:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    65
  assumes f: "f {} = 0"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    66
  shows  "(\<lambda>n. \<Sum>i<n. f (binaryset A B i)) ----> f A + f B"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    67
proof -
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    68
  have "(\<lambda>n. \<Sum>i < Suc (Suc n). f (binaryset A B i)) = (\<lambda>n. f A + f B)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    69
    proof
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    70
      fix n
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    71
      show "(\<Sum>i < Suc (Suc n). f (binaryset A B i)) = f A + f B"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    72
        by (induct n)  (auto simp add: binaryset_def f)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    73
    qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    74
  moreover
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    75
  have "... ----> f A + f B" by (rule tendsto_const)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    76
  ultimately
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    77
  have "(\<lambda>n. \<Sum>i< Suc (Suc n). f (binaryset A B i)) ----> f A + f B"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    78
    by metis
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    79
  hence "(\<lambda>n. \<Sum>i< n+2. f (binaryset A B i)) ----> f A + f B"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    80
    by simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    81
  thus ?thesis by (rule LIMSEQ_offset [where k=2])
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    82
qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    83
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    84
lemma binaryset_sums:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    85
  assumes f: "f {} = 0"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    86
  shows  "(\<lambda>n. f (binaryset A B n)) sums (f A + f B)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    87
    by (simp add: sums_def LIMSEQ_binaryset [where f=f, OF f] atLeast0LessThan)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    88
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    89
lemma suminf_binaryset_eq:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    90
  fixes f :: "'a set \<Rightarrow> 'b::{comm_monoid_add, t2_space}"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    91
  shows "f {} = 0 \<Longrightarrow> (\<Sum>n. f (binaryset A B n)) = f A + f B"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    92
  by (metis binaryset_sums sums_unique)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    93
56994
8d5e5ec1cac3 fixed document generation for HOL-Probability
hoelzl
parents: 56212
diff changeset
    94
subsection {* Properties of a premeasure @{term \<mu>} *}
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    95
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    96
text {*
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    97
  The definitions for @{const positive} and @{const countably_additive} should be here, by they are
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    98
  necessary to define @{typ "'a measure"} in @{theory Sigma_Algebra}.
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
    99
*}
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   100
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   101
definition additive where
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   102
  "additive M \<mu> \<longleftrightarrow> (\<forall>x\<in>M. \<forall>y\<in>M. x \<inter> y = {} \<longrightarrow> \<mu> (x \<union> y) = \<mu> x + \<mu> y)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   103
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   104
definition increasing where
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   105
  "increasing M \<mu> \<longleftrightarrow> (\<forall>x\<in>M. \<forall>y\<in>M. x \<subseteq> y \<longrightarrow> \<mu> x \<le> \<mu> y)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   106
49773
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   107
lemma positiveD1: "positive M f \<Longrightarrow> f {} = 0" by (auto simp: positive_def)
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   108
lemma positiveD2: "positive M f \<Longrightarrow> A \<in> M \<Longrightarrow> 0 \<le> f A" by (auto simp: positive_def)
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   109
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   110
lemma positiveD_empty:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   111
  "positive M f \<Longrightarrow> f {} = 0"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   112
  by (auto simp add: positive_def)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   113
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   114
lemma additiveD:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   115
  "additive M f \<Longrightarrow> x \<inter> y = {} \<Longrightarrow> x \<in> M \<Longrightarrow> y \<in> M \<Longrightarrow> f (x \<union> y) = f x + f y"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   116
  by (auto simp add: additive_def)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   117
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   118
lemma increasingD:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   119
  "increasing M f \<Longrightarrow> x \<subseteq> y \<Longrightarrow> x\<in>M \<Longrightarrow> y\<in>M \<Longrightarrow> f x \<le> f y"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   120
  by (auto simp add: increasing_def)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   121
50104
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
   122
lemma countably_additiveI[case_names countably]:
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   123
  "(\<And>A. range A \<subseteq> M \<Longrightarrow> disjoint_family A \<Longrightarrow> (\<Union>i. A i) \<in> M \<Longrightarrow> (\<Sum>i. f (A i)) = f (\<Union>i. A i))
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   124
  \<Longrightarrow> countably_additive M f"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   125
  by (simp add: countably_additive_def)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   126
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   127
lemma (in ring_of_sets) disjointed_additive:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   128
  assumes f: "positive M f" "additive M f" and A: "range A \<subseteq> M" "incseq A"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   129
  shows "(\<Sum>i\<le>n. f (disjointed A i)) = f (A n)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   130
proof (induct n)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   131
  case (Suc n)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   132
  then have "(\<Sum>i\<le>Suc n. f (disjointed A i)) = f (A n) + f (disjointed A (Suc n))"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   133
    by simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   134
  also have "\<dots> = f (A n \<union> disjointed A (Suc n))"
60727
53697011b03a move disjoint sets to their own theory
hoelzl
parents: 60715
diff changeset
   135
    using A by (subst f(2)[THEN additiveD]) (auto simp: disjointed_mono)
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   136
  also have "A n \<union> disjointed A (Suc n) = A (Suc n)"
60727
53697011b03a move disjoint sets to their own theory
hoelzl
parents: 60715
diff changeset
   137
    using `incseq A` by (auto dest: incseq_SucD simp: disjointed_mono)
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   138
  finally show ?case .
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   139
qed simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   140
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   141
lemma (in ring_of_sets) additive_sum:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   142
  fixes A:: "'i \<Rightarrow> 'a set"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   143
  assumes f: "positive M f" and ad: "additive M f" and "finite S"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   144
      and A: "A`S \<subseteq> M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   145
      and disj: "disjoint_family_on A S"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   146
  shows  "(\<Sum>i\<in>S. f (A i)) = f (\<Union>i\<in>S. A i)"
53374
a14d2a854c02 tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents: 51351
diff changeset
   147
  using `finite S` disj A
a14d2a854c02 tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents: 51351
diff changeset
   148
proof induct
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   149
  case empty show ?case using f by (simp add: positive_def)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   150
next
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   151
  case (insert s S)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   152
  then have "A s \<inter> (\<Union>i\<in>S. A i) = {}"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   153
    by (auto simp add: disjoint_family_on_def neq_iff)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   154
  moreover
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   155
  have "A s \<in> M" using insert by blast
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   156
  moreover have "(\<Union>i\<in>S. A i) \<in> M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   157
    using insert `finite S` by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   158
  ultimately have "f (A s \<union> (\<Union>i\<in>S. A i)) = f (A s) + f(\<Union>i\<in>S. A i)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   159
    using ad UNION_in_sets A by (auto simp add: additive_def)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   160
  with insert show ?case using ad disjoint_family_on_mono[of S "insert s S" A]
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   161
    by (auto simp add: additive_def subset_insertI)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   162
qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   163
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   164
lemma (in ring_of_sets) additive_increasing:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   165
  assumes posf: "positive M f" and addf: "additive M f"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   166
  shows "increasing M f"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   167
proof (auto simp add: increasing_def)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   168
  fix x y
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   169
  assume xy: "x \<in> M" "y \<in> M" "x \<subseteq> y"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   170
  then have "y - x \<in> M" by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   171
  then have "0 \<le> f (y-x)" using posf[unfolded positive_def] by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   172
  then have "f x + 0 \<le> f x + f (y-x)" by (intro add_left_mono) auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   173
  also have "... = f (x \<union> (y-x))" using addf
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   174
    by (auto simp add: additive_def) (metis Diff_disjoint Un_Diff_cancel Diff xy(1,2))
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   175
  also have "... = f y"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   176
    by (metis Un_Diff_cancel Un_absorb1 xy(3))
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   177
  finally show "f x \<le> f y" by simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   178
qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   179
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 50002
diff changeset
   180
lemma (in ring_of_sets) subadditive:
635d73673b5e regularity of measures, therefore:
immler
parents: 50002
diff changeset
   181
  assumes f: "positive M f" "additive M f" and A: "range A \<subseteq> M" and S: "finite S"
635d73673b5e regularity of measures, therefore:
immler
parents: 50002
diff changeset
   182
  shows "f (\<Union>i\<in>S. A i) \<le> (\<Sum>i\<in>S. f (A i))"
635d73673b5e regularity of measures, therefore:
immler
parents: 50002
diff changeset
   183
using S
635d73673b5e regularity of measures, therefore:
immler
parents: 50002
diff changeset
   184
proof (induct S)
635d73673b5e regularity of measures, therefore:
immler
parents: 50002
diff changeset
   185
  case empty thus ?case using f by (auto simp: positive_def)
635d73673b5e regularity of measures, therefore:
immler
parents: 50002
diff changeset
   186
next
635d73673b5e regularity of measures, therefore:
immler
parents: 50002
diff changeset
   187
  case (insert x F)
60585
48fdff264eb2 tuned whitespace;
wenzelm
parents: 60580
diff changeset
   188
  hence in_M: "A x \<in> M" "(\<Union>i\<in>F. A i) \<in> M" "(\<Union>i\<in>F. A i) - A x \<in> M" using A by force+
48fdff264eb2 tuned whitespace;
wenzelm
parents: 60580
diff changeset
   189
  have subs: "(\<Union>i\<in>F. A i) - A x \<subseteq> (\<Union>i\<in>F. A i)" by auto
48fdff264eb2 tuned whitespace;
wenzelm
parents: 60580
diff changeset
   190
  have "(\<Union>i\<in>(insert x F). A i) = A x \<union> ((\<Union>i\<in>F. A i) - A x)" by auto
48fdff264eb2 tuned whitespace;
wenzelm
parents: 60580
diff changeset
   191
  hence "f (\<Union>i\<in>(insert x F). A i) = f (A x \<union> ((\<Union>i\<in>F. A i) - A x))"
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 50002
diff changeset
   192
    by simp
60585
48fdff264eb2 tuned whitespace;
wenzelm
parents: 60580
diff changeset
   193
  also have "\<dots> = f (A x) + f ((\<Union>i\<in>F. A i) - A x)"
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 50002
diff changeset
   194
    using f(2) by (rule additiveD) (insert in_M, auto)
60585
48fdff264eb2 tuned whitespace;
wenzelm
parents: 60580
diff changeset
   195
  also have "\<dots> \<le> f (A x) + f (\<Union>i\<in>F. A i)"
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 50002
diff changeset
   196
    using additive_increasing[OF f] in_M subs by (auto simp: increasing_def intro: add_left_mono)
635d73673b5e regularity of measures, therefore:
immler
parents: 50002
diff changeset
   197
  also have "\<dots> \<le> f (A x) + (\<Sum>i\<in>F. f (A i))" using insert by (auto intro: add_left_mono)
60585
48fdff264eb2 tuned whitespace;
wenzelm
parents: 60580
diff changeset
   198
  finally show "f (\<Union>i\<in>(insert x F). A i) \<le> (\<Sum>i\<in>(insert x F). f (A i))" using insert by simp
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 50002
diff changeset
   199
qed
635d73673b5e regularity of measures, therefore:
immler
parents: 50002
diff changeset
   200
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   201
lemma (in ring_of_sets) countably_additive_additive:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   202
  assumes posf: "positive M f" and ca: "countably_additive M f"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   203
  shows "additive M f"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   204
proof (auto simp add: additive_def)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   205
  fix x y
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   206
  assume x: "x \<in> M" and y: "y \<in> M" and "x \<inter> y = {}"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   207
  hence "disjoint_family (binaryset x y)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   208
    by (auto simp add: disjoint_family_on_def binaryset_def)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   209
  hence "range (binaryset x y) \<subseteq> M \<longrightarrow>
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   210
         (\<Union>i. binaryset x y i) \<in> M \<longrightarrow>
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   211
         f (\<Union>i. binaryset x y i) = (\<Sum> n. f (binaryset x y n))"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   212
    using ca
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   213
    by (simp add: countably_additive_def)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   214
  hence "{x,y,{}} \<subseteq> M \<longrightarrow> x \<union> y \<in> M \<longrightarrow>
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   215
         f (x \<union> y) = (\<Sum>n. f (binaryset x y n))"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   216
    by (simp add: range_binaryset_eq UN_binaryset_eq)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   217
  thus "f (x \<union> y) = f x + f y" using posf x y
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   218
    by (auto simp add: Un suminf_binaryset_eq positive_def)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   219
qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   220
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   221
lemma (in algebra) increasing_additive_bound:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   222
  fixes A:: "nat \<Rightarrow> 'a set" and  f :: "'a set \<Rightarrow> ereal"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   223
  assumes f: "positive M f" and ad: "additive M f"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   224
      and inc: "increasing M f"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   225
      and A: "range A \<subseteq> M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   226
      and disj: "disjoint_family A"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   227
  shows  "(\<Sum>i. f (A i)) \<le> f \<Omega>"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   228
proof (safe intro!: suminf_bound)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   229
  fix N
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   230
  note disj_N = disjoint_family_on_mono[OF _ disj, of "{..<N}"]
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   231
  have "(\<Sum>i<N. f (A i)) = f (\<Union>i\<in>{..<N}. A i)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   232
    using A by (intro additive_sum [OF f ad _ _]) (auto simp: disj_N)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   233
  also have "... \<le> f \<Omega>" using space_closed A
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   234
    by (intro increasingD[OF inc] finite_UN) auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   235
  finally show "(\<Sum>i<N. f (A i)) \<le> f \<Omega>" by simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   236
qed (insert f A, auto simp: positive_def)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   237
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   238
lemma (in ring_of_sets) countably_additiveI_finite:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   239
  assumes "finite \<Omega>" "positive M \<mu>" "additive M \<mu>"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   240
  shows "countably_additive M \<mu>"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   241
proof (rule countably_additiveI)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   242
  fix F :: "nat \<Rightarrow> 'a set" assume F: "range F \<subseteq> M" "(\<Union>i. F i) \<in> M" and disj: "disjoint_family F"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   243
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   244
  have "\<forall>i\<in>{i. F i \<noteq> {}}. \<exists>x. x \<in> F i" by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   245
  from bchoice[OF this] obtain f where f: "\<And>i. F i \<noteq> {} \<Longrightarrow> f i \<in> F i" by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   246
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   247
  have inj_f: "inj_on f {i. F i \<noteq> {}}"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   248
  proof (rule inj_onI, simp)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   249
    fix i j a b assume *: "f i = f j" "F i \<noteq> {}" "F j \<noteq> {}"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   250
    then have "f i \<in> F i" "f j \<in> F j" using f by force+
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   251
    with disj * show "i = j" by (auto simp: disjoint_family_on_def)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   252
  qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   253
  have "finite (\<Union>i. F i)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   254
    by (metis F(2) assms(1) infinite_super sets_into_space)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   255
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   256
  have F_subset: "{i. \<mu> (F i) \<noteq> 0} \<subseteq> {i. F i \<noteq> {}}"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   257
    by (auto simp: positiveD_empty[OF `positive M \<mu>`])
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   258
  moreover have fin_not_empty: "finite {i. F i \<noteq> {}}"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   259
  proof (rule finite_imageD)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   260
    from f have "f`{i. F i \<noteq> {}} \<subseteq> (\<Union>i. F i)" by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   261
    then show "finite (f`{i. F i \<noteq> {}})"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   262
      by (rule finite_subset) fact
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   263
  qed fact
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   264
  ultimately have fin_not_0: "finite {i. \<mu> (F i) \<noteq> 0}"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   265
    by (rule finite_subset)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   266
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   267
  have disj_not_empty: "disjoint_family_on F {i. F i \<noteq> {}}"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   268
    using disj by (auto simp: disjoint_family_on_def)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   269
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   270
  from fin_not_0 have "(\<Sum>i. \<mu> (F i)) = (\<Sum>i | \<mu> (F i) \<noteq> 0. \<mu> (F i))"
47761
dfe747e72fa8 moved lemmas to appropriate places
hoelzl
parents: 47694
diff changeset
   271
    by (rule suminf_finite) auto
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   272
  also have "\<dots> = (\<Sum>i | F i \<noteq> {}. \<mu> (F i))"
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57276
diff changeset
   273
    using fin_not_empty F_subset by (rule setsum.mono_neutral_left) auto
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   274
  also have "\<dots> = \<mu> (\<Union>i\<in>{i. F i \<noteq> {}}. F i)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   275
    using `positive M \<mu>` `additive M \<mu>` fin_not_empty disj_not_empty F by (intro additive_sum) auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   276
  also have "\<dots> = \<mu> (\<Union>i. F i)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   277
    by (rule arg_cong[where f=\<mu>]) auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   278
  finally show "(\<Sum>i. \<mu> (F i)) = \<mu> (\<Union>i. F i)" .
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   279
qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   280
49773
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   281
lemma (in ring_of_sets) countably_additive_iff_continuous_from_below:
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   282
  assumes f: "positive M f" "additive M f"
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   283
  shows "countably_additive M f \<longleftrightarrow>
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   284
    (\<forall>A. range A \<subseteq> M \<longrightarrow> incseq A \<longrightarrow> (\<Union>i. A i) \<in> M \<longrightarrow> (\<lambda>i. f (A i)) ----> f (\<Union>i. A i))"
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   285
  unfolding countably_additive_def
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   286
proof safe
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   287
  assume count_sum: "\<forall>A. range A \<subseteq> M \<longrightarrow> disjoint_family A \<longrightarrow> UNION UNIV A \<in> M \<longrightarrow> (\<Sum>i. f (A i)) = f (UNION UNIV A)"
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   288
  fix A :: "nat \<Rightarrow> 'a set" assume A: "range A \<subseteq> M" "incseq A" "(\<Union>i. A i) \<in> M"
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   289
  then have dA: "range (disjointed A) \<subseteq> M" by (auto simp: range_disjointed_sets)
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   290
  with count_sum[THEN spec, of "disjointed A"] A(3)
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   291
  have f_UN: "(\<Sum>i. f (disjointed A i)) = f (\<Union>i. A i)"
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   292
    by (auto simp: UN_disjointed_eq disjoint_family_disjointed)
56193
c726ecfb22b6 cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
hoelzl
parents: 56154
diff changeset
   293
  moreover have "(\<lambda>n. (\<Sum>i<n. f (disjointed A i))) ----> (\<Sum>i. f (disjointed A i))"
49773
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   294
    using f(1)[unfolded positive_def] dA
56193
c726ecfb22b6 cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
hoelzl
parents: 56154
diff changeset
   295
    by (auto intro!: summable_LIMSEQ summable_ereal_pos)
49773
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   296
  from LIMSEQ_Suc[OF this]
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   297
  have "(\<lambda>n. (\<Sum>i\<le>n. f (disjointed A i))) ----> (\<Sum>i. f (disjointed A i))"
56193
c726ecfb22b6 cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
hoelzl
parents: 56154
diff changeset
   298
    unfolding lessThan_Suc_atMost .
49773
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   299
  moreover have "\<And>n. (\<Sum>i\<le>n. f (disjointed A i)) = f (A n)"
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   300
    using disjointed_additive[OF f A(1,2)] .
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   301
  ultimately show "(\<lambda>i. f (A i)) ----> f (\<Union>i. A i)" by simp
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   302
next
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   303
  assume cont: "\<forall>A. range A \<subseteq> M \<longrightarrow> incseq A \<longrightarrow> (\<Union>i. A i) \<in> M \<longrightarrow> (\<lambda>i. f (A i)) ----> f (\<Union>i. A i)"
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   304
  fix A :: "nat \<Rightarrow> 'a set" assume A: "range A \<subseteq> M" "disjoint_family A" "(\<Union>i. A i) \<in> M"
57446
06e195515deb some lemmas about the indicator function; removed lemma sums_def2
hoelzl
parents: 57418
diff changeset
   305
  have *: "(\<Union>n. (\<Union>i<n. A i)) = (\<Union>i. A i)" by auto
06e195515deb some lemmas about the indicator function; removed lemma sums_def2
hoelzl
parents: 57418
diff changeset
   306
  have "(\<lambda>n. f (\<Union>i<n. A i)) ----> f (\<Union>i. A i)"
49773
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   307
  proof (unfold *[symmetric], intro cont[rule_format])
60585
48fdff264eb2 tuned whitespace;
wenzelm
parents: 60580
diff changeset
   308
    show "range (\<lambda>i. \<Union>i<i. A i) \<subseteq> M" "(\<Union>i. \<Union>i<i. A i) \<in> M"
49773
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   309
      using A * by auto
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   310
  qed (force intro!: incseq_SucI)
57446
06e195515deb some lemmas about the indicator function; removed lemma sums_def2
hoelzl
parents: 57418
diff changeset
   311
  moreover have "\<And>n. f (\<Union>i<n. A i) = (\<Sum>i<n. f (A i))"
49773
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   312
    using A
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   313
    by (intro additive_sum[OF f, of _ A, symmetric])
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   314
       (auto intro: disjoint_family_on_mono[where B=UNIV])
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   315
  ultimately
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   316
  have "(\<lambda>i. f (A i)) sums f (\<Union>i. A i)"
57446
06e195515deb some lemmas about the indicator function; removed lemma sums_def2
hoelzl
parents: 57418
diff changeset
   317
    unfolding sums_def by simp
49773
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   318
  from sums_unique[OF this]
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   319
  show "(\<Sum>i. f (A i)) = f (\<Union>i. A i)" by simp
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   320
qed
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   321
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   322
lemma (in ring_of_sets) continuous_from_above_iff_empty_continuous:
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   323
  assumes f: "positive M f" "additive M f"
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   324
  shows "(\<forall>A. range A \<subseteq> M \<longrightarrow> decseq A \<longrightarrow> (\<Inter>i. A i) \<in> M \<longrightarrow> (\<forall>i. f (A i) \<noteq> \<infinity>) \<longrightarrow> (\<lambda>i. f (A i)) ----> f (\<Inter>i. A i))
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   325
     \<longleftrightarrow> (\<forall>A. range A \<subseteq> M \<longrightarrow> decseq A \<longrightarrow> (\<Inter>i. A i) = {} \<longrightarrow> (\<forall>i. f (A i) \<noteq> \<infinity>) \<longrightarrow> (\<lambda>i. f (A i)) ----> 0)"
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   326
proof safe
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   327
  assume cont: "(\<forall>A. range A \<subseteq> M \<longrightarrow> decseq A \<longrightarrow> (\<Inter>i. A i) \<in> M \<longrightarrow> (\<forall>i. f (A i) \<noteq> \<infinity>) \<longrightarrow> (\<lambda>i. f (A i)) ----> f (\<Inter>i. A i))"
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   328
  fix A :: "nat \<Rightarrow> 'a set" assume A: "range A \<subseteq> M" "decseq A" "(\<Inter>i. A i) = {}" "\<forall>i. f (A i) \<noteq> \<infinity>"
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   329
  with cont[THEN spec, of A] show "(\<lambda>i. f (A i)) ----> 0"
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   330
    using `positive M f`[unfolded positive_def] by auto
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   331
next
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   332
  assume cont: "\<forall>A. range A \<subseteq> M \<longrightarrow> decseq A \<longrightarrow> (\<Inter>i. A i) = {} \<longrightarrow> (\<forall>i. f (A i) \<noteq> \<infinity>) \<longrightarrow> (\<lambda>i. f (A i)) ----> 0"
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   333
  fix A :: "nat \<Rightarrow> 'a set" assume A: "range A \<subseteq> M" "decseq A" "(\<Inter>i. A i) \<in> M" "\<forall>i. f (A i) \<noteq> \<infinity>"
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   334
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   335
  have f_mono: "\<And>a b. a \<in> M \<Longrightarrow> b \<in> M \<Longrightarrow> a \<subseteq> b \<Longrightarrow> f a \<le> f b"
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   336
    using additive_increasing[OF f] unfolding increasing_def by simp
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   337
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   338
  have decseq_fA: "decseq (\<lambda>i. f (A i))"
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   339
    using A by (auto simp: decseq_def intro!: f_mono)
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   340
  have decseq: "decseq (\<lambda>i. A i - (\<Inter>i. A i))"
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   341
    using A by (auto simp: decseq_def)
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   342
  then have decseq_f: "decseq (\<lambda>i. f (A i - (\<Inter>i. A i)))"
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   343
    using A unfolding decseq_def by (auto intro!: f_mono Diff)
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   344
  have "f (\<Inter>x. A x) \<le> f (A 0)"
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   345
    using A by (auto intro!: f_mono)
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   346
  then have f_Int_fin: "f (\<Inter>x. A x) \<noteq> \<infinity>"
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   347
    using A by auto
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   348
  { fix i
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   349
    have "f (A i - (\<Inter>i. A i)) \<le> f (A i)" using A by (auto intro!: f_mono)
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   350
    then have "f (A i - (\<Inter>i. A i)) \<noteq> \<infinity>"
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   351
      using A by auto }
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   352
  note f_fin = this
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   353
  have "(\<lambda>i. f (A i - (\<Inter>i. A i))) ----> 0"
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   354
  proof (intro cont[rule_format, OF _ decseq _ f_fin])
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   355
    show "range (\<lambda>i. A i - (\<Inter>i. A i)) \<subseteq> M" "(\<Inter>i. A i - (\<Inter>i. A i)) = {}"
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   356
      using A by auto
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   357
  qed
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   358
  from INF_Lim_ereal[OF decseq_f this]
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   359
  have "(INF n. f (A n - (\<Inter>i. A i))) = 0" .
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   360
  moreover have "(INF n. f (\<Inter>i. A i)) = f (\<Inter>i. A i)"
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   361
    by auto
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   362
  ultimately have "(INF n. f (A n - (\<Inter>i. A i)) + f (\<Inter>i. A i)) = 0 + f (\<Inter>i. A i)"
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   363
    using A(4) f_fin f_Int_fin
56212
3253aaf73a01 consolidated theorem names containing INFI and SUPR: have INF and SUP instead uniformly
haftmann
parents: 56193
diff changeset
   364
    by (subst INF_ereal_add) (auto simp: decseq_f)
49773
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   365
  moreover {
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   366
    fix n
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   367
    have "f (A n - (\<Inter>i. A i)) + f (\<Inter>i. A i) = f ((A n - (\<Inter>i. A i)) \<union> (\<Inter>i. A i))"
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   368
      using A by (subst f(2)[THEN additiveD]) auto
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   369
    also have "(A n - (\<Inter>i. A i)) \<union> (\<Inter>i. A i) = A n"
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   370
      by auto
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   371
    finally have "f (A n - (\<Inter>i. A i)) + f (\<Inter>i. A i) = f (A n)" . }
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   372
  ultimately have "(INF n. f (A n)) = f (\<Inter>i. A i)"
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   373
    by simp
51351
dd1dd470690b generalized lemmas in Extended_Real_Limits
hoelzl
parents: 51000
diff changeset
   374
  with LIMSEQ_INF[OF decseq_fA]
49773
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   375
  show "(\<lambda>i. f (A i)) ----> f (\<Inter>i. A i)" by simp
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   376
qed
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   377
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   378
lemma (in ring_of_sets) empty_continuous_imp_continuous_from_below:
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   379
  assumes f: "positive M f" "additive M f" "\<forall>A\<in>M. f A \<noteq> \<infinity>"
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   380
  assumes cont: "\<forall>A. range A \<subseteq> M \<longrightarrow> decseq A \<longrightarrow> (\<Inter>i. A i) = {} \<longrightarrow> (\<lambda>i. f (A i)) ----> 0"
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   381
  assumes A: "range A \<subseteq> M" "incseq A" "(\<Union>i. A i) \<in> M"
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   382
  shows "(\<lambda>i. f (A i)) ----> f (\<Union>i. A i)"
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   383
proof -
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   384
  have "\<forall>A\<in>M. \<exists>x. f A = ereal x"
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   385
  proof
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   386
    fix A assume "A \<in> M" with f show "\<exists>x. f A = ereal x"
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   387
      unfolding positive_def by (cases "f A") auto
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   388
  qed
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   389
  from bchoice[OF this] guess f' .. note f' = this[rule_format]
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   390
  from A have "(\<lambda>i. f ((\<Union>i. A i) - A i)) ----> 0"
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   391
    by (intro cont[rule_format]) (auto simp: decseq_def incseq_def)
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   392
  moreover
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   393
  { fix i
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   394
    have "f ((\<Union>i. A i) - A i) + f (A i) = f ((\<Union>i. A i) - A i \<union> A i)"
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   395
      using A by (intro f(2)[THEN additiveD, symmetric]) auto
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   396
    also have "(\<Union>i. A i) - A i \<union> A i = (\<Union>i. A i)"
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   397
      by auto
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   398
    finally have "f' (\<Union>i. A i) - f' (A i) = f' ((\<Union>i. A i) - A i)"
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   399
      using A by (subst (asm) (1 2 3) f') auto
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   400
    then have "f ((\<Union>i. A i) - A i) = ereal (f' (\<Union>i. A i) - f' (A i))"
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   401
      using A f' by auto }
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   402
  ultimately have "(\<lambda>i. f' (\<Union>i. A i) - f' (A i)) ----> 0"
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   403
    by (simp add: zero_ereal_def)
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   404
  then have "(\<lambda>i. f' (A i)) ----> f' (\<Union>i. A i)"
60142
3275dddf356f fixes for limits
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
   405
    by (rule Lim_transform2[OF tendsto_const])
49773
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   406
  then show "(\<lambda>i. f (A i)) ----> f (\<Union>i. A i)"
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   407
    using A by (subst (1 2) f') auto
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   408
qed
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   409
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   410
lemma (in ring_of_sets) empty_continuous_imp_countably_additive:
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   411
  assumes f: "positive M f" "additive M f" and fin: "\<forall>A\<in>M. f A \<noteq> \<infinity>"
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   412
  assumes cont: "\<And>A. range A \<subseteq> M \<Longrightarrow> decseq A \<Longrightarrow> (\<Inter>i. A i) = {} \<Longrightarrow> (\<lambda>i. f (A i)) ----> 0"
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   413
  shows "countably_additive M f"
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   414
  using countably_additive_iff_continuous_from_below[OF f]
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   415
  using empty_continuous_imp_continuous_from_below[OF f fin] cont
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   416
  by blast
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   417
56994
8d5e5ec1cac3 fixed document generation for HOL-Probability
hoelzl
parents: 56212
diff changeset
   418
subsection {* Properties of @{const emeasure} *}
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   419
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   420
lemma emeasure_positive: "positive (sets M) (emeasure M)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   421
  by (cases M) (auto simp: sets_def emeasure_def Abs_measure_inverse measure_space_def)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   422
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   423
lemma emeasure_empty[simp, intro]: "emeasure M {} = 0"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   424
  using emeasure_positive[of M] by (simp add: positive_def)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   425
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   426
lemma emeasure_nonneg[intro!]: "0 \<le> emeasure M A"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   427
  using emeasure_notin_sets[of A M] emeasure_positive[of M]
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   428
  by (cases "A \<in> sets M") (auto simp: positive_def)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   429
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   430
lemma emeasure_not_MInf[simp]: "emeasure M A \<noteq> - \<infinity>"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   431
  using emeasure_nonneg[of M A] by auto
50419
3177d0374701 add exponential and uniform distributions
hoelzl
parents: 50387
diff changeset
   432
3177d0374701 add exponential and uniform distributions
hoelzl
parents: 50387
diff changeset
   433
lemma emeasure_le_0_iff: "emeasure M A \<le> 0 \<longleftrightarrow> emeasure M A = 0"
3177d0374701 add exponential and uniform distributions
hoelzl
parents: 50387
diff changeset
   434
  using emeasure_nonneg[of M A] by auto
3177d0374701 add exponential and uniform distributions
hoelzl
parents: 50387
diff changeset
   435
3177d0374701 add exponential and uniform distributions
hoelzl
parents: 50387
diff changeset
   436
lemma emeasure_less_0_iff: "emeasure M A < 0 \<longleftrightarrow> False"
3177d0374701 add exponential and uniform distributions
hoelzl
parents: 50387
diff changeset
   437
  using emeasure_nonneg[of M A] by auto
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
   438
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
   439
lemma emeasure_single_in_space: "emeasure M {x} \<noteq> 0 \<Longrightarrow> x \<in> space M"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
   440
  using emeasure_notin_sets[of "{x}" M] by (auto dest: sets.sets_into_space)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
   441
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   442
lemma emeasure_countably_additive: "countably_additive (sets M) (emeasure M)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   443
  by (cases M) (auto simp: sets_def emeasure_def Abs_measure_inverse measure_space_def)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   444
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   445
lemma suminf_emeasure:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   446
  "range A \<subseteq> sets M \<Longrightarrow> disjoint_family A \<Longrightarrow> (\<Sum>i. emeasure M (A i)) = emeasure M (\<Union>i. A i)"
50244
de72bbe42190 qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents: 50104
diff changeset
   447
  using sets.countable_UN[of A UNIV M] emeasure_countably_additive[of M]
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   448
  by (simp add: countably_additive_def)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   449
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
   450
lemma sums_emeasure:
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
   451
  "disjoint_family F \<Longrightarrow> (\<And>i. F i \<in> sets M) \<Longrightarrow> (\<lambda>i. emeasure M (F i)) sums emeasure M (\<Union>i. F i)"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
   452
  unfolding sums_iff by (intro conjI summable_ereal_pos emeasure_nonneg suminf_emeasure) auto
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
   453
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   454
lemma emeasure_additive: "additive (sets M) (emeasure M)"
50244
de72bbe42190 qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents: 50104
diff changeset
   455
  by (metis sets.countably_additive_additive emeasure_positive emeasure_countably_additive)
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   456
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   457
lemma plus_emeasure:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   458
  "a \<in> sets M \<Longrightarrow> b \<in> sets M \<Longrightarrow> a \<inter> b = {} \<Longrightarrow> emeasure M a + emeasure M b = emeasure M (a \<union> b)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   459
  using additiveD[OF emeasure_additive] ..
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   460
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   461
lemma setsum_emeasure:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   462
  "F`I \<subseteq> sets M \<Longrightarrow> disjoint_family_on F I \<Longrightarrow> finite I \<Longrightarrow>
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   463
    (\<Sum>i\<in>I. emeasure M (F i)) = emeasure M (\<Union>i\<in>I. F i)"
50244
de72bbe42190 qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents: 50104
diff changeset
   464
  by (metis sets.additive_sum emeasure_positive emeasure_additive)
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   465
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   466
lemma emeasure_mono:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   467
  "a \<subseteq> b \<Longrightarrow> b \<in> sets M \<Longrightarrow> emeasure M a \<le> emeasure M b"
50244
de72bbe42190 qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents: 50104
diff changeset
   468
  by (metis sets.additive_increasing emeasure_additive emeasure_nonneg emeasure_notin_sets
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   469
            emeasure_positive increasingD)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   470
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   471
lemma emeasure_space:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   472
  "emeasure M A \<le> emeasure M (space M)"
50244
de72bbe42190 qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents: 50104
diff changeset
   473
  by (metis emeasure_mono emeasure_nonneg emeasure_notin_sets sets.sets_into_space sets.top)
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   474
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   475
lemma emeasure_compl:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   476
  assumes s: "s \<in> sets M" and fin: "emeasure M s \<noteq> \<infinity>"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   477
  shows "emeasure M (space M - s) = emeasure M (space M) - emeasure M s"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   478
proof -
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   479
  from s have "0 \<le> emeasure M s" by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   480
  have "emeasure M (space M) = emeasure M (s \<union> (space M - s))" using s
50244
de72bbe42190 qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents: 50104
diff changeset
   481
    by (metis Un_Diff_cancel Un_absorb1 s sets.sets_into_space)
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   482
  also have "... = emeasure M s + emeasure M (space M - s)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   483
    by (rule plus_emeasure[symmetric]) (auto simp add: s)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   484
  finally have "emeasure M (space M) = emeasure M s + emeasure M (space M - s)" .
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   485
  then show ?thesis
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   486
    using fin `0 \<le> emeasure M s`
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   487
    unfolding ereal_eq_minus_iff by (auto simp: ac_simps)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   488
qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   489
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   490
lemma emeasure_Diff:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   491
  assumes finite: "emeasure M B \<noteq> \<infinity>"
50002
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
   492
  and [measurable]: "A \<in> sets M" "B \<in> sets M" and "B \<subseteq> A"
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   493
  shows "emeasure M (A - B) = emeasure M A - emeasure M B"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   494
proof -
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   495
  have "0 \<le> emeasure M B" using assms by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   496
  have "(A - B) \<union> B = A" using `B \<subseteq> A` by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   497
  then have "emeasure M A = emeasure M ((A - B) \<union> B)" by simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   498
  also have "\<dots> = emeasure M (A - B) + emeasure M B"
50002
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
   499
    by (subst plus_emeasure[symmetric]) auto
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   500
  finally show "emeasure M (A - B) = emeasure M A - emeasure M B"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   501
    unfolding ereal_eq_minus_iff
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   502
    using finite `0 \<le> emeasure M B` by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   503
qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   504
49773
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   505
lemma Lim_emeasure_incseq:
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   506
  "range A \<subseteq> sets M \<Longrightarrow> incseq A \<Longrightarrow> (\<lambda>i. (emeasure M (A i))) ----> emeasure M (\<Union>i. A i)"
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   507
  using emeasure_countably_additive
50244
de72bbe42190 qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents: 50104
diff changeset
   508
  by (auto simp add: sets.countably_additive_iff_continuous_from_below emeasure_positive
de72bbe42190 qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents: 50104
diff changeset
   509
    emeasure_additive)
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   510
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   511
lemma incseq_emeasure:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   512
  assumes "range B \<subseteq> sets M" "incseq B"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   513
  shows "incseq (\<lambda>i. emeasure M (B i))"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   514
  using assms by (auto simp: incseq_def intro!: emeasure_mono)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   515
49773
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   516
lemma SUP_emeasure_incseq:
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   517
  assumes A: "range A \<subseteq> sets M" "incseq A"
49773
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   518
  shows "(SUP n. emeasure M (A n)) = emeasure M (\<Union>i. A i)"
51000
c9adb50f74ad use order topology for extended reals
hoelzl
parents: 50419
diff changeset
   519
  using LIMSEQ_SUP[OF incseq_emeasure, OF A] Lim_emeasure_incseq[OF A]
49773
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   520
  by (simp add: LIMSEQ_unique)
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   521
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   522
lemma decseq_emeasure:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   523
  assumes "range B \<subseteq> sets M" "decseq B"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   524
  shows "decseq (\<lambda>i. emeasure M (B i))"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   525
  using assms by (auto simp: decseq_def intro!: emeasure_mono)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   526
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   527
lemma INF_emeasure_decseq:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   528
  assumes A: "range A \<subseteq> sets M" and "decseq A"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   529
  and finite: "\<And>i. emeasure M (A i) \<noteq> \<infinity>"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   530
  shows "(INF n. emeasure M (A n)) = emeasure M (\<Inter>i. A i)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   531
proof -
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   532
  have le_MI: "emeasure M (\<Inter>i. A i) \<le> emeasure M (A 0)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   533
    using A by (auto intro!: emeasure_mono)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   534
  hence *: "emeasure M (\<Inter>i. A i) \<noteq> \<infinity>" using finite[of 0] by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   535
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   536
  have A0: "0 \<le> emeasure M (A 0)" using A by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   537
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   538
  have "emeasure M (A 0) - (INF n. emeasure M (A n)) = emeasure M (A 0) + (SUP n. - emeasure M (A n))"
56212
3253aaf73a01 consolidated theorem names containing INFI and SUPR: have INF and SUP instead uniformly
haftmann
parents: 56193
diff changeset
   539
    by (simp add: ereal_SUP_uminus minus_ereal_def)
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   540
  also have "\<dots> = (SUP n. emeasure M (A 0) - emeasure M (A n))"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   541
    unfolding minus_ereal_def using A0 assms
56212
3253aaf73a01 consolidated theorem names containing INFI and SUPR: have INF and SUP instead uniformly
haftmann
parents: 56193
diff changeset
   542
    by (subst SUP_ereal_add) (auto simp add: decseq_emeasure)
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   543
  also have "\<dots> = (SUP n. emeasure M (A 0 - A n))"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   544
    using A finite `decseq A`[unfolded decseq_def] by (subst emeasure_Diff) auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   545
  also have "\<dots> = emeasure M (\<Union>i. A 0 - A i)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   546
  proof (rule SUP_emeasure_incseq)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   547
    show "range (\<lambda>n. A 0 - A n) \<subseteq> sets M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   548
      using A by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   549
    show "incseq (\<lambda>n. A 0 - A n)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   550
      using `decseq A` by (auto simp add: incseq_def decseq_def)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   551
  qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   552
  also have "\<dots> = emeasure M (A 0) - emeasure M (\<Inter>i. A i)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   553
    using A finite * by (simp, subst emeasure_Diff) auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   554
  finally show ?thesis
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   555
    unfolding ereal_minus_eq_minus_iff using finite A0 by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   556
qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   557
61359
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61166
diff changeset
   558
lemma emeasure_INT_decseq_subset:
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61166
diff changeset
   559
  fixes F :: "nat \<Rightarrow> 'a set"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61166
diff changeset
   560
  assumes I: "I \<noteq> {}" and F: "\<And>i j. i \<in> I \<Longrightarrow> j \<in> I \<Longrightarrow> i \<le> j \<Longrightarrow> F j \<subseteq> F i"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61166
diff changeset
   561
  assumes F_sets[measurable]: "\<And>i. i \<in> I \<Longrightarrow> F i \<in> sets M"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61166
diff changeset
   562
    and fin: "\<And>i. i \<in> I \<Longrightarrow> emeasure M (F i) \<noteq> \<infinity>"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61166
diff changeset
   563
  shows "emeasure M (\<Inter>i\<in>I. F i) = (INF i:I. emeasure M (F i))"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61166
diff changeset
   564
proof cases
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61166
diff changeset
   565
  assume "finite I"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61166
diff changeset
   566
  have "(\<Inter>i\<in>I. F i) = F (Max I)"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61166
diff changeset
   567
    using I \<open>finite I\<close> by (intro antisym INF_lower INF_greatest F) auto
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61166
diff changeset
   568
  moreover have "(INF i:I. emeasure M (F i)) = emeasure M (F (Max I))"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61166
diff changeset
   569
    using I \<open>finite I\<close> by (intro antisym INF_lower INF_greatest F emeasure_mono) auto
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61166
diff changeset
   570
  ultimately show ?thesis
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61166
diff changeset
   571
    by simp
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61166
diff changeset
   572
next
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61166
diff changeset
   573
  assume "infinite I"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61166
diff changeset
   574
  def L \<equiv> "\<lambda>n. LEAST i. i \<in> I \<and> i \<ge> n"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61166
diff changeset
   575
  have L: "L n \<in> I \<and> n \<le> L n" for n
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61166
diff changeset
   576
    unfolding L_def
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61166
diff changeset
   577
  proof (rule LeastI_ex)
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61166
diff changeset
   578
    show "\<exists>x. x \<in> I \<and> n \<le> x"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61166
diff changeset
   579
      using \<open>infinite I\<close> finite_subset[of I "{..< n}"]
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61166
diff changeset
   580
      by (rule_tac ccontr) (auto simp: not_le)
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61166
diff changeset
   581
  qed
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61166
diff changeset
   582
  have L_eq[simp]: "i \<in> I \<Longrightarrow> L i = i" for i
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61166
diff changeset
   583
    unfolding L_def by (intro Least_equality) auto
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61166
diff changeset
   584
  have L_mono: "i \<le> j \<Longrightarrow> L i \<le> L j" for i j
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61166
diff changeset
   585
    using L[of j] unfolding L_def by (intro Least_le) (auto simp: L_def)
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61166
diff changeset
   586
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61166
diff changeset
   587
  have "emeasure M (\<Inter>i. F (L i)) = (INF i. emeasure M (F (L i)))"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61166
diff changeset
   588
  proof (intro INF_emeasure_decseq[symmetric])
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61166
diff changeset
   589
    show "decseq (\<lambda>i. F (L i))"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61166
diff changeset
   590
      using L by (intro antimonoI F L_mono) auto
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61166
diff changeset
   591
  qed (insert L fin, auto)
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61166
diff changeset
   592
  also have "\<dots> = (INF i:I. emeasure M (F i))"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61166
diff changeset
   593
  proof (intro antisym INF_greatest)
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61166
diff changeset
   594
    show "i \<in> I \<Longrightarrow> (INF i. emeasure M (F (L i))) \<le> emeasure M (F i)" for i
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61166
diff changeset
   595
      by (intro INF_lower2[of i]) auto
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61166
diff changeset
   596
  qed (insert L, auto intro: INF_lower)
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61166
diff changeset
   597
  also have "(\<Inter>i. F (L i)) = (\<Inter>i\<in>I. F i)"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61166
diff changeset
   598
  proof (intro antisym INF_greatest)
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61166
diff changeset
   599
    show "i \<in> I \<Longrightarrow> (\<Inter>i. F (L i)) \<subseteq> F i" for i
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61166
diff changeset
   600
      by (intro INF_lower2[of i]) auto
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61166
diff changeset
   601
  qed (insert L, auto)
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61166
diff changeset
   602
  finally show ?thesis .
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61166
diff changeset
   603
qed
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61166
diff changeset
   604
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   605
lemma Lim_emeasure_decseq:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   606
  assumes A: "range A \<subseteq> sets M" "decseq A" and fin: "\<And>i. emeasure M (A i) \<noteq> \<infinity>"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   607
  shows "(\<lambda>i. emeasure M (A i)) ----> emeasure M (\<Inter>i. A i)"
51351
dd1dd470690b generalized lemmas in Extended_Real_Limits
hoelzl
parents: 51000
diff changeset
   608
  using LIMSEQ_INF[OF decseq_emeasure, OF A]
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   609
  using INF_emeasure_decseq[OF A fin] by simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   610
60636
ee18efe9b246 add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
hoelzl
parents: 60585
diff changeset
   611
lemma emeasure_lfp'[consumes 1, case_names cont measurable]:
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
   612
  assumes "P M"
60172
423273355b55 rename continuous and down_continuous in Order_Continuity to sup_/inf_continuous; relate them with topological continuity
hoelzl
parents: 60142
diff changeset
   613
  assumes cont: "sup_continuous F"
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
   614
  assumes *: "\<And>M A. P M \<Longrightarrow> (\<And>N. P N \<Longrightarrow> Measurable.pred N A) \<Longrightarrow> Measurable.pred M (F A)"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
   615
  shows "emeasure M {x\<in>space M. lfp F x} = (SUP i. emeasure M {x\<in>space M. (F ^^ i) (\<lambda>x. False) x})"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
   616
proof -
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
   617
  have "emeasure M {x\<in>space M. lfp F x} = emeasure M (\<Union>i. {x\<in>space M. (F ^^ i) (\<lambda>x. False) x})"
60172
423273355b55 rename continuous and down_continuous in Order_Continuity to sup_/inf_continuous; relate them with topological continuity
hoelzl
parents: 60142
diff changeset
   618
    using sup_continuous_lfp[OF cont] by (auto simp add: bot_fun_def intro!: arg_cong2[where f=emeasure])
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
   619
  moreover { fix i from `P M` have "{x\<in>space M. (F ^^ i) (\<lambda>x. False) x} \<in> sets M"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
   620
    by (induct i arbitrary: M) (auto simp add: pred_def[symmetric] intro: *) }
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
   621
  moreover have "incseq (\<lambda>i. {x\<in>space M. (F ^^ i) (\<lambda>x. False) x})"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
   622
  proof (rule incseq_SucI)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
   623
    fix i
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
   624
    have "(F ^^ i) (\<lambda>x. False) \<le> (F ^^ (Suc i)) (\<lambda>x. False)"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
   625
    proof (induct i)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
   626
      case 0 show ?case by (simp add: le_fun_def)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
   627
    next
60172
423273355b55 rename continuous and down_continuous in Order_Continuity to sup_/inf_continuous; relate them with topological continuity
hoelzl
parents: 60142
diff changeset
   628
      case Suc thus ?case using monoD[OF sup_continuous_mono[OF cont] Suc] by auto
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
   629
    qed
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
   630
    then show "{x \<in> space M. (F ^^ i) (\<lambda>x. False) x} \<subseteq> {x \<in> space M. (F ^^ Suc i) (\<lambda>x. False) x}"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
   631
      by auto
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
   632
  qed
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
   633
  ultimately show ?thesis
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
   634
    by (subst SUP_emeasure_incseq) auto
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
   635
qed
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
   636
60636
ee18efe9b246 add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
hoelzl
parents: 60585
diff changeset
   637
lemma emeasure_lfp:
ee18efe9b246 add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
hoelzl
parents: 60585
diff changeset
   638
  assumes [simp]: "\<And>s. sets (M s) = sets N"
ee18efe9b246 add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
hoelzl
parents: 60585
diff changeset
   639
  assumes cont: "sup_continuous F" "sup_continuous f"
ee18efe9b246 add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
hoelzl
parents: 60585
diff changeset
   640
  assumes nonneg: "\<And>P s. 0 \<le> f P s"
ee18efe9b246 add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
hoelzl
parents: 60585
diff changeset
   641
  assumes meas: "\<And>P. Measurable.pred N P \<Longrightarrow> Measurable.pred N (F P)"
60714
ff8aa76d6d1c stronger induction assumption in lfp_transfer and emeasure_lfp
hoelzl
parents: 60636
diff changeset
   642
  assumes iter: "\<And>P s. Measurable.pred N P \<Longrightarrow> P \<le> lfp F \<Longrightarrow> emeasure (M s) {x\<in>space N. F P x} = f (\<lambda>s. emeasure (M s) {x\<in>space N. P x}) s"
60636
ee18efe9b246 add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
hoelzl
parents: 60585
diff changeset
   643
  shows "emeasure (M s) {x\<in>space N. lfp F x} = lfp f s"
ee18efe9b246 add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
hoelzl
parents: 60585
diff changeset
   644
proof (subst lfp_transfer_bounded[where \<alpha>="\<lambda>F s. emeasure (M s) {x\<in>space N. F x}" and g=f and f=F and P="Measurable.pred N", symmetric])
ee18efe9b246 add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
hoelzl
parents: 60585
diff changeset
   645
  fix C assume "incseq C" "\<And>i. Measurable.pred N (C i)"
ee18efe9b246 add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
hoelzl
parents: 60585
diff changeset
   646
  then show "(\<lambda>s. emeasure (M s) {x \<in> space N. (SUP i. C i) x}) = (SUP i. (\<lambda>s. emeasure (M s) {x \<in> space N. C i x}))"
ee18efe9b246 add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
hoelzl
parents: 60585
diff changeset
   647
    unfolding SUP_apply[abs_def]
ee18efe9b246 add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
hoelzl
parents: 60585
diff changeset
   648
    by (subst SUP_emeasure_incseq) (auto simp: mono_def fun_eq_iff intro!: arg_cong2[where f=emeasure])
ee18efe9b246 add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
hoelzl
parents: 60585
diff changeset
   649
qed (auto simp add: iter nonneg le_fun_def SUP_apply[abs_def] intro!: meas cont)
ee18efe9b246 add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
hoelzl
parents: 60585
diff changeset
   650
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   651
lemma emeasure_subadditive:
50002
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
   652
  assumes [measurable]: "A \<in> sets M" "B \<in> sets M"
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   653
  shows "emeasure M (A \<union> B) \<le> emeasure M A + emeasure M B"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   654
proof -
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   655
  from plus_emeasure[of A M "B - A"]
50002
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
   656
  have "emeasure M (A \<union> B) = emeasure M A + emeasure M (B - A)" by simp
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   657
  also have "\<dots> \<le> emeasure M A + emeasure M B"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   658
    using assms by (auto intro!: add_left_mono emeasure_mono)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   659
  finally show ?thesis .
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   660
qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   661
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   662
lemma emeasure_subadditive_finite:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   663
  assumes "finite I" "A ` I \<subseteq> sets M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   664
  shows "emeasure M (\<Union>i\<in>I. A i) \<le> (\<Sum>i\<in>I. emeasure M (A i))"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   665
using assms proof induct
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   666
  case (insert i I)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   667
  then have "emeasure M (\<Union>i\<in>insert i I. A i) = emeasure M (A i \<union> (\<Union>i\<in>I. A i))"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   668
    by simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   669
  also have "\<dots> \<le> emeasure M (A i) + emeasure M (\<Union>i\<in>I. A i)"
50002
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
   670
    using insert by (intro emeasure_subadditive) auto
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   671
  also have "\<dots> \<le> emeasure M (A i) + (\<Sum>i\<in>I. emeasure M (A i))"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   672
    using insert by (intro add_mono) auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   673
  also have "\<dots> = (\<Sum>i\<in>insert i I. emeasure M (A i))"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   674
    using insert by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   675
  finally show ?case .
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   676
qed simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   677
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   678
lemma emeasure_subadditive_countably:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   679
  assumes "range f \<subseteq> sets M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   680
  shows "emeasure M (\<Union>i. f i) \<le> (\<Sum>i. emeasure M (f i))"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   681
proof -
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   682
  have "emeasure M (\<Union>i. f i) = emeasure M (\<Union>i. disjointed f i)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   683
    unfolding UN_disjointed_eq ..
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   684
  also have "\<dots> = (\<Sum>i. emeasure M (disjointed f i))"
50244
de72bbe42190 qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents: 50104
diff changeset
   685
    using sets.range_disjointed_sets[OF assms] suminf_emeasure[of "disjointed f"]
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   686
    by (simp add:  disjoint_family_disjointed comp_def)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   687
  also have "\<dots> \<le> (\<Sum>i. emeasure M (f i))"
50244
de72bbe42190 qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents: 50104
diff changeset
   688
    using sets.range_disjointed_sets[OF assms] assms
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   689
    by (auto intro!: suminf_le_pos emeasure_mono disjointed_subset)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   690
  finally show ?thesis .
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   691
qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   692
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   693
lemma emeasure_insert:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   694
  assumes sets: "{x} \<in> sets M" "A \<in> sets M" and "x \<notin> A"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   695
  shows "emeasure M (insert x A) = emeasure M {x} + emeasure M A"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   696
proof -
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   697
  have "{x} \<inter> A = {}" using `x \<notin> A` by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   698
  from plus_emeasure[OF sets this] show ?thesis by simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   699
qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   700
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
   701
lemma emeasure_insert_ne:
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
   702
  "A \<noteq> {} \<Longrightarrow> {x} \<in> sets M \<Longrightarrow> A \<in> sets M \<Longrightarrow> x \<notin> A \<Longrightarrow> emeasure M (insert x A) = emeasure M {x} + emeasure M A"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61359
diff changeset
   703
  by (rule emeasure_insert)
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
   704
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   705
lemma emeasure_eq_setsum_singleton:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   706
  assumes "finite S" "\<And>x. x \<in> S \<Longrightarrow> {x} \<in> sets M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   707
  shows "emeasure M S = (\<Sum>x\<in>S. emeasure M {x})"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   708
  using setsum_emeasure[of "\<lambda>x. {x}" S M] assms
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   709
  by (auto simp: disjoint_family_on_def subset_eq)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   710
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   711
lemma setsum_emeasure_cover:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   712
  assumes "finite S" and "A \<in> sets M" and br_in_M: "B ` S \<subseteq> sets M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   713
  assumes A: "A \<subseteq> (\<Union>i\<in>S. B i)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   714
  assumes disj: "disjoint_family_on B S"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   715
  shows "emeasure M A = (\<Sum>i\<in>S. emeasure M (A \<inter> (B i)))"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   716
proof -
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   717
  have "(\<Sum>i\<in>S. emeasure M (A \<inter> (B i))) = emeasure M (\<Union>i\<in>S. A \<inter> (B i))"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   718
  proof (rule setsum_emeasure)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   719
    show "disjoint_family_on (\<lambda>i. A \<inter> B i) S"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   720
      using `disjoint_family_on B S`
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   721
      unfolding disjoint_family_on_def by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   722
  qed (insert assms, auto)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   723
  also have "(\<Union>i\<in>S. A \<inter> (B i)) = A"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   724
    using A by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   725
  finally show ?thesis by simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   726
qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   727
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   728
lemma emeasure_eq_0:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   729
  "N \<in> sets M \<Longrightarrow> emeasure M N = 0 \<Longrightarrow> K \<subseteq> N \<Longrightarrow> emeasure M K = 0"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   730
  by (metis emeasure_mono emeasure_nonneg order_eq_iff)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   731
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   732
lemma emeasure_UN_eq_0:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   733
  assumes "\<And>i::nat. emeasure M (N i) = 0" and "range N \<subseteq> sets M"
60585
48fdff264eb2 tuned whitespace;
wenzelm
parents: 60580
diff changeset
   734
  shows "emeasure M (\<Union>i. N i) = 0"
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   735
proof -
60585
48fdff264eb2 tuned whitespace;
wenzelm
parents: 60580
diff changeset
   736
  have "0 \<le> emeasure M (\<Union>i. N i)" using assms by auto
48fdff264eb2 tuned whitespace;
wenzelm
parents: 60580
diff changeset
   737
  moreover have "emeasure M (\<Union>i. N i) \<le> 0"
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   738
    using emeasure_subadditive_countably[OF assms(2)] assms(1) by simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   739
  ultimately show ?thesis by simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   740
qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   741
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   742
lemma measure_eqI_finite:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   743
  assumes [simp]: "sets M = Pow A" "sets N = Pow A" and "finite A"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   744
  assumes eq: "\<And>a. a \<in> A \<Longrightarrow> emeasure M {a} = emeasure N {a}"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   745
  shows "M = N"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   746
proof (rule measure_eqI)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   747
  fix X assume "X \<in> sets M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   748
  then have X: "X \<subseteq> A" by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   749
  then have "emeasure M X = (\<Sum>a\<in>X. emeasure M {a})"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   750
    using `finite A` by (subst emeasure_eq_setsum_singleton) (auto dest: finite_subset)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   751
  also have "\<dots> = (\<Sum>a\<in>X. emeasure N {a})"
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57276
diff changeset
   752
    using X eq by (auto intro!: setsum.cong)
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   753
  also have "\<dots> = emeasure N X"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   754
    using X `finite A` by (subst emeasure_eq_setsum_singleton) (auto dest: finite_subset)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   755
  finally show "emeasure M X = emeasure N X" .
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   756
qed simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   757
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   758
lemma measure_eqI_generator_eq:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   759
  fixes M N :: "'a measure" and E :: "'a set set" and A :: "nat \<Rightarrow> 'a set"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   760
  assumes "Int_stable E" "E \<subseteq> Pow \<Omega>"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   761
  and eq: "\<And>X. X \<in> E \<Longrightarrow> emeasure M X = emeasure N X"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   762
  and M: "sets M = sigma_sets \<Omega> E"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   763
  and N: "sets N = sigma_sets \<Omega> E"
49784
5e5b2da42a69 remove incseq assumption from measure_eqI_generator_eq
hoelzl
parents: 49773
diff changeset
   764
  and A: "range A \<subseteq> E" "(\<Union>i. A i) = \<Omega>" "\<And>i. emeasure M (A i) \<noteq> \<infinity>"
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   765
  shows "M = N"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   766
proof -
49773
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   767
  let ?\<mu>  = "emeasure M" and ?\<nu> = "emeasure N"
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   768
  interpret S: sigma_algebra \<Omega> "sigma_sets \<Omega> E" by (rule sigma_algebra_sigma_sets) fact
49789
e0a4cb91a8a9 add induction rule for intersection-stable sigma-sets
hoelzl
parents: 49784
diff changeset
   769
  have "space M = \<Omega>"
50244
de72bbe42190 qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents: 50104
diff changeset
   770
    using sets.top[of M] sets.space_closed[of M] S.top S.space_closed `sets M = sigma_sets \<Omega> E`
de72bbe42190 qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents: 50104
diff changeset
   771
    by blast
49789
e0a4cb91a8a9 add induction rule for intersection-stable sigma-sets
hoelzl
parents: 49784
diff changeset
   772
e0a4cb91a8a9 add induction rule for intersection-stable sigma-sets
hoelzl
parents: 49784
diff changeset
   773
  { fix F D assume "F \<in> E" and "?\<mu> F \<noteq> \<infinity>"
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   774
    then have [intro]: "F \<in> sigma_sets \<Omega> E" by auto
49773
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   775
    have "?\<nu> F \<noteq> \<infinity>" using `?\<mu> F \<noteq> \<infinity>` `F \<in> E` eq by simp
49789
e0a4cb91a8a9 add induction rule for intersection-stable sigma-sets
hoelzl
parents: 49784
diff changeset
   776
    assume "D \<in> sets M"
e0a4cb91a8a9 add induction rule for intersection-stable sigma-sets
hoelzl
parents: 49784
diff changeset
   777
    with `Int_stable E` `E \<subseteq> Pow \<Omega>` have "emeasure M (F \<inter> D) = emeasure N (F \<inter> D)"
e0a4cb91a8a9 add induction rule for intersection-stable sigma-sets
hoelzl
parents: 49784
diff changeset
   778
      unfolding M
e0a4cb91a8a9 add induction rule for intersection-stable sigma-sets
hoelzl
parents: 49784
diff changeset
   779
    proof (induct rule: sigma_sets_induct_disjoint)
e0a4cb91a8a9 add induction rule for intersection-stable sigma-sets
hoelzl
parents: 49784
diff changeset
   780
      case (basic A)
e0a4cb91a8a9 add induction rule for intersection-stable sigma-sets
hoelzl
parents: 49784
diff changeset
   781
      then have "F \<inter> A \<in> E" using `Int_stable E` `F \<in> E` by (auto simp: Int_stable_def)
e0a4cb91a8a9 add induction rule for intersection-stable sigma-sets
hoelzl
parents: 49784
diff changeset
   782
      then show ?case using eq by auto
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   783
    next
49789
e0a4cb91a8a9 add induction rule for intersection-stable sigma-sets
hoelzl
parents: 49784
diff changeset
   784
      case empty then show ?case by simp
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   785
    next
49789
e0a4cb91a8a9 add induction rule for intersection-stable sigma-sets
hoelzl
parents: 49784
diff changeset
   786
      case (compl A)
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   787
      then have **: "F \<inter> (\<Omega> - A) = F - (F \<inter> A)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   788
        and [intro]: "F \<inter> A \<in> sigma_sets \<Omega> E"
49789
e0a4cb91a8a9 add induction rule for intersection-stable sigma-sets
hoelzl
parents: 49784
diff changeset
   789
        using `F \<in> E` S.sets_into_space by (auto simp: M)
49773
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   790
      have "?\<nu> (F \<inter> A) \<le> ?\<nu> F" by (auto intro!: emeasure_mono simp: M N)
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   791
      then have "?\<nu> (F \<inter> A) \<noteq> \<infinity>" using `?\<nu> F \<noteq> \<infinity>` by auto
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   792
      have "?\<mu> (F \<inter> A) \<le> ?\<mu> F" by (auto intro!: emeasure_mono simp: M N)
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   793
      then have "?\<mu> (F \<inter> A) \<noteq> \<infinity>" using `?\<mu> F \<noteq> \<infinity>` by auto
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   794
      then have "?\<mu> (F \<inter> (\<Omega> - A)) = ?\<mu> F - ?\<mu> (F \<inter> A)" unfolding **
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   795
        using `F \<inter> A \<in> sigma_sets \<Omega> E` by (auto intro!: emeasure_Diff simp: M N)
49789
e0a4cb91a8a9 add induction rule for intersection-stable sigma-sets
hoelzl
parents: 49784
diff changeset
   796
      also have "\<dots> = ?\<nu> F - ?\<nu> (F \<inter> A)" using eq `F \<in> E` compl by simp
49773
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   797
      also have "\<dots> = ?\<nu> (F \<inter> (\<Omega> - A))" unfolding **
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   798
        using `F \<inter> A \<in> sigma_sets \<Omega> E` `?\<nu> (F \<inter> A) \<noteq> \<infinity>`
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   799
        by (auto intro!: emeasure_Diff[symmetric] simp: M N)
49789
e0a4cb91a8a9 add induction rule for intersection-stable sigma-sets
hoelzl
parents: 49784
diff changeset
   800
      finally show ?case
e0a4cb91a8a9 add induction rule for intersection-stable sigma-sets
hoelzl
parents: 49784
diff changeset
   801
        using `space M = \<Omega>` by auto
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   802
    next
49789
e0a4cb91a8a9 add induction rule for intersection-stable sigma-sets
hoelzl
parents: 49784
diff changeset
   803
      case (union A)
49773
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   804
      then have "?\<mu> (\<Union>x. F \<inter> A x) = ?\<nu> (\<Union>x. F \<inter> A x)"
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   805
        by (subst (1 2) suminf_emeasure[symmetric]) (auto simp: disjoint_family_on_def subset_eq M N)
49789
e0a4cb91a8a9 add induction rule for intersection-stable sigma-sets
hoelzl
parents: 49784
diff changeset
   806
      with A show ?case
49773
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   807
        by auto
49789
e0a4cb91a8a9 add induction rule for intersection-stable sigma-sets
hoelzl
parents: 49784
diff changeset
   808
    qed }
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   809
  note * = this
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   810
  show "M = N"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   811
  proof (rule measure_eqI)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   812
    show "sets M = sets N"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   813
      using M N by simp
49784
5e5b2da42a69 remove incseq assumption from measure_eqI_generator_eq
hoelzl
parents: 49773
diff changeset
   814
    have [simp, intro]: "\<And>i. A i \<in> sets M"
5e5b2da42a69 remove incseq assumption from measure_eqI_generator_eq
hoelzl
parents: 49773
diff changeset
   815
      using A(1) by (auto simp: subset_eq M)
49773
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
   816
    fix F assume "F \<in> sets M"
49784
5e5b2da42a69 remove incseq assumption from measure_eqI_generator_eq
hoelzl
parents: 49773
diff changeset
   817
    let ?D = "disjointed (\<lambda>i. F \<inter> A i)"
49789
e0a4cb91a8a9 add induction rule for intersection-stable sigma-sets
hoelzl
parents: 49784
diff changeset
   818
    from `space M = \<Omega>` have F_eq: "F = (\<Union>i. ?D i)"
50244
de72bbe42190 qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents: 50104
diff changeset
   819
      using `F \<in> sets M`[THEN sets.sets_into_space] A(2)[symmetric] by (auto simp: UN_disjointed_eq)
49784
5e5b2da42a69 remove incseq assumption from measure_eqI_generator_eq
hoelzl
parents: 49773
diff changeset
   820
    have [simp, intro]: "\<And>i. ?D i \<in> sets M"
50244
de72bbe42190 qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents: 50104
diff changeset
   821
      using sets.range_disjointed_sets[of "\<lambda>i. F \<inter> A i" M] `F \<in> sets M`
49784
5e5b2da42a69 remove incseq assumption from measure_eqI_generator_eq
hoelzl
parents: 49773
diff changeset
   822
      by (auto simp: subset_eq)
5e5b2da42a69 remove incseq assumption from measure_eqI_generator_eq
hoelzl
parents: 49773
diff changeset
   823
    have "disjoint_family ?D"
5e5b2da42a69 remove incseq assumption from measure_eqI_generator_eq
hoelzl
parents: 49773
diff changeset
   824
      by (auto simp: disjoint_family_disjointed)
50002
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
   825
    moreover
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
   826
    have "(\<Sum>i. emeasure M (?D i)) = (\<Sum>i. emeasure N (?D i))"
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
   827
    proof (intro arg_cong[where f=suminf] ext)
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
   828
      fix i
49784
5e5b2da42a69 remove incseq assumption from measure_eqI_generator_eq
hoelzl
parents: 49773
diff changeset
   829
      have "A i \<inter> ?D i = ?D i"
5e5b2da42a69 remove incseq assumption from measure_eqI_generator_eq
hoelzl
parents: 49773
diff changeset
   830
        by (auto simp: disjointed_def)
50002
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
   831
      then show "emeasure M (?D i) = emeasure N (?D i)"
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
   832
        using *[of "A i" "?D i", OF _ A(3)] A(1) by auto
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
   833
    qed
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
   834
    ultimately show "emeasure M F = emeasure N F"
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
   835
      by (simp add: image_subset_iff `sets M = sets N`[symmetric] F_eq[symmetric] suminf_emeasure)
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   836
  qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   837
qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   838
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   839
lemma measure_of_of_measure: "measure_of (space M) (sets M) (emeasure M) = M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   840
proof (intro measure_eqI emeasure_measure_of_sigma)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   841
  show "sigma_algebra (space M) (sets M)" ..
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   842
  show "positive (sets M) (emeasure M)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   843
    by (simp add: positive_def emeasure_nonneg)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   844
  show "countably_additive (sets M) (emeasure M)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   845
    by (simp add: emeasure_countably_additive)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   846
qed simp_all
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   847
56994
8d5e5ec1cac3 fixed document generation for HOL-Probability
hoelzl
parents: 56212
diff changeset
   848
subsection {* @{text \<mu>}-null sets *}
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   849
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   850
definition null_sets :: "'a measure \<Rightarrow> 'a set set" where
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   851
  "null_sets M = {N\<in>sets M. emeasure M N = 0}"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   852
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   853
lemma null_setsD1[dest]: "A \<in> null_sets M \<Longrightarrow> emeasure M A = 0"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   854
  by (simp add: null_sets_def)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   855
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   856
lemma null_setsD2[dest]: "A \<in> null_sets M \<Longrightarrow> A \<in> sets M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   857
  unfolding null_sets_def by simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   858
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   859
lemma null_setsI[intro]: "emeasure M A = 0 \<Longrightarrow> A \<in> sets M \<Longrightarrow> A \<in> null_sets M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   860
  unfolding null_sets_def by simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   861
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   862
interpretation null_sets: ring_of_sets "space M" "null_sets M" for M
47762
d31085f07f60 add Caratheodories theorem for semi-rings of sets
hoelzl
parents: 47761
diff changeset
   863
proof (rule ring_of_setsI)
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   864
  show "null_sets M \<subseteq> Pow (space M)"
50244
de72bbe42190 qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents: 50104
diff changeset
   865
    using sets.sets_into_space by auto
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   866
  show "{} \<in> null_sets M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   867
    by auto
53374
a14d2a854c02 tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents: 51351
diff changeset
   868
  fix A B assume null_sets: "A \<in> null_sets M" "B \<in> null_sets M"
a14d2a854c02 tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents: 51351
diff changeset
   869
  then have sets: "A \<in> sets M" "B \<in> sets M"
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   870
    by auto
53374
a14d2a854c02 tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents: 51351
diff changeset
   871
  then have *: "emeasure M (A \<union> B) \<le> emeasure M A + emeasure M B"
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   872
    "emeasure M (A - B) \<le> emeasure M A"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   873
    by (auto intro!: emeasure_subadditive emeasure_mono)
53374
a14d2a854c02 tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents: 51351
diff changeset
   874
  then have "emeasure M B = 0" "emeasure M A = 0"
a14d2a854c02 tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents: 51351
diff changeset
   875
    using null_sets by auto
a14d2a854c02 tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents: 51351
diff changeset
   876
  with sets * show "A - B \<in> null_sets M" "A \<union> B \<in> null_sets M"
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   877
    by (auto intro!: antisym)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   878
qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   879
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61359
diff changeset
   880
lemma UN_from_nat_into:
57275
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
   881
  assumes I: "countable I" "I \<noteq> {}"
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
   882
  shows "(\<Union>i\<in>I. N i) = (\<Union>i. N (from_nat_into I i))"
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   883
proof -
57275
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
   884
  have "(\<Union>i\<in>I. N i) = \<Union>(N ` range (from_nat_into I))"
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
   885
    using I by simp
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
   886
  also have "\<dots> = (\<Union>i. (N \<circ> from_nat_into I) i)"
56154
f0a927235162 more complete set of lemmas wrt. image and composition
haftmann
parents: 54417
diff changeset
   887
    by (simp only: SUP_def image_comp)
57275
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
   888
  finally show ?thesis by simp
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
   889
qed
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
   890
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
   891
lemma null_sets_UN':
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
   892
  assumes "countable I"
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
   893
  assumes "\<And>i. i \<in> I \<Longrightarrow> N i \<in> null_sets M"
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
   894
  shows "(\<Union>i\<in>I. N i) \<in> null_sets M"
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
   895
proof cases
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
   896
  assume "I = {}" then show ?thesis by simp
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
   897
next
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
   898
  assume "I \<noteq> {}"
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
   899
  show ?thesis
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
   900
  proof (intro conjI CollectI null_setsI)
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
   901
    show "(\<Union>i\<in>I. N i) \<in> sets M"
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
   902
      using assms by (intro sets.countable_UN') auto
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
   903
    have "emeasure M (\<Union>i\<in>I. N i) \<le> (\<Sum>n. emeasure M (N (from_nat_into I n)))"
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
   904
      unfolding UN_from_nat_into[OF `countable I` `I \<noteq> {}`]
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
   905
      using assms `I \<noteq> {}` by (intro emeasure_subadditive_countably) (auto intro: from_nat_into)
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
   906
    also have "(\<lambda>n. emeasure M (N (from_nat_into I n))) = (\<lambda>_. 0)"
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
   907
      using assms `I \<noteq> {}` by (auto intro: from_nat_into)
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
   908
    finally show "emeasure M (\<Union>i\<in>I. N i) = 0"
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
   909
      by (intro antisym emeasure_nonneg) simp
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
   910
  qed
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   911
qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   912
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   913
lemma null_sets_UN[intro]:
57275
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
   914
  "(\<And>i::'i::countable. N i \<in> null_sets M) \<Longrightarrow> (\<Union>i. N i) \<in> null_sets M"
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
   915
  by (rule null_sets_UN') auto
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   916
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   917
lemma null_set_Int1:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   918
  assumes "B \<in> null_sets M" "A \<in> sets M" shows "A \<inter> B \<in> null_sets M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   919
proof (intro CollectI conjI null_setsI)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   920
  show "emeasure M (A \<inter> B) = 0" using assms
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   921
    by (intro emeasure_eq_0[of B _ "A \<inter> B"]) auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   922
qed (insert assms, auto)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   923
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   924
lemma null_set_Int2:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   925
  assumes "B \<in> null_sets M" "A \<in> sets M" shows "B \<inter> A \<in> null_sets M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   926
  using assms by (subst Int_commute) (rule null_set_Int1)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   927
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   928
lemma emeasure_Diff_null_set:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   929
  assumes "B \<in> null_sets M" "A \<in> sets M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   930
  shows "emeasure M (A - B) = emeasure M A"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   931
proof -
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   932
  have *: "A - B = (A - (A \<inter> B))" by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   933
  have "A \<inter> B \<in> null_sets M" using assms by (rule null_set_Int1)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   934
  then show ?thesis
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   935
    unfolding * using assms
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   936
    by (subst emeasure_Diff) auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   937
qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   938
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   939
lemma null_set_Diff:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   940
  assumes "B \<in> null_sets M" "A \<in> sets M" shows "B - A \<in> null_sets M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   941
proof (intro CollectI conjI null_setsI)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   942
  show "emeasure M (B - A) = 0" using assms by (intro emeasure_eq_0[of B _ "B - A"]) auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   943
qed (insert assms, auto)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   944
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   945
lemma emeasure_Un_null_set:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   946
  assumes "A \<in> sets M" "B \<in> null_sets M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   947
  shows "emeasure M (A \<union> B) = emeasure M A"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   948
proof -
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   949
  have *: "A \<union> B = A \<union> (B - A)" by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   950
  have "B - A \<in> null_sets M" using assms(2,1) by (rule null_set_Diff)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   951
  then show ?thesis
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   952
    unfolding * using assms
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   953
    by (subst plus_emeasure[symmetric]) auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   954
qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   955
56994
8d5e5ec1cac3 fixed document generation for HOL-Probability
hoelzl
parents: 56212
diff changeset
   956
subsection {* The almost everywhere filter (i.e.\ quantifier) *}
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   957
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   958
definition ae_filter :: "'a measure \<Rightarrow> 'a filter" where
57276
49c51eeaa623 filters are easier to define with INF on filters.
hoelzl
parents: 57275
diff changeset
   959
  "ae_filter M = (INF N:null_sets M. principal (space M - N))"
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   960
57276
49c51eeaa623 filters are easier to define with INF on filters.
hoelzl
parents: 57275
diff changeset
   961
abbreviation almost_everywhere :: "'a measure \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> bool" where
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   962
  "almost_everywhere M P \<equiv> eventually P (ae_filter M)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   963
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   964
syntax
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   965
  "_almost_everywhere" :: "pttrn \<Rightarrow> 'a \<Rightarrow> bool \<Rightarrow> bool" ("AE _ in _. _" [0,0,10] 10)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   966
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   967
translations
57276
49c51eeaa623 filters are easier to define with INF on filters.
hoelzl
parents: 57275
diff changeset
   968
  "AE x in M. P" == "CONST almost_everywhere M (\<lambda>x. P)"
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   969
57276
49c51eeaa623 filters are easier to define with INF on filters.
hoelzl
parents: 57275
diff changeset
   970
lemma eventually_ae_filter: "eventually P (ae_filter M) \<longleftrightarrow> (\<exists>N\<in>null_sets M. {x \<in> space M. \<not> P x} \<subseteq> N)"
49c51eeaa623 filters are easier to define with INF on filters.
hoelzl
parents: 57275
diff changeset
   971
  unfolding ae_filter_def by (subst eventually_INF_base) (auto simp: eventually_principal subset_eq)
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   972
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   973
lemma AE_I':
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   974
  "N \<in> null_sets M \<Longrightarrow> {x\<in>space M. \<not> P x} \<subseteq> N \<Longrightarrow> (AE x in M. P x)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   975
  unfolding eventually_ae_filter by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   976
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   977
lemma AE_iff_null:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   978
  assumes "{x\<in>space M. \<not> P x} \<in> sets M" (is "?P \<in> sets M")
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   979
  shows "(AE x in M. P x) \<longleftrightarrow> {x\<in>space M. \<not> P x} \<in> null_sets M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   980
proof
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   981
  assume "AE x in M. P x" then obtain N where N: "N \<in> sets M" "?P \<subseteq> N" "emeasure M N = 0"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   982
    unfolding eventually_ae_filter by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   983
  have "0 \<le> emeasure M ?P" by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   984
  moreover have "emeasure M ?P \<le> emeasure M N"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   985
    using assms N(1,2) by (auto intro: emeasure_mono)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   986
  ultimately have "emeasure M ?P = 0" unfolding `emeasure M N = 0` by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   987
  then show "?P \<in> null_sets M" using assms by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   988
next
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   989
  assume "?P \<in> null_sets M" with assms show "AE x in M. P x" by (auto intro: AE_I')
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   990
qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   991
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   992
lemma AE_iff_null_sets:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   993
  "N \<in> sets M \<Longrightarrow> N \<in> null_sets M \<longleftrightarrow> (AE x in M. x \<notin> N)"
50244
de72bbe42190 qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents: 50104
diff changeset
   994
  using Int_absorb1[OF sets.sets_into_space, of N M]
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   995
  by (subst AE_iff_null) (auto simp: Int_def[symmetric])
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
   996
47761
dfe747e72fa8 moved lemmas to appropriate places
hoelzl
parents: 47694
diff changeset
   997
lemma AE_not_in:
dfe747e72fa8 moved lemmas to appropriate places
hoelzl
parents: 47694
diff changeset
   998
  "N \<in> null_sets M \<Longrightarrow> AE x in M. x \<notin> N"
dfe747e72fa8 moved lemmas to appropriate places
hoelzl
parents: 47694
diff changeset
   999
  by (metis AE_iff_null_sets null_setsD2)
dfe747e72fa8 moved lemmas to appropriate places
hoelzl
parents: 47694
diff changeset
  1000
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1001
lemma AE_iff_measurable:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1002
  "N \<in> sets M \<Longrightarrow> {x\<in>space M. \<not> P x} = N \<Longrightarrow> (AE x in M. P x) \<longleftrightarrow> emeasure M N = 0"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1003
  using AE_iff_null[of _ P] by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1004
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1005
lemma AE_E[consumes 1]:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1006
  assumes "AE x in M. P x"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1007
  obtains N where "{x \<in> space M. \<not> P x} \<subseteq> N" "emeasure M N = 0" "N \<in> sets M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1008
  using assms unfolding eventually_ae_filter by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1009
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1010
lemma AE_E2:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1011
  assumes "AE x in M. P x" "{x\<in>space M. P x} \<in> sets M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1012
  shows "emeasure M {x\<in>space M. \<not> P x} = 0" (is "emeasure M ?P = 0")
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1013
proof -
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1014
  have "{x\<in>space M. \<not> P x} = space M - {x\<in>space M. P x}" by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1015
  with AE_iff_null[of M P] assms show ?thesis by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1016
qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1017
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1018
lemma AE_I:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1019
  assumes "{x \<in> space M. \<not> P x} \<subseteq> N" "emeasure M N = 0" "N \<in> sets M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1020
  shows "AE x in M. P x"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1021
  using assms unfolding eventually_ae_filter by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1022
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1023
lemma AE_mp[elim!]:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1024
  assumes AE_P: "AE x in M. P x" and AE_imp: "AE x in M. P x \<longrightarrow> Q x"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1025
  shows "AE x in M. Q x"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1026
proof -
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1027
  from AE_P obtain A where P: "{x\<in>space M. \<not> P x} \<subseteq> A"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1028
    and A: "A \<in> sets M" "emeasure M A = 0"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1029
    by (auto elim!: AE_E)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1030
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1031
  from AE_imp obtain B where imp: "{x\<in>space M. P x \<and> \<not> Q x} \<subseteq> B"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1032
    and B: "B \<in> sets M" "emeasure M B = 0"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1033
    by (auto elim!: AE_E)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1034
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1035
  show ?thesis
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1036
  proof (intro AE_I)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1037
    have "0 \<le> emeasure M (A \<union> B)" using A B by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1038
    moreover have "emeasure M (A \<union> B) \<le> 0"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1039
      using emeasure_subadditive[of A M B] A B by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1040
    ultimately show "A \<union> B \<in> sets M" "emeasure M (A \<union> B) = 0" using A B by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1041
    show "{x\<in>space M. \<not> Q x} \<subseteq> A \<union> B"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1042
      using P imp by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1043
  qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1044
qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1045
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1046
(* depricated replace by laws about eventually *)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1047
lemma
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1048
  shows AE_iffI: "AE x in M. P x \<Longrightarrow> AE x in M. P x \<longleftrightarrow> Q x \<Longrightarrow> AE x in M. Q x"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1049
    and AE_disjI1: "AE x in M. P x \<Longrightarrow> AE x in M. P x \<or> Q x"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1050
    and AE_disjI2: "AE x in M. Q x \<Longrightarrow> AE x in M. P x \<or> Q x"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1051
    and AE_conjI: "AE x in M. P x \<Longrightarrow> AE x in M. Q x \<Longrightarrow> AE x in M. P x \<and> Q x"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1052
    and AE_conj_iff[simp]: "(AE x in M. P x \<and> Q x) \<longleftrightarrow> (AE x in M. P x) \<and> (AE x in M. Q x)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1053
  by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1054
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1055
lemma AE_impI:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1056
  "(P \<Longrightarrow> AE x in M. Q x) \<Longrightarrow> AE x in M. P \<longrightarrow> Q x"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1057
  by (cases P) auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1058
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1059
lemma AE_measure:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1060
  assumes AE: "AE x in M. P x" and sets: "{x\<in>space M. P x} \<in> sets M" (is "?P \<in> sets M")
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1061
  shows "emeasure M {x\<in>space M. P x} = emeasure M (space M)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1062
proof -
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1063
  from AE_E[OF AE] guess N . note N = this
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1064
  with sets have "emeasure M (space M) \<le> emeasure M (?P \<union> N)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1065
    by (intro emeasure_mono) auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1066
  also have "\<dots> \<le> emeasure M ?P + emeasure M N"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1067
    using sets N by (intro emeasure_subadditive) auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1068
  also have "\<dots> = emeasure M ?P" using N by simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1069
  finally show "emeasure M ?P = emeasure M (space M)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1070
    using emeasure_space[of M "?P"] by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1071
qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1072
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1073
lemma AE_space: "AE x in M. x \<in> space M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1074
  by (rule AE_I[where N="{}"]) auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1075
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1076
lemma AE_I2[simp, intro]:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1077
  "(\<And>x. x \<in> space M \<Longrightarrow> P x) \<Longrightarrow> AE x in M. P x"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1078
  using AE_space by force
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1079
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1080
lemma AE_Ball_mp:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1081
  "\<forall>x\<in>space M. P x \<Longrightarrow> AE x in M. P x \<longrightarrow> Q x \<Longrightarrow> AE x in M. Q x"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1082
  by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1083
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1084
lemma AE_cong[cong]:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1085
  "(\<And>x. x \<in> space M \<Longrightarrow> P x \<longleftrightarrow> Q x) \<Longrightarrow> (AE x in M. P x) \<longleftrightarrow> (AE x in M. Q x)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1086
  by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1087
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1088
lemma AE_all_countable:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1089
  "(AE x in M. \<forall>i. P i x) \<longleftrightarrow> (\<forall>i::'i::countable. AE x in M. P i x)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1090
proof
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1091
  assume "\<forall>i. AE x in M. P i x"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1092
  from this[unfolded eventually_ae_filter Bex_def, THEN choice]
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1093
  obtain N where N: "\<And>i. N i \<in> null_sets M" "\<And>i. {x\<in>space M. \<not> P i x} \<subseteq> N i" by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1094
  have "{x\<in>space M. \<not> (\<forall>i. P i x)} \<subseteq> (\<Union>i. {x\<in>space M. \<not> P i x})" by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1095
  also have "\<dots> \<subseteq> (\<Union>i. N i)" using N by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1096
  finally have "{x\<in>space M. \<not> (\<forall>i. P i x)} \<subseteq> (\<Union>i. N i)" .
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1097
  moreover from N have "(\<Union>i. N i) \<in> null_sets M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1098
    by (intro null_sets_UN) auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1099
  ultimately show "AE x in M. \<forall>i. P i x"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1100
    unfolding eventually_ae_filter by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1101
qed auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1102
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61359
diff changeset
  1103
lemma AE_ball_countable:
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1104
  assumes [intro]: "countable X"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1105
  shows "(AE x in M. \<forall>y\<in>X. P x y) \<longleftrightarrow> (\<forall>y\<in>X. AE x in M. P x y)"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1106
proof
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1107
  assume "\<forall>y\<in>X. AE x in M. P x y"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1108
  from this[unfolded eventually_ae_filter Bex_def, THEN bchoice]
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1109
  obtain N where N: "\<And>y. y \<in> X \<Longrightarrow> N y \<in> null_sets M" "\<And>y. y \<in> X \<Longrightarrow> {x\<in>space M. \<not> P x y} \<subseteq> N y"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1110
    by auto
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1111
  have "{x\<in>space M. \<not> (\<forall>y\<in>X. P x y)} \<subseteq> (\<Union>y\<in>X. {x\<in>space M. \<not> P x y})"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1112
    by auto
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1113
  also have "\<dots> \<subseteq> (\<Union>y\<in>X. N y)"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1114
    using N by auto
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1115
  finally have "{x\<in>space M. \<not> (\<forall>y\<in>X. P x y)} \<subseteq> (\<Union>y\<in>X. N y)" .
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1116
  moreover from N have "(\<Union>y\<in>X. N y) \<in> null_sets M"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1117
    by (intro null_sets_UN') auto
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1118
  ultimately show "AE x in M. \<forall>y\<in>X. P x y"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1119
    unfolding eventually_ae_filter by auto
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1120
qed auto
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1121
57275
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1122
lemma AE_discrete_difference:
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1123
  assumes X: "countable X"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61359
diff changeset
  1124
  assumes null: "\<And>x. x \<in> X \<Longrightarrow> emeasure M {x} = 0"
57275
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1125
  assumes sets: "\<And>x. x \<in> X \<Longrightarrow> {x} \<in> sets M"
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1126
  shows "AE x in M. x \<notin> X"
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1127
proof -
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1128
  have "(\<Union>x\<in>X. {x}) \<in> null_sets M"
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1129
    using assms by (intro null_sets_UN') auto
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1130
  from AE_not_in[OF this] show "AE x in M. x \<notin> X"
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1131
    by auto
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1132
qed
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
  1133
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1134
lemma AE_finite_all:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1135
  assumes f: "finite S" shows "(AE x in M. \<forall>i\<in>S. P i x) \<longleftrightarrow> (\<forall>i\<in>S. AE x in M. P i x)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1136
  using f by induct auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1137
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1138
lemma AE_finite_allI:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1139
  assumes "finite S"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1140
  shows "(\<And>s. s \<in> S \<Longrightarrow> AE x in M. Q s x) \<Longrightarrow> AE x in M. \<forall>s\<in>S. Q s x"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1141
  using AE_finite_all[OF `finite S`] by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1142
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1143
lemma emeasure_mono_AE:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1144
  assumes imp: "AE x in M. x \<in> A \<longrightarrow> x \<in> B"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1145
    and B: "B \<in> sets M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1146
  shows "emeasure M A \<le> emeasure M B"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1147
proof cases
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1148
  assume A: "A \<in> sets M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1149
  from imp obtain N where N: "{x\<in>space M. \<not> (x \<in> A \<longrightarrow> x \<in> B)} \<subseteq> N" "N \<in> null_sets M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1150
    by (auto simp: eventually_ae_filter)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1151
  have "emeasure M A = emeasure M (A - N)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1152
    using N A by (subst emeasure_Diff_null_set) auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1153
  also have "emeasure M (A - N) \<le> emeasure M (B - N)"
50244
de72bbe42190 qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents: 50104
diff changeset
  1154
    using N A B sets.sets_into_space by (auto intro!: emeasure_mono)
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1155
  also have "emeasure M (B - N) = emeasure M B"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1156
    using N B by (subst emeasure_Diff_null_set) auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1157
  finally show ?thesis .
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1158
qed (simp add: emeasure_nonneg emeasure_notin_sets)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1159
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1160
lemma emeasure_eq_AE:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1161
  assumes iff: "AE x in M. x \<in> A \<longleftrightarrow> x \<in> B"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1162
  assumes A: "A \<in> sets M" and B: "B \<in> sets M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1163
  shows "emeasure M A = emeasure M B"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1164
  using assms by (safe intro!: antisym emeasure_mono_AE) auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1165
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1166
lemma emeasure_Collect_eq_AE:
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1167
  "AE x in M. P x \<longleftrightarrow> Q x \<Longrightarrow> Measurable.pred M Q \<Longrightarrow> Measurable.pred M P \<Longrightarrow>
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1168
   emeasure M {x\<in>space M. P x} = emeasure M {x\<in>space M. Q x}"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1169
   by (intro emeasure_eq_AE) auto
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1170
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1171
lemma emeasure_eq_0_AE: "AE x in M. \<not> P x \<Longrightarrow> emeasure M {x\<in>space M. P x} = 0"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1172
  using AE_iff_measurable[OF _ refl, of M "\<lambda>x. \<not> P x"]
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1173
  by (cases "{x\<in>space M. P x} \<in> sets M") (simp_all add: emeasure_notin_sets)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1174
60715
ee0afee54b1d add emeasure_add_AE
hoelzl
parents: 60714
diff changeset
  1175
lemma emeasure_add_AE:
ee0afee54b1d add emeasure_add_AE
hoelzl
parents: 60714
diff changeset
  1176
  assumes [measurable]: "A \<in> sets M" "B \<in> sets M" "C \<in> sets M"
ee0afee54b1d add emeasure_add_AE
hoelzl
parents: 60714
diff changeset
  1177
  assumes 1: "AE x in M. x \<in> C \<longleftrightarrow> x \<in> A \<or> x \<in> B"
ee0afee54b1d add emeasure_add_AE
hoelzl
parents: 60714
diff changeset
  1178
  assumes 2: "AE x in M. \<not> (x \<in> A \<and> x \<in> B)"
ee0afee54b1d add emeasure_add_AE
hoelzl
parents: 60714
diff changeset
  1179
  shows "emeasure M C = emeasure M A + emeasure M B"
ee0afee54b1d add emeasure_add_AE
hoelzl
parents: 60714
diff changeset
  1180
proof -
ee0afee54b1d add emeasure_add_AE
hoelzl
parents: 60714
diff changeset
  1181
  have "emeasure M C = emeasure M (A \<union> B)"
ee0afee54b1d add emeasure_add_AE
hoelzl
parents: 60714
diff changeset
  1182
    by (rule emeasure_eq_AE) (insert 1, auto)
ee0afee54b1d add emeasure_add_AE
hoelzl
parents: 60714
diff changeset
  1183
  also have "\<dots> = emeasure M A + emeasure M (B - A)"
ee0afee54b1d add emeasure_add_AE
hoelzl
parents: 60714
diff changeset
  1184
    by (subst plus_emeasure) auto
ee0afee54b1d add emeasure_add_AE
hoelzl
parents: 60714
diff changeset
  1185
  also have "emeasure M (B - A) = emeasure M B"
ee0afee54b1d add emeasure_add_AE
hoelzl
parents: 60714
diff changeset
  1186
    by (rule emeasure_eq_AE) (insert 2, auto)
ee0afee54b1d add emeasure_add_AE
hoelzl
parents: 60714
diff changeset
  1187
  finally show ?thesis .
ee0afee54b1d add emeasure_add_AE
hoelzl
parents: 60714
diff changeset
  1188
qed
ee0afee54b1d add emeasure_add_AE
hoelzl
parents: 60714
diff changeset
  1189
56994
8d5e5ec1cac3 fixed document generation for HOL-Probability
hoelzl
parents: 56212
diff changeset
  1190
subsection {* @{text \<sigma>}-finite Measures *}
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1191
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1192
locale sigma_finite_measure =
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1193
  fixes M :: "'a measure"
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1194
  assumes sigma_finite_countable:
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1195
    "\<exists>A::'a set set. countable A \<and> A \<subseteq> sets M \<and> (\<Union>A) = space M \<and> (\<forall>a\<in>A. emeasure M a \<noteq> \<infinity>)"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1196
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1197
lemma (in sigma_finite_measure) sigma_finite:
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1198
  obtains A :: "nat \<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
  1199
  where "range A \<subseteq> sets M" "(\<Union>i. A i) = space M" "\<And>i. emeasure M (A i) \<noteq> \<infinity>"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1200
proof -
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1201
  obtain A :: "'a set set" where
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1202
    [simp]: "countable A" and
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1203
    A: "A \<subseteq> sets M" "(\<Union>A) = space M" "\<And>a. a \<in> A \<Longrightarrow> emeasure M a \<noteq> \<infinity>"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1204
    using sigma_finite_countable 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
  1205
  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
  1206
  proof cases
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1207
    assume "A = {}" with `(\<Union>A) = space M` 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
  1208
      by (intro that[of "\<lambda>_. {}"]) auto
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1209
  next
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61359
diff changeset
  1210
    assume "A \<noteq> {}"
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1211
    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
  1212
    proof
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1213
      show "range (from_nat_into A) \<subseteq> sets M"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1214
        using `A \<noteq> {}` A 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
  1215
      have "(\<Union>i. from_nat_into A i) = \<Union>A"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1216
        using range_from_nat_into[OF `A \<noteq> {}` `countable A`] 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
  1217
      with A show "(\<Union>i. from_nat_into A i) = space M"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1218
        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
  1219
    qed (intro A from_nat_into `A \<noteq> {}`)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1220
  qed
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1221
qed
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1222
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1223
lemma (in sigma_finite_measure) sigma_finite_disjoint:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1224
  obtains A :: "nat \<Rightarrow> 'a set"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1225
  where "range A \<subseteq> sets M" "(\<Union>i. A i) = space M" "\<And>i. emeasure M (A i) \<noteq> \<infinity>" "disjoint_family A"
60580
7e741e22d7fc tuned proofs;
wenzelm
parents: 60172
diff changeset
  1226
proof -
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1227
  obtain A :: "nat \<Rightarrow> 'a set" where
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1228
    range: "range A \<subseteq> sets M" and
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1229
    space: "(\<Union>i. A i) = space M" and
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1230
    measure: "\<And>i. emeasure M (A i) \<noteq> \<infinity>"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1231
    using sigma_finite by auto
60580
7e741e22d7fc tuned proofs;
wenzelm
parents: 60172
diff changeset
  1232
  show thesis
7e741e22d7fc tuned proofs;
wenzelm
parents: 60172
diff changeset
  1233
  proof (rule that[of "disjointed A"])
7e741e22d7fc tuned proofs;
wenzelm
parents: 60172
diff changeset
  1234
    show "range (disjointed A) \<subseteq> sets M"
7e741e22d7fc tuned proofs;
wenzelm
parents: 60172
diff changeset
  1235
      by (rule sets.range_disjointed_sets[OF range])
7e741e22d7fc tuned proofs;
wenzelm
parents: 60172
diff changeset
  1236
    show "(\<Union>i. disjointed A i) = space M"
7e741e22d7fc tuned proofs;
wenzelm
parents: 60172
diff changeset
  1237
      and "disjoint_family (disjointed A)"
7e741e22d7fc tuned proofs;
wenzelm
parents: 60172
diff changeset
  1238
      using disjoint_family_disjointed UN_disjointed_eq[of A] space range
7e741e22d7fc tuned proofs;
wenzelm
parents: 60172
diff changeset
  1239
      by auto
7e741e22d7fc tuned proofs;
wenzelm
parents: 60172
diff changeset
  1240
    show "emeasure M (disjointed A i) \<noteq> \<infinity>" for i
7e741e22d7fc tuned proofs;
wenzelm
parents: 60172
diff changeset
  1241
    proof -
7e741e22d7fc tuned proofs;
wenzelm
parents: 60172
diff changeset
  1242
      have "emeasure M (disjointed A i) \<le> emeasure M (A i)"
7e741e22d7fc tuned proofs;
wenzelm
parents: 60172
diff changeset
  1243
        using range disjointed_subset[of A i] by (auto intro!: emeasure_mono)
7e741e22d7fc tuned proofs;
wenzelm
parents: 60172
diff changeset
  1244
      then show ?thesis using measure[of i] by auto
7e741e22d7fc tuned proofs;
wenzelm
parents: 60172
diff changeset
  1245
    qed
7e741e22d7fc tuned proofs;
wenzelm
parents: 60172
diff changeset
  1246
  qed
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1247
qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1248
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1249
lemma (in sigma_finite_measure) sigma_finite_incseq:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1250
  obtains A :: "nat \<Rightarrow> 'a set"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1251
  where "range A \<subseteq> sets M" "(\<Union>i. A i) = space M" "\<And>i. emeasure M (A i) \<noteq> \<infinity>" "incseq A"
60580
7e741e22d7fc tuned proofs;
wenzelm
parents: 60172
diff changeset
  1252
proof -
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1253
  obtain F :: "nat \<Rightarrow> 'a set" where
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1254
    F: "range F \<subseteq> sets M" "(\<Union>i. F i) = space M" "\<And>i. emeasure M (F i) \<noteq> \<infinity>"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1255
    using sigma_finite by auto
60580
7e741e22d7fc tuned proofs;
wenzelm
parents: 60172
diff changeset
  1256
  show thesis
7e741e22d7fc tuned proofs;
wenzelm
parents: 60172
diff changeset
  1257
  proof (rule that[of "\<lambda>n. \<Union>i\<le>n. F i"])
7e741e22d7fc tuned proofs;
wenzelm
parents: 60172
diff changeset
  1258
    show "range (\<lambda>n. \<Union>i\<le>n. F i) \<subseteq> sets M"
7e741e22d7fc tuned proofs;
wenzelm
parents: 60172
diff changeset
  1259
      using F by (force simp: incseq_def)
7e741e22d7fc tuned proofs;
wenzelm
parents: 60172
diff changeset
  1260
    show "(\<Union>n. \<Union>i\<le>n. F i) = space M"
7e741e22d7fc tuned proofs;
wenzelm
parents: 60172
diff changeset
  1261
    proof -
7e741e22d7fc tuned proofs;
wenzelm
parents: 60172
diff changeset
  1262
      from F have "\<And>x. x \<in> space M \<Longrightarrow> \<exists>i. x \<in> F i" by auto
7e741e22d7fc tuned proofs;
wenzelm
parents: 60172
diff changeset
  1263
      with F show ?thesis by fastforce
7e741e22d7fc tuned proofs;
wenzelm
parents: 60172
diff changeset
  1264
    qed
60585
48fdff264eb2 tuned whitespace;
wenzelm
parents: 60580
diff changeset
  1265
    show "emeasure M (\<Union>i\<le>n. F i) \<noteq> \<infinity>" for n
60580
7e741e22d7fc tuned proofs;
wenzelm
parents: 60172
diff changeset
  1266
    proof -
60585
48fdff264eb2 tuned whitespace;
wenzelm
parents: 60580
diff changeset
  1267
      have "emeasure M (\<Union>i\<le>n. F i) \<le> (\<Sum>i\<le>n. emeasure M (F i))"
60580
7e741e22d7fc tuned proofs;
wenzelm
parents: 60172
diff changeset
  1268
        using F by (auto intro!: emeasure_subadditive_finite)
7e741e22d7fc tuned proofs;
wenzelm
parents: 60172
diff changeset
  1269
      also have "\<dots> < \<infinity>"
7e741e22d7fc tuned proofs;
wenzelm
parents: 60172
diff changeset
  1270
        using F by (auto simp: setsum_Pinfty)
7e741e22d7fc tuned proofs;
wenzelm
parents: 60172
diff changeset
  1271
      finally show ?thesis by simp
7e741e22d7fc tuned proofs;
wenzelm
parents: 60172
diff changeset
  1272
    qed
7e741e22d7fc tuned proofs;
wenzelm
parents: 60172
diff changeset
  1273
    show "incseq (\<lambda>n. \<Union>i\<le>n. F i)"
7e741e22d7fc tuned proofs;
wenzelm
parents: 60172
diff changeset
  1274
      by (force simp: incseq_def)
7e741e22d7fc tuned proofs;
wenzelm
parents: 60172
diff changeset
  1275
  qed
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1276
qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1277
56994
8d5e5ec1cac3 fixed document generation for HOL-Probability
hoelzl
parents: 56212
diff changeset
  1278
subsection {* Measure space induced by distribution of @{const measurable}-functions *}
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1279
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1280
definition distr :: "'a measure \<Rightarrow> 'b measure \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'b measure" where
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1281
  "distr M N f = measure_of (space N) (sets N) (\<lambda>A. emeasure M (f -` A \<inter> space M))"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1282
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1283
lemma
59048
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59011
diff changeset
  1284
  shows sets_distr[simp, measurable_cong]: "sets (distr M N f) = sets N"
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1285
    and space_distr[simp]: "space (distr M N f) = space N"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1286
  by (auto simp: distr_def)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1287
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1288
lemma
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1289
  shows measurable_distr_eq1[simp]: "measurable (distr Mf Nf f) Mf' = measurable Nf Mf'"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1290
    and measurable_distr_eq2[simp]: "measurable Mg' (distr Mg Ng g) = measurable Mg' Ng"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1291
  by (auto simp: measurable_def)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1292
54417
dbb8ecfe1337 add restrict_space measure
hoelzl
parents: 53374
diff changeset
  1293
lemma distr_cong:
dbb8ecfe1337 add restrict_space measure
hoelzl
parents: 53374
diff changeset
  1294
  "M = K \<Longrightarrow> sets N = sets L \<Longrightarrow> (\<And>x. x \<in> space M \<Longrightarrow> f x = g x) \<Longrightarrow> distr M N f = distr K L g"
dbb8ecfe1337 add restrict_space measure
hoelzl
parents: 53374
diff changeset
  1295
  using sets_eq_imp_space_eq[of N L] by (simp add: distr_def Int_def cong: rev_conj_cong)
dbb8ecfe1337 add restrict_space measure
hoelzl
parents: 53374
diff changeset
  1296
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1297
lemma emeasure_distr:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1298
  fixes f :: "'a \<Rightarrow> 'b"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1299
  assumes f: "f \<in> measurable M N" and A: "A \<in> sets N"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1300
  shows "emeasure (distr M N f) A = emeasure M (f -` A \<inter> space M)" (is "_ = ?\<mu> A")
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1301
  unfolding distr_def
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1302
proof (rule emeasure_measure_of_sigma)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1303
  show "positive (sets N) ?\<mu>"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1304
    by (auto simp: positive_def)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1305
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1306
  show "countably_additive (sets N) ?\<mu>"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1307
  proof (intro countably_additiveI)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1308
    fix A :: "nat \<Rightarrow> 'b set" assume "range A \<subseteq> sets N" "disjoint_family A"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1309
    then have A: "\<And>i. A i \<in> sets N" "(\<Union>i. A i) \<in> sets N" by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1310
    then have *: "range (\<lambda>i. f -` (A i) \<inter> space M) \<subseteq> sets M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1311
      using f by (auto simp: measurable_def)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1312
    moreover have "(\<Union>i. f -`  A i \<inter> space M) \<in> sets M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1313
      using * by blast
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1314
    moreover have **: "disjoint_family (\<lambda>i. f -` A i \<inter> space M)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1315
      using `disjoint_family A` by (auto simp: disjoint_family_on_def)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1316
    ultimately show "(\<Sum>i. ?\<mu> (A i)) = ?\<mu> (\<Union>i. A i)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1317
      using suminf_emeasure[OF _ **] A f
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1318
      by (auto simp: comp_def vimage_UN)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1319
  qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1320
  show "sigma_algebra (space N) (sets N)" ..
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1321
qed fact
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1322
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1323
lemma emeasure_Collect_distr:
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1324
  assumes X[measurable]: "X \<in> measurable M N" "Measurable.pred N P"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1325
  shows "emeasure (distr M N X) {x\<in>space N. P x} = emeasure M {x\<in>space M. P (X x)}"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1326
  by (subst emeasure_distr)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1327
     (auto intro!: arg_cong2[where f=emeasure] X(1)[THEN measurable_space])
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1328
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1329
lemma emeasure_lfp2[consumes 1, case_names cont f measurable]:
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1330
  assumes "P M"
60172
423273355b55 rename continuous and down_continuous in Order_Continuity to sup_/inf_continuous; relate them with topological continuity
hoelzl
parents: 60142
diff changeset
  1331
  assumes cont: "sup_continuous F"
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1332
  assumes f: "\<And>M. P M \<Longrightarrow> f \<in> measurable M' M"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1333
  assumes *: "\<And>M A. P M \<Longrightarrow> (\<And>N. P N \<Longrightarrow> Measurable.pred N A) \<Longrightarrow> Measurable.pred M (F A)"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1334
  shows "emeasure M' {x\<in>space M'. lfp F (f x)} = (SUP i. emeasure M' {x\<in>space M'. (F ^^ i) (\<lambda>x. False) (f x)})"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1335
proof (subst (1 2) emeasure_Collect_distr[symmetric, where X=f])
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1336
  show "f \<in> measurable M' M"  "f \<in> measurable M' M"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1337
    using f[OF `P M`] by auto
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1338
  { fix i show "Measurable.pred M ((F ^^ i) (\<lambda>x. False))"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1339
    using `P M` by (induction i arbitrary: M) (auto intro!: *) }
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1340
  show "Measurable.pred M (lfp F)"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1341
    using `P M` cont * by (rule measurable_lfp_coinduct[of P])
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1342
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1343
  have "emeasure (distr M' M f) {x \<in> space (distr M' M f). lfp F x} =
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1344
    (SUP i. emeasure (distr M' M f) {x \<in> space (distr M' M f). (F ^^ i) (\<lambda>x. False) x})"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1345
    using `P M`
60636
ee18efe9b246 add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
hoelzl
parents: 60585
diff changeset
  1346
  proof (coinduction arbitrary: M rule: emeasure_lfp')
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1347
    case (measurable A N) then have "\<And>N. P N \<Longrightarrow> Measurable.pred (distr M' N f) A"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1348
      by metis
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1349
    then have "\<And>N. P N \<Longrightarrow> Measurable.pred N A"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1350
      by simp
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1351
    with `P N`[THEN *] show ?case
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1352
      by auto
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1353
  qed fact
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1354
  then show "emeasure (distr M' M f) {x \<in> space M. lfp F x} =
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1355
    (SUP i. emeasure (distr M' M f) {x \<in> space M. (F ^^ i) (\<lambda>x. False) x})"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1356
   by simp
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1357
qed
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1358
50104
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1359
lemma distr_id[simp]: "distr N N (\<lambda>x. x) = N"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1360
  by (rule measure_eqI) (auto simp: emeasure_distr)
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1361
50001
382bd3173584 add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents: 49789
diff changeset
  1362
lemma measure_distr:
382bd3173584 add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents: 49789
diff changeset
  1363
  "f \<in> measurable M N \<Longrightarrow> S \<in> sets N \<Longrightarrow> measure (distr M N f) S = measure M (f -` S \<inter> space M)"
382bd3173584 add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents: 49789
diff changeset
  1364
  by (simp add: emeasure_distr measure_def)
382bd3173584 add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents: 49789
diff changeset
  1365
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1366
lemma distr_cong_AE:
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61359
diff changeset
  1367
  assumes 1: "M = K" "sets N = sets L" and
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1368
    2: "(AE x in M. f x = g x)" and "f \<in> measurable M N" and "g \<in> measurable 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
  1369
  shows "distr M N f = distr K L g"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1370
proof (rule measure_eqI)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1371
  fix A assume "A \<in> sets (distr M N f)"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1372
  with assms show "emeasure (distr M N f) A = emeasure (distr K L g) A"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1373
    by (auto simp add: emeasure_distr intro!: emeasure_eq_AE measurable_sets)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1374
qed (insert 1, simp)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1375
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1376
lemma AE_distrD:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1377
  assumes f: "f \<in> measurable M M'"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1378
    and AE: "AE x in distr M M' f. P x"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1379
  shows "AE x in M. P (f x)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1380
proof -
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1381
  from AE[THEN AE_E] guess N .
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1382
  with f show ?thesis
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1383
    unfolding eventually_ae_filter
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1384
    by (intro bexI[of _ "f -` N \<inter> space M"])
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1385
       (auto simp: emeasure_distr measurable_def)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1386
qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1387
49773
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
  1388
lemma AE_distr_iff:
50002
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
  1389
  assumes f[measurable]: "f \<in> measurable M N" and P[measurable]: "{x \<in> space N. P x} \<in> sets N"
49773
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
  1390
  shows "(AE x in distr M N f. P x) \<longleftrightarrow> (AE x in M. P (f x))"
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
  1391
proof (subst (1 2) AE_iff_measurable[OF _ refl])
50002
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
  1392
  have "f -` {x\<in>space N. \<not> P x} \<inter> space M = {x \<in> space M. \<not> P (f x)}"
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
  1393
    using f[THEN measurable_space] by auto
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
  1394
  then show "(emeasure (distr M N f) {x \<in> space (distr M N f). \<not> P x} = 0) =
49773
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
  1395
    (emeasure M {x \<in> space M. \<not> P (f x)} = 0)"
50002
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
  1396
    by (simp add: emeasure_distr)
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
  1397
qed auto
49773
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
  1398
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1399
lemma null_sets_distr_iff:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1400
  "f \<in> measurable M N \<Longrightarrow> A \<in> null_sets (distr M N f) \<longleftrightarrow> f -` A \<inter> space M \<in> null_sets M \<and> A \<in> sets N"
50002
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
  1401
  by (auto simp add: null_sets_def emeasure_distr)
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1402
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1403
lemma distr_distr:
50002
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
  1404
  "g \<in> measurable N L \<Longrightarrow> f \<in> measurable M N \<Longrightarrow> distr (distr M N f) L g = distr M L (g \<circ> f)"
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
  1405
  by (auto simp add: emeasure_distr measurable_space
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1406
           intro!: arg_cong[where f="emeasure M"] measure_eqI)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1407
56994
8d5e5ec1cac3 fixed document generation for HOL-Probability
hoelzl
parents: 56212
diff changeset
  1408
subsection {* Real measure values *}
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1409
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1410
lemma measure_nonneg: "0 \<le> measure M A"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1411
  using emeasure_nonneg[of M A] unfolding measure_def by (auto intro: real_of_ereal_pos)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1412
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1413
lemma measure_le_0_iff: "measure M X \<le> 0 \<longleftrightarrow> measure M X = 0"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1414
  using measure_nonneg[of M X] by auto
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1415
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1416
lemma measure_empty[simp]: "measure M {} = 0"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1417
  unfolding measure_def by simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1418
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1419
lemma emeasure_eq_ereal_measure:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1420
  "emeasure M A \<noteq> \<infinity> \<Longrightarrow> emeasure M A = ereal (measure M A)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1421
  using emeasure_nonneg[of M A]
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1422
  by (cases "emeasure M A") (auto simp: measure_def)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1423
61633
64e6d712af16 add lemmas
Andreas Lochbihler
parents: 61609
diff changeset
  1424
lemma max_0_ereal_measure [simp]: "max 0 (ereal (measure M X)) = ereal (measure M X)"
64e6d712af16 add lemmas
Andreas Lochbihler
parents: 61609
diff changeset
  1425
by(simp add: max_def measure_nonneg)
64e6d712af16 add lemmas
Andreas Lochbihler
parents: 61609
diff changeset
  1426
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1427
lemma measure_Union:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1428
  assumes finite: "emeasure M A \<noteq> \<infinity>" "emeasure M B \<noteq> \<infinity>"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1429
  and measurable: "A \<in> sets M" "B \<in> sets M" "A \<inter> B = {}"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1430
  shows "measure M (A \<union> B) = measure M A + measure M B"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1431
  unfolding measure_def
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1432
  using plus_emeasure[OF measurable, symmetric] finite
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1433
  by (simp add: emeasure_eq_ereal_measure)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1434
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1435
lemma measure_finite_Union:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1436
  assumes measurable: "A`S \<subseteq> sets M" "disjoint_family_on A S" "finite S"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1437
  assumes finite: "\<And>i. i \<in> S \<Longrightarrow> emeasure M (A i) \<noteq> \<infinity>"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1438
  shows "measure M (\<Union>i\<in>S. A i) = (\<Sum>i\<in>S. measure M (A i))"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1439
  unfolding measure_def
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1440
  using setsum_emeasure[OF measurable, symmetric] finite
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1441
  by (simp add: emeasure_eq_ereal_measure)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1442
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1443
lemma measure_Diff:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1444
  assumes finite: "emeasure M A \<noteq> \<infinity>"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1445
  and measurable: "A \<in> sets M" "B \<in> sets M" "B \<subseteq> A"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1446
  shows "measure M (A - B) = measure M A - measure M B"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1447
proof -
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1448
  have "emeasure M (A - B) \<le> emeasure M A" "emeasure M B \<le> emeasure M A"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1449
    using measurable by (auto intro!: emeasure_mono)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1450
  hence "measure M ((A - B) \<union> B) = measure M (A - B) + measure M B"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1451
    using measurable finite by (rule_tac measure_Union) auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1452
  thus ?thesis using `B \<subseteq> A` by (auto simp: Un_absorb2)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1453
qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1454
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1455
lemma measure_UNION:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1456
  assumes measurable: "range A \<subseteq> sets M" "disjoint_family A"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1457
  assumes finite: "emeasure M (\<Union>i. A i) \<noteq> \<infinity>"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1458
  shows "(\<lambda>i. measure M (A i)) sums (measure M (\<Union>i. A i))"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1459
proof -
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1460
  from summable_sums[OF summable_ereal_pos, of "\<lambda>i. emeasure M (A i)"]
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1461
       suminf_emeasure[OF measurable] emeasure_nonneg[of M]
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1462
  have "(\<lambda>i. emeasure M (A i)) sums (emeasure M (\<Union>i. A i))" by simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1463
  moreover
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1464
  { fix i
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1465
    have "emeasure M (A i) \<le> emeasure M (\<Union>i. A i)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1466
      using measurable by (auto intro!: emeasure_mono)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1467
    then have "emeasure M (A i) = ereal ((measure M (A i)))"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1468
      using finite by (intro emeasure_eq_ereal_measure) auto }
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1469
  ultimately show ?thesis using finite
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1470
    unfolding sums_ereal[symmetric] by (simp add: emeasure_eq_ereal_measure)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1471
qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1472
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1473
lemma measure_subadditive:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1474
  assumes measurable: "A \<in> sets M" "B \<in> sets M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1475
  and fin: "emeasure M A \<noteq> \<infinity>" "emeasure M B \<noteq> \<infinity>"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1476
  shows "(measure M (A \<union> B)) \<le> (measure M A) + (measure M B)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1477
proof -
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1478
  have "emeasure M (A \<union> B) \<noteq> \<infinity>"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1479
    using emeasure_subadditive[OF measurable] fin by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1480
  then show "(measure M (A \<union> B)) \<le> (measure M A) + (measure M B)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1481
    using emeasure_subadditive[OF measurable] fin
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1482
    by (auto simp: emeasure_eq_ereal_measure)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1483
qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1484
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1485
lemma measure_subadditive_finite:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1486
  assumes A: "finite I" "A`I \<subseteq> sets M" and fin: "\<And>i. i \<in> I \<Longrightarrow> emeasure M (A i) \<noteq> \<infinity>"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1487
  shows "measure M (\<Union>i\<in>I. A i) \<le> (\<Sum>i\<in>I. measure M (A i))"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1488
proof -
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1489
  { have "emeasure M (\<Union>i\<in>I. A i) \<le> (\<Sum>i\<in>I. emeasure M (A i))"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1490
      using emeasure_subadditive_finite[OF A] .
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1491
    also have "\<dots> < \<infinity>"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1492
      using fin by (simp add: setsum_Pinfty)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1493
    finally have "emeasure M (\<Union>i\<in>I. A i) \<noteq> \<infinity>" by simp }
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1494
  then show ?thesis
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1495
    using emeasure_subadditive_finite[OF A] fin
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1496
    unfolding measure_def by (simp add: emeasure_eq_ereal_measure suminf_ereal measure_nonneg)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1497
qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1498
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1499
lemma measure_subadditive_countably:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1500
  assumes A: "range A \<subseteq> sets M" and fin: "(\<Sum>i. emeasure M (A i)) \<noteq> \<infinity>"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1501
  shows "measure M (\<Union>i. A i) \<le> (\<Sum>i. measure M (A i))"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1502
proof -
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1503
  from emeasure_nonneg fin have "\<And>i. emeasure M (A i) \<noteq> \<infinity>" by (rule suminf_PInfty)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1504
  moreover
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1505
  { have "emeasure M (\<Union>i. A i) \<le> (\<Sum>i. emeasure M (A i))"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1506
      using emeasure_subadditive_countably[OF A] .
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1507
    also have "\<dots> < \<infinity>"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1508
      using fin by simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1509
    finally have "emeasure M (\<Union>i. A i) \<noteq> \<infinity>" by simp }
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1510
  ultimately  show ?thesis
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1511
    using emeasure_subadditive_countably[OF A] fin
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1512
    unfolding measure_def by (simp add: emeasure_eq_ereal_measure suminf_ereal measure_nonneg)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1513
qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1514
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1515
lemma measure_eq_setsum_singleton:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1516
  assumes S: "finite S" "\<And>x. x \<in> S \<Longrightarrow> {x} \<in> sets M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1517
  and fin: "\<And>x. x \<in> S \<Longrightarrow> emeasure M {x} \<noteq> \<infinity>"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1518
  shows "(measure M S) = (\<Sum>x\<in>S. (measure M {x}))"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1519
  unfolding measure_def
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1520
  using emeasure_eq_setsum_singleton[OF S] fin
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1521
  by simp (simp add: emeasure_eq_ereal_measure)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1522
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1523
lemma Lim_measure_incseq:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1524
  assumes A: "range A \<subseteq> sets M" "incseq A" and fin: "emeasure M (\<Union>i. A i) \<noteq> \<infinity>"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1525
  shows "(\<lambda>i. (measure M (A i))) ----> (measure M (\<Union>i. A i))"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1526
proof -
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1527
  have "ereal ((measure M (\<Union>i. A i))) = emeasure M (\<Union>i. A i)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1528
    using fin by (auto simp: emeasure_eq_ereal_measure)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1529
  then show ?thesis
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1530
    using Lim_emeasure_incseq[OF A]
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1531
    unfolding measure_def
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1532
    by (intro lim_real_of_ereal) simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1533
qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1534
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1535
lemma Lim_measure_decseq:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1536
  assumes A: "range A \<subseteq> sets M" "decseq A" and fin: "\<And>i. emeasure M (A i) \<noteq> \<infinity>"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1537
  shows "(\<lambda>n. measure M (A n)) ----> measure M (\<Inter>i. A i)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1538
proof -
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1539
  have "emeasure M (\<Inter>i. A i) \<le> emeasure M (A 0)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1540
    using A by (auto intro!: emeasure_mono)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1541
  also have "\<dots> < \<infinity>"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1542
    using fin[of 0] by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1543
  finally have "ereal ((measure M (\<Inter>i. A i))) = emeasure M (\<Inter>i. A i)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1544
    by (auto simp: emeasure_eq_ereal_measure)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1545
  then show ?thesis
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1546
    unfolding measure_def
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1547
    using Lim_emeasure_decseq[OF A fin]
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1548
    by (intro lim_real_of_ereal) simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1549
qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1550
56994
8d5e5ec1cac3 fixed document generation for HOL-Probability
hoelzl
parents: 56212
diff changeset
  1551
subsection {* Measure spaces with @{term "emeasure M (space M) < \<infinity>"} *}
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1552
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1553
locale finite_measure = sigma_finite_measure M for M +
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1554
  assumes finite_emeasure_space: "emeasure M (space M) \<noteq> \<infinity>"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1555
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1556
lemma finite_measureI[Pure.intro!]:
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1557
  "emeasure M (space M) \<noteq> \<infinity> \<Longrightarrow> finite_measure M"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1558
  proof qed (auto intro!: exI[of _ "{space M}"])
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1559
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1560
lemma (in finite_measure) emeasure_finite[simp, intro]: "emeasure M A \<noteq> \<infinity>"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1561
  using finite_emeasure_space emeasure_space[of M A] by auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1562
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1563
lemma (in finite_measure) emeasure_eq_measure: "emeasure M A = ereal (measure M A)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1564
  unfolding measure_def by (simp add: emeasure_eq_ereal_measure)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1565
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1566
lemma (in finite_measure) emeasure_real: "\<exists>r. 0 \<le> r \<and> emeasure M A = ereal r"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1567
  using emeasure_finite[of A] emeasure_nonneg[of M A] by (cases "emeasure M A") auto
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1568
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1569
lemma (in finite_measure) bounded_measure: "measure M A \<le> measure M (space M)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1570
  using emeasure_space[of M A] emeasure_real[of A] emeasure_real[of "space M"] by (auto simp: measure_def)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1571
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1572
lemma (in finite_measure) finite_measure_Diff:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1573
  assumes sets: "A \<in> sets M" "B \<in> sets M" and "B \<subseteq> A"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1574
  shows "measure M (A - B) = measure M A - measure M B"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1575
  using measure_Diff[OF _ assms] by simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1576
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1577
lemma (in finite_measure) finite_measure_Union:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1578
  assumes sets: "A \<in> sets M" "B \<in> sets M" and "A \<inter> B = {}"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1579
  shows "measure M (A \<union> B) = measure M A + measure M B"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1580
  using measure_Union[OF _ _ assms] by simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1581
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1582
lemma (in finite_measure) finite_measure_finite_Union:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1583
  assumes measurable: "A`S \<subseteq> sets M" "disjoint_family_on A S" "finite S"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1584
  shows "measure M (\<Union>i\<in>S. A i) = (\<Sum>i\<in>S. measure M (A i))"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1585
  using measure_finite_Union[OF assms] by simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1586
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1587
lemma (in finite_measure) finite_measure_UNION:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1588
  assumes A: "range A \<subseteq> sets M" "disjoint_family A"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1589
  shows "(\<lambda>i. measure M (A i)) sums (measure M (\<Union>i. A i))"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1590
  using measure_UNION[OF A] by simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1591
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1592
lemma (in finite_measure) finite_measure_mono:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1593
  assumes "A \<subseteq> B" "B \<in> sets M" shows "measure M A \<le> measure M B"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1594
  using emeasure_mono[OF assms] emeasure_real[of A] emeasure_real[of B] by (auto simp: measure_def)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1595
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1596
lemma (in finite_measure) finite_measure_subadditive:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1597
  assumes m: "A \<in> sets M" "B \<in> sets M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1598
  shows "measure M (A \<union> B) \<le> measure M A + measure M B"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1599
  using measure_subadditive[OF m] by simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1600
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1601
lemma (in finite_measure) finite_measure_subadditive_finite:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1602
  assumes "finite I" "A`I \<subseteq> sets M" shows "measure M (\<Union>i\<in>I. A i) \<le> (\<Sum>i\<in>I. measure M (A i))"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1603
  using measure_subadditive_finite[OF assms] by simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1604
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1605
lemma (in finite_measure) finite_measure_subadditive_countably:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1606
  assumes A: "range A \<subseteq> sets M" and sum: "summable (\<lambda>i. measure M (A i))"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1607
  shows "measure M (\<Union>i. A i) \<le> (\<Sum>i. measure M (A i))"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1608
proof -
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1609
  from `summable (\<lambda>i. measure M (A i))`
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1610
  have "(\<lambda>i. ereal (measure M (A i))) sums ereal (\<Sum>i. measure M (A i))"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1611
    by (simp add: sums_ereal) (rule summable_sums)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1612
  from sums_unique[OF this, symmetric]
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1613
       measure_subadditive_countably[OF A]
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1614
  show ?thesis by (simp add: emeasure_eq_measure)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1615
qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1616
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1617
lemma (in finite_measure) finite_measure_eq_setsum_singleton:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1618
  assumes "finite S" and *: "\<And>x. x \<in> S \<Longrightarrow> {x} \<in> sets M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1619
  shows "measure M S = (\<Sum>x\<in>S. measure M {x})"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1620
  using measure_eq_setsum_singleton[OF assms] by simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1621
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1622
lemma (in finite_measure) finite_Lim_measure_incseq:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1623
  assumes A: "range A \<subseteq> sets M" "incseq A"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1624
  shows "(\<lambda>i. measure M (A i)) ----> measure M (\<Union>i. A i)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1625
  using Lim_measure_incseq[OF A] by simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1626
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1627
lemma (in finite_measure) finite_Lim_measure_decseq:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1628
  assumes A: "range A \<subseteq> sets M" "decseq A"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1629
  shows "(\<lambda>n. measure M (A n)) ----> measure M (\<Inter>i. A i)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1630
  using Lim_measure_decseq[OF A] by simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1631
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1632
lemma (in finite_measure) finite_measure_compl:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1633
  assumes S: "S \<in> sets M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1634
  shows "measure M (space M - S) = measure M (space M) - measure M S"
50244
de72bbe42190 qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents: 50104
diff changeset
  1635
  using measure_Diff[OF _ sets.top S sets.sets_into_space] S by simp
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1636
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1637
lemma (in finite_measure) finite_measure_mono_AE:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1638
  assumes imp: "AE x in M. x \<in> A \<longrightarrow> x \<in> B" and B: "B \<in> sets M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1639
  shows "measure M A \<le> measure M B"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1640
  using assms emeasure_mono_AE[OF imp B]
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1641
  by (simp add: emeasure_eq_measure)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1642
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1643
lemma (in finite_measure) finite_measure_eq_AE:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1644
  assumes iff: "AE x in M. x \<in> A \<longleftrightarrow> x \<in> B"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1645
  assumes A: "A \<in> sets M" and B: "B \<in> sets M"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1646
  shows "measure M A = measure M B"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1647
  using assms emeasure_eq_AE[OF assms] by (simp add: emeasure_eq_measure)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1648
50104
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1649
lemma (in finite_measure) measure_increasing: "increasing M (measure M)"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1650
  by (auto intro!: finite_measure_mono simp: increasing_def)
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1651
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1652
lemma (in finite_measure) measure_zero_union:
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1653
  assumes "s \<in> sets M" "t \<in> sets M" "measure M t = 0"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1654
  shows "measure M (s \<union> t) = measure M s"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1655
using assms
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1656
proof -
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1657
  have "measure M (s \<union> t) \<le> measure M s"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1658
    using finite_measure_subadditive[of s t] assms by auto
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1659
  moreover have "measure M (s \<union> t) \<ge> measure M s"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1660
    using assms by (blast intro: finite_measure_mono)
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1661
  ultimately show ?thesis by simp
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1662
qed
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1663
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1664
lemma (in finite_measure) measure_eq_compl:
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1665
  assumes "s \<in> sets M" "t \<in> sets M"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1666
  assumes "measure M (space M - s) = measure M (space M - t)"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1667
  shows "measure M s = measure M t"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1668
  using assms finite_measure_compl by auto
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1669
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1670
lemma (in finite_measure) measure_eq_bigunion_image:
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1671
  assumes "range f \<subseteq> sets M" "range g \<subseteq> sets M"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1672
  assumes "disjoint_family f" "disjoint_family g"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1673
  assumes "\<And> n :: nat. measure M (f n) = measure M (g n)"
60585
48fdff264eb2 tuned whitespace;
wenzelm
parents: 60580
diff changeset
  1674
  shows "measure M (\<Union>i. f i) = measure M (\<Union>i. g i)"
50104
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1675
using assms
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1676
proof -
60585
48fdff264eb2 tuned whitespace;
wenzelm
parents: 60580
diff changeset
  1677
  have a: "(\<lambda> i. measure M (f i)) sums (measure M (\<Union>i. f i))"
50104
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1678
    by (rule finite_measure_UNION[OF assms(1,3)])
60585
48fdff264eb2 tuned whitespace;
wenzelm
parents: 60580
diff changeset
  1679
  have b: "(\<lambda> i. measure M (g i)) sums (measure M (\<Union>i. g i))"
50104
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1680
    by (rule finite_measure_UNION[OF assms(2,4)])
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1681
  show ?thesis using sums_unique[OF b] sums_unique[OF a] assms by simp
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1682
qed
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1683
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1684
lemma (in finite_measure) measure_countably_zero:
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1685
  assumes "range c \<subseteq> sets M"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1686
  assumes "\<And> i. measure M (c i) = 0"
60585
48fdff264eb2 tuned whitespace;
wenzelm
parents: 60580
diff changeset
  1687
  shows "measure M (\<Union>i :: nat. c i) = 0"
50104
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1688
proof (rule antisym)
60585
48fdff264eb2 tuned whitespace;
wenzelm
parents: 60580
diff changeset
  1689
  show "measure M (\<Union>i :: nat. c i) \<le> 0"
50104
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1690
    using finite_measure_subadditive_countably[OF assms(1)] by (simp add: assms(2))
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1691
qed (simp add: measure_nonneg)
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1692
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1693
lemma (in finite_measure) measure_space_inter:
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1694
  assumes events:"s \<in> sets M" "t \<in> sets M"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1695
  assumes "measure M t = measure M (space M)"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1696
  shows "measure M (s \<inter> t) = measure M s"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1697
proof -
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1698
  have "measure M ((space M - s) \<union> (space M - t)) = measure M (space M - s)"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1699
    using events assms finite_measure_compl[of "t"] by (auto intro!: measure_zero_union)
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1700
  also have "(space M - s) \<union> (space M - t) = space M - (s \<inter> t)"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1701
    by blast
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1702
  finally show "measure M (s \<inter> t) = measure M s"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1703
    using events by (auto intro!: measure_eq_compl[of "s \<inter> t" s])
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1704
qed
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1705
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1706
lemma (in finite_measure) measure_equiprobable_finite_unions:
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1707
  assumes s: "finite s" "\<And>x. x \<in> s \<Longrightarrow> {x} \<in> sets M"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1708
  assumes "\<And> x y. \<lbrakk>x \<in> s; y \<in> s\<rbrakk> \<Longrightarrow> measure M {x} = measure M {y}"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1709
  shows "measure M s = real (card s) * measure M {SOME x. x \<in> s}"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1710
proof cases
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1711
  assume "s \<noteq> {}"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1712
  then have "\<exists> x. x \<in> s" by blast
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1713
  from someI_ex[OF this] assms
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1714
  have prob_some: "\<And> x. x \<in> s \<Longrightarrow> measure M {x} = measure M {SOME y. y \<in> s}" by blast
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1715
  have "measure M s = (\<Sum> x \<in> s. measure M {x})"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1716
    using finite_measure_eq_setsum_singleton[OF s] by simp
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1717
  also have "\<dots> = (\<Sum> x \<in> s. measure M {SOME y. y \<in> s})" using prob_some by auto
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1718
  also have "\<dots> = real (card s) * measure M {(SOME x. x \<in> s)}"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61359
diff changeset
  1719
    using setsum_constant assms by simp
50104
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1720
  finally show ?thesis by simp
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1721
qed simp
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1722
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1723
lemma (in finite_measure) measure_real_sum_image_fn:
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1724
  assumes "e \<in> sets M"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1725
  assumes "\<And> x. x \<in> s \<Longrightarrow> e \<inter> f x \<in> sets M"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1726
  assumes "finite s"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1727
  assumes disjoint: "\<And> x y. \<lbrakk>x \<in> s ; y \<in> s ; x \<noteq> y\<rbrakk> \<Longrightarrow> f x \<inter> f y = {}"
60585
48fdff264eb2 tuned whitespace;
wenzelm
parents: 60580
diff changeset
  1728
  assumes upper: "space M \<subseteq> (\<Union>i \<in> s. f i)"
50104
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1729
  shows "measure M e = (\<Sum> x \<in> s. measure M (e \<inter> f x))"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1730
proof -
60585
48fdff264eb2 tuned whitespace;
wenzelm
parents: 60580
diff changeset
  1731
  have e: "e = (\<Union>i \<in> s. e \<inter> f i)"
50244
de72bbe42190 qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents: 50104
diff changeset
  1732
    using `e \<in> sets M` sets.sets_into_space upper by blast
60585
48fdff264eb2 tuned whitespace;
wenzelm
parents: 60580
diff changeset
  1733
  hence "measure M e = measure M (\<Union>i \<in> s. e \<inter> f i)" by simp
50104
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1734
  also have "\<dots> = (\<Sum> x \<in> s. measure M (e \<inter> f x))"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1735
  proof (rule finite_measure_finite_Union)
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1736
    show "finite s" by fact
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1737
    show "(\<lambda>i. e \<inter> f i)`s \<subseteq> sets M" using assms(2) by auto
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1738
    show "disjoint_family_on (\<lambda>i. e \<inter> f i) s"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1739
      using disjoint by (auto simp: disjoint_family_on_def)
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1740
  qed
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1741
  finally show ?thesis .
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1742
qed
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1743
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1744
lemma (in finite_measure) measure_exclude:
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1745
  assumes "A \<in> sets M" "B \<in> sets M"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1746
  assumes "measure M A = measure M (space M)" "A \<inter> B = {}"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1747
  shows "measure M B = 0"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50087
diff changeset
  1748
  using measure_space_inter[of B A] assms by (auto simp: ac_simps)
57235
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57137
diff changeset
  1749
lemma (in finite_measure) finite_measure_distr:
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61359
diff changeset
  1750
  assumes f: "f \<in> measurable M M'"
57235
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57137
diff changeset
  1751
  shows "finite_measure (distr M M' f)"
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57137
diff changeset
  1752
proof (rule finite_measureI)
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57137
diff changeset
  1753
  have "f -` space M' \<inter> space M = space M" using f by (auto dest: measurable_space)
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57137
diff changeset
  1754
  with f show "emeasure (distr M M' f) (space (distr M M' f)) \<noteq> \<infinity>" by (auto simp: emeasure_distr)
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57137
diff changeset
  1755
qed
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57137
diff changeset
  1756
60636
ee18efe9b246 add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
hoelzl
parents: 60585
diff changeset
  1757
lemma emeasure_gfp[consumes 1, case_names cont measurable]:
ee18efe9b246 add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
hoelzl
parents: 60585
diff changeset
  1758
  assumes sets[simp]: "\<And>s. sets (M s) = sets N"
ee18efe9b246 add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
hoelzl
parents: 60585
diff changeset
  1759
  assumes "\<And>s. finite_measure (M s)"
ee18efe9b246 add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
hoelzl
parents: 60585
diff changeset
  1760
  assumes cont: "inf_continuous F" "inf_continuous f"
ee18efe9b246 add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
hoelzl
parents: 60585
diff changeset
  1761
  assumes meas: "\<And>P. Measurable.pred N P \<Longrightarrow> Measurable.pred N (F P)"
ee18efe9b246 add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
hoelzl
parents: 60585
diff changeset
  1762
  assumes iter: "\<And>P s. Measurable.pred N P \<Longrightarrow> emeasure (M s) {x\<in>space N. F P x} = f (\<lambda>s. emeasure (M s) {x\<in>space N. P x}) s"
ee18efe9b246 add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
hoelzl
parents: 60585
diff changeset
  1763
  assumes bound: "\<And>P. f P \<le> f (\<lambda>s. emeasure (M s) (space (M s)))"
ee18efe9b246 add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
hoelzl
parents: 60585
diff changeset
  1764
  shows "emeasure (M s) {x\<in>space N. gfp F x} = gfp f s"
ee18efe9b246 add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
hoelzl
parents: 60585
diff changeset
  1765
proof (subst gfp_transfer_bounded[where \<alpha>="\<lambda>F s. emeasure (M s) {x\<in>space N. F x}" and g=f and f=F and
ee18efe9b246 add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
hoelzl
parents: 60585
diff changeset
  1766
    P="Measurable.pred N", symmetric])
ee18efe9b246 add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
hoelzl
parents: 60585
diff changeset
  1767
  interpret finite_measure "M s" for s by fact
ee18efe9b246 add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
hoelzl
parents: 60585
diff changeset
  1768
  fix C assume "decseq C" "\<And>i. Measurable.pred N (C i)"
ee18efe9b246 add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
hoelzl
parents: 60585
diff changeset
  1769
  then show "(\<lambda>s. emeasure (M s) {x \<in> space N. (INF i. C i) x}) = (INF i. (\<lambda>s. emeasure (M s) {x \<in> space N. C i x}))"
ee18efe9b246 add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
hoelzl
parents: 60585
diff changeset
  1770
    unfolding INF_apply[abs_def]
ee18efe9b246 add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
hoelzl
parents: 60585
diff changeset
  1771
    by (subst INF_emeasure_decseq) (auto simp: antimono_def fun_eq_iff intro!: arg_cong2[where f=emeasure])
ee18efe9b246 add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
hoelzl
parents: 60585
diff changeset
  1772
next
ee18efe9b246 add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
hoelzl
parents: 60585
diff changeset
  1773
  show "f x \<le> (\<lambda>s. emeasure (M s) {x \<in> space N. F top x})" for x
ee18efe9b246 add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
hoelzl
parents: 60585
diff changeset
  1774
    using bound[of x] sets_eq_imp_space_eq[OF sets] by (simp add: iter)
ee18efe9b246 add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
hoelzl
parents: 60585
diff changeset
  1775
qed (auto simp add: iter le_fun_def INF_apply[abs_def] intro!: meas cont)
ee18efe9b246 add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
hoelzl
parents: 60585
diff changeset
  1776
56994
8d5e5ec1cac3 fixed document generation for HOL-Probability
hoelzl
parents: 56212
diff changeset
  1777
subsection {* Counting space *}
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1778
49773
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
  1779
lemma strict_monoI_Suc:
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
  1780
  assumes ord [simp]: "(\<And>n. f n < f (Suc n))" shows "strict_mono f"
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
  1781
  unfolding strict_mono_def
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
  1782
proof safe
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
  1783
  fix n m :: nat assume "n < m" then show "f n < f m"
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
  1784
    by (induct m) (auto simp: less_Suc_eq intro: less_trans ord)
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
  1785
qed
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
  1786
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1787
lemma emeasure_count_space:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1788
  assumes "X \<subseteq> A" shows "emeasure (count_space A) X = (if finite X then ereal (card X) else \<infinity>)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1789
    (is "_ = ?M X")
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1790
  unfolding count_space_def
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1791
proof (rule emeasure_measure_of_sigma)
49773
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
  1792
  show "X \<in> Pow A" using `X \<subseteq> A` by auto
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1793
  show "sigma_algebra A (Pow A)" by (rule sigma_algebra_Pow)
49773
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
  1794
  show positive: "positive (Pow A) ?M"
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1795
    by (auto simp: positive_def)
49773
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
  1796
  have additive: "additive (Pow A) ?M"
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
  1797
    by (auto simp: card_Un_disjoint additive_def)
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1798
49773
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
  1799
  interpret ring_of_sets A "Pow A"
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
  1800
    by (rule ring_of_setsI) auto
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61359
diff changeset
  1801
  show "countably_additive (Pow A) ?M"
49773
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
  1802
    unfolding countably_additive_iff_continuous_from_below[OF positive additive]
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
  1803
  proof safe
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
  1804
    fix F :: "nat \<Rightarrow> 'a set" assume "incseq F"
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
  1805
    show "(\<lambda>i. ?M (F i)) ----> ?M (\<Union>i. F i)"
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
  1806
    proof cases
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
  1807
      assume "\<exists>i. \<forall>j\<ge>i. F i = F j"
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
  1808
      then guess i .. note i = this
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
  1809
      { fix j from i `incseq F` have "F j \<subseteq> F i"
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
  1810
          by (cases "i \<le> j") (auto simp: incseq_def) }
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
  1811
      then have eq: "(\<Union>i. F i) = F i"
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
  1812
        by auto
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
  1813
      with i show ?thesis
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
  1814
        by (auto intro!: Lim_eventually eventually_sequentiallyI[where c=i])
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
  1815
    next
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
  1816
      assume "\<not> (\<exists>i. \<forall>j\<ge>i. F i = F j)"
53374
a14d2a854c02 tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents: 51351
diff changeset
  1817
      then obtain f where f: "\<And>i. i \<le> f i" "\<And>i. F i \<noteq> F (f i)" by metis
a14d2a854c02 tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents: 51351
diff changeset
  1818
      then have "\<And>i. F i \<subseteq> F (f i)" using `incseq F` by (auto simp: incseq_def)
a14d2a854c02 tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents: 51351
diff changeset
  1819
      with f have *: "\<And>i. F i \<subset> F (f i)" by auto
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1820
49773
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
  1821
      have "incseq (\<lambda>i. ?M (F i))"
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
  1822
        using `incseq F` unfolding incseq_def by (auto simp: card_mono dest: finite_subset)
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
  1823
      then have "(\<lambda>i. ?M (F i)) ----> (SUP n. ?M (F n))"
51000
c9adb50f74ad use order topology for extended reals
hoelzl
parents: 50419
diff changeset
  1824
        by (rule LIMSEQ_SUP)
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1825
49773
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
  1826
      moreover have "(SUP n. ?M (F n)) = \<infinity>"
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
  1827
      proof (rule SUP_PInfty)
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
  1828
        fix n :: nat show "\<exists>k::nat\<in>UNIV. ereal n \<le> ?M (F k)"
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
  1829
        proof (induct n)
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
  1830
          case (Suc n)
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
  1831
          then guess k .. note k = this
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
  1832
          moreover have "finite (F k) \<Longrightarrow> finite (F (f k)) \<Longrightarrow> card (F k) < card (F (f k))"
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
  1833
            using `F k \<subset> F (f k)` by (simp add: psubset_card_mono)
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
  1834
          moreover have "finite (F (f k)) \<Longrightarrow> finite (F k)"
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
  1835
            using `k \<le> f k` `incseq F` by (auto simp: incseq_def dest: finite_subset)
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
  1836
          ultimately show ?case
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
  1837
            by (auto intro!: exI[of _ "f k"])
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
  1838
        qed auto
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1839
      qed
49773
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
  1840
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
  1841
      moreover
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
  1842
      have "inj (\<lambda>n. F ((f ^^ n) 0))"
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
  1843
        by (intro strict_mono_imp_inj_on strict_monoI_Suc) (simp add: *)
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
  1844
      then have 1: "infinite (range (\<lambda>i. F ((f ^^ i) 0)))"
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
  1845
        by (rule range_inj_infinite)
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
  1846
      have "infinite (Pow (\<Union>i. F i))"
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
  1847
        by (rule infinite_super[OF _ 1]) auto
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
  1848
      then have "infinite (\<Union>i. F i)"
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
  1849
        by auto
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61359
diff changeset
  1850
49773
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
  1851
      ultimately show ?thesis by auto
16907431e477 tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
hoelzl
parents: 47762
diff changeset
  1852
    qed
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1853
  qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1854
qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1855
59011
4902a2fec434 add reindex rules for distr and nn_integral on count_space
hoelzl
parents: 59000
diff changeset
  1856
lemma distr_bij_count_space:
4902a2fec434 add reindex rules for distr and nn_integral on count_space
hoelzl
parents: 59000
diff changeset
  1857
  assumes f: "bij_betw f A B"
4902a2fec434 add reindex rules for distr and nn_integral on count_space
hoelzl
parents: 59000
diff changeset
  1858
  shows "distr (count_space A) (count_space B) f = count_space B"
4902a2fec434 add reindex rules for distr and nn_integral on count_space
hoelzl
parents: 59000
diff changeset
  1859
proof (rule measure_eqI)
4902a2fec434 add reindex rules for distr and nn_integral on count_space
hoelzl
parents: 59000
diff changeset
  1860
  have f': "f \<in> measurable (count_space A) (count_space B)"
4902a2fec434 add reindex rules for distr and nn_integral on count_space
hoelzl
parents: 59000
diff changeset
  1861
    using f unfolding Pi_def bij_betw_def by auto
4902a2fec434 add reindex rules for distr and nn_integral on count_space
hoelzl
parents: 59000
diff changeset
  1862
  fix X assume "X \<in> sets (distr (count_space A) (count_space B) f)"
4902a2fec434 add reindex rules for distr and nn_integral on count_space
hoelzl
parents: 59000
diff changeset
  1863
  then have X: "X \<in> sets (count_space B)" by auto
4902a2fec434 add reindex rules for distr and nn_integral on count_space
hoelzl
parents: 59000
diff changeset
  1864
  moreover then have "f -` X \<inter> A = the_inv_into A f ` X"
4902a2fec434 add reindex rules for distr and nn_integral on count_space
hoelzl
parents: 59000
diff changeset
  1865
    using f by (auto simp: bij_betw_def subset_image_iff image_iff the_inv_into_f_f intro: the_inv_into_f_f[symmetric])
4902a2fec434 add reindex rules for distr and nn_integral on count_space
hoelzl
parents: 59000
diff changeset
  1866
  moreover have "inj_on (the_inv_into A f) B"
4902a2fec434 add reindex rules for distr and nn_integral on count_space
hoelzl
parents: 59000
diff changeset
  1867
    using X f by (auto simp: bij_betw_def inj_on_the_inv_into)
4902a2fec434 add reindex rules for distr and nn_integral on count_space
hoelzl
parents: 59000
diff changeset
  1868
  with X have "inj_on (the_inv_into A f) X"
4902a2fec434 add reindex rules for distr and nn_integral on count_space
hoelzl
parents: 59000
diff changeset
  1869
    by (auto intro: subset_inj_on)
4902a2fec434 add reindex rules for distr and nn_integral on count_space
hoelzl
parents: 59000
diff changeset
  1870
  ultimately show "emeasure (distr (count_space A) (count_space B) f) X = emeasure (count_space B) X"
4902a2fec434 add reindex rules for distr and nn_integral on count_space
hoelzl
parents: 59000
diff changeset
  1871
    using f unfolding emeasure_distr[OF f' X]
4902a2fec434 add reindex rules for distr and nn_integral on count_space
hoelzl
parents: 59000
diff changeset
  1872
    by (subst (1 2) emeasure_count_space) (auto simp: card_image dest: finite_imageD)
4902a2fec434 add reindex rules for distr and nn_integral on count_space
hoelzl
parents: 59000
diff changeset
  1873
qed simp
4902a2fec434 add reindex rules for distr and nn_integral on count_space
hoelzl
parents: 59000
diff changeset
  1874
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1875
lemma emeasure_count_space_finite[simp]:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1876
  "X \<subseteq> A \<Longrightarrow> finite X \<Longrightarrow> emeasure (count_space A) X = ereal (card X)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1877
  using emeasure_count_space[of X A] by simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1878
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1879
lemma emeasure_count_space_infinite[simp]:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1880
  "X \<subseteq> A \<Longrightarrow> infinite X \<Longrightarrow> emeasure (count_space A) X = \<infinity>"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1881
  using emeasure_count_space[of X A] by simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1882
58606
9c66f7c541fb add Giry monad
hoelzl
parents: 57447
diff changeset
  1883
lemma measure_count_space: "measure (count_space A) X = (if X \<subseteq> A then card X else 0)"
9c66f7c541fb add Giry monad
hoelzl
parents: 57447
diff changeset
  1884
  unfolding measure_def
9c66f7c541fb add Giry monad
hoelzl
parents: 57447
diff changeset
  1885
  by (cases "finite X") (simp_all add: emeasure_notin_sets)
9c66f7c541fb add Giry monad
hoelzl
parents: 57447
diff changeset
  1886
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1887
lemma emeasure_count_space_eq_0:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1888
  "emeasure (count_space A) X = 0 \<longleftrightarrow> (X \<subseteq> A \<longrightarrow> X = {})"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1889
proof cases
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1890
  assume X: "X \<subseteq> A"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1891
  then show ?thesis
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1892
  proof (intro iffI impI)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1893
    assume "emeasure (count_space A) X = 0"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1894
    with X show "X = {}"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1895
      by (subst (asm) emeasure_count_space) (auto split: split_if_asm)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1896
  qed simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1897
qed (simp add: emeasure_notin_sets)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1898
58606
9c66f7c541fb add Giry monad
hoelzl
parents: 57447
diff changeset
  1899
lemma space_empty: "space M = {} \<Longrightarrow> M = count_space {}"
9c66f7c541fb add Giry monad
hoelzl
parents: 57447
diff changeset
  1900
  by (rule measure_eqI) (simp_all add: space_empty_iff)
9c66f7c541fb add Giry monad
hoelzl
parents: 57447
diff changeset
  1901
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1902
lemma null_sets_count_space: "null_sets (count_space A) = { {} }"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1903
  unfolding null_sets_def by (auto simp add: emeasure_count_space_eq_0)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1904
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1905
lemma AE_count_space: "(AE x in count_space A. P x) \<longleftrightarrow> (\<forall>x\<in>A. P x)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1906
  unfolding eventually_ae_filter by (auto simp add: null_sets_count_space)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1907
57025
e7fd64f82876 add various lemmas
hoelzl
parents: 56994
diff changeset
  1908
lemma sigma_finite_measure_count_space_countable:
e7fd64f82876 add various lemmas
hoelzl
parents: 56994
diff changeset
  1909
  assumes A: "countable A"
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1910
  shows "sigma_finite_measure (count_space A)"
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57446
diff changeset
  1911
  proof qed (auto intro!: exI[of _ "(\<lambda>a. {a}) ` A"] simp: A)
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1912
57025
e7fd64f82876 add various lemmas
hoelzl
parents: 56994
diff changeset
  1913
lemma sigma_finite_measure_count_space:
e7fd64f82876 add various lemmas
hoelzl
parents: 56994
diff changeset
  1914
  fixes A :: "'a::countable set" shows "sigma_finite_measure (count_space A)"
e7fd64f82876 add various lemmas
hoelzl
parents: 56994
diff changeset
  1915
  by (rule sigma_finite_measure_count_space_countable) auto
e7fd64f82876 add various lemmas
hoelzl
parents: 56994
diff changeset
  1916
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1917
lemma finite_measure_count_space:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1918
  assumes [simp]: "finite A"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1919
  shows "finite_measure (count_space A)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1920
  by rule simp
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1921
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1922
lemma sigma_finite_measure_count_space_finite:
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1923
  assumes A: "finite A" shows "sigma_finite_measure (count_space A)"
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1924
proof -
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1925
  interpret finite_measure "count_space A" using A by (rule finite_measure_count_space)
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1926
  show "sigma_finite_measure (count_space A)" ..
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1927
qed
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  1928
56994
8d5e5ec1cac3 fixed document generation for HOL-Probability
hoelzl
parents: 56212
diff changeset
  1929
subsection {* Measure restricted to space *}
54417
dbb8ecfe1337 add restrict_space measure
hoelzl
parents: 53374
diff changeset
  1930
dbb8ecfe1337 add restrict_space measure
hoelzl
parents: 53374
diff changeset
  1931
lemma emeasure_restrict_space:
57025
e7fd64f82876 add various lemmas
hoelzl
parents: 56994
diff changeset
  1932
  assumes "\<Omega> \<inter> space M \<in> sets M" "A \<subseteq> \<Omega>"
54417
dbb8ecfe1337 add restrict_space measure
hoelzl
parents: 53374
diff changeset
  1933
  shows "emeasure (restrict_space M \<Omega>) A = emeasure M A"
dbb8ecfe1337 add restrict_space measure
hoelzl
parents: 53374
diff changeset
  1934
proof cases
dbb8ecfe1337 add restrict_space measure
hoelzl
parents: 53374
diff changeset
  1935
  assume "A \<in> sets M"
57025
e7fd64f82876 add various lemmas
hoelzl
parents: 56994
diff changeset
  1936
  show ?thesis
54417
dbb8ecfe1337 add restrict_space measure
hoelzl
parents: 53374
diff changeset
  1937
  proof (rule emeasure_measure_of[OF restrict_space_def])
57025
e7fd64f82876 add various lemmas
hoelzl
parents: 56994
diff changeset
  1938
    show "op \<inter> \<Omega> ` sets M \<subseteq> Pow (\<Omega> \<inter> space M)" "A \<in> sets (restrict_space M \<Omega>)"
e7fd64f82876 add various lemmas
hoelzl
parents: 56994
diff changeset
  1939
      using `A \<subseteq> \<Omega>` `A \<in> sets M` sets.space_closed by (auto simp: sets_restrict_space)
e7fd64f82876 add various lemmas
hoelzl
parents: 56994
diff changeset
  1940
    show "positive (sets (restrict_space M \<Omega>)) (emeasure M)"
54417
dbb8ecfe1337 add restrict_space measure
hoelzl
parents: 53374
diff changeset
  1941
      by (auto simp: positive_def emeasure_nonneg)
57025
e7fd64f82876 add various lemmas
hoelzl
parents: 56994
diff changeset
  1942
    show "countably_additive (sets (restrict_space M \<Omega>)) (emeasure M)"
54417
dbb8ecfe1337 add restrict_space measure
hoelzl
parents: 53374
diff changeset
  1943
    proof (rule countably_additiveI)
dbb8ecfe1337 add restrict_space measure
hoelzl
parents: 53374
diff changeset
  1944
      fix A :: "nat \<Rightarrow> _" assume "range A \<subseteq> sets (restrict_space M \<Omega>)" "disjoint_family A"
dbb8ecfe1337 add restrict_space measure
hoelzl
parents: 53374
diff changeset
  1945
      with assms have "\<And>i. A i \<in> sets M" "\<And>i. A i \<subseteq> space M" "disjoint_family A"
57025
e7fd64f82876 add various lemmas
hoelzl
parents: 56994
diff changeset
  1946
        by (fastforce simp: sets_restrict_space_iff[OF assms(1)] image_subset_iff
e7fd64f82876 add various lemmas
hoelzl
parents: 56994
diff changeset
  1947
                      dest: sets.sets_into_space)+
e7fd64f82876 add various lemmas
hoelzl
parents: 56994
diff changeset
  1948
      then show "(\<Sum>i. emeasure M (A i)) = emeasure M (\<Union>i. A i)"
54417
dbb8ecfe1337 add restrict_space measure
hoelzl
parents: 53374
diff changeset
  1949
        by (subst suminf_emeasure) (auto simp: disjoint_family_subset)
dbb8ecfe1337 add restrict_space measure
hoelzl
parents: 53374
diff changeset
  1950
    qed
dbb8ecfe1337 add restrict_space measure
hoelzl
parents: 53374
diff changeset
  1951
  qed
dbb8ecfe1337 add restrict_space measure
hoelzl
parents: 53374
diff changeset
  1952
next
dbb8ecfe1337 add restrict_space measure
hoelzl
parents: 53374
diff changeset
  1953
  assume "A \<notin> sets M"
dbb8ecfe1337 add restrict_space measure
hoelzl
parents: 53374
diff changeset
  1954
  moreover with assms have "A \<notin> sets (restrict_space M \<Omega>)"
dbb8ecfe1337 add restrict_space measure
hoelzl
parents: 53374
diff changeset
  1955
    by (simp add: sets_restrict_space_iff)
dbb8ecfe1337 add restrict_space measure
hoelzl
parents: 53374
diff changeset
  1956
  ultimately show ?thesis
dbb8ecfe1337 add restrict_space measure
hoelzl
parents: 53374
diff changeset
  1957
    by (simp add: emeasure_notin_sets)
dbb8ecfe1337 add restrict_space measure
hoelzl
parents: 53374
diff changeset
  1958
qed
dbb8ecfe1337 add restrict_space measure
hoelzl
parents: 53374
diff changeset
  1959
57137
f174712d0a84 better support for restrict_space
hoelzl
parents: 57025
diff changeset
  1960
lemma measure_restrict_space:
f174712d0a84 better support for restrict_space
hoelzl
parents: 57025
diff changeset
  1961
  assumes "\<Omega> \<inter> space M \<in> sets M" "A \<subseteq> \<Omega>"
f174712d0a84 better support for restrict_space
hoelzl
parents: 57025
diff changeset
  1962
  shows "measure (restrict_space M \<Omega>) A = measure M A"
f174712d0a84 better support for restrict_space
hoelzl
parents: 57025
diff changeset
  1963
  using emeasure_restrict_space[OF assms] by (simp add: measure_def)
f174712d0a84 better support for restrict_space
hoelzl
parents: 57025
diff changeset
  1964
f174712d0a84 better support for restrict_space
hoelzl
parents: 57025
diff changeset
  1965
lemma AE_restrict_space_iff:
f174712d0a84 better support for restrict_space
hoelzl
parents: 57025
diff changeset
  1966
  assumes "\<Omega> \<inter> space M \<in> sets M"
f174712d0a84 better support for restrict_space
hoelzl
parents: 57025
diff changeset
  1967
  shows "(AE x in restrict_space M \<Omega>. P x) \<longleftrightarrow> (AE x in M. x \<in> \<Omega> \<longrightarrow> P x)"
f174712d0a84 better support for restrict_space
hoelzl
parents: 57025
diff changeset
  1968
proof -
f174712d0a84 better support for restrict_space
hoelzl
parents: 57025
diff changeset
  1969
  have ex_cong: "\<And>P Q f. (\<And>x. P x \<Longrightarrow> Q x) \<Longrightarrow> (\<And>x. Q x \<Longrightarrow> P (f x)) \<Longrightarrow> (\<exists>x. P x) \<longleftrightarrow> (\<exists>x. Q x)"
f174712d0a84 better support for restrict_space
hoelzl
parents: 57025
diff changeset
  1970
    by auto
f174712d0a84 better support for restrict_space
hoelzl
parents: 57025
diff changeset
  1971
  { fix X assume X: "X \<in> sets M" "emeasure M X = 0"
f174712d0a84 better support for restrict_space
hoelzl
parents: 57025
diff changeset
  1972
    then have "emeasure M (\<Omega> \<inter> space M \<inter> X) \<le> emeasure M X"
f174712d0a84 better support for restrict_space
hoelzl
parents: 57025
diff changeset
  1973
      by (intro emeasure_mono) auto
f174712d0a84 better support for restrict_space
hoelzl
parents: 57025
diff changeset
  1974
    then have "emeasure M (\<Omega> \<inter> space M \<inter> X) = 0"
f174712d0a84 better support for restrict_space
hoelzl
parents: 57025
diff changeset
  1975
      using X by (auto intro!: antisym) }
f174712d0a84 better support for restrict_space
hoelzl
parents: 57025
diff changeset
  1976
  with assms show ?thesis
f174712d0a84 better support for restrict_space
hoelzl
parents: 57025
diff changeset
  1977
    unfolding eventually_ae_filter
f174712d0a84 better support for restrict_space
hoelzl
parents: 57025
diff changeset
  1978
    by (auto simp add: space_restrict_space null_sets_def sets_restrict_space_iff
f174712d0a84 better support for restrict_space
hoelzl
parents: 57025
diff changeset
  1979
                       emeasure_restrict_space cong: conj_cong
f174712d0a84 better support for restrict_space
hoelzl
parents: 57025
diff changeset
  1980
             intro!: ex_cong[where f="\<lambda>X. (\<Omega> \<inter> space M) \<inter> X"])
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61359
diff changeset
  1981
qed
57137
f174712d0a84 better support for restrict_space
hoelzl
parents: 57025
diff changeset
  1982
57025
e7fd64f82876 add various lemmas
hoelzl
parents: 56994
diff changeset
  1983
lemma restrict_restrict_space:
e7fd64f82876 add various lemmas
hoelzl
parents: 56994
diff changeset
  1984
  assumes "A \<inter> space M \<in> sets M" "B \<inter> space M \<in> sets M"
e7fd64f82876 add various lemmas
hoelzl
parents: 56994
diff changeset
  1985
  shows "restrict_space (restrict_space M A) B = restrict_space M (A \<inter> B)" (is "?l = ?r")
e7fd64f82876 add various lemmas
hoelzl
parents: 56994
diff changeset
  1986
proof (rule measure_eqI[symmetric])
e7fd64f82876 add various lemmas
hoelzl
parents: 56994
diff changeset
  1987
  show "sets ?r = sets ?l"
e7fd64f82876 add various lemmas
hoelzl
parents: 56994
diff changeset
  1988
    unfolding sets_restrict_space image_comp by (intro image_cong) auto
e7fd64f82876 add various lemmas
hoelzl
parents: 56994
diff changeset
  1989
next
e7fd64f82876 add various lemmas
hoelzl
parents: 56994
diff changeset
  1990
  fix X assume "X \<in> sets (restrict_space M (A \<inter> B))"
e7fd64f82876 add various lemmas
hoelzl
parents: 56994
diff changeset
  1991
  then obtain Y where "Y \<in> sets M" "X = Y \<inter> A \<inter> B"
e7fd64f82876 add various lemmas
hoelzl
parents: 56994
diff changeset
  1992
    by (auto simp: sets_restrict_space)
e7fd64f82876 add various lemmas
hoelzl
parents: 56994
diff changeset
  1993
  with assms sets.Int[OF assms] show "emeasure ?r X = emeasure ?l X"
e7fd64f82876 add various lemmas
hoelzl
parents: 56994
diff changeset
  1994
    by (subst (1 2) emeasure_restrict_space)
e7fd64f82876 add various lemmas
hoelzl
parents: 56994
diff changeset
  1995
       (auto simp: space_restrict_space sets_restrict_space_iff emeasure_restrict_space ac_simps)
e7fd64f82876 add various lemmas
hoelzl
parents: 56994
diff changeset
  1996
qed
e7fd64f82876 add various lemmas
hoelzl
parents: 56994
diff changeset
  1997
e7fd64f82876 add various lemmas
hoelzl
parents: 56994
diff changeset
  1998
lemma restrict_count_space: "restrict_space (count_space B) A = count_space (A \<inter> B)"
54417
dbb8ecfe1337 add restrict_space measure
hoelzl
parents: 53374
diff changeset
  1999
proof (rule measure_eqI)
57025
e7fd64f82876 add various lemmas
hoelzl
parents: 56994
diff changeset
  2000
  show "sets (restrict_space (count_space B) A) = sets (count_space (A \<inter> B))"
e7fd64f82876 add various lemmas
hoelzl
parents: 56994
diff changeset
  2001
    by (subst sets_restrict_space) auto
54417
dbb8ecfe1337 add restrict_space measure
hoelzl
parents: 53374
diff changeset
  2002
  moreover fix X assume "X \<in> sets (restrict_space (count_space B) A)"
57025
e7fd64f82876 add various lemmas
hoelzl
parents: 56994
diff changeset
  2003
  ultimately have "X \<subseteq> A \<inter> B" by auto
e7fd64f82876 add various lemmas
hoelzl
parents: 56994
diff changeset
  2004
  then show "emeasure (restrict_space (count_space B) A) X = emeasure (count_space (A \<inter> B)) X"
54417
dbb8ecfe1337 add restrict_space measure
hoelzl
parents: 53374
diff changeset
  2005
    by (cases "finite X") (auto simp add: emeasure_restrict_space)
dbb8ecfe1337 add restrict_space measure
hoelzl
parents: 53374
diff changeset
  2006
qed
dbb8ecfe1337 add restrict_space measure
hoelzl
parents: 53374
diff changeset
  2007
60063
81835db730e8 add lemmas about restrict_space
Andreas Lochbihler
parents: 59593
diff changeset
  2008
lemma sigma_finite_measure_restrict_space:
81835db730e8 add lemmas about restrict_space
Andreas Lochbihler
parents: 59593
diff changeset
  2009
  assumes "sigma_finite_measure M"
81835db730e8 add lemmas about restrict_space
Andreas Lochbihler
parents: 59593
diff changeset
  2010
  and A: "A \<in> sets M"
81835db730e8 add lemmas about restrict_space
Andreas Lochbihler
parents: 59593
diff changeset
  2011
  shows "sigma_finite_measure (restrict_space M A)"
81835db730e8 add lemmas about restrict_space
Andreas Lochbihler
parents: 59593
diff changeset
  2012
proof -
81835db730e8 add lemmas about restrict_space
Andreas Lochbihler
parents: 59593
diff changeset
  2013
  interpret sigma_finite_measure M by fact
81835db730e8 add lemmas about restrict_space
Andreas Lochbihler
parents: 59593
diff changeset
  2014
  from sigma_finite_countable obtain C
81835db730e8 add lemmas about restrict_space
Andreas Lochbihler
parents: 59593
diff changeset
  2015
    where C: "countable C" "C \<subseteq> sets M" "(\<Union>C) = space M" "\<forall>a\<in>C. emeasure M a \<noteq> \<infinity>"
81835db730e8 add lemmas about restrict_space
Andreas Lochbihler
parents: 59593
diff changeset
  2016
    by blast
81835db730e8 add lemmas about restrict_space
Andreas Lochbihler
parents: 59593
diff changeset
  2017
  let ?C = "op \<inter> A ` C"
81835db730e8 add lemmas about restrict_space
Andreas Lochbihler
parents: 59593
diff changeset
  2018
  from C have "countable ?C" "?C \<subseteq> sets (restrict_space M A)" "(\<Union>?C) = space (restrict_space M A)"
81835db730e8 add lemmas about restrict_space
Andreas Lochbihler
parents: 59593
diff changeset
  2019
    by(auto simp add: sets_restrict_space space_restrict_space)
81835db730e8 add lemmas about restrict_space
Andreas Lochbihler
parents: 59593
diff changeset
  2020
  moreover {
81835db730e8 add lemmas about restrict_space
Andreas Lochbihler
parents: 59593
diff changeset
  2021
    fix a
81835db730e8 add lemmas about restrict_space
Andreas Lochbihler
parents: 59593
diff changeset
  2022
    assume "a \<in> ?C"
81835db730e8 add lemmas about restrict_space
Andreas Lochbihler
parents: 59593
diff changeset
  2023
    then obtain a' where "a = A \<inter> a'" "a' \<in> C" ..
81835db730e8 add lemmas about restrict_space
Andreas Lochbihler
parents: 59593
diff changeset
  2024
    then have "emeasure (restrict_space M A) a \<le> emeasure M a'"
81835db730e8 add lemmas about restrict_space
Andreas Lochbihler
parents: 59593
diff changeset
  2025
      using A C by(auto simp add: emeasure_restrict_space intro: emeasure_mono)
81835db730e8 add lemmas about restrict_space
Andreas Lochbihler
parents: 59593
diff changeset
  2026
    also have "\<dots> < \<infinity>" using C(4)[rule_format, of a'] \<open>a' \<in> C\<close> by simp
81835db730e8 add lemmas about restrict_space
Andreas Lochbihler
parents: 59593
diff changeset
  2027
    finally have "emeasure (restrict_space M A) a \<noteq> \<infinity>" by simp }
81835db730e8 add lemmas about restrict_space
Andreas Lochbihler
parents: 59593
diff changeset
  2028
  ultimately show ?thesis
81835db730e8 add lemmas about restrict_space
Andreas Lochbihler
parents: 59593
diff changeset
  2029
    by unfold_locales (rule exI conjI|assumption|blast)+
81835db730e8 add lemmas about restrict_space
Andreas Lochbihler
parents: 59593
diff changeset
  2030
qed
81835db730e8 add lemmas about restrict_space
Andreas Lochbihler
parents: 59593
diff changeset
  2031
81835db730e8 add lemmas about restrict_space
Andreas Lochbihler
parents: 59593
diff changeset
  2032
lemma finite_measure_restrict_space:
81835db730e8 add lemmas about restrict_space
Andreas Lochbihler
parents: 59593
diff changeset
  2033
  assumes "finite_measure M"
81835db730e8 add lemmas about restrict_space
Andreas Lochbihler
parents: 59593
diff changeset
  2034
  and A: "A \<in> sets M"
81835db730e8 add lemmas about restrict_space
Andreas Lochbihler
parents: 59593
diff changeset
  2035
  shows "finite_measure (restrict_space M A)"
81835db730e8 add lemmas about restrict_space
Andreas Lochbihler
parents: 59593
diff changeset
  2036
proof -
81835db730e8 add lemmas about restrict_space
Andreas Lochbihler
parents: 59593
diff changeset
  2037
  interpret finite_measure M by fact
81835db730e8 add lemmas about restrict_space
Andreas Lochbihler
parents: 59593
diff changeset
  2038
  show ?thesis
81835db730e8 add lemmas about restrict_space
Andreas Lochbihler
parents: 59593
diff changeset
  2039
    by(rule finite_measureI)(simp add: emeasure_restrict_space A space_restrict_space)
81835db730e8 add lemmas about restrict_space
Andreas Lochbihler
parents: 59593
diff changeset
  2040
qed
81835db730e8 add lemmas about restrict_space
Andreas Lochbihler
parents: 59593
diff changeset
  2041
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61359
diff changeset
  2042
lemma restrict_distr:
57137
f174712d0a84 better support for restrict_space
hoelzl
parents: 57025
diff changeset
  2043
  assumes [measurable]: "f \<in> measurable M N"
f174712d0a84 better support for restrict_space
hoelzl
parents: 57025
diff changeset
  2044
  assumes [simp]: "\<Omega> \<inter> space N \<in> sets N" and restrict: "f \<in> space M \<rightarrow> \<Omega>"
f174712d0a84 better support for restrict_space
hoelzl
parents: 57025
diff changeset
  2045
  shows "restrict_space (distr M N f) \<Omega> = distr M (restrict_space N \<Omega>) f"
f174712d0a84 better support for restrict_space
hoelzl
parents: 57025
diff changeset
  2046
  (is "?l = ?r")
f174712d0a84 better support for restrict_space
hoelzl
parents: 57025
diff changeset
  2047
proof (rule measure_eqI)
f174712d0a84 better support for restrict_space
hoelzl
parents: 57025
diff changeset
  2048
  fix A assume "A \<in> sets (restrict_space (distr M N f) \<Omega>)"
f174712d0a84 better support for restrict_space
hoelzl
parents: 57025
diff changeset
  2049
  with restrict show "emeasure ?l A = emeasure ?r A"
f174712d0a84 better support for restrict_space
hoelzl
parents: 57025
diff changeset
  2050
    by (subst emeasure_distr)
f174712d0a84 better support for restrict_space
hoelzl
parents: 57025
diff changeset
  2051
       (auto simp: sets_restrict_space_iff emeasure_restrict_space emeasure_distr
f174712d0a84 better support for restrict_space
hoelzl
parents: 57025
diff changeset
  2052
             intro!: measurable_restrict_space2)
f174712d0a84 better support for restrict_space
hoelzl
parents: 57025
diff changeset
  2053
qed (simp add: sets_restrict_space)
f174712d0a84 better support for restrict_space
hoelzl
parents: 57025
diff changeset
  2054
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  2055
lemma measure_eqI_restrict_generator:
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  2056
  assumes E: "Int_stable E" "E \<subseteq> Pow \<Omega>" "\<And>X. X \<in> E \<Longrightarrow> emeasure M X = emeasure N X"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  2057
  assumes sets_eq: "sets M = sets N" and \<Omega>: "\<Omega> \<in> sets M"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  2058
  assumes "sets (restrict_space M \<Omega>) = sigma_sets \<Omega> E"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61359
diff changeset
  2059
  assumes "sets (restrict_space N \<Omega>) = sigma_sets \<Omega> E"
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  2060
  assumes ae: "AE x in M. x \<in> \<Omega>" "AE x in N. x \<in> \<Omega>"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  2061
  assumes A: "countable A" "A \<noteq> {}" "A \<subseteq> E" "\<Union>A = \<Omega>" "\<And>a. a \<in> A \<Longrightarrow> emeasure M a \<noteq> \<infinity>"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  2062
  shows "M = N"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  2063
proof (rule measure_eqI)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  2064
  fix X assume X: "X \<in> sets M"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  2065
  then have "emeasure M X = emeasure (restrict_space M \<Omega>) (X \<inter> \<Omega>)"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  2066
    using ae \<Omega> by (auto simp add: emeasure_restrict_space intro!: emeasure_eq_AE)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  2067
  also have "restrict_space M \<Omega> = restrict_space N \<Omega>"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  2068
  proof (rule measure_eqI_generator_eq)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  2069
    fix X assume "X \<in> E"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  2070
    then show "emeasure (restrict_space M \<Omega>) X = emeasure (restrict_space N \<Omega>) X"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  2071
      using E \<Omega> by (subst (1 2) emeasure_restrict_space) (auto simp: sets_eq sets_eq[THEN sets_eq_imp_space_eq])
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  2072
  next
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  2073
    show "range (from_nat_into A) \<subseteq> E" "(\<Union>i. from_nat_into A i) = \<Omega>"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  2074
      unfolding Sup_image_eq[symmetric, where f="from_nat_into A"] using A by auto
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  2075
  next
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  2076
    fix i
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  2077
    have "emeasure (restrict_space M \<Omega>) (from_nat_into A i) = emeasure M (from_nat_into A i)"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  2078
      using A \<Omega> by (subst emeasure_restrict_space)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  2079
                   (auto simp: sets_eq sets_eq[THEN sets_eq_imp_space_eq] intro: from_nat_into)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  2080
    with A show "emeasure (restrict_space M \<Omega>) (from_nat_into A i) \<noteq> \<infinity>"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  2081
      by (auto intro: from_nat_into)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  2082
  qed fact+
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  2083
  also have "emeasure (restrict_space N \<Omega>) (X \<inter> \<Omega>) = emeasure N X"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  2084
    using X ae \<Omega> by (auto simp add: emeasure_restrict_space sets_eq intro!: emeasure_eq_AE)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  2085
  finally show "emeasure M X = emeasure N X" .
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  2086
qed fact
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  2087
59425
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59048
diff changeset
  2088
subsection {* Null measure *}
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59048
diff changeset
  2089
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59048
diff changeset
  2090
definition "null_measure M = sigma (space M) (sets M)"
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59048
diff changeset
  2091
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59048
diff changeset
  2092
lemma space_null_measure[simp]: "space (null_measure M) = space M"
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59048
diff changeset
  2093
  by (simp add: null_measure_def)
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59048
diff changeset
  2094
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61359
diff changeset
  2095
lemma sets_null_measure[simp, measurable_cong]: "sets (null_measure M) = sets M"
59425
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59048
diff changeset
  2096
  by (simp add: null_measure_def)
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59048
diff changeset
  2097
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59048
diff changeset
  2098
lemma emeasure_null_measure[simp]: "emeasure (null_measure M) X = 0"
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59048
diff changeset
  2099
  by (cases "X \<in> sets M", rule emeasure_measure_of)
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59048
diff changeset
  2100
     (auto simp: positive_def countably_additive_def emeasure_notin_sets null_measure_def
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59048
diff changeset
  2101
           dest: sets.sets_into_space)
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59048
diff changeset
  2102
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59048
diff changeset
  2103
lemma measure_null_measure[simp]: "measure (null_measure M) X = 0"
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59048
diff changeset
  2104
  by (simp add: measure_def)
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59048
diff changeset
  2105
61633
64e6d712af16 add lemmas
Andreas Lochbihler
parents: 61609
diff changeset
  2106
lemma null_measure_idem [simp]: "null_measure (null_measure M) = null_measure M"
64e6d712af16 add lemmas
Andreas Lochbihler
parents: 61609
diff changeset
  2107
by(rule measure_eqI) simp_all
64e6d712af16 add lemmas
Andreas Lochbihler
parents: 61609
diff changeset
  2108
60772
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2109
subsection \<open>Measures form a chain-complete partial order\<close>
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2110
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2111
instantiation measure :: (type) order_bot
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2112
begin
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2113
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2114
definition bot_measure :: "'a measure" where
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2115
  "bot_measure = sigma {} {{}}"
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2116
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2117
lemma space_bot[simp]: "space bot = {}"
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2118
  unfolding bot_measure_def by (rule space_measure_of) auto
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2119
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2120
lemma sets_bot[simp]: "sets bot = {{}}"
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2121
  unfolding bot_measure_def by (subst sets_measure_of) auto
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2122
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2123
lemma emeasure_bot[simp]: "emeasure bot = (\<lambda>x. 0)"
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2124
  unfolding bot_measure_def by (rule emeasure_sigma)
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2125
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2126
inductive less_eq_measure :: "'a measure \<Rightarrow> 'a measure \<Rightarrow> bool" where
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2127
  "sets N = sets M \<Longrightarrow> (\<And>A. A \<in> sets M \<Longrightarrow> emeasure M A \<le> emeasure N A) \<Longrightarrow> less_eq_measure M N"
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2128
| "less_eq_measure bot N"
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2129
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2130
definition less_measure :: "'a measure \<Rightarrow> 'a measure \<Rightarrow> bool" where
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2131
  "less_measure M N \<longleftrightarrow> (M \<le> N \<and> \<not> N \<le> M)"
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2132
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2133
instance
61166
5976fe402824 renamed method "goals" to "goal_cases" to emphasize its meaning;
wenzelm
parents: 60772
diff changeset
  2134
proof (standard, goal_cases)
60772
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2135
  case 1 then show ?case
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2136
    unfolding less_measure_def ..
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2137
next
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2138
  case (2 M) then show ?case
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2139
    by (intro less_eq_measure.intros) auto
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2140
next
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2141
  case (3 M N L) then show ?case
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2142
    apply (safe elim!: less_eq_measure.cases)
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2143
    apply (simp_all add: less_eq_measure.intros)
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2144
    apply (rule less_eq_measure.intros)
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2145
    apply simp
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2146
    apply (blast intro: order_trans) []
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2147
    unfolding less_eq_measure.simps
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2148
    apply (rule disjI2)
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2149
    apply simp
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2150
    apply (rule measure_eqI)
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2151
    apply (auto intro!: antisym)
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2152
    done
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2153
next
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2154
  case (4 M N) then show ?case
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2155
    apply (safe elim!: less_eq_measure.cases intro!: measure_eqI)
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2156
    apply simp
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2157
    apply simp
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2158
    apply (blast intro: antisym)
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2159
    apply (simp)
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2160
    apply (blast intro: antisym)
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2161
    apply simp
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2162
    done
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2163
qed (rule less_eq_measure.intros)
47694
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  2164
end
05663f75964c reworked Probability theory
hoelzl
parents:
diff changeset
  2165
60772
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2166
lemma le_emeasureD: "M \<le> N \<Longrightarrow> emeasure M A \<le> emeasure N A"
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2167
  by (cases "A \<in> sets M") (auto elim!: less_eq_measure.cases simp: emeasure_notin_sets)
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2168
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2169
lemma le_sets: "N \<le> M \<Longrightarrow> sets N \<le> sets M"
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2170
  unfolding less_eq_measure.simps by auto
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2171
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2172
instantiation measure :: (type) ccpo
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2173
begin
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2174
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2175
definition Sup_measure :: "'a measure set \<Rightarrow> 'a measure" where
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2176
  "Sup_measure A = measure_of (SUP a:A. space a) (SUP a:A. sets a) (SUP a:A. emeasure a)"
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2177
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2178
lemma
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2179
  assumes A: "Complete_Partial_Order.chain op \<le> A" and a: "a \<noteq> bot" "a \<in> A"
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2180
  shows space_Sup: "space (Sup A) = space a"
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2181
    and sets_Sup: "sets (Sup A) = sets a"
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2182
proof -
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2183
  have sets: "(SUP a:A. sets a) = sets a"
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2184
  proof (intro antisym SUP_least)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61359
diff changeset
  2185
    fix a' show "a' \<in> A \<Longrightarrow> sets a' \<subseteq> sets a"
60772
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2186
      using a chainD[OF A, of a a'] by (auto elim!: less_eq_measure.cases)
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2187
  qed (insert \<open>a\<in>A\<close>, auto)
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2188
  have space: "(SUP a:A. space a) = space a"
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2189
  proof (intro antisym SUP_least)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61359
diff changeset
  2190
    fix a' show "a' \<in> A \<Longrightarrow> space a' \<subseteq> space a"
60772
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2191
      using a chainD[OF A, of a a'] by (intro sets_le_imp_space_le) (auto elim!: less_eq_measure.cases)
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2192
  qed (insert \<open>a\<in>A\<close>, auto)
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2193
  show "space (Sup A) = space a"
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2194
    unfolding Sup_measure_def sets space sets.space_measure_of_eq ..
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2195
  show "sets (Sup A) = sets a"
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2196
    unfolding Sup_measure_def sets space sets.sets_measure_of_eq ..
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2197
qed
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2198
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2199
lemma emeasure_Sup:
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2200
  assumes A: "Complete_Partial_Order.chain op \<le> A" "A \<noteq> {}"
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2201
  assumes "X \<in> sets (Sup A)"
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2202
  shows "emeasure (Sup A) X = (SUP a:A. emeasure a) X"
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2203
proof (rule emeasure_measure_of[OF Sup_measure_def])
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2204
  show "countably_additive (sets (Sup A)) (SUP a:A. emeasure a)"
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2205
    unfolding countably_additive_def
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2206
  proof safe
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2207
    fix F :: "nat \<Rightarrow> 'a set" assume F: "range F \<subseteq> sets (Sup A)" "disjoint_family F"
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2208
    show "(\<Sum>i. (SUP a:A. emeasure a) (F i)) = SUPREMUM A emeasure (UNION UNIV F)"
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2209
      unfolding SUP_apply
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2210
    proof (subst suminf_SUP_eq_directed)
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2211
      fix N i j assume "i \<in> A" "j \<in> A"
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2212
      with A(1)
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2213
      show "\<exists>k\<in>A. \<forall>n\<in>N. emeasure i (F n) \<le> emeasure k (F n) \<and> emeasure j (F n) \<le> emeasure k (F n)"
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2214
        by (blast elim: chainE dest: le_emeasureD)
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2215
    next
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2216
      show "(SUP n:A. \<Sum>i. emeasure n (F i)) = (SUP y:A. emeasure y (UNION UNIV F))"
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2217
      proof (intro SUP_cong refl)
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2218
        fix a assume "a \<in> A" then show "(\<Sum>i. emeasure a (F i)) = emeasure a (UNION UNIV F)"
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2219
          using sets_Sup[OF A(1), of a] F by (cases "a = bot") (auto simp: suminf_emeasure)
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2220
      qed
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2221
    qed (insert F \<open>A \<noteq> {}\<close>, auto simp: suminf_emeasure intro!: SUP_cong)
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2222
  qed
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2223
qed (insert \<open>A \<noteq> {}\<close> \<open>X \<in> sets (Sup A)\<close>, auto simp: positive_def dest: sets.sets_into_space intro: SUP_upper2)
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2224
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2225
instance
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2226
proof
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2227
  fix A and x :: "'a measure" assume A: "Complete_Partial_Order.chain op \<le> A" and x: "x \<in> A"
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2228
  show "x \<le> Sup A"
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2229
  proof cases
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2230
    assume "x \<noteq> bot"
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2231
    show ?thesis
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2232
    proof
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2233
      show "sets (Sup A) = sets x"
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2234
        using A \<open>x \<noteq> bot\<close> x by (rule sets_Sup)
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2235
      with x show "\<And>a. a \<in> sets x \<Longrightarrow> emeasure x a \<le> emeasure (Sup A) a"
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2236
        by (subst emeasure_Sup[OF A]) (auto intro: SUP_upper)
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2237
    qed
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2238
  qed simp
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2239
next
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2240
  fix A and x :: "'a measure" assume A: "Complete_Partial_Order.chain op \<le> A" and x: "\<And>z. z \<in> A \<Longrightarrow> z \<le> x"
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2241
  consider "A = {}" | "A = {bot}" | x where "x\<in>A" "x \<noteq> bot"
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2242
    by blast
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2243
  then show "Sup A \<le> x"
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2244
  proof cases
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2245
    assume "A = {}"
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2246
    moreover have "Sup ({}::'a measure set) = bot"
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2247
      by (auto simp add: Sup_measure_def sigma_sets_empty_eq intro!: measure_eqI)
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2248
    ultimately show ?thesis
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2249
      by simp
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2250
  next
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2251
    assume "A = {bot}"
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2252
    moreover have "Sup ({bot}::'a measure set) = bot"
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2253
      by (auto simp add: Sup_measure_def sigma_sets_empty_eq intro!: measure_eqI)
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2254
     ultimately show ?thesis
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2255
      by simp
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2256
  next
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2257
    fix a assume "a \<in> A" "a \<noteq> bot"
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2258
    then have "a \<le> x" "x \<noteq> bot" "a \<noteq> bot"
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2259
      using x[OF \<open>a \<in> A\<close>] by (auto simp: bot_unique)
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2260
    then have "sets x = sets a"
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2261
      by (auto elim: less_eq_measure.cases)
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2262
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2263
    show "Sup A \<le> x"
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2264
    proof (rule less_eq_measure.intros)
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2265
      show "sets x = sets (Sup A)"
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2266
        by (subst sets_Sup[OF A \<open>a \<noteq> bot\<close> \<open>a \<in> A\<close>]) fact
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2267
    next
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2268
      fix X assume "X \<in> sets (Sup A)"
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2269
      then have "emeasure (Sup A) X \<le> (SUP a:A. emeasure a X)"
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2270
        using \<open>a\<in>A\<close> by (subst emeasure_Sup[OF A _]) auto
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2271
      also have "\<dots> \<le> emeasure x X"
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2272
        by (intro SUP_least le_emeasureD x)
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2273
      finally show "emeasure (Sup A) X \<le> emeasure x X" .
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2274
    qed
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2275
  qed
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2276
qed
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2277
end
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2278
61633
64e6d712af16 add lemmas
Andreas Lochbihler
parents: 61609
diff changeset
  2279
lemma
64e6d712af16 add lemmas
Andreas Lochbihler
parents: 61609
diff changeset
  2280
  assumes A: "Complete_Partial_Order.chain op \<le> (f ` A)" and a: "a \<in> A" "f a \<noteq> bot"
64e6d712af16 add lemmas
Andreas Lochbihler
parents: 61609
diff changeset
  2281
  shows space_SUP: "space (SUP M:A. f M) = space (f a)"
64e6d712af16 add lemmas
Andreas Lochbihler
parents: 61609
diff changeset
  2282
    and sets_SUP: "sets (SUP M:A. f M) = sets (f a)"
64e6d712af16 add lemmas
Andreas Lochbihler
parents: 61609
diff changeset
  2283
unfolding SUP_def by(rule space_Sup[OF A a(2) imageI[OF a(1)]] sets_Sup[OF A a(2) imageI[OF a(1)]])+
64e6d712af16 add lemmas
Andreas Lochbihler
parents: 61609
diff changeset
  2284
64e6d712af16 add lemmas
Andreas Lochbihler
parents: 61609
diff changeset
  2285
lemma emeasure_SUP:
64e6d712af16 add lemmas
Andreas Lochbihler
parents: 61609
diff changeset
  2286
  assumes A: "Complete_Partial_Order.chain op \<le> (f ` A)" "A \<noteq> {}"
64e6d712af16 add lemmas
Andreas Lochbihler
parents: 61609
diff changeset
  2287
  assumes "X \<in> sets (SUP M:A. f M)"
64e6d712af16 add lemmas
Andreas Lochbihler
parents: 61609
diff changeset
  2288
  shows "emeasure (SUP M:A. f M) X = (SUP M:A. emeasure (f M)) X"
64e6d712af16 add lemmas
Andreas Lochbihler
parents: 61609
diff changeset
  2289
using \<open>X \<in> _\<close> unfolding SUP_def by(subst emeasure_Sup[OF A(1)]; simp add: A)
64e6d712af16 add lemmas
Andreas Lochbihler
parents: 61609
diff changeset
  2290
60772
a0cfa9050fa8 Measures form a CCPO
hoelzl
parents: 60727
diff changeset
  2291
end