src/HOL/Probability/Probability_Measure.thy
author hoelzl
Fri, 30 Sep 2016 16:08:38 +0200
changeset 64008 17a20ca86d62
parent 63886 685fb01256af
child 64267 b9a1486e79be
permissions -rw-r--r--
HOL-Probability: more about probability, prepare for Markov processes in the AFP
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
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61605
diff changeset
     6
section \<open>Probability measure\<close>
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
63627
6ddb43c6b711 rename HOL-Multivariate_Analysis to HOL-Analysis.
hoelzl
parents: 63626
diff changeset
     9
  imports "~~/src/HOL/Analysis/Analysis"
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 +
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
    13
  assumes emeasure_space_1: "emeasure 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!]:
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
    16
  assumes *: "emeasure M (space M) = 1"
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
    17
  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
    18
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
    19
  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
    20
  proof
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
    21
    show "emeasure M (space M) \<noteq> \<infinity>" using * by simp
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
    22
  qed
61169
4de9ff3ea29a tuned proofs -- less legacy;
wenzelm
parents: 61125
diff changeset
    23
  show "prob_space M" by standard fact
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
    24
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
    25
59425
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59356
diff changeset
    26
lemma prob_space_imp_sigma_finite: "prob_space M \<Longrightarrow> sigma_finite_measure M"
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59356
diff changeset
    27
  unfolding prob_space_def finite_measure_def by simp
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59356
diff changeset
    28
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
    29
abbreviation (in prob_space) "events \<equiv> sets M"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
    30
abbreviation (in prob_space) "prob \<equiv> measure M"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
    31
abbreviation (in prob_space) "random_variable M' X \<equiv> X \<in> measurable M M'"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51683
diff changeset
    32
abbreviation (in prob_space) "expectation \<equiv> integral\<^sup>L M"
57235
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
    33
abbreviation (in prob_space) "variance X \<equiv> integral\<^sup>L M (\<lambda>x. (X x - expectation X)\<^sup>2)"
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
    34
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57418
diff changeset
    35
lemma (in prob_space) finite_measure [simp]: "finite_measure M"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57418
diff changeset
    36
  by unfold_locales
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57418
diff changeset
    37
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
    38
lemma (in prob_space) prob_space_distr:
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
    39
  assumes f: "f \<in> measurable M M'" shows "prob_space (distr M M' f)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
    40
proof (rule prob_spaceI)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
    41
  have "f -` space M' \<inter> space M = space M" using f by (auto dest: measurable_space)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
    42
  with f show "emeasure (distr M M' f) (space (distr M M' f)) = 1"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
    43
    by (auto simp: emeasure_distr emeasure_space_1)
43339
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
    44
qed
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
    45
61359
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
    46
lemma prob_space_distrD:
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
    47
  assumes f: "f \<in> measurable M N" and M: "prob_space (distr M N f)" shows "prob_space M"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
    48
proof
61605
1bf7b186542e qualifier is mandatory by default;
wenzelm
parents: 61565
diff changeset
    49
  interpret M: prob_space "distr M N f" by fact
61359
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
    50
  have "f -` space N \<inter> space M = space M"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
    51
    using f[THEN measurable_space] by auto
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
    52
  then show "emeasure M (space M) = 1"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
    53
    using M.emeasure_space_1 by (simp add: emeasure_distr[OF f])
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
    54
qed
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
    55
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
    56
lemma (in prob_space) prob_space: "prob (space M) = 1"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
    57
  using emeasure_space_1 unfolding measure_def by (simp add: one_ennreal.rep_eq)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    58
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    59
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
    60
  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
    61
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
    62
lemma (in prob_space) not_empty: "space M \<noteq> {}"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
    63
  using prob_space by auto
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
    64
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
    65
lemma (in prob_space) emeasure_eq_1_AE:
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
    66
  "S \<in> sets M \<Longrightarrow> AE x in M. x \<in> S \<Longrightarrow> emeasure M S = 1"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
    67
  by (subst emeasure_eq_AE[where B="space M"]) (auto simp: emeasure_space_1)
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
    68
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
    69
lemma (in prob_space) emeasure_le_1: "emeasure M S \<le> 1"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
    70
  unfolding ennreal_1[symmetric] emeasure_eq_measure by (subst ennreal_le_iff) auto
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
    71
63099
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63040
diff changeset
    72
lemma (in prob_space) emeasure_ge_1_iff: "emeasure M A \<ge> 1 \<longleftrightarrow> emeasure M A = 1"
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63040
diff changeset
    73
  by (rule iffI, intro antisym emeasure_le_1) simp_all
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63040
diff changeset
    74
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
    75
lemma (in prob_space) AE_iff_emeasure_eq_1:
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
    76
  assumes [measurable]: "Measurable.pred M P"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
    77
  shows "(AE x in M. P x) \<longleftrightarrow> emeasure M {x\<in>space M. P x} = 1"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
    78
proof -
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
    79
  have *: "{x \<in> space M. \<not> P x} = space M - {x\<in>space M. P x}"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
    80
    by auto
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
    81
  show ?thesis
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
    82
    by (auto simp add: ennreal_minus_eq_0 * emeasure_compl emeasure_space_1 AE_iff_measurable[OF _ refl]
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
    83
             intro: antisym emeasure_le_1)
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
    84
qed
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
    85
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
    86
lemma (in prob_space) measure_le_1: "emeasure M X \<le> 1"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
    87
  using emeasure_space[of M X] by (simp add: emeasure_space_1)
42950
6e5c2a3c69da move lemmas to Extended_Reals and Extended_Real_Limits
hoelzl
parents: 42902
diff changeset
    88
63099
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63040
diff changeset
    89
lemma (in prob_space) measure_ge_1_iff: "measure M A \<ge> 1 \<longleftrightarrow> measure M A = 1"
63626
44ce6b524ff3 move measure theory from HOL-Probability to HOL-Multivariate_Analysis
hoelzl
parents: 63099
diff changeset
    90
  by (auto intro!: antisym)
63099
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63040
diff changeset
    91
43339
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
    92
lemma (in prob_space) AE_I_eq_1:
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
    93
  assumes "emeasure M {x\<in>space M. P x} = 1" "{x\<in>space M. P x} \<in> sets M"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
    94
  shows "AE x in M. P x"
43339
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
    95
proof (rule AE_I)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
    96
  show "emeasure M (space M - {x \<in> space M. P x}) = 0"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
    97
    using assms emeasure_space_1 by (simp add: emeasure_compl)
43339
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
    98
qed (insert assms, auto)
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
    99
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
   100
lemma prob_space_restrict_space:
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
   101
  "S \<in> sets M \<Longrightarrow> emeasure M S = 1 \<Longrightarrow> prob_space (restrict_space M S)"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
   102
  by (intro prob_spaceI)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
   103
     (simp add: emeasure_restrict_space space_restrict_space)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
   104
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   105
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
   106
  assumes A: "A \<in> events"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36624
diff changeset
   107
  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
   108
  using finite_measure_compl[OF A] by (simp add: prob_space)
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
   109
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   110
lemma (in prob_space) AE_in_set_eq_1:
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   111
  assumes A[measurable]: "A \<in> events" shows "(AE x in M. x \<in> A) \<longleftrightarrow> prob A = 1"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   112
proof -
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   113
  have *: "{x\<in>space M. x \<in> A} = A"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   114
    using A[THEN sets.sets_into_space] by auto
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   115
  show ?thesis
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   116
    by (subst AE_iff_emeasure_eq_1) (auto simp: emeasure_eq_measure *)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   117
qed
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   118
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   119
lemma (in prob_space) AE_False: "(AE x in M. False) \<longleftrightarrow> False"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   120
proof
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   121
  assume "AE x in M. False"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   122
  then have "AE x in M. x \<in> {}" by simp
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   123
  then show False
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   124
    by (subst (asm) AE_in_set_eq_1) auto
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   125
qed simp
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   126
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   127
lemma (in prob_space) AE_prob_1:
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   128
  assumes "prob A = 1" shows "AE x in M. x \<in> A"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   129
proof -
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61605
diff changeset
   130
  from \<open>prob A = 1\<close> have "A \<in> events"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   131
    by (metis measure_notin_sets zero_neq_one)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   132
  with AE_in_set_eq_1 assms show ?thesis by simp
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   133
qed
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   134
50098
98abff4a775b rules for AE and prob
hoelzl
parents: 50003
diff changeset
   135
lemma (in prob_space) AE_const[simp]: "(AE x in M. P) \<longleftrightarrow> P"
98abff4a775b rules for AE and prob
hoelzl
parents: 50003
diff changeset
   136
  by (cases P) (auto simp: AE_False)
98abff4a775b rules for AE and prob
hoelzl
parents: 50003
diff changeset
   137
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
   138
lemma (in prob_space) ae_filter_bot: "ae_filter M \<noteq> bot"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
   139
  by (simp add: trivial_limit_def)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
   140
50098
98abff4a775b rules for AE and prob
hoelzl
parents: 50003
diff changeset
   141
lemma (in prob_space) AE_contr:
98abff4a775b rules for AE and prob
hoelzl
parents: 50003
diff changeset
   142
  assumes ae: "AE \<omega> in M. P \<omega>" "AE \<omega> in M. \<not> P \<omega>"
98abff4a775b rules for AE and prob
hoelzl
parents: 50003
diff changeset
   143
  shows False
98abff4a775b rules for AE and prob
hoelzl
parents: 50003
diff changeset
   144
proof -
98abff4a775b rules for AE and prob
hoelzl
parents: 50003
diff changeset
   145
  from ae have "AE \<omega> in M. False" by eventually_elim auto
98abff4a775b rules for AE and prob
hoelzl
parents: 50003
diff changeset
   146
  then show False by auto
98abff4a775b rules for AE and prob
hoelzl
parents: 50003
diff changeset
   147
qed
98abff4a775b rules for AE and prob
hoelzl
parents: 50003
diff changeset
   148
57025
e7fd64f82876 add various lemmas
hoelzl
parents: 56996
diff changeset
   149
lemma (in prob_space) integral_ge_const:
e7fd64f82876 add various lemmas
hoelzl
parents: 56996
diff changeset
   150
  fixes c :: real
e7fd64f82876 add various lemmas
hoelzl
parents: 56996
diff changeset
   151
  shows "integrable M f \<Longrightarrow> (AE x in M. c \<le> f x) \<Longrightarrow> c \<le> (\<integral>x. f x \<partial>M)"
e7fd64f82876 add various lemmas
hoelzl
parents: 56996
diff changeset
   152
  using integral_mono_AE[of M "\<lambda>x. c" f] prob_space by simp
e7fd64f82876 add various lemmas
hoelzl
parents: 56996
diff changeset
   153
e7fd64f82876 add various lemmas
hoelzl
parents: 56996
diff changeset
   154
lemma (in prob_space) integral_le_const:
e7fd64f82876 add various lemmas
hoelzl
parents: 56996
diff changeset
   155
  fixes c :: real
e7fd64f82876 add various lemmas
hoelzl
parents: 56996
diff changeset
   156
  shows "integrable M f \<Longrightarrow> (AE x in M. f x \<le> c) \<Longrightarrow> (\<integral>x. f x \<partial>M) \<le> c"
e7fd64f82876 add various lemmas
hoelzl
parents: 56996
diff changeset
   157
  using integral_mono_AE[of M f "\<lambda>x. c"] prob_space by simp
e7fd64f82876 add various lemmas
hoelzl
parents: 56996
diff changeset
   158
e7fd64f82876 add various lemmas
hoelzl
parents: 56996
diff changeset
   159
lemma (in prob_space) nn_integral_ge_const:
e7fd64f82876 add various lemmas
hoelzl
parents: 56996
diff changeset
   160
  "(AE x in M. c \<le> f x) \<Longrightarrow> c \<le> (\<integral>\<^sup>+x. f x \<partial>M)"
e7fd64f82876 add various lemmas
hoelzl
parents: 56996
diff changeset
   161
  using nn_integral_mono_AE[of "\<lambda>x. c" f M] emeasure_space_1
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   162
  by (simp split: if_split_asm)
57025
e7fd64f82876 add various lemmas
hoelzl
parents: 56996
diff changeset
   163
43339
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   164
lemma (in prob_space) expectation_less:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56166
diff changeset
   165
  fixes X :: "_ \<Rightarrow> real"
43339
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   166
  assumes [simp]: "integrable M X"
49786
f33d5f009627 continuous version of entropy_le
hoelzl
parents: 49783
diff changeset
   167
  assumes gt: "AE x in M. X x < b"
43339
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   168
  shows "expectation X < b"
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   169
proof -
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   170
  have "expectation X < expectation (\<lambda>x. b)"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   171
    using gt emeasure_space_1
43340
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 43339
diff changeset
   172
    by (intro integral_less_AE_space) auto
43339
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   173
  then show ?thesis using prob_space by simp
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   174
qed
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   175
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   176
lemma (in prob_space) expectation_greater:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56166
diff changeset
   177
  fixes X :: "_ \<Rightarrow> real"
43339
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   178
  assumes [simp]: "integrable M X"
49786
f33d5f009627 continuous version of entropy_le
hoelzl
parents: 49783
diff changeset
   179
  assumes gt: "AE x in M. a < X x"
43339
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   180
  shows "a < expectation X"
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   181
proof -
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   182
  have "expectation (\<lambda>x. a) < expectation X"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   183
    using gt emeasure_space_1
43340
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 43339
diff changeset
   184
    by (intro integral_less_AE_space) auto
43339
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   185
  then show ?thesis using prob_space by simp
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   186
qed
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   187
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   188
lemma (in prob_space) jensens_inequality:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56166
diff changeset
   189
  fixes q :: "real \<Rightarrow> real"
49786
f33d5f009627 continuous version of entropy_le
hoelzl
parents: 49783
diff changeset
   190
  assumes X: "integrable M X" "AE x in M. X x \<in> I"
43339
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   191
  assumes I: "I = {a <..< b} \<or> I = {a <..} \<or> I = {..< b} \<or> I = UNIV"
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   192
  assumes q: "integrable M (\<lambda>x. q (X x))" "convex_on I q"
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   193
  shows "q (expectation X) \<le> expectation (\<lambda>x. q (X x))"
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   194
proof -
46731
5302e932d1e5 avoid undeclared variables in let bindings;
wenzelm
parents: 45934
diff changeset
   195
  let ?F = "\<lambda>x. Inf ((\<lambda>t. (q x - q t) / (x - t)) ` ({x<..} \<inter> I))"
49786
f33d5f009627 continuous version of entropy_le
hoelzl
parents: 49783
diff changeset
   196
  from X(2) AE_False have "I \<noteq> {}" by auto
43339
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   197
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   198
  from I have "open I" by auto
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   199
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   200
  note I
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   201
  moreover
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   202
  { assume "I \<subseteq> {a <..}"
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   203
    with X have "a < expectation X"
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   204
      by (intro expectation_greater) auto }
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   205
  moreover
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   206
  { assume "I \<subseteq> {..< b}"
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   207
    with X have "expectation X < b"
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   208
      by (intro expectation_less) auto }
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   209
  ultimately have "expectation X \<in> I"
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   210
    by (elim disjE)  (auto simp: subset_eq)
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   211
  moreover
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   212
  { fix y assume y: "y \<in> I"
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61605
diff changeset
   213
    with q(2) \<open>open I\<close> have "Sup ((\<lambda>x. q x + ?F x * (y - x)) ` I) = q y"
62343
24106dc44def prefer abbreviations for compound operators INFIMUM and SUPREMUM
haftmann
parents: 62083
diff changeset
   214
      by (auto intro!: cSup_eq_maximum convex_le_Inf_differential image_eqI [OF _ y] simp: interior_open) }
43339
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   215
  ultimately have "q (expectation X) = Sup ((\<lambda>x. q x + ?F x * (expectation X - x)) ` I)"
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   216
    by simp
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   217
  also have "\<dots> \<le> expectation (\<lambda>w. q (X w))"
