src/HOL/Probability/Probability_Measure.thy
author hoelzl
Fri, 20 May 2011 16:23:03 +0200
changeset 43763 a61e30bfd0bc
parent 43729 b02349e70d5a
child 43769 e8dbf90a2f3b
permissions -rw-r--r--
add lemma prob_finite_product
hoelzl@43019
     1
(*  Title:      HOL/Probability/Probability_Measure.thy
hoelzl@42922
     2
    Author:     Johannes Hölzl, TU München
hoelzl@42922
     3
    Author:     Armin Heller, TU München
hoelzl@42922
     4
*)
hoelzl@42922
     5
hoelzl@43019
     6
header {*Probability measure*}
hoelzl@42922
     7
hoelzl@43019
     8
theory Probability_Measure
hoelzl@43017
     9
imports Lebesgue_Integration Radon_Nikodym Finite_Product_Measure
hoelzl@35582
    10
begin
hoelzl@35582
    11
hoelzl@42852
    12
lemma real_of_extreal_inverse[simp]:
hoelzl@42852
    13
  fixes X :: extreal
hoelzl@41102
    14
  shows "real (inverse X) = 1 / real X"
hoelzl@41102
    15
  by (cases X) (auto simp: inverse_eq_divide)
hoelzl@41102
    16
hoelzl@42852
    17
lemma real_of_extreal_le_0[simp]: "real (X :: extreal) \<le> 0 \<longleftrightarrow> (X \<le> 0 \<or> X = \<infinity>)"
hoelzl@41102
    18
  by (cases X) auto
hoelzl@41102
    19
hoelzl@42852
    20
lemma abs_real_of_extreal[simp]: "\<bar>real (X :: extreal)\<bar> = real \<bar>X\<bar>"
hoelzl@41102
    21
  by (cases X) auto
hoelzl@41102
    22
hoelzl@42852
    23
lemma zero_less_real_of_extreal: "0 < real X \<longleftrightarrow> (0 < X \<and> X \<noteq> \<infinity>)"
hoelzl@42852
    24
  by (cases X) auto
hoelzl@42852
    25
hoelzl@42852
    26
lemma real_of_extreal_le_1: fixes X :: extreal shows "X \<le> 1 \<Longrightarrow> real X \<le> 1"
hoelzl@42852
    27
  by (cases X) (auto simp: one_extreal_def)
hoelzl@42852
    28
hoelzl@35582
    29
locale prob_space = measure_space +
hoelzl@42553
    30
  assumes measure_space_1: "measure M (space M) = 1"
hoelzl@38902
    31
hoelzl@38902
    32
sublocale prob_space < finite_measure
hoelzl@38902
    33
proof
hoelzl@42852
    34
  from measure_space_1 show "\<mu> (space M) \<noteq> \<infinity>" by simp
hoelzl@38902
    35
qed
hoelzl@38902
    36
hoelzl@41102
    37
abbreviation (in prob_space) "events \<equiv> sets M"
hoelzl@42852
    38
abbreviation (in prob_space) "prob \<equiv> \<mu>'"
hoelzl@41102
    39
abbreviation (in prob_space) "random_variable M' X \<equiv> sigma_algebra M' \<and> X \<in> measurable M M'"
hoelzl@42553
    40
abbreviation (in prob_space) "expectation \<equiv> integral\<^isup>L M"
hoelzl@35582
    41
hoelzl@41102
    42
definition (in prob_space)
hoelzl@35582
    43
  "indep A B \<longleftrightarrow> A \<in> events \<and> B \<in> events \<and> prob (A \<inter> B) = prob A * prob B"
hoelzl@35582
    44
hoelzl@41102
    45
definition (in prob_space)
hoelzl@35582
    46
  "indep_families F G \<longleftrightarrow> (\<forall> A \<in> F. \<forall> B \<in> G. indep A B)"
hoelzl@35582
    47
hoelzl@41102
    48
definition (in prob_space)
hoelzl@42852
    49
  "distribution X A = \<mu>' (X -` A \<inter> space M)"
hoelzl@35582
    50
hoelzl@41102
    51
abbreviation (in prob_space)
hoelzl@36612
    52
  "joint_distribution X Y \<equiv> distribution (\<lambda>x. (X x, Y x))"
hoelzl@36612
    53
hoelzl@42852
    54
declare (in finite_measure) positive_measure'[intro, simp]
hoelzl@42852
    55
hoelzl@39331
    56
lemma (in prob_space) distribution_cong:
hoelzl@39331
    57
  assumes "\<And>x. x \<in> space M \<Longrightarrow> X x = Y x"
hoelzl@39331
    58
  shows "distribution X = distribution Y"
nipkow@39535
    59
  unfolding distribution_def fun_eq_iff
hoelzl@42852
    60
  using assms by (auto intro!: arg_cong[where f="\<mu>'"])
hoelzl@39331
    61
hoelzl@39331
    62
lemma (in prob_space) joint_distribution_cong:
hoelzl@39331
    63
  assumes "\<And>x. x \<in> space M \<Longrightarrow> X x = X' x"
hoelzl@39331
    64
  assumes "\<And>x. x \<in> space M \<Longrightarrow> Y x = Y' x"
hoelzl@39331
    65
  shows "joint_distribution X Y = joint_distribution X' Y'"
nipkow@39535
    66
  unfolding distribution_def fun_eq_iff
hoelzl@42852
    67
  using assms by (auto intro!: arg_cong[where f="\<mu>'"])
hoelzl@39331
    68
hoelzl@41102
    69
lemma (in prob_space) distribution_id[simp]:
hoelzl@42852
    70
  "N \<in> events \<Longrightarrow> distribution (\<lambda>x. x) N = prob N"
hoelzl@42852
    71
  by (auto simp: distribution_def intro!: arg_cong[where f=prob])
hoelzl@41102
    72
hoelzl@41102
    73
lemma (in prob_space) prob_space: "prob (space M) = 1"
hoelzl@42852
    74
  using measure_space_1 unfolding \<mu>'_def by (simp add: one_extreal_def)
hoelzl@35582
    75
hoelzl@42852
    76
lemma (in prob_space) prob_le_1[simp, intro]: "prob A \<le> 1"
hoelzl@42852
    77
  using bounded_measure[of A] by (simp add: prob_space)
hoelzl@42852
    78
hoelzl@42852
    79
lemma (in prob_space) distribution_positive[simp, intro]:
hoelzl@42852
    80
  "0 \<le> distribution X A" unfolding distribution_def by auto
hoelzl@42852
    81
hoelzl@42852
    82
lemma (in prob_space) joint_distribution_remove[simp]:
hoelzl@42852
    83
    "joint_distribution X X {(x, x)} = distribution X {x}"
