src/HOL/Probability/Probability_Space.thy
author huffman
Wed, 15 Sep 2010 12:54:17 -0700
changeset 39805 16c53975ae1a
parent 39302 d7728f65b353
child 40859 de0b30e6c2d2
permissions -rw-r--r--
clean up definition of compile_pat function
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
     1
theory Probability_Space
39083
e46acc0ea1fe introduced integration on subalgebras
hoelzl
parents: 38656
diff changeset
     2
imports Lebesgue_Integration Radon_Nikodym
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
     3
begin
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
     4
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
     5
locale prob_space = measure_space +
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
     6
  assumes measure_space_1: "\<mu> (space M) = 1"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
     7
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
     8
sublocale prob_space < finite_measure
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
     9
proof
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
    10
  from measure_space_1 show "\<mu> (space M) \<noteq> \<omega>" by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
    11
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
    12
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
    13
context prob_space
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
    14
begin
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
    15
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
    16
abbreviation "events \<equiv> sets M"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
    17
abbreviation "prob \<equiv> \<lambda>A. real (\<mu> A)"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
    18
abbreviation "prob_preserving \<equiv> measure_preserving"
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
    19
abbreviation "random_variable \<equiv> \<lambda> s X. X \<in> measurable M s"
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
    20
abbreviation "expectation \<equiv> integral"
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
    21
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
    22
definition
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
    23
  "indep A B \<longleftrightarrow> A \<in> events \<and> B \<in> events \<and> prob (A \<inter> B) = prob A * prob B"
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
    24
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
    25
definition
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
    26
  "indep_families F G \<longleftrightarrow> (\<forall> A \<in> F. \<forall> B \<in> G. indep A B)"
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
    27
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
    28
definition
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
    29
  "distribution X = (\<lambda>s. \<mu> ((X -` s) \<inter> (space M)))"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
    30
36624
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
    31
abbreviation
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
    32
  "joint_distribution X Y \<equiv> distribution (\<lambda>x. (X x, Y x))"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
    33
39097
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
    34
lemma (in prob_space) distribution_cong:
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
    35
  assumes "\<And>x. x \<in> space M \<Longrightarrow> X x = Y x"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
    36
  shows "distribution X = distribution Y"
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 39198
diff changeset
    37
  unfolding distribution_def fun_eq_iff
39097
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
    38
  using assms by (auto intro!: arg_cong[where f="\<mu>"])
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
    39
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
    40
lemma (in prob_space) joint_distribution_cong:
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
    41
  assumes "\<And>x. x \<in> space M \<Longrightarrow> X x = X' x"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
    42
  assumes "\<And>x. x \<in> space M \<Longrightarrow> Y x = Y' x"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
    43
  shows "joint_distribution X Y = joint_distribution X' Y'"
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 39198
diff changeset
    44
  unfolding distribution_def fun_eq_iff
39097
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
    45
  using assms by (auto intro!: arg_cong[where f="\<mu>"])
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
    46
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
    47
lemma prob_space: "prob (space M) = 1"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
    48
  unfolding measure_space_1 by simp
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
    49
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
    50
lemma measure_le_1[simp, intro]:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
    51
  assumes "A \<in> events" shows "\<mu> A \<le> 1"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
    52
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
    53
  have "\<mu> A \<le> \<mu> (space M)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
    54
    using assms sets_into_space by(auto intro!: measure_mono)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
    55
  also note measure_space_1
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
    56
  finally show ?thesis .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
    57
qed
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
    58
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
    59
lemma prob_compl:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
    60
  assumes "A \<in> events"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
    61
  shows "prob (space M - A) = 1 - prob A"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
    62
  using `A \<in> events`[THEN sets_into_space] `A \<in> events` measure_space_1
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
    63
  by (subst real_finite_measure_Diff) auto
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
    64
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
    65
lemma indep_space:
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
    66
  assumes "s \<in> events"
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
    67
  shows "indep (space M) s"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
    68
  using assms prob_space by (simp add: indep_def)
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
    69
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
    70
lemma prob_space_increasing: "increasing M prob"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
    71
  by (auto intro!: real_measure_mono simp: increasing_def)
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
    72
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
    73
lemma prob_zero_union:
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
    74
  assumes "s \<in> events" "t \<in> events" "prob t = 0"
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
    75
  shows "prob (s \<union> t) = prob s"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
    76
using assms
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
    77
proof -
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
    78
  have "prob (s \<union> t) \<le> prob s"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
    79
    using real_finite_measure_subadditive[of s t] assms by auto
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
    80
  moreover have "prob (s \<union> t) \<ge> prob s"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
    81
    using assms by (blast intro: real_measure_mono)
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
    82
  ultimately show ?thesis by simp
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
    83
qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
    84
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
    85
lemma prob_eq_compl:
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
    86
  assumes "s \<in> events" "t \<in> events"
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
    87
  assumes "prob (space M - s) = prob (space M - t)"
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
    88
  shows "prob s = prob t"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
    89
  using assms prob_compl by auto
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
    90
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
    91
lemma prob_one_inter:
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
    92
  assumes events:"s \<in> events" "t \<in> events"
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
    93
  assumes "prob t = 1"
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
    94
  shows "prob (s \<inter> t) = prob s"
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
    95
proof -
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
    96
  have "prob ((space M - s) \<union> (space M - t)) = prob (space M - s)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
    97
    using events assms  prob_compl[of "t"] by (auto intro!: prob_zero_union)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
    98
  also have "(space M - s) \<union> (space M - t) = space M - (s \<inter> t)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
    99
    by blast
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   100
  finally show "prob (s \<inter> t) = prob s"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   101
    using events by (auto intro!: prob_eq_compl[of "s \<inter> t" s])
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   102
qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   103
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   104
lemma prob_eq_bigunion_image:
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   105
  assumes "range f \<subseteq> events" "range g \<subseteq> events"
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   106
  assumes "disjoint_family f" "disjoint_family g"
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   107
  assumes "\<And> n :: nat. prob (f n) = prob (g n)"
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   108
  shows "(prob (\<Union> i. f i) = prob (\<Union> i. g i))"
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   109
using assms
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   110
proof -
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   111
  have a: "(\<lambda> i. prob (f i)) sums (prob (\<Union> i. f i))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   112
    by (rule real_finite_measure_UNION[OF assms(1,3)])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   113
  have b: "(\<lambda> i. prob (g i)) sums (prob (\<Union> i. g i))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   114
    by (rule real_finite_measure_UNION[OF assms(2,4)])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   115
  show ?thesis using sums_unique[OF b] sums_unique[OF a] assms by simp
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   116
qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   117
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   118
lemma prob_countably_zero:
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   119
  assumes "range c \<subseteq> events"
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   120
  assumes "\<And> i. prob (c i) = 0"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   121
  shows "prob (\<Union> i :: nat. c i) = 0"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   122