51475
ebf9d4fd00ba introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
hoelzl
parents: 50419
diff changeset
   218
  proof (rule cSup_least)
43339
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   219
    show "(\<lambda>x. q x + ?F x * (expectation X - x)) ` I \<noteq> {}"
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61605
diff changeset
   220
      using \<open>I \<noteq> {}\<close> by auto
43339
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   221
  next
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   222
    fix k assume "k \<in> (\<lambda>x. q x + ?F x * (expectation X - x)) ` I"
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   223
    then guess x .. note x = this
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   224
    have "q x + ?F x * (expectation X  - x) = expectation (\<lambda>w. q x + ?F x * (X w - x))"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   225
      using prob_space by (simp add: X)
43339
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   226
    also have "\<dots> \<le> expectation (\<lambda>w. q (X w))"
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61605
diff changeset
   227
      using \<open>x \<in> I\<close> \<open>open I\<close> X(2)
63886
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63627
diff changeset
   228
      apply (intro integral_mono_AE Bochner_Integration.integrable_add Bochner_Integration.integrable_mult_right Bochner_Integration.integrable_diff
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56166
diff changeset
   229
                integrable_const X q)
61810
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61808
diff changeset
   230
      apply (elim eventually_mono)
49786
f33d5f009627 continuous version of entropy_le
hoelzl
parents: 49783
diff changeset
   231
      apply (intro convex_le_Inf_differential)
f33d5f009627 continuous version of entropy_le
hoelzl
parents: 49783
diff changeset
   232
      apply (auto simp: interior_open q)
f33d5f009627 continuous version of entropy_le
hoelzl
parents: 49783
diff changeset
   233
      done
43339
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   234
    finally show "k \<le> expectation (\<lambda>w. q (X w))" using x by auto
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   235
  qed
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   236
  finally show "q (expectation X) \<le> expectation (\<lambda>x. q (X x))" .
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   237
qed
9ba256ad6781 jensens inequality
hoelzl
parents: 42991
diff changeset
   238
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61605
diff changeset
   239
subsection  \<open>Introduce binder for probability\<close>
50001
382bd3173584 add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents: 49795
diff changeset
   240
382bd3173584 add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents: 49795
diff changeset
   241
syntax
59356
e6f5b1bbcb01 allow line breaks in probability syntax
Andreas Lochbihler
parents: 59353
diff changeset
   242
  "_prob" :: "pttrn \<Rightarrow> logic \<Rightarrow> logic \<Rightarrow> logic" ("('\<P>'((/_ in _./ _)'))")
50001
382bd3173584 add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents: 49795
diff changeset
   243
382bd3173584 add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents: 49795
diff changeset
   244
translations
382bd3173584 add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents: 49795
diff changeset
   245
  "\<P>(x in M. P)" => "CONST measure M {x \<in> CONST space M. P}"
382bd3173584 add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents: 49795
diff changeset
   246
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61605
diff changeset
   247
print_translation \<open>
58764
ca2f59aef665 add print translation for probability notation \<P>
Andreas Lochbihler
parents: 57447
diff changeset
   248
  let
ca2f59aef665 add print translation for probability notation \<P>
Andreas Lochbihler
parents: 57447
diff changeset
   249
    fun to_pattern (Const (@{const_syntax Pair}, _) $ l $ r) =
ca2f59aef665 add print translation for probability notation \<P>
Andreas Lochbihler
parents: 57447
diff changeset
   250
      Syntax.const @{const_syntax Pair} :: to_pattern l @ to_pattern r
ca2f59aef665 add print translation for probability notation \<P>
Andreas Lochbihler
parents: 57447
diff changeset
   251
    | to_pattern (t as (Const (@{syntax_const "_bound"}, _)) $ _) = [t]
ca2f59aef665 add print translation for probability notation \<P>
Andreas Lochbihler
parents: 57447
diff changeset
   252
ca2f59aef665 add print translation for probability notation \<P>
Andreas Lochbihler
parents: 57447
diff changeset
   253
    fun mk_pattern ((t, n) :: xs) = mk_patterns n xs |>> curry list_comb t
ca2f59aef665 add print translation for probability notation \<P>
Andreas Lochbihler
parents: 57447
diff changeset
   254
    and mk_patterns 0 xs = ([], xs)
ca2f59aef665 add print translation for probability notation \<P>
Andreas Lochbihler
parents: 57447
diff changeset
   255
    | mk_patterns n xs =
ca2f59aef665 add print translation for probability notation \<P>
Andreas Lochbihler
parents: 57447
diff changeset
   256
      let
ca2f59aef665 add print translation for probability notation \<P>
Andreas Lochbihler
parents: 57447
diff changeset
   257
        val (t, xs') = mk_pattern xs
ca2f59aef665 add print translation for probability notation \<P>
Andreas Lochbihler
parents: 57447
diff changeset
   258
        val (ts, xs'') = mk_patterns (n - 1) xs'
ca2f59aef665 add print translation for probability notation \<P>
Andreas Lochbihler
parents: 57447
diff changeset
   259
      in
ca2f59aef665 add print translation for probability notation \<P>
Andreas Lochbihler
parents: 57447
diff changeset
   260
        (t :: ts, xs'')
ca2f59aef665 add print translation for probability notation \<P>
Andreas Lochbihler
parents: 57447
diff changeset
   261
      end
ca2f59aef665 add print translation for probability notation \<P>
Andreas Lochbihler
parents: 57447
diff changeset
   262
ca2f59aef665 add print translation for probability notation \<P>
Andreas Lochbihler
parents: 57447
diff changeset
   263
    fun unnest_tuples
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   264
      (Const (@{syntax_const "_pattern"}, _) $
58764
ca2f59aef665 add print translation for probability notation \<P>
Andreas Lochbihler
parents: 57447
diff changeset
   265
        t1 $
ca2f59aef665 add print translation for probability notation \<P>
Andreas Lochbihler
parents: 57447
diff changeset
   266
        (t as (Const (@{syntax_const "_pattern"}, _) $ _ $ _)))
ca2f59aef665 add print translation for probability notation \<P>
Andreas Lochbihler
parents: 57447
diff changeset
   267
      = let
ca2f59aef665 add print translation for probability notation \<P>
Andreas Lochbihler
parents: 57447
diff changeset
   268
        val (_ $ t2 $ t3) = unnest_tuples t
ca2f59aef665 add print translation for probability notation \<P>
Andreas Lochbihler
parents: 57447
diff changeset
   269
      in
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   270
        Syntax.const @{syntax_const "_pattern"} $
58764
ca2f59aef665 add print translation for probability notation \<P>
Andreas Lochbihler
parents: 57447
diff changeset
   271
          unnest_tuples t1 $
ca2f59aef665 add print translation for probability notation \<P>
Andreas Lochbihler
parents: 57447
diff changeset
   272
          (Syntax.const @{syntax_const "_patterns"} $ t2 $ t3)
ca2f59aef665 add print translation for probability notation \<P>
Andreas Lochbihler
parents: 57447
diff changeset
   273
      end
ca2f59aef665 add print translation for probability notation \<P>
Andreas Lochbihler
parents: 57447
diff changeset
   274
    | unnest_tuples pat = pat
ca2f59aef665 add print translation for probability notation \<P>
Andreas Lochbihler
parents: 57447
diff changeset
   275
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   276
    fun tr' [sig_alg, Const (@{const_syntax Collect}, _) $ t] =
58764
ca2f59aef665 add print translation for probability notation \<P>
Andreas Lochbihler
parents: 57447
diff changeset
   277
      let
ca2f59aef665 add print translation for probability notation \<P>
Andreas Lochbihler
parents: 57447
diff changeset
   278
        val bound_dummyT = Const (@{syntax_const "_bound"}, dummyT)
ca2f59aef665 add print translation for probability notation \<P>
Andreas Lochbihler
parents: 57447
diff changeset
   279
ca2f59aef665 add print translation for probability notation \<P>
Andreas Lochbihler
parents: 57447
diff changeset
   280
        fun go pattern elem
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   281
          (Const (@{const_syntax "conj"}, _) $
58764
ca2f59aef665 add print translation for probability notation \<P>
Andreas Lochbihler
parents: 57447
diff changeset
   282
            (Const (@{const_syntax Set.member}, _) $ elem' $ (Const (@{const_syntax space}, _) $ sig_alg')) $
ca2f59aef665 add print translation for probability notation \<P>
Andreas Lochbihler
parents: 57447
diff changeset
   283
            u)
ca2f59aef665 add print translation for probability notation \<P>
Andreas Lochbihler
parents: 57447
diff changeset
   284
          = let
ca2f59aef665 add print translation for probability notation \<P>
Andreas Lochbihler
parents: 57447
diff changeset
   285
              val _ = if sig_alg aconv sig_alg' andalso to_pattern elem' = rev elem then () else raise Match;
ca2f59aef665 add print translation for probability notation \<P>
Andreas Lochbihler
parents: 57447
diff changeset
   286
              val (pat, rest) = mk_pattern (rev pattern);
ca2f59aef665 add print translation for probability notation \<P>
Andreas Lochbihler
parents: 57447
diff changeset
   287
              val _ = case rest of [] => () | _ => raise Match
ca2f59aef665 add print translation for probability notation \<P>
Andreas Lochbihler
parents: 57447
diff changeset
   288
            in
ca2f59aef665 add print translation for probability notation \<P>
Andreas Lochbihler
parents: 57447
diff changeset
   289
              Syntax.const @{syntax_const "_prob"} $ unnest_tuples pat $ sig_alg $ u
ca2f59aef665 add print translation for probability notation \<P>
Andreas Lochbihler
parents: 57447
diff changeset
   290
            end
ca2f59aef665 add print translation for probability notation \<P>
Andreas Lochbihler
parents: 57447
diff changeset
   291
        | go pattern elem (Abs abs) =
ca2f59aef665 add print translation for probability notation \<P>
Andreas Lochbihler
parents: 57447
diff changeset
   292
            let
ca2f59aef665 add print translation for probability notation \<P>
Andreas Lochbihler
parents: 57447
diff changeset
   293
              val (x as (_ $ tx), t) = Syntax_Trans.atomic_abs_tr' abs
ca2f59aef665 add print translation for probability notation \<P>
Andreas Lochbihler
parents: 57447
diff changeset
   294
            in
ca2f59aef665 add print translation for probability notation \<P>
Andreas Lochbihler
parents: 57447
diff changeset
   295
              go ((x, 0) :: pattern) (bound_dummyT $ tx :: elem) t
ca2f59aef665 add print translation for probability notation \<P>
Andreas Lochbihler
parents: 57447
diff changeset
   296
            end
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61359
diff changeset
   297
        | go pattern elem (Const (@{const_syntax case_prod}, _) $ t) =
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   298
            go
58764
ca2f59aef665 add print translation for probability notation \<P>
Andreas Lochbihler
parents: 57447
diff changeset
   299
              ((Syntax.const @{syntax_const "_pattern"}, 2) :: pattern)
ca2f59aef665 add print translation for probability notation \<P>
Andreas Lochbihler
parents: 57447
diff changeset
   300
              (Syntax.const @{const_syntax Pair} :: elem)
ca2f59aef665 add print translation for probability notation \<P>
Andreas Lochbihler
parents: 57447
diff changeset
   301
              t
ca2f59aef665 add print translation for probability notation \<P>
Andreas Lochbihler
parents: 57447
diff changeset
   302
      in
ca2f59aef665 add print translation for probability notation \<P>
Andreas Lochbihler
parents: 57447
diff changeset
   303
        go [] [] t
ca2f59aef665 add print translation for probability notation \<P>
Andreas Lochbihler
parents: 57447
diff changeset
   304
      end
ca2f59aef665 add print translation for probability notation \<P>
Andreas Lochbihler
parents: 57447
diff changeset
   305
  in
