src/HOL/Probability/Probability_Space.thy
author hoelzl
Tue, 22 Mar 2011 20:06:10 +0100
changeset 42067 66c8281349ec
parent 41981 cdf7693bbe08
child 42146 5b52c6a9c627
permissions -rw-r--r--
standardized headers
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
42067
66c8281349ec standardized headers
hoelzl
parents: 41981
diff changeset
     1
(*  Title:      HOL/Probability/Probability_Space.thy
66c8281349ec standardized headers
hoelzl
parents: 41981
diff changeset
     2
    Author:     Johannes Hölzl, TU München
66c8281349ec standardized headers
hoelzl
parents: 41981
diff changeset
     3
    Author:     Armin Heller, TU München
66c8281349ec standardized headers
hoelzl
parents: 41981
diff changeset
     4
*)
66c8281349ec standardized headers
hoelzl
parents: 41981
diff changeset
     5
66c8281349ec standardized headers
hoelzl
parents: 41981
diff changeset
     6
header {*Probability spaces*}
66c8281349ec standardized headers
hoelzl
parents: 41981
diff changeset
     7
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
     8
theory Probability_Space
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
     9
imports Lebesgue_Integration Radon_Nikodym Product_Measure
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
    10
begin
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
    11
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    12
lemma real_of_extreal_inverse[simp]:
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    13
  fixes X :: extreal
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
    14
  shows "real (inverse X) = 1 / real X"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
    15
  by (cases X) (auto simp: inverse_eq_divide)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
    16
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    17
lemma real_of_extreal_le_0[simp]: "real (X :: extreal) \<le> 0 \<longleftrightarrow> (X \<le> 0 \<or> X = \<infinity>)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    18
  by (cases X) auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    19
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    20
lemma abs_real_of_extreal[simp]: "\<bar>real (X :: extreal)\<bar> = real \<bar>X\<bar>"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
    21
  by (cases X) auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
    22
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    23
lemma zero_less_real_of_extreal: "0 < real X \<longleftrightarrow> (0 < X \<and> X \<noteq> \<infinity>)"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
    24
  by (cases X) auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
    25
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    26
lemma real_of_extreal_le_1: fixes X :: extreal shows "X \<le> 1 \<Longrightarrow> real X \<le> 1"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    27
  by (cases X) (auto simp: one_extreal_def)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    28
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
    29
locale prob_space = measure_space +
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
    30
  assumes measure_space_1: "measure M (space M) = 1"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
    31
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
    32
sublocale prob_space < finite_measure
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
    33
proof
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    34
  from measure_space_1 show "\<mu> (space M) \<noteq> \<infinity>" by simp
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
    35
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
    36
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
    37
abbreviation (in prob_space) "events \<equiv> sets M"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    38
abbreviation (in prob_space) "prob \<equiv> \<mu>'"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
    39
abbreviation (in prob_space) "prob_preserving \<equiv> measure_preserving"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
    40
abbreviation (in prob_space) "random_variable M' X \<equiv> sigma_algebra M' \<and> X \<in> measurable M M'"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
    41
abbreviation (in prob_space) "expectation \<equiv> integral\<^isup>L M"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
    42
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
    43
definition (in prob_space)
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
    44
  "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
    45
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
    46
definition (in prob_space)
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
    47
  "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
    48
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
    49
definition (in prob_space)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    50
  "distribution X A = \<mu>' (X -` A \<inter> space M)"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
    51
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
    52
abbreviation (in prob_space)
36624
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
    53
  "joint_distribution X Y \<equiv> distribution (\<lambda>x. (X x, Y x))"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
    54
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    55
declare (in finite_measure) positive_measure'[intro, simp]
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    56
39097
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
    57
lemma (in prob_space) distribution_cong:
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
    58
  assumes "\<And>x. x \<in> space M \<Longrightarrow> X x = Y x"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
    59
  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
    60
  unfolding distribution_def fun_eq_iff
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    61
  using assms by (auto intro!: arg_cong[where f="\<mu>'"])
39097
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
    62
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
    63
lemma (in prob_space) joint_distribution_cong:
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
    64
  assumes "\<And>x. x \<in> space M \<Longrightarrow> X x = X' x"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
    65
  assumes "\<And>x. x \<in> space M \<Longrightarrow> Y x = Y' x"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
    66
  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
    67
  unfolding distribution_def fun_eq_iff
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    68
  using assms by (auto intro!: arg_cong[where f="\<mu>'"])
39097
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
    69
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
    70
lemma (in prob_space) distribution_id[simp]:
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    71
  "N \<in> events \<Longrightarrow> distribution (\<lambda>x. x) N = prob N"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    72
  by (auto simp: distribution_def intro!: arg_cong[where f=prob])
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
    73
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
    74