hoelzl@42852
    84
  unfolding distribution_def by (auto intro!: arg_cong[where f=\<mu>'])
hoelzl@42852
    85
hoelzl@42852
    86
lemma (in prob_space) distribution_1:
hoelzl@42852
    87
  "distribution X A \<le> 1"
hoelzl@42852
    88
  unfolding distribution_def by simp
hoelzl@35582
    89
hoelzl@41102
    90
lemma (in prob_space) prob_compl:
hoelzl@42852
    91
  assumes A: "A \<in> events"
hoelzl@38902
    92
  shows "prob (space M - A) = 1 - prob A"
hoelzl@42852
    93
  using finite_measure_compl[OF A] by (simp add: prob_space)
hoelzl@35582
    94
hoelzl@42852
    95
lemma (in prob_space) indep_space: "s \<in> events \<Longrightarrow> indep (space M) s"
hoelzl@42852
    96
  by (simp add: indep_def prob_space)
hoelzl@35582
    97
hoelzl@41102
    98
lemma (in prob_space) prob_space_increasing: "increasing M prob"
hoelzl@42852
    99
  by (auto intro!: finite_measure_mono simp: increasing_def)
hoelzl@35582
   100
hoelzl@41102
   101
lemma (in prob_space) prob_zero_union:
hoelzl@35582
   102
  assumes "s \<in> events" "t \<in> events" "prob t = 0"
hoelzl@35582
   103
  shows "prob (s \<union> t) = prob s"
hoelzl@38902
   104
using assms
hoelzl@35582
   105
proof -
hoelzl@35582
   106
  have "prob (s \<union> t) \<le> prob s"
hoelzl@42852
   107
    using finite_measure_subadditive[of s t] assms by auto
hoelzl@35582
   108
  moreover have "prob (s \<union> t) \<ge> prob s"
hoelzl@42852
   109
    using assms by (blast intro: finite_measure_mono)
hoelzl@35582
   110
  ultimately show ?thesis by simp
hoelzl@35582
   111
qed
hoelzl@35582
   112
hoelzl@41102
   113
lemma (in prob_space) prob_eq_compl:
hoelzl@35582
   114
  assumes "s \<in> events" "t \<in> events"
hoelzl@35582
   115
  assumes "prob (space M - s) = prob (space M - t)"
hoelzl@35582
   116
  shows "prob s = prob t"
hoelzl@38902
   117
  using assms prob_compl by auto
hoelzl@35582
   118
hoelzl@41102
   119
lemma (in prob_space) prob_one_inter:
hoelzl@35582
   120
  assumes events:"s \<in> events" "t \<in> events"
hoelzl@35582
   121
  assumes "prob t = 1"
hoelzl@35582
   122
  shows "prob (s \<inter> t) = prob s"
hoelzl@35582
   123
proof -
hoelzl@38902
   124
  have "prob ((space M - s) \<union> (space M - t)) = prob (space M - s)"
hoelzl@38902
   125
    using events assms  prob_compl[of "t"] by (auto intro!: prob_zero_union)
hoelzl@38902
   126
  also have "(space M - s) \<union> (space M - t) = space M - (s \<inter> t)"
hoelzl@38902
   127
    by blast
hoelzl@38902
   128
  finally show "prob (s \<inter> t) = prob s"
hoelzl@38902
   129
    using events by (auto intro!: prob_eq_compl[of "s \<inter> t" s])
hoelzl@35582
   130
qed
hoelzl@35582
   131
hoelzl@41102
   132
lemma (in prob_space) prob_eq_bigunion_image:
hoelzl@35582
   133
  assumes "range f \<subseteq> events" "range g \<subseteq> events"
hoelzl@35582
   134
  assumes "disjoint_family f" "disjoint_family g"
hoelzl@35582
   135
  assumes "\<And> n :: nat. prob (f n) = prob (g n)"
hoelzl@35582
   136
  shows "(prob (\<Union> i. f i) = prob (\<Union> i. g i))"
hoelzl@35582
   137
using assms
hoelzl@35582
   138
proof -
hoelzl@38902
   139
  have a: "(\<lambda> i. prob (f i)) sums (prob (\<Union> i. f i))"
hoelzl@42852
   140
    by (rule finite_measure_UNION[OF assms(1,3)])
hoelzl@38902
   141
  have b: "(\<lambda> i. prob (g i)) sums (prob (\<Union> i. g i))"
hoelzl@42852
   142
    by (rule finite_measure_UNION[OF assms(2,4)])
hoelzl@38902
   143
  show ?thesis using sums_unique[OF b] sums_unique[OF a] assms by simp
hoelzl@35582
   144
qed
hoelzl@35582
   145
hoelzl@41102
   146
lemma (in prob_space) prob_countably_zero:
hoelzl@35582
   147
  assumes "range c \<subseteq> events"
hoelzl@35582
   148
  assumes "\<And> i. prob (c i) = 0"
hoelzl@38902
   149
  shows "prob (\<Union> i :: nat. c i) = 0"
hoelzl@38902
   150
proof (rule antisym)
hoelzl@38902
   151
  show "prob (\<Union> i :: nat. c i) \<le> 0"
hoelzl@42852
   152
    using finite_measure_countably_subadditive[OF assms(1)]
hoelzl@38902
   153
    by (simp add: assms(2) suminf_zero summable_zero)
hoelzl@42852
   154
qed simp
hoelzl@35582
   155
hoelzl@41102
   156
lemma (in prob_space) indep_sym:
hoelzl@35582
   157
   "indep a b \<Longrightarrow> indep b a"
hoelzl@35582
   158
unfolding indep_def using Int_commute[of a b] by auto
hoelzl@35582
   159
hoelzl@41102
   160
lemma (in prob_space) indep_refl:
hoelzl@35582
   161
  assumes "a \<in> events"
hoelzl@35582
   162
  shows "indep a a = (prob a = 0) \<or> (prob a = 1)"
hoelzl@35582
   163
using assms unfolding indep_def by auto
hoelzl@35582
   164
hoelzl@41102
   165
lemma (in prob_space) prob_equiprobable_finite_unions:
hoelzl@38902
   166
  assumes "s \<in> events"
hoelzl@38902
   167
  assumes s_finite: "finite s" "\<And>x. x \<in> s \<Longrightarrow> {x} \<in> events"
hoelzl@35582
   168
  assumes "\<And> x y. \<lbrakk>x \<in> s; y \<in> s\<rbrakk> \<Longrightarrow> (prob {x} = prob {y})"
hoelzl@38902
   169
  shows "prob s = real (card s) * prob {SOME x. x \<in> s}"
hoelzl@35582
   170
proof (cases "s = {}")
hoelzl@38902
   171
  case False hence "\<exists> x. x \<in> s" by blast
hoelzl@35582
   172
  from someI_ex[OF this] assms
hoelzl@35582
   173
  have prob_some: "\<And> x. x \<in> s \<Longrightarrow> prob {x} = prob {SOME y. y \<in> s}" by blast
hoelzl@35582
   174
  have "prob s = (\<Sum> x \<in> s. prob {x})"
hoelzl@42852
   175
    using finite_measure_finite_singleton[OF s_finite] by simp
hoelzl@35582
   176
  also have "\<dots> = (\<Sum> x \<in> s. prob {SOME y. y \<in> s})" using prob_some by auto
hoelzl@38902
   177
  also have "\<dots> = real (card s) * prob {(SOME x. x \<in> s)}"
hoelzl@38902
   178
    using setsum_constant assms by (simp add: real_eq_of_nat)
hoelzl@35582
   179
  finally show ?thesis by simp
hoelzl@38902
   180
qed simp
hoelzl@35582
   181
hoelzl@41102
   182
lemma (in prob_space) prob_real_sum_image_fn:
hoelzl@35582
   183
  assumes "e \<in> events"
hoelzl@35582
   184
  assumes "\<And> x. x \<in> s \<Longrightarrow> e \<inter> f x \<in> events"
hoelzl@35582
   185
  assumes "finite s"
hoelzl@38902
   186
  assumes disjoint: "\<And> x y. \<lbrakk>x \<in> s ; y \<in> s ; x \<noteq> y\<rbrakk> \<Longrightarrow> f x \<inter> f y = {}"
hoelzl@38902
   187
  assumes upper: "space M \<subseteq> (\<Union> i \<in> s. f i)"
hoelzl@35582
   188
  shows "prob e = (\<Sum> x \<in> s. prob (e \<inter> f x))"
hoelzl@35582
   189
proof -
hoelzl@38902
   190
  have e: "e = (\<Union> i \<in> s. e \<inter> f i)"
hoelzl@38902
   191
    using `e \<in> events` sets_into_space upper by blast
hoelzl@38902
   192
  hence "prob e = prob (\<Union> i \<in> s. e \<inter> f i)" by simp
hoelzl@38902
   193
  also have "\<dots> = (\<Sum> x \<in> s. prob (e \<inter> f x))"
hoelzl@42852
   194
  proof (rule finite_measure_finite_Union)
hoelzl@38902
   195
    show "finite s" by fact
hoelzl@38902
   196
    show "\<And>i. i \<in> s \<Longrightarrow> e \<inter> f i \<in> events" by fact
hoelzl@38902
   197
    show "disjoint_family_on (\<lambda>i. e \<inter> f i) s"
hoelzl@38902
   198
      using disjoint by (auto simp: disjoint_family_on_def)
hoelzl@38902
   199
  qed
hoelzl@38902
   200
  finally show ?thesis .
hoelzl@35582
   201
qed
hoelzl@35582
   202
hoelzl@43069
   203
lemma (in prob_space) prob_space_vimage:
hoelzl@43069
   204
  assumes S: "sigma_algebra S"
hoelzl@43069
   205
  assumes T: "T \<in> measure_preserving M S"
hoelzl@43069
   206
  shows "prob_space S"
hoelzl@43069
   207
proof -
hoelzl@43069
   208
  interpret S: measure_space S
hoelzl@43069
   209
    using S and T by (rule measure_space_vimage)
hoelzl@43069
   210
  show ?thesis
hoelzl@43069
   211
  proof
hoelzl@43069
   212
    from T[THEN measure_preservingD2]
hoelzl@43069
   213
    have "T -` space S \<inter> space M = space M"
hoelzl@43069
   214
      by (auto simp: measurable_def)
hoelzl@43069
   215
    with T[THEN measure_preservingD, of "space S", symmetric]
hoelzl@43069
   216
    show  "measure S (space S) = 1"
hoelzl@43069
   217
      using measure_space_1 by simp
hoelzl@43069
   218
  qed
hoelzl@43069
   219
qed
hoelzl@43069
   220
hoelzl@41102
   221
lemma (in prob_space) distribution_prob_space:
hoelzl@43069
   222
  assumes X: "random_variable S X"
hoelzl@43069
   223
  shows "prob_space (S\<lparr>measure := extreal \<circ> distribution X\<rparr>)" (is "prob_space ?S")
hoelzl@43069
   224
proof (rule prob_space_vimage)
hoelzl@43069
   225
  show "X \<in> measure_preserving M ?S"
hoelzl@43069
   226
    using X
hoelzl@43069
   227
    unfolding measure_preserving_def distribution_def_raw
hoelzl@43069
   228
    by (auto simp: finite_measure_eq measurable_sets)
hoelzl@43069
   229
  show "sigma_algebra ?S" using X by simp
hoelzl@35582
   230
qed
hoelzl@35582
   231
hoelzl@41102
   232
lemma (in prob_space) AE_distribution:
hoelzl@42852
   233
  assumes X: "random_variable MX X" and "AE x in MX\<lparr>measure := extreal \<circ> distribution X\<rparr>. Q x"
hoelzl@41102
   234
  shows "AE x. Q (X x)"
hoelzl@41102
   235
proof -
hoelzl@42852
   236
  interpret X: prob_space "MX\<lparr>measure := extreal \<circ> distribution X\<rparr>" using X by (rule distribution_prob_space)
hoelzl@41102
   237
  obtain N where N: "N \<in> sets MX" "distribution X N = 0" "{x\<in>space MX. \<not> Q x} \<subseteq> N"
hoelzl@41102
   238
    using assms unfolding X.almost_everywhere_def by auto
hoelzl@42852
   239
  from X[unfolded measurable_def] N show "AE x. Q (X x)"
hoelzl@42852
   240
    by (intro AE_I'[where N="X -` N \<inter> space M"])
hoelzl@42852
   241
       (auto simp: finite_measure_eq distribution_def measurable_sets)
hoelzl@41102
   242
qed
hoelzl@41102
   243
hoelzl@42852
   244
lemma (in prob_space) distribution_eq_integral:
hoelzl@42852
   245
  "random_variable S X \<Longrightarrow> A \<in> sets S \<Longrightarrow> distribution X A = expectation (indicator (X -` A \<inter> space M))"
hoelzl@42852
   246
  using finite_measure_eq[of "X -` A \<inter> space M"]
hoelzl@42852
   247
  by (auto simp: measurable_sets distribution_def)
hoelzl@35582
   248
hoelzl@42852
   249
lemma (in prob_space) distribution_eq_translated_integral:
hoelzl@42852
   250
  assumes "random_variable S X" "A \<in> sets S"
hoelzl@42852
   251
  shows "distribution X A = integral\<^isup>P (S\<lparr>measure := extreal \<circ> distribution X\<rparr>) (indicator A)"
hoelzl@35582
   252
proof -
hoelzl@42852
   253
  interpret S: prob_space "S\<lparr>measure := extreal \<circ> distribution X\<rparr>"
hoelzl@42553
   254
    using assms(1) by (rule distribution_prob_space)
hoelzl@35582
   255
  show ?thesis
hoelzl@42852
   256
    using S.positive_integral_indicator(1)[of A] assms by simp
hoelzl@35582
   257
qed
hoelzl@35582
   258
hoelzl@41102
   259
lemma (in prob_space) finite_expectation1:
hoelzl@41102
   260
  assumes f: "finite (X`space M)" and rv: "random_variable borel X"
hoelzl@42852
   261
  shows "expectation X = (\<Sum>r \<in> X ` space M. r * prob (X -` {r} \<inter> space M))" (is "_ = ?r")
hoelzl@42852
   262
proof (subst integral_on_finite)
hoelzl@42852
   263
  show "X \<in> borel_measurable M" "finite (X`space M)" using assms by auto
hoelzl@42852
   264
  show "(\<Sum> r \<in> X ` space M. r * real (\<mu> (X -` {r} \<inter> space M))) = ?r"
hoelzl@42852
   265
    "\<And>x. \<mu> (X -` {x} \<inter> space M) \<noteq> \<infinity>"
hoelzl@42852
   266
    using finite_measure_eq[OF borel_measurable_vimage, of X] rv by auto
hoelzl@38902
   267
qed
hoelzl@35582
   268
hoelzl@41102
   269
lemma (in prob_space) finite_expectation:
hoelzl@42553
   270
  assumes "finite (X`space M)" "random_variable borel X"
hoelzl@42852
   271
  shows "expectation X = (\<Sum> r \<in> X ` (space M). r * distribution X {r})"
hoelzl@38902
   272
  using assms unfolding distribution_def using finite_expectation1 by auto
hoelzl@38902
   273
hoelzl@41102
   274
lemma (in prob_space) prob_x_eq_1_imp_prob_y_eq_0:
hoelzl@35582
   275
  assumes "{x} \<in> events"
hoelzl@38902
   276
  assumes "prob {x} = 1"
hoelzl@35582
   277
  assumes "{y} \<in> events"
hoelzl@35582
   278
  assumes "y \<noteq> x"
hoelzl@35582
   279
  shows "prob {y} = 0"
hoelzl@35582
   280
  using prob_one_inter[of "{y}" "{x}"] assms by auto
hoelzl@35582
   281
hoelzl@41102
   282
lemma (in prob_space) distribution_empty[simp]: "distribution X {} = 0"
hoelzl@38902
   283
  unfolding distribution_def by simp
hoelzl@38902
   284
hoelzl@41102
   285
lemma (in prob_space) distribution_space[simp]: "distribution X (X ` space M) = 1"
hoelzl@38902
   286
proof -
hoelzl@38902
   287
  have "X -` X ` space M \<inter> space M = space M" by auto
hoelzl@42852
   288
  thus ?thesis unfolding distribution_def by (simp add: prob_space)
hoelzl@38902
   289
qed
hoelzl@38902
   290
hoelzl@41102
   291
lemma (in prob_space) distribution_one:
hoelzl@41102
   292
  assumes "random_variable M' X" and "A \<in> sets M'"
hoelzl@38902
   293
  shows "distribution X A \<le> 1"
hoelzl@38902
   294
proof -
hoelzl@42852
   295
  have "distribution X A \<le> \<mu>' (space M)" unfolding distribution_def
hoelzl@42852
   296
    using assms[unfolded measurable_def] by (auto intro!: finite_measure_mono)
hoelzl@42852
   297
  thus ?thesis by (simp add: prob_space)
hoelzl@38902
   298
qed
hoelzl@38902
   299
hoelzl@41102
   300
lemma (in prob_space) distribution_x_eq_1_imp_distribution_y_eq_0:
hoelzl@35582
   301
  assumes X: "random_variable \<lparr>space = X ` (space M), sets = Pow (X ` (space M))\<rparr> X"
hoelzl@38902
   302
    (is "random_variable ?S X")
hoelzl@38902
   303
  assumes "distribution X {x} = 1"
hoelzl@35582
   304
  assumes "y \<noteq> x"
hoelzl@35582
   305
  shows "distribution X {y} = 0"
hoelzl@42553
   306
proof cases
hoelzl@42553
   307
  { fix x have "X -` {x} \<inter> space M \<in> sets M"
hoelzl@42553
   308
    proof cases
hoelzl@42553
   309
      assume "x \<in> X`space M" with X show ?thesis
hoelzl@42553
   310
        by (auto simp: measurable_def image_iff)
hoelzl@42553
   311
    next
hoelzl@42553
   312
      assume "x \<notin> X`space M" then have "X -` {x} \<inter> space M = {}" by auto
hoelzl@42553
   313
      then show ?thesis by auto
hoelzl@42553
   314
    qed } note single = this
hoelzl@42553
   315
  have "X -` {x} \<inter> space M - X -` {y} \<inter> space M = X -` {x} \<inter> space M"
hoelzl@42553
   316
    "X -` {y} \<inter> space M \<inter> (X -` {x} \<inter> space M) = {}"
hoelzl@42553
   317
    using `y \<noteq> x` by auto
hoelzl@42852
   318
  with finite_measure_inter_full_set[OF single single, of x y] assms(2)
hoelzl@42852
   319
  show ?thesis by (auto simp: distribution_def prob_space)
hoelzl@42553
   320
next
hoelzl@42553
   321
  assume "{y} \<notin> sets ?S"
hoelzl@42553
   322
  then have "X -` {y} \<inter> space M = {}" by auto
hoelzl@42553
   323
  thus "distribution X {y} = 0" unfolding distribution_def by auto
hoelzl@35582
   324
qed
hoelzl@35582
   325
hoelzl@41102
   326
lemma (in prob_space) joint_distribution_Times_le_fst:
hoelzl@41102
   327
  assumes X: "random_variable MX X" and Y: "random_variable MY Y"
hoelzl@41102
   328
    and A: "A \<in> sets MX" and B: "B \<in> sets MY"
hoelzl@41102
   329
  shows "joint_distribution X Y (A \<times> B) \<le> distribution X A"
hoelzl@41102
   330
  unfolding distribution_def
hoelzl@42852
   331
proof (intro finite_measure_mono)
hoelzl@41102
   332
  show "(\<lambda>x. (X x, Y x)) -` (A \<times> B) \<inter> space M \<subseteq> X -` A \<inter> space M" by force
hoelzl@41102
   333
  show "X -` A \<inter> space M \<in> events"
hoelzl@41102
   334
    using X A unfolding measurable_def by simp
hoelzl@41102
   335
  have *: "(\<lambda>x. (X x, Y x)) -` (A \<times> B) \<inter> space M =
hoelzl@41102
   336
    (X -` A \<inter> space M) \<inter> (Y -` B \<inter> space M)" by auto
hoelzl@41102
   337
qed
hoelzl@41102
   338
hoelzl@41102
   339
lemma (in prob_space) joint_distribution_commute:
hoelzl@41102
   340
  "joint_distribution X Y x = joint_distribution Y X ((\<lambda>(x,y). (y,x))`x)"
hoelzl@42852
   341
  unfolding distribution_def by (auto intro!: arg_cong[where f=\<mu>'])
hoelzl@41102
   342
hoelzl@41102
   343
lemma (in prob_space) joint_distribution_Times_le_snd:
hoelzl@41102
   344
  assumes X: "random_variable MX X" and Y: "random_variable MY Y"
hoelzl@41102
   345
    and A: "A \<in> sets MX" and B: "B \<in> sets MY"
hoelzl@41102
   346
  shows "joint_distribution X Y (A \<times> B) \<le> distribution Y B"
hoelzl@41102
   347
  using assms
hoelzl@41102
   348
  by (subst joint_distribution_commute)
hoelzl@41102
   349
     (simp add: swap_product joint_distribution_Times_le_fst)
hoelzl@41102
   350
hoelzl@41102
   351
lemma (in prob_space) random_variable_pairI:
hoelzl@41102
   352
  assumes "random_variable MX X"
hoelzl@41102
   353
  assumes "random_variable MY Y"
hoelzl@42553
   354
  shows "random_variable (MX \<Otimes>\<^isub>M MY) (\<lambda>x. (X x, Y x))"
hoelzl@41102
   355
proof
hoelzl@41102
   356
  interpret MX: sigma_algebra MX using assms by simp
hoelzl@41102
   357
  interpret MY: sigma_algebra MY using assms by simp
hoelzl@41102
   358
  interpret P: pair_sigma_algebra MX MY by default
hoelzl@42553
   359
  show "sigma_algebra (MX \<Otimes>\<^isub>M MY)" by default
hoelzl@41102
   360
  have sa: "sigma_algebra M" by default
hoelzl@42553
   361
  show "(\<lambda>x. (X x, Y x)) \<in> measurable M (MX \<Otimes>\<^isub>M MY)"
hoelzl@41343
   362
    unfolding P.measurable_pair_iff[OF sa] using assms by (simp add: comp_def)
hoelzl@41102
   363
qed
hoelzl@41102
   364
hoelzl@41102
   365
lemma (in prob_space) joint_distribution_commute_singleton:
hoelzl@41102
   366
  "joint_distribution X Y {(x, y)} = joint_distribution Y X {(y, x)}"
hoelzl@42852
   367
  unfolding distribution_def by (auto intro!: arg_cong[where f=\<mu>'])
hoelzl@41102
   368
hoelzl@41102
   369
lemma (in prob_space) joint_distribution_assoc_singleton:
hoelzl@41102
   370
  "joint_distribution X (\<lambda>x. (Y x, Z x)) {(x, y, z)} =
hoelzl@41102
   371
   joint_distribution (\<lambda>x. (X x, Y x)) Z {((x, y), z)}"
hoelzl@42852
   372
  unfolding distribution_def by (auto intro!: arg_cong[where f=\<mu>'])
hoelzl@41102
   373
hoelzl@42553
   374
locale pair_prob_space = M1: prob_space M1 + M2: prob_space M2 for M1 M2
hoelzl@41102
   375
hoelzl@42553
   376
sublocale pair_prob_space \<subseteq> pair_sigma_finite M1 M2 by default
hoelzl@41102
   377
hoelzl@42553
   378
sublocale pair_prob_space \<subseteq> P: prob_space P
hoelzl@42553
   379
by default (simp add: pair_measure_times M1.measure_space_1 M2.measure_space_1 space_pair_measure)
hoelzl@41102
   380
hoelzl@41102
   381
lemma countably_additiveI[case_names countably]:
hoelzl@41102
   382
  assumes "\<And>A. \<lbrakk> range A \<subseteq> sets M ; disjoint_family A ; (\<Union>i. A i) \<in> sets M\<rbrakk> \<Longrightarrow>
hoelzl@42852
   383
    (\<Sum>n. \<mu> (A n)) = \<mu> (\<Union>i. A i)"
hoelzl@41102
   384
  shows "countably_additive M \<mu>"
hoelzl@41102
   385
  using assms unfolding countably_additive_def by auto
hoelzl@41102
   386
hoelzl@41102
   387
lemma (in prob_space) joint_distribution_prob_space:
hoelzl@41102
   388
  assumes "random_variable MX X" "random_variable MY Y"
hoelzl@42852
   389
  shows "prob_space ((MX \<Otimes>\<^isub>M MY) \<lparr> measure := extreal \<circ> joint_distribution X Y\<rparr>)"
hoelzl@42553
   390
  using random_variable_pairI[OF assms] by (rule distribution_prob_space)
hoelzl@41102
   391
hoelzl@41102
   392
section "Probability spaces on finite sets"
hoelzl@35582
   393
hoelzl@35974
   394
locale finite_prob_space = prob_space + finite_measure_space
hoelzl@35974
   395
hoelzl@41102
   396
abbreviation (in prob_space) "finite_random_variable M' X \<equiv> finite_sigma_algebra M' \<and> X \<in> measurable M M'"
hoelzl@41102
   397
hoelzl@41102
   398
lemma (in prob_space) finite_random_variableD:
hoelzl@41102
   399
  assumes "finite_random_variable M' X" shows "random_variable M' X"
hoelzl@41102
   400
proof -
hoelzl@41102
   401
  interpret M': finite_sigma_algebra M' using assms by simp
hoelzl@41102
   402
  then show "random_variable M' X" using assms by simp default
hoelzl@41102
   403
qed
hoelzl@41102
   404
hoelzl@41102
   405
lemma (in prob_space) distribution_finite_prob_space:
hoelzl@41102
   406
  assumes "finite_random_variable MX X"
hoelzl@42852
   407
  shows "finite_prob_space (MX\<lparr>measure := extreal \<circ> distribution X\<rparr>)"
hoelzl@41102
   408
proof -
hoelzl@42852
   409
  interpret X: prob_space "MX\<lparr>measure := extreal \<circ> distribution X\<rparr>"
hoelzl@41102
   410
    using assms[THEN finite_random_variableD] by (rule distribution_prob_space)
hoelzl@41102
   411
  interpret MX: finite_sigma_algebra MX
hoelzl@42553
   412
    using assms by auto
hoelzl@42852
   413
  show ?thesis by default (simp_all add: MX.finite_space)
hoelzl@41102
   414
qed
hoelzl@41102
   415
hoelzl@41102
   416
lemma (in prob_space) simple_function_imp_finite_random_variable[simp, intro]:
hoelzl@42553
   417
  assumes "simple_function M X"
hoelzl@42553
   418
  shows "finite_random_variable \<lparr> space = X`space M, sets = Pow (X`space M), \<dots> = x \<rparr> X"
hoelzl@42553
   419
    (is "finite_random_variable ?X _")
hoelzl@41102
   420
proof (intro conjI)
hoelzl@41102
   421
  have [simp]: "finite (X ` space M)" using assms unfolding simple_function_def by simp
hoelzl@42553
   422
  interpret X: sigma_algebra ?X by (rule sigma_algebra_Pow)
hoelzl@42553
   423
  show "finite_sigma_algebra ?X"
hoelzl@41102
   424
    by default auto
hoelzl@42553
   425
  show "X \<in> measurable M ?X"
hoelzl@41102
   426
  proof (unfold measurable_def, clarsimp)
hoelzl@41102
   427
    fix A assume A: "A \<subseteq> X`space M"
hoelzl@41102
   428
    then have "finite A" by (rule finite_subset) simp
hoelzl@41102
   429
    then have "X -` (\<Union>a\<in>A. {a}) \<inter> space M \<in> events"
hoelzl@41102
   430
      unfolding vimage_UN UN_extend_simps
hoelzl@41102
   431
      apply (rule finite_UN)
hoelzl@41102
   432
      using A assms unfolding simple_function_def by auto
hoelzl@41102
   433
    then show "X -` A \<inter> space M \<in> events" by simp
hoelzl@41102
   434
  qed
hoelzl@41102
   435
qed
hoelzl@41102
   436
hoelzl@41102
   437
lemma (in prob_space) simple_function_imp_random_variable[simp, intro]:
hoelzl@42553
   438
  assumes "simple_function M X"
hoelzl@42553
   439
  shows "random_variable \<lparr> space = X`space M, sets = Pow (X`space M), \<dots> = ext \<rparr> X"
hoelzl@42553
   440
  using simple_function_imp_finite_random_variable[OF assms, of ext]
hoelzl@41102
   441
  by (auto dest!: finite_random_variableD)
hoelzl@41102
   442
hoelzl@41102
   443
lemma (in prob_space) sum_over_space_real_distribution:
hoelzl@42852
   444
  "simple_function M X \<Longrightarrow> (\<Sum>x\<in>X`space M. distribution X {x}) = 1"
hoelzl@41102
   445
  unfolding distribution_def prob_space[symmetric]
hoelzl@42852
   446
  by (subst finite_measure_finite_Union[symmetric])
hoelzl@41102
   447
     (auto simp add: disjoint_family_on_def simple_function_def
hoelzl@41102
   448
           intro!: arg_cong[where f=prob])
hoelzl@41102
   449
hoelzl@41102
   450
lemma (in prob_space) finite_random_variable_pairI:
hoelzl@41102
   451
  assumes "finite_random_variable MX X"
hoelzl@41102
   452
  assumes "finite_random_variable MY Y"
hoelzl@42553
   453
  shows "finite_random_variable (MX \<Otimes>\<^isub>M MY) (\<lambda>x. (X x, Y x))"
hoelzl@41102
   454
proof
hoelzl@41102
   455
  interpret MX: finite_sigma_algebra MX using assms by simp
hoelzl@41102
   456
  interpret MY: finite_sigma_algebra MY using assms by simp
hoelzl@41102
   457
  interpret P: pair_finite_sigma_algebra MX MY by default
hoelzl@42553
   458
  show "finite_sigma_algebra (MX \<Otimes>\<^isub>M MY)" by default
hoelzl@41102
   459
  have sa: "sigma_algebra M" by default
hoelzl@42553
   460
  show "(\<lambda>x. (X x, Y x)) \<in> measurable M (MX \<Otimes>\<^isub>M MY)"
hoelzl@41343
   461
    unfolding P.measurable_pair_iff[OF sa] using assms by (simp add: comp_def)
hoelzl@41102
   462
qed
hoelzl@41102
   463
hoelzl@41102
   464
lemma (in prob_space) finite_random_variable_imp_sets:
hoelzl@41102
   465
  "finite_random_variable MX X \<Longrightarrow> x \<in> space MX \<Longrightarrow> {x} \<in> sets MX"
hoelzl@41102
   466
  unfolding finite_sigma_algebra_def finite_sigma_algebra_axioms_def by simp
hoelzl@41102
   467
hoelzl@42852
   468
lemma (in prob_space) finite_random_variable_measurable:
hoelzl@41102
   469
  assumes X: "finite_random_variable MX X" shows "X -` A \<inter> space M \<in> events"
hoelzl@41102
   470
proof -
hoelzl@41102
   471
  interpret X: finite_sigma_algebra MX using X by simp
hoelzl@41102
   472
  from X have vimage: "\<And>A. A \<subseteq> space MX \<Longrightarrow> X -` A \<inter> space M \<in> events" and
hoelzl@41102
   473
    "X \<in> space M \<rightarrow> space MX"
hoelzl@41102
   474
    by (auto simp: measurable_def)
hoelzl@41102
   475
  then have *: "X -` A \<inter> space M = X -` (A \<inter> space MX) \<inter> space M"
hoelzl@41102
   476
    by auto
hoelzl@41102
   477
  show "X -` A \<inter> space M \<in> events"
hoelzl@41102
   478
    unfolding * by (intro vimage) auto
hoelzl@41102
   479
qed
hoelzl@41102
   480
hoelzl@41102
   481
lemma (in prob_space) joint_distribution_finite_Times_le_fst:
hoelzl@41102
   482
  assumes X: "finite_random_variable MX X" and Y: "finite_random_variable MY Y"
hoelzl@41102
   483
  shows "joint_distribution X Y (A \<times> B) \<le> distribution X A"
hoelzl@41102
   484
  unfolding distribution_def
hoelzl@42852
   485
proof (intro finite_measure_mono)
hoelzl@41102
   486
  show "(\<lambda>x. (X x, Y x)) -` (A \<times> B) \<inter> space M \<subseteq> X -` A \<inter> space M" by force
hoelzl@41102
   487
  show "X -` A \<inter> space M \<in> events"
hoelzl@42852
   488
    using finite_random_variable_measurable[OF X] .
hoelzl@41102
   489
  have *: "(\<lambda>x. (X x, Y x)) -` (A \<times> B) \<inter> space M =
hoelzl@41102
   490
    (X -` A \<inter> space M) \<inter> (Y -` B \<inter> space M)" by auto
hoelzl@41102
   491
qed
hoelzl@41102
   492
hoelzl@41102
   493
lemma (in prob_space) joint_distribution_finite_Times_le_snd:
hoelzl@41102
   494
  assumes X: "finite_random_variable MX X" and Y: "finite_random_variable MY Y"
hoelzl@41102
   495
  shows "joint_distribution X Y (A \<times> B) \<le> distribution Y B"
hoelzl@41102
   496
  using assms
hoelzl@41102
   497
  by (subst joint_distribution_commute)
hoelzl@41102
   498
     (simp add: swap_product joint_distribution_finite_Times_le_fst)
hoelzl@41102
   499
hoelzl@41102
   500
lemma (in prob_space) finite_distribution_order:
hoelzl@42852
   501
  fixes MX :: "('c, 'd) measure_space_scheme" and MY :: "('e, 'f) measure_space_scheme"
hoelzl@41102
   502
  assumes "finite_random_variable MX X" "finite_random_variable MY Y"
hoelzl@41102
   503
  shows "r \<le> joint_distribution X Y {(x, y)} \<Longrightarrow> r \<le> distribution X {x}"
hoelzl@41102
   504
    and "r \<le> joint_distribution X Y {(x, y)} \<Longrightarrow> r \<le> distribution Y {y}"
hoelzl@41102
   505
    and "r < joint_distribution X Y {(x, y)} \<Longrightarrow> r < distribution X {x}"
hoelzl@41102
   506
    and "r < joint_distribution X Y {(x, y)} \<Longrightarrow> r < distribution Y {y}"
hoelzl@41102
   507
    and "distribution X {x} = 0 \<Longrightarrow> joint_distribution X Y {(x, y)} = 0"
hoelzl@41102
   508
    and "distribution Y {y} = 0 \<Longrightarrow> joint_distribution X Y {(x, y)} = 0"
hoelzl@41102
   509
  using joint_distribution_finite_Times_le_snd[OF assms, of "{x}" "{y}"]
hoelzl@41102
   510
  using joint_distribution_finite_Times_le_fst[OF assms, of "{x}" "{y}"]
hoelzl@42852
   511
  by (auto intro: antisym)
hoelzl@41102
   512
hoelzl@41102
   513
lemma (in prob_space) setsum_joint_distribution:
hoelzl@41102
   514
  assumes X: "finite_random_variable MX X"
hoelzl@41102
   515
  assumes Y: "random_variable MY Y" "B \<in> sets MY"
hoelzl@41102
   516
  shows "(\<Sum>a\<in>space MX. joint_distribution X Y ({a} \<times> B)) = distribution Y B"
hoelzl@41102
   517
  unfolding distribution_def
hoelzl@42852
   518
proof (subst finite_measure_finite_Union[symmetric])
hoelzl@41102
   519
  interpret MX: finite_sigma_algebra MX using X by auto
hoelzl@41102
   520
  show "finite (space MX)" using MX.finite_space .
hoelzl@41102
   521
  let "?d i" = "(\<lambda>x. (X x, Y x)) -` ({i} \<times> B) \<inter> space M"
hoelzl@41102
   522
  { fix i assume "i \<in> space MX"
hoelzl@41102
   523
    moreover have "?d i = (X -` {i} \<inter> space M) \<inter> (Y -` B \<inter> space M)" by auto
hoelzl@41102
   524
    ultimately show "?d i \<in> events"
hoelzl@41102
   525
      using measurable_sets[of X M MX] measurable_sets[of Y M MY B] X Y
hoelzl@41102
   526
      using MX.sets_eq_Pow by auto }
hoelzl@41102
   527
  show "disjoint_family_on ?d (space MX)" by (auto simp: disjoint_family_on_def)
hoelzl@42852
   528
  show "\<mu>' (\<Union>i\<in>space MX. ?d i) = \<mu>' (Y -` B \<inter> space M)"
hoelzl@42852
   529
    using X[unfolded measurable_def] by (auto intro!: arg_cong[where f=\<mu>'])
hoelzl@41102
   530
qed
hoelzl@41102
   531
hoelzl@41102
   532
lemma (in prob_space) setsum_joint_distribution_singleton:
hoelzl@41102
   533
  assumes X: "finite_random_variable MX X"
hoelzl@41102
   534
  assumes Y: "finite_random_variable MY Y" "b \<in> space MY"
hoelzl@41102
   535
  shows "(\<Sum>a\<in>space MX. joint_distribution X Y {(a, b)}) = distribution Y {b}"
hoelzl@41102
   536
  using setsum_joint_distribution[OF X
hoelzl@41102
   537
    finite_random_variableD[OF Y(1)]
hoelzl@41102
   538
    finite_random_variable_imp_sets[OF Y]] by simp
hoelzl@41102
   539
hoelzl@42553
   540
locale pair_finite_prob_space = M1: finite_prob_space M1 + M2: finite_prob_space M2 for M1 M2
hoelzl@41102
   541
hoelzl@42553
   542
sublocale pair_finite_prob_space \<subseteq> pair_prob_space M1 M2 by default
hoelzl@42553
   543
sublocale pair_finite_prob_space \<subseteq> pair_finite_space M1 M2  by default
hoelzl@42553
   544
sublocale pair_finite_prob_space \<subseteq> finite_prob_space P by default
hoelzl@41102
   545
hoelzl@43728
   546
locale product_finite_prob_space =
hoelzl@43728
   547
  fixes M :: "'i \<Rightarrow> ('a,'b) measure_space_scheme"
hoelzl@43728
   548
    and I :: "'i set"
hoelzl@43728
   549
  assumes finite_space: "\<And>i. finite_prob_space (M i)" and finite_index: "finite I"
hoelzl@43728
   550
hoelzl@43728
   551
sublocale product_finite_prob_space \<subseteq> M!: finite_prob_space "M i" using finite_space .
hoelzl@43728
   552
sublocale product_finite_prob_space \<subseteq> finite_product_sigma_finite M I by default (rule finite_index)
hoelzl@43728
   553
sublocale product_finite_prob_space \<subseteq> prob_space "Pi\<^isub>M I M"
hoelzl@43728
   554
proof
hoelzl@43728
   555
  show "\<mu> (space P) = 1"
hoelzl@43728
   556
    using measure_times[OF M.top] M.measure_space_1
hoelzl@43728
   557
    by (simp add: setprod_1 space_product_algebra)
hoelzl@43728
   558
qed
hoelzl@43728
   559
hoelzl@43728
   560
lemma funset_eq_UN_fun_upd_I:
hoelzl@43728
   561
  assumes *: "\<And>f. f \<in> F (insert a A) \<Longrightarrow> f(a := d) \<in> F A"
hoelzl@43728
   562
  and **: "\<And>f. f \<in> F (insert a A) \<Longrightarrow> f a \<in> G (f(a:=d))"
hoelzl@43728
   563
  and ***: "\<And>f x. \<lbrakk> f \<in> F A ; x \<in> G f \<rbrakk> \<Longrightarrow> f(a:=x) \<in> F (insert a A)"
hoelzl@43728
   564
  shows "F (insert a A) = (\<Union>f\<in>F A. fun_upd f a ` (G f))"
hoelzl@43728
   565
proof safe
hoelzl@43728
   566
  fix f assume f: "f \<in> F (insert a A)"
hoelzl@43728
   567
  show "f \<in> (\<Union>f\<in>F A. fun_upd f a ` G f)"
hoelzl@43728
   568
  proof (rule UN_I[of "f(a := d)"])
hoelzl@43728
   569
    show "f(a := d) \<in> F A" using *[OF f] .
hoelzl@43728
   570
    show "f \<in> fun_upd (f(a:=d)) a ` G (f(a:=d))"
hoelzl@43728
   571
    proof (rule image_eqI[of _ _ "f a"])
hoelzl@43728
   572
      show "f a \<in> G (f(a := d))" using **[OF f] .
hoelzl@43728
   573
    qed simp
hoelzl@43728
   574
  qed
hoelzl@43728
   575
next
hoelzl@43728
   576
  fix f x assume "f \<in> F A" "x \<in> G f"
hoelzl@43728
   577
  from ***[OF this] show "f(a := x) \<in> F (insert a A)" .
hoelzl@43728
   578
qed
hoelzl@43728
   579
hoelzl@43728
   580
lemma extensional_funcset_insert_eq[simp]:
hoelzl@43728
   581
  assumes "a \<notin> A"
hoelzl@43728
   582
  shows "extensional (insert a A) \<inter> (insert a A \<rightarrow> B) = (\<Union>f \<in> extensional A \<inter> (A \<rightarrow> B). (\<lambda>b. f(a := b)) ` B)"
hoelzl@43728
   583
  apply (rule funset_eq_UN_fun_upd_I)
hoelzl@43728
   584
  using assms
hoelzl@43728
   585
  by (auto intro!: inj_onI dest: inj_onD split: split_if_asm simp: extensional_def)
hoelzl@43728
   586
hoelzl@43728
   587
lemma finite_extensional_funcset[simp, intro]:
hoelzl@43728
   588
  assumes "finite A" "finite B"
hoelzl@43728
   589
  shows "finite (extensional A \<inter> (A \<rightarrow> B))"
hoelzl@43728
   590
  using assms by induct (auto simp: extensional_funcset_insert_eq)
hoelzl@43728
   591
hoelzl@43728
   592
lemma finite_PiE[simp, intro]:
hoelzl@43728
   593
  assumes fin: "finite A" "\<And>i. i \<in> A \<Longrightarrow> finite (B i)"
hoelzl@43728
   594
  shows "finite (Pi\<^isub>E A B)"
hoelzl@43728
   595
proof -
hoelzl@43728
   596
  have *: "(Pi\<^isub>E A B) \<subseteq> extensional A \<inter> (A \<rightarrow> (\<Union>i\<in>A. B i))" by auto
hoelzl@43728
   597
  show ?thesis
hoelzl@43728
   598
    using fin by (intro finite_subset[OF *] finite_extensional_funcset) auto
hoelzl@43728
   599
qed
hoelzl@43728
   600
hoelzl@43763
   601
lemma (in product_finite_prob_space) singleton_eq_product:
hoelzl@43763
   602
  assumes x: "x \<in> space P" shows "{x} = (\<Pi>\<^isub>E i\<in>I. {x i})"
hoelzl@43763
   603
proof (safe intro!: ext[of _ x])
hoelzl@43763
   604
  fix y i assume *: "y \<in> (\<Pi> i\<in>I. {x i})" "y \<in> extensional I"
hoelzl@43763
   605
  with x show "y i = x i"
hoelzl@43763
   606
    by (cases "i \<in> I") (auto simp: extensional_def)
hoelzl@43763
   607
qed (insert x, auto)
hoelzl@43763
   608
hoelzl@43728
   609
sublocale product_finite_prob_space \<subseteq> finite_prob_space "Pi\<^isub>M I M"
hoelzl@43728
   610
proof
hoelzl@43728
   611
  show "finite (space P)"
hoelzl@43728
   612
    using finite_index M.finite_space by auto
hoelzl@43728
   613
hoelzl@43728
   614
  { fix x assume "x \<in> space P"
hoelzl@43763
   615
    with this[THEN singleton_eq_product]
hoelzl@43763
   616
    have "{x} \<in> sets P"
hoelzl@43728
   617
      by (auto intro!: in_P) }
hoelzl@43728
   618
  note x_in_P = this
hoelzl@43728
   619
hoelzl@43728
   620
  have "Pow (space P) \<subseteq> sets P"
hoelzl@43728
   621
  proof
hoelzl@43728
   622
    fix X assume "X \<in> Pow (space P)"
hoelzl@43728
   623
    moreover then have "finite X"
hoelzl@43728
   624
      using `finite (space P)` by (blast intro: finite_subset)
hoelzl@43728
   625
    ultimately have "(\<Union>x\<in>X. {x}) \<in> sets P"
hoelzl@43728
   626
      by (intro finite_UN x_in_P) auto
hoelzl@43728
   627
    then show "X \<in> sets P" by simp
hoelzl@43728
   628
  qed
hoelzl@43728
   629
  with space_closed show [simp]: "sets P = Pow (space P)" ..
hoelzl@43728
   630
hoelzl@43728
   631
  { fix x assume "x \<in> space P"
hoelzl@43728
   632
    from this top have "\<mu> {x} \<le> \<mu> (space P)" by (intro measure_mono) auto
hoelzl@43728
   633
    then show "\<mu> {x} \<noteq> \<infinity>"
hoelzl@43728
   634
      using measure_space_1 by auto }
hoelzl@43728
   635
qed
hoelzl@43728
   636
hoelzl@43728
   637
lemma (in product_finite_prob_space) measure_finite_times:
hoelzl@43728
   638
  "(\<And>i. i \<in> I \<Longrightarrow> X i \<subseteq> space (M i)) \<Longrightarrow> \<mu> (\<Pi>\<^isub>E i\<in>I. X i) = (\<Prod>i\<in>I. M.\<mu> i (X i))"
hoelzl@43728
   639
  by (rule measure_times) simp
hoelzl@43728
   640
hoelzl@43763
   641
lemma (in product_finite_prob_space) measure_singleton_times:
hoelzl@43763
   642
  assumes x: "x \<in> space P" shows "\<mu> {x} = (\<Prod>i\<in>I. M.\<mu> i {x i})"
hoelzl@43763
   643
  unfolding singleton_eq_product[OF x] using x
hoelzl@43763
   644
  by (intro measure_finite_times) auto
hoelzl@43763
   645
hoelzl@43763
   646
lemma (in product_finite_prob_space) prob_finite_times:
hoelzl@43728
   647
  assumes X: "\<And>i. i \<in> I \<Longrightarrow> X i \<subseteq> space (M i)"
hoelzl@43728
   648
  shows "prob (\<Pi>\<^isub>E i\<in>I. X i) = (\<Prod>i\<in>I. M.prob i (X i))"
hoelzl@43728
   649
proof -
hoelzl@43728
   650
  have "extreal (\<mu>' (\<Pi>\<^isub>E i\<in>I. X i)) = \<mu> (\<Pi>\<^isub>E i\<in>I. X i)"
hoelzl@43728
   651
    using X by (intro finite_measure_eq[symmetric] in_P) auto
hoelzl@43728
   652
  also have "\<dots> = (\<Prod>i\<in>I. M.\<mu> i (X i))"
hoelzl@43728
   653
    using measure_finite_times X by simp
hoelzl@43728
   654
  also have "\<dots> = extreal (\<Prod>i\<in>I. M.\<mu>' i (X i))"
hoelzl@43728
   655
    using X by (simp add: M.finite_measure_eq setprod_extreal)
hoelzl@43728
   656
  finally show ?thesis by simp
hoelzl@43728
   657
qed
hoelzl@43728
   658
hoelzl@43763
   659
lemma (in product_finite_prob_space) prob_singleton_times:
hoelzl@43763
   660
  assumes x: "x \<in> space P"
hoelzl@43763
   661
  shows "prob {x} = (\<Prod>i\<in>I. M.prob i {x i})"
hoelzl@43763
   662
  unfolding singleton_eq_product[OF x] using x
hoelzl@43763
   663
  by (intro prob_finite_times) auto
hoelzl@43763
   664
hoelzl@43763
   665
lemma (in product_finite_prob_space) prob_finite_product:
hoelzl@43763
   666
  "A \<subseteq> space P \<Longrightarrow> prob A = (\<Sum>x\<in>A. \<Prod>i\<in>I. M.prob i {x i})"
hoelzl@43763
   667
  by (auto simp add: finite_measure_singleton prob_singleton_times
hoelzl@43763
   668
           simp del: space_product_algebra
hoelzl@43763
   669
           intro!: setsum_cong prob_singleton_times)
hoelzl@43763
   670
hoelzl@41102
   671
lemma (in prob_space) joint_distribution_finite_prob_space:
hoelzl@41102
   672
  assumes X: "finite_random_variable MX X"
hoelzl@41102
   673
  assumes Y: "finite_random_variable MY Y"
hoelzl@42852
   674
  shows "finite_prob_space ((MX \<Otimes>\<^isub>M MY)\<lparr> measure := extreal \<circ> joint_distribution X Y\<rparr>)"
hoelzl@42553
   675
  by (intro distribution_finite_prob_space finite_random_variable_pairI X Y)
hoelzl@41102
   676
hoelzl@36612
   677
lemma finite_prob_space_eq:
hoelzl@42553
   678
  "finite_prob_space M \<longleftrightarrow> finite_measure_space M \<and> measure M (space M) = 1"
hoelzl@36612
   679
  unfolding finite_prob_space_def finite_measure_space_def prob_space_def prob_space_axioms_def
hoelzl@36612
   680
  by auto
hoelzl@36612
   681
hoelzl@36612
   682
lemma (in prob_space) not_empty: "space M \<noteq> {}"
hoelzl@42852
   683
  using prob_space empty_measure' by auto
hoelzl@36612
   684
hoelzl@38902
   685
lemma (in finite_prob_space) sum_over_space_eq_1: "(\<Sum>x\<in>space M. \<mu> {x}) = 1"
hoelzl@38902
   686
  using measure_space_1 sum_over_space by simp
hoelzl@36612
   687
hoelzl@36612
   688
lemma (in finite_prob_space) joint_distribution_restriction_fst:
hoelzl@36612
   689
  "joint_distribution X Y A \<le> distribution X (fst ` A)"
hoelzl@36612
   690
  unfolding distribution_def
hoelzl@42852
   691
proof (safe intro!: finite_measure_mono)
hoelzl@36612
   692
  fix x assume "x \<in> space M" and *: "(X x, Y x) \<in> A"
hoelzl@36612
   693
  show "x \<in> X -` fst ` A"
hoelzl@36612
   694
    by (auto intro!: image_eqI[OF _ *])
hoelzl@36612
   695
qed (simp_all add: sets_eq_Pow)
hoelzl@36612
   696
hoelzl@36612
   697
lemma (in finite_prob_space) joint_distribution_restriction_snd:
hoelzl@36612
   698
  "joint_distribution X Y A \<le> distribution Y (snd ` A)"
hoelzl@36612
   699
  unfolding distribution_def
hoelzl@42852
   700
proof (safe intro!: finite_measure_mono)
hoelzl@36612
   701
  fix x assume "x \<in> space M" and *: "(X x, Y x) \<in> A"
hoelzl@36612
   702
  show "x \<in> Y -` snd ` A"
hoelzl@36612
   703
    by (auto intro!: image_eqI[OF _ *])
hoelzl@36612
   704
qed (simp_all add: sets_eq_Pow)
hoelzl@36612
   705
hoelzl@36612
   706
lemma (in finite_prob_space) distribution_order:
hoelzl@36612
   707
  shows "0 \<le> distribution X x'"
hoelzl@36612
   708
  and "(distribution X x' \<noteq> 0) \<longleftrightarrow> (0 < distribution X x')"
hoelzl@36612
   709
  and "r \<le> joint_distribution X Y {(x, y)} \<Longrightarrow> r \<le> distribution X {x}"
hoelzl@36612
   710
  and "r \<le> joint_distribution X Y {(x, y)} \<Longrightarrow> r \<le> distribution Y {y}"
hoelzl@36612
   711
  and "r < joint_distribution X Y {(x, y)} \<Longrightarrow> r < distribution X {x}"
hoelzl@36612
   712
  and "r < joint_distribution X Y {(x, y)} \<Longrightarrow> r < distribution Y {y}"
hoelzl@36612
   713
  and "distribution X {x} = 0 \<Longrightarrow> joint_distribution X Y {(x, y)} = 0"
hoelzl@36612
   714
  and "distribution Y {y} = 0 \<Longrightarrow> joint_distribution X Y {(x, y)} = 0"
hoelzl@42852
   715
  using
hoelzl@36612
   716
    joint_distribution_restriction_fst[of X Y "{(x, y)}"]
hoelzl@36612
   717
    joint_distribution_restriction_snd[of X Y "{(x, y)}"]
hoelzl@42852
   718
  by (auto intro: antisym)
hoelzl@36612
   719
hoelzl@39331
   720
lemma (in finite_prob_space) distribution_mono:
hoelzl@39331
   721
  assumes "\<And>t. \<lbrakk> t \<in> space M ; X t \<in> x \<rbrakk> \<Longrightarrow> Y t \<in> y"
hoelzl@39331
   722
  shows "distribution X x \<le> distribution Y y"
hoelzl@39331
   723
  unfolding distribution_def
hoelzl@42852
   724
  using assms by (auto simp: sets_eq_Pow intro!: finite_measure_mono)
hoelzl@35974
   725
hoelzl@39331
   726
lemma (in finite_prob_space) distribution_mono_gt_0:
hoelzl@39331
   727
  assumes gt_0: "0 < distribution X x"
hoelzl@39331
   728
  assumes *: "\<And>t. \<lbrakk> t \<in> space M ; X t \<in> x \<rbrakk> \<Longrightarrow> Y t \<in> y"
hoelzl@39331
   729
  shows "0 < distribution Y y"
hoelzl@39331
   730
  by (rule less_le_trans[OF gt_0 distribution_mono]) (rule *)
hoelzl@39331
   731
hoelzl@39331
   732
lemma (in finite_prob_space) sum_over_space_distrib:
hoelzl@39331
   733
  "(\<Sum>x\<in>X`space M. distribution X {x}) = 1"
hoelzl@42852
   734
  unfolding distribution_def prob_space[symmetric] using finite_space
hoelzl@42852
   735
  by (subst finite_measure_finite_Union[symmetric])
hoelzl@42852
   736
     (auto simp add: disjoint_family_on_def sets_eq_Pow
hoelzl@42852
   737
           intro!: arg_cong[where f=\<mu>'])
hoelzl@39331
   738
hoelzl@39331
   739
lemma (in finite_prob_space) finite_sum_over_space_eq_1:
hoelzl@42852
   740
  "(\<Sum>x\<in>space M. prob {x}) = 1"
hoelzl@42852
   741
  using prob_space finite_space
hoelzl@42852
   742
  by (subst (asm) finite_measure_finite_singleton) auto
hoelzl@39331
   743
hoelzl@39331
   744
lemma (in prob_space) distribution_remove_const:
hoelzl@39331
   745
  shows "joint_distribution X (\<lambda>x. ()) {(x, ())} = distribution X {x}"
hoelzl@39331
   746
  and "joint_distribution (\<lambda>x. ()) X {((), x)} = distribution X {x}"
hoelzl@39331
   747
  and "joint_distribution X (\<lambda>x. (Y x, ())) {(x, y, ())} = joint_distribution X Y {(x, y)}"
hoelzl@39331
   748
  and "joint_distribution X (\<lambda>x. ((), Y x)) {(x, (), y)} = joint_distribution X Y {(x, y)}"
hoelzl@39331
   749
  and "distribution (\<lambda>x. ()) {()} = 1"
hoelzl@42852
   750
  by (auto intro!: arg_cong[where f=\<mu>'] simp: distribution_def prob_space[symmetric])
hoelzl@39331
   751
hoelzl@39331
   752
lemma (in finite_prob_space) setsum_distribution_gen:
hoelzl@39331
   753
  assumes "Z -` {c} \<inter> space M = (\<Union>x \<in> X`space M. Y -` {f x}) \<inter> space M"
hoelzl@39331
   754
  and "inj_on f (X`space M)"
hoelzl@39331
   755
  shows "(\<Sum>x \<in> X`space M. distribution Y {f x}) = distribution Z {c}"
hoelzl@39331
   756
  unfolding distribution_def assms
hoelzl@39331
   757
  using finite_space assms
hoelzl@42852
   758
  by (subst finite_measure_finite_Union[symmetric])
hoelzl@39331
   759
     (auto simp add: disjoint_family_on_def sets_eq_Pow inj_on_def
hoelzl@39331
   760
      intro!: arg_cong[where f=prob])
hoelzl@39331
   761
hoelzl@39331
   762
lemma (in finite_prob_space) setsum_distribution:
hoelzl@39331
   763
  "(\<Sum>x \<in> X`space M. joint_distribution X Y {(x, y)}) = distribution Y {y}"
hoelzl@39331
   764
  "(\<Sum>y \<in> Y`space M. joint_distribution X Y {(x, y)}) = distribution X {x}"
hoelzl@39331
   765
  "(\<Sum>x \<in> X`space M. joint_distribution X (\<lambda>x. (Y x, Z x)) {(x, y, z)}) = joint_distribution Y Z {(y, z)}"
hoelzl@39331
   766
  "(\<Sum>y \<in> Y`space M. joint_distribution X (\<lambda>x. (Y x, Z x)) {(x, y, z)}) = joint_distribution X Z {(x, z)}"
hoelzl@39331
   767
  "(\<Sum>z \<in> Z`space M. joint_distribution X (\<lambda>x. (Y x, Z x)) {(x, y, z)}) = joint_distribution X Y {(x, y)}"
hoelzl@39331
   768
  by (auto intro!: inj_onI setsum_distribution_gen)
hoelzl@39331
   769
hoelzl@39331
   770
lemma (in finite_prob_space) uniform_prob:
hoelzl@39331
   771
  assumes "x \<in> space M"
hoelzl@39331
   772
  assumes "\<And> x y. \<lbrakk>x \<in> space M ; y \<in> space M\<rbrakk> \<Longrightarrow> prob {x} = prob {y}"
hoelzl@42852
   773
  shows "prob {x} = 1 / card (space M)"
hoelzl@39331
   774
proof -
hoelzl@39331
   775
  have prob_x: "\<And> y. y \<in> space M \<Longrightarrow> prob {y} = prob {x}"
hoelzl@39331
   776
    using assms(2)[OF _ `x \<in> space M`] by blast
hoelzl@39331
   777
  have "1 = prob (space M)"
hoelzl@39331
   778
    using prob_space by auto
hoelzl@39331
   779
  also have "\<dots> = (\<Sum> x \<in> space M. prob {x})"
hoelzl@42852
   780
    using finite_measure_finite_Union[of "space M" "\<lambda> x. {x}", simplified]
hoelzl@39331
   781
      sets_eq_Pow inj_singleton[unfolded inj_on_def, rule_format]
hoelzl@39331
   782
      finite_space unfolding disjoint_family_on_def  prob_space[symmetric]
hoelzl@39331
   783
    by (auto simp add:setsum_restrict_set)
hoelzl@39331
   784
  also have "\<dots> = (\<Sum> y \<in> space M. prob {x})"
hoelzl@39331
   785
    using prob_x by auto
hoelzl@39331
   786
  also have "\<dots> = real_of_nat (card (space M)) * prob {x}" by simp
hoelzl@39331
   787
  finally have one: "1 = real (card (space M)) * prob {x}"
hoelzl@39331
   788
    using real_eq_of_nat by auto
hoelzl@39331
   789
  hence two: "real (card (space M)) \<noteq> 0" by fastsimp 
hoelzl@39331
   790
  from one have three: "prob {x} \<noteq> 0" by fastsimp
hoelzl@39331
   791
  thus ?thesis using one two three divide_cancel_right
hoelzl@39331
   792
    by (auto simp:field_simps)
hoelzl@35974
   793
qed
hoelzl@35974
   794
hoelzl@39326
   795
lemma (in prob_space) prob_space_subalgebra:
hoelzl@41793
   796
  assumes "sigma_algebra N" "sets N \<subseteq> sets M" "space N = space M"
hoelzl@42553
   797
    and "\<And>A. A \<in> sets N \<Longrightarrow> measure N A = \<mu> A"
hoelzl@42553
   798
  shows "prob_space N"
hoelzl@39326
   799
proof -
hoelzl@42553
   800
  interpret N: measure_space N
hoelzl@42553
   801
    by (rule measure_space_subalgebra[OF assms])
hoelzl@39326
   802
  show ?thesis
hoelzl@42553
   803
  proof qed (insert assms(4)[OF N.top], simp add: assms measure_space_1)
hoelzl@39326
   804
qed
hoelzl@39326
   805
hoelzl@39326
   806
lemma (in prob_space) prob_space_of_restricted_space:
hoelzl@42852
   807
  assumes "\<mu> A \<noteq> 0" "A \<in> sets M"
hoelzl@42553
   808
  shows "prob_space (restricted_space A \<lparr>measure := \<lambda>S. \<mu> S / \<mu> A\<rparr>)"
hoelzl@42553
   809
    (is "prob_space ?P")
hoelzl@42553
   810
proof -
hoelzl@42553
   811
  interpret A: measure_space "restricted_space A"
hoelzl@39326
   812
    using `A \<in> sets M` by (rule restricted_measure_space)
hoelzl@42553
   813
  interpret A': sigma_algebra ?P
hoelzl@42553
   814
    by (rule A.sigma_algebra_cong) auto
hoelzl@42553
   815
  show "prob_space ?P"
hoelzl@39326
   816
  proof
hoelzl@42553
   817
    show "measure ?P (space ?P) = 1"
hoelzl@42852
   818
      using real_measure[OF `A \<in> events`] `\<mu> A \<noteq> 0` by auto
hoelzl@42852
   819
    show "positive ?P (measure ?P)"
hoelzl@42852
   820
    proof (simp add: positive_def, safe)
hoelzl@42852
   821
      show "0 / \<mu> A = 0" using `\<mu> A \<noteq> 0` by (cases "\<mu> A") (auto simp: zero_extreal_def)
hoelzl@42852
   822
      fix B assume "B \<in> events"
hoelzl@42852
   823
      with real_measure[of "A \<inter> B"] real_measure[OF `A \<in> events`] `A \<in> sets M`
hoelzl@42852
   824
      show "0 \<le> \<mu> (A \<inter> B) / \<mu> A" by (auto simp: Int)
hoelzl@42852
   825
    qed
hoelzl@42852
   826
    show "countably_additive ?P (measure ?P)"
hoelzl@42852
   827
    proof (simp add: countably_additive_def, safe)
hoelzl@42852
   828
      fix B and F :: "nat \<Rightarrow> 'a set"
hoelzl@42852
   829
      assume F: "range F \<subseteq> op \<inter> A ` events" "disjoint_family F"
hoelzl@42852
   830
      { fix i
hoelzl@42852
   831
        from F have "F i \<in> op \<inter> A ` events" by auto
hoelzl@42852
   832
        with `A \<in> events` have "F i \<in> events" by auto }
hoelzl@42852
   833
      moreover then have "range F \<subseteq> events" by auto
hoelzl@42852
   834
      moreover have "\<And>S. \<mu> S / \<mu> A = inverse (\<mu> A) * \<mu> S"
hoelzl@42852
   835
        by (simp add: mult_commute divide_extreal_def)
hoelzl@42852
   836
      moreover have "0 \<le> inverse (\<mu> A)"
hoelzl@42852
   837
        using real_measure[OF `A \<in> events`] by auto
hoelzl@42852
   838
      ultimately show "(\<Sum>i. \<mu> (F i) / \<mu> A) = \<mu> (\<Union>i. F i) / \<mu> A"
hoelzl@42852
   839
        using measure_countably_additive[of F] F
hoelzl@42852
   840
        by (auto simp: suminf_cmult_extreal)
hoelzl@42852
   841
    qed
hoelzl@39326
   842
  qed
hoelzl@39326
   843
qed
hoelzl@39326
   844
hoelzl@39326
   845
lemma finite_prob_spaceI:
hoelzl@42852
   846
  assumes "finite (space M)" "sets M = Pow(space M)"
hoelzl@42852
   847
    and "measure M (space M) = 1" "measure M {} = 0" "\<And>A. A \<subseteq> space M \<Longrightarrow> 0 \<le> measure M A"
hoelzl@42553
   848
    and "\<And>A B. A\<subseteq>space M \<Longrightarrow> B\<subseteq>space M \<Longrightarrow> A \<inter> B = {} \<Longrightarrow> measure M (A \<union> B) = measure M A + measure M B"
hoelzl@42553
   849
  shows "finite_prob_space M"
hoelzl@39326
   850
  unfolding finite_prob_space_eq
hoelzl@39326
   851
proof
hoelzl@42553
   852
  show "finite_measure_space M" using assms
hoelzl@42852
   853
    by (auto intro!: finite_measure_spaceI)
hoelzl@42553
   854
  show "measure M (space M) = 1" by fact
hoelzl@39326
   855
qed
hoelzl@36612
   856
hoelzl@36612
   857
lemma (in finite_prob_space) finite_measure_space:
hoelzl@39331
   858
  fixes X :: "'a \<Rightarrow> 'x"
hoelzl@42852
   859
  shows "finite_measure_space \<lparr>space = X ` space M, sets = Pow (X ` space M), measure = extreal \<circ> distribution X\<rparr>"
hoelzl@42553
   860
    (is "finite_measure_space ?S")
hoelzl@39326
   861
proof (rule finite_measure_spaceI, simp_all)
hoelzl@36612
   862
  show "finite (X ` space M)" using finite_space by simp
hoelzl@39331
   863
next
hoelzl@39331
   864
  fix A B :: "'x set" assume "A \<inter> B = {}"
hoelzl@39331
   865
  then show "distribution X (A \<union> B) = distribution X A + distribution X B"
hoelzl@39331
   866
    unfolding distribution_def
hoelzl@42852
   867
    by (subst finite_measure_Union[symmetric])
hoelzl@42852
   868
       (auto intro!: arg_cong[where f=\<mu>'] simp: sets_eq_Pow)
hoelzl@36612
   869
qed
hoelzl@36612
   870
hoelzl@39331
   871
lemma (in finite_prob_space) finite_prob_space_of_images:
hoelzl@42852
   872
  "finite_prob_space \<lparr> space = X ` space M, sets = Pow (X ` space M), measure = extreal \<circ> distribution X \<rparr>"
hoelzl@42852
   873
  by (simp add: finite_prob_space_eq finite_measure_space measure_space_1 one_extreal_def)
hoelzl@39331
   874
hoelzl@39330
   875
lemma (in finite_prob_space) finite_product_measure_space:
hoelzl@39331
   876
  fixes X :: "'a \<Rightarrow> 'x" and Y :: "'a \<Rightarrow> 'y"
hoelzl@39330
   877
  assumes "finite s1" "finite s2"
hoelzl@42852
   878
  shows "finite_measure_space \<lparr> space = s1 \<times> s2, sets = Pow (s1 \<times> s2), measure = extreal \<circ> joint_distribution X Y\<rparr>"
hoelzl@42553
   879
    (is "finite_measure_space ?M")
hoelzl@39331
   880
proof (rule finite_measure_spaceI, simp_all)
hoelzl@39331
   881
  show "finite (s1 \<times> s2)"
hoelzl@39330
   882
    using assms by auto
hoelzl@39331
   883
next
hoelzl@39331
   884
  fix A B :: "('x*'y) set" assume "A \<inter> B = {}"
hoelzl@39331
   885
  then show "joint_distribution X Y (A \<union> B) = joint_distribution X Y A + joint_distribution X Y B"
hoelzl@39331
   886
    unfolding distribution_def
hoelzl@42852
   887
    by (subst finite_measure_Union[symmetric])
hoelzl@42852
   888
       (auto intro!: arg_cong[where f=\<mu>'] simp: sets_eq_Pow)
hoelzl@39330
   889
qed
hoelzl@39330
   890
hoelzl@39331
   891
lemma (in finite_prob_space) finite_product_measure_space_of_images:
hoelzl@39330
   892
  shows "finite_measure_space \<lparr> space = X ` space M \<times> Y ` space M,
hoelzl@42553
   893
                                sets = Pow (X ` space M \<times> Y ` space M),
hoelzl@42852
   894
                                measure = extreal \<circ> joint_distribution X Y \<rparr>"
hoelzl@39330
   895
  using finite_space by (auto intro!: finite_product_measure_space)
hoelzl@39330
   896
hoelzl@41102
   897
lemma (in finite_prob_space) finite_product_prob_space_of_images:
hoelzl@42553
   898
  "finite_prob_space \<lparr> space = X ` space M \<times> Y ` space M, sets = Pow (X ` space M \<times> Y ` space M),
hoelzl@42852
   899
                       measure = extreal \<circ> joint_distribution X Y \<rparr>"
hoelzl@42553
   900
  (is "finite_prob_space ?S")
hoelzl@42852
   901
proof (simp add: finite_prob_space_eq finite_product_measure_space_of_images one_extreal_def)
hoelzl@41102
   902
  have "X -` X ` space M \<inter> Y -` Y ` space M \<inter> space M = space M" by auto
hoelzl@41102
   903
  thus "joint_distribution X Y (X ` space M \<times> Y ` space M) = 1"
hoelzl@41102
   904
    by (simp add: distribution_def prob_space vimage_Times comp_def measure_space_1)
hoelzl@41102
   905
qed
hoelzl@41102
   906
hoelzl@39319
   907
section "Conditional Expectation and Probability"
hoelzl@39319
   908
hoelzl@39319
   909
lemma (in prob_space) conditional_expectation_exists:
hoelzl@42852
   910
  fixes X :: "'a \<Rightarrow> extreal" and N :: "('a, 'b) measure_space_scheme"
hoelzl@42852
   911
  assumes borel: "X \<in> borel_measurable M" "AE x. 0 \<le> X x"
hoelzl@42553
   912
  and N: "sigma_algebra N" "sets N \<subseteq> sets M" "space N = space M" "\<And>A. measure N A = \<mu> A"
hoelzl@42852
   913
  shows "\<exists>Y\<in>borel_measurable N. (\<forall>x. 0 \<le> Y x) \<and> (\<forall>C\<in>sets N.
hoelzl@42852
   914
      (\<integral>\<^isup>+x. Y x * indicator C x \<partial>M) = (\<integral>\<^isup>+x. X x * indicator C x \<partial>M))"
hoelzl@39317
   915
proof -
hoelzl@42553
   916
  note N(4)[simp]
hoelzl@42553
   917
  interpret P: prob_space N
hoelzl@41793
   918
    using prob_space_subalgebra[OF N] .
hoelzl@39317
   919
hoelzl@39317
   920
  let "?f A" = "\<lambda>x. X x * indicator A x"
hoelzl@42553
   921
  let "?Q A" = "integral\<^isup>P M (?f A)"
hoelzl@39317
   922
hoelzl@39317
   923
  from measure_space_density[OF borel]
hoelzl@42553
   924
  have Q: "measure_space (N\<lparr> measure := ?Q \<rparr>)"
hoelzl@42553
   925
    apply (rule measure_space.measure_space_subalgebra[of "M\<lparr> measure := ?Q \<rparr>"])
hoelzl@42553
   926
    using N by (auto intro!: P.sigma_algebra_cong)
hoelzl@42553
   927
  then interpret Q: measure_space "N\<lparr> measure := ?Q \<rparr>" .
hoelzl@39317
   928
hoelzl@39317
   929
  have "P.absolutely_continuous ?Q"
hoelzl@39317
   930
    unfolding P.absolutely_continuous_def
hoelzl@41793
   931
  proof safe
hoelzl@42553
   932
    fix A assume "A \<in> sets N" "P.\<mu> A = 0"
hoelzl@42852
   933
    then have f_borel: "?f A \<in> borel_measurable M" "AE x. x \<notin> A"
hoelzl@42852
   934
      using borel N by (auto intro!: borel_measurable_indicator AE_not_in)
hoelzl@42852
   935
    then show "?Q A = 0"
hoelzl@42852
   936
      by (auto simp add: positive_integral_0_iff_AE)
hoelzl@39317
   937
  qed
hoelzl@39317
   938
  from P.Radon_Nikodym[OF Q this]
hoelzl@42852
   939
  obtain Y where Y: "Y \<in> borel_measurable N" "\<And>x. 0 \<le> Y x"
hoelzl@42553
   940
    "\<And>A. A \<in> sets N \<Longrightarrow> ?Q A =(\<integral>\<^isup>+x. Y x * indicator A x \<partial>N)"
hoelzl@39317
   941
    by blast
hoelzl@41793
   942
  with N(2) show ?thesis
hoelzl@42852
   943
    by (auto intro!: bexI[OF _ Y(1)] simp: positive_integral_subalgebra[OF _ _ N(2,3,4,1)])
hoelzl@39317
   944
qed
hoelzl@39317
   945
hoelzl@39319
   946
definition (in prob_space)
hoelzl@42852
   947
  "conditional_expectation N X = (SOME Y. Y\<in>borel_measurable N \<and> (\<forall>x. 0 \<le> Y x)
hoelzl@42553
   948
    \<and> (\<forall>C\<in>sets N. (\<integral>\<^isup>+x. Y x * indicator C x\<partial>M) = (\<integral>\<^isup>+x. X x * indicator C x\<partial>M)))"
hoelzl@39319
   949
hoelzl@39319
   950
abbreviation (in prob_space)
hoelzl@39326
   951
  "conditional_prob N A \<equiv> conditional_expectation N (indicator A)"
hoelzl@39319
   952
hoelzl@39319
   953
lemma (in prob_space)
hoelzl@42852
   954
  fixes X :: "'a \<Rightarrow> extreal" and N :: "('a, 'b) measure_space_scheme"
hoelzl@42852
   955
  assumes borel: "X \<in> borel_measurable M" "AE x. 0 \<le> X x"
hoelzl@42553
   956
  and N: "sigma_algebra N" "sets N \<subseteq> sets M" "space N = space M" "\<And>A. measure N A = \<mu> A"
hoelzl@39319
   957
  shows borel_measurable_conditional_expectation:
hoelzl@41793
   958
    "conditional_expectation N X \<in> borel_measurable N"
hoelzl@41793
   959
  and conditional_expectation: "\<And>C. C \<in> sets N \<Longrightarrow>
hoelzl@42553
   960
      (\<integral>\<^isup>+x. conditional_expectation N X x * indicator C x \<partial>M) =
hoelzl@42553
   961
      (\<integral>\<^isup>+x. X x * indicator C x \<partial>M)"
hoelzl@41793
   962
   (is "\<And>C. C \<in> sets N \<Longrightarrow> ?eq C")
hoelzl@39319
   963
proof -
hoelzl@39319
   964
  note CE = conditional_expectation_exists[OF assms, unfolded Bex_def]
hoelzl@41793
   965
  then show "conditional_expectation N X \<in> borel_measurable N"
hoelzl@39319
   966
    unfolding conditional_expectation_def by (rule someI2_ex) blast
hoelzl@39319
   967
hoelzl@41793
   968
  from CE show "\<And>C. C \<in> sets N \<Longrightarrow> ?eq C"
hoelzl@39319
   969
    unfolding conditional_expectation_def by (rule someI2_ex) blast
hoelzl@39319
   970
qed
hoelzl@39319
   971
hoelzl@42852
   972
lemma (in sigma_algebra) factorize_measurable_function_pos:
hoelzl@42852
   973
  fixes Z :: "'a \<Rightarrow> extreal" and Y :: "'a \<Rightarrow> 'c"
hoelzl@39325
   974
  assumes "sigma_algebra M'" and "Y \<in> measurable M M'" "Z \<in> borel_measurable M"
hoelzl@42852
   975
  assumes Z: "Z \<in> borel_measurable (sigma_algebra.vimage_algebra M' (space M) Y)"
hoelzl@42852
   976
  shows "\<exists>g\<in>borel_measurable M'. \<forall>x\<in>space M. max 0 (Z x) = g (Y x)"
hoelzl@42852
   977
proof -
hoelzl@39325
   978
  interpret M': sigma_algebra M' by fact
hoelzl@39325
   979
  have Y: "Y \<in> space M \<rightarrow> space M'" using assms unfolding measurable_def by auto
hoelzl@39325
   980
  from M'.sigma_algebra_vimage[OF this]
hoelzl@39325
   981
  interpret va: sigma_algebra "M'.vimage_algebra (space M) Y" .
hoelzl@39325
   982
hoelzl@42852
   983
  from va.borel_measurable_implies_simple_function_sequence'[OF Z] guess f . note f = this
hoelzl@39325
   984
hoelzl@42553
   985
  have "\<forall>i. \<exists>g. simple_function M' g \<and> (\<forall>x\<in>space M. f i x = g (Y x))"
hoelzl@39325
   986
  proof
hoelzl@39325
   987
    fix i
hoelzl@42852
   988
    from f(1)[of i] have "finite (f i`space M)" and B_ex:
hoelzl@39325
   989
      "\<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"
hoelzl@42553
   990
      unfolding simple_function_def by auto
hoelzl@39325
   991
    from B_ex[THEN bchoice] guess B .. note B = this
hoelzl@39325
   992
hoelzl@39325
   993
    let ?g = "\<lambda>x. \<Sum>z\<in>f i`space M. z * indicator (B z) x"
hoelzl@39325
   994
hoelzl@42553
   995
    show "\<exists>g. simple_function M' g \<and> (\<forall>x\<in>space M. f i x = g (Y x))"
hoelzl@39325
   996
    proof (intro exI[of _ ?g] conjI ballI)
hoelzl@42553
   997
      show "simple_function M' ?g" using B by auto
hoelzl@39325
   998
hoelzl@39325
   999
      fix x assume "x \<in> space M"
hoelzl@42852
  1000
      then have "\<And>z. z \<in> f i`space M \<Longrightarrow> indicator (B z) (Y x) = (indicator (f i -` {z} \<inter> space M) x::extreal)"
hoelzl@39325
  1001
        unfolding indicator_def using B by auto
hoelzl@42852
  1002
      then show "f i x = ?g (Y x)" using `x \<in> space M` f(1)[of i]
hoelzl@39325
  1003
        by (subst va.simple_function_indicator_representation) auto
hoelzl@39325
  1004
    qed
hoelzl@39325
  1005
  qed
hoelzl@39325
  1006
  from choice[OF this] guess g .. note g = this
hoelzl@39325
  1007
hoelzl@42852
  1008
  show ?thesis
hoelzl@39325
  1009
  proof (intro ballI bexI)
hoelzl@41345
  1010
    show "(\<lambda>x. SUP i. g i x) \<in> borel_measurable M'"
hoelzl@39325
  1011
      using g by (auto intro: M'.borel_measurable_simple_function)
hoelzl@39325
  1012
    fix x assume "x \<in> space M"
hoelzl@42852
  1013
    have "max 0 (Z x) = (SUP i. f i x)" using f by simp
hoelzl@42852
  1014
    also have "\<dots> = (SUP i. g i (Y x))"
hoelzl@39325
  1015
      using g `x \<in> space M` by simp
hoelzl@42852
  1016
    finally show "max 0 (Z x) = (SUP i. g i (Y x))" .
hoelzl@42852
  1017
  qed
hoelzl@42852
  1018
qed
hoelzl@42852
  1019
hoelzl@42852
  1020
lemma extreal_0_le_iff_le_0[simp]:
hoelzl@42852
  1021
  fixes a :: extreal shows "0 \<le> -a \<longleftrightarrow> a \<le> 0"
hoelzl@42852
  1022
  by (cases rule: extreal2_cases[of a]) auto
hoelzl@42852
  1023
hoelzl@42852
  1024
lemma (in sigma_algebra) factorize_measurable_function:
hoelzl@42852
  1025
  fixes Z :: "'a \<Rightarrow> extreal" and Y :: "'a \<Rightarrow> 'c"
hoelzl@42852
  1026
  assumes "sigma_algebra M'" and "Y \<in> measurable M M'" "Z \<in> borel_measurable M"
hoelzl@42852
  1027
  shows "Z \<in> borel_measurable (sigma_algebra.vimage_algebra M' (space M) Y)
hoelzl@42852
  1028
    \<longleftrightarrow> (\<exists>g\<in>borel_measurable M'. \<forall>x\<in>space M. Z x = g (Y x))"
hoelzl@42852
  1029
proof safe
hoelzl@42852
  1030
  interpret M': sigma_algebra M' by fact
hoelzl@42852
  1031
  have Y: "Y \<in> space M \<rightarrow> space M'" using assms unfolding measurable_def by auto
hoelzl@42852
  1032
  from M'.sigma_algebra_vimage[OF this]
hoelzl@42852
  1033
  interpret va: sigma_algebra "M'.vimage_algebra (space M) Y" .
hoelzl@42852
  1034
hoelzl@42852
  1035
  { fix g :: "'c \<Rightarrow> extreal" assume "g \<in> borel_measurable M'"
hoelzl@42852
  1036
    with M'.measurable_vimage_algebra[OF Y]
hoelzl@42852
  1037
    have "g \<circ> Y \<in> borel_measurable (M'.vimage_algebra (space M) Y)"
hoelzl@42852
  1038
      by (rule measurable_comp)
hoelzl@42852
  1039
    moreover assume "\<forall>x\<in>space M. Z x = g (Y x)"
hoelzl@42852
  1040
    then have "Z \<in> borel_measurable (M'.vimage_algebra (space M) Y) \<longleftrightarrow>
hoelzl@42852
  1041
       g \<circ> Y \<in> borel_measurable (M'.vimage_algebra (space M) Y)"
hoelzl@42852
  1042
       by (auto intro!: measurable_cong)
hoelzl@42852
  1043
    ultimately show "Z \<in> borel_measurable (M'.vimage_algebra (space M) Y)"
hoelzl@42852
  1044
      by simp }
hoelzl@42852
  1045
hoelzl@42852
  1046
  assume Z: "Z \<in> borel_measurable (M'.vimage_algebra (space M) Y)"
hoelzl@42852
  1047
  with assms have "(\<lambda>x. - Z x) \<in> borel_measurable M"
hoelzl@42852
  1048
    "(\<lambda>x. - Z x) \<in> borel_measurable (M'.vimage_algebra (space M) Y)"
hoelzl@42852
  1049
    by auto
hoelzl@42852
  1050
  from factorize_measurable_function_pos[OF assms(1,2) this] guess n .. note n = this
hoelzl@42852
  1051
  from factorize_measurable_function_pos[OF assms Z] guess p .. note p = this
hoelzl@42852
  1052
  let "?g x" = "p x - n x"
hoelzl@42852
  1053
  show "\<exists>g\<in>borel_measurable M'. \<forall>x\<in>space M. Z x = g (Y x)"
hoelzl@42852
  1054
  proof (intro bexI ballI)
hoelzl@42852
  1055
    show "?g \<in> borel_measurable M'" using p n by auto
hoelzl@42852
  1056
    fix x assume "x \<in> space M"
hoelzl@42852
  1057
    then have "p (Y x) = max 0 (Z x)" "n (Y x) = max 0 (- Z x)"
hoelzl@42852
  1058
      using p n by auto
hoelzl@42852
  1059
    then show "Z x = ?g (Y x)"
hoelzl@42852
  1060
      by (auto split: split_max)
hoelzl@39325
  1061
  qed
hoelzl@39325
  1062
qed
hoelzl@39324
  1063
hoelzl@43729
  1064
subsection "Bernoulli space"
hoelzl@43729
  1065
hoelzl@43729
  1066
definition "bernoulli_space p = \<lparr> space = UNIV, sets = UNIV,
hoelzl@43729
  1067
  measure = extreal \<circ> setsum (\<lambda>b. if b then min 1 (max 0 p) else 1 - min 1 (max 0 p)) \<rparr>"
hoelzl@43729
  1068
hoelzl@43729
  1069
interpretation bernoulli: finite_prob_space "bernoulli_space p" for p
hoelzl@43729
  1070
  by (rule finite_prob_spaceI)
hoelzl@43729
  1071
     (auto simp: bernoulli_space_def UNIV_bool one_extreal_def setsum_Un_disjoint intro!: setsum_nonneg)
hoelzl@43729
  1072
hoelzl@43729
  1073
lemma bernoulli_measure:
hoelzl@43729
  1074
  "0 \<le> p \<Longrightarrow> p \<le> 1 \<Longrightarrow> bernoulli.prob p B = (\<Sum>b\<in>B. if b then p else 1 - p)"
hoelzl@43729
  1075
  unfolding bernoulli.\<mu>'_def unfolding bernoulli_space_def by (auto intro!: setsum_cong)
hoelzl@43729
  1076
hoelzl@43729
  1077
lemma bernoulli_measure_True: "0 \<le> p \<Longrightarrow> p \<le> 1 \<Longrightarrow> bernoulli.prob p {True} = p"
hoelzl@43729
  1078
  and bernoulli_measure_False: "0 \<le> p \<Longrightarrow> p \<le> 1 \<Longrightarrow> bernoulli.prob p {False} = 1 - p"
hoelzl@43729
  1079
  unfolding bernoulli_measure by simp_all
hoelzl@43729
  1080
hoelzl@35582
  1081
end