ca2f59aef665 add print translation for probability notation \<P>
Andreas Lochbihler
parents: 57447
diff changeset
   306
    [(@{const_syntax Sigma_Algebra.measure}, K tr')]
ca2f59aef665 add print translation for probability notation \<P>
Andreas Lochbihler
parents: 57447
diff changeset
   307
  end
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61605
diff changeset
   308
\<close>
58764
ca2f59aef665 add print translation for probability notation \<P>
Andreas Lochbihler
parents: 57447
diff changeset
   309
50001
382bd3173584 add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents: 49795
diff changeset
   310
definition
382bd3173584 add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents: 49795
diff changeset
   311
  "cond_prob M P Q = \<P>(\<omega> in M. P \<omega> \<and> Q \<omega>) / \<P>(\<omega> in M. Q \<omega>)"
382bd3173584 add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents: 49795
diff changeset
   312
382bd3173584 add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents: 49795
diff changeset
   313
syntax
382bd3173584 add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents: 49795
diff changeset
   314
  "_conditional_prob" :: "pttrn \<Rightarrow> logic \<Rightarrow> logic \<Rightarrow> logic \<Rightarrow> logic" ("('\<P>'(_ in _. _ \<bar>/ _'))")
382bd3173584 add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents: 49795
diff changeset
   315
382bd3173584 add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents: 49795
diff changeset
   316
translations
382bd3173584 add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents: 49795
diff changeset
   317
  "\<P>(x in M. P \<bar> Q)" => "CONST cond_prob M (\<lambda>x. P) (\<lambda>x. Q)"
382bd3173584 add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents: 49795
diff changeset
   318
382bd3173584 add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents: 49795
diff changeset
   319
lemma (in prob_space) AE_E_prob:
382bd3173584 add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents: 49795
diff changeset
   320
  assumes ae: "AE x in M. P x"
382bd3173584 add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents: 49795
diff changeset
   321
  obtains S where "S \<subseteq> {x \<in> space M. P x}" "S \<in> events" "prob S = 1"
382bd3173584 add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents: 49795
diff changeset
   322
proof -
382bd3173584 add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents: 49795
diff changeset
   323
  from ae[THEN AE_E] guess N .
382bd3173584 add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents: 49795
diff changeset
   324
  then show thesis
382bd3173584 add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents: 49795
diff changeset
   325
    by (intro that[of "space M - N"])
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   326
       (auto simp: prob_compl prob_space emeasure_eq_measure measure_nonneg)
50001
382bd3173584 add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents: 49795
diff changeset
   327
qed
382bd3173584 add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents: 49795
diff changeset
   328
382bd3173584 add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents: 49795
diff changeset
   329
lemma (in prob_space) prob_neg: "{x\<in>space M. P x} \<in> events \<Longrightarrow> \<P>(x in M. \<not> P x) = 1 - \<P>(x in M. P x)"
382bd3173584 add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents: 49795
diff changeset
   330
  by (auto intro!: arg_cong[where f=prob] simp add: prob_compl[symmetric])
382bd3173584 add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents: 49795
diff changeset
   331
382bd3173584 add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents: 49795
diff changeset
   332
lemma (in prob_space) prob_eq_AE:
382bd3173584 add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents: 49795
diff changeset
   333
  "(AE x in M. P x \<longleftrightarrow> Q x) \<Longrightarrow> {x\<in>space M. P x} \<in> events \<Longrightarrow> {x\<in>space M. Q x} \<in> events \<Longrightarrow> \<P>(x in M. P x) = \<P>(x in M. Q x)"
382bd3173584 add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents: 49795
diff changeset
   334
  by (rule finite_measure_eq_AE) auto
382bd3173584 add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents: 49795
diff changeset
   335
382bd3173584 add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents: 49795
diff changeset
   336
lemma (in prob_space) prob_eq_0_AE:
382bd3173584 add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents: 49795
diff changeset
   337
  assumes not: "AE x in M. \<not> P x" shows "\<P>(x in M. P x) = 0"
382bd3173584 add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents: 49795
diff changeset
   338
proof cases
382bd3173584 add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents: 49795
diff changeset
   339
  assume "{x\<in>space M. P x} \<in> events"
382bd3173584 add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents: 49795
diff changeset
   340
  with not have "\<P>(x in M. P x) = \<P>(x in M. False)"
382bd3173584 add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents: 49795
diff changeset
   341
    by (intro prob_eq_AE) auto
382bd3173584 add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents: 49795
diff changeset
   342
  then show ?thesis by simp
382bd3173584 add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents: 49795
diff changeset
   343
qed (simp add: measure_notin_sets)
382bd3173584 add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents: 49795
diff changeset
   344
50098
98abff4a775b rules for AE and prob
hoelzl
parents: 50003
diff changeset
   345
lemma (in prob_space) prob_Collect_eq_0:
98abff4a775b rules for AE and prob
hoelzl
parents: 50003
diff changeset
   346
  "{x \<in> space M. P x} \<in> sets M \<Longrightarrow> \<P>(x in M. P x) = 0 \<longleftrightarrow> (AE x in M. \<not> P x)"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   347
  using AE_iff_measurable[OF _ refl, of M "\<lambda>x. \<not> P x"] by (simp add: emeasure_eq_measure measure_nonneg)
50098
98abff4a775b rules for AE and prob
hoelzl
parents: 50003
diff changeset
   348
98abff4a775b rules for AE and prob
hoelzl
parents: 50003
diff changeset
   349
lemma (in prob_space) prob_Collect_eq_1:
98abff4a775b rules for AE and prob
hoelzl
parents: 50003
diff changeset
   350
  "{x \<in> space M. P x} \<in> sets M \<Longrightarrow> \<P>(x in M. P x) = 1 \<longleftrightarrow> (AE x in M. P x)"
98abff4a775b rules for AE and prob
hoelzl
parents: 50003
diff changeset
   351
  using AE_in_set_eq_1[of "{x\<in>space M. P x}"] by simp
98abff4a775b rules for AE and prob
hoelzl
parents: 50003
diff changeset
   352
98abff4a775b rules for AE and prob
hoelzl
parents: 50003
diff changeset
   353
lemma (in prob_space) prob_eq_0:
98abff4a775b rules for AE and prob
hoelzl
parents: 50003
diff changeset
   354
  "A \<in> sets M \<Longrightarrow> prob A = 0 \<longleftrightarrow> (AE x in M. x \<notin> A)"
98abff4a775b rules for AE and prob
hoelzl
parents: 50003
diff changeset
   355
  using AE_iff_measurable[OF _ refl, of M "\<lambda>x. x \<notin> A"]
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   356
  by (auto simp add: emeasure_eq_measure Int_def[symmetric] measure_nonneg)
50098
98abff4a775b rules for AE and prob
hoelzl
parents: 50003
diff changeset
   357
98abff4a775b rules for AE and prob
hoelzl
parents: 50003
diff changeset
   358
lemma (in prob_space) prob_eq_1:
98abff4a775b rules for AE and prob
hoelzl
parents: 50003
diff changeset
   359
  "A \<in> sets M \<Longrightarrow> prob A = 1 \<longleftrightarrow> (AE x in M. x \<in> A)"
98abff4a775b rules for AE and prob
hoelzl
parents: 50003
diff changeset
   360
  using AE_in_set_eq_1[of A] by simp
98abff4a775b rules for AE and prob
hoelzl
parents: 50003
diff changeset
   361
50001
382bd3173584 add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents: 49795
diff changeset
   362
lemma (in prob_space) prob_sums:
382bd3173584 add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents: 49795
diff changeset
   363
  assumes P: "\<And>n. {x\<in>space M. P n x} \<in> events"
382bd3173584 add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents: 49795
diff changeset
   364
  assumes Q: "{x\<in>space M. Q x} \<in> events"
382bd3173584 add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents: 49795
diff changeset
   365
  assumes ae: "AE x in M. (\<forall>n. P n x \<longrightarrow> Q x) \<and> (Q x \<longrightarrow> (\<exists>!n. P n x))"
382bd3173584 add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents: 49795
diff changeset
   366
  shows "(\<lambda>n. \<P>(x in M. P n x)) sums \<P>(x in M. Q x)"
382bd3173584 add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents: 49795
diff changeset
   367
proof -
382bd3173584 add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents: 49795
diff changeset
   368
  from ae[THEN AE_E_prob] guess S . note S = this
382bd3173584 add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents: 49795
diff changeset
   369
  then have disj: "disjoint_family (\<lambda>n. {x\<in>space M. P n x} \<inter> S)"
382bd3173584 add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents: 49795
diff changeset
   370
    by (auto simp: disjoint_family_on_def)
382bd3173584 add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents: 49795
diff changeset
   371
  from S have ae_S:
382bd3173584 add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents: 49795
diff changeset
   372
    "AE x in M. x \<in> {x\<in>space M. Q x} \<longleftrightarrow> x \<in> (\<Union>n. {x\<in>space M. P n x} \<inter> S)"
382bd3173584 add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents: 49795
diff changeset
   373
    "\<And>n. AE x in M. x \<in> {x\<in>space M. P n x} \<longleftrightarrow> x \<in> {x\<in>space M. P n x} \<inter> S"
382bd3173584 add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents: 49795
diff changeset
   374
    using ae by (auto dest!: AE_prob_1)
382bd3173584 add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents: 49795
diff changeset
   375
  from ae_S have *:
382bd3173584 add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents: 49795
diff changeset
   376
    "\<P>(x in M. Q x) = prob (\<Union>n. {x\<in>space M. P n x} \<inter> S)"
382bd3173584 add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents: 49795
diff changeset
   377
    using P Q S by (intro finite_measure_eq_AE) auto
382bd3173584 add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents: 49795
diff changeset
   378
  from ae_S have **:
382bd3173584 add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents: 49795
diff changeset
   379
    "\<And>n. \<P>(x in M. P n x) = prob ({x\<in>space M. P n x} \<inter> S)"
382bd3173584 add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents: 49795
diff changeset
   380
    using P Q S by (intro finite_measure_eq_AE) auto
382bd3173584 add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents: 49795
diff changeset
   381
  show ?thesis
382bd3173584 add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents: 49795
diff changeset
   382
    unfolding * ** using S P disj
382bd3173584 add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents: 49795
diff changeset
   383
    by (intro finite_measure_UNION) auto
382bd3173584 add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents: 49795
diff changeset
   384
qed
382bd3173584 add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents: 49795
diff changeset
   385
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
   386
lemma (in prob_space) prob_setsum:
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
   387
  assumes [simp, intro]: "finite I"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
   388
  assumes P: "\<And>n. n \<in> I \<Longrightarrow> {x\<in>space M. P n x} \<in> events"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
   389
  assumes Q: "{x\<in>space M. Q x} \<in> events"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
   390
  assumes ae: "AE x in M. (\<forall>n\<in>I. P n x \<longrightarrow> Q x) \<and> (Q x \<longrightarrow> (\<exists>!n\<in>I. P n x))"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
   391
  shows "\<P>(x in M. Q x) = (\<Sum>n\<in>I. \<P>(x in M. P n x))"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
   392
proof -
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
   393
  from ae[THEN AE_E_prob] guess S . note S = this
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
   394
  then have disj: "disjoint_family_on (\<lambda>n. {x\<in>space M. P n x} \<inter> S) I"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
   395
    by (auto simp: disjoint_family_on_def)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
   396
  from S have ae_S:
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
   397
    "AE x in M. x \<in> {x\<in>space M. Q x} \<longleftrightarrow> x \<in> (\<Union>n\<in>I. {x\<in>space M. P n x} \<inter> S)"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
   398
    "\<And>n. n \<in> I \<Longrightarrow> AE x in M. x \<in> {x\<in>space M. P n x} \<longleftrightarrow> x \<in> {x\<in>space M. P n x} \<inter> S"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
   399
    using ae by (auto dest!: AE_prob_1)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
   400
  from ae_S have *:
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
   401
    "\<P>(x in M. Q x) = prob (\<Union>n\<in>I. {x\<in>space M. P n x} \<inter> S)"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
   402
    using P Q S by (intro finite_measure_eq_AE) (auto intro!: sets.Int)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
   403
  from ae_S have **:
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
   404
    "\<And>n. n \<in> I \<Longrightarrow> \<P>(x in M. P n x) = prob ({x\<in>space M. P n x} \<inter> S)"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
   405
    using P Q S by (intro finite_measure_eq_AE) auto
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
   406
  show ?thesis
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
   407
    using S P disj
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
   408
    by (auto simp add: * ** simp del: UN_simps intro!: finite_measure_finite_Union)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
   409
qed
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
   410
54418
3b8e33d1a39a measure of a countable union
hoelzl
parents: 53015
diff changeset
   411
lemma (in prob_space) prob_EX_countable:
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   412
  assumes sets: "\<And>i. i \<in> I \<Longrightarrow> {x\<in>space M. P i x} \<in> sets M" and I: "countable I"
54418
3b8e33d1a39a measure of a countable union
hoelzl
parents: 53015
diff changeset
   413
  assumes disj: "AE x in M. \<forall>i\<in>I. \<forall>j\<in>I. P i x \<longrightarrow> P j x \<longrightarrow> i = j"
3b8e33d1a39a measure of a countable union
hoelzl
parents: 53015
diff changeset
   414
  shows "\<P>(x in M. \<exists>i\<in>I. P i x) = (\<integral>\<^sup>+i. \<P>(x in M. P i x) \<partial>count_space I)"
3b8e33d1a39a measure of a countable union
hoelzl
parents: 53015
diff changeset
   415
proof -
3b8e33d1a39a measure of a countable union
hoelzl
parents: 53015
diff changeset
   416
  let ?N= "\<lambda>x. \<exists>!i\<in>I. P i x"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   417
  have "ennreal (\<P>(x in M. \<exists>i\<in>I. P i x)) = \<P>(x in M. (\<exists>i\<in>I. P i x \<and> ?N x))"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   418
    unfolding ennreal_inj[OF measure_nonneg measure_nonneg]
54418
3b8e33d1a39a measure of a countable union
hoelzl
parents: 53015
diff changeset
   419
  proof (rule prob_eq_AE)
3b8e33d1a39a measure of a countable union
hoelzl
parents: 53015
diff changeset
   420
    show "AE x in M. (\<exists>i\<in>I. P i x) = (\<exists>i\<in>I. P i x \<and> ?N x)"
3b8e33d1a39a measure of a countable union
hoelzl
parents: 53015
diff changeset
   421
      using disj by eventually_elim blast
3b8e33d1a39a measure of a countable union
hoelzl
parents: 53015
diff changeset
   422
  qed (auto intro!: sets.sets_Collect_countable_Ex' sets.sets_Collect_conj sets.sets_Collect_countable_Ex1' I sets)+
3b8e33d1a39a measure of a countable union
hoelzl
parents: 53015
diff changeset
   423
  also have "\<P>(x in M. (\<exists>i\<in>I. P i x \<and> ?N x)) = emeasure M (\<Union>i\<in>I. {x\<in>space M. P i x \<and> ?N x})"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   424
    unfolding emeasure_eq_measure by (auto intro!: arg_cong[where f=prob] simp: measure_nonneg)
54418
3b8e33d1a39a measure of a countable union
hoelzl
parents: 53015
diff changeset
   425
  also have "\<dots> = (\<integral>\<^sup>+i. emeasure M {x\<in>space M. P i x \<and> ?N x} \<partial>count_space I)"
3b8e33d1a39a measure of a countable union
hoelzl
parents: 53015
diff changeset
   426
    by (rule emeasure_UN_countable)
3b8e33d1a39a measure of a countable union
hoelzl
parents: 53015
diff changeset
   427
       (auto intro!: sets.sets_Collect_countable_Ex' sets.sets_Collect_conj sets.sets_Collect_countable_Ex1' I sets
3b8e33d1a39a measure of a countable union
hoelzl
parents: 53015
diff changeset
   428
             simp: disjoint_family_on_def)
3b8e33d1a39a measure of a countable union
hoelzl
parents: 53015
diff changeset
   429
  also have "\<dots> = (\<integral>\<^sup>+i. \<P>(x in M. P i x) \<partial>count_space I)"
3b8e33d1a39a measure of a countable union
hoelzl
parents: 53015
diff changeset
   430
    unfolding emeasure_eq_measure using disj
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   431
    by (intro nn_integral_cong ennreal_inj[THEN iffD2] prob_eq_AE)
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   432
       (auto intro!: sets.sets_Collect_countable_Ex' sets.sets_Collect_conj sets.sets_Collect_countable_Ex1' I sets measure_nonneg)+
54418
3b8e33d1a39a measure of a countable union
hoelzl
parents: 53015
diff changeset
   433
  finally show ?thesis .
3b8e33d1a39a measure of a countable union
hoelzl
parents: 53015
diff changeset
   434
qed
3b8e33d1a39a measure of a countable union
hoelzl
parents: 53015
diff changeset
   435
50001
382bd3173584 add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents: 49795
diff changeset
   436
lemma (in prob_space) cond_prob_eq_AE:
382bd3173584 add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents: 49795
diff changeset
   437
  assumes P: "AE x in M. Q x \<longrightarrow> P x \<longleftrightarrow> P' x" "{x\<in>space M. P x} \<in> events" "{x\<in>space M. P' x} \<in> events"
382bd3173584 add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents: 49795
diff changeset
   438
  assumes Q: "AE x in M. Q x \<longleftrightarrow> Q' x" "{x\<in>space M. Q x} \<in> events" "{x\<in>space M. Q' x} \<in> events"
382bd3173584 add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents: 49795
diff changeset
   439
  shows "cond_prob M P Q = cond_prob M P' Q'"
382bd3173584 add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents: 49795
diff changeset
   440
  using P Q
50244
de72bbe42190 qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents: 50104
diff changeset
   441
  by (auto simp: cond_prob_def intro!: arg_cong2[where f="op /"] prob_eq_AE sets.sets_Collect_conj)
50001
382bd3173584 add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents: 49795
diff changeset
   442
382bd3173584 add syntax and a.e.-rules for (conditional) probability on predicates
hoelzl
parents: 49795
diff changeset
   443
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   444
lemma (in prob_space) joint_distribution_Times_le_fst:
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   445
  "random_variable MX X \<Longrightarrow> random_variable MY Y \<Longrightarrow> A \<in> sets MX \<Longrightarrow> B \<in> sets MY
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51683
diff changeset
   446
    \<Longrightarrow> emeasure (distr M (MX \<Otimes>\<^sub>M MY) (\<lambda>x. (X x, Y x))) (A \<times> B) \<le> emeasure (distr M MX X) A"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   447
  by (auto simp: emeasure_distr measurable_pair_iff comp_def intro!: emeasure_mono measurable_sets)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   448
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   449
lemma (in prob_space) joint_distribution_Times_le_snd:
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   450
  "random_variable MX X \<Longrightarrow> random_variable MY Y \<Longrightarrow> A \<in> sets MX \<Longrightarrow> B \<in> sets MY
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51683
diff changeset
   451
    \<Longrightarrow> emeasure (distr M (MX \<Otimes>\<^sub>M MY) (\<lambda>x. (X x, Y x))) (A \<times> B) \<le> emeasure (distr M MY Y) B"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   452
  by (auto simp: emeasure_distr measurable_pair_iff comp_def intro!: emeasure_mono measurable_sets)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   453
57235
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   454
lemma (in prob_space) variance_eq:
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   455
  fixes X :: "'a \<Rightarrow> real"
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   456
  assumes [simp]: "integrable M X"
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   457
  assumes [simp]: "integrable M (\<lambda>x. (X x)\<^sup>2)"
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   458
  shows "variance X = expectation (\<lambda>x. (X x)\<^sup>2) - (expectation X)\<^sup>2"
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   459
  by (simp add: field_simps prob_space power2_diff power2_eq_square[symmetric])
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   460
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   461
lemma (in prob_space) variance_positive: "0 \<le> variance (X::'a \<Rightarrow> real)"
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   462
  by (intro integral_nonneg_AE) (auto intro!: integral_nonneg_AE)
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   463
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57418
diff changeset
   464
lemma (in prob_space) variance_mean_zero:
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57418
diff changeset
   465
  "expectation X = 0 \<Longrightarrow> variance X = expectation (\<lambda>x. (X x)^2)"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57418
diff changeset
   466
  by simp
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57418
diff changeset
   467
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
   468
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
   469
61565
352c73a689da Qualifiers in locale expressions default to mandatory regardless of the command.
ballarin
parents: 61424
diff changeset
   470
sublocale pair_prob_space \<subseteq> P?: prob_space "M1 \<Otimes>\<^sub>M M2"
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
   471
proof
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51683
diff changeset
   472
  show "emeasure (M1 \<Otimes>\<^sub>M M2) (space (M1 \<Otimes>\<^sub>M M2)) = 1"
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47694
diff changeset
   473
    by (simp add: M2.emeasure_pair_measure_Times M1.emeasure_space_1 M2.emeasure_space_1 space_pair_measure)
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
   474
qed
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   475
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   476
locale product_prob_space = product_sigma_finite M for M :: "'i \<Rightarrow> 'a measure" +
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
   477
  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
   478
  assumes prob_space: "\<And>i. prob_space (M i)"
42988
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42981
diff changeset
   479
61565
352c73a689da Qualifiers in locale expressions default to mandatory regardless of the command.
ballarin
parents: 61424
diff changeset
   480
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
   481
  by (rule prob_space)
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42981
diff changeset
   482
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
   483
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
   484
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51683
diff changeset
   485
sublocale finite_product_prob_space \<subseteq> prob_space "\<Pi>\<^sub>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
   486
proof
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51683
diff changeset
   487
  show "emeasure (\<Pi>\<^sub>M i\<in>I. M i) (space (\<Pi>\<^sub>M i\<in>I. M i)) = 1"
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57275
diff changeset
   488
    by (simp add: measure_times M.emeasure_space_1 setprod.neutral_const space_PiM)
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
   489
qed
42988
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42981
diff changeset
   490
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42981
diff changeset
   491
lemma (in finite_product_prob_space) prob_times:
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42981
diff changeset
   492
  assumes X: "\<And>i. i \<in> I \<Longrightarrow> X i \<in> sets (M i)"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51683
diff changeset
   493
  shows "prob (\<Pi>\<^sub>E i\<in>I. X i) = (\<Prod>i\<in>I. M.prob i (X i))"
42988
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42981
diff changeset
   494
proof -
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   495
  have "ennreal (measure (\<Pi>\<^sub>M i\<in>I. M i) (\<Pi>\<^sub>E i\<in>I. X i)) = emeasure (\<Pi>\<^sub>M i\<in>I. M i) (\<Pi>\<^sub>E i\<in>I. X i)"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   496
    using X by (simp add: emeasure_eq_measure)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   497
  also have "\<dots> = (\<Prod>i\<in>I. emeasure (M i) (X i))"
42988
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42981
diff changeset
   498
    using measure_times X by simp
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   499
  also have "\<dots> = ennreal (\<Prod>i\<in>I. measure (M i) (X i))"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   500
    using X by (simp add: M.emeasure_eq_measure setprod_ennreal measure_nonneg)
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   501
  finally show ?thesis by (simp add: measure_nonneg setprod_nonneg)
42859
d9dfc733f25c add product of probability spaces with finite cardinality
hoelzl
parents: 42858
diff changeset
   502
qed
d9dfc733f25c add product of probability spaces with finite cardinality
hoelzl
parents: 42858
diff changeset
   503
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61605
diff changeset
   504
subsection \<open>Distributions\<close>
42892
a61e30bfd0bc add lemma prob_finite_product
hoelzl
parents: 42860
diff changeset
   505
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   506
definition distributed :: "'a measure \<Rightarrow> 'b measure \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> ('b \<Rightarrow> ennreal) \<Rightarrow> bool"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   507
where
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   508
  "distributed M N X f \<longleftrightarrow>
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   509
  distr M N X = density N f \<and> f \<in> borel_measurable N \<and> X \<in> measurable M N"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   510
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   511
lemma
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   512
  assumes "distributed M N X f"
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   513
  shows distributed_distr_eq_density: "distr M N X = density N f"
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   514
    and distributed_measurable: "X \<in> measurable M N"
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   515
    and distributed_borel_measurable: "f \<in> borel_measurable N"
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   516
  using assms by (simp_all add: distributed_def)
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   517
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   518
lemma
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   519
  assumes D: "distributed M N X f"
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   520
  shows distributed_measurable'[measurable_dest]:
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   521
      "g \<in> measurable L M \<Longrightarrow> (\<lambda>x. X (g x)) \<in> measurable L N"
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   522
    and distributed_borel_measurable'[measurable_dest]:
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   523
      "h \<in> measurable L N \<Longrightarrow> (\<lambda>x. f (h x)) \<in> borel_measurable L"
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   524
  using distributed_measurable[OF D] distributed_borel_measurable[OF D]
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   525
  by simp_all
39097
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   526
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   527
lemma distributed_real_measurable:
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   528
  "(\<And>x. x \<in> space N \<Longrightarrow> 0 \<le> f x) \<Longrightarrow> distributed M N X (\<lambda>x. ennreal (f x)) \<Longrightarrow> f \<in> borel_measurable N"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   529
  by (simp_all add: distributed_def)
35977
30d42bfd0174 Added finite measure space.
hoelzl
parents: 35929
diff changeset
   530
59353
f0707dc3d9aa measurability prover: removed app splitting, replaced by more powerful destruction rules
hoelzl
parents: 59000
diff changeset
   531
lemma distributed_real_measurable':
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   532
  "(\<And>x. x \<in> space N \<Longrightarrow> 0 \<le> f x) \<Longrightarrow> distributed M N X (\<lambda>x. ennreal (f x)) \<Longrightarrow>
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   533
    h \<in> measurable L N \<Longrightarrow> (\<lambda>x. f (h x)) \<in> borel_measurable L"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   534
  using distributed_real_measurable[measurable] by simp
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   535
59353
f0707dc3d9aa measurability prover: removed app splitting, replaced by more powerful destruction rules
hoelzl
parents: 59000
diff changeset
   536
lemma joint_distributed_measurable1:
f0707dc3d9aa measurability prover: removed app splitting, replaced by more powerful destruction rules
hoelzl
parents: 59000
diff changeset
   537
  "distributed M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x)) f \<Longrightarrow> h1 \<in> measurable N M \<Longrightarrow> (\<lambda>x. X (h1 x)) \<in> measurable N S"
f0707dc3d9aa measurability prover: removed app splitting, replaced by more powerful destruction rules
hoelzl
parents: 59000
diff changeset
   538
  by simp
f0707dc3d9aa measurability prover: removed app splitting, replaced by more powerful destruction rules
hoelzl
parents: 59000
diff changeset
   539
f0707dc3d9aa measurability prover: removed app splitting, replaced by more powerful destruction rules
hoelzl
parents: 59000
diff changeset
   540
lemma joint_distributed_measurable2:
f0707dc3d9aa measurability prover: removed app splitting, replaced by more powerful destruction rules
hoelzl
parents: 59000
diff changeset
   541
  "distributed M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x)) f \<Longrightarrow> h2 \<in> measurable N M \<Longrightarrow> (\<lambda>x. Y (h2 x)) \<in> measurable N T"
f0707dc3d9aa measurability prover: removed app splitting, replaced by more powerful destruction rules
hoelzl
parents: 59000
diff changeset
   542
  by simp
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   543
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   544
lemma distributed_count_space:
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   545
  assumes X: "distributed M (count_space A) X P" and a: "a \<in> A" and A: "finite A"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   546
  shows "P a = emeasure M (X -` {a} \<inter> space M)"
39097
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   547
proof -
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   548
  have "emeasure M (X -` {a} \<inter> space M) = emeasure (distr M (count_space A) X) {a}"
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   549
    using X a A by (simp add: emeasure_distr)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   550
  also have "\<dots> = emeasure (density (count_space A) P) {a}"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   551
    using X by (simp add: distributed_distr_eq_density)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51683
diff changeset
   552
  also have "\<dots> = (\<integral>\<^sup>+x. P a * indicator {a} x \<partial>count_space A)"
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
   553
    using X a by (auto simp add: emeasure_density distributed_def indicator_def intro!: nn_integral_cong)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   554
  also have "\<dots> = P a"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   555
    using X a by (subst nn_integral_cmult_indicator) (auto simp: distributed_def one_ennreal_def[symmetric] AE_count_space)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   556
  finally show ?thesis ..
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39091
diff changeset
   557
qed
35977
30d42bfd0174 Added finite measure space.
hoelzl
parents: 35929
diff changeset
   558
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   559
lemma distributed_cong_density:
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   560
  "(AE x in N. f x = g x) \<Longrightarrow> g \<in> borel_measurable N \<Longrightarrow> f \<in> borel_measurable N \<Longrightarrow>
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   561
    distributed M N X f \<longleftrightarrow> distributed M N X g"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   562
  by (auto simp: distributed_def intro!: density_cong)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   563
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   564
lemma (in prob_space) distributed_imp_emeasure_nonzero:
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   565
  assumes X: "distributed M MX X Px"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   566
  shows "emeasure MX {x \<in> space MX. Px x \<noteq> 0} \<noteq> 0"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   567
proof
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   568
  note Px = distributed_borel_measurable[OF X]
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   569
  interpret X: prob_space "distr M MX X"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   570
    using distributed_measurable[OF X] by (rule prob_space_distr)
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   571
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   572
  assume "emeasure MX {x \<in> space MX. Px x \<noteq> 0} = 0"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   573
  with Px have "AE x in MX. Px x = 0"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   574
    by (intro AE_I[OF subset_refl]) (auto simp: borel_measurable_ennreal_iff)
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   575
  moreover
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   576
  from X.emeasure_space_1 have "(\<integral>\<^sup>+x. Px x \<partial>MX) = 1"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   577
    unfolding distributed_distr_eq_density[OF X] using Px
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   578
    by (subst (asm) emeasure_density)
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   579
       (auto simp: borel_measurable_ennreal_iff intro!: integral_cong cong: nn_integral_cong)
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   580
  ultimately show False
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   581
    by (simp add: nn_integral_cong_AE)
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   582
qed
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   583
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   584
lemma subdensity:
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   585
  assumes T: "T \<in> measurable P Q"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   586
  assumes f: "distributed M P X f"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   587
  assumes g: "distributed M Q Y g"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   588
  assumes Y: "Y = T \<circ> X"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   589
  shows "AE x in P. g (T x) = 0 \<longrightarrow> f x = 0"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   590
proof -
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   591
  have "{x\<in>space Q. g x = 0} \<in> null_sets (distr M Q (T \<circ> X))"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   592
    using g Y by (auto simp: null_sets_density_iff distributed_def)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   593
  also have "distr M Q (T \<circ> X) = distr (distr M P X) Q T"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   594
    using T f[THEN distributed_measurable] by (rule distr_distr[symmetric])
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   595
  finally have "T -` {x\<in>space Q. g x = 0} \<inter> space P \<in> null_sets (distr M P X)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   596
    using T by (subst (asm) null_sets_distr_iff) auto
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   597
  also have "T -` {x\<in>space Q. g x = 0} \<inter> space P = {x\<in>space P. g (T x) = 0}"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   598
    using T by (auto dest: measurable_space)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   599
  finally show ?thesis
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   600
    using f g by (auto simp add: null_sets_density_iff distributed_def)
35977
30d42bfd0174 Added finite measure space.
hoelzl
parents: 35929
diff changeset
   601
qed
30d42bfd0174 Added finite measure space.
hoelzl
parents: 35929
diff changeset
   602
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   603
lemma subdensity_real:
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   604
  fixes g :: "'a \<Rightarrow> real" and f :: "'b \<Rightarrow> real"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   605
  assumes T: "T \<in> measurable P Q"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   606
  assumes f: "distributed M P X f"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   607
  assumes g: "distributed M Q Y g"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   608
  assumes Y: "Y = T \<circ> X"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   609
  shows "(AE x in P. 0 \<le> g (T x)) \<Longrightarrow> (AE x in P. 0 \<le> f x) \<Longrightarrow> AE x in P. g (T x) = 0 \<longrightarrow> f x = 0"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   610
  using subdensity[OF T, of M X "\<lambda>x. ennreal (f x)" Y "\<lambda>x. ennreal (g x)"] assms
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   611
  by auto
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   612
49788
3c10763f5cb4 show and use distributed_swap and distributed_jointI
hoelzl
parents: 49786
diff changeset
   613
lemma distributed_emeasure:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51683
diff changeset
   614
  "distributed M N X f \<Longrightarrow> A \<in> sets N \<Longrightarrow> emeasure M (X -` A \<inter> space M) = (\<integral>\<^sup>+x. f x * indicator A x \<partial>N)"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   615
  by (auto simp: distributed_distr_eq_density[symmetric] emeasure_density[symmetric] emeasure_distr)
49788
3c10763f5cb4 show and use distributed_swap and distributed_jointI
hoelzl
parents: 49786
diff changeset
   616
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
   617
lemma distributed_nn_integral:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51683
diff changeset
   618
  "distributed M N X f \<Longrightarrow> g \<in> borel_measurable N \<Longrightarrow> (\<integral>\<^sup>+x. f x * g x \<partial>N) = (\<integral>\<^sup>+x. g (X x) \<partial>M)"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   619
  by (auto simp: distributed_distr_eq_density[symmetric] nn_integral_density[symmetric] nn_integral_distr)
49788
3c10763f5cb4 show and use distributed_swap and distributed_jointI
hoelzl
parents: 49786
diff changeset
   620
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   621
lemma distributed_integral:
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   622
  "distributed M N X f \<Longrightarrow> g \<in> borel_measurable N \<Longrightarrow> (\<And>x. x \<in> space N \<Longrightarrow> 0 \<le> f x) \<Longrightarrow>
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   623
    (\<integral>x. f x * g x \<partial>N) = (\<integral>x. g (X x) \<partial>M)"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   624
  supply distributed_real_measurable[measurable]
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   625
  by (auto simp: distributed_distr_eq_density[symmetric] integral_real_density[symmetric] integral_distr)
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   626
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   627
lemma distributed_transform_integral:
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   628
  assumes Px: "distributed M N X Px" "\<And>x. x \<in> space N \<Longrightarrow> 0 \<le> Px x"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   629
  assumes "distributed M P Y Py" "\<And>x. x \<in> space P \<Longrightarrow> 0 \<le> Py x"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   630
  assumes Y: "Y = T \<circ> X" and T: "T \<in> measurable N P" and f: "f \<in> borel_measurable P"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   631
  shows "(\<integral>x. Py x * f x \<partial>P) = (\<integral>x. Px x * f (T x) \<partial>N)"
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
   632
proof -
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   633
  have "(\<integral>x. Py x * f x \<partial>P) = (\<integral>x. f (Y x) \<partial>M)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   634
    by (rule distributed_integral) fact+
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   635
  also have "\<dots> = (\<integral>x. f (T (X x)) \<partial>M)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   636
    using Y by simp
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   637
  also have "\<dots> = (\<integral>x. Px x * f (T x) \<partial>N)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   638
    using measurable_comp[OF T f] Px by (intro distributed_integral[symmetric]) (auto simp: comp_def)
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
   639
  finally show ?thesis .
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 39091
diff changeset
   640
qed
36624
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   641
49788
3c10763f5cb4 show and use distributed_swap and distributed_jointI
hoelzl
parents: 49786
diff changeset
   642
lemma (in prob_space) distributed_unique:
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   643
  assumes Px: "distributed M S X Px"
49788
3c10763f5cb4 show and use distributed_swap and distributed_jointI
hoelzl
parents: 49786
diff changeset
   644
  assumes Py: "distributed M S X Py"
3c10763f5cb4 show and use distributed_swap and distributed_jointI
hoelzl
parents: 49786
diff changeset
   645
  shows "AE x in S. Px x = Py x"
3c10763f5cb4 show and use distributed_swap and distributed_jointI
hoelzl
parents: 49786
diff changeset
   646
proof -
3c10763f5cb4 show and use distributed_swap and distributed_jointI
hoelzl
parents: 49786
diff changeset
   647
  interpret X: prob_space "distr M S X"
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   648
    using Px by (intro prob_space_distr) simp
49788
3c10763f5cb4 show and use distributed_swap and distributed_jointI
hoelzl
parents: 49786
diff changeset
   649
  have "sigma_finite_measure (distr M S X)" ..
3c10763f5cb4 show and use distributed_swap and distributed_jointI
hoelzl
parents: 49786
diff changeset
   650
  with sigma_finite_density_unique[of Px S Py ] Px Py
3c10763f5cb4 show and use distributed_swap and distributed_jointI
hoelzl
parents: 49786
diff changeset
   651
  show ?thesis
3c10763f5cb4 show and use distributed_swap and distributed_jointI
hoelzl
parents: 49786
diff changeset
   652
    by (auto simp: distributed_def)
3c10763f5cb4 show and use distributed_swap and distributed_jointI
hoelzl
parents: 49786
diff changeset
   653
qed
3c10763f5cb4 show and use distributed_swap and distributed_jointI
hoelzl
parents: 49786
diff changeset
   654
3c10763f5cb4 show and use distributed_swap and distributed_jointI
hoelzl
parents: 49786
diff changeset
   655
lemma (in prob_space) distributed_jointI:
3c10763f5cb4 show and use distributed_swap and distributed_jointI
hoelzl
parents: 49786
diff changeset
   656
  assumes "sigma_finite_measure S" "sigma_finite_measure T"
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   657
  assumes X[measurable]: "X \<in> measurable M S" and Y[measurable]: "Y \<in> measurable M T"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51683
diff changeset
   658
  assumes [measurable]: "f \<in> borel_measurable (S \<Otimes>\<^sub>M T)" and f: "AE x in S \<Otimes>\<^sub>M T. 0 \<le> f x"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   659
  assumes eq: "\<And>A B. A \<in> sets S \<Longrightarrow> B \<in> sets T \<Longrightarrow>
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51683
diff changeset
   660
    emeasure M {x \<in> space M. X x \<in> A \<and> Y x \<in> B} = (\<integral>\<^sup>+x. (\<integral>\<^sup>+y. f (x, y) * indicator B y \<partial>T) * indicator A x \<partial>S)"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51683
diff changeset
   661
  shows "distributed M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x)) f"
49788
3c10763f5cb4 show and use distributed_swap and distributed_jointI
hoelzl
parents: 49786
diff changeset
   662
  unfolding distributed_def
3c10763f5cb4 show and use distributed_swap and distributed_jointI
hoelzl
parents: 49786
diff changeset
   663
proof safe
3c10763f5cb4 show and use distributed_swap and distributed_jointI
hoelzl
parents: 49786
diff changeset
   664
  interpret S: sigma_finite_measure S by fact
3c10763f5cb4 show and use distributed_swap and distributed_jointI
hoelzl
parents: 49786
diff changeset
   665
  interpret T: sigma_finite_measure T by fact
61169
4de9ff3ea29a tuned proofs -- less legacy;
wenzelm
parents: 61125
diff changeset
   666
  interpret ST: pair_sigma_finite S T ..
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   667
49788
3c10763f5cb4 show and use distributed_swap and distributed_jointI
hoelzl
parents: 49786
diff changeset
   668
  from ST.sigma_finite_up_in_pair_measure_generator guess F :: "nat \<Rightarrow> ('b \<times> 'c) set" .. note F = this
3c10763f5cb4 show and use distributed_swap and distributed_jointI
hoelzl
parents: 49786
diff changeset
   669
  let ?E = "{a \<times> b |a b. a \<in> sets S \<and> b \<in> sets T}"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51683
diff changeset
   670
  let ?P = "S \<Otimes>\<^sub>M T"
49788
3c10763f5cb4 show and use distributed_swap and distributed_jointI
hoelzl
parents: 49786
diff changeset
   671
  show "distr M ?P (\<lambda>x. (X x, Y x)) = density ?P f" (is "?L = ?R")
3c10763f5cb4 show and use distributed_swap and distributed_jointI
hoelzl
parents: 49786
diff changeset
   672
  proof (rule measure_eqI_generator_eq[OF Int_stable_pair_measure_generator[of S T]])
3c10763f5cb4 show and use distributed_swap and distributed_jointI
hoelzl
parents: 49786
diff changeset
   673
    show "?E \<subseteq> Pow (space ?P)"
50244
de72bbe42190 qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents: 50104
diff changeset
   674
      using sets.space_closed[of S] sets.space_closed[of T] by (auto simp: space_pair_measure)
49788
3c10763f5cb4 show and use distributed_swap and distributed_jointI
hoelzl
parents: 49786
diff changeset
   675
    show "sets ?L = sigma_sets (space ?P) ?E"
3c10763f5cb4 show and use distributed_swap and distributed_jointI
hoelzl
parents: 49786
diff changeset
   676
      by (simp add: sets_pair_measure space_pair_measure)
3c10763f5cb4 show and use distributed_swap and distributed_jointI
hoelzl
parents: 49786
diff changeset
   677
    then show "sets ?R = sigma_sets (space ?P) ?E"
3c10763f5cb4 show and use distributed_swap and distributed_jointI
hoelzl
parents: 49786
diff changeset
   678
      by simp
3c10763f5cb4 show and use distributed_swap and distributed_jointI
hoelzl
parents: 49786
diff changeset
   679
  next
3c10763f5cb4 show and use distributed_swap and distributed_jointI
hoelzl
parents: 49786
diff changeset
   680
    interpret L: prob_space ?L
3c10763f5cb4 show and use distributed_swap and distributed_jointI
hoelzl
parents: 49786
diff changeset
   681
      by (rule prob_space_distr) (auto intro!: measurable_Pair)
3c10763f5cb4 show and use distributed_swap and distributed_jointI
hoelzl
parents: 49786
diff changeset
   682
    show "range F \<subseteq> ?E" "(\<Union>i. F i) = space ?P" "\<And>i. emeasure ?L (F i) \<noteq> \<infinity>"
3c10763f5cb4 show and use distributed_swap and distributed_jointI
hoelzl
parents: 49786
diff changeset
   683
      using F by (auto simp: space_pair_measure)
3c10763f5cb4 show and use distributed_swap and distributed_jointI
hoelzl
parents: 49786
diff changeset
   684
  next
3c10763f5cb4 show and use distributed_swap and distributed_jointI
hoelzl
parents: 49786
diff changeset
   685
    fix E assume "E \<in> ?E"
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   686
    then obtain A B where E[simp]: "E = A \<times> B"
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   687
      and A[measurable]: "A \<in> sets S" and B[measurable]: "B \<in> sets T" by auto
49788
3c10763f5cb4 show and use distributed_swap and distributed_jointI
hoelzl
parents: 49786
diff changeset
   688
    have "emeasure ?L E = emeasure M {x \<in> space M. X x \<in> A \<and> Y x \<in> B}"
3c10763f5cb4 show and use distributed_swap and distributed_jointI
hoelzl
parents: 49786
diff changeset
   689
      by (auto intro!: arg_cong[where f="emeasure M"] simp add: emeasure_distr measurable_Pair)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51683
diff changeset
   690
    also have "\<dots> = (\<integral>\<^sup>+x. (\<integral>\<^sup>+y. (f (x, y) * indicator B y) * indicator A x \<partial>T) \<partial>S)"
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
   691
      using f by (auto simp add: eq nn_integral_multc intro!: nn_integral_cong)
49788
3c10763f5cb4 show and use distributed_swap and distributed_jointI
hoelzl
parents: 49786
diff changeset
   692
    also have "\<dots> = emeasure ?R E"
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
   693
      by (auto simp add: emeasure_density T.nn_integral_fst[symmetric]
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
   694
               intro!: nn_integral_cong split: split_indicator)
49788
3c10763f5cb4 show and use distributed_swap and distributed_jointI
hoelzl
parents: 49786
diff changeset
   695
    finally show "emeasure ?L E = emeasure ?R E" .
3c10763f5cb4 show and use distributed_swap and distributed_jointI
hoelzl
parents: 49786
diff changeset
   696
  qed
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   697
qed (auto simp: f)
49788
3c10763f5cb4 show and use distributed_swap and distributed_jointI
hoelzl
parents: 49786
diff changeset
   698
3c10763f5cb4 show and use distributed_swap and distributed_jointI
hoelzl
parents: 49786
diff changeset
   699
lemma (in prob_space) distributed_swap:
3c10763f5cb4 show and use distributed_swap and distributed_jointI
hoelzl
parents: 49786
diff changeset
   700
  assumes "sigma_finite_measure S" "sigma_finite_measure T"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51683
diff changeset
   701
  assumes Pxy: "distributed M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x)) Pxy"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51683
diff changeset
   702
  shows "distributed M (T \<Otimes>\<^sub>M S) (\<lambda>x. (Y x, X x)) (\<lambda>(x, y). Pxy (y, x))"
49788
3c10763f5cb4 show and use distributed_swap and distributed_jointI
hoelzl
parents: 49786
diff changeset
   703
proof -
3c10763f5cb4 show and use distributed_swap and distributed_jointI
hoelzl
parents: 49786
diff changeset
   704
  interpret S: sigma_finite_measure S by fact
3c10763f5cb4 show and use distributed_swap and distributed_jointI
hoelzl
parents: 49786
diff changeset
   705
  interpret T: sigma_finite_measure T by fact
61169
4de9ff3ea29a tuned proofs -- less legacy;
wenzelm
parents: 61125
diff changeset
   706
  interpret ST: pair_sigma_finite S T ..
4de9ff3ea29a tuned proofs -- less legacy;
wenzelm
parents: 61125
diff changeset
   707
  interpret TS: pair_sigma_finite T S ..
49788
3c10763f5cb4 show and use distributed_swap and distributed_jointI
hoelzl
parents: 49786
diff changeset
   708
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   709
  note Pxy[measurable]
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   710
  show ?thesis
49788
3c10763f5cb4 show and use distributed_swap and distributed_jointI
hoelzl
parents: 49786
diff changeset
   711
    apply (subst TS.distr_pair_swap)
3c10763f5cb4 show and use distributed_swap and distributed_jointI
hoelzl
parents: 49786
diff changeset
   712
    unfolding distributed_def
3c10763f5cb4 show and use distributed_swap and distributed_jointI
hoelzl
parents: 49786
diff changeset
   713
  proof safe
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51683
diff changeset
   714
    let ?D = "distr (S \<Otimes>\<^sub>M T) (T \<Otimes>\<^sub>M S) (\<lambda>(x, y). (y, x))"
49788
3c10763f5cb4 show and use distributed_swap and distributed_jointI
hoelzl
parents: 49786
diff changeset
   715
    show 1: "(\<lambda>(x, y). Pxy (y, x)) \<in> borel_measurable ?D"
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   716
      by auto
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51683
diff changeset
   717
    show 2: "random_variable (distr (S \<Otimes>\<^sub>M T) (T \<Otimes>\<^sub>M S) (\<lambda>(x, y). (y, x))) (\<lambda>x. (Y x, X x))"
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   718
      using Pxy by auto
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51683
diff changeset
   719
    { fix A assume A: "A \<in> sets (T \<Otimes>\<^sub>M S)"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51683
diff changeset
   720
      let ?B = "(\<lambda>(x, y). (y, x)) -` A \<inter> space (S \<Otimes>\<^sub>M T)"
50244
de72bbe42190 qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents: 50104
diff changeset
   721
      from sets.sets_into_space[OF A]
49788
3c10763f5cb4 show and use distributed_swap and distributed_jointI
hoelzl
parents: 49786
diff changeset
   722
      have "emeasure M ((\<lambda>x. (Y x, X x)) -` A \<inter> space M) =
3c10763f5cb4 show and use distributed_swap and distributed_jointI
hoelzl
parents: 49786
diff changeset
   723
        emeasure M ((\<lambda>x. (X x, Y x)) -` ?B \<inter> space M)"
3c10763f5cb4 show and use distributed_swap and distributed_jointI
hoelzl
parents: 49786
diff changeset
   724
        by (auto intro!: arg_cong2[where f=emeasure] simp: space_pair_measure)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51683
diff changeset
   725
      also have "\<dots> = (\<integral>\<^sup>+ x. Pxy x * indicator ?B x \<partial>(S \<Otimes>\<^sub>M T))"
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   726
        using Pxy A by (intro distributed_emeasure) auto
49788
3c10763f5cb4 show and use distributed_swap and distributed_jointI
hoelzl
parents: 49786
diff changeset
   727
      finally have "emeasure M ((\<lambda>x. (Y x, X x)) -` A \<inter> space M) =
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51683
diff changeset
   728
        (\<integral>\<^sup>+ x. Pxy x * indicator A (snd x, fst x) \<partial>(S \<Otimes>\<^sub>M T))"
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
   729
        by (auto intro!: nn_integral_cong split: split_indicator) }
49788
3c10763f5cb4 show and use distributed_swap and distributed_jointI
hoelzl
parents: 49786
diff changeset
   730
    note * = this
3c10763f5cb4 show and use distributed_swap and distributed_jointI
hoelzl
parents: 49786
diff changeset
   731
    show "distr M ?D (\<lambda>x. (Y x, X x)) = density ?D (\<lambda>(x, y). Pxy (y, x))"
3c10763f5cb4 show and use distributed_swap and distributed_jointI
hoelzl
parents: 49786
diff changeset
   732
      apply (intro measure_eqI)
3c10763f5cb4 show and use distributed_swap and distributed_jointI
hoelzl
parents: 49786
diff changeset
   733
      apply (simp_all add: emeasure_distr[OF 2] emeasure_density[OF 1])
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
   734
      apply (subst nn_integral_distr)
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   735
      apply (auto intro!: * simp: comp_def split_beta)
49788
3c10763f5cb4 show and use distributed_swap and distributed_jointI
hoelzl
parents: 49786
diff changeset
   736
      done
3c10763f5cb4 show and use distributed_swap and distributed_jointI
hoelzl
parents: 49786
diff changeset
   737
  qed
36624
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   738
qed
25153c08655e Cleanup information theory
hoelzl
parents: 35977
diff changeset
   739
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   740
lemma (in prob_space) distr_marginal1:
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   741
  assumes "sigma_finite_measure S" "sigma_finite_measure T"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51683
diff changeset
   742
  assumes Pxy: "distributed M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x)) Pxy"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51683
diff changeset
   743
  defines "Px \<equiv> \<lambda>x. (\<integral>\<^sup>+z. Pxy (x, z) \<partial>T)"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   744
  shows "distributed M S X Px"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   745
  unfolding distributed_def
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   746
proof safe
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   747
  interpret S: sigma_finite_measure S by fact
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   748
  interpret T: sigma_finite_measure T by fact
61169
4de9ff3ea29a tuned proofs -- less legacy;
wenzelm
parents: 61125
diff changeset
   749
  interpret ST: pair_sigma_finite S T ..
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   750
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   751
  note Pxy[measurable]
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   752
  show X: "X \<in> measurable M S" by simp
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   753
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   754
  show borel: "Px \<in> borel_measurable S"
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
   755
    by (auto intro!: T.nn_integral_fst simp: Px_def)
39097
943c7b348524 Moved lemmas to appropriate locations
hoelzl
parents: 39096
diff changeset
   756
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51683
diff changeset
   757
  interpret Pxy: prob_space "distr M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x))"
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   758
    by (intro prob_space_distr) simp
49788
3c10763f5cb4 show and use distributed_swap and distributed_jointI
hoelzl
parents: 49786
diff changeset
   759
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   760
  show "distr M S X = density S Px"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   761
  proof (rule measure_eqI)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   762
    fix A assume A: "A \<in> sets (distr M S X)"
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   763
    with X measurable_space[of Y M T]
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51683
diff changeset
   764
    have "emeasure (distr M S X) A = emeasure (distr M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x))) (A \<times> space T)"
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   765
      by (auto simp add: emeasure_distr intro!: arg_cong[where f="emeasure M"])
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51683
diff changeset
   766
    also have "\<dots> = emeasure (density (S \<Otimes>\<^sub>M T) Pxy) (A \<times> space T)"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   767
      using Pxy by (simp add: distributed_def)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51683
diff changeset
   768
    also have "\<dots> = \<integral>\<^sup>+ x. \<integral>\<^sup>+ y. Pxy (x, y) * indicator (A \<times> space T) (x, y) \<partial>T \<partial>S"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   769
      using A borel Pxy
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
   770
      by (simp add: emeasure_density T.nn_integral_fst[symmetric])
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51683
diff changeset
   771
    also have "\<dots> = \<integral>\<^sup>+ x. Px x * indicator A x \<partial>S"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   772
    proof (rule nn_integral_cong)
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   773
      fix x assume "x \<in> space S"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   774
      moreover have eq: "\<And>y. y \<in> space T \<Longrightarrow> indicator (A \<times> space T) (x, y) = indicator A x"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   775
        by (auto simp: indicator_def)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51683
diff changeset
   776
      ultimately have "(\<integral>\<^sup>+ y. Pxy (x, y) * indicator (A \<times> space T) (x, y) \<partial>T) = (\<integral>\<^sup>+ y. Pxy (x, y) \<partial>T) * indicator A x"
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
   777
        by (simp add: eq nn_integral_multc cong: nn_integral_cong)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51683
diff changeset
   778
      also have "(\<integral>\<^sup>+ y. Pxy (x, y) \<partial>T) = Px x"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   779
        by (simp add: Px_def)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51683
diff changeset
   780
      finally show "(\<integral>\<^sup>+ y. Pxy (x, y) * indicator (A \<times> space T) (x, y) \<partial>T) = Px x * indicator A x" .
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   781
    qed
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   782
    finally show "emeasure (distr M S X) A = emeasure (density S Px) A"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   783
      using A borel Pxy by (simp add: emeasure_density)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   784
  qed simp
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   785
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   786
49788
3c10763f5cb4 show and use distributed_swap and distributed_jointI
hoelzl
parents: 49786
diff changeset
   787
lemma (in prob_space) distr_marginal2:
3c10763f5cb4 show and use distributed_swap and distributed_jointI
hoelzl
parents: 49786
diff changeset
   788
  assumes S: "sigma_finite_measure S" and T: "sigma_finite_measure T"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51683
diff changeset
   789
  assumes Pxy: "distributed M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x)) Pxy"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51683
diff changeset
   790
  shows "distributed M T Y (\<lambda>y. (\<integral>\<^sup>+x. Pxy (x, y) \<partial>S))"
49788
3c10763f5cb4 show and use distributed_swap and distributed_jointI
hoelzl
parents: 49786
diff changeset
   791
  using distr_marginal1[OF T S distributed_swap[OF S T]] Pxy by simp
3c10763f5cb4 show and use distributed_swap and distributed_jointI
hoelzl
parents: 49786
diff changeset
   792
3c10763f5cb4 show and use distributed_swap and distributed_jointI
hoelzl
parents: 49786
diff changeset
   793
lemma (in prob_space) distributed_marginal_eq_joint1:
3c10763f5cb4 show and use distributed_swap and distributed_jointI
hoelzl
parents: 49786
diff changeset
   794
  assumes T: "sigma_finite_measure T"
3c10763f5cb4 show and use distributed_swap and distributed_jointI
hoelzl
parents: 49786
diff changeset
   795
  assumes S: "sigma_finite_measure S"
3c10763f5cb4 show and use distributed_swap and distributed_jointI
hoelzl
parents: 49786
diff changeset
   796
  assumes Px: "distributed M S X Px"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51683
diff changeset
   797
  assumes Pxy: "distributed M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x)) Pxy"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51683
diff changeset
   798
  shows "AE x in S. Px x = (\<integral>\<^sup>+y. Pxy (x, y) \<partial>T)"
49788
3c10763f5cb4 show and use distributed_swap and distributed_jointI
hoelzl
parents: 49786
diff changeset
   799
  using Px distr_marginal1[OF S T Pxy] by (rule distributed_unique)
3c10763f5cb4 show and use distributed_swap and distributed_jointI
hoelzl
parents: 49786
diff changeset
   800
3c10763f5cb4 show and use distributed_swap and distributed_jointI
hoelzl
parents: 49786
diff changeset
   801
lemma (in prob_space) distributed_marginal_eq_joint2:
3c10763f5cb4 show and use distributed_swap and distributed_jointI
hoelzl
parents: 49786
diff changeset
   802
  assumes T: "sigma_finite_measure T"
3c10763f5cb4 show and use distributed_swap and distributed_jointI
hoelzl
parents: 49786
diff changeset
   803
  assumes S: "sigma_finite_measure S"
3c10763f5cb4 show and use distributed_swap and distributed_jointI
hoelzl
parents: 49786
diff changeset
   804
  assumes Py: "distributed M T Y Py"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51683
diff changeset
   805
  assumes Pxy: "distributed M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x)) Pxy"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51683
diff changeset
   806
  shows "AE y in T. Py y = (\<integral>\<^sup>+x. Pxy (x, y) \<partial>S)"
49788
3c10763f5cb4 show and use distributed_swap and distributed_jointI
hoelzl
parents: 49786
diff changeset
   807
  using Py distr_marginal2[OF S T Pxy] by (rule distributed_unique)
3c10763f5cb4 show and use distributed_swap and distributed_jointI
hoelzl
parents: 49786
diff changeset
   808
49795
9f2fb9b25a77 joint distribution of independent variables
hoelzl
parents: 49788
diff changeset
   809
lemma (in prob_space) distributed_joint_indep':
9f2fb9b25a77 joint distribution of independent variables
hoelzl
parents: 49788
diff changeset
   810
  assumes S: "sigma_finite_measure S" and T: "sigma_finite_measure T"
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   811
  assumes X[measurable]: "distributed M S X Px" and Y[measurable]: "distributed M T Y Py"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51683
diff changeset
   812
  assumes indep: "distr M S X \<Otimes>\<^sub>M distr M T Y = distr M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x))"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51683
diff changeset
   813
  shows "distributed M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x)) (\<lambda>(x, y). Px x * Py y)"
49795
9f2fb9b25a77 joint distribution of independent variables
hoelzl
parents: 49788
diff changeset
   814
  unfolding distributed_def
9f2fb9b25a77 joint distribution of independent variables
hoelzl
parents: 49788
diff changeset
   815
proof safe
9f2fb9b25a77 joint distribution of independent variables
hoelzl
parents: 49788
diff changeset
   816
  interpret S: sigma_finite_measure S by fact
9f2fb9b25a77 joint distribution of independent variables
hoelzl
parents: 49788
diff changeset
   817
  interpret T: sigma_finite_measure T by fact
61169
4de9ff3ea29a tuned proofs -- less legacy;
wenzelm
parents: 61125
diff changeset
   818
  interpret ST: pair_sigma_finite S T ..
49795
9f2fb9b25a77 joint distribution of independent variables
hoelzl
parents: 49788
diff changeset
   819
9f2fb9b25a77 joint distribution of independent variables
hoelzl
parents: 49788
diff changeset
   820
  interpret X: prob_space "density S Px"
9f2fb9b25a77 joint distribution of independent variables
hoelzl
parents: 49788
diff changeset
   821
    unfolding distributed_distr_eq_density[OF X, symmetric]
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   822
    by (rule prob_space_distr) simp
49795
9f2fb9b25a77 joint distribution of independent variables
hoelzl
parents: 49788
diff changeset
   823
  have sf_X: "sigma_finite_measure (density S Px)" ..
9f2fb9b25a77 joint distribution of independent variables
hoelzl
parents: 49788
diff changeset
   824
9f2fb9b25a77 joint distribution of independent variables
hoelzl
parents: 49788
diff changeset
   825
  interpret Y: prob_space "density T Py"
9f2fb9b25a77 joint distribution of independent variables
hoelzl
parents: 49788
diff changeset
   826
    unfolding distributed_distr_eq_density[OF Y, symmetric]
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50002
diff changeset
   827
    by (rule prob_space_distr) simp
49795
9f2fb9b25a77 joint distribution of independent variables
hoelzl
parents: 49788
diff changeset
   828
  have sf_Y: "sigma_finite_measure (density T Py)" ..
9f2fb9b25a77 joint distribution of independent variables
hoelzl
parents: 49788
diff changeset
   829
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51683
diff changeset
   830
  show "distr M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x)) = density (S \<Otimes>\<^sub>M T) (\<lambda>(x, y). Px x * Py y)"
49795
9f2fb9b25a77 joint distribution of independent variables
hoelzl
parents: 49788
diff changeset
   831
    unfolding indep[symmetric] distributed_distr_eq_density[OF X] distributed_distr_eq_density[OF Y]
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   832
    using distributed_borel_measurable[OF X]
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   833
    using distributed_borel_measurable[OF Y]
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   834
    by (rule pair_measure_density[OF _ _ T sf_Y])
49795
9f2fb9b25a77 joint distribution of independent variables
hoelzl
parents: 49788
diff changeset
   835
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51683
diff changeset
   836
  show "random_variable (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x))" by auto
49795
9f2fb9b25a77 joint distribution of independent variables
hoelzl
parents: 49788
diff changeset
   837
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51683
diff changeset
   838
  show Pxy: "(\<lambda>(x, y). Px x * Py y) \<in> borel_measurable (S \<Otimes>\<^sub>M T)" by auto
49795
9f2fb9b25a77 joint distribution of independent variables
hoelzl
parents: 49788
diff changeset
   839
qed
9f2fb9b25a77 joint distribution of independent variables
hoelzl
parents: 49788
diff changeset
   840
57235
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   841
lemma distributed_integrable:
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   842
  "distributed M N X f \<Longrightarrow> g \<in> borel_measurable N \<Longrightarrow> (\<And>x. x \<in> space N \<Longrightarrow> 0 \<le> f x) \<Longrightarrow>
57235
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   843
    integrable N (\<lambda>x. f x * g x) \<longleftrightarrow> integrable M (\<lambda>x. g (X x))"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   844
  supply distributed_real_measurable[measurable]
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   845
  by (auto simp: distributed_distr_eq_density[symmetric] integrable_real_density[symmetric] integrable_distr_eq)
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   846
57235
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   847
lemma distributed_transform_integrable:
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   848
  assumes Px: "distributed M N X Px" "\<And>x. x \<in> space N \<Longrightarrow> 0 \<le> Px x"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   849
  assumes "distributed M P Y Py" "\<And>x. x \<in> space P \<Longrightarrow> 0 \<le> Py x"
57235
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   850
  assumes Y: "Y = (\<lambda>x. T (X x))" and T: "T \<in> measurable N P" and f: "f \<in> borel_measurable P"
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   851
  shows "integrable P (\<lambda>x. Py x * f x) \<longleftrightarrow> integrable N (\<lambda>x. Px x * f (T x))"
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   852
proof -
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   853
  have "integrable P (\<lambda>x. Py x * f x) \<longleftrightarrow> integrable M (\<lambda>x. f (Y x))"
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   854
    by (rule distributed_integrable) fact+
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   855
  also have "\<dots> \<longleftrightarrow> integrable M (\<lambda>x. f (T (X x)))"
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   856
    using Y by simp
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   857
  also have "\<dots> \<longleftrightarrow> integrable N (\<lambda>x. Px x * f (T x))"
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   858
    using measurable_comp[OF T f] Px by (intro distributed_integrable[symmetric]) (auto simp: comp_def)
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   859
  finally show ?thesis .
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   860
qed
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   861
57275
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
   862
lemma distributed_integrable_var:
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
   863
  fixes X :: "'a \<Rightarrow> real"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   864
  shows "distributed M lborel X (\<lambda>x. ennreal (f x)) \<Longrightarrow> (\<And>x. 0 \<le> f x) \<Longrightarrow>
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   865
    integrable lborel (\<lambda>x. f x * x) \<Longrightarrow> integrable M X"
57275
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
   866
  using distributed_integrable[of M lborel X f "\<lambda>x. x"] by simp
0ddb5b755cdc moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57235
diff changeset
   867
57235
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   868
lemma (in prob_space) distributed_variance:
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   869
  fixes f::"real \<Rightarrow> real"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   870
  assumes D: "distributed M lborel X f" and [simp]: "\<And>x. 0 \<le> f x"
57235
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   871
  shows "variance X = (\<integral>x. x\<^sup>2 * f (x + expectation X) \<partial>lborel)"
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   872
proof (subst distributed_integral[OF D, symmetric])
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   873
  show "(\<integral> x. f x * (x - expectation X)\<^sup>2 \<partial>lborel) = (\<integral> x. x\<^sup>2 * f (x + expectation X) \<partial>lborel)"
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   874
    by (subst lborel_integral_real_affine[where c=1 and t="expectation X"])  (auto simp: ac_simps)
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   875
qed simp_all
57235
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   876
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   877
lemma (in prob_space) variance_affine:
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   878
  fixes f::"real \<Rightarrow> real"
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   879
  assumes [arith]: "b \<noteq> 0"
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   880
  assumes D[intro]: "distributed M lborel X f"
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   881
  assumes [simp]: "prob_space (density lborel f)"
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   882
  assumes I[simp]: "integrable M X"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   883
  assumes I2[simp]: "integrable M (\<lambda>x. (X x)\<^sup>2)"
57235
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   884
  shows "variance (\<lambda>x. a + b * X x) = b\<^sup>2 * variance X"
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   885
  by (subst variance_eq)
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   886
     (auto simp: power2_sum power_mult_distrib prob_space variance_eq right_diff_distrib)
b0b9a10e4bf4 properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents: 57025
diff changeset
   887
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   888
definition
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   889
  "simple_distributed M X f \<longleftrightarrow>
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   890
    (\<forall>x. 0 \<le> f x) \<and>
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   891
    distributed M (count_space (X`space M)) X (\<lambda>x. ennreal (f x)) \<and>
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   892
    finite (X`space M)"
42902
e8dbf90a2f3b Add restricted borel measure to {0 .. 1}
hoelzl
parents: 42892
diff changeset
   893
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   894
lemma simple_distributed_nonneg[dest]: "simple_distributed M X f \<Longrightarrow> 0 \<le> f x"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   895
  by (auto simp: simple_distributed_def)
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   896
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   897
lemma simple_distributed:
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   898
  "simple_distributed M X Px \<Longrightarrow> distributed M (count_space (X`space M)) X Px"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   899
  unfolding simple_distributed_def by auto
42902
e8dbf90a2f3b Add restricted borel measure to {0 .. 1}
hoelzl
parents: 42892
diff changeset
   900
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   901
lemma simple_distributed_finite[dest]: "simple_distributed M X P \<Longrightarrow> finite (X`space M)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   902
  by (simp add: simple_distributed_def)
42902
e8dbf90a2f3b Add restricted borel measure to {0 .. 1}
hoelzl
parents: 42892
diff changeset
   903
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   904
lemma (in prob_space) distributed_simple_function_superset:
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   905
  assumes X: "simple_function M X" "\<And>x. x \<in> X ` space M \<Longrightarrow> P x = measure M (X -` {x} \<inter> space M)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   906
  assumes A: "X`space M \<subseteq> A" "finite A"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   907
  defines "S \<equiv> count_space A" and "P' \<equiv> (\<lambda>x. if x \<in> X`space M then P x else 0)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   908
  shows "distributed M S X P'"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   909
  unfolding distributed_def
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   910
proof safe
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   911
  show "(\<lambda>x. ennreal (P' x)) \<in> borel_measurable S" unfolding S_def by simp
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   912
  show "distr M S X = density S P'"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   913
  proof (rule measure_eqI_finite)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   914
    show "sets (distr M S X) = Pow A" "sets (density S P') = Pow A"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   915
      using A unfolding S_def by auto
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   916
    show "finite A" by fact
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   917
    fix a assume a: "a \<in> A"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   918
    then have "a \<notin> X`space M \<Longrightarrow> X -` {a} \<inter> space M = {}" by auto
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   919
    with A a X have "emeasure (distr M S X) {a} = P' a"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   920
      by (subst emeasure_distr)
50002
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
   921
         (auto simp add: S_def P'_def simple_functionD emeasure_eq_measure measurable_count_space_eq2
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   922
               intro!: arg_cong[where f=prob])
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   923
    also have "\<dots> = (\<integral>\<^sup>+x. ennreal (P' a) * indicator {a} x \<partial>S)"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   924
      using A X a
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
   925
      by (subst nn_integral_cmult_indicator)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   926
         (auto simp: S_def P'_def simple_distributed_def simple_functionD measure_nonneg)
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   927
    also have "\<dots> = (\<integral>\<^sup>+x. ennreal (P' x) * indicator {a} x \<partial>S)"
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
   928
      by (auto simp: indicator_def intro!: nn_integral_cong)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   929
    also have "\<dots> = emeasure (density S P') {a}"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   930
      using a A by (intro emeasure_density[symmetric]) (auto simp: S_def)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   931
    finally show "emeasure (distr M S X) {a} = emeasure (density S P') {a}" .
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   932
  qed
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   933
  show "random_variable S X"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   934
    using X(1) A by (auto simp: measurable_def simple_functionD S_def)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   935
qed
42902
e8dbf90a2f3b Add restricted borel measure to {0 .. 1}
hoelzl
parents: 42892
diff changeset
   936
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   937
lemma (in prob_space) simple_distributedI:
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   938
  assumes X: "simple_function M X"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   939
    "\<And>x. 0 \<le> P x"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   940
    "\<And>x. x \<in> X ` space M \<Longrightarrow> P x = measure M (X -` {x} \<inter> space M)"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   941
  shows "simple_distributed M X P"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   942
  unfolding simple_distributed_def
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   943
proof (safe intro!: X)
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   944
  have "distributed M (count_space (X ` space M)) X (\<lambda>x. ennreal (if x \<in> X`space M then P x else 0))"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   945
    (is "?A")
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   946
    using simple_functionD[OF X(1)] by (intro distributed_simple_function_superset[OF X(1,3)]) auto
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   947
  also have "?A \<longleftrightarrow> distributed M (count_space (X ` space M)) X (\<lambda>x. ennreal (P x))"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   948
    by (rule distributed_cong_density) auto
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   949
  finally show "\<dots>" .
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   950
qed (rule simple_functionD[OF X(1)])
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   951
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   952
lemma simple_distributed_joint_finite:
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   953
  assumes X: "simple_distributed M (\<lambda>x. (X x, Y x)) Px"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   954
  shows "finite (X ` space M)" "finite (Y ` space M)"
42902
e8dbf90a2f3b Add restricted borel measure to {0 .. 1}
hoelzl
parents: 42892
diff changeset
   955
proof -
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   956
  have "finite ((\<lambda>x. (X x, Y x)) ` space M)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   957
    using X by (auto simp: simple_distributed_def simple_functionD)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   958
  then have "finite (fst ` (\<lambda>x. (X x, Y x)) ` space M)" "finite (snd ` (\<lambda>x. (X x, Y x)) ` space M)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   959
    by auto
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   960
  then show fin: "finite (X ` space M)" "finite (Y ` space M)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   961
    by (auto simp: image_image)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   962
qed
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   963
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   964
lemma simple_distributed_joint2_finite:
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   965
  assumes X: "simple_distributed M (\<lambda>x. (X x, Y x, Z x)) Px"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   966
  shows "finite (X ` space M)" "finite (Y ` space M)" "finite (Z ` space M)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   967
proof -
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   968
  have "finite ((\<lambda>x. (X x, Y x, Z x)) ` space M)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   969
    using X by (auto simp: simple_distributed_def simple_functionD)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   970
  then have "finite (fst ` (\<lambda>x. (X x, Y x, Z x)) ` space M)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   971
    "finite ((fst \<circ> snd) ` (\<lambda>x. (X x, Y x, Z x)) ` space M)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   972
    "finite ((snd \<circ> snd) ` (\<lambda>x. (X x, Y x, Z x)) ` space M)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   973
    by auto
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   974
  then show fin: "finite (X ` space M)" "finite (Y ` space M)" "finite (Z ` space M)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   975
    by (auto simp: image_image)
42902
e8dbf90a2f3b Add restricted borel measure to {0 .. 1}
hoelzl
parents: 42892
diff changeset
   976
qed
e8dbf90a2f3b Add restricted borel measure to {0 .. 1}
hoelzl
parents: 42892
diff changeset
   977
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   978
lemma simple_distributed_simple_function:
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   979
  "simple_distributed M X Px \<Longrightarrow> simple_function M X"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   980
  unfolding simple_distributed_def distributed_def
50002
ce0d316b5b44 add measurability prover; add support for Borel sets
hoelzl
parents: 50001
diff changeset
   981
  by (auto simp: simple_function_def measurable_count_space_eq2)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   982
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   983
lemma simple_distributed_measure:
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   984
  "simple_distributed M X P \<Longrightarrow> a \<in> X`space M \<Longrightarrow> P a = measure M (X -` {a} \<inter> space M)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   985
  using distributed_count_space[of M "X`space M" X P a, symmetric]
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   986
  by (auto simp: simple_distributed_def measure_def)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   987
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   988
lemma (in prob_space) simple_distributed_joint:
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   989
  assumes X: "simple_distributed M (\<lambda>x. (X x, Y x)) Px"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51683
diff changeset
   990
  defines "S \<equiv> count_space (X`space M) \<Otimes>\<^sub>M count_space (Y`space M)"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   991
  defines "P \<equiv> (\<lambda>x. if x \<in> (\<lambda>x. (X x, Y x))`space M then Px x else 0)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   992
  shows "distributed M S (\<lambda>x. (X x, Y x)) P"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   993
proof -
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   994
  from simple_distributed_joint_finite[OF X, simp]
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   995
  have S_eq: "S = count_space (X`space M \<times> Y`space M)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   996
    by (simp add: S_def pair_measure_count_space)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   997
  show ?thesis
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   998
    unfolding S_eq P_def
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   999
  proof (rule distributed_simple_function_superset)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1000
    show "simple_function M (\<lambda>x. (X x, Y x))"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1001
      using X by (rule simple_distributed_simple_function)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1002
    fix x assume "x \<in> (\<lambda>x. (X x, Y x)) ` space M"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1003
    from simple_distributed_measure[OF X this]
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1004
    show "Px x = prob ((\<lambda>x. (X x, Y x)) -` {x} \<inter> space M)" .
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1005
  qed auto
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1006
qed
42860
b02349e70d5a add Bernoulli space
hoelzl
parents: 42859
diff changeset
  1007
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1008
lemma (in prob_space) simple_distributed_joint2:
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1009
  assumes X: "simple_distributed M (\<lambda>x. (X x, Y x, Z x)) Px"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51683
diff changeset
  1010
  defines "S \<equiv> count_space (X`space M) \<Otimes>\<^sub>M count_space (Y`space M) \<Otimes>\<^sub>M count_space (Z`space M)"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1011
  defines "P \<equiv> (\<lambda>x. if x \<in> (\<lambda>x. (X x, Y x, Z x))`space M then Px x else 0)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1012
  shows "distributed M S (\<lambda>x. (X x, Y x, Z x)) P"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1013
proof -
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1014
  from simple_distributed_joint2_finite[OF X, simp]
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1015
  have S_eq: "S = count_space (X`space M \<times> Y`space M \<times> Z`space M)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1016
    by (simp add: S_def pair_measure_count_space)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1017
  show ?thesis
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1018
    unfolding S_eq P_def
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1019
  proof (rule distributed_simple_function_superset)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1020
    show "simple_function M (\<lambda>x. (X x, Y x, Z x))"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1021
      using X by (rule simple_distributed_simple_function)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1022
    fix x assume "x \<in> (\<lambda>x. (X x, Y x, Z x)) ` space M"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1023
    from simple_distributed_measure[OF X this]
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1024
    show "Px x = prob ((\<lambda>x. (X x, Y x, Z x)) -` {x} \<inter> space M)" .
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1025
  qed auto
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1026
qed
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1027
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1028
lemma (in prob_space) simple_distributed_setsum_space:
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1029
  assumes X: "simple_distributed M X f"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1030
  shows "setsum f (X`space M) = 1"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1031
proof -
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1032
  from X have "setsum f (X`space M) = prob (\<Union>i\<in>X`space M. X -` {i} \<inter> space M)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1033
    by (subst finite_measure_finite_Union)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1034
       (auto simp add: disjoint_family_on_def simple_distributed_measure simple_distributed_simple_function simple_functionD
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57275
diff changeset
  1035
             intro!: setsum.cong arg_cong[where f="prob"])
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1036
  also have "\<dots> = prob (space M)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1037
    by (auto intro!: arg_cong[where f=prob])
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1038
  finally show ?thesis
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1039
    using emeasure_space_1 by (simp add: emeasure_eq_measure)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1040
qed
42860
b02349e70d5a add Bernoulli space
hoelzl
parents: 42859
diff changeset
  1041
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1042
lemma (in prob_space) distributed_marginal_eq_joint_simple:
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1043
  assumes Px: "simple_function M X"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1044
  assumes Py: "simple_distributed M Y Py"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1045
  assumes Pxy: "simple_distributed M (\<lambda>x. (X x, Y x)) Pxy"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1046
  assumes y: "y \<in> Y`space M"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1047
  shows "Py y = (\<Sum>x\<in>X`space M. if (x, y) \<in> (\<lambda>x. (X x, Y x)) ` space M then Pxy (x, y) else 0)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1048
proof -
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1049
  note Px = simple_distributedI[OF Px measure_nonneg refl]
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1050
  have "AE y in count_space (Y ` space M). ennreal (Py y) =
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1051
       \<integral>\<^sup>+ x. ennreal (if (x, y) \<in> (\<lambda>x. (X x, Y x)) ` space M then Pxy (x, y) else 0) \<partial>count_space (X ` space M)"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1052
    using sigma_finite_measure_count_space_finite sigma_finite_measure_count_space_finite
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1053
      simple_distributed[OF Py] simple_distributed_joint[OF Pxy]
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1054
    by (rule distributed_marginal_eq_joint2)
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1055
       (auto intro: Py Px simple_distributed_finite)
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1056
  then have "ennreal (Py y) =
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1057
    (\<Sum>x\<in>X`space M. ennreal (if (x, y) \<in> (\<lambda>x. (X x, Y x)) ` space M then Pxy (x, y) else 0))"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1058
    using y Px[THEN simple_distributed_finite]
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1059
    by (auto simp: AE_count_space nn_integral_count_space_finite)
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1060
  also have "\<dots> = (\<Sum>x\<in>X`space M. if (x, y) \<in> (\<lambda>x. (X x, Y x)) ` space M then Pxy (x, y) else 0)"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1061
    using Pxy by (intro setsum_ennreal) auto
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1062
  finally show ?thesis
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1063
    using simple_distributed_nonneg[OF Py] simple_distributed_nonneg[OF Pxy]
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1064
    by (subst (asm) ennreal_inj) (auto intro!: setsum_nonneg)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1065
qed
42860
b02349e70d5a add Bernoulli space
hoelzl
parents: 42859
diff changeset
  1066
50419
3177d0374701 add exponential and uniform distributions
hoelzl
parents: 50244
diff changeset
  1067
lemma distributedI_real:
3177d0374701 add exponential and uniform distributions
hoelzl
parents: 50244
diff changeset
  1068
  fixes f :: "'a \<Rightarrow> real"
3177d0374701 add exponential and uniform distributions
hoelzl
parents: 50244
diff changeset
  1069
  assumes gen: "sets M1 = sigma_sets (space M1) E" and "Int_stable E"
3177d0374701 add exponential and uniform distributions
hoelzl
parents: 50244
diff changeset
  1070
    and A: "range A \<subseteq> E" "(\<Union>i::nat. A i) = space M1" "\<And>i. emeasure (distr M M1 X) (A i) \<noteq> \<infinity>"
3177d0374701 add exponential and uniform distributions
hoelzl
parents: 50244
diff changeset
  1071
    and X: "X \<in> measurable M M1"
3177d0374701 add exponential and uniform distributions
hoelzl
parents: 50244
diff changeset
  1072
    and f: "f \<in> borel_measurable M1" "AE x in M1. 0 \<le> f x"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51683
diff changeset
  1073
    and eq: "\<And>A. A \<in> E \<Longrightarrow> emeasure M (X -` A \<inter> space M) = (\<integral>\<^sup>+ x. f x * indicator A x \<partial>M1)"
50419
3177d0374701 add exponential and uniform distributions
hoelzl
parents: 50244
diff changeset
  1074
  shows "distributed M M1 X f"
3177d0374701 add exponential and uniform distributions
hoelzl
parents: 50244
diff changeset
  1075
  unfolding distributed_def
3177d0374701 add exponential and uniform distributions
hoelzl
parents: 50244
diff changeset
  1076
proof (intro conjI)
3177d0374701 add exponential and uniform distributions
hoelzl
parents: 50244
diff changeset
  1077
  show "distr M M1 X = density M1 f"
3177d0374701 add exponential and uniform distributions
hoelzl
parents: 50244
diff changeset
  1078
  proof (rule measure_eqI_generator_eq[where A=A])
3177d0374701 add exponential and uniform distributions
hoelzl
parents: 50244
diff changeset
  1079
    { fix A assume A: "A \<in> E"
3177d0374701 add exponential and uniform distributions
hoelzl
parents: 50244
diff changeset
  1080
      then have "A \<in> sigma_sets (space M1) E" by auto
3177d0374701 add exponential and uniform distributions
hoelzl
parents: 50244
diff changeset
  1081
      then have "A \<in> sets M1"
3177d0374701 add exponential and uniform distributions
hoelzl
parents: 50244
diff changeset
  1082
        using gen by simp
3177d0374701 add exponential and uniform distributions
hoelzl
parents: 50244
diff changeset
  1083
      with f A eq[of A] X show "emeasure (distr M M1 X) A = emeasure (density M1 f) A"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1084
        by (auto simp add: emeasure_distr emeasure_density ennreal_indicator
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1085
                 intro!: nn_integral_cong split: split_indicator) }
50419
3177d0374701 add exponential and uniform distributions
hoelzl
parents: 50244
diff changeset
  1086
    note eq_E = this
3177d0374701 add exponential and uniform distributions
hoelzl
parents: 50244
diff changeset
  1087
    show "Int_stable E" by fact
3177d0374701 add exponential and uniform distributions
hoelzl
parents: 50244
diff changeset
  1088
    { fix e assume "e \<in> E"
3177d0374701 add exponential and uniform distributions
hoelzl
parents: 50244
diff changeset
  1089
      then have "e \<in> sigma_sets (space M1) E" by auto
3177d0374701 add exponential and uniform distributions
hoelzl
parents: 50244
diff changeset
  1090
      then have "e \<in> sets M1" unfolding gen .
3177d0374701 add exponential and uniform distributions
hoelzl
parents: 50244
diff changeset
  1091
      then have "e \<subseteq> space M1" by (rule sets.sets_into_space) }
3177d0374701 add exponential and uniform distributions
hoelzl
parents: 50244
diff changeset
  1092
    then show "E \<subseteq> Pow (space M1)" by auto
3177d0374701 add exponential and uniform distributions
hoelzl
parents: 50244
diff changeset
  1093
    show "sets (distr M M1 X) = sigma_sets (space M1) E"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1094
      "sets (density M1 (\<lambda>x. ennreal (f x))) = sigma_sets (space M1) E"
50419
3177d0374701 add exponential and uniform distributions
hoelzl
parents: 50244
diff changeset
  1095
      unfolding gen[symmetric] by auto
3177d0374701 add exponential and uniform distributions
hoelzl
parents: 50244
diff changeset
  1096
  qed fact+
3177d0374701 add exponential and uniform distributions
hoelzl
parents: 50244
diff changeset
  1097
qed (insert X f, auto)
3177d0374701 add exponential and uniform distributions
hoelzl
parents: 50244
diff changeset
  1098
3177d0374701 add exponential and uniform distributions
hoelzl
parents: 50244
diff changeset
  1099
lemma distributedI_borel_atMost:
3177d0374701 add exponential and uniform distributions
hoelzl
parents: 50244
diff changeset
  1100
  fixes f :: "real \<Rightarrow> real"
3177d0374701 add exponential and uniform distributions
hoelzl
parents: 50244
diff changeset
  1101
  assumes [measurable]: "X \<in> borel_measurable M"
3177d0374701 add exponential and uniform distributions
hoelzl
parents: 50244
diff changeset
  1102
    and [measurable]: "f \<in> borel_measurable borel" and f[simp]: "AE x in lborel. 0 \<le> f x"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1103
    and g_eq: "\<And>a. (\<integral>\<^sup>+x. f x * indicator {..a} x \<partial>lborel)  = ennreal (g a)"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1104
    and M_eq: "\<And>a. emeasure M {x\<in>space M. X x \<le> a} = ennreal (g a)"
50419
3177d0374701 add exponential and uniform distributions
hoelzl
parents: 50244
diff changeset
  1105
  shows "distributed M lborel X f"
3177d0374701 add exponential and uniform distributions
hoelzl
parents: 50244
diff changeset
  1106
proof (rule distributedI_real)
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57418
diff changeset
  1107
  show "sets (lborel::real measure) = sigma_sets (space lborel) (range atMost)"
50419
3177d0374701 add exponential and uniform distributions
hoelzl
parents: 50244
diff changeset
  1108
    by (simp add: borel_eq_atMost)
3177d0374701 add exponential and uniform distributions
hoelzl
parents: 50244
diff changeset
  1109
  show "Int_stable (range atMost :: real set set)"
3177d0374701 add exponential and uniform distributions
hoelzl
parents: 50244
diff changeset
  1110
    by (auto simp: Int_stable_def)
3177d0374701 add exponential and uniform distributions
hoelzl
parents: 50244
diff changeset
  1111
  have vimage_eq: "\<And>a. (X -` {..a} \<inter> space M) = {x\<in>space M. X x \<le> a}" by auto
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62975
diff changeset
  1112
  define A where "A i = {.. real i}" for i :: nat
50419
3177d0374701 add exponential and uniform distributions
hoelzl
parents: 50244
diff changeset
  1113
  then show "range A \<subseteq> range atMost" "(\<Union>i. A i) = space lborel"
3177d0374701 add exponential and uniform distributions
hoelzl
parents: 50244
diff changeset
  1114
    "\<And>i. emeasure (distr M lborel X) (A i) \<noteq> \<infinity>"
3177d0374701 add exponential and uniform distributions
hoelzl
parents: 50244
diff changeset
  1115
    by (auto simp: real_arch_simple emeasure_distr vimage_eq M_eq)
3177d0374701 add exponential and uniform distributions
hoelzl
parents: 50244
diff changeset
  1116
3177d0374701 add exponential and uniform distributions
hoelzl
parents: 50244
diff changeset
  1117
  fix A :: "real set" assume "A \<in> range atMost"
3177d0374701 add exponential and uniform distributions
hoelzl
parents: 50244
diff changeset
  1118
  then obtain a where A: "A = {..a}" by auto
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51683
diff changeset
  1119
  show "emeasure M (X -` A \<inter> space M) = (\<integral>\<^sup>+x. f x * indicator A x \<partial>lborel)"
50419
3177d0374701 add exponential and uniform distributions
hoelzl
parents: 50244
diff changeset
  1120
    unfolding vimage_eq A M_eq g_eq ..
3177d0374701 add exponential and uniform distributions
hoelzl
parents: 50244
diff changeset
  1121
qed auto
3177d0374701 add exponential and uniform distributions
hoelzl
parents: 50244
diff changeset
  1122
3177d0374701 add exponential and uniform distributions
hoelzl
parents: 50244
diff changeset
  1123
lemma (in prob_space) uniform_distributed_params:
3177d0374701 add exponential and uniform distributions
hoelzl
parents: 50244
diff changeset
  1124
  assumes X: "distributed M MX X (\<lambda>x. indicator A x / measure MX A)"
3177d0374701 add exponential and uniform distributions
hoelzl
parents: 50244
diff changeset
  1125
  shows "A \<in> sets MX" "measure MX A \<noteq> 0"
3177d0374701 add exponential and uniform distributions
hoelzl
parents: 50244
diff changeset
  1126
proof -
3177d0374701 add exponential and uniform distributions
hoelzl
parents: 50244
diff changeset
  1127
  interpret X: prob_space "distr M MX X"
3177d0374701 add exponential and uniform distributions
hoelzl
parents: 50244
diff changeset
  1128
    using distributed_measurable[OF X] by (rule prob_space_distr)
3177d0374701 add exponential and uniform distributions
hoelzl
parents: 50244
diff changeset
  1129
3177d0374701 add exponential and uniform distributions
hoelzl
parents: 50244
diff changeset
  1130
  show "measure MX A \<noteq> 0"
3177d0374701 add exponential and uniform distributions
hoelzl
parents: 50244
diff changeset
  1131
  proof
3177d0374701 add exponential and uniform distributions
hoelzl
parents: 50244
diff changeset
  1132
    assume "measure MX A = 0"
3177d0374701 add exponential and uniform distributions
hoelzl
parents: 50244
diff changeset
  1133
    with X.emeasure_space_1 X.prob_space distributed_distr_eq_density[OF X]
3177d0374701 add exponential and uniform distributions
hoelzl
parents: 50244
diff changeset
  1134
    show False
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1135
      by (simp add: emeasure_density zero_ennreal_def[symmetric])
50419
3177d0374701 add exponential and uniform distributions
hoelzl
parents: 50244
diff changeset
  1136
  qed
3177d0374701 add exponential and uniform distributions
hoelzl
parents: 50244
diff changeset
  1137
  with measure_notin_sets[of A MX] show "A \<in> sets MX"
3177d0374701 add exponential and uniform distributions
hoelzl
parents: 50244
diff changeset
  1138
    by blast
3177d0374701 add exponential and uniform distributions
hoelzl
parents: 50244
diff changeset
  1139
qed
3177d0374701 add exponential and uniform distributions
hoelzl
parents: 50244
diff changeset
  1140
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1141
lemma prob_space_uniform_measure:
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1142
  assumes A: "emeasure M A \<noteq> 0" "emeasure M A \<noteq> \<infinity>"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1143
  shows "prob_space (uniform_measure M A)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1144
proof
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1145
  show "emeasure (uniform_measure M A) (space (uniform_measure M A)) = 1"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1146
    using emeasure_uniform_measure[OF emeasure_neq_0_sets[OF A(1)], of "space M"]
50244
de72bbe42190 qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents: 50104
diff changeset
  1147
    using sets.sets_into_space[OF emeasure_neq_0_sets[OF A(1)]] A
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1148
    by (simp add: Int_absorb2 less_top)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1149
qed
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1150
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1151
lemma prob_space_uniform_count_measure: "finite A \<Longrightarrow> A \<noteq> {} \<Longrightarrow> prob_space (uniform_count_measure A)"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1152
  by standard (auto simp: emeasure_uniform_count_measure space_uniform_count_measure one_ennreal_def)
42860
b02349e70d5a add Bernoulli space
hoelzl
parents: 42859
diff changeset
  1153
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1154
lemma (in prob_space) measure_uniform_measure_eq_cond_prob:
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1155
  assumes [measurable]: "Measurable.pred M P" "Measurable.pred M Q"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1156
  shows "\<P>(x in uniform_measure M {x\<in>space M. Q x}. P x) = \<P>(x in M. P x \<bar> Q x)"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1157
proof cases
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1158
  assume Q: "measure M {x\<in>space M. Q x} = 0"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1159
  then have *: "AE x in M. \<not> Q x"
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1160
    by (simp add: prob_eq_0)
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1161
  then have "density M (\<lambda>x. indicator {x \<in> space M. Q x} x / emeasure M {x \<in> space M. Q x}) = density M (\<lambda>x. 0)"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1162
    by (intro density_cong) auto
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1163
  with * show ?thesis
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1164
    unfolding uniform_measure_def
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1165
    by (simp add: emeasure_density measure_def cond_prob_def emeasure_eq_0_AE)
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1166
next
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1167
  assume Q: "measure M {x\<in>space M. Q x} \<noteq> 0"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1168
  then show "\<P>(x in uniform_measure M {x \<in> space M. Q x}. P x) = cond_prob M P Q"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1169
    by (subst measure_uniform_measure)
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1170
       (auto simp: emeasure_eq_measure cond_prob_def measure_nonneg intro!: arg_cong[where f=prob])
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1171
qed
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1172
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1173
lemma prob_space_point_measure:
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1174
  "finite S \<Longrightarrow> (\<And>s. s \<in> S \<Longrightarrow> 0 \<le> p s) \<Longrightarrow> (\<Sum>s\<in>S. p s) = 1 \<Longrightarrow> prob_space (point_measure S p)"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1175
  by (rule prob_spaceI) (simp add: space_point_measure emeasure_point_measure_finite)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58876
diff changeset
  1176
61359
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
  1177
lemma (in prob_space) distr_pair_fst: "distr (N \<Otimes>\<^sub>M M) N fst = N"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
  1178
proof (intro measure_eqI)
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
  1179
  fix A assume A: "A \<in> sets (distr (N \<Otimes>\<^sub>M M) N fst)"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
  1180
  from A have "emeasure (distr (N \<Otimes>\<^sub>M M) N fst) A = emeasure (N \<Otimes>\<^sub>M M) (A \<times> space M)"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
  1181
    by (auto simp add: emeasure_distr space_pair_measure dest: sets.sets_into_space intro!: arg_cong2[where f=emeasure])
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
  1182
  with A show "emeasure (distr (N \<Otimes>\<^sub>M M) N fst) A = emeasure N A"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
  1183
    by (simp add: emeasure_pair_measure_Times emeasure_space_1)
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
  1184
qed simp
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
  1185
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
  1186
lemma (in product_prob_space) distr_reorder:
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
  1187
  assumes "inj_on t J" "t \<in> J \<rightarrow> K" "finite K"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
  1188
  shows "distr (PiM K M) (Pi\<^sub>M J (\<lambda>x. M (t x))) (\<lambda>\<omega>. \<lambda>n\<in>J. \<omega> (t n)) = PiM J (\<lambda>x. M (t x))"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
  1189
proof (rule product_sigma_finite.PiM_eqI)
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
  1190
  show "product_sigma_finite (\<lambda>x. M (t x))" ..
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
  1191
  have "t`J \<subseteq> K" using assms by auto
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
  1192
  then show [simp]: "finite J"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
  1193
    by (rule finite_imageD[OF finite_subset]) fact+
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
  1194
  fix A assume A: "\<And>i. i \<in> J \<Longrightarrow> A i \<in> sets (M (t i))"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
  1195
  moreover have "((\<lambda>\<omega>. \<lambda>n\<in>J. \<omega> (t n)) -` Pi\<^sub>E J A \<inter> space (Pi\<^sub>M K M)) =
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
  1196
    (\<Pi>\<^sub>E i\<in>K. if i \<in> t`J then A (the_inv_into J t i) else space (M i))"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
  1197
    using A A[THEN sets.sets_into_space] \<open>t \<in> J \<rightarrow> K\<close> \<open>inj_on t J\<close>
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
  1198
    by (subst prod_emb_Pi[symmetric]) (auto simp: space_PiM PiE_iff the_inv_into_f_f prod_emb_def)
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
  1199
  ultimately show "distr (Pi\<^sub>M K M) (Pi\<^sub>M J (\<lambda>x. M (t x))) (\<lambda>\<omega>. \<lambda>n\<in>J. \<omega> (t n)) (Pi\<^sub>E J A) = (\<Prod>i\<in>J. M (t i) (A i))"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
  1200
    using assms
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
  1201
    apply (subst emeasure_distr)
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
  1202
    apply (auto intro!: sets_PiM_I_finite simp: Pi_iff)
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
  1203
    apply (subst emeasure_PiM)
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
  1204
    apply (auto simp: the_inv_into_f_f \<open>inj_on t J\<close> setprod.reindex[OF \<open>inj_on t J\<close>]
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
  1205
      if_distrib[where f="emeasure (M _)"] setprod.If_cases emeasure_space_1 Int_absorb1 \<open>t`J \<subseteq> K\<close>)
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
  1206
    done
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
  1207
qed simp
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
  1208
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
  1209
lemma (in product_prob_space) distr_restrict:
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
  1210
  "J \<subseteq> K \<Longrightarrow> finite K \<Longrightarrow> (\<Pi>\<^sub>M i\<in>J. M i) = distr (\<Pi>\<^sub>M i\<in>K. M i) (\<Pi>\<^sub>M i\<in>J. M i) (\<lambda>f. restrict f J)"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
  1211
  using distr_reorder[of "\<lambda>x. x" J K] by (simp add: Pi_iff subset_eq)
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
  1212
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
  1213
lemma (in product_prob_space) emeasure_prod_emb[simp]:
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
  1214
  assumes L: "J \<subseteq> L" "finite L" and X: "X \<in> sets (Pi\<^sub>M J M)"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
  1215
  shows "emeasure (Pi\<^sub>M L M) (prod_emb L M J X) = emeasure (Pi\<^sub>M J M) X"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
  1216
  by (subst distr_restrict[OF L])
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
  1217
     (simp add: prod_emb_def space_PiM emeasure_distr measurable_restrict_subset L X)
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
  1218
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
  1219
lemma emeasure_distr_restrict:
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
  1220
  assumes "I \<subseteq> K" and Q[measurable_cong]: "sets Q = sets (PiM K M)" and A[measurable]: "A \<in> sets (PiM I M)"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
  1221
  shows "emeasure (distr Q (PiM I M) (\<lambda>\<omega>. restrict \<omega> I)) A = emeasure Q (prod_emb K M I A)"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
  1222
  using \<open>I\<subseteq>K\<close> sets_eq_imp_space_eq[OF Q]
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
  1223
  by (subst emeasure_distr)
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
  1224
     (auto simp: measurable_cong_sets[OF Q] prod_emb_def space_PiM[symmetric] intro!: measurable_restrict)
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
  1225
63626
44ce6b524ff3 move measure theory from HOL-Probability to HOL-Multivariate_Analysis
hoelzl
parents: 63099
diff changeset
  1226
lemma (in prob_space) prob_space_completion: "prob_space (completion M)"
44ce6b524ff3 move measure theory from HOL-Probability to HOL-Multivariate_Analysis
hoelzl
parents: 63099
diff changeset
  1227
  by (rule prob_spaceI) (simp add: emeasure_space_1)
44ce6b524ff3 move measure theory from HOL-Probability to HOL-Multivariate_Analysis
hoelzl
parents: 63099
diff changeset
  1228
35582
b16d99a72dc9 Add Lebesgue integral and probability space.
hoelzl
parents:
diff changeset
  1229
end