src/HOL/Probability/Probability_Measure.thy
author wenzelm
Sat, 07 Apr 2012 16:41:59 +0200
changeset 47389 e8552cba702d
parent 46905 6b1c0a80a57a
child 47694 05663f75964c
permissions -rw-r--r--
explicit checks stable_finished_theory/stable_command allow parallel asynchronous command transactions; tuned;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
42148
d596e7bb251f rename Probability_Space to Probability_Measure
hoelzl
parents: 42146
diff changeset
     1
(*  Title:      HOL/Probability/Probability_Measure.thy
42067
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
42148
d596e7bb251f rename Probability_Space to Probability_Measure
hoelzl
parents: 42146
diff changeset
     6
header {*Probability measure*}
42067
66c8281349ec standardized headers
hoelzl
parents: 41981
diff changeset
     7
42148
d596e7bb251f rename Probability_Space to Probability_Measure
hoelzl
parents: 42146
diff changeset
     8
theory Probability_Measure
43556
0d78c8d31d0d move conditional expectation to its own theory file
hoelzl
parents: 43340
diff changeset
     9
imports Lebesgue_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
45777
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
    12
locale prob_space = finite_measure +
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
    13
  assumes measure_space_1: "measure M (space M) = 1"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
    14
45777
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
    15
lemma prob_spaceI[Pure.intro!]:
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
    16
  assumes "measure_space M"
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
    17
  assumes *: "measure M (space M) = 1"
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
    18
  shows "prob_space M"
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
    19
proof -
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
    20
  interpret finite_measure M
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
    21
  proof
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
    22
    show "measure_space M" by fact
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
    23
    show "measure M (space M) \<noteq> \<infinity>" using * by simp 
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
    24
  qed
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
    25
  show "prob_space M" by default fact
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
    26
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
    27
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
    28
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
    29
abbreviation (in prob_space) "prob \<equiv> \<mu>'"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
    30
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
    31
abbreviation (in prob_space) "expectation \<equiv> integral\<^isup>L M"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
    32
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
    33
definition (in prob_space)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    34
  "distribution X A = \<mu>' (X -` A \<inter> space M)"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
    35
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
    36
abbreviation (in prob_space)
36624
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
    37
  "joint_distribution X Y \<equiv> distribution (\<lambda>x. (X x, Y x))"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
    38
43339
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
    39
lemma (in prob_space) prob_space_cong:
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
    40
  assumes "\<And>A. A \<in> sets M \<Longrightarrow> measure N A = \<mu> A" "space N = space M" "sets N = sets M"
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
    41
  shows "prob_space N"
45777
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
    42
proof
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
    43
  show "measure_space N" by (intro measure_space_cong assms)
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
    44
  show "measure N (space N) = 1"
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
    45
    using measure_space_1 assms by simp
43339
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
    46
qed
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
    47
39097
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
    48
lemma (in prob_space) distribution_cong:
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
    49
  assumes "\<And>x. x \<in> space M \<Longrightarrow> X x = Y x"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
    50
  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
    51
  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
    52
  using assms by (auto intro!: arg_cong[where f="\<mu>'"])
39097
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
    53
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
    54
lemma (in prob_space) joint_distribution_cong:
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
    55
  assumes "\<And>x. x \<in> space M \<Longrightarrow> X x = X' x"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
    56
  assumes "\<And>x. x \<in> space M \<Longrightarrow> Y x = Y' x"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
    57
  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
    58
  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
    59
  using assms by (auto intro!: arg_cong[where f="\<mu>'"])
39097
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
    60
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
    61
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
    62
  "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
    63
  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
    64
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
    65
lemma (in prob_space) prob_space: "prob (space M) = 1"
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43658
diff changeset
    66
  using measure_space_1 unfolding \<mu>'_def by (simp add: one_ereal_def)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    67
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    68
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
    69
  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
    70
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    71
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
    72
  "0 \<le> distribution X A" unfolding distribution_def by auto
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
    73
43339
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
    74
lemma (in prob_space) not_zero_less_distribution[simp]:
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
    75
  "(\<not> 0 < distribution X A) \<longleftrightarrow> distribution X A = 0"
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
    76
  using distribution_positive[of X A] by arith
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
    77
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    78
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
    79
    "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
    80
  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
    81
45712
852597248663 moved theorems about distribution to the definition; removed oopsed-lemma
hoelzl
parents: 45711
diff changeset
    82
lemma (in prob_space) distribution_unit[simp]: "distribution (\<lambda>x. ()) {()} = 1"
852597248663 moved theorems about distribution to the definition; removed oopsed-lemma
hoelzl
parents: 45711
diff changeset
    83
  unfolding distribution_def using prob_space by auto
852597248663 moved theorems about distribution to the definition; removed oopsed-lemma
hoelzl
parents: 45711
diff changeset
    84
852597248663 moved theorems about distribution to the definition; removed oopsed-lemma
hoelzl
parents: 45711
diff changeset
    85
lemma (in prob_space) joint_distribution_unit[simp]: "distribution (\<lambda>x. (X x, ())) {(a, ())} = distribution X {a}"
852597248663 moved theorems about distribution to the definition; removed oopsed-lemma
hoelzl
parents: 45711
diff changeset
    86
  unfolding distribution_def by (auto intro!: arg_cong[where f=\<mu>'])
852597248663 moved theorems about distribution to the definition; removed oopsed-lemma
hoelzl
parents: 45711
diff changeset
    87
43339
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
    88
lemma (in prob_space) not_empty: "space M \<noteq> {}"
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
    89
  using prob_space empty_measure' by auto
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
    90
42950
6e5c2a3c69da move lemmas to Extended_Reals and Extended_Real_Limits
hoelzl
parents: 42902
diff changeset
    91
lemma (in prob_space) measure_le_1: "X \<in> sets M \<Longrightarrow> \<mu> X \<le> 1"
6e5c2a3c69da move lemmas to Extended_Reals and Extended_Real_Limits
hoelzl
parents: 42902
diff changeset
    92
  unfolding measure_space_1[symmetric]
6e5c2a3c69da move lemmas to Extended_Reals and Extended_Real_Limits
hoelzl
parents: 42902
diff changeset
    93
  using sets_into_space
6e5c2a3c69da move lemmas to Extended_Reals and Extended_Real_Limits
hoelzl
parents: 42902
diff changeset
    94
  by (intro measure_mono) auto
6e5c2a3c69da move lemmas to Extended_Reals and Extended_Real_Limits
hoelzl
parents: 42902
diff changeset
    95
43339
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
    96
lemma (in prob_space) AE_I_eq_1:
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
    97
  assumes "\<mu> {x\<in>space M. P x} = 1" "{x\<in>space M. P x} \<in> sets M"
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
    98
  shows "AE x. P x"
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
    99
proof (rule AE_I)
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   100
  show "\<mu> (space M - {x \<in> space M. P x}) = 0"
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   101
    using assms measure_space_1 by (simp add: measure_compl)
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   102
qed (insert assms, auto)
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   103
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   104
lemma (in prob_space) distribution_1:
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   105
  "distribution X A \<le> 1"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   106
  unfolding distribution_def by simp
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   107
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   108
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
   109
  assumes A: "A \<in> events"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   110
  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
   111
  using finite_measure_compl[OF A] by (simp add: prob_space)
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   112
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   113
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
   114
  by (auto intro!: finite_measure_mono simp: increasing_def)
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   115
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   116
lemma (in prob_space) prob_zero_union:
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   117
  assumes "s \<in> events" "t \<in> events" "prob t = 0"
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   118
  shows "prob (s \<union> t) = prob s"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   119
using assms
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   120
proof -
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   121
  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
   122
    using finite_measure_subadditive[of s t] assms by auto
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   123
  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
   124
    using assms by (blast intro: finite_measure_mono)
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   125
  ultimately show ?thesis by simp
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   126
qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   127
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   128
lemma (in prob_space) prob_eq_compl:
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   129
  assumes "s \<in> events" "t \<in> events"
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   130
  assumes "prob (space M - s) = prob (space M - t)"
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   131
  shows "prob s = prob t"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   132
  using assms prob_compl by auto
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   133
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   134
lemma (in prob_space) prob_one_inter:
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   135
  assumes events:"s \<in> events" "t \<in> events"
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   136
  assumes "prob t = 1"
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   137
  shows "prob (s \<inter> t) = prob s"
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   138
proof -
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   139
  have "prob ((space M - s) \<union> (space M - t)) = prob (space M - s)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   140
    using events assms  prob_compl[of "t"] by (auto intro!: prob_zero_union)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   141
  also have "(space M - s) \<union> (space M - t) = space M - (s \<inter> t)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   142
    by blast
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   143
  finally show "prob (s \<inter> t) = prob s"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   144
    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
   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_eq_bigunion_image:
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   148
  assumes "range f \<subseteq> events" "range g \<subseteq> events"
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   149
  assumes "disjoint_family f" "disjoint_family g"
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   150
  assumes "\<And> n :: nat. prob (f n) = prob (g n)"
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   151
  shows "(prob (\<Union> i. f i) = prob (\<Union> i. g i))"
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   152
using assms
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   153
proof -
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   154
  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
   155
    by (rule finite_measure_UNION[OF assms(1,3)])
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   156
  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
   157
    by (rule finite_measure_UNION[OF assms(2,4)])
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   158
  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
   159
qed
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) prob_countably_zero:
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   162
  assumes "range c \<subseteq> events"
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   163
  assumes "\<And> i. prob (c i) = 0"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   164
  shows "prob (\<Union> i :: nat. c i) = 0"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   165
proof (rule antisym)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   166
  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
   167
    using finite_measure_countably_subadditive[OF assms(1)]
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   168
    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
   169
qed simp
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   170
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   171
lemma (in prob_space) prob_equiprobable_finite_unions:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   172
  assumes "s \<in> events"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   173
  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
   174
  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
   175
  shows "prob s = real (card s) * prob {SOME x. x \<in> s}"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   176
proof (cases "s = {}")
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   177
  case False hence "\<exists> x. x \<in> s" by blast
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   178
  from someI_ex[OF this] assms
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   179
  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
   180
  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
   181
    using finite_measure_finite_singleton[OF s_finite] by simp
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   182
  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
   183
  also have "\<dots> = real (card s) * prob {(SOME x. x \<in> s)}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   184
    using setsum_constant assms by (simp add: real_eq_of_nat)
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   185
  finally show ?thesis by simp
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   186
qed simp
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   187
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   188
lemma (in prob_space) prob_real_sum_image_fn:
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   189
  assumes "e \<in> events"
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   190
  assumes "\<And> x. x \<in> s \<Longrightarrow> e \<inter> f x \<in> events"
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   191
  assumes "finite s"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   192
  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
   193
  assumes upper: "space M \<subseteq> (\<Union> i \<in> s. f i)"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   194
  shows "prob e = (\<Sum> x \<in> s. prob (e \<inter> f x))"
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   195
proof -
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   196
  have e: "e = (\<Union> i \<in> s. e \<inter> f i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   197
    using `e \<in> events` sets_into_space upper by blast
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   198
  hence "prob e = prob (\<Union> i \<in> s. e \<inter> f i)" by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   199
  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
   200
  proof (rule finite_measure_finite_Union)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   201
    show "finite s" by fact
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   202
    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
   203
    show "disjoint_family_on (\<lambda>i. e \<inter> f i) s"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   204
      using disjoint by (auto simp: disjoint_family_on_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   205
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   206
  finally show ?thesis .
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   207
qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   208
42199
aded34119213 add prob_space_vimage
hoelzl
parents: 42148
diff changeset
   209
lemma (in prob_space) prob_space_vimage:
aded34119213 add prob_space_vimage
hoelzl
parents: 42148
diff changeset
   210
  assumes S: "sigma_algebra S"
aded34119213 add prob_space_vimage
hoelzl
parents: 42148
diff changeset
   211
  assumes T: "T \<in> measure_preserving M S"
aded34119213 add prob_space_vimage
hoelzl
parents: 42148
diff changeset
   212
  shows "prob_space S"
45777
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
   213
proof
42199
aded34119213 add prob_space_vimage
hoelzl
parents: 42148
diff changeset
   214
  interpret S: measure_space S
aded34119213 add prob_space_vimage
hoelzl
parents: 42148
diff changeset
   215
    using S and T by (rule measure_space_vimage)
45777
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
   216
  show "measure_space S" ..
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
   217
  
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
   218
  from T[THEN measure_preservingD2]
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
   219
  have "T -` space S \<inter> space M = space M"
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
   220
    by (auto simp: measurable_def)
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
   221
  with T[THEN measure_preservingD, of "space S", symmetric]
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
   222
  show  "measure S (space S) = 1"
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
   223
    using measure_space_1 by simp
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   224
qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   225
42988
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42981
diff changeset
   226
lemma prob_space_unique_Int_stable:
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42981
diff changeset
   227
  fixes E :: "('a, 'b) algebra_scheme" and A :: "nat \<Rightarrow> 'a set"
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42981
diff changeset
   228
  assumes E: "Int_stable E" "space E \<in> sets E"
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42981
diff changeset
   229
  and M: "prob_space M" "space M = space E" "sets M = sets (sigma E)"
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42981
diff changeset
   230
  and N: "prob_space N" "space N = space E" "sets N = sets (sigma E)"
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42981
diff changeset
   231
  and eq: "\<And>X. X \<in> sets E \<Longrightarrow> finite_measure.\<mu>' M X = finite_measure.\<mu>' N X"
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42981
diff changeset
   232
  assumes "X \<in> sets (sigma E)"
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42981
diff changeset
   233
  shows "finite_measure.\<mu>' M X = finite_measure.\<mu>' N X"
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42981
diff changeset
   234
proof -
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42981
diff changeset
   235
  interpret M!: prob_space M by fact
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42981
diff changeset
   236
  interpret N!: prob_space N by fact
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42981
diff changeset
   237
  have "measure M X = measure N X"
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42981
diff changeset
   238
  proof (rule measure_unique_Int_stable[OF `Int_stable E`])
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42981
diff changeset
   239
    show "range (\<lambda>i. space M) \<subseteq> sets E" "incseq (\<lambda>i. space M)" "(\<Union>i. space M) = space E"
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42981
diff changeset
   240
      using E M N by auto
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42981
diff changeset
   241
    show "\<And>i. M.\<mu> (space M) \<noteq> \<infinity>"
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42981
diff changeset
   242
      using M.measure_space_1 by simp
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42981
diff changeset
   243
    show "measure_space \<lparr>space = space E, sets = sets (sigma E), measure_space.measure = M.\<mu>\<rparr>"
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42981
diff changeset
   244
      using E M N by (auto intro!: M.measure_space_cong)
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42981
diff changeset
   245
    show "measure_space \<lparr>space = space E, sets = sets (sigma E), measure_space.measure = N.\<mu>\<rparr>"
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42981
diff changeset
   246
      using E M N by (auto intro!: N.measure_space_cong)
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42981
diff changeset
   247
    { fix X assume "X \<in> sets E"
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42981
diff changeset
   248
      then have "X \<in> sets (sigma E)"
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42981
diff changeset
   249
        by (auto simp: sets_sigma sigma_sets.Basic)
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42981
diff changeset
   250
      with eq[OF `X \<in> sets E`] M N show "M.\<mu> X = N.\<mu> X"
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42981
diff changeset
   251
        by (simp add: M.finite_measure_eq N.finite_measure_eq) }
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42981
diff changeset
   252
  qed fact
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42981
diff changeset
   253
  with `X \<in> sets (sigma E)` M N show ?thesis
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42981
diff changeset
   254
    by (simp add: M.finite_measure_eq N.finite_measure_eq)
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42981
diff changeset
   255
qed
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42981
diff changeset
   256
42199
aded34119213 add prob_space_vimage
hoelzl
parents: 42148
diff changeset
   257
lemma (in prob_space) distribution_prob_space:
aded34119213 add prob_space_vimage
hoelzl
parents: 42148
diff changeset
   258
  assumes X: "random_variable S X"
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43658
diff changeset
   259
  shows "prob_space (S\<lparr>measure := ereal \<circ> distribution X\<rparr>)" (is "prob_space ?S")
42199
aded34119213 add prob_space_vimage
hoelzl
parents: 42148
diff changeset
   260
proof (rule prob_space_vimage)
aded34119213 add prob_space_vimage
hoelzl
parents: 42148
diff changeset
   261
  show "X \<in> measure_preserving M ?S"
aded34119213 add prob_space_vimage
hoelzl
parents: 42148
diff changeset
   262
    using X
46905
6b1c0a80a57a prefer abs_def over def_raw;
wenzelm
parents: 46898
diff changeset
   263
    unfolding measure_preserving_def distribution_def [abs_def]
42199
aded34119213 add prob_space_vimage
hoelzl
parents: 42148
diff changeset
   264
    by (auto simp: finite_measure_eq measurable_sets)
aded34119213 add prob_space_vimage
hoelzl
parents: 42148
diff changeset
   265
  show "sigma_algebra ?S" using X by simp
aded34119213 add prob_space_vimage
hoelzl
parents: 42148
diff changeset
   266
qed
aded34119213 add prob_space_vimage
hoelzl
parents: 42148
diff changeset
   267
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   268
lemma (in prob_space) AE_distribution:
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43658
diff changeset
   269
  assumes X: "random_variable MX X" and "AE x in MX\<lparr>measure := ereal \<circ> distribution X\<rparr>. Q x"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   270
  shows "AE x. Q (X x)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   271
proof -
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43658
diff changeset
   272
  interpret X: prob_space "MX\<lparr>measure := ereal \<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
   273
  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
   274
    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
   275
  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
   276
    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
   277
       (auto simp: finite_measure_eq distribution_def measurable_sets)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   278
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   279
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   280
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
   281
  "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
   282
  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
   283
  by (auto simp: measurable_sets distribution_def)
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   284
43339
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   285
lemma (in prob_space) expectation_less:
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   286
  assumes [simp]: "integrable M X"
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   287
  assumes gt: "\<forall>x\<in>space M. X x < b"
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   288
  shows "expectation X < b"
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   289
proof -
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   290
  have "expectation X < expectation (\<lambda>x. b)"
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   291
    using gt measure_space_1
43340
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 43339
diff changeset
   292
    by (intro integral_less_AE_space) auto
43339
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   293
  then show ?thesis using prob_space by simp
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   294
qed
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   295
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   296
lemma (in prob_space) expectation_greater:
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   297
  assumes [simp]: "integrable M X"
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   298
  assumes gt: "\<forall>x\<in>space M. a < X x"
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   299
  shows "a < expectation X"
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   300
proof -
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   301
  have "expectation (\<lambda>x. a) < expectation X"
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   302
    using gt measure_space_1
43340
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 43339
diff changeset
   303
    by (intro integral_less_AE_space) auto
43339
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   304
  then show ?thesis using prob_space by simp
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   305
qed
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   306
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   307
lemma convex_le_Inf_differential:
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   308
  fixes f :: "real \<Rightarrow> real"
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   309
  assumes "convex_on I f"
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   310
  assumes "x \<in> interior I" "y \<in> I"
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   311
  shows "f y \<ge> f x + Inf ((\<lambda>t. (f x - f t) / (x - t)) ` ({x<..} \<inter> I)) * (y - x)"
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   312
    (is "_ \<ge> _ + Inf (?F x) * (y - x)")
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   313
proof -
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   314
  show ?thesis
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   315
  proof (cases rule: linorder_cases)
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   316
    assume "x < y"
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   317
    moreover
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   318
    have "open (interior I)" by auto
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   319
    from openE[OF this `x \<in> interior I`] guess e . note e = this
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   320
    moreover def t \<equiv> "min (x + e / 2) ((x + y) / 2)"
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   321
    ultimately have "x < t" "t < y" "t \<in> ball x e"
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   322
      by (auto simp: mem_ball dist_real_def field_simps split: split_min)
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   323
    with `x \<in> interior I` e interior_subset[of I] have "t \<in> I" "x \<in> I" by auto
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   324
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   325
    have "open (interior I)" by auto
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   326
    from openE[OF this `x \<in> interior I`] guess e .
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   327
    moreover def K \<equiv> "x - e / 2"
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   328
    with `0 < e` have "K \<in> ball x e" "K < x" by (auto simp: mem_ball dist_real_def)
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   329
    ultimately have "K \<in> I" "K < x" "x \<in> I"
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   330
      using interior_subset[of I] `x \<in> interior I` by auto
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   331
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   332
    have "Inf (?F x) \<le> (f x - f y) / (x - y)"
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   333
    proof (rule Inf_lower2)
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   334
      show "(f x - f t) / (x - t) \<in> ?F x"
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   335
        using `t \<in> I` `x < t` by auto
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   336
      show "(f x - f t) / (x - t) \<le> (f x - f y) / (x - y)"
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   337
        using `convex_on I f` `x \<in> I` `y \<in> I` `x < t` `t < y` by (rule convex_on_diff)
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   338
    next
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   339
      fix y assume "y \<in> ?F x"
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   340
      with order_trans[OF convex_on_diff[OF `convex_on I f` `K \<in> I` _ `K < x` _]]
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   341
      show "(f K - f x) / (K - x) \<le> y" by auto
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   342
    qed
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   343
    then show ?thesis
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   344
      using `x < y` by (simp add: field_simps)
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   345
  next
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   346
    assume "y < x"
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   347
    moreover
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   348
    have "open (interior I)" by auto
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   349
    from openE[OF this `x \<in> interior I`] guess e . note e = this
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   350
    moreover def t \<equiv> "x + e / 2"
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   351
    ultimately have "x < t" "t \<in> ball x e"
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   352
      by (auto simp: mem_ball dist_real_def field_simps)
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   353
    with `x \<in> interior I` e interior_subset[of I] have "t \<in> I" "x \<in> I" by auto
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   354
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   355
    have "(f x - f y) / (x - y) \<le> Inf (?F x)"
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   356
    proof (rule Inf_greatest)
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   357
      have "(f x - f y) / (x - y) = (f y - f x) / (y - x)"
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   358
        using `y < x` by (auto simp: field_simps)
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   359
      also
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   360
      fix z  assume "z \<in> ?F x"
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   361
      with order_trans[OF convex_on_diff[OF `convex_on I f` `y \<in> I` _ `y < x`]]
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   362
      have "(f y - f x) / (y - x) \<le> z" by auto
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   363
      finally show "(f x - f y) / (x - y) \<le> z" .
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   364
    next
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   365
      have "open (interior I)" by auto
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   366
      from openE[OF this `x \<in> interior I`] guess e . note e = this
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   367
      then have "x + e / 2 \<in> ball x e" by (auto simp: mem_ball dist_real_def)
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   368
      with e interior_subset[of I] have "x + e / 2 \<in> {x<..} \<inter> I" by auto
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   369
      then show "?F x \<noteq> {}" by blast
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   370
    qed
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   371
    then show ?thesis
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   372
      using `y < x` by (simp add: field_simps)
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   373
  qed simp
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   374
qed
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   375
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   376
lemma (in prob_space) jensens_inequality:
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   377
  fixes a b :: real
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   378
  assumes X: "integrable M X" "X ` space M \<subseteq> I"
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   379
  assumes I: "I = {a <..< b} \<or> I = {a <..} \<or> I = {..< b} \<or> I = UNIV"
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   380
  assumes q: "integrable M (\<lambda>x. q (X x))" "convex_on I q"
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   381
  shows "q (expectation X) \<le> expectation (\<lambda>x. q (X x))"
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   382
proof -
46731
5302e932d1e5 avoid undeclared variables in let bindings;
wenzelm
parents: 45934
diff changeset
   383
  let ?F = "\<lambda>x. Inf ((\<lambda>t. (q x - q t) / (x - t)) ` ({x<..} \<inter> I))"
43339
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   384
  from not_empty X(2) have "I \<noteq> {}" by auto
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   385
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   386
  from I have "open I" by auto
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   387
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   388
  note I
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   389
  moreover
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   390
  { assume "I \<subseteq> {a <..}"
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   391
    with X have "a < expectation X"
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   392
      by (intro expectation_greater) auto }
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   393
  moreover
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   394
  { assume "I \<subseteq> {..< b}"
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   395
    with X have "expectation X < b"
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   396
      by (intro expectation_less) auto }
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   397
  ultimately have "expectation X \<in> I"
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   398
    by (elim disjE)  (auto simp: subset_eq)
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   399
  moreover
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   400
  { fix y assume y: "y \<in> I"
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   401
    with q(2) `open I` have "Sup ((\<lambda>x. q x + ?F x * (y - x)) ` I) = q y"
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   402
      by (auto intro!: Sup_eq_maximum convex_le_Inf_differential image_eqI[OF _ y] simp: interior_open) }
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   403
  ultimately have "q (expectation X) = Sup ((\<lambda>x. q x + ?F x * (expectation X - x)) ` I)"
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   404
    by simp
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   405
  also have "\<dots> \<le> expectation (\<lambda>w. q (X w))"
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   406
  proof (rule Sup_least)
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   407
    show "(\<lambda>x. q x + ?F x * (expectation X - x)) ` I \<noteq> {}"
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   408
      using `I \<noteq> {}` by auto
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   409
  next
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   410
    fix k assume "k \<in> (\<lambda>x. q x + ?F x * (expectation X - x)) ` I"
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   411
    then guess x .. note x = this
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   412
    have "q x + ?F x * (expectation X  - x) = expectation (\<lambda>w. q x + ?F x * (X w - x))"
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   413
      using prob_space
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   414
      by (simp add: integral_add integral_cmult integral_diff lebesgue_integral_const X)
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   415
    also have "\<dots> \<le> expectation (\<lambda>w. q (X w))"
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   416
      using `x \<in> I` `open I` X(2)
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   417
      by (intro integral_mono integral_add integral_cmult integral_diff
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   418
                lebesgue_integral_const X q convex_le_Inf_differential)
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   419
         (auto simp: interior_open)
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   420
    finally show "k \<le> expectation (\<lambda>w. q (X w))" using x by auto
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   421
  qed
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   422
  finally show "q (expectation X) \<le> expectation (\<lambda>x. q (X x))" .
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   423
qed
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   424
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   425
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
   426
  assumes "random_variable S X" "A \<in> sets S"
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43658
diff changeset
   427
  shows "distribution X A = integral\<^isup>P (S\<lparr>measure := ereal \<circ> distribution X\<rparr>) (indicator A)"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   428
proof -
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43658
diff changeset
   429
  interpret S: prob_space "S\<lparr>measure := ereal \<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
   430
    using assms(1) by (rule distribution_prob_space)
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   431
  show ?thesis
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   432
    using S.positive_integral_indicator(1)[of A] assms by simp
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   433
qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   434
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   435
lemma (in prob_space) finite_expectation1:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   436
  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
   437
  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
   438
proof (subst integral_on_finite)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   439
  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
   440
  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
   441
    "\<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
   442
    using finite_measure_eq[OF borel_measurable_vimage, of X] rv by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   443
qed
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   444
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   445
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
   446
  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
   447
  shows "expectation X = (\<Sum> r \<in> X ` (space M). r * distribution X {r})"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   448
  using assms unfolding distribution_def using finite_expectation1 by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   449
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   450
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
   451
  assumes "{x} \<in> events"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   452
  assumes "prob {x} = 1"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   453
  assumes "{y} \<in> events"
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   454
  assumes "y \<noteq> x"
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   455
  shows "prob {y} = 0"
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   456
  using prob_one_inter[of "{y}" "{x}"] assms by auto
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   457
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   458
lemma (in prob_space) distribution_empty[simp]: "distribution X {} = 0"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   459
  unfolding distribution_def by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   460
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   461
lemma (in prob_space) distribution_space[simp]: "distribution X (X ` space M) = 1"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   462
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   463
  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
   464
  thus ?thesis unfolding distribution_def by (simp add: prob_space)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   465
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   466
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   467
lemma (in prob_space) distribution_one:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   468
  assumes "random_variable M' X" and "A \<in> sets M'"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   469
  shows "distribution X A \<le> 1"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   470
proof -
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   471
  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
   472
    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
   473
  thus ?thesis by (simp add: prob_space)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   474
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   475
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   476
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
   477
  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
   478
    (is "random_variable ?S X")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   479
  assumes "distribution X {x} = 1"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   480
  assumes "y \<noteq> x"
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   481
  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
   482
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
   483
  { 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
   484
    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
   485
      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
   486
        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
   487
    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
   488
      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
   489
      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
   490
    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
   491
  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
   492
    "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
   493
    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
   494
  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
   495
  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
   496
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
   497
  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
   498
  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
   499
  thus "distribution X {y} = 0" unfolding distribution_def by auto
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   500
qed
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   501
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   502
lemma (in prob_space) joint_distribution_Times_le_fst:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   503
  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
   504
    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
   505
  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
   506
  unfolding distribution_def
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   507
proof (intro finite_measure_mono)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   508
  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
   509
  show "X -` A \<inter> space M \<in> events"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   510
    using X A unfolding measurable_def by simp
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   511
  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
   512
    (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
   513
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   514
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   515
lemma (in prob_space) joint_distribution_commute:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   516
  "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
   517
  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
   518
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   519
lemma (in prob_space) joint_distribution_Times_le_snd:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   520
  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
   521
    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
   522
  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
   523
  using assms
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   524
  by (subst joint_distribution_commute)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   525
     (simp add: swap_product joint_distribution_Times_le_fst)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   526
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   527
lemma (in prob_space) random_variable_pairI:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   528
  assumes "random_variable MX X"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   529
  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
   530
  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
   531
proof
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   532
  interpret MX: sigma_algebra MX using assms by simp
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   533
  interpret MY: sigma_algebra MY using assms by simp
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   534
  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
   535
  show "sigma_algebra (MX \<Otimes>\<^isub>M MY)" by default
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   536
  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
   537
  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
   538
    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
   539
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   540
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   541
lemma (in prob_space) joint_distribution_commute_singleton:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   542
  "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
   543
  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
   544
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   545
lemma (in prob_space) joint_distribution_assoc_singleton:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   546
  "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
   547
   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
   548
  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
   549
45777
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
   550
locale pair_prob_space = pair_sigma_finite M1 M2 + M1: prob_space M1 + M2: prob_space M2 for M1 M2
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
   551
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
   552
sublocale pair_prob_space \<subseteq> P: prob_space P
45777
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
   553
proof
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
   554
  show "measure_space P" ..
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
   555
  show "measure P (space P) = 1"
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
   556
    by (simp add: pair_measure_times M1.measure_space_1 M2.measure_space_1 space_pair_measure)
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
   557
qed
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   558
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   559
lemma countably_additiveI[case_names countably]:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   560
  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
   561
    (\<Sum>n. \<mu> (A n)) = \<mu> (\<Union>i. A i)"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   562
  shows "countably_additive M \<mu>"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   563
  using assms unfolding countably_additive_def by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   564
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   565
lemma (in prob_space) joint_distribution_prob_space:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   566
  assumes "random_variable MX X" "random_variable MY Y"
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43658
diff changeset
   567
  shows "prob_space ((MX \<Otimes>\<^isub>M MY) \<lparr> measure := ereal \<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
   568
  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
   569
45777
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
   570
locale product_prob_space = product_sigma_finite M for M :: "'i \<Rightarrow> ('a, 'b) measure_space_scheme" +
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
   571
  fixes I :: "'i set"
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
   572
  assumes prob_space: "\<And>i. prob_space (M i)"
42988
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42981
diff changeset
   573
45777
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
   574
sublocale product_prob_space \<subseteq> M: prob_space "M i" for i
42988
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42981
diff changeset
   575
  by (rule prob_space)
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42981
diff changeset
   576
45777
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
   577
locale finite_product_prob_space = finite_product_sigma_finite M I + product_prob_space M I for M I
42988
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42981
diff changeset
   578
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42981
diff changeset
   579
sublocale finite_product_prob_space \<subseteq> prob_space "\<Pi>\<^isub>M i\<in>I. M i"
45777
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
   580
proof
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
   581
  show "measure_space P" ..
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
   582
  show "measure P (space P) = 1"
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
   583
    by (simp add: measure_times M.measure_space_1 setprod_1)
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
   584
qed
42988
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42981
diff changeset
   585
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42981
diff changeset
   586
lemma (in finite_product_prob_space) prob_times:
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42981
diff changeset
   587
  assumes X: "\<And>i. i \<in> I \<Longrightarrow> X i \<in> sets (M i)"
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42981
diff changeset
   588
  shows "prob (\<Pi>\<^isub>E i\<in>I. X i) = (\<Prod>i\<in>I. M.prob i (X i))"
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42981
diff changeset
   589
proof -
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43658
diff changeset
   590
  have "ereal (\<mu>' (\<Pi>\<^isub>E i\<in>I. X i)) = \<mu> (\<Pi>\<^isub>E i\<in>I. X i)"
42988
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42981
diff changeset
   591
    using X by (intro finite_measure_eq[symmetric] in_P) auto
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42981
diff changeset
   592
  also have "\<dots> = (\<Prod>i\<in>I. M.\<mu> i (X i))"
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42981
diff changeset
   593
    using measure_times X by simp
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43658
diff changeset
   594
  also have "\<dots> = ereal (\<Prod>i\<in>I. M.\<mu>' i (X i))"
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43658
diff changeset
   595
    using X by (simp add: M.finite_measure_eq setprod_ereal)
42988
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42981
diff changeset
   596
  finally show ?thesis by simp
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42981
diff changeset
   597
qed
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42981
diff changeset
   598
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42981
diff changeset
   599
lemma (in prob_space) random_variable_restrict:
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42981
diff changeset
   600
  assumes I: "finite I"
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42981
diff changeset
   601
  assumes X: "\<And>i. i \<in> I \<Longrightarrow> random_variable (N i) (X i)"
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42981
diff changeset
   602
  shows "random_variable (Pi\<^isub>M I N) (\<lambda>x. \<lambda>i\<in>I. X i x)"
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42981
diff changeset
   603
proof
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42981
diff changeset
   604
  { fix i assume "i \<in> I"
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42981
diff changeset
   605
    with X interpret N: sigma_algebra "N i" by simp
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42981
diff changeset
   606
    have "sets (N i) \<subseteq> Pow (space (N i))" by (rule N.space_closed) }
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42981
diff changeset
   607
  note N_closed = this
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42981
diff changeset
   608
  then show "sigma_algebra (Pi\<^isub>M I N)"
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42981
diff changeset
   609
    by (simp add: product_algebra_def)
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42981
diff changeset
   610
       (intro sigma_algebra_sigma product_algebra_generator_sets_into_space)
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42981
diff changeset
   611
  show "(\<lambda>x. \<lambda>i\<in>I. X i x) \<in> measurable M (Pi\<^isub>M I N)"
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42981
diff changeset
   612
    using X by (intro measurable_restrict[OF I N_closed]) auto
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42981
diff changeset
   613
qed
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42981
diff changeset
   614
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   615
section "Probability spaces on finite sets"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   616
35977
30d42bfd0174 Added finite measure space.
hoelzl
parents: 35929
diff changeset
   617
locale finite_prob_space = prob_space + finite_measure_space
30d42bfd0174 Added finite measure space.
hoelzl
parents: 35929
diff changeset
   618
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   619
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
   620
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   621
lemma (in prob_space) finite_random_variableD:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   622
  assumes "finite_random_variable M' X" shows "random_variable M' X"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   623
proof -
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   624
  interpret M': finite_sigma_algebra M' using assms by simp
46898
1570b30ee040 tuned proofs -- eliminated pointless chaining of facts after 'interpret';
wenzelm
parents: 46731
diff changeset
   625
  show "random_variable M' X" using assms by simp default
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   626
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   627
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   628
lemma (in prob_space) distribution_finite_prob_space:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   629
  assumes "finite_random_variable MX X"
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43658
diff changeset
   630
  shows "finite_prob_space (MX\<lparr>measure := ereal \<circ> distribution X\<rparr>)"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   631
proof -
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43658
diff changeset
   632
  interpret X: prob_space "MX\<lparr>measure := ereal \<circ> distribution X\<rparr>"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   633
    using assms[THEN finite_random_variableD] by (rule distribution_prob_space)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   634
  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
   635
    using assms by auto
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   636
  show ?thesis by default (simp_all add: MX.finite_space)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   637
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   638
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   639
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
   640
  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
   641
  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
   642
    (is "finite_random_variable ?X _")
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   643
proof (intro conjI)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   644
  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
   645
  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
   646
  show "finite_sigma_algebra ?X"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   647
    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
   648
  show "X \<in> measurable M ?X"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   649
  proof (unfold measurable_def, clarsimp)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   650
    fix A assume A: "A \<subseteq> X`space M"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   651
    then have "finite A" by (rule finite_subset) simp
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   652
    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
   653
      unfolding vimage_UN UN_extend_simps
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   654
      apply (rule finite_UN)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   655
      using A assms unfolding simple_function_def by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   656
    then show "X -` A \<inter> space M \<in> events" by simp
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   657
  qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   658
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   659
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   660
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
   661
  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
   662
  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
   663
  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
   664
  by (auto dest!: finite_random_variableD)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   665
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   666
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
   667
  "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
   668
  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
   669
  by (subst finite_measure_finite_Union[symmetric])
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   670
     (auto simp add: disjoint_family_on_def simple_function_def
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   671
           intro!: arg_cong[where f=prob])
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   672
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   673
lemma (in prob_space) finite_random_variable_pairI:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   674
  assumes "finite_random_variable MX X"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   675
  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
   676
  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
   677
proof
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   678
  interpret MX: finite_sigma_algebra MX using assms by simp
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   679
  interpret MY: finite_sigma_algebra MY using assms by simp
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   680
  interpret P: pair_finite_sigma_algebra MX MY by default
45777
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
   681
  show "finite_sigma_algebra (MX \<Otimes>\<^isub>M MY)" ..
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   682
  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
   683
  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
   684
    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
   685
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   686
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   687
lemma (in prob_space) finite_random_variable_imp_sets:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   688
  "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
   689
  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
   690
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   691
lemma (in prob_space) finite_random_variable_measurable:
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   692
  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
   693
proof -
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   694
  interpret X: finite_sigma_algebra MX using X by simp
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   695
  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
   696
    "X \<in> space M \<rightarrow> space MX"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   697
    by (auto simp: measurable_def)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   698
  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
   699
    by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   700
  show "X -` A \<inter> space M \<in> events"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   701
    unfolding * by (intro vimage) auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   702
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   703
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   704
lemma (in prob_space) joint_distribution_finite_Times_le_fst:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   705
  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
   706
  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
   707
  unfolding distribution_def
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   708
proof (intro finite_measure_mono)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   709
  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
   710
  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
   711
    using finite_random_variable_measurable[OF X] .
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   712
  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
   713
    (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
   714
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   715
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   716
lemma (in prob_space) joint_distribution_finite_Times_le_snd:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   717
  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
   718
  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
   719
  using assms
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   720
  by (subst joint_distribution_commute)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   721
     (simp add: swap_product joint_distribution_finite_Times_le_fst)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   722
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   723
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
   724
  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
   725
  assumes "finite_random_variable MX X" "finite_random_variable MY Y"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   726
  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
   727
    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
   728
    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
   729
    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
   730
    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
   731
    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
   732
  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
   733
  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
   734
  by (auto intro: antisym)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   735
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   736
lemma (in prob_space) setsum_joint_distribution:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   737
  assumes X: "finite_random_variable MX X"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   738
  assumes Y: "random_variable MY Y" "B \<in> sets MY"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   739
  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
   740
  unfolding distribution_def
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   741
proof (subst finite_measure_finite_Union[symmetric])
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   742
  interpret MX: finite_sigma_algebra MX using X by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   743
  show "finite (space MX)" using MX.finite_space .
46731
5302e932d1e5 avoid undeclared variables in let bindings;
wenzelm
parents: 45934
diff changeset
   744
  let ?d = "\<lambda>i. (\<lambda>x. (X x, Y x)) -` ({i} \<times> B) \<inter> space M"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   745
  { fix i assume "i \<in> space MX"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   746
    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
   747
    ultimately show "?d i \<in> events"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   748
      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
   749
      using MX.sets_eq_Pow by auto }
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   750
  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
   751
  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
   752
    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
   753
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   754
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   755
lemma (in prob_space) setsum_joint_distribution_singleton:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   756
  assumes X: "finite_random_variable MX X"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   757
  assumes Y: "finite_random_variable MY Y" "b \<in> space MY"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   758
  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
   759
  using setsum_joint_distribution[OF X
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   760
    finite_random_variableD[OF Y(1)]
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   761
    finite_random_variable_imp_sets[OF Y]] by simp
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   762
45712
852597248663 moved theorems about distribution to the definition; removed oopsed-lemma
hoelzl
parents: 45711
diff changeset
   763
lemma (in prob_space) setsum_distribution:
852597248663 moved theorems about distribution to the definition; removed oopsed-lemma
hoelzl
parents: 45711
diff changeset
   764
  assumes X: "finite_random_variable MX X" shows "(\<Sum>a\<in>space MX. distribution X {a}) = 1"
852597248663 moved theorems about distribution to the definition; removed oopsed-lemma
hoelzl
parents: 45711
diff changeset
   765
  using setsum_joint_distribution[OF assms, of "\<lparr> space = UNIV, sets = Pow UNIV \<rparr>" "\<lambda>x. ()" "{()}"]
852597248663 moved theorems about distribution to the definition; removed oopsed-lemma
hoelzl
parents: 45711
diff changeset
   766
  using sigma_algebra_Pow[of "UNIV::unit set" "()"] by simp
852597248663 moved theorems about distribution to the definition; removed oopsed-lemma
hoelzl
parents: 45711
diff changeset
   767
45777
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
   768
locale pair_finite_prob_space = pair_prob_space M1 M2 + pair_finite_space M1 M2 + 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
   769
45777
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
   770
sublocale pair_finite_prob_space \<subseteq> finite_prob_space P by default
42859
d9dfc733f25c add product of probability spaces with finite cardinality
hoelzl
parents: 42858
diff changeset
   771
d9dfc733f25c add product of probability spaces with finite cardinality
hoelzl
parents: 42858
diff changeset
   772
lemma funset_eq_UN_fun_upd_I:
d9dfc733f25c add product of probability spaces with finite cardinality
hoelzl
parents: 42858
diff changeset
   773
  assumes *: "\<And>f. f \<in> F (insert a A) \<Longrightarrow> f(a := d) \<in> F A"
d9dfc733f25c add product of probability spaces with finite cardinality
hoelzl
parents: 42858
diff changeset
   774
  and **: "\<And>f. f \<in> F (insert a A) \<Longrightarrow> f a \<in> G (f(a:=d))"
d9dfc733f25c add product of probability spaces with finite cardinality
hoelzl
parents: 42858
diff changeset
   775
  and ***: "\<And>f x. \<lbrakk> f \<in> F A ; x \<in> G f \<rbrakk> \<Longrightarrow> f(a:=x) \<in> F (insert a A)"
d9dfc733f25c add product of probability spaces with finite cardinality
hoelzl
parents: 42858
diff changeset
   776
  shows "F (insert a A) = (\<Union>f\<in>F A. fun_upd f a ` (G f))"
d9dfc733f25c add product of probability spaces with finite cardinality
hoelzl
parents: 42858
diff changeset
   777
proof safe
d9dfc733f25c add product of probability spaces with finite cardinality
hoelzl
parents: 42858
diff changeset
   778
  fix f assume f: "f \<in> F (insert a A)"
d9dfc733f25c add product of probability spaces with finite cardinality
hoelzl
parents: 42858
diff changeset
   779
  show "f \<in> (\<Union>f\<in>F A. fun_upd f a ` G f)"
d9dfc733f25c add product of probability spaces with finite cardinality
hoelzl
parents: 42858
diff changeset
   780
  proof (rule UN_I[of "f(a := d)"])
d9dfc733f25c add product of probability spaces with finite cardinality
hoelzl
parents: 42858
diff changeset
   781
    show "f(a := d) \<in> F A" using *[OF f] .
d9dfc733f25c add product of probability spaces with finite cardinality
hoelzl
parents: 42858
diff changeset
   782
    show "f \<in> fun_upd (f(a:=d)) a ` G (f(a:=d))"
d9dfc733f25c add product of probability spaces with finite cardinality
hoelzl
parents: 42858
diff changeset
   783
    proof (rule image_eqI[of _ _ "f a"])
d9dfc733f25c add product of probability spaces with finite cardinality
hoelzl
parents: 42858
diff changeset
   784
      show "f a \<in> G (f(a := d))" using **[OF f] .
d9dfc733f25c add product of probability spaces with finite cardinality
hoelzl
parents: 42858
diff changeset
   785
    qed simp
d9dfc733f25c add product of probability spaces with finite cardinality
hoelzl
parents: 42858
diff changeset
   786
  qed
d9dfc733f25c add product of probability spaces with finite cardinality
hoelzl
parents: 42858
diff changeset
   787
next
d9dfc733f25c add product of probability spaces with finite cardinality
hoelzl
parents: 42858
diff changeset
   788
  fix f x assume "f \<in> F A" "x \<in> G f"
d9dfc733f25c add product of probability spaces with finite cardinality
hoelzl
parents: 42858
diff changeset
   789
  from ***[OF this] show "f(a := x) \<in> F (insert a A)" .
d9dfc733f25c add product of probability spaces with finite cardinality
hoelzl
parents: 42858
diff changeset
   790
qed
d9dfc733f25c add product of probability spaces with finite cardinality
hoelzl
parents: 42858
diff changeset
   791
d9dfc733f25c add product of probability spaces with finite cardinality
hoelzl
parents: 42858
diff changeset
   792
lemma extensional_funcset_insert_eq[simp]:
d9dfc733f25c add product of probability spaces with finite cardinality
hoelzl
parents: 42858
diff changeset
   793
  assumes "a \<notin> A"
d9dfc733f25c add product of probability spaces with finite cardinality
hoelzl
parents: 42858
diff changeset
   794
  shows "extensional (insert a A) \<inter> (insert a A \<rightarrow> B) = (\<Union>f \<in> extensional A \<inter> (A \<rightarrow> B). (\<lambda>b. f(a := b)) ` B)"
d9dfc733f25c add product of probability spaces with finite cardinality
hoelzl
parents: 42858
diff changeset
   795
  apply (rule funset_eq_UN_fun_upd_I)
d9dfc733f25c add product of probability spaces with finite cardinality
hoelzl
parents: 42858
diff changeset
   796
  using assms
d9dfc733f25c add product of probability spaces with finite cardinality
hoelzl
parents: 42858
diff changeset
   797
  by (auto intro!: inj_onI dest: inj_onD split: split_if_asm simp: extensional_def)
d9dfc733f25c add product of probability spaces with finite cardinality
hoelzl
parents: 42858
diff changeset
   798
d9dfc733f25c add product of probability spaces with finite cardinality
hoelzl
parents: 42858
diff changeset
   799
lemma finite_extensional_funcset[simp, intro]:
d9dfc733f25c add product of probability spaces with finite cardinality
hoelzl
parents: 42858
diff changeset
   800
  assumes "finite A" "finite B"
d9dfc733f25c add product of probability spaces with finite cardinality
hoelzl
parents: 42858
diff changeset
   801
  shows "finite (extensional A \<inter> (A \<rightarrow> B))"
d9dfc733f25c add product of probability spaces with finite cardinality
hoelzl
parents: 42858
diff changeset
   802
  using assms by induct (auto simp: extensional_funcset_insert_eq)
d9dfc733f25c add product of probability spaces with finite cardinality
hoelzl
parents: 42858
diff changeset
   803
d9dfc733f25c add product of probability spaces with finite cardinality
hoelzl
parents: 42858
diff changeset
   804
lemma finite_PiE[simp, intro]:
d9dfc733f25c add product of probability spaces with finite cardinality
hoelzl
parents: 42858
diff changeset
   805
  assumes fin: "finite A" "\<And>i. i \<in> A \<Longrightarrow> finite (B i)"
d9dfc733f25c add product of probability spaces with finite cardinality
hoelzl
parents: 42858
diff changeset
   806
  shows "finite (Pi\<^isub>E A B)"
d9dfc733f25c add product of probability spaces with finite cardinality
hoelzl
parents: 42858
diff changeset
   807
proof -
d9dfc733f25c add product of probability spaces with finite cardinality
hoelzl
parents: 42858
diff changeset
   808
  have *: "(Pi\<^isub>E A B) \<subseteq> extensional A \<inter> (A \<rightarrow> (\<Union>i\<in>A. B i))" by auto
d9dfc733f25c add product of probability spaces with finite cardinality
hoelzl
parents: 42858
diff changeset
   809
  show ?thesis
d9dfc733f25c add product of probability spaces with finite cardinality
hoelzl
parents: 42858
diff changeset
   810
    using fin by (intro finite_subset[OF *] finite_extensional_funcset) auto
d9dfc733f25c add product of probability spaces with finite cardinality
hoelzl
parents: 42858
diff changeset
   811
qed
d9dfc733f25c add product of probability spaces with finite cardinality
hoelzl
parents: 42858
diff changeset
   812
45777
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
   813
locale finite_product_finite_prob_space = finite_product_prob_space M I for M I +
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
   814
  assumes finite_space: "\<And>i. finite_prob_space (M i)"
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
   815
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
   816
sublocale finite_product_finite_prob_space \<subseteq> M!: finite_prob_space "M i" using finite_space .
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
   817
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
   818
lemma (in finite_product_finite_prob_space) singleton_eq_product:
42892
a61e30bfd0bc add lemma prob_finite_product
hoelzl
parents: 42860
diff changeset
   819
  assumes x: "x \<in> space P" shows "{x} = (\<Pi>\<^isub>E i\<in>I. {x i})"
a61e30bfd0bc add lemma prob_finite_product
hoelzl
parents: 42860
diff changeset
   820
proof (safe intro!: ext[of _ x])
a61e30bfd0bc add lemma prob_finite_product
hoelzl
parents: 42860
diff changeset
   821
  fix y i assume *: "y \<in> (\<Pi> i\<in>I. {x i})" "y \<in> extensional I"
a61e30bfd0bc add lemma prob_finite_product
hoelzl
parents: 42860
diff changeset
   822
  with x show "y i = x i"
a61e30bfd0bc add lemma prob_finite_product
hoelzl
parents: 42860
diff changeset
   823
    by (cases "i \<in> I") (auto simp: extensional_def)
a61e30bfd0bc add lemma prob_finite_product
hoelzl
parents: 42860
diff changeset
   824
qed (insert x, auto)
a61e30bfd0bc add lemma prob_finite_product
hoelzl
parents: 42860
diff changeset
   825
45777
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
   826
sublocale finite_product_finite_prob_space \<subseteq> finite_prob_space "Pi\<^isub>M I M"
42859
d9dfc733f25c add product of probability spaces with finite cardinality
hoelzl
parents: 42858
diff changeset
   827
proof
d9dfc733f25c add product of probability spaces with finite cardinality
hoelzl
parents: 42858
diff changeset
   828
  show "finite (space P)"
d9dfc733f25c add product of probability spaces with finite cardinality
hoelzl
parents: 42858
diff changeset
   829
    using finite_index M.finite_space by auto
d9dfc733f25c add product of probability spaces with finite cardinality
hoelzl
parents: 42858
diff changeset
   830
d9dfc733f25c add product of probability spaces with finite cardinality
hoelzl
parents: 42858
diff changeset
   831
  { fix x assume "x \<in> space P"
42892
a61e30bfd0bc add lemma prob_finite_product
hoelzl
parents: 42860
diff changeset
   832
    with this[THEN singleton_eq_product]
a61e30bfd0bc add lemma prob_finite_product
hoelzl
parents: 42860
diff changeset
   833
    have "{x} \<in> sets P"
42859
d9dfc733f25c add product of probability spaces with finite cardinality
hoelzl
parents: 42858
diff changeset
   834
      by (auto intro!: in_P) }
d9dfc733f25c add product of probability spaces with finite cardinality
hoelzl
parents: 42858
diff changeset
   835
  note x_in_P = this
d9dfc733f25c add product of probability spaces with finite cardinality
hoelzl
parents: 42858
diff changeset
   836
d9dfc733f25c add product of probability spaces with finite cardinality
hoelzl
parents: 42858
diff changeset
   837
  have "Pow (space P) \<subseteq> sets P"
d9dfc733f25c add product of probability spaces with finite cardinality
hoelzl
parents: 42858
diff changeset
   838
  proof
d9dfc733f25c add product of probability spaces with finite cardinality
hoelzl
parents: 42858
diff changeset
   839
    fix X assume "X \<in> Pow (space P)"
d9dfc733f25c add product of probability spaces with finite cardinality
hoelzl
parents: 42858
diff changeset
   840
    moreover then have "finite X"
d9dfc733f25c add product of probability spaces with finite cardinality
hoelzl
parents: 42858
diff changeset
   841
      using `finite (space P)` by (blast intro: finite_subset)
d9dfc733f25c add product of probability spaces with finite cardinality
hoelzl
parents: 42858
diff changeset
   842
    ultimately have "(\<Union>x\<in>X. {x}) \<in> sets P"
d9dfc733f25c add product of probability spaces with finite cardinality
hoelzl
parents: 42858
diff changeset
   843
      by (intro finite_UN x_in_P) auto
d9dfc733f25c add product of probability spaces with finite cardinality
hoelzl
parents: 42858
diff changeset
   844
    then show "X \<in> sets P" by simp
d9dfc733f25c add product of probability spaces with finite cardinality
hoelzl
parents: 42858
diff changeset
   845
  qed
d9dfc733f25c add product of probability spaces with finite cardinality
hoelzl
parents: 42858
diff changeset
   846
  with space_closed show [simp]: "sets P = Pow (space P)" ..
d9dfc733f25c add product of probability spaces with finite cardinality
hoelzl
parents: 42858
diff changeset
   847
qed
d9dfc733f25c add product of probability spaces with finite cardinality
hoelzl
parents: 42858
diff changeset
   848
45777
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
   849
lemma (in finite_product_finite_prob_space) measure_finite_times:
42859
d9dfc733f25c add product of probability spaces with finite cardinality
hoelzl
parents: 42858
diff changeset
   850
  "(\<And>i. i \<in> I \<Longrightarrow> X i \<subseteq> space (M i)) \<Longrightarrow> \<mu> (\<Pi>\<^isub>E i\<in>I. X i) = (\<Prod>i\<in>I. M.\<mu> i (X i))"
d9dfc733f25c add product of probability spaces with finite cardinality
hoelzl
parents: 42858
diff changeset
   851
  by (rule measure_times) simp
d9dfc733f25c add product of probability spaces with finite cardinality
hoelzl
parents: 42858
diff changeset
   852
45777
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
   853
lemma (in finite_product_finite_prob_space) measure_singleton_times:
42892
a61e30bfd0bc add lemma prob_finite_product
hoelzl
parents: 42860
diff changeset
   854
  assumes x: "x \<in> space P" shows "\<mu> {x} = (\<Prod>i\<in>I. M.\<mu> i {x i})"
a61e30bfd0bc add lemma prob_finite_product
hoelzl
parents: 42860
diff changeset
   855
  unfolding singleton_eq_product[OF x] using x
a61e30bfd0bc add lemma prob_finite_product
hoelzl
parents: 42860
diff changeset
   856
  by (intro measure_finite_times) auto
a61e30bfd0bc add lemma prob_finite_product
hoelzl
parents: 42860
diff changeset
   857
45777
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
   858
lemma (in finite_product_finite_prob_space) prob_finite_times:
42859
d9dfc733f25c add product of probability spaces with finite cardinality
hoelzl
parents: 42858
diff changeset
   859
  assumes X: "\<And>i. i \<in> I \<Longrightarrow> X i \<subseteq> space (M i)"
d9dfc733f25c add product of probability spaces with finite cardinality
hoelzl
parents: 42858
diff changeset
   860
  shows "prob (\<Pi>\<^isub>E i\<in>I. X i) = (\<Prod>i\<in>I. M.prob i (X i))"
d9dfc733f25c add product of probability spaces with finite cardinality
hoelzl
parents: 42858
diff changeset
   861
proof -
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43658
diff changeset
   862
  have "ereal (\<mu>' (\<Pi>\<^isub>E i\<in>I. X i)) = \<mu> (\<Pi>\<^isub>E i\<in>I. X i)"
42859
d9dfc733f25c add product of probability spaces with finite cardinality
hoelzl
parents: 42858
diff changeset
   863
    using X by (intro finite_measure_eq[symmetric] in_P) auto
d9dfc733f25c add product of probability spaces with finite cardinality
hoelzl
parents: 42858
diff changeset
   864
  also have "\<dots> = (\<Prod>i\<in>I. M.\<mu> i (X i))"
d9dfc733f25c add product of probability spaces with finite cardinality
hoelzl
parents: 42858
diff changeset
   865
    using measure_finite_times X by simp
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43658
diff changeset
   866
  also have "\<dots> = ereal (\<Prod>i\<in>I. M.\<mu>' i (X i))"
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43658
diff changeset
   867
    using X by (simp add: M.finite_measure_eq setprod_ereal)
42859
d9dfc733f25c add product of probability spaces with finite cardinality
hoelzl
parents: 42858
diff changeset
   868
  finally show ?thesis by simp
d9dfc733f25c add product of probability spaces with finite cardinality
hoelzl
parents: 42858
diff changeset
   869
qed
d9dfc733f25c add product of probability spaces with finite cardinality
hoelzl
parents: 42858
diff changeset
   870
45777
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
   871
lemma (in finite_product_finite_prob_space) prob_singleton_times:
42892
a61e30bfd0bc add lemma prob_finite_product
hoelzl
parents: 42860
diff changeset
   872
  assumes x: "x \<in> space P"
a61e30bfd0bc add lemma prob_finite_product
hoelzl
parents: 42860
diff changeset
   873
  shows "prob {x} = (\<Prod>i\<in>I. M.prob i {x i})"
a61e30bfd0bc add lemma prob_finite_product
hoelzl
parents: 42860
diff changeset
   874
  unfolding singleton_eq_product[OF x] using x
a61e30bfd0bc add lemma prob_finite_product
hoelzl
parents: 42860
diff changeset
   875
  by (intro prob_finite_times) auto
a61e30bfd0bc add lemma prob_finite_product
hoelzl
parents: 42860
diff changeset
   876
45777
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
   877
lemma (in finite_product_finite_prob_space) prob_finite_product:
42892
a61e30bfd0bc add lemma prob_finite_product
hoelzl
parents: 42860
diff changeset
   878
  "A \<subseteq> space P \<Longrightarrow> prob A = (\<Sum>x\<in>A. \<Prod>i\<in>I. M.prob i {x i})"
a61e30bfd0bc add lemma prob_finite_product
hoelzl
parents: 42860
diff changeset
   879
  by (auto simp add: finite_measure_singleton prob_singleton_times
a61e30bfd0bc add lemma prob_finite_product
hoelzl
parents: 42860
diff changeset
   880
           simp del: space_product_algebra
a61e30bfd0bc add lemma prob_finite_product
hoelzl
parents: 42860
diff changeset
   881
           intro!: setsum_cong prob_singleton_times)
a61e30bfd0bc add lemma prob_finite_product
hoelzl
parents: 42860
diff changeset
   882
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   883
lemma (in prob_space) joint_distribution_finite_prob_space:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   884
  assumes X: "finite_random_variable MX X"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   885
  assumes Y: "finite_random_variable MY Y"
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43658
diff changeset
   886
  shows "finite_prob_space ((MX \<Otimes>\<^isub>M MY)\<lparr> measure := ereal \<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
   887
  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
   888
36624
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   889
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
   890
  "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
   891
  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
   892
  by auto
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   893
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   894
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
   895
  using measure_space_1 sum_over_space by simp
36624
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   896
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   897
lemma (in finite_prob_space) joint_distribution_restriction_fst:
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   898
  "joint_distribution X Y A \<le> distribution X (fst ` A)"
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   899
  unfolding distribution_def
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   900
proof (safe intro!: finite_measure_mono)
36624
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   901
  fix x assume "x \<in> space M" and *: "(X x, Y x) \<in> A"
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   902
  show "x \<in> X -` fst ` A"
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   903
    by (auto intro!: image_eqI[OF _ *])
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   904
qed (simp_all add: sets_eq_Pow)
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   905
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   906
lemma (in finite_prob_space) joint_distribution_restriction_snd:
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   907
  "joint_distribution X Y A \<le> distribution Y (snd ` A)"
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   908
  unfolding distribution_def
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   909
proof (safe intro!: finite_measure_mono)
36624
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   910
  fix x assume "x \<in> space M" and *: "(X x, Y x) \<in> A"
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   911
  show "x \<in> Y -` snd ` A"
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   912
    by (auto intro!: image_eqI[OF _ *])
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   913
qed (simp_all add: sets_eq_Pow)
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   914
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   915
lemma (in finite_prob_space) distribution_order:
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   916
  shows "0 \<le> distribution X x'"
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   917
  and "(distribution X x' \<noteq> 0) \<longleftrightarrow> (0 < distribution X x')"
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   918
  and "r \<le> joint_distribution X Y {(x, y)} \<Longrightarrow> r \<le> distribution X {x}"
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   919
  and "r \<le> joint_distribution X Y {(x, y)} \<Longrightarrow> r \<le> distribution Y {y}"
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   920
  and "r < joint_distribution X Y {(x, y)} \<Longrightarrow> r < distribution X {x}"
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   921
  and "r < joint_distribution X Y {(x, y)} \<Longrightarrow> r < distribution Y {y}"
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   922
  and "distribution X {x} = 0 \<Longrightarrow> joint_distribution X Y {(x, y)} = 0"
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   923
  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
   924
  using
36624
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   925
    joint_distribution_restriction_fst[of X Y "{(x, y)}"]
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   926
    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
   927
  by (auto intro: antisym)
36624
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   928
39097
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   929
lemma (in finite_prob_space) distribution_mono:
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   930
  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
   931
  shows "distribution X x \<le> distribution Y y"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   932
  unfolding distribution_def
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   933
  using assms by (auto simp: sets_eq_Pow intro!: finite_measure_mono)
39097
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   934
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   935
lemma (in finite_prob_space) distribution_mono_gt_0:
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   936
  assumes gt_0: "0 < distribution X x"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   937
  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
   938
  shows "0 < distribution Y y"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   939
  by (rule less_le_trans[OF gt_0 distribution_mono]) (rule *)
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   940
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   941
lemma (in finite_prob_space) sum_over_space_distrib:
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   942
  "(\<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
   943
  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
   944
  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
   945
     (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
   946
           intro!: arg_cong[where f=\<mu>'])
39097
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   947
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   948
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
   949
  "(\<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
   950
  using prob_space finite_space
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   951
  by (subst (asm) finite_measure_finite_singleton) auto
39097
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   952
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   953
lemma (in prob_space) distribution_remove_const:
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   954
  shows "joint_distribution X (\<lambda>x. ()) {(x, ())} = distribution X {x}"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   955
  and "joint_distribution (\<lambda>x. ()) X {((), x)} = distribution X {x}"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   956
  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
   957
  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
   958
  and "distribution (\<lambda>x. ()) {()} = 1"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   959
  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
   960
39097
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   961
lemma (in finite_prob_space) setsum_distribution_gen:
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   962
  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
   963
  and "inj_on f (X`space M)"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   964
  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
   965
  unfolding distribution_def assms
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   966
  using finite_space assms
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   967
  by (subst finite_measure_finite_Union[symmetric])
39097
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   968
     (auto simp add: disjoint_family_on_def sets_eq_Pow inj_on_def
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   969
      intro!: arg_cong[where f=prob])
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   970
45711
a2c32e196d49 rename finite_prob_space.setsum_distribution, it collides with prob_space.setsum_distribution
hoelzl
parents: 44890
diff changeset
   971
lemma (in finite_prob_space) setsum_distribution_cut:
39097
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   972
  "(\<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
   973
  "(\<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
   974
  "(\<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
   975
  "(\<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
   976
  "(\<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
   977
  by (auto intro!: inj_onI setsum_distribution_gen)
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   978
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   979
lemma (in finite_prob_space) uniform_prob:
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   980
  assumes "x \<in> space M"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   981
  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
   982
  shows "prob {x} = 1 / card (space M)"
39097
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   983
proof -
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   984
  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
   985
    using assms(2)[OF _ `x \<in> space M`] by blast
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   986
  have "1 = prob (space M)"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   987
    using prob_space by auto
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   988
  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
   989
    using finite_measure_finite_Union[of "space M" "\<lambda> x. {x}", simplified]
39097
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   990
      sets_eq_Pow inj_singleton[unfolded inj_on_def, rule_format]
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   991
      finite_space unfolding disjoint_family_on_def  prob_space[symmetric]
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   992
    by (auto simp add:setsum_restrict_set)
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   993
  also have "\<dots> = (\<Sum> y \<in> space M. prob {x})"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   994
    using prob_x by auto
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   995
  also have "\<dots> = real_of_nat (card (space M)) * prob {x}" by simp
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   996
  finally have one: "1 = real (card (space M)) * prob {x}"
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   997
    using real_eq_of_nat by auto
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 43920
diff changeset
   998
  hence two: "real (card (space M)) \<noteq> 0" by fastforce
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 43920
diff changeset
   999
  from one have three: "prob {x} \<noteq> 0" by fastforce
39097
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
  1000
  thus ?thesis using one two three divide_cancel_right
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
  1001
    by (auto simp:field_simps)
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39091
diff changeset
  1002
qed
35977
30d42bfd0174 Added finite measure space.
hoelzl
parents: 35929
diff changeset
  1003
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39091
diff changeset
  1004
lemma (in prob_space) prob_space_subalgebra:
41545
9c869baf1c66 tuned formalization of subalgebra
hoelzl
parents: 41544
diff changeset
  1005
  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
  1006
    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
  1007
  shows "prob_space N"
45777
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
  1008
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
  1009
  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
  1010
    by (rule measure_space_subalgebra[OF assms])
45777
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
  1011
  show "measure_space N" ..
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
  1012
  show "measure N (space N) = 1"
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
  1013
    using assms(4)[OF N.top] by (simp add: assms measure_space_1)
35977
30d42bfd0174 Added finite measure space.
hoelzl
parents: 35929
diff changeset
  1014
qed
30d42bfd0174 Added finite measure space.
hoelzl
parents: 35929
diff changeset
  1015
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39091
diff changeset
  1016
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
  1017
  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
  1018
  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
  1019
    (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
  1020
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
  1021
  interpret A: measure_space "restricted_space A"
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39091
diff changeset
  1022
    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
  1023
  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
  1024
    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
  1025
  show "prob_space ?P"
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39091
diff changeset
  1026
  proof
45777
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
  1027
    show "measure_space ?P"
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
  1028
    proof
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
  1029
      show "positive ?P (measure ?P)"
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
  1030
      proof (simp add: positive_def, safe)
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
  1031
        fix B assume "B \<in> events"
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
  1032
        with real_measure[of "A \<inter> B"] real_measure[OF `A \<in> events`] `A \<in> sets M`
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
  1033
        show "0 \<le> \<mu> (A \<inter> B) / \<mu> A" by (auto simp: Int)
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
  1034
      qed
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
  1035
      show "countably_additive ?P (measure ?P)"
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
  1036
      proof (simp add: countably_additive_def, safe)
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
  1037
        fix B and F :: "nat \<Rightarrow> 'a set"
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
  1038
        assume F: "range F \<subseteq> op \<inter> A ` events" "disjoint_family F"
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
  1039
        { fix i
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
  1040
          from F have "F i \<in> op \<inter> A ` events" by auto
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
  1041
          with `A \<in> events` have "F i \<in> events" by auto }
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
  1042
        moreover then have "range F \<subseteq> events" by auto
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
  1043
        moreover have "\<And>S. \<mu> S / \<mu> A = inverse (\<mu> A) * \<mu> S"
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
  1044
          by (simp add: mult_commute divide_ereal_def)
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
  1045
        moreover have "0 \<le> inverse (\<mu> A)"
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
  1046
          using real_measure[OF `A \<in> events`] by auto
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
  1047
        ultimately show "(\<Sum>i. \<mu> (F i) / \<mu> A) = \<mu> (\<Union>i. F i) / \<mu> A"
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
  1048
          using measure_countably_additive[of F] F
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
  1049
          by (auto simp: suminf_cmult_ereal)
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
  1050
      qed
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
  1051
    qed
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
  1052
    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
  1053
      using real_measure[OF `A \<in> events`] `\<mu> A \<noteq> 0` by auto
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39091
diff changeset
  1054
  qed
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39091
diff changeset
  1055
qed
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39091
diff changeset
  1056
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39091
diff changeset
  1057
lemma finite_prob_spaceI:
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1058
  assumes "finite (space M)" "sets M = Pow(space M)"
45777
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
  1059
    and 1: "measure M (space M) = 1" and "\<And>x. x \<in> space M \<Longrightarrow> 0 \<le> measure M {x}"
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
  1060
    and add: "\<And>A B. A \<subseteq> space M \<Longrightarrow> measure M A = (\<Sum>x\<in>A. measure M {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
  1061
  shows "finite_prob_space M"
45777
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
  1062
proof -
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
  1063
  interpret finite_measure_space M
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
  1064
  proof
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
  1065
    show "measure M (space M) \<noteq> \<infinity>" using 1 by simp
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
  1066
  qed fact+
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
  1067
  show ?thesis by default fact
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
  1068
qed
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
  1069
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
  1070
lemma (in finite_prob_space) distribution_eq_setsum:
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
  1071
  "distribution X A = (\<Sum>x\<in>A \<inter> X ` space M. distribution X {x})"
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
  1072
proof -
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
  1073
  have *: "X -` A \<inter> space M = (\<Union>x\<in>A \<inter> X ` space M. X -` {x} \<inter> space M)"
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
  1074
    by auto
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
  1075
  then show "distribution X A = (\<Sum>x\<in>A \<inter> X ` space M. distribution X {x})"
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
  1076
    using finite_space unfolding distribution_def *
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
  1077
    by (intro finite_measure_finite_Union)
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
  1078
       (auto simp: disjoint_family_on_def)
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
  1079
qed
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
  1080
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
  1081
lemma (in finite_prob_space) distribution_eq_setsum_finite:
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
  1082
  assumes "finite A"
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
  1083
  shows "distribution X A = (\<Sum>x\<in>A. distribution X {x})"
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
  1084
proof -
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
  1085
  note distribution_eq_setsum[of X A]
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
  1086
  also have "(\<Sum>x\<in>A \<inter> X ` space M. distribution X {x}) = (\<Sum>x\<in>A. distribution X {x})"
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
  1087
  proof (intro setsum_mono_zero_cong_left ballI)
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
  1088
    fix i assume "i\<in>A - A \<inter> X ` space M"
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
  1089
    then have "X -` {i} \<inter> space M = {}" by auto
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
  1090
    then show "distribution X {i} = 0"
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
  1091
      by (simp add: distribution_def)
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
  1092
  next
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
  1093
    show "finite A" by fact
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
  1094
  qed simp_all
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
  1095
  finally show ?thesis .
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39091
diff changeset
  1096
qed
36624
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
  1097
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
  1098
lemma (in finite_prob_space) finite_measure_space:
39097
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
  1099
  fixes X :: "'a \<Rightarrow> 'x"
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43658
diff changeset
  1100
  shows "finite_measure_space \<lparr>space = X ` space M, sets = Pow (X ` space M), measure = ereal \<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
  1101
    (is "finite_measure_space ?S")
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39091
diff changeset
  1102
proof (rule finite_measure_spaceI, simp_all)
36624
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
  1103
  show "finite (X ` space M)" using finite_space by simp
39097
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
  1104
next
45777
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
  1105
  fix A assume "A \<subseteq> X ` space M"
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
  1106
  then show "distribution X A = (\<Sum>x\<in>A. distribution X {x})"
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
  1107
    by (subst distribution_eq_setsum) (simp add: Int_absorb2)
36624
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
  1108
qed
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
  1109
39097
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
  1110
lemma (in finite_prob_space) finite_prob_space_of_images:
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43658
diff changeset
  1111
  "finite_prob_space \<lparr> space = X ` space M, sets = Pow (X ` space M), measure = ereal \<circ> distribution X \<rparr>"
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43658
diff changeset
  1112
  by (simp add: finite_prob_space_eq finite_measure_space measure_space_1 one_ereal_def)
39097
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
  1113
39096
hoelzl
parents: 39092
diff changeset
  1114
lemma (in finite_prob_space) finite_product_measure_space:
39097
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
  1115
  fixes X :: "'a \<Rightarrow> 'x" and Y :: "'a \<Rightarrow> 'y"
39096
hoelzl
parents: 39092
diff changeset
  1116
  assumes "finite s1" "finite s2"
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43658
diff changeset
  1117
  shows "finite_measure_space \<lparr> space = s1 \<times> s2, sets = Pow (s1 \<times> s2), measure = ereal \<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
  1118
    (is "finite_measure_space ?M")
39097
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
  1119
proof (rule finite_measure_spaceI, simp_all)
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
  1120
  show "finite (s1 \<times> s2)"
39096
hoelzl
parents: 39092
diff changeset
  1121
    using assms by auto
39097
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
  1122
next
45777
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
  1123
  fix A assume "A \<subseteq> (s1 \<times> s2)"
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
  1124
  with assms show "joint_distribution X Y A = (\<Sum>x\<in>A. joint_distribution X Y {x})"
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
  1125
    by (intro distribution_eq_setsum_finite) (auto dest: finite_subset)
39096
hoelzl
parents: 39092
diff changeset
  1126
qed
hoelzl
parents: 39092
diff changeset
  1127
39097
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
  1128
lemma (in finite_prob_space) finite_product_measure_space_of_images:
39096
hoelzl
parents: 39092
diff changeset
  1129
  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
  1130
                                sets = Pow (X ` space M \<times> Y ` space M),
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43658
diff changeset
  1131
                                measure = ereal \<circ> joint_distribution X Y \<rparr>"
39096
hoelzl
parents: 39092
diff changeset
  1132
  using finite_space by (auto intro!: finite_product_measure_space)
hoelzl
parents: 39092
diff changeset
  1133
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
  1134
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
  1135
  "finite_prob_space \<lparr> space = X ` space M \<times> Y ` space M, sets = Pow (X ` space M \<times> Y ` space M),
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43658
diff changeset
  1136
                       measure = ereal \<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
  1137
  (is "finite_prob_space ?S")
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43658
diff changeset
  1138
proof (simp add: finite_prob_space_eq finite_product_measure_space_of_images one_ereal_def)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
  1139
  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
  1140
  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
  1141
    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
  1142
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
  1143
43658
0d96ec6ec33b the borel probability measure is easier to handle with {0 ..< 1} (coverable by disjoint intervals {_ ..< _})
hoelzl
parents: 43556
diff changeset
  1144
subsection "Borel Measure on {0 ..< 1}"
42902
e8dbf90a2f3b Add restricted borel measure to {0 .. 1}
hoelzl
parents: 42892
diff changeset
  1145
e8dbf90a2f3b Add restricted borel measure to {0 .. 1}
hoelzl
parents: 42892
diff changeset
  1146
definition pborel :: "real measure_space" where
43658
0d96ec6ec33b the borel probability measure is easier to handle with {0 ..< 1} (coverable by disjoint intervals {_ ..< _})
hoelzl
parents: 43556
diff changeset
  1147
  "pborel = lborel.restricted_space {0 ..< 1}"
42902
e8dbf90a2f3b Add restricted borel measure to {0 .. 1}
hoelzl
parents: 42892
diff changeset
  1148
e8dbf90a2f3b Add restricted borel measure to {0 .. 1}
hoelzl
parents: 42892
diff changeset
  1149
lemma space_pborel[simp]:
43658
0d96ec6ec33b the borel probability measure is easier to handle with {0 ..< 1} (coverable by disjoint intervals {_ ..< _})
hoelzl
parents: 43556
diff changeset
  1150
  "space pborel = {0 ..< 1}"
42902
e8dbf90a2f3b Add restricted borel measure to {0 .. 1}
hoelzl
parents: 42892
diff changeset
  1151
  unfolding pborel_def by auto
e8dbf90a2f3b Add restricted borel measure to {0 .. 1}
hoelzl
parents: 42892
diff changeset
  1152
e8dbf90a2f3b Add restricted borel measure to {0 .. 1}
hoelzl
parents: 42892
diff changeset
  1153
lemma sets_pborel:
43658
0d96ec6ec33b the borel probability measure is easier to handle with {0 ..< 1} (coverable by disjoint intervals {_ ..< _})
hoelzl
parents: 43556
diff changeset
  1154
  "A \<in> sets pborel \<longleftrightarrow> A \<in> sets borel \<and> A \<subseteq> { 0 ..< 1}"
42902
e8dbf90a2f3b Add restricted borel measure to {0 .. 1}
hoelzl
parents: 42892
diff changeset
  1155
  unfolding pborel_def by auto
e8dbf90a2f3b Add restricted borel measure to {0 .. 1}
hoelzl
parents: 42892
diff changeset
  1156
e8dbf90a2f3b Add restricted borel measure to {0 .. 1}
hoelzl
parents: 42892
diff changeset
  1157
lemma in_pborel[intro, simp]:
43658
0d96ec6ec33b the borel probability measure is easier to handle with {0 ..< 1} (coverable by disjoint intervals {_ ..< _})
hoelzl
parents: 43556
diff changeset
  1158
  "A \<subseteq> {0 ..< 1} \<Longrightarrow> A \<in> sets borel \<Longrightarrow> A \<in> sets pborel"
42902
e8dbf90a2f3b Add restricted borel measure to {0 .. 1}
hoelzl
parents: 42892
diff changeset
  1159
  unfolding pborel_def by auto
e8dbf90a2f3b Add restricted borel measure to {0 .. 1}
hoelzl
parents: 42892
diff changeset
  1160
e8dbf90a2f3b Add restricted borel measure to {0 .. 1}
hoelzl
parents: 42892
diff changeset
  1161
interpretation pborel: measure_space pborel
43658
0d96ec6ec33b the borel probability measure is easier to handle with {0 ..< 1} (coverable by disjoint intervals {_ ..< _})
hoelzl
parents: 43556
diff changeset
  1162
  using lborel.restricted_measure_space[of "{0 ..< 1}"]
42902
e8dbf90a2f3b Add restricted borel measure to {0 .. 1}
hoelzl
parents: 42892
diff changeset
  1163
  by (simp add: pborel_def)
e8dbf90a2f3b Add restricted borel measure to {0 .. 1}
hoelzl
parents: 42892
diff changeset
  1164
e8dbf90a2f3b Add restricted borel measure to {0 .. 1}
hoelzl
parents: 42892
diff changeset
  1165
interpretation pborel: prob_space pborel
45777
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
  1166
proof
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
  1167
  show "measure pborel (space pborel) = 1"
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
  1168
    by (simp add: one_ereal_def pborel_def)
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 45712
diff changeset
  1169
qed default
42902
e8dbf90a2f3b Add restricted borel measure to {0 .. 1}
hoelzl
parents: 42892
diff changeset
  1170
43658
0d96ec6ec33b the borel probability measure is easier to handle with {0 ..< 1} (coverable by disjoint intervals {_ ..< _})
hoelzl
parents: 43556
diff changeset
  1171
lemma pborel_prob: "pborel.prob A = (if A \<in> sets borel \<and> A \<subseteq> {0 ..< 1} then real (lborel.\<mu> A) else 0)"
42902
e8dbf90a2f3b Add restricted borel measure to {0 .. 1}
hoelzl
parents: 42892
diff changeset
  1172
  unfolding pborel.\<mu>'_def by (auto simp: pborel_def)
e8dbf90a2f3b Add restricted borel measure to {0 .. 1}
hoelzl
parents: 42892
diff changeset
  1173
e8dbf90a2f3b Add restricted borel measure to {0 .. 1}
hoelzl
parents: 42892
diff changeset
  1174
lemma pborel_singelton[simp]: "pborel.prob {a} = 0"
e8dbf90a2f3b Add restricted borel measure to {0 .. 1}
hoelzl
parents: 42892
diff changeset
  1175
  by (auto simp: pborel_prob)
e8dbf90a2f3b Add restricted borel measure to {0 .. 1}
hoelzl
parents: 42892
diff changeset
  1176
e8dbf90a2f3b Add restricted borel measure to {0 .. 1}
hoelzl
parents: 42892
diff changeset
  1177
lemma
43658
0d96ec6ec33b the borel probability measure is easier to handle with {0 ..< 1} (coverable by disjoint intervals {_ ..< _})
hoelzl
parents: 43556
diff changeset
  1178
  shows pborel_atLeastAtMost[simp]: "pborel.\<mu>' {a .. b} = (if 0 \<le> a \<and> a \<le> b \<and> b < 1 then b - a else 0)"
42902
e8dbf90a2f3b Add restricted borel measure to {0 .. 1}
hoelzl
parents: 42892
diff changeset
  1179
    and pborel_atLeastLessThan[simp]: "pborel.\<mu>' {a ..< b} = (if 0 \<le> a \<and> a \<le> b \<and> b \<le> 1 then b - a else 0)"
43658
0d96ec6ec33b the borel probability measure is easier to handle with {0 ..< 1} (coverable by disjoint intervals {_ ..< _})
hoelzl
parents: 43556
diff changeset
  1180
    and pborel_greaterThanAtMost[simp]: "pborel.\<mu>' {a <.. b} = (if 0 \<le> a \<and> a \<le> b \<and> b < 1 then b - a else 0)"
42902
e8dbf90a2f3b Add restricted borel measure to {0 .. 1}
hoelzl
parents: 42892
diff changeset
  1181
    and pborel_greaterThanLessThan[simp]: "pborel.\<mu>' {a <..< b} = (if 0 \<le> a \<and> a \<le> b \<and> b \<le> 1 then b - a else 0)"
43658
0d96ec6ec33b the borel probability measure is easier to handle with {0 ..< 1} (coverable by disjoint intervals {_ ..< _})
hoelzl
parents: 43556
diff changeset
  1182
  unfolding pborel_prob
0d96ec6ec33b the borel probability measure is easier to handle with {0 ..< 1} (coverable by disjoint intervals {_ ..< _})
hoelzl
parents: 43556
diff changeset
  1183
  by (auto simp: atLeastAtMost_subseteq_atLeastLessThan_iff
0d96ec6ec33b the borel probability measure is easier to handle with {0 ..< 1} (coverable by disjoint intervals {_ ..< _})
hoelzl
parents: 43556
diff changeset
  1184
    greaterThanAtMost_subseteq_atLeastLessThan_iff greaterThanLessThan_subseteq_atLeastLessThan_iff)
42902
e8dbf90a2f3b Add restricted borel measure to {0 .. 1}
hoelzl
parents: 42892
diff changeset
  1185
e8dbf90a2f3b Add restricted borel measure to {0 .. 1}
hoelzl
parents: 42892
diff changeset
  1186
lemma pborel_lebesgue_measure:
e8dbf90a2f3b Add restricted borel measure to {0 .. 1}
hoelzl
parents: 42892
diff changeset
  1187
  "A \<in> sets pborel \<Longrightarrow> pborel.prob A = real (measure lebesgue A)"
e8dbf90a2f3b Add restricted borel measure to {0 .. 1}
hoelzl
parents: 42892
diff changeset
  1188
  by (simp add: sets_pborel pborel_prob)
e8dbf90a2f3b Add restricted borel measure to {0 .. 1}
hoelzl
parents: 42892
diff changeset
  1189
e8dbf90a2f3b Add restricted borel measure to {0 .. 1}
hoelzl
parents: 42892
diff changeset
  1190
lemma pborel_alt:
e8dbf90a2f3b Add restricted borel measure to {0 .. 1}
hoelzl
parents: 42892
diff changeset
  1191
  "pborel = sigma \<lparr>
43658
0d96ec6ec33b the borel probability measure is easier to handle with {0 ..< 1} (coverable by disjoint intervals {_ ..< _})
hoelzl
parents: 43556
diff changeset
  1192
    space = {0..<1},
0d96ec6ec33b the borel probability measure is easier to handle with {0 ..< 1} (coverable by disjoint intervals {_ ..< _})
hoelzl
parents: 43556
diff changeset
  1193
    sets = range (\<lambda>(x,y). {x..<y} \<inter> {0..<1}),
42902
e8dbf90a2f3b Add restricted borel measure to {0 .. 1}
hoelzl
parents: 42892
diff changeset
  1194
    measure = measure lborel \<rparr>" (is "_ = ?R")
e8dbf90a2f3b Add restricted borel measure to {0 .. 1}
hoelzl
parents: 42892
diff changeset
  1195
proof -
43658
0d96ec6ec33b the borel probability measure is easier to handle with {0 ..< 1} (coverable by disjoint intervals {_ ..< _})
hoelzl
parents: 43556
diff changeset
  1196
  have *: "{0..<1::real} \<in> sets borel" by auto
0d96ec6ec33b the borel probability measure is easier to handle with {0 ..< 1} (coverable by disjoint intervals {_ ..< _})
hoelzl
parents: 43556
diff changeset
  1197
  have **: "op \<inter> {0..<1::real} ` range (\<lambda>(x, y). {x..<y}) = range (\<lambda>(x,y). {x..<y} \<inter> {0..<1})"
42902
e8dbf90a2f3b Add restricted borel measure to {0 .. 1}
hoelzl
parents: 42892
diff changeset
  1198
    unfolding image_image by (intro arg_cong[where f=range]) auto
43658
0d96ec6ec33b the borel probability measure is easier to handle with {0 ..< 1} (coverable by disjoint intervals {_ ..< _})
hoelzl
parents: 43556
diff changeset
  1199
  have "pborel = algebra.restricted_space (sigma \<lparr>space=UNIV, sets=range (\<lambda>(a, b). {a ..< b :: real}),
0d96ec6ec33b the borel probability measure is easier to handle with {0 ..< 1} (coverable by disjoint intervals {_ ..< _})
hoelzl
parents: 43556
diff changeset
  1200
    measure = measure pborel\<rparr>) {0 ..< 1}"
0d96ec6ec33b the borel probability measure is easier to handle with {0 ..< 1} (coverable by disjoint intervals {_ ..< _})
hoelzl
parents: 43556
diff changeset
  1201
    by (simp add: sigma_def lebesgue_def pborel_def borel_eq_atLeastLessThan lborel_def)
42902
e8dbf90a2f3b Add restricted borel measure to {0 .. 1}
hoelzl
parents: 42892
diff changeset
  1202
  also have "\<dots> = ?R"
e8dbf90a2f3b Add restricted borel measure to {0 .. 1}
hoelzl
parents: 42892
diff changeset
  1203
    by (subst restricted_sigma)
e8dbf90a2f3b Add restricted borel measure to {0 .. 1}
hoelzl
parents: 42892
diff changeset
  1204
       (simp_all add: sets_sigma sigma_sets.Basic ** pborel_def image_eqI[of _ _ "(0,1)"])
e8dbf90a2f3b Add restricted borel measure to {0 .. 1}
hoelzl
parents: 42892
diff changeset
  1205
  finally show ?thesis .
e8dbf90a2f3b Add restricted borel measure to {0 .. 1}
hoelzl
parents: 42892
diff changeset
  1206
qed
e8dbf90a2f3b Add restricted borel measure to {0 .. 1}
hoelzl
parents: 42892
diff changeset
  1207
42860
b02349e70d5a add Bernoulli space
hoelzl
parents: 42859
diff changeset
  1208
subsection "Bernoulli space"
b02349e70d5a add Bernoulli space
hoelzl
parents: 42859
diff changeset
  1209
b02349e70d5a add Bernoulli space
hoelzl
parents: 42859
diff changeset
  1210
definition "bernoulli_space p = \<lparr> space = UNIV, sets = UNIV,
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43658
diff changeset
  1211
  measure = ereal \<circ> setsum (\<lambda>b. if b then min 1 (max 0 p) else 1 - min 1 (max 0 p)) \<rparr>"
42860
b02349e70d5a add Bernoulli space
hoelzl
parents: 42859
diff changeset
  1212
b02349e70d5a add Bernoulli space
hoelzl
parents: 42859
diff changeset
  1213
interpretation bernoulli: finite_prob_space "bernoulli_space p" for p
b02349e70d5a add Bernoulli space
hoelzl
parents: 42859
diff changeset
  1214
  by (rule finite_prob_spaceI)
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43658
diff changeset
  1215
     (auto simp: bernoulli_space_def UNIV_bool one_ereal_def setsum_Un_disjoint intro!: setsum_nonneg)
42860
b02349e70d5a add Bernoulli space
hoelzl
parents: 42859
diff changeset
  1216
b02349e70d5a add Bernoulli space
hoelzl
parents: 42859
diff changeset
  1217
lemma bernoulli_measure:
b02349e70d5a add Bernoulli space
hoelzl
parents: 42859
diff changeset
  1218
  "0 \<le> p \<Longrightarrow> p \<le> 1 \<Longrightarrow> bernoulli.prob p B = (\<Sum>b\<in>B. if b then p else 1 - p)"
b02349e70d5a add Bernoulli space
hoelzl
parents: 42859
diff changeset
  1219
  unfolding bernoulli.\<mu>'_def unfolding bernoulli_space_def by (auto intro!: setsum_cong)
b02349e70d5a add Bernoulli space
hoelzl
parents: 42859
diff changeset
  1220
b02349e70d5a add Bernoulli space
hoelzl
parents: 42859
diff changeset
  1221
lemma bernoulli_measure_True: "0 \<le> p \<Longrightarrow> p \<le> 1 \<Longrightarrow> bernoulli.prob p {True} = p"
b02349e70d5a add Bernoulli space
hoelzl
parents: 42859
diff changeset
  1222
  and bernoulli_measure_False: "0 \<le> p \<Longrightarrow> p \<le> 1 \<Longrightarrow> bernoulli.prob p {False} = 1 - p"
b02349e70d5a add Bernoulli space
hoelzl
parents: 42859
diff changeset
  1223
  unfolding bernoulli_measure by simp_all
b02349e70d5a add Bernoulli space
hoelzl
parents: 42859
diff changeset
  1224
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1225
end