proof (rule antisym)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   123
  show "prob (\<Union> i :: nat. c i) \<le> 0"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   124
    using real_finite_measurable_countably_subadditive[OF assms(1)]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   125
    by (simp add: assms(2) suminf_zero summable_zero)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   126
  show "0 \<le> prob (\<Union> i :: nat. c i)" by (rule real_pinfreal_nonneg)
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   127
qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   128
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   129
lemma indep_sym:
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   130
   "indep a b \<Longrightarrow> indep b a"
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   131
unfolding indep_def using Int_commute[of a b] by auto
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   132
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   133
lemma indep_refl:
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   134
  assumes "a \<in> events"
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   135
  shows "indep a a = (prob a = 0) \<or> (prob a = 1)"
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   136
using assms unfolding indep_def by auto
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   137
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   138
lemma prob_equiprobable_finite_unions:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   139
  assumes "s \<in> events"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   140
  assumes s_finite: "finite s" "\<And>x. x \<in> s \<Longrightarrow> {x} \<in> events"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   141
  assumes "\<And> x y. \<lbrakk>x \<in> s; y \<in> s\<rbrakk> \<Longrightarrow> (prob {x} = prob {y})"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   142
  shows "prob s = real (card s) * prob {SOME x. x \<in> s}"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   143
proof (cases "s = {}")
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   144
  case False hence "\<exists> x. x \<in> s" by blast
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   145
  from someI_ex[OF this] assms
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   146
  have prob_some: "\<And> x. x \<in> s \<Longrightarrow> prob {x} = prob {SOME y. y \<in> s}" by blast
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   147
  have "prob s = (\<Sum> x \<in> s. prob {x})"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   148
    using real_finite_measure_finite_singelton[OF s_finite] by simp
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   149
  also have "\<dots> = (\<Sum> x \<in> s. prob {SOME y. y \<in> s})" using prob_some by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   150
  also have "\<dots> = real (card s) * prob {(SOME x. x \<in> s)}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   151
    using setsum_constant assms by (simp add: real_eq_of_nat)
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   152
  finally show ?thesis by simp
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   153
qed simp
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   154
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   155
lemma prob_real_sum_image_fn:
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   156
  assumes "e \<in> events"
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   157
  assumes "\<And> x. x \<in> s \<Longrightarrow> e \<inter> f x \<in> events"
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   158
  assumes "finite s"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   159
  assumes disjoint: "\<And> x y. \<lbrakk>x \<in> s ; y \<in> s ; x \<noteq> y\<rbrakk> \<Longrightarrow> f x \<inter> f y = {}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   160
  assumes upper: "space M \<subseteq> (\<Union> i \<in> s. f i)"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   161
  shows "prob e = (\<Sum> x \<in> s. prob (e \<inter> f x))"
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   162
proof -
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   163
  have e: "e = (\<Union> i \<in> s. e \<inter> f i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   164
    using `e \<in> events` sets_into_space upper by blast
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   165
  hence "prob e = prob (\<Union> i \<in> s. e \<inter> f i)" by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   166
  also have "\<dots> = (\<Sum> x \<in> s. prob (e \<inter> f x))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   167
  proof (rule real_finite_measure_finite_Union)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   168
    show "finite s" by fact
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   169
    show "\<And>i. i \<in> s \<Longrightarrow> e \<inter> f i \<in> events" by fact
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   170
    show "disjoint_family_on (\<lambda>i. e \<inter> f i) s"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   171
      using disjoint by (auto simp: disjoint_family_on_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   172
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   173
  finally show ?thesis .
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   174
qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   175
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   176
lemma distribution_prob_space:
39089
df379a447753 vimage of measurable function is a measure space
hoelzl
parents: 39085
diff changeset
   177
  assumes S: "sigma_algebra S" "random_variable S X"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   178
  shows "prob_space S (distribution X)"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   179
proof -
39089
df379a447753 vimage of measurable function is a measure space
hoelzl
parents: 39085
diff changeset
   180
  interpret S: measure_space S "distribution X"
df379a447753 vimage of measurable function is a measure space
hoelzl
parents: 39085
diff changeset
   181
    using measure_space_vimage[OF S(2,1)] unfolding distribution_def .
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   182
  show ?thesis
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   183
  proof
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   184
    have "X -` space S \<inter> space M = space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   185
      using `random_variable S X` by (auto simp: measurable_def)
39089
df379a447753 vimage of measurable function is a measure space
hoelzl
parents: 39085
diff changeset
   186
    then show "distribution X (space S) = 1"
df379a447753 vimage of measurable function is a measure space
hoelzl
parents: 39085
diff changeset
   187
      using measure_space_1 by (simp add: distribution_def)
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   188
  qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   189
qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   190
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   191
lemma distribution_lebesgue_thm1:
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   192
  assumes "random_variable s X"
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   193
  assumes "A \<in> sets s"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   194
  shows "real (distribution X A) = expectation (indicator (X -` A \<inter> space M))"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   195
unfolding distribution_def
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   196
using assms unfolding measurable_def
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   197
using integral_indicator by auto
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   198
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   199
lemma distribution_lebesgue_thm2:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   200
  assumes "sigma_algebra S" "random_variable S X" and "A \<in> sets S"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   201
  shows "distribution X A =
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   202
    measure_space.positive_integral S (distribution X) (indicator A)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   203
  (is "_ = measure_space.positive_integral _ ?D _")
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   204
proof -
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   205
  interpret S: prob_space S "distribution X" using assms(1,2) by (rule distribution_prob_space)
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   206
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   207
  show ?thesis
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   208
    using S.positive_integral_indicator(1)
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   209
    using assms unfolding distribution_def by auto
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   210
qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   211
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   212
lemma finite_expectation1:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   213
  assumes "finite (X`space M)" and rv: "random_variable borel_space X"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   214
  shows "expectation X = (\<Sum> r \<in> X ` space M. r * prob (X -` {r} \<inter> space M))"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   215
proof (rule integral_on_finite(2)[OF assms(2,1)])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   216
  fix x have "X -` {x} \<inter> space M \<in> sets M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   217
    using rv unfolding measurable_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   218
  thus "\<mu> (X -` {x} \<inter> space M) \<noteq> \<omega>" using finite_measure by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   219
qed
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   220
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   221
lemma finite_expectation:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   222
  assumes "finite (space M)" "random_variable borel_space X"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   223
  shows "expectation X = (\<Sum> r \<in> X ` (space M). r * real (distribution X {r}))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   224
  using assms unfolding distribution_def using finite_expectation1 by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   225
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   226
lemma prob_x_eq_1_imp_prob_y_eq_0:
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   227
  assumes "{x} \<in> events"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   228
  assumes "prob {x} = 1"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   229
  assumes "{y} \<in> events"
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   230
  assumes "y \<noteq> x"
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   231
  shows "prob {y} = 0"
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   232
  using prob_one_inter[of "{y}" "{x}"] assms by auto
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   233
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   234
lemma distribution_empty[simp]: "distribution X {} = 0"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   235
  unfolding distribution_def by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   236
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   237
lemma distribution_space[simp]: "distribution X (X ` space M) = 1"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   238
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   239
  have "X -` X ` space M \<inter> space M = space M" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   240
  thus ?thesis unfolding distribution_def by (simp add: measure_space_1)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   241
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   242
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   243
lemma distribution_one:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   244
  assumes "random_variable M X" and "A \<in> events"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   245
  shows "distribution X A \<le> 1"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   246
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   247
  have "distribution X A \<le> \<mu> (space M)" unfolding distribution_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   248
    using assms[unfolded measurable_def] by (auto intro!: measure_mono)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   249
  thus ?thesis by (simp add: measure_space_1)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   250
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   251
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   252
lemma distribution_finite:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   253
  assumes "random_variable M X" and "A \<in> events"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   254
  shows "distribution X A \<noteq> \<omega>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   255
  using distribution_one[OF assms] by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   256
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   257
lemma distribution_x_eq_1_imp_distribution_y_eq_0:
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   258
  assumes X: "random_variable \<lparr>space = X ` (space M), sets = Pow (X ` (space M))\<rparr> X"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   259
    (is "random_variable ?S X")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   260
  assumes "distribution X {x} = 1"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   261
  assumes "y \<noteq> x"
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   262
  shows "distribution X {y} = 0"
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   263
proof -
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   264
  have "sigma_algebra ?S" by (rule sigma_algebra_Pow)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   265
  from distribution_prob_space[OF this X]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   266
  interpret S: prob_space ?S "distribution X" by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   267
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   268
  have x: "{x} \<in> sets ?S"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   269
  proof (rule ccontr)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   270
    assume "{x} \<notin> sets ?S"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   271
    hence "X -` {x} \<inter> space M = {}" by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   272
    thus "False" using assms unfolding distribution_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   273
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   274
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   275
  have [simp]: "{y} \<inter> {x} = {}" "{x} - {y} = {x}" using `y \<noteq> x` by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   276
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   277
  show ?thesis
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   278
  proof cases
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   279
    assume "{y} \<in> sets ?S"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   280
    with `{x} \<in> sets ?S` assms show "distribution X {y} = 0"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   281
      using S.measure_inter_full_set[of "{y}" "{x}"]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   282
      by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   283
  next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   284
    assume "{y} \<notin> sets ?S"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   285
    hence "X -` {y} \<inter> space M = {}" by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   286
    thus "distribution X {y} = 0" unfolding distribution_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   287
  qed
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   288
qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   289
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   290
end
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   291
35977
30d42bfd0174 Added finite measure space.
hoelzl
parents: 35929
diff changeset
   292
locale finite_prob_space = prob_space + finite_measure_space
30d42bfd0174 Added finite measure space.
hoelzl
parents: 35929
diff changeset
   293
36624
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   294
lemma finite_prob_space_eq:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   295
  "finite_prob_space M \<mu> \<longleftrightarrow> finite_measure_space M \<mu> \<and> \<mu> (space M) = 1"
36624
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   296
  unfolding finite_prob_space_def finite_measure_space_def prob_space_def prob_space_axioms_def
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   297
  by auto
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   298
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   299
lemma (in prob_space) not_empty: "space M \<noteq> {}"
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   300
  using prob_space empty_measure by auto
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   301
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   302
lemma (in finite_prob_space) sum_over_space_eq_1: "(\<Sum>x\<in>space M. \<mu> {x}) = 1"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   303
  using measure_space_1 sum_over_space by simp
36624
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   304
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   305
lemma (in finite_prob_space) positive_distribution: "0 \<le> distribution X x"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   306
  unfolding distribution_def by simp
36624
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   307
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   308
lemma (in finite_prob_space) joint_distribution_restriction_fst:
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   309
  "joint_distribution X Y A \<le> distribution X (fst ` A)"
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   310
  unfolding distribution_def
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   311
proof (safe intro!: measure_mono)
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   312
  fix x assume "x \<in> space M" and *: "(X x, Y x) \<in> A"
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   313
  show "x \<in> X -` fst ` A"
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   314
    by (auto intro!: image_eqI[OF _ *])
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   315
qed (simp_all add: sets_eq_Pow)
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   316
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   317
lemma (in finite_prob_space) joint_distribution_restriction_snd:
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   318
  "joint_distribution X Y A \<le> distribution Y (snd ` A)"
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   319
  unfolding distribution_def
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   320
proof (safe intro!: measure_mono)
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   321
  fix x assume "x \<in> space M" and *: "(X x, Y x) \<in> A"
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   322
  show "x \<in> Y -` snd ` A"
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   323
    by (auto intro!: image_eqI[OF _ *])
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   324
qed (simp_all add: sets_eq_Pow)
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   325
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   326
lemma (in finite_prob_space) distribution_order:
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   327
  shows "0 \<le> distribution X x'"
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   328
  and "(distribution X x' \<noteq> 0) \<longleftrightarrow> (0 < distribution X x')"
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   329
  and "r \<le> joint_distribution X Y {(x, y)} \<Longrightarrow> r \<le> distribution X {x}"
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   330
  and "r \<le> joint_distribution X Y {(x, y)} \<Longrightarrow> r \<le> distribution Y {y}"
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   331
  and "r < joint_distribution X Y {(x, y)} \<Longrightarrow> r < distribution X {x}"
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   332
  and "r < joint_distribution X Y {(x, y)} \<Longrightarrow> r < distribution Y {y}"
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   333
  and "distribution X {x} = 0 \<Longrightarrow> joint_distribution X Y {(x, y)} = 0"
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   334
  and "distribution Y {y} = 0 \<Longrightarrow> joint_distribution X Y {(x, y)} = 0"
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   335
  using positive_distribution[of X x']
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   336
    positive_distribution[of "\<lambda>x. (X x, Y x)" "{(x, y)}"]
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   337
    joint_distribution_restriction_fst[of X Y "{(x, y)}"]
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   338
    joint_distribution_restriction_snd[of X Y "{(x, y)}"]
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   339
  by auto
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   340
39097
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   341
lemma (in finite_prob_space) distribution_mono:
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   342
  assumes "\<And>t. \<lbrakk> t \<in> space M ; X t \<in> x \<rbrakk> \<Longrightarrow> Y t \<in> y"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   343
  shows "distribution X x \<le> distribution Y y"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   344
  unfolding distribution_def
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   345
  using assms by (auto simp: sets_eq_Pow intro!: measure_mono)
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   346
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   347
lemma (in finite_prob_space) distribution_mono_gt_0:
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   348
  assumes gt_0: "0 < distribution X x"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   349
  assumes *: "\<And>t. \<lbrakk> t \<in> space M ; X t \<in> x \<rbrakk> \<Longrightarrow> Y t \<in> y"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   350
  shows "0 < distribution Y y"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   351
  by (rule less_le_trans[OF gt_0 distribution_mono]) (rule *)
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   352
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   353
lemma (in finite_prob_space) sum_over_space_distrib:
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   354
  "(\<Sum>x\<in>X`space M. distribution X {x}) = 1"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   355
  unfolding distribution_def measure_space_1[symmetric] using finite_space
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   356
  by (subst measure_finitely_additive'')
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   357
     (auto simp add: disjoint_family_on_def sets_eq_Pow intro!: arg_cong[where f=\<mu>])
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   358
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   359
lemma (in finite_prob_space) sum_over_space_real_distribution:
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   360
  "(\<Sum>x\<in>X`space M. real (distribution X {x})) = 1"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   361
  unfolding distribution_def prob_space[symmetric] using finite_space
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   362
  by (subst real_finite_measure_finite_Union[symmetric])
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   363
     (auto simp add: disjoint_family_on_def sets_eq_Pow intro!: arg_cong[where f=prob])
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   364
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   365
lemma (in finite_prob_space) finite_sum_over_space_eq_1:
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   366
  "(\<Sum>x\<in>space M. real (\<mu> {x})) = 1"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   367
  using sum_over_space_eq_1 finite_measure by (simp add: real_of_pinfreal_setsum sets_eq_Pow)
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   368
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   369
lemma (in finite_prob_space) distribution_finite:
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   370
  "distribution X A \<noteq> \<omega>"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   371
  using finite_measure[of "X -` A \<inter> space M"]
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   372
  unfolding distribution_def sets_eq_Pow by auto
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   373
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   374
lemma (in finite_prob_space) real_distribution_gt_0[simp]:
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   375
  "0 < real (distribution Y y) \<longleftrightarrow>  0 < distribution Y y"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   376
  using assms by (auto intro!: real_pinfreal_pos distribution_finite)
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   377
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   378
lemma (in finite_prob_space) real_distribution_mult_pos_pos:
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   379
  assumes "0 < distribution Y y"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   380
  and "0 < distribution X x"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   381
  shows "0 < real (distribution Y y * distribution X x)"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   382
  unfolding real_of_pinfreal_mult[symmetric]
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   383
  using assms by (auto intro!: mult_pos_pos)
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   384
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   385
lemma (in finite_prob_space) real_distribution_divide_pos_pos:
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   386
  assumes "0 < distribution Y y"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   387
  and "0 < distribution X x"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   388
  shows "0 < real (distribution Y y / distribution X x)"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   389
  unfolding divide_pinfreal_def real_of_pinfreal_mult[symmetric]
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   390
  using assms distribution_finite[of X x] by (cases "distribution X x") (auto intro!: mult_pos_pos)
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   391
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   392
lemma (in finite_prob_space) real_distribution_mult_inverse_pos_pos:
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   393
  assumes "0 < distribution Y y"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   394
  and "0 < distribution X x"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   395
  shows "0 < real (distribution Y y * inverse (distribution X x))"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   396
  unfolding divide_pinfreal_def real_of_pinfreal_mult[symmetric]
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   397
  using assms distribution_finite[of X x] by (cases "distribution X x") (auto intro!: mult_pos_pos)
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   398
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   399
lemma (in prob_space) distribution_remove_const:
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   400
  shows "joint_distribution X (\<lambda>x. ()) {(x, ())} = distribution X {x}"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   401
  and "joint_distribution (\<lambda>x. ()) X {((), x)} = distribution X {x}"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   402
  and "joint_distribution X (\<lambda>x. (Y x, ())) {(x, y, ())} = joint_distribution X Y {(x, y)}"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   403
  and "joint_distribution X (\<lambda>x. ((), Y x)) {(x, (), y)} = joint_distribution X Y {(x, y)}"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   404
  and "distribution (\<lambda>x. ()) {()} = 1"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   405
  unfolding measure_space_1[symmetric]
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   406
  by (auto intro!: arg_cong[where f="\<mu>"] simp: distribution_def)
35977
30d42bfd0174 Added finite measure space.
hoelzl
parents: 35929
diff changeset
   407
39097
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   408
lemma (in finite_prob_space) setsum_distribution_gen:
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   409
  assumes "Z -` {c} \<inter> space M = (\<Union>x \<in> X`space M. Y -` {f x}) \<inter> space M"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   410
  and "inj_on f (X`space M)"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   411
  shows "(\<Sum>x \<in> X`space M. distribution Y {f x}) = distribution Z {c}"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   412
  unfolding distribution_def assms
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   413
  using finite_space assms
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   414
  by (subst measure_finitely_additive'')
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   415
     (auto simp add: disjoint_family_on_def sets_eq_Pow inj_on_def
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   416
      intro!: arg_cong[where f=prob])
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   417
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   418
lemma (in finite_prob_space) setsum_distribution:
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   419
  "(\<Sum>x \<in> X`space M. joint_distribution X Y {(x, y)}) = distribution Y {y}"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   420
  "(\<Sum>y \<in> Y`space M. joint_distribution X Y {(x, y)}) = distribution X {x}"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   421
  "(\<Sum>x \<in> X`space M. joint_distribution X (\<lambda>x. (Y x, Z x)) {(x, y, z)}) = joint_distribution Y Z {(y, z)}"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   422
  "(\<Sum>y \<in> Y`space M. joint_distribution X (\<lambda>x. (Y x, Z x)) {(x, y, z)}) = joint_distribution X Z {(x, z)}"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   423
  "(\<Sum>z \<in> Z`space M. joint_distribution X (\<lambda>x. (Y x, Z x)) {(x, y, z)}) = joint_distribution X Y {(x, y)}"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   424
  by (auto intro!: inj_onI setsum_distribution_gen)
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   425
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   426
lemma (in finite_prob_space) setsum_real_distribution_gen:
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   427
  assumes "Z -` {c} \<inter> space M = (\<Union>x \<in> X`space M. Y -` {f x}) \<inter> space M"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   428
  and "inj_on f (X`space M)"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   429
  shows "(\<Sum>x \<in> X`space M. real (distribution Y {f x})) = real (distribution Z {c})"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   430
  unfolding distribution_def assms
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   431
  using finite_space assms
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   432
  by (subst real_finite_measure_finite_Union[symmetric])
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   433
     (auto simp add: disjoint_family_on_def sets_eq_Pow inj_on_def
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   434
        intro!: arg_cong[where f=prob])
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   435
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   436
lemma (in finite_prob_space) setsum_real_distribution:
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   437
  "(\<Sum>x \<in> X`space M. real (joint_distribution X Y {(x, y)})) = real (distribution Y {y})"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   438
  "(\<Sum>y \<in> Y`space M. real (joint_distribution X Y {(x, y)})) = real (distribution X {x})"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   439
  "(\<Sum>x \<in> X`space M. real (joint_distribution X (\<lambda>x. (Y x, Z x)) {(x, y, z)})) = real (joint_distribution Y Z {(y, z)})"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   440
  "(\<Sum>y \<in> Y`space M. real (joint_distribution X (\<lambda>x. (Y x, Z x)) {(x, y, z)})) = real (joint_distribution X Z {(x, z)})"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   441
  "(\<Sum>z \<in> Z`space M. real (joint_distribution X (\<lambda>x. (Y x, Z x)) {(x, y, z)})) = real (joint_distribution X Y {(x, y)})"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   442
  by (auto intro!: inj_onI setsum_real_distribution_gen)
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   443
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   444
lemma (in finite_prob_space) real_distribution_order:
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   445
  shows "r \<le> real (joint_distribution X Y {(x, y)}) \<Longrightarrow> r \<le> real (distribution X {x})"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   446
  and "r \<le> real (joint_distribution X Y {(x, y)}) \<Longrightarrow> r \<le> real (distribution Y {y})"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   447
  and "r < real (joint_distribution X Y {(x, y)}) \<Longrightarrow> r < real (distribution X {x})"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   448
  and "r < real (joint_distribution X Y {(x, y)}) \<Longrightarrow> r < real (distribution Y {y})"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   449
  and "distribution X {x} = 0 \<Longrightarrow> real (joint_distribution X Y {(x, y)}) = 0"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   450
  and "distribution Y {y} = 0 \<Longrightarrow> real (joint_distribution X Y {(x, y)}) = 0"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   451
  using real_of_pinfreal_mono[OF distribution_finite joint_distribution_restriction_fst, of X Y "{(x, y)}"]
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   452
  using real_of_pinfreal_mono[OF distribution_finite joint_distribution_restriction_snd, of X Y "{(x, y)}"]
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   453
  using real_pinfreal_nonneg[of "joint_distribution X Y {(x, y)}"]
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   454
  by auto
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   455
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   456
lemma (in prob_space) joint_distribution_remove[simp]:
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   457
    "joint_distribution X X {(x, x)} = distribution X {x}"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   458
  unfolding distribution_def by (auto intro!: arg_cong[where f="\<mu>"])
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   459
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   460
lemma (in finite_prob_space) distribution_1:
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   461
  "distribution X A \<le> 1"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   462
  unfolding distribution_def measure_space_1[symmetric]
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   463
  by (auto intro!: measure_mono simp: sets_eq_Pow)
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   464
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   465
lemma (in finite_prob_space) real_distribution_1:
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   466
  "real (distribution X A) \<le> 1"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   467
  unfolding real_pinfreal_1[symmetric]
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   468
  by (rule real_of_pinfreal_mono[OF _ distribution_1]) simp
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   469
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   470
lemma (in finite_prob_space) uniform_prob:
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   471
  assumes "x \<in> space M"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   472
  assumes "\<And> x y. \<lbrakk>x \<in> space M ; y \<in> space M\<rbrakk> \<Longrightarrow> prob {x} = prob {y}"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   473
  shows "prob {x} = 1 / real (card (space M))"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   474
proof -
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   475
  have prob_x: "\<And> y. y \<in> space M \<Longrightarrow> prob {y} = prob {x}"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   476
    using assms(2)[OF _ `x \<in> space M`] by blast
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   477
  have "1 = prob (space M)"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   478
    using prob_space by auto
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   479
  also have "\<dots> = (\<Sum> x \<in> space M. prob {x})"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   480
    using real_finite_measure_finite_Union[of "space M" "\<lambda> x. {x}", simplified]
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   481
      sets_eq_Pow inj_singleton[unfolded inj_on_def, rule_format]
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   482
      finite_space unfolding disjoint_family_on_def  prob_space[symmetric]
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   483
    by (auto simp add:setsum_restrict_set)
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   484
  also have "\<dots> = (\<Sum> y \<in> space M. prob {x})"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   485
    using prob_x by auto
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   486
  also have "\<dots> = real_of_nat (card (space M)) * prob {x}" by simp
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   487
  finally have one: "1 = real (card (space M)) * prob {x}"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   488
    using real_eq_of_nat by auto
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   489
  hence two: "real (card (space M)) \<noteq> 0" by fastsimp 
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   490
  from one have three: "prob {x} \<noteq> 0" by fastsimp
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   491
  thus ?thesis using one two three divide_cancel_right
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   492
    by (auto simp:field_simps)
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39091
diff changeset
   493
qed
35977
30d42bfd0174 Added finite measure space.
hoelzl
parents: 35929
diff changeset
   494
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39091
diff changeset
   495
lemma (in prob_space) prob_space_subalgebra:
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39091
diff changeset
   496
  assumes "N \<subseteq> sets M" "sigma_algebra (M\<lparr> sets := N \<rparr>)"
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39091
diff changeset
   497
  shows "prob_space (M\<lparr> sets := N \<rparr>) \<mu>"
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39091
diff changeset
   498
proof -
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39091
diff changeset
   499
  interpret N: measure_space "M\<lparr> sets := N \<rparr>" \<mu>
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39091
diff changeset
   500
    using measure_space_subalgebra[OF assms] .
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39091
diff changeset
   501
  show ?thesis
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39091
diff changeset
   502
    proof qed (simp add: measure_space_1)
35977
30d42bfd0174 Added finite measure space.
hoelzl
parents: 35929
diff changeset
   503
qed
30d42bfd0174 Added finite measure space.
hoelzl
parents: 35929
diff changeset
   504
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39091
diff changeset
   505
lemma (in prob_space) prob_space_of_restricted_space:
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39091
diff changeset
   506
  assumes "\<mu> A \<noteq> 0" "\<mu> A \<noteq> \<omega>" "A \<in> sets M"
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39091
diff changeset
   507
  shows "prob_space (restricted_space A) (\<lambda>S. \<mu> S / \<mu> A)"
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39091
diff changeset
   508
  unfolding prob_space_def prob_space_axioms_def
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39091
diff changeset
   509
proof
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39091
diff changeset
   510
  show "\<mu> (space (restricted_space A)) / \<mu> A = 1"
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39091
diff changeset
   511
    using `\<mu> A \<noteq> 0` `\<mu> A \<noteq> \<omega>` by (auto simp: pinfreal_noteq_omega_Ex)
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39091
diff changeset
   512
  have *: "\<And>S. \<mu> S / \<mu> A = inverse (\<mu> A) * \<mu> S" by (simp add: mult_commute)
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39091
diff changeset
   513
  interpret A: measure_space "restricted_space A" \<mu>
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39091
diff changeset
   514
    using `A \<in> sets M` by (rule restricted_measure_space)
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39091
diff changeset
   515
  show "measure_space (restricted_space A) (\<lambda>S. \<mu> S / \<mu> A)"
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39091
diff changeset
   516
  proof
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39091
diff changeset
   517
    show "\<mu> {} / \<mu> A = 0" by auto
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39091
diff changeset
   518
    show "countably_additive (restricted_space A) (\<lambda>S. \<mu> S / \<mu> A)"
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39091
diff changeset
   519
        unfolding countably_additive_def psuminf_cmult_right *
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39091
diff changeset
   520
        using A.measure_countably_additive by auto
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39091
diff changeset
   521
  qed
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39091
diff changeset
   522
qed
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39091
diff changeset
   523
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39091
diff changeset
   524
lemma finite_prob_spaceI:
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39091
diff changeset
   525
  assumes "finite (space M)" "sets M = Pow(space M)" "\<mu> (space M) = 1" "\<mu> {} = 0"
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39091
diff changeset
   526
    and "\<And>A B. A\<subseteq>space M \<Longrightarrow> B\<subseteq>space M \<Longrightarrow> A \<inter> B = {} \<Longrightarrow> \<mu> (A \<union> B) = \<mu> A + \<mu> B"
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39091
diff changeset
   527
  shows "finite_prob_space M \<mu>"
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39091
diff changeset
   528
  unfolding finite_prob_space_eq
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39091
diff changeset
   529
proof
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39091
diff changeset
   530
  show "finite_measure_space M \<mu>" using assms
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39091
diff changeset
   531
     by (auto intro!: finite_measure_spaceI)
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39091
diff changeset
   532
  show "\<mu> (space M) = 1" by fact
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39091
diff changeset
   533
qed
36624
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   534
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   535
lemma (in finite_prob_space) finite_measure_space:
39097
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   536
  fixes X :: "'a \<Rightarrow> 'x"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   537
  shows "finite_measure_space \<lparr>space = X ` space M, sets = Pow (X ` space M)\<rparr> (distribution X)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   538
    (is "finite_measure_space ?S _")
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39091
diff changeset
   539
proof (rule finite_measure_spaceI, simp_all)
36624
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   540
  show "finite (X ` space M)" using finite_space by simp
39097
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   541
next
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   542
  fix A B :: "'x set" assume "A \<inter> B = {}"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   543
  then show "distribution X (A \<union> B) = distribution X A + distribution X B"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   544
    unfolding distribution_def
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   545
    by (subst measure_additive)
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   546
       (auto intro!: arg_cong[where f=\<mu>] simp: sets_eq_Pow)
36624
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   547
qed
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   548
39097
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   549
lemma (in finite_prob_space) finite_prob_space_of_images:
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   550
  "finite_prob_space \<lparr> space = X ` space M, sets = Pow (X ` space M)\<rparr> (distribution X)"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   551
  by (simp add: finite_prob_space_eq finite_measure_space)
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   552
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   553
lemma (in prob_space) joint_distribution_commute:
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   554
  "joint_distribution X Y x = joint_distribution Y X ((\<lambda>(x,y). (y,x))`x)"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   555
  unfolding distribution_def by (auto intro!: arg_cong[where f=\<mu>])
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   556
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   557
lemma (in finite_prob_space) real_distribution_order':
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   558
  shows "real (distribution X {x}) = 0 \<Longrightarrow> real (joint_distribution X Y {(x, y)}) = 0"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   559
  and "real (distribution Y {y}) = 0 \<Longrightarrow> real (joint_distribution X Y {(x, y)}) = 0"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   560
  using real_of_pinfreal_mono[OF distribution_finite joint_distribution_restriction_fst, of X Y "{(x, y)}"]
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   561
  using real_of_pinfreal_mono[OF distribution_finite joint_distribution_restriction_snd, of X Y "{(x, y)}"]
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   562
  using real_pinfreal_nonneg[of "joint_distribution X Y {(x, y)}"]
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   563
  by auto
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   564
39096
hoelzl
parents: 39092
diff changeset
   565
lemma (in finite_prob_space) finite_product_measure_space:
39097
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   566
  fixes X :: "'a \<Rightarrow> 'x" and Y :: "'a \<Rightarrow> 'y"
39096
hoelzl
parents: 39092
diff changeset
   567
  assumes "finite s1" "finite s2"
hoelzl
parents: 39092
diff changeset
   568
  shows "finite_measure_space \<lparr> space = s1 \<times> s2, sets = Pow (s1 \<times> s2)\<rparr> (joint_distribution X Y)"
hoelzl
parents: 39092
diff changeset
   569
    (is "finite_measure_space ?M ?D")
39097
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   570
proof (rule finite_measure_spaceI, simp_all)
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   571
  show "finite (s1 \<times> s2)"
39096
hoelzl
parents: 39092
diff changeset
   572
    using assms by auto
39097
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   573
  show "joint_distribution X Y (s1\<times>s2) \<noteq> \<omega>"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   574
    using distribution_finite .
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   575
next
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   576
  fix A B :: "('x*'y) set" assume "A \<inter> B = {}"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   577
  then show "joint_distribution X Y (A \<union> B) = joint_distribution X Y A + joint_distribution X Y B"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   578
    unfolding distribution_def
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   579
    by (subst measure_additive)
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   580
       (auto intro!: arg_cong[where f=\<mu>] simp: sets_eq_Pow)
39096
hoelzl
parents: 39092
diff changeset
   581
qed
hoelzl
parents: 39092
diff changeset
   582
39097
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   583
lemma (in finite_prob_space) finite_product_measure_space_of_images:
39096
hoelzl
parents: 39092
diff changeset
   584
  shows "finite_measure_space \<lparr> space = X ` space M \<times> Y ` space M,
hoelzl
parents: 39092
diff changeset
   585
                                sets = Pow (X ` space M \<times> Y ` space M) \<rparr>
hoelzl
parents: 39092
diff changeset
   586
                              (joint_distribution X Y)"
hoelzl
parents: 39092
diff changeset
   587
  using finite_space by (auto intro!: finite_product_measure_space)
hoelzl
parents: 39092
diff changeset
   588
39085
8b7c009da23c added definition of conditional expectation
hoelzl
parents: 39084
diff changeset
   589
section "Conditional Expectation and Probability"
8b7c009da23c added definition of conditional expectation
hoelzl
parents: 39084
diff changeset
   590
8b7c009da23c added definition of conditional expectation
hoelzl
parents: 39084
diff changeset
   591
lemma (in prob_space) conditional_expectation_exists:
39083
e46acc0ea1fe introduced integration on subalgebras
hoelzl
parents: 38656
diff changeset
   592
  fixes X :: "'a \<Rightarrow> pinfreal"
e46acc0ea1fe introduced integration on subalgebras
hoelzl
parents: 38656
diff changeset
   593
  assumes borel: "X \<in> borel_measurable M"
e46acc0ea1fe introduced integration on subalgebras
hoelzl
parents: 38656
diff changeset
   594
  and N_subalgebra: "N \<subseteq> sets M" "sigma_algebra (M\<lparr> sets := N \<rparr>)"
e46acc0ea1fe introduced integration on subalgebras
hoelzl
parents: 38656
diff changeset
   595
  shows "\<exists>Y\<in>borel_measurable (M\<lparr> sets := N \<rparr>). \<forall>C\<in>N.
e46acc0ea1fe introduced integration on subalgebras
hoelzl
parents: 38656
diff changeset
   596
      positive_integral (\<lambda>x. Y x * indicator C x) = positive_integral (\<lambda>x. X x * indicator C x)"
e46acc0ea1fe introduced integration on subalgebras
hoelzl
parents: 38656
diff changeset
   597
proof -
e46acc0ea1fe introduced integration on subalgebras
hoelzl
parents: 38656
diff changeset
   598
  interpret P: prob_space "M\<lparr> sets := N \<rparr>" \<mu>
e46acc0ea1fe introduced integration on subalgebras
hoelzl
parents: 38656
diff changeset
   599
    using prob_space_subalgebra[OF N_subalgebra] .
e46acc0ea1fe introduced integration on subalgebras
hoelzl
parents: 38656
diff changeset
   600
e46acc0ea1fe introduced integration on subalgebras
hoelzl
parents: 38656
diff changeset
   601
  let "?f A" = "\<lambda>x. X x * indicator A x"
e46acc0ea1fe introduced integration on subalgebras
hoelzl
parents: 38656
diff changeset
   602
  let "?Q A" = "positive_integral (?f A)"
e46acc0ea1fe introduced integration on subalgebras
hoelzl
parents: 38656
diff changeset
   603
e46acc0ea1fe introduced integration on subalgebras
hoelzl
parents: 38656
diff changeset
   604
  from measure_space_density[OF borel]
e46acc0ea1fe introduced integration on subalgebras
hoelzl
parents: 38656
diff changeset
   605
  have Q: "measure_space (M\<lparr> sets := N \<rparr>) ?Q"
e46acc0ea1fe introduced integration on subalgebras
hoelzl
parents: 38656
diff changeset
   606
    by (rule measure_space.measure_space_subalgebra[OF _ N_subalgebra])
e46acc0ea1fe introduced integration on subalgebras
hoelzl
parents: 38656
diff changeset
   607
  then interpret Q: measure_space "M\<lparr> sets := N \<rparr>" ?Q .
e46acc0ea1fe introduced integration on subalgebras
hoelzl
parents: 38656
diff changeset
   608
e46acc0ea1fe introduced integration on subalgebras
hoelzl
parents: 38656
diff changeset
   609
  have "P.absolutely_continuous ?Q"
e46acc0ea1fe introduced integration on subalgebras
hoelzl
parents: 38656
diff changeset
   610
    unfolding P.absolutely_continuous_def
e46acc0ea1fe introduced integration on subalgebras
hoelzl
parents: 38656
diff changeset
   611
  proof (safe, simp)
e46acc0ea1fe introduced integration on subalgebras
hoelzl
parents: 38656
diff changeset
   612
    fix A assume "A \<in> N" "\<mu> A = 0"
e46acc0ea1fe introduced integration on subalgebras
hoelzl
parents: 38656
diff changeset
   613
    moreover then have f_borel: "?f A \<in> borel_measurable M"
e46acc0ea1fe introduced integration on subalgebras
hoelzl
parents: 38656
diff changeset
   614
      using borel N_subalgebra by (auto intro: borel_measurable_indicator)
e46acc0ea1fe introduced integration on subalgebras
hoelzl
parents: 38656
diff changeset
   615
    moreover have "{x\<in>space M. ?f A x \<noteq> 0} = (?f A -` {0<..} \<inter> space M) \<inter> A"
e46acc0ea1fe introduced integration on subalgebras
hoelzl
parents: 38656
diff changeset
   616
      by (auto simp: indicator_def)
e46acc0ea1fe introduced integration on subalgebras
hoelzl
parents: 38656
diff changeset
   617
    moreover have "\<mu> \<dots> \<le> \<mu> A"
e46acc0ea1fe introduced integration on subalgebras
hoelzl
parents: 38656
diff changeset
   618
      using `A \<in> N` N_subalgebra f_borel
e46acc0ea1fe introduced integration on subalgebras
hoelzl
parents: 38656
diff changeset
   619
      by (auto intro!: measure_mono Int[of _ A] measurable_sets)
e46acc0ea1fe introduced integration on subalgebras
hoelzl
parents: 38656
diff changeset
   620
    ultimately show "?Q A = 0"
e46acc0ea1fe introduced integration on subalgebras
hoelzl
parents: 38656
diff changeset
   621
      by (simp add: positive_integral_0_iff)
e46acc0ea1fe introduced integration on subalgebras
hoelzl
parents: 38656
diff changeset
   622
  qed
e46acc0ea1fe introduced integration on subalgebras
hoelzl
parents: 38656
diff changeset
   623
  from P.Radon_Nikodym[OF Q this]
e46acc0ea1fe introduced integration on subalgebras
hoelzl
parents: 38656
diff changeset
   624
  obtain Y where Y: "Y \<in> borel_measurable (M\<lparr>sets := N\<rparr>)"
e46acc0ea1fe introduced integration on subalgebras
hoelzl
parents: 38656
diff changeset
   625
    "\<And>A. A \<in> sets (M\<lparr>sets:=N\<rparr>) \<Longrightarrow> ?Q A = P.positive_integral (\<lambda>x. Y x * indicator A x)"
e46acc0ea1fe introduced integration on subalgebras
hoelzl
parents: 38656
diff changeset
   626
    by blast
39084
7a6ecce97661 proved existence of conditional expectation
hoelzl
parents: 39083
diff changeset
   627
  with N_subalgebra show ?thesis
7a6ecce97661 proved existence of conditional expectation
hoelzl
parents: 39083
diff changeset
   628
    by (auto intro!: bexI[OF _ Y(1)])
39083
e46acc0ea1fe introduced integration on subalgebras
hoelzl
parents: 38656
diff changeset
   629
qed
e46acc0ea1fe introduced integration on subalgebras
hoelzl
parents: 38656
diff changeset
   630
39085
8b7c009da23c added definition of conditional expectation
hoelzl
parents: 39084
diff changeset
   631
definition (in prob_space)
8b7c009da23c added definition of conditional expectation
hoelzl
parents: 39084
diff changeset
   632
  "conditional_expectation N X = (SOME Y. Y\<in>borel_measurable (M\<lparr>sets:=N\<rparr>)
8b7c009da23c added definition of conditional expectation
hoelzl
parents: 39084
diff changeset
   633
    \<and> (\<forall>C\<in>N. positive_integral (\<lambda>x. Y x * indicator C x) = positive_integral (\<lambda>x. X x * indicator C x)))"
8b7c009da23c added definition of conditional expectation
hoelzl
parents: 39084
diff changeset
   634
8b7c009da23c added definition of conditional expectation
hoelzl
parents: 39084
diff changeset
   635
abbreviation (in prob_space)
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39091
diff changeset
   636
  "conditional_prob N A \<equiv> conditional_expectation N (indicator A)"
39085
8b7c009da23c added definition of conditional expectation
hoelzl
parents: 39084
diff changeset
   637
8b7c009da23c added definition of conditional expectation
hoelzl
parents: 39084
diff changeset
   638
lemma (in prob_space)
8b7c009da23c added definition of conditional expectation
hoelzl
parents: 39084
diff changeset
   639
  fixes X :: "'a \<Rightarrow> pinfreal"
8b7c009da23c added definition of conditional expectation
hoelzl
parents: 39084
diff changeset
   640
  assumes borel: "X \<in> borel_measurable M"
8b7c009da23c added definition of conditional expectation
hoelzl
parents: 39084
diff changeset
   641
  and N_subalgebra: "N \<subseteq> sets M" "sigma_algebra (M\<lparr> sets := N \<rparr>)"
8b7c009da23c added definition of conditional expectation
hoelzl
parents: 39084
diff changeset
   642
  shows borel_measurable_conditional_expectation:
8b7c009da23c added definition of conditional expectation
hoelzl
parents: 39084
diff changeset
   643
    "conditional_expectation N X \<in> borel_measurable (M\<lparr> sets := N \<rparr>)"
8b7c009da23c added definition of conditional expectation
hoelzl
parents: 39084
diff changeset
   644
  and conditional_expectation: "\<And>C. C \<in> N \<Longrightarrow>
8b7c009da23c added definition of conditional expectation
hoelzl
parents: 39084
diff changeset
   645
      positive_integral (\<lambda>x. conditional_expectation N X x * indicator C x) =
8b7c009da23c added definition of conditional expectation
hoelzl
parents: 39084
diff changeset
   646
      positive_integral (\<lambda>x. X x * indicator C x)"
8b7c009da23c added definition of conditional expectation
hoelzl
parents: 39084
diff changeset
   647
   (is "\<And>C. C \<in> N \<Longrightarrow> ?eq C")
8b7c009da23c added definition of conditional expectation
hoelzl
parents: 39084
diff changeset
   648
proof -
8b7c009da23c added definition of conditional expectation
hoelzl
parents: 39084
diff changeset
   649
  note CE = conditional_expectation_exists[OF assms, unfolded Bex_def]
8b7c009da23c added definition of conditional expectation
hoelzl
parents: 39084
diff changeset
   650
  then show "conditional_expectation N X \<in> borel_measurable (M\<lparr> sets := N \<rparr>)"
8b7c009da23c added definition of conditional expectation
hoelzl
parents: 39084
diff changeset
   651
    unfolding conditional_expectation_def by (rule someI2_ex) blast
8b7c009da23c added definition of conditional expectation
hoelzl
parents: 39084
diff changeset
   652
8b7c009da23c added definition of conditional expectation
hoelzl
parents: 39084
diff changeset
   653
  from CE show "\<And>C. C\<in>N \<Longrightarrow> ?eq C"
8b7c009da23c added definition of conditional expectation
hoelzl
parents: 39084
diff changeset
   654
    unfolding conditional_expectation_def by (rule someI2_ex) blast
8b7c009da23c added definition of conditional expectation
hoelzl
parents: 39084
diff changeset
   655
qed
8b7c009da23c added definition of conditional expectation
hoelzl
parents: 39084
diff changeset
   656
39091
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   657
lemma (in sigma_algebra) factorize_measurable_function:
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   658
  fixes Z :: "'a \<Rightarrow> pinfreal" and Y :: "'a \<Rightarrow> 'c"
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   659
  assumes "sigma_algebra M'" and "Y \<in> measurable M M'" "Z \<in> borel_measurable M"
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   660
  shows "Z \<in> borel_measurable (sigma_algebra.vimage_algebra M' (space M) Y)
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   661
    \<longleftrightarrow> (\<exists>g\<in>borel_measurable M'. \<forall>x\<in>space M. Z x = g (Y x))"
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   662
proof safe
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   663
  interpret M': sigma_algebra M' by fact
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   664
  have Y: "Y \<in> space M \<rightarrow> space M'" using assms unfolding measurable_def by auto
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   665
  from M'.sigma_algebra_vimage[OF this]
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   666
  interpret va: sigma_algebra "M'.vimage_algebra (space M) Y" .
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   667
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   668
  { fix g :: "'c \<Rightarrow> pinfreal" assume "g \<in> borel_measurable M'"
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   669
    with M'.measurable_vimage_algebra[OF Y]
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   670
    have "g \<circ> Y \<in> borel_measurable (M'.vimage_algebra (space M) Y)"
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   671
      by (rule measurable_comp)
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   672
    moreover assume "\<forall>x\<in>space M. Z x = g (Y x)"
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   673
    then have "Z \<in> borel_measurable (M'.vimage_algebra (space M) Y) \<longleftrightarrow>
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   674
       g \<circ> Y \<in> borel_measurable (M'.vimage_algebra (space M) Y)"
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   675
       by (auto intro!: measurable_cong)
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   676
    ultimately show "Z \<in> borel_measurable (M'.vimage_algebra (space M) Y)"
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   677
      by simp }
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   678
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   679
  assume "Z \<in> borel_measurable (M'.vimage_algebra (space M) Y)"
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   680
  from va.borel_measurable_implies_simple_function_sequence[OF this]
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   681
  obtain f where f: "\<And>i. va.simple_function (f i)" and "f \<up> Z" by blast
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   682
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   683
  have "\<forall>i. \<exists>g. M'.simple_function g \<and> (\<forall>x\<in>space M. f i x = g (Y x))"
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   684
  proof
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   685
    fix i
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   686
    from f[of i] have "finite (f i`space M)" and B_ex:
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   687
      "\<forall>z\<in>(f i)`space M. \<exists>B. B \<in> sets M' \<and> (f i) -` {z} \<inter> space M = Y -` B \<inter> space M"
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   688
      unfolding va.simple_function_def by auto
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   689
    from B_ex[THEN bchoice] guess B .. note B = this
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   690
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   691
    let ?g = "\<lambda>x. \<Sum>z\<in>f i`space M. z * indicator (B z) x"
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   692
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   693
    show "\<exists>g. M'.simple_function g \<and> (\<forall>x\<in>space M. f i x = g (Y x))"
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   694
    proof (intro exI[of _ ?g] conjI ballI)
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   695
      show "M'.simple_function ?g" using B by auto
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   696
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   697
      fix x assume "x \<in> space M"
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   698
      then have "\<And>z. z \<in> f i`space M \<Longrightarrow> indicator (B z) (Y x) = (indicator (f i -` {z} \<inter> space M) x::pinfreal)"
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   699
        unfolding indicator_def using B by auto
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   700
      then show "f i x = ?g (Y x)" using `x \<in> space M` f[of i]
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   701
        by (subst va.simple_function_indicator_representation) auto
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   702
    qed
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   703
  qed
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   704
  from choice[OF this] guess g .. note g = this
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   705
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   706
  show "\<exists>g\<in>borel_measurable M'. \<forall>x\<in>space M. Z x = g (Y x)"
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   707
  proof (intro ballI bexI)
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   708
    show "(SUP i. g i) \<in> borel_measurable M'"
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   709
      using g by (auto intro: M'.borel_measurable_simple_function)
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   710
    fix x assume "x \<in> space M"
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   711
    have "Z x = (SUP i. f i) x" using `f \<up> Z` unfolding isoton_def by simp
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   712
    also have "\<dots> = (SUP i. g i) (Y x)" unfolding SUPR_fun_expand
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   713
      using g `x \<in> space M` by simp
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   714
    finally show "Z x = (SUP i. g i) (Y x)" .
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   715
  qed
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   716
qed
39090
a2d38b8b693e Introduced sigma algebra generated by function preimages.
hoelzl
parents: 39089
diff changeset
   717
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   718
end