lemma (in prob_space) prob_space: "prob (space M) = 1"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    75
  using measure_space_1 unfolding \<mu>'_def by (simp add: one_extreal_def)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    76
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    77
lemma (in prob_space) prob_le_1[simp, intro]: "prob A \<le> 1"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    78
  using bounded_measure[of A] by (simp add: prob_space)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    79
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    80
lemma (in prob_space) distribution_positive[simp, intro]:
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    81
  "0 \<le> distribution X A" unfolding distribution_def by auto
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
    82
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    83
lemma (in prob_space) joint_distribution_remove[simp]:
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    84
    "joint_distribution X X {(x, x)} = distribution X {x}"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    85
  unfolding distribution_def by (auto intro!: arg_cong[where f=\<mu>'])
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    86
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    87
lemma (in prob_space) distribution_1:
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    88
  "distribution X A \<le> 1"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    89
  unfolding distribution_def by simp
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
    90
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
    91
lemma (in prob_space) prob_compl:
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    92
  assumes A: "A \<in> events"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
    93
  shows "prob (space M - A) = 1 - prob A"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    94
  using finite_measure_compl[OF A] by (simp add: prob_space)
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
    95
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    96
lemma (in prob_space) indep_space: "s \<in> events \<Longrightarrow> indep (space M) s"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    97
  by (simp add: indep_def prob_space)
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
    98
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
    99
lemma (in prob_space) prob_space_increasing: "increasing M prob"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   100
  by (auto intro!: finite_measure_mono simp: increasing_def)
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   101
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   102
lemma (in prob_space) prob_zero_union:
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   103
  assumes "s \<in> events" "t \<in> events" "prob t = 0"
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   104
  shows "prob (s \<union> t) = prob s"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   105
using assms
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   106
proof -
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   107
  have "prob (s \<union> t) \<le> prob s"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   108
    using finite_measure_subadditive[of s t] assms by auto
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   109
  moreover have "prob (s \<union> t) \<ge> prob s"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   110
    using assms by (blast intro: finite_measure_mono)
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   111
  ultimately show ?thesis by simp
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   112
qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   113
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   114
lemma (in prob_space) prob_eq_compl:
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   115
  assumes "s \<in> events" "t \<in> events"
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   116
  assumes "prob (space M - s) = prob (space M - t)"
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   117
  shows "prob s = prob t"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   118
  using assms prob_compl by auto
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   119
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   120
lemma (in prob_space) prob_one_inter:
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   121
  assumes events:"s \<in> events" "t \<in> events"
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   122
  assumes "prob t = 1"
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   123
  shows "prob (s \<inter> t) = prob s"
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   124
proof -
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   125
  have "prob ((space M - s) \<union> (space M - t)) = prob (space M - s)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   126
    using events assms  prob_compl[of "t"] by (auto intro!: prob_zero_union)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   127
  also have "(space M - s) \<union> (space M - t) = space M - (s \<inter> t)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   128
    by blast
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   129
  finally show "prob (s \<inter> t) = prob s"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   130
    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
   131
qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   132
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   133
lemma (in prob_space) prob_eq_bigunion_image:
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   134
  assumes "range f \<subseteq> events" "range g \<subseteq> events"
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   135
  assumes "disjoint_family f" "disjoint_family g"
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   136
  assumes "\<And> n :: nat. prob (f n) = prob (g n)"
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   137
  shows "(prob (\<Union> i. f i) = prob (\<Union> i. g i))"
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   138
using assms
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   139
proof -
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   140
  have a: "(\<lambda> i. prob (f i)) sums (prob (\<Union> i. f i))"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   141
    by (rule finite_measure_UNION[OF assms(1,3)])
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   142
  have b: "(\<lambda> i. prob (g i)) sums (prob (\<Union> i. g i))"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   143
    by (rule finite_measure_UNION[OF assms(2,4)])
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   144
  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
   145
qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   146
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   147
lemma (in prob_space) prob_countably_zero:
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   148
  assumes "range c \<subseteq> events"
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   149
  assumes "\<And> i. prob (c i) = 0"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   150
  shows "prob (\<Union> i :: nat. c i) = 0"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   151
proof (rule antisym)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   152
  show "prob (\<Union> i :: nat. c i) \<le> 0"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   153
    using finite_measure_countably_subadditive[OF assms(1)]
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   154
    by (simp add: assms(2) suminf_zero summable_zero)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   155
qed simp
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   156
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   157
lemma (in prob_space) indep_sym:
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   158
   "indep a b \<Longrightarrow> indep b a"
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   159
unfolding indep_def using Int_commute[of a b] by auto
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   160
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   161
lemma (in prob_space) indep_refl:
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   162
  assumes "a \<in> events"
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   163
  shows "indep a a = (prob a = 0) \<or> (prob a = 1)"
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   164
using assms unfolding indep_def by auto
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   165
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   166
lemma (in prob_space) prob_equiprobable_finite_unions:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   167
  assumes "s \<in> events"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   168
  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
   169
  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
   170
  shows "prob s = real (card s) * prob {SOME x. x \<in> s}"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   171
proof (cases "s = {}")
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   172
  case False hence "\<exists> x. x \<in> s" by blast
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   173
  from someI_ex[OF this] assms
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   174
  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
   175
  have "prob s = (\<Sum> x \<in> s. prob {x})"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   176
    using finite_measure_finite_singleton[OF s_finite] by simp
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   177
  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
   178
  also have "\<dots> = real (card s) * prob {(SOME x. x \<in> s)}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   179
    using setsum_constant assms by (simp add: real_eq_of_nat)
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   180
  finally show ?thesis by simp
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   181
qed simp
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   182
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   183
lemma (in prob_space) prob_real_sum_image_fn:
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   184
  assumes "e \<in> events"
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   185
  assumes "\<And> x. x \<in> s \<Longrightarrow> e \<inter> f x \<in> events"
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   186
  assumes "finite s"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   187
  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
   188
  assumes upper: "space M \<subseteq> (\<Union> i \<in> s. f i)"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   189
  shows "prob e = (\<Sum> x \<in> s. prob (e \<inter> f x))"
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   190
proof -
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   191
  have e: "e = (\<Union> i \<in> s. e \<inter> f i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   192
    using `e \<in> events` sets_into_space upper by blast
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   193
  hence "prob e = prob (\<Union> i \<in> s. e \<inter> f i)" by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   194
  also have "\<dots> = (\<Sum> x \<in> s. prob (e \<inter> f x))"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   195
  proof (rule finite_measure_finite_Union)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   196
    show "finite s" by fact
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   197
    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
   198
    show "disjoint_family_on (\<lambda>i. e \<inter> f i) s"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   199
      using disjoint by (auto simp: disjoint_family_on_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   200
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   201
  finally show ?thesis .
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   202
qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   203
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   204
lemma (in prob_space) distribution_prob_space:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   205
  assumes "random_variable S X"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   206
  shows "prob_space (S\<lparr>measure := extreal \<circ> distribution X\<rparr>)"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   207
proof -
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   208
  interpret S: measure_space "S\<lparr>measure := extreal \<circ> distribution X\<rparr>"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   209
  proof (rule measure_space.measure_space_cong)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   210
    show "measure_space (S\<lparr> measure := \<lambda>A. \<mu> (X -` A \<inter> space M) \<rparr>)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   211
      using assms by (auto intro!: measure_space_vimage simp: measure_preserving_def)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   212
  qed (insert assms, auto simp add: finite_measure_eq distribution_def measurable_sets)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   213
  show ?thesis
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   214
  proof (default, simp)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   215
    have "X -` space S \<inter> space M = space M"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   216
      using `random_variable S X` by (auto simp: measurable_def)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   217
    then show "extreal (distribution X (space S)) = 1"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   218
      by (simp add: distribution_def one_extreal_def prob_space)
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   219
  qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   220
qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   221
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   222
lemma (in prob_space) AE_distribution:
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   223
  assumes X: "random_variable MX X" and "AE x in MX\<lparr>measure := extreal \<circ> distribution X\<rparr>. Q x"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   224
  shows "AE x. Q (X x)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   225
proof -
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   226
  interpret X: prob_space "MX\<lparr>measure := extreal \<circ> distribution X\<rparr>" using X by (rule distribution_prob_space)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   227
  obtain N where N: "N \<in> sets MX" "distribution X N = 0" "{x\<in>space MX. \<not> Q x} \<subseteq> N"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   228
    using assms unfolding X.almost_everywhere_def by auto
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   229
  from X[unfolded measurable_def] N show "AE x. Q (X x)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   230
    by (intro AE_I'[where N="X -` N \<inter> space M"])
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   231
       (auto simp: finite_measure_eq distribution_def measurable_sets)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   232
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   233
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   234
lemma (in prob_space) distribution_eq_integral:
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   235
  "random_variable S X \<Longrightarrow> A \<in> sets S \<Longrightarrow> distribution X A = expectation (indicator (X -` A \<inter> space M))"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   236
  using finite_measure_eq[of "X -` A \<inter> space M"]
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   237
  by (auto simp: measurable_sets distribution_def)
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   238
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   239
lemma (in prob_space) distribution_eq_translated_integral:
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   240
  assumes "random_variable S X" "A \<in> sets S"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   241
  shows "distribution X A = integral\<^isup>P (S\<lparr>measure := extreal \<circ> distribution X\<rparr>) (indicator A)"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   242
proof -
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   243
  interpret S: prob_space "S\<lparr>measure := extreal \<circ> distribution X\<rparr>"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   244
    using assms(1) by (rule distribution_prob_space)
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   245
  show ?thesis
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   246
    using S.positive_integral_indicator(1)[of A] assms by simp
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   247
qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   248
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   249
lemma (in prob_space) finite_expectation1:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   250
  assumes f: "finite (X`space M)" and rv: "random_variable borel X"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   251
  shows "expectation X = (\<Sum>r \<in> X ` space M. r * prob (X -` {r} \<inter> space M))" (is "_ = ?r")
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   252
proof (subst integral_on_finite)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   253
  show "X \<in> borel_measurable M" "finite (X`space M)" using assms by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   254
  show "(\<Sum> r \<in> X ` space M. r * real (\<mu> (X -` {r} \<inter> space M))) = ?r"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   255
    "\<And>x. \<mu> (X -` {x} \<inter> space M) \<noteq> \<infinity>"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   256
    using finite_measure_eq[OF borel_measurable_vimage, of X] rv by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   257
qed
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   258
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   259
lemma (in prob_space) finite_expectation:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   260
  assumes "finite (X`space M)" "random_variable borel X"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   261
  shows "expectation X = (\<Sum> r \<in> X ` (space M). r * distribution X {r})"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   262
  using assms unfolding distribution_def using finite_expectation1 by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   263
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   264
lemma (in prob_space) prob_x_eq_1_imp_prob_y_eq_0:
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   265
  assumes "{x} \<in> events"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   266
  assumes "prob {x} = 1"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   267
  assumes "{y} \<in> events"
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   268
  assumes "y \<noteq> x"
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   269
  shows "prob {y} = 0"
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   270
  using prob_one_inter[of "{y}" "{x}"] assms by auto
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   271
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   272
lemma (in prob_space) distribution_empty[simp]: "distribution X {} = 0"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   273
  unfolding distribution_def by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   274
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   275
lemma (in prob_space) distribution_space[simp]: "distribution X (X ` space M) = 1"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   276
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   277
  have "X -` X ` space M \<inter> space M = space M" by auto
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   278
  thus ?thesis unfolding distribution_def by (simp add: prob_space)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   279
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   280
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   281
lemma (in prob_space) distribution_one:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   282
  assumes "random_variable M' X" and "A \<in> sets M'"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   283
  shows "distribution X A \<le> 1"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   284
proof -
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   285
  have "distribution X A \<le> \<mu>' (space M)" unfolding distribution_def
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   286
    using assms[unfolded measurable_def] by (auto intro!: finite_measure_mono)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   287
  thus ?thesis by (simp add: prob_space)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   288
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   289
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   290
lemma (in prob_space) distribution_x_eq_1_imp_distribution_y_eq_0:
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   291
  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
   292
    (is "random_variable ?S X")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   293
  assumes "distribution X {x} = 1"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   294
  assumes "y \<noteq> x"
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   295
  shows "distribution X {y} = 0"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   296
proof cases
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   297
  { fix x have "X -` {x} \<inter> space M \<in> sets M"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   298
    proof cases
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   299
      assume "x \<in> X`space M" with X show ?thesis
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   300
        by (auto simp: measurable_def image_iff)
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   301
    next
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   302
      assume "x \<notin> X`space M" then have "X -` {x} \<inter> space M = {}" by auto
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   303
      then show ?thesis by auto
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   304
    qed } note single = this
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   305
  have "X -` {x} \<inter> space M - X -` {y} \<inter> space M = X -` {x} \<inter> space M"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   306
    "X -` {y} \<inter> space M \<inter> (X -` {x} \<inter> space M) = {}"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   307
    using `y \<noteq> x` by auto
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   308
  with finite_measure_inter_full_set[OF single single, of x y] assms(2)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   309
  show ?thesis by (auto simp: distribution_def prob_space)
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   310
next
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   311
  assume "{y} \<notin> sets ?S"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   312
  then have "X -` {y} \<inter> space M = {}" by auto
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   313
  thus "distribution X {y} = 0" unfolding distribution_def by auto
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   314
qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   315
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   316
lemma (in prob_space) joint_distribution_Times_le_fst:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   317
  assumes X: "random_variable MX X" and Y: "random_variable MY Y"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   318
    and A: "A \<in> sets MX" and B: "B \<in> sets MY"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   319
  shows "joint_distribution X Y (A \<times> B) \<le> distribution X A"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   320
  unfolding distribution_def
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   321
proof (intro finite_measure_mono)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   322
  show "(\<lambda>x. (X x, Y x)) -` (A \<times> B) \<inter> space M \<subseteq> X -` A \<inter> space M" by force
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   323
  show "X -` A \<inter> space M \<in> events"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   324
    using X A unfolding measurable_def by simp
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   325
  have *: "(\<lambda>x. (X x, Y x)) -` (A \<times> B) \<inter> space M =
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   326
    (X -` A \<inter> space M) \<inter> (Y -` B \<inter> space M)" by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   327
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   328
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   329
lemma (in prob_space) joint_distribution_commute:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   330
  "joint_distribution X Y x = joint_distribution Y X ((\<lambda>(x,y). (y,x))`x)"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   331
  unfolding distribution_def by (auto intro!: arg_cong[where f=\<mu>'])
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   332
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   333
lemma (in prob_space) joint_distribution_Times_le_snd:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   334
  assumes X: "random_variable MX X" and Y: "random_variable MY Y"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   335
    and A: "A \<in> sets MX" and B: "B \<in> sets MY"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   336
  shows "joint_distribution X Y (A \<times> B) \<le> distribution Y B"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   337
  using assms
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   338
  by (subst joint_distribution_commute)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   339
     (simp add: swap_product joint_distribution_Times_le_fst)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   340
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   341
lemma (in prob_space) random_variable_pairI:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   342
  assumes "random_variable MX X"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   343
  assumes "random_variable MY Y"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   344
  shows "random_variable (MX \<Otimes>\<^isub>M MY) (\<lambda>x. (X x, Y x))"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   345
proof
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   346
  interpret MX: sigma_algebra MX using assms by simp
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   347
  interpret MY: sigma_algebra MY using assms by simp
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   348
  interpret P: pair_sigma_algebra MX MY by default
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   349
  show "sigma_algebra (MX \<Otimes>\<^isub>M MY)" by default
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   350
  have sa: "sigma_algebra M" by default
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   351
  show "(\<lambda>x. (X x, Y x)) \<in> measurable M (MX \<Otimes>\<^isub>M MY)"
41095
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   352
    unfolding P.measurable_pair_iff[OF sa] using assms by (simp add: comp_def)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   353
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   354
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   355
lemma (in prob_space) joint_distribution_commute_singleton:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   356
  "joint_distribution X Y {(x, y)} = joint_distribution Y X {(y, x)}"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   357
  unfolding distribution_def by (auto intro!: arg_cong[where f=\<mu>'])
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   358
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   359
lemma (in prob_space) joint_distribution_assoc_singleton:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   360
  "joint_distribution X (\<lambda>x. (Y x, Z x)) {(x, y, z)} =
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   361
   joint_distribution (\<lambda>x. (X x, Y x)) Z {((x, y), z)}"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   362
  unfolding distribution_def by (auto intro!: arg_cong[where f=\<mu>'])
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   363
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   364
locale pair_prob_space = M1: prob_space M1 + M2: prob_space M2 for M1 M2
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   365
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   366
sublocale pair_prob_space \<subseteq> pair_sigma_finite M1 M2 by default
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   367
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   368
sublocale pair_prob_space \<subseteq> P: prob_space P
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   369
by default (simp add: pair_measure_times M1.measure_space_1 M2.measure_space_1 space_pair_measure)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   370
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   371
lemma countably_additiveI[case_names countably]:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   372
  assumes "\<And>A. \<lbrakk> range A \<subseteq> sets M ; disjoint_family A ; (\<Union>i. A i) \<in> sets M\<rbrakk> \<Longrightarrow>
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   373
    (\<Sum>n. \<mu> (A n)) = \<mu> (\<Union>i. A i)"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   374
  shows "countably_additive M \<mu>"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   375
  using assms unfolding countably_additive_def by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   376
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   377
lemma (in prob_space) joint_distribution_prob_space:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   378
  assumes "random_variable MX X" "random_variable MY Y"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   379
  shows "prob_space ((MX \<Otimes>\<^isub>M MY) \<lparr> measure := extreal \<circ> joint_distribution X Y\<rparr>)"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   380
  using random_variable_pairI[OF assms] by (rule distribution_prob_space)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   381
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   382
section "Probability spaces on finite sets"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   383
35977
30d42bfd0174 Added finite measure space.
hoelzl
parents: 35929
diff changeset
   384
locale finite_prob_space = prob_space + finite_measure_space
30d42bfd0174 Added finite measure space.
hoelzl
parents: 35929
diff changeset
   385
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   386
abbreviation (in prob_space) "finite_random_variable M' X \<equiv> finite_sigma_algebra M' \<and> X \<in> measurable M M'"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   387
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   388
lemma (in prob_space) finite_random_variableD:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   389
  assumes "finite_random_variable M' X" shows "random_variable M' X"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   390
proof -
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   391
  interpret M': finite_sigma_algebra M' using assms by simp
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   392
  then show "random_variable M' X" using assms by simp default
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   393
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   394
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   395
lemma (in prob_space) distribution_finite_prob_space:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   396
  assumes "finite_random_variable MX X"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   397
  shows "finite_prob_space (MX\<lparr>measure := extreal \<circ> distribution X\<rparr>)"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   398
proof -
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   399
  interpret X: prob_space "MX\<lparr>measure := extreal \<circ> distribution X\<rparr>"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   400
    using assms[THEN finite_random_variableD] by (rule distribution_prob_space)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   401
  interpret MX: finite_sigma_algebra MX
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   402
    using assms by auto
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   403
  show ?thesis by default (simp_all add: MX.finite_space)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   404
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   405
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   406
lemma (in prob_space) simple_function_imp_finite_random_variable[simp, intro]:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   407
  assumes "simple_function M X"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   408
  shows "finite_random_variable \<lparr> space = X`space M, sets = Pow (X`space M), \<dots> = x \<rparr> X"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   409
    (is "finite_random_variable ?X _")
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   410
proof (intro conjI)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   411
  have [simp]: "finite (X ` space M)" using assms unfolding simple_function_def by simp
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   412
  interpret X: sigma_algebra ?X by (rule sigma_algebra_Pow)
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   413
  show "finite_sigma_algebra ?X"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   414
    by default auto
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   415
  show "X \<in> measurable M ?X"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   416
  proof (unfold measurable_def, clarsimp)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   417
    fix A assume A: "A \<subseteq> X`space M"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   418
    then have "finite A" by (rule finite_subset) simp
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   419
    then have "X -` (\<Union>a\<in>A. {a}) \<inter> space M \<in> events"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   420
      unfolding vimage_UN UN_extend_simps
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   421
      apply (rule finite_UN)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   422
      using A assms unfolding simple_function_def by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   423
    then show "X -` A \<inter> space M \<in> events" by simp
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   424
  qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   425
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   426
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   427
lemma (in prob_space) simple_function_imp_random_variable[simp, intro]:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   428
  assumes "simple_function M X"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   429
  shows "random_variable \<lparr> space = X`space M, sets = Pow (X`space M), \<dots> = ext \<rparr> X"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   430
  using simple_function_imp_finite_random_variable[OF assms, of ext]
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   431
  by (auto dest!: finite_random_variableD)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   432
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   433
lemma (in prob_space) sum_over_space_real_distribution:
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   434
  "simple_function M X \<Longrightarrow> (\<Sum>x\<in>X`space M. distribution X {x}) = 1"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   435
  unfolding distribution_def prob_space[symmetric]
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   436
  by (subst finite_measure_finite_Union[symmetric])
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   437
     (auto simp add: disjoint_family_on_def simple_function_def
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   438
           intro!: arg_cong[where f=prob])
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   439
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   440
lemma (in prob_space) finite_random_variable_pairI:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   441
  assumes "finite_random_variable MX X"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   442
  assumes "finite_random_variable MY Y"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   443
  shows "finite_random_variable (MX \<Otimes>\<^isub>M MY) (\<lambda>x. (X x, Y x))"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   444
proof
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   445
  interpret MX: finite_sigma_algebra MX using assms by simp
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   446
  interpret MY: finite_sigma_algebra MY using assms by simp
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   447
  interpret P: pair_finite_sigma_algebra MX MY by default
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   448
  show "finite_sigma_algebra (MX \<Otimes>\<^isub>M MY)" by default
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   449
  have sa: "sigma_algebra M" by default
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   450
  show "(\<lambda>x. (X x, Y x)) \<in> measurable M (MX \<Otimes>\<^isub>M MY)"
41095
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41023
diff changeset
   451
    unfolding P.measurable_pair_iff[OF sa] using assms by (simp add: comp_def)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   452
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   453
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   454
lemma (in prob_space) finite_random_variable_imp_sets:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   455
  "finite_random_variable MX X \<Longrightarrow> x \<in> space MX \<Longrightarrow> {x} \<in> sets MX"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   456
  unfolding finite_sigma_algebra_def finite_sigma_algebra_axioms_def by simp
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   457
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   458
lemma (in prob_space) finite_random_variable_measurable:
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   459
  assumes X: "finite_random_variable MX X" shows "X -` A \<inter> space M \<in> events"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   460
proof -
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   461
  interpret X: finite_sigma_algebra MX using X by simp
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   462
  from X have vimage: "\<And>A. A \<subseteq> space MX \<Longrightarrow> X -` A \<inter> space M \<in> events" and
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   463
    "X \<in> space M \<rightarrow> space MX"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   464
    by (auto simp: measurable_def)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   465
  then have *: "X -` A \<inter> space M = X -` (A \<inter> space MX) \<inter> space M"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   466
    by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   467
  show "X -` A \<inter> space M \<in> events"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   468
    unfolding * by (intro vimage) auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   469
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   470
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   471
lemma (in prob_space) joint_distribution_finite_Times_le_fst:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   472
  assumes X: "finite_random_variable MX X" and Y: "finite_random_variable MY Y"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   473
  shows "joint_distribution X Y (A \<times> B) \<le> distribution X A"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   474
  unfolding distribution_def
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   475
proof (intro finite_measure_mono)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   476
  show "(\<lambda>x. (X x, Y x)) -` (A \<times> B) \<inter> space M \<subseteq> X -` A \<inter> space M" by force
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   477
  show "X -` A \<inter> space M \<in> events"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   478
    using finite_random_variable_measurable[OF X] .
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   479
  have *: "(\<lambda>x. (X x, Y x)) -` (A \<times> B) \<inter> space M =
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   480
    (X -` A \<inter> space M) \<inter> (Y -` B \<inter> space M)" by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   481
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   482
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   483
lemma (in prob_space) joint_distribution_finite_Times_le_snd:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   484
  assumes X: "finite_random_variable MX X" and Y: "finite_random_variable MY Y"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   485
  shows "joint_distribution X Y (A \<times> B) \<le> distribution Y B"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   486
  using assms
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   487
  by (subst joint_distribution_commute)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   488
     (simp add: swap_product joint_distribution_finite_Times_le_fst)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   489
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   490
lemma (in prob_space) finite_distribution_order:
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   491
  fixes MX :: "('c, 'd) measure_space_scheme" and MY :: "('e, 'f) measure_space_scheme"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   492
  assumes "finite_random_variable MX X" "finite_random_variable MY Y"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   493
  shows "r \<le> joint_distribution X Y {(x, y)} \<Longrightarrow> r \<le> distribution X {x}"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   494
    and "r \<le> joint_distribution X Y {(x, y)} \<Longrightarrow> r \<le> distribution Y {y}"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   495
    and "r < joint_distribution X Y {(x, y)} \<Longrightarrow> r < distribution X {x}"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   496
    and "r < joint_distribution X Y {(x, y)} \<Longrightarrow> r < distribution Y {y}"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   497
    and "distribution X {x} = 0 \<Longrightarrow> joint_distribution X Y {(x, y)} = 0"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   498
    and "distribution Y {y} = 0 \<Longrightarrow> joint_distribution X Y {(x, y)} = 0"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   499
  using joint_distribution_finite_Times_le_snd[OF assms, of "{x}" "{y}"]
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   500
  using joint_distribution_finite_Times_le_fst[OF assms, of "{x}" "{y}"]
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   501
  by (auto intro: antisym)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   502
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   503
lemma (in prob_space) setsum_joint_distribution:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   504
  assumes X: "finite_random_variable MX X"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   505
  assumes Y: "random_variable MY Y" "B \<in> sets MY"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   506
  shows "(\<Sum>a\<in>space MX. joint_distribution X Y ({a} \<times> B)) = distribution Y B"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   507
  unfolding distribution_def
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   508
proof (subst finite_measure_finite_Union[symmetric])
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   509
  interpret MX: finite_sigma_algebra MX using X by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   510
  show "finite (space MX)" using MX.finite_space .
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   511
  let "?d i" = "(\<lambda>x. (X x, Y x)) -` ({i} \<times> B) \<inter> space M"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   512
  { fix i assume "i \<in> space MX"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   513
    moreover have "?d i = (X -` {i} \<inter> space M) \<inter> (Y -` B \<inter> space M)" by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   514
    ultimately show "?d i \<in> events"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   515
      using measurable_sets[of X M MX] measurable_sets[of Y M MY B] X Y
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   516
      using MX.sets_eq_Pow by auto }
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   517
  show "disjoint_family_on ?d (space MX)" by (auto simp: disjoint_family_on_def)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   518
  show "\<mu>' (\<Union>i\<in>space MX. ?d i) = \<mu>' (Y -` B \<inter> space M)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   519
    using X[unfolded measurable_def] by (auto intro!: arg_cong[where f=\<mu>'])
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   520
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   521
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   522
lemma (in prob_space) setsum_joint_distribution_singleton:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   523
  assumes X: "finite_random_variable MX X"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   524
  assumes Y: "finite_random_variable MY Y" "b \<in> space MY"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   525
  shows "(\<Sum>a\<in>space MX. joint_distribution X Y {(a, b)}) = distribution Y {b}"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   526
  using setsum_joint_distribution[OF X
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   527
    finite_random_variableD[OF Y(1)]
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   528
    finite_random_variable_imp_sets[OF Y]] by simp
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   529
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   530
locale pair_finite_prob_space = M1: finite_prob_space M1 + M2: finite_prob_space M2 for M1 M2
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   531
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   532
sublocale pair_finite_prob_space \<subseteq> pair_prob_space M1 M2 by default
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   533
sublocale pair_finite_prob_space \<subseteq> pair_finite_space M1 M2  by default
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   534
sublocale pair_finite_prob_space \<subseteq> finite_prob_space P by default
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   535
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   536
lemma (in prob_space) joint_distribution_finite_prob_space:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   537
  assumes X: "finite_random_variable MX X"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   538
  assumes Y: "finite_random_variable MY Y"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   539
  shows "finite_prob_space ((MX \<Otimes>\<^isub>M MY)\<lparr> measure := extreal \<circ> joint_distribution X Y\<rparr>)"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   540
  by (intro distribution_finite_prob_space finite_random_variable_pairI X Y)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   541
36624
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   542
lemma finite_prob_space_eq:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   543
  "finite_prob_space M \<longleftrightarrow> finite_measure_space M \<and> measure M (space M) = 1"
36624
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   544
  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
   545
  by auto
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   546
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   547
lemma (in prob_space) not_empty: "space M \<noteq> {}"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   548
  using prob_space empty_measure' by auto
36624
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   549
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   550
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
   551
  using measure_space_1 sum_over_space by simp
36624
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   552
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   553
lemma (in finite_prob_space) joint_distribution_restriction_fst:
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   554
  "joint_distribution X Y A \<le> distribution X (fst ` A)"
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   555
  unfolding distribution_def
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   556
proof (safe intro!: finite_measure_mono)
36624
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   557
  fix x assume "x \<in> space M" and *: "(X x, Y x) \<in> A"
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   558
  show "x \<in> X -` fst ` A"
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   559
    by (auto intro!: image_eqI[OF _ *])
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   560
qed (simp_all add: sets_eq_Pow)
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   561
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   562
lemma (in finite_prob_space) joint_distribution_restriction_snd:
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   563
  "joint_distribution X Y A \<le> distribution Y (snd ` A)"
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   564
  unfolding distribution_def
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   565
proof (safe intro!: finite_measure_mono)
36624
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   566
  fix x assume "x \<in> space M" and *: "(X x, Y x) \<in> A"
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   567
  show "x \<in> Y -` snd ` A"
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   568
    by (auto intro!: image_eqI[OF _ *])
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   569
qed (simp_all add: sets_eq_Pow)
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   570
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   571
lemma (in finite_prob_space) distribution_order:
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   572
  shows "0 \<le> distribution X x'"
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   573
  and "(distribution X x' \<noteq> 0) \<longleftrightarrow> (0 < distribution X x')"
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   574
  and "r \<le> joint_distribution X Y {(x, y)} \<Longrightarrow> r \<le> distribution X {x}"
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   575
  and "r \<le> joint_distribution X Y {(x, y)} \<Longrightarrow> r \<le> distribution Y {y}"
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   576
  and "r < joint_distribution X Y {(x, y)} \<Longrightarrow> r < distribution X {x}"
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   577
  and "r < joint_distribution X Y {(x, y)} \<Longrightarrow> r < distribution Y {y}"
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   578
  and "distribution X {x} = 0 \<Longrightarrow> joint_distribution X Y {(x, y)} = 0"
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   579
  and "distribution Y {y} = 0 \<Longrightarrow> joint_distribution X Y {(x, y)} = 0"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   580
  using
36624
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   581
    joint_distribution_restriction_fst[of X Y "{(x, y)}"]
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   582
    joint_distribution_restriction_snd[of X Y "{(x, y)}"]
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   583
  by (auto intro: antisym)
36624
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   584
39097
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   585
lemma (in finite_prob_space) distribution_mono:
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   586
  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
   587
  shows "distribution X x \<le> distribution Y y"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   588
  unfolding distribution_def
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   589
  using assms by (auto simp: sets_eq_Pow intro!: finite_measure_mono)
39097
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   590
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   591
lemma (in finite_prob_space) distribution_mono_gt_0:
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   592
  assumes gt_0: "0 < distribution X x"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   593
  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
   594
  shows "0 < distribution Y y"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   595
  by (rule less_le_trans[OF gt_0 distribution_mono]) (rule *)
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   596
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   597
lemma (in finite_prob_space) sum_over_space_distrib:
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   598
  "(\<Sum>x\<in>X`space M. distribution X {x}) = 1"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   599
  unfolding distribution_def prob_space[symmetric] using finite_space
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   600
  by (subst finite_measure_finite_Union[symmetric])
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   601
     (auto simp add: disjoint_family_on_def sets_eq_Pow
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   602
           intro!: arg_cong[where f=\<mu>'])
39097
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   603
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   604
lemma (in finite_prob_space) sum_over_space_real_distribution:
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   605
  "(\<Sum>x\<in>X`space M. distribution X {x}) = 1"
39097
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   606
  unfolding distribution_def prob_space[symmetric] using finite_space
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   607
  by (subst finite_measure_finite_Union[symmetric])
39097
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   608
     (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
   609
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   610
lemma (in finite_prob_space) finite_sum_over_space_eq_1:
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   611
  "(\<Sum>x\<in>space M. prob {x}) = 1"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   612
  using prob_space finite_space
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   613
  by (subst (asm) finite_measure_finite_singleton) auto
39097
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   614
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   615
lemma (in prob_space) distribution_remove_const:
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   616
  shows "joint_distribution X (\<lambda>x. ()) {(x, ())} = distribution X {x}"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   617
  and "joint_distribution (\<lambda>x. ()) X {((), x)} = distribution X {x}"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   618
  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
   619
  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
   620
  and "distribution (\<lambda>x. ()) {()} = 1"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   621
  by (auto intro!: arg_cong[where f=\<mu>'] simp: distribution_def prob_space[symmetric])
35977
30d42bfd0174 Added finite measure space.
hoelzl
parents: 35929
diff changeset
   622
39097
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   623
lemma (in finite_prob_space) setsum_distribution_gen:
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   624
  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
   625
  and "inj_on f (X`space M)"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   626
  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
   627
  unfolding distribution_def assms
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   628
  using finite_space assms
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   629
  by (subst finite_measure_finite_Union[symmetric])
39097
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   630
     (auto simp add: disjoint_family_on_def sets_eq_Pow inj_on_def
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   631
      intro!: arg_cong[where f=prob])
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   632
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   633
lemma (in finite_prob_space) setsum_distribution:
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   634
  "(\<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
   635
  "(\<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
   636
  "(\<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
   637
  "(\<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
   638
  "(\<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
   639
  by (auto intro!: inj_onI setsum_distribution_gen)
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   640
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   641
lemma (in finite_prob_space) uniform_prob:
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   642
  assumes "x \<in> space M"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   643
  assumes "\<And> x y. \<lbrakk>x \<in> space M ; y \<in> space M\<rbrakk> \<Longrightarrow> prob {x} = prob {y}"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   644
  shows "prob {x} = 1 / card (space M)"
39097
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   645
proof -
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   646
  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
   647
    using assms(2)[OF _ `x \<in> space M`] by blast
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   648
  have "1 = prob (space M)"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   649
    using prob_space by auto
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   650
  also have "\<dots> = (\<Sum> x \<in> space M. prob {x})"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   651
    using finite_measure_finite_Union[of "space M" "\<lambda> x. {x}", simplified]
39097
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   652
      sets_eq_Pow inj_singleton[unfolded inj_on_def, rule_format]
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   653
      finite_space unfolding disjoint_family_on_def  prob_space[symmetric]
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   654
    by (auto simp add:setsum_restrict_set)
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   655
  also have "\<dots> = (\<Sum> y \<in> space M. prob {x})"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   656
    using prob_x by auto
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   657
  also have "\<dots> = real_of_nat (card (space M)) * prob {x}" by simp
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   658
  finally have one: "1 = real (card (space M)) * prob {x}"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   659
    using real_eq_of_nat by auto
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   660
  hence two: "real (card (space M)) \<noteq> 0" by fastsimp 
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   661
  from one have three: "prob {x} \<noteq> 0" by fastsimp
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   662
  thus ?thesis using one two three divide_cancel_right
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   663
    by (auto simp:field_simps)
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39091
diff changeset
   664
qed
35977
30d42bfd0174 Added finite measure space.
hoelzl
parents: 35929
diff changeset
   665
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39091
diff changeset
   666
lemma (in prob_space) prob_space_subalgebra:
41545
9c869baf1c66 tuned formalization of subalgebra
hoelzl
parents: 41544
diff changeset
   667
  assumes "sigma_algebra N" "sets N \<subseteq> sets M" "space N = space M"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   668
    and "\<And>A. A \<in> sets N \<Longrightarrow> measure N A = \<mu> A"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   669
  shows "prob_space N"
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39091
diff changeset
   670
proof -
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   671
  interpret N: measure_space N
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   672
    by (rule measure_space_subalgebra[OF assms])
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39091
diff changeset
   673
  show ?thesis
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   674
  proof qed (insert assms(4)[OF N.top], simp add: assms measure_space_1)
35977
30d42bfd0174 Added finite measure space.
hoelzl
parents: 35929
diff changeset
   675
qed
30d42bfd0174 Added finite measure space.
hoelzl
parents: 35929
diff changeset
   676
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39091
diff changeset
   677
lemma (in prob_space) prob_space_of_restricted_space:
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   678
  assumes "\<mu> A \<noteq> 0" "A \<in> sets M"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   679
  shows "prob_space (restricted_space A \<lparr>measure := \<lambda>S. \<mu> S / \<mu> A\<rparr>)"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   680
    (is "prob_space ?P")
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   681
proof -
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   682
  interpret A: measure_space "restricted_space A"
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39091
diff changeset
   683
    using `A \<in> sets M` by (rule restricted_measure_space)
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   684
  interpret A': sigma_algebra ?P
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   685
    by (rule A.sigma_algebra_cong) auto
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   686
  show "prob_space ?P"
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39091
diff changeset
   687
  proof
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   688
    show "measure ?P (space ?P) = 1"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   689
      using real_measure[OF `A \<in> events`] `\<mu> A \<noteq> 0` by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   690
    show "positive ?P (measure ?P)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   691
    proof (simp add: positive_def, safe)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   692
      show "0 / \<mu> A = 0" using `\<mu> A \<noteq> 0` by (cases "\<mu> A") (auto simp: zero_extreal_def)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   693
      fix B assume "B \<in> events"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   694
      with real_measure[of "A \<inter> B"] real_measure[OF `A \<in> events`] `A \<in> sets M`
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   695
      show "0 \<le> \<mu> (A \<inter> B) / \<mu> A" by (auto simp: Int)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   696
    qed
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   697
    show "countably_additive ?P (measure ?P)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   698
    proof (simp add: countably_additive_def, safe)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   699
      fix B and F :: "nat \<Rightarrow> 'a set"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   700
      assume F: "range F \<subseteq> op \<inter> A ` events" "disjoint_family F"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   701
      { fix i
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   702
        from F have "F i \<in> op \<inter> A ` events" by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   703
        with `A \<in> events` have "F i \<in> events" by auto }
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   704
      moreover then have "range F \<subseteq> events" by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   705
      moreover have "\<And>S. \<mu> S / \<mu> A = inverse (\<mu> A) * \<mu> S"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   706
        by (simp add: mult_commute divide_extreal_def)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   707
      moreover have "0 \<le> inverse (\<mu> A)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   708
        using real_measure[OF `A \<in> events`] by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   709
      ultimately show "(\<Sum>i. \<mu> (F i) / \<mu> A) = \<mu> (\<Union>i. F i) / \<mu> A"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   710
        using measure_countably_additive[of F] F
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   711
        by (auto simp: suminf_cmult_extreal)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   712
    qed
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39091
diff changeset
   713
  qed
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39091
diff changeset
   714
qed
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39091
diff changeset
   715
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39091
diff changeset
   716
lemma finite_prob_spaceI:
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   717
  assumes "finite (space M)" "sets M = Pow(space M)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   718
    and "measure M (space M) = 1" "measure M {} = 0" "\<And>A. A \<subseteq> space M \<Longrightarrow> 0 \<le> measure M A"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   719
    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"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   720
  shows "finite_prob_space M"
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39091
diff changeset
   721
  unfolding finite_prob_space_eq
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39091
diff changeset
   722
proof
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   723
  show "finite_measure_space M" using assms
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   724
    by (auto intro!: finite_measure_spaceI)
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   725
  show "measure M (space M) = 1" by fact
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39091
diff changeset
   726
qed
36624
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   727
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   728
lemma (in finite_prob_space) finite_measure_space:
39097
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   729
  fixes X :: "'a \<Rightarrow> 'x"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   730
  shows "finite_measure_space \<lparr>space = X ` space M, sets = Pow (X ` space M), measure = extreal \<circ> distribution X\<rparr>"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   731
    (is "finite_measure_space ?S")
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39091
diff changeset
   732
proof (rule finite_measure_spaceI, simp_all)
36624
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   733
  show "finite (X ` space M)" using finite_space by simp
39097
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   734
next
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   735
  fix A B :: "'x set" assume "A \<inter> B = {}"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   736
  then show "distribution X (A \<union> B) = distribution X A + distribution X B"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   737
    unfolding distribution_def
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   738
    by (subst finite_measure_Union[symmetric])
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   739
       (auto intro!: arg_cong[where f=\<mu>'] simp: sets_eq_Pow)
36624
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   740
qed
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   741
39097
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   742
lemma (in finite_prob_space) finite_prob_space_of_images:
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   743
  "finite_prob_space \<lparr> space = X ` space M, sets = Pow (X ` space M), measure = extreal \<circ> distribution X \<rparr>"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   744
  by (simp add: finite_prob_space_eq finite_measure_space measure_space_1 one_extreal_def)
39097
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   745
39096
hoelzl
parents: 39092
diff changeset
   746
lemma (in finite_prob_space) finite_product_measure_space:
39097
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   747
  fixes X :: "'a \<Rightarrow> 'x" and Y :: "'a \<Rightarrow> 'y"
39096
hoelzl
parents: 39092
diff changeset
   748
  assumes "finite s1" "finite s2"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   749
  shows "finite_measure_space \<lparr> space = s1 \<times> s2, sets = Pow (s1 \<times> s2), measure = extreal \<circ> joint_distribution X Y\<rparr>"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   750
    (is "finite_measure_space ?M")
39097
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   751
proof (rule finite_measure_spaceI, simp_all)
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   752
  show "finite (s1 \<times> s2)"
39096
hoelzl
parents: 39092
diff changeset
   753
    using assms by auto
39097
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   754
next
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   755
  fix A B :: "('x*'y) set" assume "A \<inter> B = {}"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   756
  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
   757
    unfolding distribution_def
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   758
    by (subst finite_measure_Union[symmetric])
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   759
       (auto intro!: arg_cong[where f=\<mu>'] simp: sets_eq_Pow)
39096
hoelzl
parents: 39092
diff changeset
   760
qed
hoelzl
parents: 39092
diff changeset
   761
39097
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   762
lemma (in finite_prob_space) finite_product_measure_space_of_images:
39096
hoelzl
parents: 39092
diff changeset
   763
  shows "finite_measure_space \<lparr> space = X ` space M \<times> Y ` space M,
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   764
                                sets = Pow (X ` space M \<times> Y ` space M),
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   765
                                measure = extreal \<circ> joint_distribution X Y \<rparr>"
39096
hoelzl
parents: 39092
diff changeset
   766
  using finite_space by (auto intro!: finite_product_measure_space)
hoelzl
parents: 39092
diff changeset
   767
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   768
lemma (in finite_prob_space) finite_product_prob_space_of_images:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   769
  "finite_prob_space \<lparr> space = X ` space M \<times> Y ` space M, sets = Pow (X ` space M \<times> Y ` space M),
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   770
                       measure = extreal \<circ> joint_distribution X Y \<rparr>"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   771
  (is "finite_prob_space ?S")
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   772
proof (simp add: finite_prob_space_eq finite_product_measure_space_of_images one_extreal_def)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   773
  have "X -` X ` space M \<inter> Y -` Y ` space M \<inter> space M = space M" by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   774
  thus "joint_distribution X Y (X ` space M \<times> Y ` space M) = 1"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   775
    by (simp add: distribution_def prob_space vimage_Times comp_def measure_space_1)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   776
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   777
39085
8b7c009da23c added definition of conditional expectation
hoelzl
parents: 39084
diff changeset
   778
section "Conditional Expectation and Probability"
8b7c009da23c added definition of conditional expectation
hoelzl
parents: 39084
diff changeset
   779
8b7c009da23c added definition of conditional expectation
hoelzl
parents: 39084
diff changeset
   780
lemma (in prob_space) conditional_expectation_exists:
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   781
  fixes X :: "'a \<Rightarrow> extreal" and N :: "('a, 'b) measure_space_scheme"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   782
  assumes borel: "X \<in> borel_measurable M" "AE x. 0 \<le> X x"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   783
  and N: "sigma_algebra N" "sets N \<subseteq> sets M" "space N = space M" "\<And>A. measure N A = \<mu> A"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   784
  shows "\<exists>Y\<in>borel_measurable N. (\<forall>x. 0 \<le> Y x) \<and> (\<forall>C\<in>sets N.
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   785
      (\<integral>\<^isup>+x. Y x * indicator C x \<partial>M) = (\<integral>\<^isup>+x. X x * indicator C x \<partial>M))"
39083
e46acc0ea1fe introduced integration on subalgebras
hoelzl
parents: 38656
diff changeset
   786
proof -
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   787
  note N(4)[simp]
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   788
  interpret P: prob_space N
41545
9c869baf1c66 tuned formalization of subalgebra
hoelzl
parents: 41544
diff changeset
   789
    using prob_space_subalgebra[OF N] .
39083
e46acc0ea1fe introduced integration on subalgebras
hoelzl
parents: 38656
diff changeset
   790
e46acc0ea1fe introduced integration on subalgebras
hoelzl
parents: 38656
diff changeset
   791
  let "?f A" = "\<lambda>x. X x * indicator A x"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   792
  let "?Q A" = "integral\<^isup>P M (?f A)"
39083
e46acc0ea1fe introduced integration on subalgebras
hoelzl
parents: 38656
diff changeset
   793
e46acc0ea1fe introduced integration on subalgebras
hoelzl
parents: 38656
diff changeset
   794
  from measure_space_density[OF borel]
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   795
  have Q: "measure_space (N\<lparr> measure := ?Q \<rparr>)"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   796
    apply (rule measure_space.measure_space_subalgebra[of "M\<lparr> measure := ?Q \<rparr>"])
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   797
    using N by (auto intro!: P.sigma_algebra_cong)
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   798
  then interpret Q: measure_space "N\<lparr> measure := ?Q \<rparr>" .
39083
e46acc0ea1fe introduced integration on subalgebras
hoelzl
parents: 38656
diff changeset
   799
e46acc0ea1fe introduced integration on subalgebras
hoelzl
parents: 38656
diff changeset
   800
  have "P.absolutely_continuous ?Q"
e46acc0ea1fe introduced integration on subalgebras
hoelzl
parents: 38656
diff changeset
   801
    unfolding P.absolutely_continuous_def
41545
9c869baf1c66 tuned formalization of subalgebra
hoelzl
parents: 41544
diff changeset
   802
  proof safe
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   803
    fix A assume "A \<in> sets N" "P.\<mu> A = 0"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   804
    then have f_borel: "?f A \<in> borel_measurable M" "AE x. x \<notin> A"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   805
      using borel N by (auto intro!: borel_measurable_indicator AE_not_in)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   806
    then show "?Q A = 0"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   807
      by (auto simp add: positive_integral_0_iff_AE)
39083
e46acc0ea1fe introduced integration on subalgebras
hoelzl
parents: 38656
diff changeset
   808
  qed
e46acc0ea1fe introduced integration on subalgebras
hoelzl
parents: 38656
diff changeset
   809
  from P.Radon_Nikodym[OF Q this]
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   810
  obtain Y where Y: "Y \<in> borel_measurable N" "\<And>x. 0 \<le> Y x"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   811
    "\<And>A. A \<in> sets N \<Longrightarrow> ?Q A =(\<integral>\<^isup>+x. Y x * indicator A x \<partial>N)"
39083
e46acc0ea1fe introduced integration on subalgebras
hoelzl
parents: 38656
diff changeset
   812
    by blast
41545
9c869baf1c66 tuned formalization of subalgebra
hoelzl
parents: 41544
diff changeset
   813
  with N(2) show ?thesis
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   814
    by (auto intro!: bexI[OF _ Y(1)] simp: positive_integral_subalgebra[OF _ _ N(2,3,4,1)])
39083
e46acc0ea1fe introduced integration on subalgebras
hoelzl
parents: 38656
diff changeset
   815
qed
e46acc0ea1fe introduced integration on subalgebras
hoelzl
parents: 38656
diff changeset
   816
39085
8b7c009da23c added definition of conditional expectation
hoelzl
parents: 39084
diff changeset
   817
definition (in prob_space)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   818
  "conditional_expectation N X = (SOME Y. Y\<in>borel_measurable N \<and> (\<forall>x. 0 \<le> Y x)
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   819
    \<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)))"
39085
8b7c009da23c added definition of conditional expectation
hoelzl
parents: 39084
diff changeset
   820
8b7c009da23c added definition of conditional expectation
hoelzl
parents: 39084
diff changeset
   821
abbreviation (in prob_space)
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39091
diff changeset
   822
  "conditional_prob N A \<equiv> conditional_expectation N (indicator A)"
39085
8b7c009da23c added definition of conditional expectation
hoelzl
parents: 39084
diff changeset
   823
8b7c009da23c added definition of conditional expectation
hoelzl
parents: 39084
diff changeset
   824
lemma (in prob_space)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   825
  fixes X :: "'a \<Rightarrow> extreal" and N :: "('a, 'b) measure_space_scheme"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   826
  assumes borel: "X \<in> borel_measurable M" "AE x. 0 \<le> X x"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   827
  and N: "sigma_algebra N" "sets N \<subseteq> sets M" "space N = space M" "\<And>A. measure N A = \<mu> A"
39085
8b7c009da23c added definition of conditional expectation
hoelzl
parents: 39084
diff changeset
   828
  shows borel_measurable_conditional_expectation:
41545
9c869baf1c66 tuned formalization of subalgebra
hoelzl
parents: 41544
diff changeset
   829
    "conditional_expectation N X \<in> borel_measurable N"
9c869baf1c66 tuned formalization of subalgebra
hoelzl
parents: 41544
diff changeset
   830
  and conditional_expectation: "\<And>C. C \<in> sets N \<Longrightarrow>
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   831
      (\<integral>\<^isup>+x. conditional_expectation N X x * indicator C x \<partial>M) =
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   832
      (\<integral>\<^isup>+x. X x * indicator C x \<partial>M)"
41545
9c869baf1c66 tuned formalization of subalgebra
hoelzl
parents: 41544
diff changeset
   833
   (is "\<And>C. C \<in> sets N \<Longrightarrow> ?eq C")
39085
8b7c009da23c added definition of conditional expectation
hoelzl
parents: 39084
diff changeset
   834
proof -
8b7c009da23c added definition of conditional expectation
hoelzl
parents: 39084
diff changeset
   835
  note CE = conditional_expectation_exists[OF assms, unfolded Bex_def]
41545
9c869baf1c66 tuned formalization of subalgebra
hoelzl
parents: 41544
diff changeset
   836
  then show "conditional_expectation N X \<in> borel_measurable N"
39085
8b7c009da23c added definition of conditional expectation
hoelzl
parents: 39084
diff changeset
   837
    unfolding conditional_expectation_def by (rule someI2_ex) blast
8b7c009da23c added definition of conditional expectation
hoelzl
parents: 39084
diff changeset
   838
41545
9c869baf1c66 tuned formalization of subalgebra
hoelzl
parents: 41544
diff changeset
   839
  from CE show "\<And>C. C \<in> sets N \<Longrightarrow> ?eq C"
39085
8b7c009da23c added definition of conditional expectation
hoelzl
parents: 39084
diff changeset
   840
    unfolding conditional_expectation_def by (rule someI2_ex) blast
8b7c009da23c added definition of conditional expectation
hoelzl
parents: 39084
diff changeset
   841
qed
8b7c009da23c added definition of conditional expectation
hoelzl
parents: 39084
diff changeset
   842
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   843
lemma (in sigma_algebra) factorize_measurable_function_pos:
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   844
  fixes Z :: "'a \<Rightarrow> extreal" and Y :: "'a \<Rightarrow> 'c"
39091
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   845
  assumes "sigma_algebra M'" and "Y \<in> measurable M M'" "Z \<in> borel_measurable M"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   846
  assumes Z: "Z \<in> borel_measurable (sigma_algebra.vimage_algebra M' (space M) Y)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   847
  shows "\<exists>g\<in>borel_measurable M'. \<forall>x\<in>space M. max 0 (Z x) = g (Y x)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   848
proof -
39091
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   849
  interpret M': sigma_algebra M' by fact
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   850
  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
   851
  from M'.sigma_algebra_vimage[OF this]
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   852
  interpret va: sigma_algebra "M'.vimage_algebra (space M) Y" .
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   853
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   854
  from va.borel_measurable_implies_simple_function_sequence'[OF Z] guess f . note f = this
39091
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   855
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   856
  have "\<forall>i. \<exists>g. simple_function M' g \<and> (\<forall>x\<in>space M. f i x = g (Y x))"
39091
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   857
  proof
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   858
    fix i
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   859
    from f(1)[of i] have "finite (f i`space M)" and B_ex:
39091
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   860
      "\<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"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   861
      unfolding simple_function_def by auto
39091
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   862
    from B_ex[THEN bchoice] guess B .. note B = this
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   863
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   864
    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
   865
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   866
    show "\<exists>g. simple_function M' g \<and> (\<forall>x\<in>space M. f i x = g (Y x))"
39091
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   867
    proof (intro exI[of _ ?g] conjI ballI)
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   868
      show "simple_function M' ?g" using B by auto
39091
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   869
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   870
      fix x assume "x \<in> space M"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   871
      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)"
39091
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   872
        unfolding indicator_def using B by auto
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   873
      then show "f i x = ?g (Y x)" using `x \<in> space M` f(1)[of i]
39091
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   874
        by (subst va.simple_function_indicator_representation) auto
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   875
    qed
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   876
  qed
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   877
  from choice[OF this] guess g .. note g = this
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   878
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   879
  show ?thesis
39091
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   880
  proof (intro ballI bexI)
41097
a1abfa4e2b44 use SUPR_ and INFI_apply instead of SUPR_, INFI_fun_expand
hoelzl
parents: 41095
diff changeset
   881
    show "(\<lambda>x. SUP i. g i x) \<in> borel_measurable M'"
39091
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   882
      using g by (auto intro: M'.borel_measurable_simple_function)
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   883
    fix x assume "x \<in> space M"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   884
    have "max 0 (Z x) = (SUP i. f i x)" using f by simp
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   885
    also have "\<dots> = (SUP i. g i (Y x))"
39091
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   886
      using g `x \<in> space M` by simp
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   887
    finally show "max 0 (Z x) = (SUP i. g i (Y x))" .
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   888
  qed
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   889
qed
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   890
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   891
lemma extreal_0_le_iff_le_0[simp]:
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   892
  fixes a :: extreal shows "0 \<le> -a \<longleftrightarrow> a \<le> 0"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   893
  by (cases rule: extreal2_cases[of a]) auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   894
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   895
lemma (in sigma_algebra) factorize_measurable_function:
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   896
  fixes Z :: "'a \<Rightarrow> extreal" and Y :: "'a \<Rightarrow> 'c"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   897
  assumes "sigma_algebra M'" and "Y \<in> measurable M M'" "Z \<in> borel_measurable M"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   898
  shows "Z \<in> borel_measurable (sigma_algebra.vimage_algebra M' (space M) Y)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   899
    \<longleftrightarrow> (\<exists>g\<in>borel_measurable M'. \<forall>x\<in>space M. Z x = g (Y x))"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   900
proof safe
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   901
  interpret M': sigma_algebra M' by fact
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   902
  have Y: "Y \<in> space M \<rightarrow> space M'" using assms unfolding measurable_def by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   903
  from M'.sigma_algebra_vimage[OF this]
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   904
  interpret va: sigma_algebra "M'.vimage_algebra (space M) Y" .
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   905
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   906
  { fix g :: "'c \<Rightarrow> extreal" assume "g \<in> borel_measurable M'"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   907
    with M'.measurable_vimage_algebra[OF Y]
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   908
    have "g \<circ> Y \<in> borel_measurable (M'.vimage_algebra (space M) Y)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   909
      by (rule measurable_comp)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   910
    moreover assume "\<forall>x\<in>space M. Z x = g (Y x)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   911
    then have "Z \<in> borel_measurable (M'.vimage_algebra (space M) Y) \<longleftrightarrow>
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   912
       g \<circ> Y \<in> borel_measurable (M'.vimage_algebra (space M) Y)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   913
       by (auto intro!: measurable_cong)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   914
    ultimately show "Z \<in> borel_measurable (M'.vimage_algebra (space M) Y)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   915
      by simp }
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   916
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   917
  assume Z: "Z \<in> borel_measurable (M'.vimage_algebra (space M) Y)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   918
  with assms have "(\<lambda>x. - Z x) \<in> borel_measurable M"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   919
    "(\<lambda>x. - Z x) \<in> borel_measurable (M'.vimage_algebra (space M) Y)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   920
    by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   921
  from factorize_measurable_function_pos[OF assms(1,2) this] guess n .. note n = this
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   922
  from factorize_measurable_function_pos[OF assms Z] guess p .. note p = this
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   923
  let "?g x" = "p x - n x"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   924
  show "\<exists>g\<in>borel_measurable M'. \<forall>x\<in>space M. Z x = g (Y x)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   925
  proof (intro bexI ballI)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   926
    show "?g \<in> borel_measurable M'" using p n by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   927
    fix x assume "x \<in> space M"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   928
    then have "p (Y x) = max 0 (Z x)" "n (Y x) = max 0 (- Z x)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   929
      using p n by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   930
    then show "Z x = ?g (Y x)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   931
      by (auto split: split_max)
39091
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   932
  qed
11314c196e11 factorizable measurable functions
hoelzl
parents: 39090
diff changeset
   933
qed
39090
a2d38b8b693e Introduced sigma algebra generated by function preimages.
hoelzl
parents: 39089
diff changeset
   934
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   935
end