src/HOL/Probability/Probability_Mass_Function.thy
author wenzelm
Fri, 20 Sep 2024 19:51:08 +0200
changeset 80914 d97fdabd9e2b
parent 77434 da41823d09a7
child 81995 d67dadd69d07
permissions -rw-r--r--
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
58606
9c66f7c541fb add Giry monad
hoelzl
parents: 58587
diff changeset
     1
(*  Title:      HOL/Probability/Probability_Mass_Function.thy
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59665
diff changeset
     2
    Author:     Johannes Hölzl, TU München
59023
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
     3
    Author:     Andreas Lochbihler, ETH Zurich
73253
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
     4
    Author:     Manuel Eberl, TU München
59023
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
     5
*)
58606
9c66f7c541fb add Giry monad
hoelzl
parents: 58587
diff changeset
     6
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
     7
section \<open> Probability mass function \<close>
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
     8
58587
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
     9
theory Probability_Mass_Function
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    10
imports
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    11
  Giry_Monad
66453
cc19f7ca2ed6 session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents: 66192
diff changeset
    12
  "HOL-Library.Multiset"
58587
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
    13
begin
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
    14
77434
da41823d09a7 Imported a theorem about Infinite_Sum. Importing this theory a bit earlier is causing syntactic ambiguities with Infinite_Set_Sum however; no_notation needed
paulson <lp15@cam.ac.uk>
parents: 75625
diff changeset
    15
text \<open>Conflicting notation from \<^theory>\<open>HOL-Analysis.Infinite_Sum\<close>\<close>
80914
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 77434
diff changeset
    16
no_notation Infinite_Sum.abs_summable_on (infixr \<open>abs'_summable'_on\<close> 46)
77434
da41823d09a7 Imported a theorem about Infinite_Sum. Importing this theory a bit earlier is causing syntactic ambiguities with Infinite_Set_Sum however; no_notation needed
paulson <lp15@cam.ac.uk>
parents: 75625
diff changeset
    17
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
    18
lemma AE_emeasure_singleton:
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
    19
  assumes x: "emeasure M {x} \<noteq> 0" and ae: "AE x in M. P x" shows "P x"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
    20
proof -
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
    21
  from x have x_M: "{x} \<in> sets M"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
    22
    by (auto intro: emeasure_notin_sets)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
    23
  from ae obtain N where N: "{x\<in>space M. \<not> P x} \<subseteq> N" "emeasure M N = 0" "N \<in> sets M"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
    24
    by (auto elim: AE_E)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
    25
  { assume "\<not> P x"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
    26
    with x_M[THEN sets.sets_into_space] N have "emeasure M {x} \<le> emeasure M N"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
    27
      by (intro emeasure_mono) auto
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
    28
    with x N have False
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
    29
      by (auto simp:) }
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
    30
  then show "P x" by auto
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
    31
qed
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
    32
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
    33
lemma AE_measure_singleton: "measure M {x} \<noteq> 0 \<Longrightarrow> AE x in M. P x \<Longrightarrow> P x"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
    34
  by (metis AE_emeasure_singleton measure_def emeasure_empty measure_empty)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
    35
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    36
lemma (in finite_measure) AE_support_countable:
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    37
  assumes [simp]: "sets M = UNIV"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    38
  shows "(AE x in M. measure M {x} \<noteq> 0) \<longleftrightarrow> (\<exists>S. countable S \<and> (AE x in M. x \<in> S))"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    39
proof
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    40
  assume "\<exists>S. countable S \<and> (AE x in M. x \<in> S)"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    41
  then obtain S where S[intro]: "countable S" and ae: "AE x in M. x \<in> S"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    42
    by auto
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59665
diff changeset
    43
  then have "emeasure M (\<Union>x\<in>{x\<in>S. emeasure M {x} \<noteq> 0}. {x}) =
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    44
    (\<integral>\<^sup>+ x. emeasure M {x} * indicator {x\<in>S. emeasure M {x} \<noteq> 0} x \<partial>count_space UNIV)"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    45
    by (subst emeasure_UN_countable)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    46
       (auto simp: disjoint_family_on_def nn_integral_restrict_space[symmetric] restrict_count_space)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    47
  also have "\<dots> = (\<integral>\<^sup>+ x. emeasure M {x} * indicator S x \<partial>count_space UNIV)"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    48
    by (auto intro!: nn_integral_cong split: split_indicator)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    49
  also have "\<dots> = emeasure M (\<Union>x\<in>S. {x})"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    50
    by (subst emeasure_UN_countable)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    51
       (auto simp: disjoint_family_on_def nn_integral_restrict_space[symmetric] restrict_count_space)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    52
  also have "\<dots> = emeasure M (space M)"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    53
    using ae by (intro emeasure_eq_AE) auto
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    54
  finally have "emeasure M {x \<in> space M. x\<in>S \<and> emeasure M {x} \<noteq> 0} = emeasure M (space M)"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    55
    by (simp add: emeasure_single_in_space cong: rev_conj_cong)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    56
  with finite_measure_compl[of "{x \<in> space M. x\<in>S \<and> emeasure M {x} \<noteq> 0}"]
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    57
  have "AE x in M. x \<in> S \<and> emeasure M {x} \<noteq> 0"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
    58
    by (intro AE_I[OF order_refl]) (auto simp: emeasure_eq_measure measure_nonneg set_diff_eq cong: conj_cong)
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    59
  then show "AE x in M. measure M {x} \<noteq> 0"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    60
    by (auto simp: emeasure_eq_measure)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    61
qed (auto intro!: exI[of _ "{x. measure M {x} \<noteq> 0}"] countable_support)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    62
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
    63
subsection \<open> PMF as measure \<close>
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    64
58587
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
    65
typedef 'a pmf = "{M :: 'a measure. prob_space M \<and> sets M = UNIV \<and> (AE x in M. measure M {x} \<noteq> 0)}"
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
    66
  morphisms measure_pmf Abs_pmf
58606
9c66f7c541fb add Giry monad
hoelzl
parents: 58587
diff changeset
    67
  by (intro exI[of _ "uniform_measure (count_space UNIV) {undefined}"])
9c66f7c541fb add Giry monad
hoelzl
parents: 58587
diff changeset
    68
     (auto intro!: prob_space_uniform_measure AE_uniform_measureI)
58587
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
    69
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
    70
declare [[coercion measure_pmf]]
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
    71
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
    72
lemma prob_space_measure_pmf: "prob_space (measure_pmf p)"
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
    73
  using pmf.measure_pmf[of p] by auto
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
    74
61605
1bf7b186542e qualifier is mandatory by default;
wenzelm
parents: 61424
diff changeset
    75
interpretation measure_pmf: prob_space "measure_pmf M" for M
58587
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
    76
  by (rule prob_space_measure_pmf)
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
    77
61605
1bf7b186542e qualifier is mandatory by default;
wenzelm
parents: 61424
diff changeset
    78
interpretation measure_pmf: subprob_space "measure_pmf M" for M
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    79
  by (rule prob_space_imp_subprob_space) unfold_locales
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    80
59048
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59024
diff changeset
    81
lemma subprob_space_measure_pmf: "subprob_space (measure_pmf x)"
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59024
diff changeset
    82
  by unfold_locales
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59024
diff changeset
    83
58587
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
    84
locale pmf_as_measure
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
    85
begin
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
    86
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
    87
setup_lifting type_definition_pmf
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
    88
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
    89
end
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
    90
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
    91
context
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
    92
begin
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
    93
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
    94
interpretation pmf_as_measure .
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
    95
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
    96
lemma sets_measure_pmf[simp]: "sets (measure_pmf p) = UNIV"
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59665
diff changeset
    97
  by transfer blast
58587
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
    98
59048
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59024
diff changeset
    99
lemma sets_measure_pmf_count_space[measurable_cong]:
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59024
diff changeset
   100
  "sets (measure_pmf M) = sets (count_space UNIV)"
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   101
  by simp
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   102
58587
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   103
lemma space_measure_pmf[simp]: "space (measure_pmf p) = UNIV"
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   104
  using sets_eq_imp_space_eq[of "measure_pmf p" "count_space UNIV"] by simp
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   105
61634
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
   106
lemma measure_pmf_UNIV [simp]: "measure (measure_pmf p) UNIV = 1"
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
   107
using measure_pmf.prob_space[of p] by simp
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
   108
59048
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59024
diff changeset
   109
lemma measure_pmf_in_subprob_algebra[measurable (raw)]: "measure_pmf x \<in> space (subprob_algebra (count_space UNIV))"
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59024
diff changeset
   110
  by (simp add: space_subprob_algebra subprob_space_measure_pmf)
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59024
diff changeset
   111
58587
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   112
lemma measurable_pmf_measure1[simp]: "measurable (M :: 'a pmf) N = UNIV \<rightarrow> space N"
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   113
  by (auto simp: measurable_def)
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   114
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   115
lemma measurable_pmf_measure2[simp]: "measurable N (M :: 'a pmf) = measurable N (count_space UNIV)"
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   116
  by (intro measurable_cong_sets) simp_all
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   117
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   118
lemma measurable_pair_restrict_pmf2:
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   119
  assumes "countable A"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   120
  assumes [measurable]: "\<And>y. y \<in> A \<Longrightarrow> (\<lambda>x. f (x, y)) \<in> measurable M L"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   121
  shows "f \<in> measurable (M \<Otimes>\<^sub>M restrict_space (measure_pmf N) A) L" (is "f \<in> measurable ?M _")
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   122
proof -
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   123
  have [measurable_cong]: "sets (restrict_space (count_space UNIV) A) = sets (count_space A)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   124
    by (simp add: restrict_count_space)
58587
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   125
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   126
  show ?thesis
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   127
    by (intro measurable_compose_countable'[where f="\<lambda>a b. f (fst b, a)" and g=snd and I=A,
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61169
diff changeset
   128
                                            unfolded prod.collapse] assms)
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   129
        measurable
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   130
qed
58587
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   131
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   132
lemma measurable_pair_restrict_pmf1:
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   133
  assumes "countable A"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   134
  assumes [measurable]: "\<And>x. x \<in> A \<Longrightarrow> (\<lambda>y. f (x, y)) \<in> measurable N L"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   135
  shows "f \<in> measurable (restrict_space (measure_pmf M) A \<Otimes>\<^sub>M N) L"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   136
proof -
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   137
  have [measurable_cong]: "sets (restrict_space (count_space UNIV) A) = sets (count_space A)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   138
    by (simp add: restrict_count_space)
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   139
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   140
  show ?thesis
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   141
    by (intro measurable_compose_countable'[where f="\<lambda>a b. f (a, snd b)" and g=fst and I=A,
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61169
diff changeset
   142
                                            unfolded prod.collapse] assms)
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   143
        measurable
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   144
qed
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   145
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   146
lift_definition pmf :: "'a pmf \<Rightarrow> 'a \<Rightarrow> real" is "\<lambda>M x. measure M {x}" .
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   147
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   148
lift_definition set_pmf :: "'a pmf \<Rightarrow> 'a set" is "\<lambda>M. {x. measure M {x} \<noteq> 0}" .
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   149
declare [[coercion set_pmf]]
58587
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   150
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   151
lemma AE_measure_pmf: "AE x in (M::'a pmf). x \<in> M"
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   152
  by transfer simp
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   153
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   154
lemma emeasure_pmf_single_eq_zero_iff:
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   155
  fixes M :: "'a pmf"
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   156
  shows "emeasure M {y} = 0 \<longleftrightarrow> y \<notin> M"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   157
  unfolding set_pmf.rep_eq by (simp add: measure_pmf.emeasure_eq_measure)
58587
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   158
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   159
lemma AE_measure_pmf_iff: "(AE x in measure_pmf M. P x) \<longleftrightarrow> (\<forall>y\<in>M. P y)"
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   160
  using AE_measure_singleton[of M] AE_measure_pmf[of M]
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   161
  by (auto simp: set_pmf.rep_eq)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   162
61634
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
   163
lemma AE_pmfI: "(\<And>y. y \<in> set_pmf M \<Longrightarrow> P y) \<Longrightarrow> almost_everywhere (measure_pmf M) P"
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
   164
by(simp add: AE_measure_pmf_iff)
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
   165
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   166
lemma countable_set_pmf [simp]: "countable (set_pmf p)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   167
  by transfer (metis prob_space.finite_measure finite_measure.countable_support)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   168
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   169
lemma pmf_positive: "x \<in> set_pmf p \<Longrightarrow> 0 < pmf p x"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   170
  by transfer (simp add: less_le)
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   171
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   172
lemma pmf_nonneg[simp]: "0 \<le> pmf p x"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   173
  by transfer simp
63886
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63882
diff changeset
   174
63099
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
   175
lemma pmf_not_neg [simp]: "\<not>pmf p x < 0"
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
   176
  by (simp add: not_less pmf_nonneg)
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
   177
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
   178
lemma pmf_pos [simp]: "pmf p x \<noteq> 0 \<Longrightarrow> pmf p x > 0"
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
   179
  using pmf_nonneg[of p x] by linarith
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   180
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   181
lemma pmf_le_1: "pmf p x \<le> 1"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   182
  by (simp add: pmf.rep_eq)
58587
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   183
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   184
lemma set_pmf_not_empty: "set_pmf M \<noteq> {}"
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   185
  using AE_measure_pmf[of M] by (intro notI) simp
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   186
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   187
lemma set_pmf_iff: "x \<in> set_pmf M \<longleftrightarrow> pmf M x \<noteq> 0"
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   188
  by transfer simp
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   189
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   190
lemma pmf_positive_iff: "0 < pmf p x \<longleftrightarrow> x \<in> set_pmf p"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   191
  unfolding less_le by (simp add: set_pmf_iff)
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   192
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   193
lemma set_pmf_eq: "set_pmf M = {x. pmf M x \<noteq> 0}"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   194
  by (auto simp: set_pmf_iff)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   195
63099
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
   196
lemma set_pmf_eq': "set_pmf p = {x. pmf p x > 0}"
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
   197
proof safe
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
   198
  fix x assume "x \<in> set_pmf p"
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
   199
  hence "pmf p x \<noteq> 0" by (auto simp: set_pmf_eq)
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
   200
  with pmf_nonneg[of p x] show "pmf p x > 0" by simp
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
   201
qed (auto simp: set_pmf_eq)
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
   202
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   203
lemma emeasure_pmf_single:
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   204
  fixes M :: "'a pmf"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   205
  shows "emeasure M {x} = pmf M x"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   206
  by transfer (simp add: finite_measure.emeasure_eq_measure[OF prob_space.finite_measure])
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   207
60068
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
   208
lemma measure_pmf_single: "measure (measure_pmf M) {x} = pmf M x"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   209
  using emeasure_pmf_single[of M x] by(simp add: measure_pmf.emeasure_eq_measure pmf_nonneg measure_nonneg)
60068
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
   210
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   211
lemma emeasure_measure_pmf_finite: "finite S \<Longrightarrow> emeasure (measure_pmf M) S = (\<Sum>s\<in>S. pmf M s)"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64008
diff changeset
   212
  by (subst emeasure_eq_sum_singleton) (auto simp: emeasure_pmf_single pmf_nonneg)
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   213
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64008
diff changeset
   214
lemma measure_measure_pmf_finite: "finite S \<Longrightarrow> measure (measure_pmf M) S = sum (pmf M) S"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   215
  using emeasure_measure_pmf_finite[of S M]
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64008
diff changeset
   216
  by (simp add: measure_pmf.emeasure_eq_measure measure_nonneg sum_nonneg pmf_nonneg)
59023
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
   217
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64008
diff changeset
   218
lemma sum_pmf_eq_1:
63099
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
   219
  assumes "finite A" "set_pmf p \<subseteq> A"
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
   220
  shows   "(\<Sum>x\<in>A. pmf p x) = 1"
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
   221
proof -
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
   222
  have "(\<Sum>x\<in>A. pmf p x) = measure_pmf.prob p A"
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
   223
    by (simp add: measure_measure_pmf_finite assms)
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
   224
  also from assms have "\<dots> = 1"
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
   225
    by (subst measure_pmf.prob_eq_1) (auto simp: AE_measure_pmf_iff)
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
   226
  finally show ?thesis .
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
   227
qed
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
   228
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   229
lemma nn_integral_measure_pmf_support:
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   230
  fixes f :: "'a \<Rightarrow> ennreal"
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   231
  assumes f: "finite A" and nn: "\<And>x. x \<in> A \<Longrightarrow> 0 \<le> f x" "\<And>x. x \<in> set_pmf M \<Longrightarrow> x \<notin> A \<Longrightarrow> f x = 0"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   232
  shows "(\<integral>\<^sup>+x. f x \<partial>measure_pmf M) = (\<Sum>x\<in>A. f x * pmf M x)"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   233
proof -
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   234
  have "(\<integral>\<^sup>+x. f x \<partial>M) = (\<integral>\<^sup>+x. f x * indicator A x \<partial>M)"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   235
    using nn by (intro nn_integral_cong_AE) (auto simp: AE_measure_pmf_iff split: split_indicator)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   236
  also have "\<dots> = (\<Sum>x\<in>A. f x * emeasure M {x})"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   237
    using assms by (intro nn_integral_indicator_finite) auto
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   238
  finally show ?thesis
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   239
    by (simp add: emeasure_measure_pmf_finite)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   240
qed
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   241
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   242
lemma nn_integral_measure_pmf_finite:
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   243
  fixes f :: "'a \<Rightarrow> ennreal"
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   244
  assumes f: "finite (set_pmf M)" and nn: "\<And>x. x \<in> set_pmf M \<Longrightarrow> 0 \<le> f x"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   245
  shows "(\<integral>\<^sup>+x. f x \<partial>measure_pmf M) = (\<Sum>x\<in>set_pmf M. f x * pmf M x)"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   246
  using assms by (intro nn_integral_measure_pmf_support) auto
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   247
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   248
lemma integrable_measure_pmf_finite:
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   249
  fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   250
  shows "finite (set_pmf M) \<Longrightarrow> integrable M f"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   251
  by (auto intro!: integrableI_bounded simp: nn_integral_measure_pmf_finite ennreal_mult_less_top)
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   252
64008
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   253
lemma integral_measure_pmf_real:
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   254
  assumes [simp]: "finite A" and "\<And>a. a \<in> set_pmf M \<Longrightarrow> f a \<noteq> 0 \<Longrightarrow> a \<in> A"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   255
  shows "(\<integral>x. f x \<partial>measure_pmf M) = (\<Sum>a\<in>A. f a * pmf M a)"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   256
proof -
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   257
  have "(\<integral>x. f x \<partial>measure_pmf M) = (\<integral>x. f x * indicator A x \<partial>measure_pmf M)"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   258
    using assms(2) by (intro integral_cong_AE) (auto split: split_indicator simp: AE_measure_pmf_iff)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   259
  also have "\<dots> = (\<Sum>a\<in>A. f a * pmf M a)"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   260
    by (subst integral_indicator_finite_real)
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   261
       (auto simp: measure_def emeasure_measure_pmf_finite pmf_nonneg)
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   262
  finally show ?thesis .
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   263
qed
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   264
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   265
lemma integrable_pmf: "integrable (count_space X) (pmf M)"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   266
proof -
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   267
  have " (\<integral>\<^sup>+ x. pmf M x \<partial>count_space X) = (\<integral>\<^sup>+ x. pmf M x \<partial>count_space (M \<inter> X))"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   268
    by (auto simp add: nn_integral_count_space_indicator set_pmf_iff intro!: nn_integral_cong split: split_indicator)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   269
  then have "integrable (count_space X) (pmf M) = integrable (count_space (M \<inter> X)) (pmf M)"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   270
    by (simp add: integrable_iff_bounded pmf_nonneg)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   271
  then show ?thesis
59023
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
   272
    by (simp add: pmf.rep_eq measure_pmf.integrable_measure disjoint_family_on_def)
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   273
qed
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   274
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   275
lemma integral_pmf: "(\<integral>x. pmf M x \<partial>count_space X) = measure M X"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   276
proof -
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   277
  have "(\<integral>x. pmf M x \<partial>count_space X) = (\<integral>\<^sup>+x. pmf M x \<partial>count_space X)"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   278
    by (simp add: pmf_nonneg integrable_pmf nn_integral_eq_integral)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   279
  also have "\<dots> = (\<integral>\<^sup>+x. emeasure M {x} \<partial>count_space (X \<inter> M))"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   280
    by (auto intro!: nn_integral_cong_AE split: split_indicator
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   281
             simp: pmf.rep_eq measure_pmf.emeasure_eq_measure nn_integral_count_space_indicator
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   282
                   AE_count_space set_pmf_iff)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   283
  also have "\<dots> = emeasure M (X \<inter> M)"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   284
    by (rule emeasure_countable_singleton[symmetric]) (auto intro: countable_set_pmf)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   285
  also have "\<dots> = emeasure M X"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   286
    by (auto intro!: emeasure_eq_AE simp: AE_measure_pmf_iff)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   287
  finally show ?thesis
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   288
    by (simp add: measure_pmf.emeasure_eq_measure measure_nonneg integral_nonneg pmf_nonneg)
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   289
qed
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   290
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   291
lemma integral_pmf_restrict:
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   292
  "(f::'a \<Rightarrow> 'b::{banach, second_countable_topology}) \<in> borel_measurable (count_space UNIV) \<Longrightarrow>
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   293
    (\<integral>x. f x \<partial>measure_pmf M) = (\<integral>x. f x \<partial>restrict_space M M)"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   294
  by (auto intro!: integral_cong_AE simp add: integral_restrict_space AE_measure_pmf_iff)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   295
58587
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   296
lemma emeasure_pmf: "emeasure (M::'a pmf) M = 1"
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   297
proof -
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   298
  have "emeasure (M::'a pmf) M = emeasure (M::'a pmf) (space M)"
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   299
    by (intro emeasure_eq_AE) (simp_all add: AE_measure_pmf)
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   300
  then show ?thesis
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   301
    using measure_pmf.emeasure_space_1 by simp
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   302
qed
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   303
59490
f71732294f29 tune proof
Andreas Lochbihler
parents: 59475
diff changeset
   304
lemma emeasure_pmf_UNIV [simp]: "emeasure (measure_pmf M) UNIV = 1"
f71732294f29 tune proof
Andreas Lochbihler
parents: 59475
diff changeset
   305
using measure_pmf.emeasure_space_1[of M] by simp
f71732294f29 tune proof
Andreas Lochbihler
parents: 59475
diff changeset
   306
59023
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
   307
lemma in_null_sets_measure_pmfI:
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
   308
  "A \<inter> set_pmf p = {} \<Longrightarrow> A \<in> null_sets (measure_pmf p)"
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
   309
using emeasure_eq_0_AE[where ?P="\<lambda>x. x \<in> A" and M="measure_pmf p"]
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
   310
by(auto simp add: null_sets_def AE_measure_pmf_iff)
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
   311
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   312
lemma measure_subprob: "measure_pmf M \<in> space (subprob_algebra (count_space UNIV))"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   313
  by (simp add: space_subprob_algebra subprob_space_measure_pmf)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   314
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   315
subsection \<open> Monad Interpretation \<close>
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   316
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   317
lemma measurable_measure_pmf[measurable]:
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   318
  "(\<lambda>x. measure_pmf (M x)) \<in> measurable (count_space UNIV) (subprob_algebra (count_space UNIV))"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   319
  by (auto simp: space_subprob_algebra intro!: prob_space_imp_subprob_space) unfold_locales
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   320
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   321
lemma bind_measure_pmf_cong:
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   322
  assumes "\<And>x. A x \<in> space (subprob_algebra N)" "\<And>x. B x \<in> space (subprob_algebra N)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   323
  assumes "\<And>i. i \<in> set_pmf x \<Longrightarrow> A i = B i"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   324
  shows "bind (measure_pmf x) A = bind (measure_pmf x) B"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   325
proof (rule measure_eqI)
62026
ea3b1b0413b4 more symbols;
wenzelm
parents: 61808
diff changeset
   326
  show "sets (measure_pmf x \<bind> A) = sets (measure_pmf x \<bind> B)"
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   327
    using assms by (subst (1 2) sets_bind) (auto simp: space_subprob_algebra)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   328
next
62026
ea3b1b0413b4 more symbols;
wenzelm
parents: 61808
diff changeset
   329
  fix X assume "X \<in> sets (measure_pmf x \<bind> A)"
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   330
  then have X: "X \<in> sets N"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   331
    using assms by (subst (asm) sets_bind) (auto simp: space_subprob_algebra)
62026
ea3b1b0413b4 more symbols;
wenzelm
parents: 61808
diff changeset
   332
  show "emeasure (measure_pmf x \<bind> A) X = emeasure (measure_pmf x \<bind> B) X"
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   333
    using assms
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   334
    by (subst (1 2) emeasure_bind[where N=N, OF _ _ X])
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   335
       (auto intro!: nn_integral_cong_AE simp: AE_measure_pmf_iff)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   336
qed
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   337
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   338
lift_definition bind_pmf :: "'a pmf \<Rightarrow> ('a \<Rightarrow> 'b pmf ) \<Rightarrow> 'b pmf" is bind
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   339
proof (clarify, intro conjI)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   340
  fix f :: "'a measure" and g :: "'a \<Rightarrow> 'b measure"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   341
  assume "prob_space f"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   342
  then interpret f: prob_space f .
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   343
  assume "sets f = UNIV" and ae_f: "AE x in f. measure f {x} \<noteq> 0"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   344
  then have s_f[simp]: "sets f = sets (count_space UNIV)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   345
    by simp
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   346
  assume g: "\<And>x. prob_space (g x) \<and> sets (g x) = UNIV \<and> (AE y in g x. measure (g x) {y} \<noteq> 0)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   347
  then have g: "\<And>x. prob_space (g x)" and s_g[simp]: "\<And>x. sets (g x) = sets (count_space UNIV)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   348
    and ae_g: "\<And>x. AE y in g x. measure (g x) {y} \<noteq> 0"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   349
    by auto
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   350
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   351
  have [measurable]: "g \<in> measurable f (subprob_algebra (count_space UNIV))"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   352
    by (auto simp: measurable_def space_subprob_algebra prob_space_imp_subprob_space g)
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59665
diff changeset
   353
62026
ea3b1b0413b4 more symbols;
wenzelm
parents: 61808
diff changeset
   354
  show "prob_space (f \<bind> g)"
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   355
    using g by (intro f.prob_space_bind[where S="count_space UNIV"]) auto
62026
ea3b1b0413b4 more symbols;
wenzelm
parents: 61808
diff changeset
   356
  then interpret fg: prob_space "f \<bind> g" .
ea3b1b0413b4 more symbols;
wenzelm
parents: 61808
diff changeset
   357
  show [simp]: "sets (f \<bind> g) = UNIV"
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   358
    using sets_eq_imp_space_eq[OF s_f]
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   359
    by (subst sets_bind[where N="count_space UNIV"]) auto
62026
ea3b1b0413b4 more symbols;
wenzelm
parents: 61808
diff changeset
   360
  show "AE x in f \<bind> g. measure (f \<bind> g) {x} \<noteq> 0"
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   361
    apply (simp add: fg.prob_eq_0 AE_bind[where B="count_space UNIV"])
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   362
    using ae_f
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   363
    apply eventually_elim
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   364
    using ae_g
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   365
    apply eventually_elim
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   366
    apply (auto dest: AE_measure_singleton)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   367
    done
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   368
qed
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   369
63099
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
   370
adhoc_overloading Monad_Syntax.bind bind_pmf
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
   371
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   372
lemma ennreal_pmf_bind: "pmf (bind_pmf N f) i = (\<integral>\<^sup>+x. pmf (f x) i \<partial>measure_pmf N)"
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   373
  unfolding pmf.rep_eq bind_pmf.rep_eq
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   374
  by (auto simp: measure_pmf.measure_bind[where N="count_space UNIV"] measure_subprob measure_nonneg
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   375
           intro!: nn_integral_eq_integral[symmetric] measure_pmf.integrable_const_bound[where B=1])
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   376
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   377
lemma pmf_bind: "pmf (bind_pmf N f) i = (\<integral>x. pmf (f x) i \<partial>measure_pmf N)"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   378
  using ennreal_pmf_bind[of N f i]
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   379
  by (subst (asm) nn_integral_eq_integral)
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   380
     (auto simp: pmf_nonneg pmf_le_1 pmf_nonneg integral_nonneg
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   381
           intro!: nn_integral_eq_integral[symmetric] measure_pmf.integrable_const_bound[where B=1])
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   382
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   383
lemma bind_pmf_const[simp]: "bind_pmf M (\<lambda>x. c) = c"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   384
  by transfer (simp add: bind_const' prob_space_imp_subprob_space)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   385
59665
37adca7fd48f add set_pmf lemmas to simpset
hoelzl
parents: 59664
diff changeset
   386
lemma set_bind_pmf[simp]: "set_pmf (bind_pmf M N) = (\<Union>M\<in>set_pmf M. set_pmf (N M))"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   387
proof -
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   388
  have "set_pmf (bind_pmf M N) = {x. ennreal (pmf (bind_pmf M N) x) \<noteq> 0}"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   389
    by (simp add: set_pmf_eq pmf_nonneg)
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   390
  also have "\<dots> = (\<Union>M\<in>set_pmf M. set_pmf (N M))"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   391
    unfolding ennreal_pmf_bind
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   392
    by (subst nn_integral_0_iff_AE) (auto simp: AE_measure_pmf_iff pmf_nonneg set_pmf_eq)
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   393
  finally show ?thesis .
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   394
qed
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   395
63099
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
   396
lemma bind_pmf_cong [fundef_cong]:
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   397
  assumes "p = q"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   398
  shows "(\<And>x. x \<in> set_pmf q \<Longrightarrow> f x = g x) \<Longrightarrow> bind_pmf p f = bind_pmf q g"
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61634
diff changeset
   399
  unfolding \<open>p = q\<close>[symmetric] measure_pmf_inject[symmetric] bind_pmf.rep_eq
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   400
  by (auto simp: AE_measure_pmf_iff Pi_iff space_subprob_algebra subprob_space_measure_pmf
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   401
                 sets_bind[where N="count_space UNIV"] emeasure_bind[where N="count_space UNIV"]
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   402
           intro!: nn_integral_cong_AE measure_eqI)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   403
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   404
lemma bind_pmf_cong_simp:
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   405
  "p = q \<Longrightarrow> (\<And>x. x \<in> set_pmf q =simp=> f x = g x) \<Longrightarrow> bind_pmf p f = bind_pmf q g"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   406
  by (simp add: simp_implies_def cong: bind_pmf_cong)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   407
62026
ea3b1b0413b4 more symbols;
wenzelm
parents: 61808
diff changeset
   408
lemma measure_pmf_bind: "measure_pmf (bind_pmf M f) = (measure_pmf M \<bind> (\<lambda>x. measure_pmf (f x)))"
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   409
  by transfer simp
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   410
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   411
lemma nn_integral_bind_pmf[simp]: "(\<integral>\<^sup>+x. f x \<partial>bind_pmf M N) = (\<integral>\<^sup>+x. \<integral>\<^sup>+y. f y \<partial>N x \<partial>M)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   412
  using measurable_measure_pmf[of N]
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   413
  unfolding measure_pmf_bind
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   414
  apply (intro nn_integral_bind[where B="count_space UNIV"])
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   415
  apply auto
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   416
  done
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   417
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   418
lemma emeasure_bind_pmf[simp]: "emeasure (bind_pmf M N) X = (\<integral>\<^sup>+x. emeasure (N x) X \<partial>M)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   419
  using measurable_measure_pmf[of N]
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   420
  unfolding measure_pmf_bind
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   421
  by (subst emeasure_bind[where N="count_space UNIV"]) auto
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59665
diff changeset
   422
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   423
lift_definition return_pmf :: "'a \<Rightarrow> 'a pmf" is "return (count_space UNIV)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   424
  by (auto intro!: prob_space_return simp: AE_return measure_return)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   425
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   426
lemma bind_return_pmf: "bind_pmf (return_pmf x) f = f x"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   427
  by transfer
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   428
     (auto intro!: prob_space_imp_subprob_space bind_return[where N="count_space UNIV"]
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   429
           simp: space_subprob_algebra)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   430
59665
37adca7fd48f add set_pmf lemmas to simpset
hoelzl
parents: 59664
diff changeset
   431
lemma set_return_pmf[simp]: "set_pmf (return_pmf x) = {x}"
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   432
  by transfer (auto simp add: measure_return split: split_indicator)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   433
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   434
lemma bind_return_pmf': "bind_pmf N return_pmf = N"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   435
proof (transfer, clarify)
62026
ea3b1b0413b4 more symbols;
wenzelm
parents: 61808
diff changeset
   436
  fix N :: "'a measure" assume "sets N = UNIV" then show "N \<bind> return (count_space UNIV) = N"
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   437
    by (subst return_sets_cong[where N=N]) (simp_all add: bind_return')
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   438
qed
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   439
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   440
lemma bind_assoc_pmf: "bind_pmf (bind_pmf A B) C = bind_pmf A (\<lambda>x. bind_pmf (B x) C)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   441
  by transfer
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   442
     (auto intro!: bind_assoc[where N="count_space UNIV" and R="count_space UNIV"]
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   443
           simp: measurable_def space_subprob_algebra prob_space_imp_subprob_space)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   444
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   445
definition "map_pmf f M = bind_pmf M (\<lambda>x. return_pmf (f x))"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   446
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   447
lemma map_bind_pmf: "map_pmf f (bind_pmf M g) = bind_pmf M (\<lambda>x. map_pmf f (g x))"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   448
  by (simp add: map_pmf_def bind_assoc_pmf)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   449
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   450
lemma bind_map_pmf: "bind_pmf (map_pmf f M) g = bind_pmf M (\<lambda>x. g (f x))"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   451
  by (simp add: map_pmf_def bind_assoc_pmf bind_return_pmf)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   452
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   453
lemma map_pmf_transfer[transfer_rule]:
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67226
diff changeset
   454
  "rel_fun (=) (rel_fun cr_pmf cr_pmf) (\<lambda>f M. distr M (count_space UNIV) f) map_pmf"
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   455
proof -
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67226
diff changeset
   456
  have "rel_fun (=) (rel_fun pmf_as_measure.cr_pmf pmf_as_measure.cr_pmf)
62026
ea3b1b0413b4 more symbols;
wenzelm
parents: 61808
diff changeset
   457
     (\<lambda>f M. M \<bind> (return (count_space UNIV) o f)) map_pmf"
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59665
diff changeset
   458
    unfolding map_pmf_def[abs_def] comp_def by transfer_prover
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   459
  then show ?thesis
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   460
    by (force simp: rel_fun_def cr_pmf_def bind_return_distr)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   461
qed
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   462
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   463
lemma map_pmf_rep_eq:
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   464
  "measure_pmf (map_pmf f M) = distr (measure_pmf M) (count_space UNIV) f"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   465
  unfolding map_pmf_def bind_pmf.rep_eq comp_def return_pmf.rep_eq
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   466
  using bind_return_distr[of M f "count_space UNIV"] by (simp add: comp_def)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   467
58587
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   468
lemma map_pmf_id[simp]: "map_pmf id = id"
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   469
  by (rule, transfer) (auto simp: emeasure_distr measurable_def intro!: measure_eqI)
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   470
59053
43e07797269b tuned proof that pmfs are bnfs
hoelzl
parents: 59052
diff changeset
   471
lemma map_pmf_ident[simp]: "map_pmf (\<lambda>x. x) = (\<lambda>x. x)"
43e07797269b tuned proof that pmfs are bnfs
hoelzl
parents: 59052
diff changeset
   472
  using map_pmf_id unfolding id_def .
43e07797269b tuned proof that pmfs are bnfs
hoelzl
parents: 59052
diff changeset
   473
58587
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   474
lemma map_pmf_compose: "map_pmf (f \<circ> g) = map_pmf f \<circ> map_pmf g"
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59665
diff changeset
   475
  by (rule, transfer) (simp add: distr_distr[symmetric, where N="count_space UNIV"] measurable_def)
58587
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   476
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   477
lemma map_pmf_comp: "map_pmf f (map_pmf g M) = map_pmf (\<lambda>x. f (g x)) M"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   478
  using map_pmf_compose[of f g] by (simp add: comp_def)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   479
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   480
lemma map_pmf_cong: "p = q \<Longrightarrow> (\<And>x. x \<in> set_pmf q \<Longrightarrow> f x = g x) \<Longrightarrow> map_pmf f p = map_pmf g q"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   481
  unfolding map_pmf_def by (rule bind_pmf_cong) auto
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   482
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67226
diff changeset
   483
lemma pmf_set_map: "set_pmf \<circ> map_pmf f = (`) f \<circ> set_pmf"
59665
37adca7fd48f add set_pmf lemmas to simpset
hoelzl
parents: 59664
diff changeset
   484
  by (auto simp add: comp_def fun_eq_iff map_pmf_def)
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   485
59665
37adca7fd48f add set_pmf lemmas to simpset
hoelzl
parents: 59664
diff changeset
   486
lemma set_map_pmf[simp]: "set_pmf (map_pmf f M) = f`set_pmf M"
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   487
  using pmf_set_map[of f] by (auto simp: comp_def fun_eq_iff)
58587
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   488
59002
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   489
lemma emeasure_map_pmf[simp]: "emeasure (map_pmf f M) X = emeasure M (f -` X)"
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   490
  unfolding map_pmf_rep_eq by (subst emeasure_distr) auto
59002
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   491
61634
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
   492
lemma measure_map_pmf[simp]: "measure (map_pmf f M) X = measure M (f -` X)"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   493
using emeasure_map_pmf[of f M X] by(simp add: measure_pmf.emeasure_eq_measure measure_nonneg)
61634
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
   494
59002
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   495
lemma nn_integral_map_pmf[simp]: "(\<integral>\<^sup>+x. f x \<partial>map_pmf g M) = (\<integral>\<^sup>+x. f (g x) \<partial>M)"
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   496
  unfolding map_pmf_rep_eq by (intro nn_integral_distr) auto
59002
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   497
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   498
lemma ennreal_pmf_map: "pmf (map_pmf f p) x = (\<integral>\<^sup>+ y. indicator (f -` {x}) y \<partial>measure_pmf p)"
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   499
proof (transfer fixing: f x)
59023
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
   500
  fix p :: "'b measure"
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
   501
  presume "prob_space p"
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
   502
  then interpret prob_space p .
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
   503
  presume "sets p = UNIV"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   504
  then show "ennreal (measure (distr p (count_space UNIV) f) {x}) = integral\<^sup>N p (indicator (f -` {x}))"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   505
    by(simp add: measure_distr measurable_def emeasure_eq_measure)
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   506
qed simp_all
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   507
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   508
lemma pmf_map: "pmf (map_pmf f p) x = measure p (f -` {x})"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   509
proof (transfer fixing: f x)
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   510
  fix p :: "'b measure"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   511
  presume "prob_space p"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   512
  then interpret prob_space p .
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   513
  presume "sets p = UNIV"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   514
  then show "measure (distr p (count_space UNIV) f) {x} = measure p (f -` {x})"
59023
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
   515
    by(simp add: measure_distr measurable_def emeasure_eq_measure)
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
   516
qed simp_all
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
   517
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
   518
lemma nn_integral_pmf: "(\<integral>\<^sup>+ x. pmf p x \<partial>count_space A) = emeasure (measure_pmf p) A"
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
   519
proof -
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
   520
  have "(\<integral>\<^sup>+ x. pmf p x \<partial>count_space A) = (\<integral>\<^sup>+ x. pmf p x \<partial>count_space (A \<inter> set_pmf p))"
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
   521
    by(auto simp add: nn_integral_count_space_indicator indicator_def set_pmf_iff intro: nn_integral_cong)
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
   522
  also have "\<dots> = emeasure (measure_pmf p) (\<Union>x\<in>A \<inter> set_pmf p. {x})"
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
   523
    by(subst emeasure_UN_countable)(auto simp add: emeasure_pmf_single disjoint_family_on_def)
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
   524
  also have "\<dots> = emeasure (measure_pmf p) ((\<Union>x\<in>A \<inter> set_pmf p. {x}) \<union> {x. x \<in> A \<and> x \<notin> set_pmf p})"
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
   525
    by(rule emeasure_Un_null_set[symmetric])(auto intro: in_null_sets_measure_pmfI)
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
   526
  also have "\<dots> = emeasure (measure_pmf p) A"
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
   527
    by(auto intro: arg_cong2[where f=emeasure])
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
   528
  finally show ?thesis .
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
   529
qed
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
   530
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   531
lemma integral_map_pmf[simp]:
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   532
  fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   533
  shows "integral\<^sup>L (map_pmf g p) f = integral\<^sup>L p (\<lambda>x. f (g x))"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   534
  by (simp add: integral_distr map_pmf_rep_eq)
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   535
73253
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
   536
lemma integrable_map_pmf_eq [simp]:
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
   537
  fixes g :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
75625
0dd3ac5fdbaa tuned BNF bounds for function space and bounded sets; NEWS and CONTRIBUTORS
traytel
parents: 75624
diff changeset
   538
  shows "integrable (map_pmf f p) g \<longleftrightarrow> integrable (measure_pmf p) (\<lambda>x. g (f x))"
73253
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
   539
  by (subst map_pmf_rep_eq, subst integrable_distr_eq) auto
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
   540
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
   541
lemma integrable_map_pmf [intro]:
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
   542
  fixes g :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
   543
  shows "integrable (measure_pmf p) (\<lambda>x. g (f x)) \<Longrightarrow> integrable (map_pmf f p) g"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
   544
  by (subst integrable_map_pmf_eq)
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
   545
66568
52b5cf533fd6 Connecting PMFs to infinite sums
eberlm <eberlm@in.tum.de>
parents: 66453
diff changeset
   546
lemma pmf_abs_summable [intro]: "pmf p abs_summable_on A"
67486
Lars Hupel <lars.hupel@mytum.de>
parents: 67399
diff changeset
   547
  by (rule abs_summable_on_subset[OF _ subset_UNIV])
66568
52b5cf533fd6 Connecting PMFs to infinite sums
eberlm <eberlm@in.tum.de>
parents: 66453
diff changeset
   548
     (auto simp:  abs_summable_on_def integrable_iff_bounded nn_integral_pmf)
52b5cf533fd6 Connecting PMFs to infinite sums
eberlm <eberlm@in.tum.de>
parents: 66453
diff changeset
   549
52b5cf533fd6 Connecting PMFs to infinite sums
eberlm <eberlm@in.tum.de>
parents: 66453
diff changeset
   550
lemma measure_pmf_conv_infsetsum: "measure (measure_pmf p) A = infsetsum (pmf p) A"
67486
Lars Hupel <lars.hupel@mytum.de>
parents: 67399
diff changeset
   551
  unfolding infsetsum_def by (simp add: integral_eq_nn_integral nn_integral_pmf measure_def)
66568
52b5cf533fd6 Connecting PMFs to infinite sums
eberlm <eberlm@in.tum.de>
parents: 66453
diff changeset
   552
52b5cf533fd6 Connecting PMFs to infinite sums
eberlm <eberlm@in.tum.de>
parents: 66453
diff changeset
   553
lemma infsetsum_pmf_eq_1:
52b5cf533fd6 Connecting PMFs to infinite sums
eberlm <eberlm@in.tum.de>
parents: 66453
diff changeset
   554
  assumes "set_pmf p \<subseteq> A"
52b5cf533fd6 Connecting PMFs to infinite sums
eberlm <eberlm@in.tum.de>
parents: 66453
diff changeset
   555
  shows   "infsetsum (pmf p) A = 1"
52b5cf533fd6 Connecting PMFs to infinite sums
eberlm <eberlm@in.tum.de>
parents: 66453
diff changeset
   556
proof -
52b5cf533fd6 Connecting PMFs to infinite sums
eberlm <eberlm@in.tum.de>
parents: 66453
diff changeset
   557
  have "infsetsum (pmf p) A = lebesgue_integral (count_space UNIV) (pmf p)"
67977
557ea2740125 Probability builds with new definitions
paulson <lp15@cam.ac.uk>
parents: 67601
diff changeset
   558
    using assms unfolding infsetsum_altdef set_lebesgue_integral_def
66568
52b5cf533fd6 Connecting PMFs to infinite sums
eberlm <eberlm@in.tum.de>
parents: 66453
diff changeset
   559
    by (intro Bochner_Integration.integral_cong) (auto simp: indicator_def set_pmf_eq)
52b5cf533fd6 Connecting PMFs to infinite sums
eberlm <eberlm@in.tum.de>
parents: 66453
diff changeset
   560
  also have "\<dots> = 1"
52b5cf533fd6 Connecting PMFs to infinite sums
eberlm <eberlm@in.tum.de>
parents: 66453
diff changeset
   561
    by (subst integral_eq_nn_integral) (auto simp: nn_integral_pmf)
52b5cf533fd6 Connecting PMFs to infinite sums
eberlm <eberlm@in.tum.de>
parents: 66453
diff changeset
   562
  finally show ?thesis .
52b5cf533fd6 Connecting PMFs to infinite sums
eberlm <eberlm@in.tum.de>
parents: 66453
diff changeset
   563
qed
52b5cf533fd6 Connecting PMFs to infinite sums
eberlm <eberlm@in.tum.de>
parents: 66453
diff changeset
   564
60068
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
   565
lemma map_return_pmf [simp]: "map_pmf f (return_pmf x) = return_pmf (f x)"
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   566
  by transfer (simp add: distr_return)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   567
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   568
lemma map_pmf_const[simp]: "map_pmf (\<lambda>_. c) M = return_pmf c"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   569
  by transfer (auto simp: prob_space.distr_const)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   570
60068
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
   571
lemma pmf_return [simp]: "pmf (return_pmf x) y = indicator {y} x"
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   572
  by transfer (simp add: measure_return)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   573
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   574
lemma nn_integral_return_pmf[simp]: "0 \<le> f x \<Longrightarrow> (\<integral>\<^sup>+x. f x \<partial>return_pmf x) = f x"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   575
  unfolding return_pmf.rep_eq by (intro nn_integral_return) auto
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   576
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   577
lemma emeasure_return_pmf[simp]: "emeasure (return_pmf x) X = indicator X x"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   578
  unfolding return_pmf.rep_eq by (intro emeasure_return) auto
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   579
63099
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
   580
lemma measure_return_pmf [simp]: "measure_pmf.prob (return_pmf x) A = indicator A x"
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
   581
proof -
63886
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63882
diff changeset
   582
  have "ennreal (measure_pmf.prob (return_pmf x) A) =
63099
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
   583
          emeasure (measure_pmf (return_pmf x)) A"
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
   584
    by (simp add: measure_pmf.emeasure_eq_measure)
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
   585
  also have "\<dots> = ennreal (indicator A x)" by (simp add: ennreal_indicator)
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
   586
  finally show ?thesis by simp
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
   587
qed
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
   588
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   589
lemma return_pmf_inj[simp]: "return_pmf x = return_pmf y \<longleftrightarrow> x = y"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   590
  by (metis insertI1 set_return_pmf singletonD)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   591
59665
37adca7fd48f add set_pmf lemmas to simpset
hoelzl
parents: 59664
diff changeset
   592
lemma map_pmf_eq_return_pmf_iff:
37adca7fd48f add set_pmf lemmas to simpset
hoelzl
parents: 59664
diff changeset
   593
  "map_pmf f p = return_pmf x \<longleftrightarrow> (\<forall>y \<in> set_pmf p. f y = x)"
37adca7fd48f add set_pmf lemmas to simpset
hoelzl
parents: 59664
diff changeset
   594
proof
37adca7fd48f add set_pmf lemmas to simpset
hoelzl
parents: 59664
diff changeset
   595
  assume "map_pmf f p = return_pmf x"
37adca7fd48f add set_pmf lemmas to simpset
hoelzl
parents: 59664
diff changeset
   596
  then have "set_pmf (map_pmf f p) = set_pmf (return_pmf x)" by simp
37adca7fd48f add set_pmf lemmas to simpset
hoelzl
parents: 59664
diff changeset
   597
  then show "\<forall>y \<in> set_pmf p. f y = x" by auto
37adca7fd48f add set_pmf lemmas to simpset
hoelzl
parents: 59664
diff changeset
   598
next
37adca7fd48f add set_pmf lemmas to simpset
hoelzl
parents: 59664
diff changeset
   599
  assume "\<forall>y \<in> set_pmf p. f y = x"
37adca7fd48f add set_pmf lemmas to simpset
hoelzl
parents: 59664
diff changeset
   600
  then show "map_pmf f p = return_pmf x"
37adca7fd48f add set_pmf lemmas to simpset
hoelzl
parents: 59664
diff changeset
   601
    unfolding map_pmf_const[symmetric, of _ p] by (intro map_pmf_cong) auto
37adca7fd48f add set_pmf lemmas to simpset
hoelzl
parents: 59664
diff changeset
   602
qed
37adca7fd48f add set_pmf lemmas to simpset
hoelzl
parents: 59664
diff changeset
   603
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   604
definition "pair_pmf A B = bind_pmf A (\<lambda>x. bind_pmf B (\<lambda>y. return_pmf (x, y)))"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   605
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   606
lemma pmf_pair: "pmf (pair_pmf M N) (a, b) = pmf M a * pmf N b"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   607
  unfolding pair_pmf_def pmf_bind pmf_return
64008
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   608
  apply (subst integral_measure_pmf_real[where A="{b}"])
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   609
  apply (auto simp: indicator_eq_0_iff)
64008
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   610
  apply (subst integral_measure_pmf_real[where A="{a}"])
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64008
diff changeset
   611
  apply (auto simp: indicator_eq_0_iff sum_nonneg_eq_0_iff pmf_nonneg)
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   612
  done
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   613
59665
37adca7fd48f add set_pmf lemmas to simpset
hoelzl
parents: 59664
diff changeset
   614
lemma set_pair_pmf[simp]: "set_pmf (pair_pmf A B) = set_pmf A \<times> set_pmf B"
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   615
  unfolding pair_pmf_def set_bind_pmf set_return_pmf by auto
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   616
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   617
lemma measure_pmf_in_subprob_space[measurable (raw)]:
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   618
  "measure_pmf M \<in> space (subprob_algebra (count_space UNIV))"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   619
  by (simp add: space_subprob_algebra) intro_locales
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   620
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   621
lemma nn_integral_pair_pmf': "(\<integral>\<^sup>+x. f x \<partial>pair_pmf A B) = (\<integral>\<^sup>+a. \<integral>\<^sup>+b. f (a, b) \<partial>B \<partial>A)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   622
proof -
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   623
  have "(\<integral>\<^sup>+x. f x \<partial>pair_pmf A B) = (\<integral>\<^sup>+x. f x * indicator (A \<times> B) x \<partial>pair_pmf A B)"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   624
    by (auto simp: AE_measure_pmf_iff intro!: nn_integral_cong_AE)
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   625
  also have "\<dots> = (\<integral>\<^sup>+a. \<integral>\<^sup>+b. f (a, b) * indicator (A \<times> B) (a, b) \<partial>B \<partial>A)"
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   626
    by (simp add: pair_pmf_def)
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   627
  also have "\<dots> = (\<integral>\<^sup>+a. \<integral>\<^sup>+b. f (a, b) \<partial>B \<partial>A)"
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   628
    by (auto intro!: nn_integral_cong_AE simp: AE_measure_pmf_iff)
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   629
  finally show ?thesis .
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   630
qed
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   631
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   632
lemma bind_pair_pmf:
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   633
  assumes M[measurable]: "M \<in> measurable (count_space UNIV \<Otimes>\<^sub>M count_space UNIV) (subprob_algebra N)"
62026
ea3b1b0413b4 more symbols;
wenzelm
parents: 61808
diff changeset
   634
  shows "measure_pmf (pair_pmf A B) \<bind> M = (measure_pmf A \<bind> (\<lambda>x. measure_pmf B \<bind> (\<lambda>y. M (x, y))))"
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   635
    (is "?L = ?R")
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   636
proof (rule measure_eqI)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   637
  have M'[measurable]: "M \<in> measurable (pair_pmf A B) (subprob_algebra N)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   638
    using M[THEN measurable_space] by (simp_all add: space_pair_measure)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   639
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   640
  note measurable_bind[where N="count_space UNIV", measurable]
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   641
  note measure_pmf_in_subprob_space[simp]
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   642
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   643
  have sets_eq_N: "sets ?L = N"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   644
    by (subst sets_bind[OF sets_kernel[OF M']]) auto
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   645
  show "sets ?L = sets ?R"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   646
    using measurable_space[OF M]
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   647
    by (simp add: sets_eq_N space_pair_measure space_subprob_algebra)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   648
  fix X assume "X \<in> sets ?L"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   649
  then have X[measurable]: "X \<in> sets N"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   650
    unfolding sets_eq_N .
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   651
  then show "emeasure ?L X = emeasure ?R X"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   652
    apply (simp add: emeasure_bind[OF _ M' X])
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   653
    apply (simp add: nn_integral_bind[where B="count_space UNIV"] pair_pmf_def measure_pmf_bind[of A]
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   654
                     nn_integral_measure_pmf_finite)
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   655
    apply (subst emeasure_bind[OF _ _ X])
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   656
    apply measurable
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   657
    apply (subst emeasure_bind[OF _ _ X])
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   658
    apply measurable
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   659
    done
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   660
qed
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   661
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   662
lemma map_fst_pair_pmf: "map_pmf fst (pair_pmf A B) = A"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   663
  by (simp add: pair_pmf_def map_pmf_def bind_assoc_pmf bind_return_pmf bind_return_pmf')
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   665
lemma map_snd_pair_pmf: "map_pmf snd (pair_pmf A B) = B"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   666
  by (simp add: pair_pmf_def map_pmf_def bind_assoc_pmf bind_return_pmf bind_return_pmf')
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   667
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   668
lemma nn_integral_pmf':
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   669
  "inj_on f A \<Longrightarrow> (\<integral>\<^sup>+x. pmf p (f x) \<partial>count_space A) = emeasure p (f ` A)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   670
  by (subst nn_integral_bij_count_space[where g=f and B="f`A"])
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   671
     (auto simp: bij_betw_def nn_integral_pmf)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   672
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   673
lemma pmf_le_0_iff[simp]: "pmf M p \<le> 0 \<longleftrightarrow> pmf M p = 0"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   674
  using pmf_nonneg[of M p] by arith
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   675
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   676
lemma min_pmf_0[simp]: "min (pmf M p) 0 = 0" "min 0 (pmf M p) = 0"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   677
  using pmf_nonneg[of M p] by arith+
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   678
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   679
lemma pmf_eq_0_set_pmf: "pmf M p = 0 \<longleftrightarrow> p \<notin> set_pmf M"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   680
  unfolding set_pmf_iff by simp
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   681
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   682
lemma pmf_map_inj: "inj_on f (set_pmf M) \<Longrightarrow> x \<in> set_pmf M \<Longrightarrow> pmf (map_pmf f M) (f x) = pmf M x"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   683
  by (auto simp: pmf.rep_eq map_pmf_rep_eq measure_distr AE_measure_pmf_iff inj_onD
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   684
           intro!: measure_pmf.finite_measure_eq_AE)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   685
73253
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
   686
lemma pair_return_pmf [simp]: "pair_pmf (return_pmf x) (return_pmf y) = return_pmf (x, y)"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
   687
  by (auto simp: pair_pmf_def bind_return_pmf)
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
   688
60068
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
   689
lemma pmf_map_inj': "inj f \<Longrightarrow> pmf (map_pmf f M) (f x) = pmf M x"
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
   690
apply(cases "x \<in> set_pmf M")
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
   691
 apply(simp add: pmf_map_inj[OF subset_inj_on])
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
   692
apply(simp add: pmf_eq_0_set_pmf[symmetric])
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
   693
apply(auto simp add: pmf_eq_0_set_pmf dest: injD)
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
   694
done
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
   695
73253
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
   696
lemma expectation_pair_pmf_fst [simp]:
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
   697
  fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
   698
  shows "measure_pmf.expectation (pair_pmf p q) (\<lambda>x. f (fst x)) = measure_pmf.expectation p f"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
   699
proof -
75625
0dd3ac5fdbaa tuned BNF bounds for function space and bounded sets; NEWS and CONTRIBUTORS
traytel
parents: 75624
diff changeset
   700
  have "measure_pmf.expectation (pair_pmf p q) (\<lambda>x. f (fst x)) =
73253
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
   701
          measure_pmf.expectation (map_pmf fst (pair_pmf p q)) f" by simp
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
   702
  also have "map_pmf fst (pair_pmf p q) = p"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
   703
    by (simp add: map_fst_pair_pmf)
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
   704
  finally show ?thesis .
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
   705
qed
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
   706
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
   707
lemma expectation_pair_pmf_snd [simp]:
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
   708
  fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
   709
  shows "measure_pmf.expectation (pair_pmf p q) (\<lambda>x. f (snd x)) = measure_pmf.expectation q f"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
   710
proof -
75625
0dd3ac5fdbaa tuned BNF bounds for function space and bounded sets; NEWS and CONTRIBUTORS
traytel
parents: 75624
diff changeset
   711
  have "measure_pmf.expectation (pair_pmf p q) (\<lambda>x. f (snd x)) =
73253
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
   712
          measure_pmf.expectation (map_pmf snd (pair_pmf p q)) f" by simp
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
   713
  also have "map_pmf snd (pair_pmf p q) = q"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
   714
    by (simp add: map_snd_pair_pmf)
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
   715
  finally show ?thesis .
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
   716
qed
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
   717
60068
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
   718
lemma pmf_map_outside: "x \<notin> f ` set_pmf M \<Longrightarrow> pmf (map_pmf f M) x = 0"
64008
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   719
  unfolding pmf_eq_0_set_pmf by simp
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   720
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   721
lemma measurable_set_pmf[measurable]: "Measurable.pred (count_space UNIV) (\<lambda>x. x \<in> set_pmf M)"
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   722
  by simp
60068
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
   723
65395
7504569a73c7 moved material from AFP to distribution
eberlm <eberlm@in.tum.de>
parents: 64634
diff changeset
   724
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   725
subsection \<open> PMFs as function \<close>
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   726
58587
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   727
context
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   728
  fixes f :: "'a \<Rightarrow> real"
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   729
  assumes nonneg: "\<And>x. 0 \<le> f x"
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   730
  assumes prob: "(\<integral>\<^sup>+x. f x \<partial>count_space UNIV) = 1"
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   731
begin
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   732
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   733
lift_definition embed_pmf :: "'a pmf" is "density (count_space UNIV) (ennreal \<circ> f)"
58587
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   734
proof (intro conjI)
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   735
  have *[simp]: "\<And>x y. ennreal (f y) * indicator {x} y = ennreal (f x) * indicator {x} y"
58587
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   736
    by (simp split: split_indicator)
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   737
  show "AE x in density (count_space UNIV) (ennreal \<circ> f).
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   738
    measure (density (count_space UNIV) (ennreal \<circ> f)) {x} \<noteq> 0"
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents: 59053
diff changeset
   739
    by (simp add: AE_density nonneg measure_def emeasure_density max_def)
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   740
  show "prob_space (density (count_space UNIV) (ennreal \<circ> f))"
61169
4de9ff3ea29a tuned proofs -- less legacy;
wenzelm
parents: 60602
diff changeset
   741
    by standard (simp add: emeasure_density prob)
58587
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   742
qed simp
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   743
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   744
lemma pmf_embed_pmf: "pmf embed_pmf x = f x"
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   745
proof transfer
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   746
  have *[simp]: "\<And>x y. ennreal (f y) * indicator {x} y = ennreal (f x) * indicator {x} y"
58587
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   747
    by (simp split: split_indicator)
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   748
  fix x show "measure (density (count_space UNIV) (ennreal \<circ> f)) {x} = f x"
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents: 59053
diff changeset
   749
    by transfer (simp add: measure_def emeasure_density nonneg max_def)
58587
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   750
qed
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   751
60068
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
   752
lemma set_embed_pmf: "set_pmf embed_pmf = {x. f x \<noteq> 0}"
63092
a949b2a5f51d eliminated use of empty "assms";
wenzelm
parents: 63040
diff changeset
   753
by(auto simp add: set_pmf_eq pmf_embed_pmf)
60068
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
   754
58587
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   755
end
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   756
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   757
lemma embed_pmf_transfer:
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   758
  "rel_fun (eq_onp (\<lambda>f. (\<forall>x. 0 \<le> f x) \<and> (\<integral>\<^sup>+x. ennreal (f x) \<partial>count_space UNIV) = 1)) pmf_as_measure.cr_pmf (\<lambda>f. density (count_space UNIV) (ennreal \<circ> f)) embed_pmf"
58587
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   759
  by (auto simp: rel_fun_def eq_onp_def embed_pmf.transfer)
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   760
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   761
lemma measure_pmf_eq_density: "measure_pmf p = density (count_space UNIV) (pmf p)"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   762
proof (transfer, elim conjE)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   763
  fix M :: "'a measure" assume [simp]: "sets M = UNIV" and ae: "AE x in M. measure M {x} \<noteq> 0"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   764
  assume "prob_space M" then interpret prob_space M .
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   765
  show "M = density (count_space UNIV) (\<lambda>x. ennreal (measure M {x}))"
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   766
  proof (rule measure_eqI)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   767
    fix A :: "'a set"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   768
    have "(\<integral>\<^sup>+ x. ennreal (measure M {x}) * indicator A x \<partial>count_space UNIV) =
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   769
      (\<integral>\<^sup>+ x. emeasure M {x} * indicator (A \<inter> {x. measure M {x} \<noteq> 0}) x \<partial>count_space UNIV)"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   770
      by (auto intro!: nn_integral_cong simp: emeasure_eq_measure split: split_indicator)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   771
    also have "\<dots> = (\<integral>\<^sup>+ x. emeasure M {x} \<partial>count_space (A \<inter> {x. measure M {x} \<noteq> 0}))"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   772
      by (subst nn_integral_restrict_space[symmetric]) (auto simp: restrict_count_space)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   773
    also have "\<dots> = emeasure M (\<Union>x\<in>(A \<inter> {x. measure M {x} \<noteq> 0}). {x})"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   774
      by (intro emeasure_UN_countable[symmetric] countable_Int2 countable_support)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   775
         (auto simp: disjoint_family_on_def)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   776
    also have "\<dots> = emeasure M A"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   777
      using ae by (intro emeasure_eq_AE) auto
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   778
    finally show " emeasure M A = emeasure (density (count_space UNIV) (\<lambda>x. ennreal (measure M {x}))) A"
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   779
      using emeasure_space_1 by (simp add: emeasure_density)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   780
  qed simp
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   781
qed
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   782
58587
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   783
lemma td_pmf_embed_pmf:
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   784
  "type_definition pmf embed_pmf {f::'a \<Rightarrow> real. (\<forall>x. 0 \<le> f x) \<and> (\<integral>\<^sup>+x. ennreal (f x) \<partial>count_space UNIV) = 1}"
58587
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   785
  unfolding type_definition_def
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   786
proof safe
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   787
  fix p :: "'a pmf"
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   788
  have "(\<integral>\<^sup>+ x. 1 \<partial>measure_pmf p) = 1"
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   789
    using measure_pmf.emeasure_space_1[of p] by simp
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   790
  then show *: "(\<integral>\<^sup>+ x. ennreal (pmf p x) \<partial>count_space UNIV) = 1"
58587
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   791
    by (simp add: measure_pmf_eq_density nn_integral_density pmf_nonneg del: nn_integral_const)
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   792
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   793
  show "embed_pmf (pmf p) = p"
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   794
    by (intro measure_pmf_inject[THEN iffD1])
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   795
       (simp add: * embed_pmf.rep_eq pmf_nonneg measure_pmf_eq_density[of p] comp_def)
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   796
next
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   797
  fix f :: "'a \<Rightarrow> real" assume "\<forall>x. 0 \<le> f x" "(\<integral>\<^sup>+x. f x \<partial>count_space UNIV) = 1"
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   798
  then show "pmf (embed_pmf f) = f"
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   799
    by (auto intro!: pmf_embed_pmf)
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   800
qed (rule pmf_nonneg)
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   801
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   802
end
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   803
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   804
lemma nn_integral_measure_pmf: "(\<integral>\<^sup>+ x. f x \<partial>measure_pmf p) = \<integral>\<^sup>+ x. ennreal (pmf p x) * f x \<partial>count_space UNIV"
60068
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
   805
by(simp add: measure_pmf_eq_density nn_integral_density pmf_nonneg)
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
   806
64008
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   807
lemma integral_measure_pmf:
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   808
  fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   809
  assumes A: "finite A"
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   810
  shows "(\<And>a. a \<in> set_pmf M \<Longrightarrow> f a \<noteq> 0 \<Longrightarrow> a \<in> A) \<Longrightarrow> (LINT x|M. f x) = (\<Sum>a\<in>A. pmf M a *\<^sub>R f a)"
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   811
  unfolding measure_pmf_eq_density
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   812
  apply (simp add: integral_density)
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   813
  apply (subst lebesgue_integral_count_space_finite_support)
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64008
diff changeset
   814
  apply (auto intro!: finite_subset[OF _ \<open>finite A\<close>] sum.mono_neutral_left simp: pmf_eq_0_set_pmf)
64008
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   815
  done
67486
Lars Hupel <lars.hupel@mytum.de>
parents: 67399
diff changeset
   816
65395
7504569a73c7 moved material from AFP to distribution
eberlm <eberlm@in.tum.de>
parents: 64634
diff changeset
   817
lemma expectation_return_pmf [simp]:
7504569a73c7 moved material from AFP to distribution
eberlm <eberlm@in.tum.de>
parents: 64634
diff changeset
   818
  fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
7504569a73c7 moved material from AFP to distribution
eberlm <eberlm@in.tum.de>
parents: 64634
diff changeset
   819
  shows "measure_pmf.expectation (return_pmf x) f = f x"
7504569a73c7 moved material from AFP to distribution
eberlm <eberlm@in.tum.de>
parents: 64634
diff changeset
   820
  by (subst integral_measure_pmf[of "{x}"]) simp_all
7504569a73c7 moved material from AFP to distribution
eberlm <eberlm@in.tum.de>
parents: 64634
diff changeset
   821
7504569a73c7 moved material from AFP to distribution
eberlm <eberlm@in.tum.de>
parents: 64634
diff changeset
   822
lemma pmf_expectation_bind:
7504569a73c7 moved material from AFP to distribution
eberlm <eberlm@in.tum.de>
parents: 64634
diff changeset
   823
  fixes p :: "'a pmf" and f :: "'a \<Rightarrow> 'b pmf"
7504569a73c7 moved material from AFP to distribution
eberlm <eberlm@in.tum.de>
parents: 64634
diff changeset
   824
    and  h :: "'b \<Rightarrow> 'c::{banach, second_countable_topology}"
7504569a73c7 moved material from AFP to distribution
eberlm <eberlm@in.tum.de>
parents: 64634
diff changeset
   825
  assumes "finite A" "\<And>x. x \<in> A \<Longrightarrow> finite (set_pmf (f x))" "set_pmf p \<subseteq> A"
7504569a73c7 moved material from AFP to distribution
eberlm <eberlm@in.tum.de>
parents: 64634
diff changeset
   826
  shows "measure_pmf.expectation (p \<bind> f) h =
7504569a73c7 moved material from AFP to distribution
eberlm <eberlm@in.tum.de>
parents: 64634
diff changeset
   827
           (\<Sum>a\<in>A. pmf p a *\<^sub>R measure_pmf.expectation (f a) h)"
7504569a73c7 moved material from AFP to distribution
eberlm <eberlm@in.tum.de>
parents: 64634
diff changeset
   828
proof -
7504569a73c7 moved material from AFP to distribution
eberlm <eberlm@in.tum.de>
parents: 64634
diff changeset
   829
  have "measure_pmf.expectation (p \<bind> f) h = (\<Sum>a\<in>(\<Union>x\<in>A. set_pmf (f x)). pmf (p \<bind> f) a *\<^sub>R h a)"
7504569a73c7 moved material from AFP to distribution
eberlm <eberlm@in.tum.de>
parents: 64634
diff changeset
   830
    using assms by (intro integral_measure_pmf) auto
7504569a73c7 moved material from AFP to distribution
eberlm <eberlm@in.tum.de>
parents: 64634
diff changeset
   831
  also have "\<dots> = (\<Sum>x\<in>(\<Union>x\<in>A. set_pmf (f x)). (\<Sum>a\<in>A. (pmf p a * pmf (f a) x) *\<^sub>R h x))"
7504569a73c7 moved material from AFP to distribution
eberlm <eberlm@in.tum.de>
parents: 64634
diff changeset
   832
  proof (intro sum.cong refl, goal_cases)
7504569a73c7 moved material from AFP to distribution
eberlm <eberlm@in.tum.de>
parents: 64634
diff changeset
   833
    case (1 x)
7504569a73c7 moved material from AFP to distribution
eberlm <eberlm@in.tum.de>
parents: 64634
diff changeset
   834
    thus ?case
67486
Lars Hupel <lars.hupel@mytum.de>
parents: 67399
diff changeset
   835
      by (subst pmf_bind, subst integral_measure_pmf[of A])
65395
7504569a73c7 moved material from AFP to distribution
eberlm <eberlm@in.tum.de>
parents: 64634
diff changeset
   836
         (insert assms, auto simp: scaleR_sum_left)
7504569a73c7 moved material from AFP to distribution
eberlm <eberlm@in.tum.de>
parents: 64634
diff changeset
   837
  qed
7504569a73c7 moved material from AFP to distribution
eberlm <eberlm@in.tum.de>
parents: 64634
diff changeset
   838
  also have "\<dots> = (\<Sum>j\<in>A. pmf p j *\<^sub>R (\<Sum>i\<in>(\<Union>x\<in>A. set_pmf (f x)). pmf (f j) i *\<^sub>R h i))"
66804
3f9bb52082c4 avoid name clashes on interpretation of abstract locales
haftmann
parents: 66568
diff changeset
   839
    by (subst sum.swap) (simp add: scaleR_sum_right)
65395
7504569a73c7 moved material from AFP to distribution
eberlm <eberlm@in.tum.de>
parents: 64634
diff changeset
   840
  also have "\<dots> = (\<Sum>j\<in>A. pmf p j *\<^sub>R measure_pmf.expectation (f j) h)"
7504569a73c7 moved material from AFP to distribution
eberlm <eberlm@in.tum.de>
parents: 64634
diff changeset
   841
  proof (intro sum.cong refl, goal_cases)
7504569a73c7 moved material from AFP to distribution
eberlm <eberlm@in.tum.de>
parents: 64634
diff changeset
   842
    case (1 x)
7504569a73c7 moved material from AFP to distribution
eberlm <eberlm@in.tum.de>
parents: 64634
diff changeset
   843
    thus ?case
67486
Lars Hupel <lars.hupel@mytum.de>
parents: 67399
diff changeset
   844
      by (subst integral_measure_pmf[of "(\<Union>x\<in>A. set_pmf (f x))"])
65395
7504569a73c7 moved material from AFP to distribution
eberlm <eberlm@in.tum.de>
parents: 64634
diff changeset
   845
         (insert assms, auto simp: scaleR_sum_left)
7504569a73c7 moved material from AFP to distribution
eberlm <eberlm@in.tum.de>
parents: 64634
diff changeset
   846
  qed
7504569a73c7 moved material from AFP to distribution
eberlm <eberlm@in.tum.de>
parents: 64634
diff changeset
   847
  finally show ?thesis .
7504569a73c7 moved material from AFP to distribution
eberlm <eberlm@in.tum.de>
parents: 64634
diff changeset
   848
qed
64008
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   849
67226
ec32cdaab97b isabelle update_cartouches -c -t;
wenzelm
parents: 66804
diff changeset
   850
lemma continuous_on_LINT_pmf: \<comment> \<open>This is dominated convergence!?\<close>
64008
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   851
  fixes f :: "'i \<Rightarrow> 'a::topological_space \<Rightarrow> 'b::{banach, second_countable_topology}"
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   852
  assumes f: "\<And>i. i \<in> set_pmf M \<Longrightarrow> continuous_on A (f i)"
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   853
    and bnd: "\<And>a i. a \<in> A \<Longrightarrow> i \<in> set_pmf M \<Longrightarrow> norm (f i a) \<le> B"
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   854
  shows "continuous_on A (\<lambda>a. LINT i|M. f i a)"
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   855
proof cases
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   856
  assume "finite M" with f show ?thesis
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   857
    using integral_measure_pmf[OF \<open>finite M\<close>]
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   858
    by (subst integral_measure_pmf[OF \<open>finite M\<close>])
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64008
diff changeset
   859
       (auto intro!: continuous_on_sum continuous_on_scaleR continuous_on_const)
64008
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   860
next
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   861
  assume "infinite M"
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   862
  let ?f = "\<lambda>i x. pmf (map_pmf (to_nat_on M) M) i *\<^sub>R f (from_nat_into M i) x"
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   863
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   864
  show ?thesis
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   865
  proof (rule uniform_limit_theorem)
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   866
    show "\<forall>\<^sub>F n in sequentially. continuous_on A (\<lambda>a. \<Sum>i<n. ?f i a)"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64008
diff changeset
   867
      by (intro always_eventually allI continuous_on_sum continuous_on_scaleR continuous_on_const f
64008
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   868
                from_nat_into set_pmf_not_empty)
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   869
    show "uniform_limit A (\<lambda>n a. \<Sum>i<n. ?f i a) (\<lambda>a. LINT i|M. f i a) sequentially"
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   870
    proof (subst uniform_limit_cong[where g="\<lambda>n a. \<Sum>i<n. ?f i a"])
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   871
      fix a assume "a \<in> A"
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   872
      have 1: "(LINT i|M. f i a) = (LINT i|map_pmf (to_nat_on M) M. f (from_nat_into M i) a)"
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   873
        by (auto intro!: integral_cong_AE AE_pmfI)
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   874
      have 2: "\<dots> = (LINT i|count_space UNIV. pmf (map_pmf (to_nat_on M) M) i *\<^sub>R f (from_nat_into M i) a)"
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   875
        by (simp add: measure_pmf_eq_density integral_density)
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   876
      have "(\<lambda>n. ?f n a) sums (LINT i|M. f i a)"
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   877
        unfolding 1 2
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   878
      proof (intro sums_integral_count_space_nat)
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   879
        have A: "integrable M (\<lambda>i. f i a)"
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   880
          using \<open>a\<in>A\<close> by (auto intro!: measure_pmf.integrable_const_bound AE_pmfI bnd)
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   881
        have "integrable (map_pmf (to_nat_on M) M) (\<lambda>i. f (from_nat_into M i) a)"
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   882
          by (auto simp add: map_pmf_rep_eq integrable_distr_eq intro!: AE_pmfI integrable_cong_AE_imp[OF A])
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   883
        then show "integrable (count_space UNIV) (\<lambda>n. ?f n a)"
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   884
          by (simp add: measure_pmf_eq_density integrable_density)
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   885
      qed
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   886
      then show "(LINT i|M. f i a) = (\<Sum> n. ?f n a)"
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   887
        by (simp add: sums_unique)
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   888
    next
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   889
      show "uniform_limit A (\<lambda>n a. \<Sum>i<n. ?f i a) (\<lambda>a. (\<Sum> n. ?f n a)) sequentially"
69529
4ab9657b3257 capitalize proper names in lemma names
nipkow
parents: 69313
diff changeset
   890
      proof (rule Weierstrass_m_test)
64008
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   891
        fix n a assume "a\<in>A"
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   892
        then show "norm (?f n a) \<le> pmf (map_pmf (to_nat_on M) M) n * B"
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   893
          using bnd by (auto intro!: mult_mono simp: from_nat_into set_pmf_not_empty)
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   894
      next
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   895
        have "integrable (map_pmf (to_nat_on M) M) (\<lambda>n. B)"
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   896
          by auto
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   897
        then show "summable (\<lambda>n. pmf (map_pmf (to_nat_on (set_pmf M)) M) n * B)"
70532
fcf3b891ccb1 new material; rotated premises of Lim_transform_eventually
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   898
          by (fastforce simp add: measure_pmf_eq_density integrable_density integrable_count_space_nat_iff summable_mult2)
64008
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   899
      qed
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   900
    qed simp
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   901
  qed simp
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   902
qed
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   903
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   904
lemma continuous_on_LBINT:
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   905
  fixes f :: "real \<Rightarrow> real"
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   906
  assumes f: "\<And>b. a \<le> b \<Longrightarrow> set_integrable lborel {a..b} f"
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   907
  shows "continuous_on UNIV (\<lambda>b. LBINT x:{a..b}. f x)"
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   908
proof (subst set_borel_integral_eq_integral)
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   909
  { fix b :: real assume "a \<le> b"
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   910
    from f[OF this] have "continuous_on {a..b} (\<lambda>b. integral {a..b} f)"
66192
e5b84854baa4 A few renamings and several tidied-up proofs
paulson <lp15@cam.ac.uk>
parents: 66089
diff changeset
   911
      by (intro indefinite_integral_continuous_1 set_borel_integral_eq_integral) }
64008
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   912
  note * = this
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   913
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   914
  have "continuous_on (\<Union>b\<in>{a..}. {a <..< b}) (\<lambda>b. integral {a..b} f)"
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   915
  proof (intro continuous_on_open_UN)
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   916
    show "b \<in> {a..} \<Longrightarrow> continuous_on {a<..<b} (\<lambda>b. integral {a..b} f)" for b
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   917
      using *[of b] by (rule continuous_on_subset) auto
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   918
  qed simp
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   919
  also have "(\<Union>b\<in>{a..}. {a <..< b}) = {a <..}"
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   920
    by (auto simp: lt_ex gt_ex less_imp_le) (simp add: Bex_def less_imp_le gt_ex cong: rev_conj_cong)
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   921
  finally have "continuous_on {a+1 ..} (\<lambda>b. integral {a..b} f)"
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   922
    by (rule continuous_on_subset) auto
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   923
  moreover have "continuous_on {a..a+1} (\<lambda>b. integral {a..b} f)"
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   924
    by (rule *) simp
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   925
  moreover
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   926
  have "x \<le> a \<Longrightarrow> {a..x} = (if a = x then {a} else {})" for x
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   927
    by auto
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   928
  then have "continuous_on {..a} (\<lambda>b. integral {a..b} f)"
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   929
    by (subst continuous_on_cong[OF refl, where g="\<lambda>x. 0"]) (auto intro!: continuous_on_const)
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   930
  ultimately have "continuous_on ({..a} \<union> {a..a+1} \<union> {a+1 ..}) (\<lambda>b. integral {a..b} f)"
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   931
    by (intro continuous_on_closed_Un) auto
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   932
  also have "{..a} \<union> {a..a+1} \<union> {a+1 ..} = UNIV"
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   933
    by auto
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   934
  finally show "continuous_on UNIV (\<lambda>b. integral {a..b} f)"
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   935
    by auto
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   936
next
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   937
  show "set_integrable lborel {a..b} f" for b
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   938
    using f by (cases "a \<le> b") auto
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   939
qed
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63973
diff changeset
   940
58587
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   941
locale pmf_as_function
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   942
begin
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   943
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   944
setup_lifting td_pmf_embed_pmf
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   945
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59665
diff changeset
   946
lemma set_pmf_transfer[transfer_rule]:
58730
b3fd0628f849 add transfer rule for set_pmf
hoelzl
parents: 58606
diff changeset
   947
  assumes "bi_total A"
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59665
diff changeset
   948
  shows "rel_fun (pcr_pmf A) (rel_set A) (\<lambda>f. {x. f x \<noteq> 0}) set_pmf"
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61634
diff changeset
   949
  using \<open>bi_total A\<close>
58730
b3fd0628f849 add transfer rule for set_pmf
hoelzl
parents: 58606
diff changeset
   950
  by (auto simp: pcr_pmf_def cr_pmf_def rel_fun_def rel_set_def bi_total_def Bex_def set_pmf_iff)
b3fd0628f849 add transfer rule for set_pmf
hoelzl
parents: 58606
diff changeset
   951
     metis+
b3fd0628f849 add transfer rule for set_pmf
hoelzl
parents: 58606
diff changeset
   952
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   953
end
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   954
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   955
context
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   956
begin
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   957
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   958
interpretation pmf_as_function .
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   959
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   960
lemma pmf_eqI: "(\<And>i. pmf M i = pmf N i) \<Longrightarrow> M = N"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   961
  by transfer auto
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   962
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   963
lemma pmf_eq_iff: "M = N \<longleftrightarrow> (\<forall>i. pmf M i = pmf N i)"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   964
  by (auto intro: pmf_eqI)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   965
63099
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
   966
lemma pmf_neq_exists_less:
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
   967
  assumes "M \<noteq> N"
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
   968
  shows   "\<exists>x. pmf M x < pmf N x"
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
   969
proof (rule ccontr)
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
   970
  assume "\<not>(\<exists>x. pmf M x < pmf N x)"
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
   971
  hence ge: "pmf M x \<ge> pmf N x" for x by (auto simp: not_less)
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
   972
  from assms obtain x where "pmf M x \<noteq> pmf N x" by (auto simp: pmf_eq_iff)
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
   973
  with ge[of x] have gt: "pmf M x > pmf N x" by simp
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
   974
  have "1 = measure (measure_pmf M) UNIV" by simp
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
   975
  also have "\<dots> = measure (measure_pmf N) {x} + measure (measure_pmf N) (UNIV - {x})"
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
   976
    by (subst measure_pmf.finite_measure_Union [symmetric]) simp_all
63886
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63882
diff changeset
   977
  also from gt have "measure (measure_pmf N) {x} < measure (measure_pmf M) {x}"
63099
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
   978
    by (simp add: measure_pmf_single)
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
   979
  also have "measure (measure_pmf N) (UNIV - {x}) \<le> measure (measure_pmf M) (UNIV - {x})"
63886
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63882
diff changeset
   980
    by (subst (1 2) integral_pmf [symmetric])
63099
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
   981
       (intro integral_mono integrable_pmf, simp_all add: ge)
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
   982
  also have "measure (measure_pmf M) {x} + \<dots> = 1"
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
   983
    by (subst measure_pmf.finite_measure_Union [symmetric]) simp_all
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
   984
  finally show False by simp_all
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
   985
qed
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
   986
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   987
lemma bind_commute_pmf: "bind_pmf A (\<lambda>x. bind_pmf B (C x)) = bind_pmf B (\<lambda>y. bind_pmf A (\<lambda>x. C x y))"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   988
  unfolding pmf_eq_iff pmf_bind
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   989
proof
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   990
  fix i
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   991
  interpret B: prob_space "restrict_space B B"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   992
    by (intro prob_space_restrict_space measure_pmf.emeasure_eq_1_AE)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   993
       (auto simp: AE_measure_pmf_iff)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   994
  interpret A: prob_space "restrict_space A A"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   995
    by (intro prob_space_restrict_space measure_pmf.emeasure_eq_1_AE)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   996
       (auto simp: AE_measure_pmf_iff)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   997
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   998
  interpret AB: pair_prob_space "restrict_space A A" "restrict_space B B"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   999
    by unfold_locales
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1000
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1001
  have "(\<integral> x. \<integral> y. pmf (C x y) i \<partial>B \<partial>A) = (\<integral> x. (\<integral> y. pmf (C x y) i \<partial>restrict_space B B) \<partial>A)"
63886
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63882
diff changeset
  1002
    by (rule Bochner_Integration.integral_cong) (auto intro!: integral_pmf_restrict)
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1003
  also have "\<dots> = (\<integral> x. (\<integral> y. pmf (C x y) i \<partial>restrict_space B B) \<partial>restrict_space A A)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1004
    by (intro integral_pmf_restrict B.borel_measurable_lebesgue_integral measurable_pair_restrict_pmf2
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1005
              countable_set_pmf borel_measurable_count_space)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1006
  also have "\<dots> = (\<integral> y. \<integral> x. pmf (C x y) i \<partial>restrict_space A A \<partial>restrict_space B B)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1007
    by (rule AB.Fubini_integral[symmetric])
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1008
       (auto intro!: AB.integrable_const_bound[where B=1] measurable_pair_restrict_pmf2
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1009
             simp: pmf_nonneg pmf_le_1 measurable_restrict_space1)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1010
  also have "\<dots> = (\<integral> y. \<integral> x. pmf (C x y) i \<partial>restrict_space A A \<partial>B)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1011
    by (intro integral_pmf_restrict[symmetric] A.borel_measurable_lebesgue_integral measurable_pair_restrict_pmf2
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1012
              countable_set_pmf borel_measurable_count_space)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1013
  also have "\<dots> = (\<integral> y. \<integral> x. pmf (C x y) i \<partial>A \<partial>B)"
63886
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63882
diff changeset
  1014
    by (rule Bochner_Integration.integral_cong) (auto intro!: integral_pmf_restrict[symmetric])
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1015
  finally show "(\<integral> x. \<integral> y. pmf (C x y) i \<partial>B \<partial>A) = (\<integral> y. \<integral> x. pmf (C x y) i \<partial>A \<partial>B)" .
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1016
qed
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1017
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1018
lemma pair_map_pmf1: "pair_pmf (map_pmf f A) B = map_pmf (apfst f) (pair_pmf A B)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1019
proof (safe intro!: pmf_eqI)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1020
  fix a :: "'a" and b :: "'b"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1021
  have [simp]: "\<And>c d. indicator (apfst f -` {(a, b)}) (c, d) = indicator (f -` {a}) c * (indicator {b} d::ennreal)"
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1022
    by (auto split: split_indicator)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1023
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1024
  have "ennreal (pmf (pair_pmf (map_pmf f A) B) (a, b)) =
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1025
         ennreal (pmf (map_pmf (apfst f) (pair_pmf A B)) (a, b))"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1026
    unfolding pmf_pair ennreal_pmf_map
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1027
    by (simp add: nn_integral_pair_pmf' max_def emeasure_pmf_single nn_integral_multc pmf_nonneg
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1028
                  emeasure_map_pmf[symmetric] ennreal_mult del: emeasure_map_pmf)
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1029
  then show "pmf (pair_pmf (map_pmf f A) B) (a, b) = pmf (map_pmf (apfst f) (pair_pmf A B)) (a, b)"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1030
    by (simp add: pmf_nonneg)
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1031
qed
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1032
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1033
lemma pair_map_pmf2: "pair_pmf A (map_pmf f B) = map_pmf (apsnd f) (pair_pmf A B)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1034
proof (safe intro!: pmf_eqI)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1035
  fix a :: "'a" and b :: "'b"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1036
  have [simp]: "\<And>c d. indicator (apsnd f -` {(a, b)}) (c, d) = indicator {a} c * (indicator (f -` {b}) d::ennreal)"
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1037
    by (auto split: split_indicator)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1038
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1039
  have "ennreal (pmf (pair_pmf A (map_pmf f B)) (a, b)) =
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1040
         ennreal (pmf (map_pmf (apsnd f) (pair_pmf A B)) (a, b))"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1041
    unfolding pmf_pair ennreal_pmf_map
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1042
    by (simp add: nn_integral_pair_pmf' max_def emeasure_pmf_single nn_integral_cmult nn_integral_multc pmf_nonneg
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1043
                  emeasure_map_pmf[symmetric] ennreal_mult del: emeasure_map_pmf)
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1044
  then show "pmf (pair_pmf A (map_pmf f B)) (a, b) = pmf (map_pmf (apsnd f) (pair_pmf A B)) (a, b)"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1045
    by (simp add: pmf_nonneg)
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1046
qed
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1047
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1048
lemma map_pair: "map_pmf (\<lambda>(a, b). (f a, g b)) (pair_pmf A B) = pair_pmf (map_pmf f A) (map_pmf g B)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1049
  by (simp add: pair_map_pmf2 pair_map_pmf1 map_pmf_comp split_beta')
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1050
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1051
end
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1052
61634
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1053
lemma pair_return_pmf1: "pair_pmf (return_pmf x) y = map_pmf (Pair x) y"
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1054
by(simp add: pair_pmf_def bind_return_pmf map_pmf_def)
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1055
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1056
lemma pair_return_pmf2: "pair_pmf x (return_pmf y) = map_pmf (\<lambda>x. (x, y)) x"
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1057
by(simp add: pair_pmf_def bind_return_pmf map_pmf_def)
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1058
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1059
lemma pair_pair_pmf: "pair_pmf (pair_pmf u v) w = map_pmf (\<lambda>(x, (y, z)). ((x, y), z)) (pair_pmf u (pair_pmf v w))"
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1060
by(simp add: pair_pmf_def bind_return_pmf map_pmf_def bind_assoc_pmf)
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1061
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1062
lemma pair_commute_pmf: "pair_pmf x y = map_pmf (\<lambda>(x, y). (y, x)) (pair_pmf y x)"
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1063
unfolding pair_pmf_def by(subst bind_commute_pmf)(simp add: map_pmf_def bind_assoc_pmf bind_return_pmf)
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1064
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1065
lemma set_pmf_subset_singleton: "set_pmf p \<subseteq> {x} \<longleftrightarrow> p = return_pmf x"
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1066
proof(intro iffI pmf_eqI)
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1067
  fix i
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1068
  assume x: "set_pmf p \<subseteq> {x}"
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1069
  hence *: "set_pmf p = {x}" using set_pmf_not_empty[of p] by auto
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1070
  have "ennreal (pmf p x) = \<integral>\<^sup>+ i. indicator {x} i \<partial>p" by(simp add: emeasure_pmf_single)
61634
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1071
  also have "\<dots> = \<integral>\<^sup>+ i. 1 \<partial>p" by(rule nn_integral_cong_AE)(simp add: AE_measure_pmf_iff * )
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1072
  also have "\<dots> = 1" by simp
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1073
  finally show "pmf p i = pmf (return_pmf x) i" using x
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1074
    by(auto split: split_indicator simp add: pmf_eq_0_set_pmf)
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1075
qed auto
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1076
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1077
lemma bind_eq_return_pmf:
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1078
  "bind_pmf p f = return_pmf x \<longleftrightarrow> (\<forall>y\<in>set_pmf p. f y = return_pmf x)"
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1079
  (is "?lhs \<longleftrightarrow> ?rhs")
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1080
proof(intro iffI strip)
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1081
  fix y
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1082
  assume y: "y \<in> set_pmf p"
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1083
  assume "?lhs"
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1084
  hence "set_pmf (bind_pmf p f) = {x}" by simp
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1085
  hence "(\<Union>y\<in>set_pmf p. set_pmf (f y)) = {x}" by simp
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1086
  hence "set_pmf (f y) \<subseteq> {x}" using y by auto
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1087
  thus "f y = return_pmf x" by(simp add: set_pmf_subset_singleton)
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1088
next
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1089
  assume *: ?rhs
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1090
  show ?lhs
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1091
  proof(rule pmf_eqI)
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1092
    fix i
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1093
    have "ennreal (pmf (bind_pmf p f) i) = \<integral>\<^sup>+ y. ennreal (pmf (f y) i) \<partial>p"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1094
      by (simp add: ennreal_pmf_bind)
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1095
    also have "\<dots> = \<integral>\<^sup>+ y. ennreal (pmf (return_pmf x) i) \<partial>p"
61634
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1096
      by(rule nn_integral_cong_AE)(simp add: AE_measure_pmf_iff * )
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1097
    also have "\<dots> = ennreal (pmf (return_pmf x) i)"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1098
      by simp
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1099
    finally show "pmf (bind_pmf p f) i = pmf (return_pmf x) i"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1100
      by (simp add: pmf_nonneg)
61634
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1101
  qed
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1102
qed
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1103
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1104
lemma pmf_False_conv_True: "pmf p False = 1 - pmf p True"
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1105
proof -
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1106
  have "pmf p False + pmf p True = measure p {False} + measure p {True}"
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1107
    by(simp add: measure_pmf_single)
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1108
  also have "\<dots> = measure p ({False} \<union> {True})"
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1109
    by(subst measure_pmf.finite_measure_Union) simp_all
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1110
  also have "{False} \<union> {True} = space p" by auto
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1111
  finally show ?thesis by simp
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1112
qed
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1113
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1114
lemma pmf_True_conv_False: "pmf p True = 1 - pmf p False"
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1115
by(simp add: pmf_False_conv_True)
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1116
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1117
subsection \<open> Conditional Probabilities \<close>
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1118
59670
dee043d19729 generalized bind_cond_pmf_cancel
hoelzl
parents: 59667
diff changeset
  1119
lemma measure_pmf_zero_iff: "measure (measure_pmf p) s = 0 \<longleftrightarrow> set_pmf p \<inter> s = {}"
dee043d19729 generalized bind_cond_pmf_cancel
hoelzl
parents: 59667
diff changeset
  1120
  by (subst measure_pmf.prob_eq_0) (auto simp: AE_measure_pmf_iff)
dee043d19729 generalized bind_cond_pmf_cancel
hoelzl
parents: 59667
diff changeset
  1121
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1122
context
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1123
  fixes p :: "'a pmf" and s :: "'a set"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1124
  assumes not_empty: "set_pmf p \<inter> s \<noteq> {}"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1125
begin
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1126
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1127
interpretation pmf_as_measure .
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1128
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1129
lemma emeasure_measure_pmf_not_zero: "emeasure (measure_pmf p) s \<noteq> 0"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1130
proof
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1131
  assume "emeasure (measure_pmf p) s = 0"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1132
  then have "AE x in measure_pmf p. x \<notin> s"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1133
    by (rule AE_I[rotated]) auto
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1134
  with not_empty show False
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1135
    by (auto simp: AE_measure_pmf_iff)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1136
qed
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1137
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1138
lemma measure_measure_pmf_not_zero: "measure (measure_pmf p) s \<noteq> 0"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1139
  using emeasure_measure_pmf_not_zero by (simp add: measure_pmf.emeasure_eq_measure measure_nonneg)
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1140
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1141
lift_definition cond_pmf :: "'a pmf" is
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1142
  "uniform_measure (measure_pmf p) s"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1143
proof (intro conjI)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1144
  show "prob_space (uniform_measure (measure_pmf p) s)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1145
    by (intro prob_space_uniform_measure) (auto simp: emeasure_measure_pmf_not_zero)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1146
  show "AE x in uniform_measure (measure_pmf p) s. measure (uniform_measure (measure_pmf p) s) {x} \<noteq> 0"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1147
    by (simp add: emeasure_measure_pmf_not_zero measure_measure_pmf_not_zero AE_uniform_measure
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1148
                  AE_measure_pmf_iff set_pmf.rep_eq less_top[symmetric])
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1149
qed simp
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1150
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1151
lemma pmf_cond: "pmf cond_pmf x = (if x \<in> s then pmf p x / measure p s else 0)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1152
  by transfer (simp add: emeasure_measure_pmf_not_zero pmf.rep_eq)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1153
59665
37adca7fd48f add set_pmf lemmas to simpset
hoelzl
parents: 59664
diff changeset
  1154
lemma set_cond_pmf[simp]: "set_pmf cond_pmf = set_pmf p \<inter> s"
62390
842917225d56 more canonical names
nipkow
parents: 62324
diff changeset
  1155
  by (auto simp add: set_pmf_iff pmf_cond measure_measure_pmf_not_zero split: if_split_asm)
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1156
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1157
end
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1158
63099
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  1159
lemma measure_pmf_posI: "x \<in> set_pmf p \<Longrightarrow> x \<in> A \<Longrightarrow> measure_pmf.prob p A > 0"
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  1160
  using measure_measure_pmf_not_zero[of p A] by (subst zero_less_measure_iff) blast
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  1161
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1162
lemma cond_map_pmf:
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1163
  assumes "set_pmf p \<inter> f -` s \<noteq> {}"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1164
  shows "cond_pmf (map_pmf f p) s = map_pmf f (cond_pmf p (f -` s))"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1165
proof -
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1166
  have *: "set_pmf (map_pmf f p) \<inter> s \<noteq> {}"
59665
37adca7fd48f add set_pmf lemmas to simpset
hoelzl
parents: 59664
diff changeset
  1167
    using assms by auto
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1168
  { fix x
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1169
    have "ennreal (pmf (map_pmf f (cond_pmf p (f -` s))) x) =
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1170
      emeasure p (f -` s \<inter> f -` {x}) / emeasure p (f -` s)"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1171
      unfolding ennreal_pmf_map cond_pmf.rep_eq[OF assms] by (simp add: nn_integral_uniform_measure)
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1172
    also have "f -` s \<inter> f -` {x} = (if x \<in> s then f -` {x} else {})"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1173
      by auto
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1174
    also have "emeasure p (if x \<in> s then f -` {x} else {}) / emeasure p (f -` s) =
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1175
      ennreal (pmf (cond_pmf (map_pmf f p) s) x)"
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1176
      using measure_measure_pmf_not_zero[OF *]
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1177
      by (simp add: pmf_cond[OF *] ennreal_pmf_map measure_pmf.emeasure_eq_measure
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1178
                    divide_ennreal pmf_nonneg measure_nonneg zero_less_measure_iff pmf_map)
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1179
    finally have "ennreal (pmf (cond_pmf (map_pmf f p) s) x) = ennreal (pmf (map_pmf f (cond_pmf p (f -` s))) x)"
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1180
      by simp }
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1181
  then show ?thesis
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1182
    by (intro pmf_eqI) (simp add: pmf_nonneg)
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1183
qed
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1184
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1185
lemma bind_cond_pmf_cancel:
59670
dee043d19729 generalized bind_cond_pmf_cancel
hoelzl
parents: 59667
diff changeset
  1186
  assumes [simp]: "\<And>x. x \<in> set_pmf p \<Longrightarrow> set_pmf q \<inter> {y. R x y} \<noteq> {}"
dee043d19729 generalized bind_cond_pmf_cancel
hoelzl
parents: 59667
diff changeset
  1187
  assumes [simp]: "\<And>y. y \<in> set_pmf q \<Longrightarrow> set_pmf p \<inter> {x. R x y} \<noteq> {}"
dee043d19729 generalized bind_cond_pmf_cancel
hoelzl
parents: 59667
diff changeset
  1188
  assumes [simp]: "\<And>x y. x \<in> set_pmf p \<Longrightarrow> y \<in> set_pmf q \<Longrightarrow> R x y \<Longrightarrow> measure q {y. R x y} = measure p {x. R x y}"
dee043d19729 generalized bind_cond_pmf_cancel
hoelzl
parents: 59667
diff changeset
  1189
  shows "bind_pmf p (\<lambda>x. cond_pmf q {y. R x y}) = q"
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1190
proof (rule pmf_eqI)
59670
dee043d19729 generalized bind_cond_pmf_cancel
hoelzl
parents: 59667
diff changeset
  1191
  fix i
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1192
  have "ennreal (pmf (bind_pmf p (\<lambda>x. cond_pmf q {y. R x y})) i) =
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1193
    (\<integral>\<^sup>+x. ennreal (pmf q i / measure p {x. R x i}) * ennreal (indicator {x. R x i} x) \<partial>p)"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1194
    by (auto simp add: ennreal_pmf_bind AE_measure_pmf_iff pmf_cond pmf_eq_0_set_pmf pmf_nonneg measure_nonneg
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1195
             intro!: nn_integral_cong_AE)
59670
dee043d19729 generalized bind_cond_pmf_cancel
hoelzl
parents: 59667
diff changeset
  1196
  also have "\<dots> = (pmf q i * measure p {x. R x i}) / measure p {x. R x i}"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1197
    by (simp add: pmf_nonneg measure_nonneg zero_ennreal_def[symmetric] ennreal_indicator
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1198
                  nn_integral_cmult measure_pmf.emeasure_eq_measure ennreal_mult[symmetric])
59670
dee043d19729 generalized bind_cond_pmf_cancel
hoelzl
parents: 59667
diff changeset
  1199
  also have "\<dots> = pmf q i"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1200
    by (cases "pmf q i = 0")
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1201
       (simp_all add: pmf_eq_0_set_pmf measure_measure_pmf_not_zero pmf_nonneg)
59670
dee043d19729 generalized bind_cond_pmf_cancel
hoelzl
parents: 59667
diff changeset
  1202
  finally show "pmf (bind_pmf p (\<lambda>x. cond_pmf q {y. R x y})) i = pmf q i"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1203
    by (simp add: pmf_nonneg)
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1204
qed
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1205
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1206
subsection \<open> Relator \<close>
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1207
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1208
inductive rel_pmf :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'a pmf \<Rightarrow> 'b pmf \<Rightarrow> bool"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1209
for R p q
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1210
where
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59665
diff changeset
  1211
  "\<lbrakk> \<And>x y. (x, y) \<in> set_pmf pq \<Longrightarrow> R x y;
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1212
     map_pmf fst pq = p; map_pmf snd pq = q \<rbrakk>
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1213
  \<Longrightarrow> rel_pmf R p q"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1214
59681
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1215
lemma rel_pmfI:
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1216
  assumes R: "rel_set R (set_pmf p) (set_pmf q)"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1217
  assumes eq: "\<And>x y. x \<in> set_pmf p \<Longrightarrow> y \<in> set_pmf q \<Longrightarrow> R x y \<Longrightarrow>
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1218
    measure p {x. R x y} = measure q {y. R x y}"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1219
  shows "rel_pmf R p q"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1220
proof
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1221
  let ?pq = "bind_pmf p (\<lambda>x. bind_pmf (cond_pmf q {y. R x y}) (\<lambda>y. return_pmf (x, y)))"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1222
  have "\<And>x. x \<in> set_pmf p \<Longrightarrow> set_pmf q \<inter> {y. R x y} \<noteq> {}"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1223
    using R by (auto simp: rel_set_def)
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1224
  then show "\<And>x y. (x, y) \<in> set_pmf ?pq \<Longrightarrow> R x y"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1225
    by auto
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1226
  show "map_pmf fst ?pq = p"
60068
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
  1227
    by (simp add: map_bind_pmf bind_return_pmf')
59681
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1228
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1229
  show "map_pmf snd ?pq = q"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1230
    using R eq
60068
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
  1231
    apply (simp add: bind_cond_pmf_cancel map_bind_pmf bind_return_pmf')
59681
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1232
    apply (rule bind_cond_pmf_cancel)
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1233
    apply (auto simp: rel_set_def)
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1234
    done
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1235
qed
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1236
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1237
lemma rel_pmf_imp_rel_set: "rel_pmf R p q \<Longrightarrow> rel_set R (set_pmf p) (set_pmf q)"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1238
  by (force simp add: rel_pmf.simps rel_set_def)
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1239
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1240
lemma rel_pmfD_measure:
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1241
  assumes rel_R: "rel_pmf R p q" and R: "\<And>a b. R a b \<Longrightarrow> R a y \<longleftrightarrow> R x b"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1242
  assumes "x \<in> set_pmf p" "y \<in> set_pmf q"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1243
  shows "measure p {x. R x y} = measure q {y. R x y}"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1244
proof -
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1245
  from rel_R obtain pq where pq: "\<And>x y. (x, y) \<in> set_pmf pq \<Longrightarrow> R x y"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1246
    and eq: "p = map_pmf fst pq" "q = map_pmf snd pq"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1247
    by (auto elim: rel_pmf.cases)
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1248
  have "measure p {x. R x y} = measure pq {x. R (fst x) y}"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1249
    by (simp add: eq map_pmf_rep_eq measure_distr)
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1250
  also have "\<dots> = measure pq {y. R x (snd y)}"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1251
    by (intro measure_pmf.finite_measure_eq_AE)
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1252
       (auto simp: AE_measure_pmf_iff R dest!: pq)
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1253
  also have "\<dots> = measure q {y. R x y}"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1254
    by (simp add: eq map_pmf_rep_eq measure_distr)
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1255
  finally show "measure p {x. R x y} = measure q {y. R x y}" .
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1256
qed
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1257
61634
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1258
lemma rel_pmf_measureD:
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1259
  assumes "rel_pmf R p q"
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1260
  shows "measure (measure_pmf p) A \<le> measure (measure_pmf q) {y. \<exists>x\<in>A. R x y}" (is "?lhs \<le> ?rhs")
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1261
using assms
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1262
proof cases
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1263
  fix pq
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1264
  assume R: "\<And>x y. (x, y) \<in> set_pmf pq \<Longrightarrow> R x y"
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1265
    and p[symmetric]: "map_pmf fst pq = p"
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1266
    and q[symmetric]: "map_pmf snd pq = q"
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1267
  have "?lhs = measure (measure_pmf pq) (fst -` A)" by(simp add: p)
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1268
  also have "\<dots> \<le> measure (measure_pmf pq) {y. \<exists>x\<in>A. R x (snd y)}"
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1269
    by(rule measure_pmf.finite_measure_mono_AE)(auto 4 3 simp add: AE_measure_pmf_iff dest: R)
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1270
  also have "\<dots> = ?rhs" by(simp add: q)
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1271
  finally show ?thesis .
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1272
qed
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1273
59681
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1274
lemma rel_pmf_iff_measure:
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1275
  assumes "symp R" "transp R"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1276
  shows "rel_pmf R p q \<longleftrightarrow>
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1277
    rel_set R (set_pmf p) (set_pmf q) \<and>
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1278
    (\<forall>x\<in>set_pmf p. \<forall>y\<in>set_pmf q. R x y \<longrightarrow> measure p {x. R x y} = measure q {y. R x y})"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1279
  by (safe intro!: rel_pmf_imp_rel_set rel_pmfI)
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1280
     (auto intro!: rel_pmfD_measure dest: sympD[OF \<open>symp R\<close>] transpD[OF \<open>transp R\<close>])
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1281
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1282
lemma quotient_rel_set_disjoint:
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1283
  "equivp R \<Longrightarrow> C \<in> UNIV // {(x, y). R x y} \<Longrightarrow> rel_set R A B \<Longrightarrow> A \<inter> C = {} \<longleftrightarrow> B \<inter> C = {}"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1284
  using in_quotient_imp_closed[of UNIV "{(x, y). R x y}" C]
59681
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1285
  by (auto 0 0 simp: equivp_equiv rel_set_def set_eq_iff elim: equivpE)
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1286
     (blast dest: equivp_symp)+
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1287
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1288
lemma quotientD: "equiv X R \<Longrightarrow> A \<in> X // R \<Longrightarrow> x \<in> A \<Longrightarrow> A = R `` {x}"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1289
  by (metis Image_singleton_iff equiv_class_eq_iff quotientE)
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1290
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1291
lemma rel_pmf_iff_equivp:
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1292
  assumes "equivp R"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1293
  shows "rel_pmf R p q \<longleftrightarrow> (\<forall>C\<in>UNIV // {(x, y). R x y}. measure p C = measure q C)"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1294
    (is "_ \<longleftrightarrow>   (\<forall>C\<in>_//?R. _)")
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1295
proof (subst rel_pmf_iff_measure, safe)
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1296
  show "symp R" "transp R"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1297
    using assms by (auto simp: equivp_reflp_symp_transp)
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1298
next
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1299
  fix C assume C: "C \<in> UNIV // ?R" and R: "rel_set R (set_pmf p) (set_pmf q)"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1300
  assume eq: "\<forall>x\<in>set_pmf p. \<forall>y\<in>set_pmf q. R x y \<longrightarrow> measure p {x. R x y} = measure q {y. R x y}"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1301
59681
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1302
  show "measure p C = measure q C"
63540
f8652d0534fa tuned proofs -- avoid unstructured calculation;
wenzelm
parents: 63343
diff changeset
  1303
  proof (cases "p \<inter> C = {}")
f8652d0534fa tuned proofs -- avoid unstructured calculation;
wenzelm
parents: 63343
diff changeset
  1304
    case True
f8652d0534fa tuned proofs -- avoid unstructured calculation;
wenzelm
parents: 63343
diff changeset
  1305
    then have "q \<inter> C = {}"
59681
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1306
      using quotient_rel_set_disjoint[OF assms C R] by simp
63540
f8652d0534fa tuned proofs -- avoid unstructured calculation;
wenzelm
parents: 63343
diff changeset
  1307
    with True show ?thesis
59681
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1308
      unfolding measure_pmf_zero_iff[symmetric] by simp
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1309
  next
63540
f8652d0534fa tuned proofs -- avoid unstructured calculation;
wenzelm
parents: 63343
diff changeset
  1310
    case False
f8652d0534fa tuned proofs -- avoid unstructured calculation;
wenzelm
parents: 63343
diff changeset
  1311
    then have "q \<inter> C \<noteq> {}"
59681
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1312
      using quotient_rel_set_disjoint[OF assms C R] by simp
63540
f8652d0534fa tuned proofs -- avoid unstructured calculation;
wenzelm
parents: 63343
diff changeset
  1313
    with False obtain x y where in_set: "x \<in> set_pmf p" "y \<in> set_pmf q" and in_C: "x \<in> C" "y \<in> C"
59681
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1314
      by auto
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1315
    then have "R x y"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1316
      using in_quotient_imp_in_rel[of UNIV ?R C x y] C assms
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1317
      by (simp add: equivp_equiv)
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1318
    with in_set eq have "measure p {x. R x y} = measure q {y. R x y}"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1319
      by auto
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1320
    moreover have "{y. R x y} = C"
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61634
diff changeset
  1321
      using assms \<open>x \<in> C\<close> C quotientD[of UNIV ?R C x] by (simp add: equivp_equiv)
59681
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1322
    moreover have "{x. R x y} = C"
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61634
diff changeset
  1323
      using assms \<open>y \<in> C\<close> C quotientD[of UNIV "?R" C y] sympD[of R]
59681
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1324
      by (auto simp add: equivp_equiv elim: equivpE)
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1325
    ultimately show ?thesis
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1326
      by auto
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1327
  qed
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1328
next
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1329
  assume eq: "\<forall>C\<in>UNIV // ?R. measure p C = measure q C"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1330
  show "rel_set R (set_pmf p) (set_pmf q)"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1331
    unfolding rel_set_def
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1332
  proof safe
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1333
    fix x assume x: "x \<in> set_pmf p"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1334
    have "{y. R x y} \<in> UNIV // ?R"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1335
      by (auto simp: quotient_def)
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1336
    with eq have *: "measure q {y. R x y} = measure p {y. R x y}"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1337
      by auto
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1338
    have "measure q {y. R x y} \<noteq> 0"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1339
      using x assms unfolding * by (auto simp: measure_pmf_zero_iff set_eq_iff dest: equivp_reflp)
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1340
    then show "\<exists>y\<in>set_pmf q. R x y"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1341
      unfolding measure_pmf_zero_iff by auto
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1342
  next
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1343
    fix y assume y: "y \<in> set_pmf q"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1344
    have "{x. R x y} \<in> UNIV // ?R"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1345
      using assms by (auto simp: quotient_def dest: equivp_symp)
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1346
    with eq have *: "measure p {x. R x y} = measure q {x. R x y}"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1347
      by auto
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1348
    have "measure p {x. R x y} \<noteq> 0"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1349
      using y assms unfolding * by (auto simp: measure_pmf_zero_iff set_eq_iff dest: equivp_reflp)
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1350
    then show "\<exists>x\<in>set_pmf p. R x y"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1351
      unfolding measure_pmf_zero_iff by auto
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1352
  qed
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1353
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1354
  fix x y assume "x \<in> set_pmf p" "y \<in> set_pmf q" "R x y"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1355
  have "{y. R x y} \<in> UNIV // ?R" "{x. R x y} = {y. R x y}"
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61634
diff changeset
  1356
    using assms \<open>R x y\<close> by (auto simp: quotient_def dest: equivp_symp equivp_transp)
59681
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1357
  with eq show "measure p {x. R x y} = measure q {y. R x y}"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1358
    by auto
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1359
qed
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1360
75624
22d1c5f2b9f4 strict bounds for BNFs (by Jan van Brügge)
traytel
parents: 73253
diff changeset
  1361
bnf pmf: "'a pmf" map: map_pmf sets: set_pmf bd : "card_suc natLeq" rel: rel_pmf
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1362
proof -
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1363
  show "map_pmf id = id" by (rule map_pmf_id)
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59665
diff changeset
  1364
  show "\<And>f g. map_pmf (f \<circ> g) = map_pmf f \<circ> map_pmf g" by (rule map_pmf_compose)
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1365
  show "\<And>f g::'a \<Rightarrow> 'b. \<And>p. (\<And>x. x \<in> set_pmf p \<Longrightarrow> f x = g x) \<Longrightarrow> map_pmf f p = map_pmf g p"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1366
    by (intro map_pmf_cong refl)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1367
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67226
diff changeset
  1368
  show "\<And>f::'a \<Rightarrow> 'b. set_pmf \<circ> map_pmf f = (`) f \<circ> set_pmf"
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1369
    by (rule pmf_set_map)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1370
75624
22d1c5f2b9f4 strict bounds for BNFs (by Jan van Brügge)
traytel
parents: 73253
diff changeset
  1371
  show "card_order (card_suc natLeq)" using natLeq_card_order by (rule card_order_card_suc)
22d1c5f2b9f4 strict bounds for BNFs (by Jan van Brügge)
traytel
parents: 73253
diff changeset
  1372
  show "BNF_Cardinal_Arithmetic.cinfinite (card_suc natLeq)"
22d1c5f2b9f4 strict bounds for BNFs (by Jan van Brügge)
traytel
parents: 73253
diff changeset
  1373
    using natLeq_Cinfinite natLeq_card_order Cinfinite_card_suc by blast
22d1c5f2b9f4 strict bounds for BNFs (by Jan van Brügge)
traytel
parents: 73253
diff changeset
  1374
  show "regularCard (card_suc natLeq)" using natLeq_card_order natLeq_Cinfinite
75625
0dd3ac5fdbaa tuned BNF bounds for function space and bounded sets; NEWS and CONTRIBUTORS
traytel
parents: 75624
diff changeset
  1375
    by (rule regularCard_card_suc)
75624
22d1c5f2b9f4 strict bounds for BNFs (by Jan van Brügge)
traytel
parents: 73253
diff changeset
  1376
22d1c5f2b9f4 strict bounds for BNFs (by Jan van Brügge)
traytel
parents: 73253
diff changeset
  1377
  show "(card_of (set_pmf p), card_suc natLeq) \<in> ordLess" for p :: "'s pmf"
60595
804dfdc82835 premises in 'show' are treated like 'assume';
wenzelm
parents: 60495
diff changeset
  1378
  proof -
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1379
    have "(card_of (set_pmf p), card_of (UNIV :: nat set)) \<in> ordLeq"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1380
      by (rule card_of_ordLeqI[where f="to_nat_on (set_pmf p)"])
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1381
         (auto intro: countable_set_pmf)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1382
    also have "(card_of (UNIV :: nat set), natLeq) \<in> ordLeq"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1383
      by (metis Field_natLeq card_of_least natLeq_Well_order)
75624
22d1c5f2b9f4 strict bounds for BNFs (by Jan van Brügge)
traytel
parents: 73253
diff changeset
  1384
    finally show ?thesis using card_suc_greater natLeq_card_order ordLeq_ordLess_trans by blast
60595
804dfdc82835 premises in 'show' are treated like 'assume';
wenzelm
parents: 60495
diff changeset
  1385
  qed
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1386
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62083
diff changeset
  1387
  show "\<And>R. rel_pmf R = (\<lambda>x y. \<exists>z. set_pmf z \<subseteq> {(x, y). R x y} \<and>
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62083
diff changeset
  1388
    map_pmf fst z = x \<and> map_pmf snd z = y)"
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 62083
diff changeset
  1389
     by (auto simp add: fun_eq_iff rel_pmf.simps)
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1390
60595
804dfdc82835 premises in 'show' are treated like 'assume';
wenzelm
parents: 60495
diff changeset
  1391
  show "rel_pmf R OO rel_pmf S \<le> rel_pmf (R OO S)"
804dfdc82835 premises in 'show' are treated like 'assume';
wenzelm
parents: 60495
diff changeset
  1392
    for R :: "'a \<Rightarrow> 'b \<Rightarrow> bool" and S :: "'b \<Rightarrow> 'c \<Rightarrow> bool"
804dfdc82835 premises in 'show' are treated like 'assume';
wenzelm
parents: 60495
diff changeset
  1393
  proof -
804dfdc82835 premises in 'show' are treated like 'assume';
wenzelm
parents: 60495
diff changeset
  1394
    { fix p q r
804dfdc82835 premises in 'show' are treated like 'assume';
wenzelm
parents: 60495
diff changeset
  1395
      assume pq: "rel_pmf R p q"
804dfdc82835 premises in 'show' are treated like 'assume';
wenzelm
parents: 60495
diff changeset
  1396
        and qr:"rel_pmf S q r"
804dfdc82835 premises in 'show' are treated like 'assume';
wenzelm
parents: 60495
diff changeset
  1397
      from pq obtain pq where pq: "\<And>x y. (x, y) \<in> set_pmf pq \<Longrightarrow> R x y"
804dfdc82835 premises in 'show' are treated like 'assume';
wenzelm
parents: 60495
diff changeset
  1398
        and p: "p = map_pmf fst pq" and q: "q = map_pmf snd pq" by cases auto
804dfdc82835 premises in 'show' are treated like 'assume';
wenzelm
parents: 60495
diff changeset
  1399
      from qr obtain qr where qr: "\<And>y z. (y, z) \<in> set_pmf qr \<Longrightarrow> S y z"
804dfdc82835 premises in 'show' are treated like 'assume';
wenzelm
parents: 60495
diff changeset
  1400
        and q': "q = map_pmf fst qr" and r: "r = map_pmf snd qr" by cases auto
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1401
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62975
diff changeset
  1402
      define pr where "pr =
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62975
diff changeset
  1403
        bind_pmf pq (\<lambda>xy. bind_pmf (cond_pmf qr {yz. fst yz = snd xy})
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62975
diff changeset
  1404
          (\<lambda>yz. return_pmf (fst xy, snd yz)))"
60595
804dfdc82835 premises in 'show' are treated like 'assume';
wenzelm
parents: 60495
diff changeset
  1405
      have pr_welldefined: "\<And>y. y \<in> q \<Longrightarrow> qr \<inter> {yz. fst yz = y} \<noteq> {}"
804dfdc82835 premises in 'show' are treated like 'assume';
wenzelm
parents: 60495
diff changeset
  1406
        by (force simp: q')
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1407
60595
804dfdc82835 premises in 'show' are treated like 'assume';
wenzelm
parents: 60495
diff changeset
  1408
      have "rel_pmf (R OO S) p r"
804dfdc82835 premises in 'show' are treated like 'assume';
wenzelm
parents: 60495
diff changeset
  1409
      proof (rule rel_pmf.intros)
804dfdc82835 premises in 'show' are treated like 'assume';
wenzelm
parents: 60495
diff changeset
  1410
        fix x z assume "(x, z) \<in> pr"
804dfdc82835 premises in 'show' are treated like 'assume';
wenzelm
parents: 60495
diff changeset
  1411
        then have "\<exists>y. (x, y) \<in> pq \<and> (y, z) \<in> qr"
804dfdc82835 premises in 'show' are treated like 'assume';
wenzelm
parents: 60495
diff changeset
  1412
          by (auto simp: q pr_welldefined pr_def split_beta)
804dfdc82835 premises in 'show' are treated like 'assume';
wenzelm
parents: 60495
diff changeset
  1413
        with pq qr show "(R OO S) x z"
804dfdc82835 premises in 'show' are treated like 'assume';
wenzelm
parents: 60495
diff changeset
  1414
          by blast
804dfdc82835 premises in 'show' are treated like 'assume';
wenzelm
parents: 60495
diff changeset
  1415
      next
804dfdc82835 premises in 'show' are treated like 'assume';
wenzelm
parents: 60495
diff changeset
  1416
        have "map_pmf snd pr = map_pmf snd (bind_pmf q (\<lambda>y. cond_pmf qr {yz. fst yz = y}))"
804dfdc82835 premises in 'show' are treated like 'assume';
wenzelm
parents: 60495
diff changeset
  1417
          by (simp add: pr_def q split_beta bind_map_pmf map_pmf_def[symmetric] map_bind_pmf map_pmf_comp)
804dfdc82835 premises in 'show' are treated like 'assume';
wenzelm
parents: 60495
diff changeset
  1418
        then show "map_pmf snd pr = r"
804dfdc82835 premises in 'show' are treated like 'assume';
wenzelm
parents: 60495
diff changeset
  1419
          unfolding r q' bind_map_pmf by (subst (asm) bind_cond_pmf_cancel) (auto simp: eq_commute)
804dfdc82835 premises in 'show' are treated like 'assume';
wenzelm
parents: 60495
diff changeset
  1420
      qed (simp add: pr_def map_bind_pmf split_beta map_pmf_def[symmetric] p map_pmf_comp)
804dfdc82835 premises in 'show' are treated like 'assume';
wenzelm
parents: 60495
diff changeset
  1421
    }
804dfdc82835 premises in 'show' are treated like 'assume';
wenzelm
parents: 60495
diff changeset
  1422
    then show ?thesis
804dfdc82835 premises in 'show' are treated like 'assume';
wenzelm
parents: 60495
diff changeset
  1423
      by(auto simp add: le_fun_def)
804dfdc82835 premises in 'show' are treated like 'assume';
wenzelm
parents: 60495
diff changeset
  1424
  qed
75624
22d1c5f2b9f4 strict bounds for BNFs (by Jan van Brügge)
traytel
parents: 73253
diff changeset
  1425
qed
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1426
61634
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1427
lemma map_pmf_idI: "(\<And>x. x \<in> set_pmf p \<Longrightarrow> f x = x) \<Longrightarrow> map_pmf f p = p"
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1428
by(simp cong: pmf.map_cong)
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1429
59665
37adca7fd48f add set_pmf lemmas to simpset
hoelzl
parents: 59664
diff changeset
  1430
lemma rel_pmf_conj[simp]:
37adca7fd48f add set_pmf lemmas to simpset
hoelzl
parents: 59664
diff changeset
  1431
  "rel_pmf (\<lambda>x y. P \<and> Q x y) x y \<longleftrightarrow> P \<and> rel_pmf Q x y"
37adca7fd48f add set_pmf lemmas to simpset
hoelzl
parents: 59664
diff changeset
  1432
  "rel_pmf (\<lambda>x y. Q x y \<and> P) x y \<longleftrightarrow> P \<and> rel_pmf Q x y"
37adca7fd48f add set_pmf lemmas to simpset
hoelzl
parents: 59664
diff changeset
  1433
  using set_pmf_not_empty by (fastforce simp: pmf.in_rel subset_eq)+
37adca7fd48f add set_pmf lemmas to simpset
hoelzl
parents: 59664
diff changeset
  1434
37adca7fd48f add set_pmf lemmas to simpset
hoelzl
parents: 59664
diff changeset
  1435
lemma rel_pmf_top[simp]: "rel_pmf top = top"
37adca7fd48f add set_pmf lemmas to simpset
hoelzl
parents: 59664
diff changeset
  1436
  by (auto simp: pmf.in_rel[abs_def] fun_eq_iff map_fst_pair_pmf map_snd_pair_pmf
37adca7fd48f add set_pmf lemmas to simpset
hoelzl
parents: 59664
diff changeset
  1437
           intro: exI[of _ "pair_pmf x y" for x y])
37adca7fd48f add set_pmf lemmas to simpset
hoelzl
parents: 59664
diff changeset
  1438
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1439
lemma rel_pmf_return_pmf1: "rel_pmf R (return_pmf x) M \<longleftrightarrow> (\<forall>a\<in>M. R x a)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1440
proof safe
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1441
  fix a assume "a \<in> M" "rel_pmf R (return_pmf x) M"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1442
  then obtain pq where *: "\<And>a b. (a, b) \<in> set_pmf pq \<Longrightarrow> R a b"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1443
    and eq: "return_pmf x = map_pmf fst pq" "M = map_pmf snd pq"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1444
    by (force elim: rel_pmf.cases)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1445
  moreover have "set_pmf (return_pmf x) = {x}"
59665
37adca7fd48f add set_pmf lemmas to simpset
hoelzl
parents: 59664
diff changeset
  1446
    by simp
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61634
diff changeset
  1447
  with \<open>a \<in> M\<close> have "(x, a) \<in> pq"
59665
37adca7fd48f add set_pmf lemmas to simpset
hoelzl
parents: 59664
diff changeset
  1448
    by (force simp: eq)
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1449
  with * show "R x a"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1450
    by auto
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1451
qed (auto intro!: rel_pmf.intros[where pq="pair_pmf (return_pmf x) M"]
59665
37adca7fd48f add set_pmf lemmas to simpset
hoelzl
parents: 59664
diff changeset
  1452
          simp: map_fst_pair_pmf map_snd_pair_pmf)
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1453
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1454
lemma rel_pmf_return_pmf2: "rel_pmf R M (return_pmf x) \<longleftrightarrow> (\<forall>a\<in>M. R a x)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1455
  by (subst pmf.rel_flip[symmetric]) (simp add: rel_pmf_return_pmf1)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1456
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1457
lemma rel_return_pmf[simp]: "rel_pmf R (return_pmf x1) (return_pmf x2) = R x1 x2"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1458
  unfolding rel_pmf_return_pmf2 set_return_pmf by simp
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1459
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1460
lemma rel_pmf_False[simp]: "rel_pmf (\<lambda>x y. False) x y = False"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1461
  unfolding pmf.in_rel fun_eq_iff using set_pmf_not_empty by fastforce
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1462
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1463
lemma rel_pmf_rel_prod:
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1464
  "rel_pmf (rel_prod R S) (pair_pmf A A') (pair_pmf B B') \<longleftrightarrow> rel_pmf R A B \<and> rel_pmf S A' B'"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1465
proof safe
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1466
  assume "rel_pmf (rel_prod R S) (pair_pmf A A') (pair_pmf B B')"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1467
  then obtain pq where pq: "\<And>a b c d. ((a, c), (b, d)) \<in> set_pmf pq \<Longrightarrow> R a b \<and> S c d"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1468
    and eq: "map_pmf fst pq = pair_pmf A A'" "map_pmf snd pq = pair_pmf B B'"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1469
    by (force elim: rel_pmf.cases)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1470
  show "rel_pmf R A B"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1471
  proof (rule rel_pmf.intros)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1472
    let ?f = "\<lambda>(a, b). (fst a, fst b)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1473
    have [simp]: "(\<lambda>x. fst (?f x)) = fst o fst" "(\<lambda>x. snd (?f x)) = fst o snd"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1474
      by auto
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1475
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1476
    show "map_pmf fst (map_pmf ?f pq) = A"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1477
      by (simp add: map_pmf_comp pmf.map_comp[symmetric] eq map_fst_pair_pmf)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1478
    show "map_pmf snd (map_pmf ?f pq) = B"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1479
      by (simp add: map_pmf_comp pmf.map_comp[symmetric] eq map_fst_pair_pmf)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1480
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1481
    fix a b assume "(a, b) \<in> set_pmf (map_pmf ?f pq)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1482
    then obtain c d where "((a, c), (b, d)) \<in> set_pmf pq"
59665
37adca7fd48f add set_pmf lemmas to simpset
hoelzl
parents: 59664
diff changeset
  1483
      by auto
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1484
    from pq[OF this] show "R a b" ..
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1485
  qed
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1486
  show "rel_pmf S A' B'"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1487
  proof (rule rel_pmf.intros)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1488
    let ?f = "\<lambda>(a, b). (snd a, snd b)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1489
    have [simp]: "(\<lambda>x. fst (?f x)) = snd o fst" "(\<lambda>x. snd (?f x)) = snd o snd"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1490
      by auto
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1491
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1492
    show "map_pmf fst (map_pmf ?f pq) = A'"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1493
      by (simp add: map_pmf_comp pmf.map_comp[symmetric] eq map_snd_pair_pmf)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1494
    show "map_pmf snd (map_pmf ?f pq) = B'"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1495
      by (simp add: map_pmf_comp pmf.map_comp[symmetric] eq map_snd_pair_pmf)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1496
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1497
    fix c d assume "(c, d) \<in> set_pmf (map_pmf ?f pq)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1498
    then obtain a b where "((a, c), (b, d)) \<in> set_pmf pq"
59665
37adca7fd48f add set_pmf lemmas to simpset
hoelzl
parents: 59664
diff changeset
  1499
      by auto
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1500
    from pq[OF this] show "S c d" ..
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1501
  qed
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1502
next
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1503
  assume "rel_pmf R A B" "rel_pmf S A' B'"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1504
  then obtain Rpq Spq
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1505
    where Rpq: "\<And>a b. (a, b) \<in> set_pmf Rpq \<Longrightarrow> R a b"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1506
        "map_pmf fst Rpq = A" "map_pmf snd Rpq = B"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1507
      and Spq: "\<And>a b. (a, b) \<in> set_pmf Spq \<Longrightarrow> S a b"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1508
        "map_pmf fst Spq = A'" "map_pmf snd Spq = B'"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1509
    by (force elim: rel_pmf.cases)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1510
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1511
  let ?f = "(\<lambda>((a, c), (b, d)). ((a, b), (c, d)))"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1512
  let ?pq = "map_pmf ?f (pair_pmf Rpq Spq)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1513
  have [simp]: "(\<lambda>x. fst (?f x)) = (\<lambda>(a, b). (fst a, fst b))" "(\<lambda>x. snd (?f x)) = (\<lambda>(a, b). (snd a, snd b))"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1514
    by auto
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1515
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1516
  show "rel_pmf (rel_prod R S) (pair_pmf A A') (pair_pmf B B')"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1517
    by (rule rel_pmf.intros[where pq="?pq"])
59665
37adca7fd48f add set_pmf lemmas to simpset
hoelzl
parents: 59664
diff changeset
  1518
       (auto simp: map_snd_pair_pmf map_fst_pair_pmf map_pmf_comp Rpq Spq
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1519
                   map_pair)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1520
qed
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1521
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59665
diff changeset
  1522
lemma rel_pmf_reflI:
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1523
  assumes "\<And>x. x \<in> set_pmf p \<Longrightarrow> P x x"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1524
  shows "rel_pmf P p p"
59665
37adca7fd48f add set_pmf lemmas to simpset
hoelzl
parents: 59664
diff changeset
  1525
  by (rule rel_pmf.intros[where pq="map_pmf (\<lambda>x. (x, x)) p"])
37adca7fd48f add set_pmf lemmas to simpset
hoelzl
parents: 59664
diff changeset
  1526
     (auto simp add: pmf.map_comp o_def assms)
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1527
61634
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1528
lemma rel_pmf_bij_betw:
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1529
  assumes f: "bij_betw f (set_pmf p) (set_pmf q)"
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1530
  and eq: "\<And>x. x \<in> set_pmf p \<Longrightarrow> pmf p x = pmf q (f x)"
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1531
  shows "rel_pmf (\<lambda>x y. f x = y) p q"
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1532
proof(rule rel_pmf.intros)
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1533
  let ?pq = "map_pmf (\<lambda>x. (x, f x)) p"
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1534
  show "map_pmf fst ?pq = p" by(simp add: pmf.map_comp o_def)
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1535
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1536
  have "map_pmf f p = q"
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1537
  proof(rule pmf_eqI)
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1538
    fix i
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1539
    show "pmf (map_pmf f p) i = pmf q i"
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1540
    proof(cases "i \<in> set_pmf q")
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1541
      case True
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1542
      with f obtain j where "i = f j" "j \<in> set_pmf p"
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1543
        by(auto simp add: bij_betw_def image_iff)
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1544
      thus ?thesis using f by(simp add: bij_betw_def pmf_map_inj eq)
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1545
    next
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1546
      case False thus ?thesis
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1547
        by(subst pmf_map_outside)(auto simp add: set_pmf_iff eq[symmetric])
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1548
    qed
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1549
  qed
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1550
  then show "map_pmf snd ?pq = q" by(simp add: pmf.map_comp o_def)
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1551
qed auto
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1552
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1553
context
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1554
begin
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1555
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1556
interpretation pmf_as_measure .
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1557
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1558
definition "join_pmf M = bind_pmf M (\<lambda>x. x)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1559
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1560
lemma bind_eq_join_pmf: "bind_pmf M f = join_pmf (map_pmf f M)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1561
  unfolding join_pmf_def bind_map_pmf ..
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1562
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1563
lemma join_eq_bind_pmf: "join_pmf M = bind_pmf M id"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1564
  by (simp add: join_pmf_def id_def)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1565
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1566
lemma pmf_join: "pmf (join_pmf N) i = (\<integral>M. pmf M i \<partial>measure_pmf N)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1567
  unfolding join_pmf_def pmf_bind ..
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1568
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1569
lemma ennreal_pmf_join: "ennreal (pmf (join_pmf N) i) = (\<integral>\<^sup>+M. pmf M i \<partial>measure_pmf N)"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1570
  unfolding join_pmf_def ennreal_pmf_bind ..
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1571
59665
37adca7fd48f add set_pmf lemmas to simpset
hoelzl
parents: 59664
diff changeset
  1572
lemma set_pmf_join_pmf[simp]: "set_pmf (join_pmf f) = (\<Union>p\<in>set_pmf f. set_pmf p)"
37adca7fd48f add set_pmf lemmas to simpset
hoelzl
parents: 59664
diff changeset
  1573
  by (simp add: join_pmf_def)
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1574
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1575
lemma join_return_pmf: "join_pmf (return_pmf M) = M"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1576
  by (simp add: integral_return pmf_eq_iff pmf_join return_pmf.rep_eq)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1577
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1578
lemma map_join_pmf: "map_pmf f (join_pmf AA) = join_pmf (map_pmf (map_pmf f) AA)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1579
  by (simp add: join_pmf_def map_pmf_def bind_assoc_pmf bind_return_pmf)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1580
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1581
lemma join_map_return_pmf: "join_pmf (map_pmf return_pmf A) = A"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1582
  by (simp add: join_pmf_def map_pmf_def bind_assoc_pmf bind_return_pmf bind_return_pmf')
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1583
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1584
end
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1585
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1586
lemma rel_pmf_joinI:
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1587
  assumes "rel_pmf (rel_pmf P) p q"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1588
  shows "rel_pmf P (join_pmf p) (join_pmf q)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1589
proof -
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1590
  from assms obtain pq where p: "p = map_pmf fst pq"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1591
    and q: "q = map_pmf snd pq"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1592
    and P: "\<And>x y. (x, y) \<in> set_pmf pq \<Longrightarrow> rel_pmf P x y"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1593
    by cases auto
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59665
diff changeset
  1594
  from P obtain PQ
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1595
    where PQ: "\<And>x y a b. \<lbrakk> (x, y) \<in> set_pmf pq; (a, b) \<in> set_pmf (PQ x y) \<rbrakk> \<Longrightarrow> P a b"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1596
    and x: "\<And>x y. (x, y) \<in> set_pmf pq \<Longrightarrow> map_pmf fst (PQ x y) = x"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1597
    and y: "\<And>x y. (x, y) \<in> set_pmf pq \<Longrightarrow> map_pmf snd (PQ x y) = y"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1598
    by(metis rel_pmf.simps)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1599
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1600
  let ?r = "bind_pmf pq (\<lambda>(x, y). PQ x y)"
59665
37adca7fd48f add set_pmf lemmas to simpset
hoelzl
parents: 59664
diff changeset
  1601
  have "\<And>a b. (a, b) \<in> set_pmf ?r \<Longrightarrow> P a b" by (auto intro: PQ)
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1602
  moreover have "map_pmf fst ?r = join_pmf p" "map_pmf snd ?r = join_pmf q"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1603
    by (simp_all add: p q x y join_pmf_def map_bind_pmf bind_map_pmf split_def cong: bind_pmf_cong)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1604
  ultimately show ?thesis ..
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1605
qed
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1606
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1607
lemma rel_pmf_bindI:
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1608
  assumes pq: "rel_pmf R p q"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1609
  and fg: "\<And>x y. R x y \<Longrightarrow> rel_pmf P (f x) (g y)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1610
  shows "rel_pmf P (bind_pmf p f) (bind_pmf q g)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1611
  unfolding bind_eq_join_pmf
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1612
  by (rule rel_pmf_joinI)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1613
     (auto simp add: pmf.rel_map intro: pmf.rel_mono[THEN le_funD, THEN le_funD, THEN le_boolD, THEN mp, OF _ pq] fg)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1614
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61634
diff changeset
  1615
text \<open>
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69529
diff changeset
  1616
  Proof that \<^const>\<open>rel_pmf\<close> preserves orders.
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59665
diff changeset
  1617
  Antisymmetry proof follows Thm. 1 in N. Saheb-Djahromi, Cpo's of measures for nondeterminism,
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59665
diff changeset
  1618
  Theoretical Computer Science 12(1):19--37, 1980,
67601
b34be3010273 use preferred resolver according to DOI Handbook §3.8
Lars Hupel <lars.hupel@mytum.de>
parents: 67489
diff changeset
  1619
  \<^url>\<open>https://doi.org/10.1016/0304-3975(80)90003-1\<close>
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61634
diff changeset
  1620
\<close>
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1621
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59665
diff changeset
  1622
lemma
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1623
  assumes *: "rel_pmf R p q"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1624
  and refl: "reflp R" and trans: "transp R"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1625
  shows measure_Ici: "measure p {y. R x y} \<le> measure q {y. R x y}" (is ?thesis1)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1626
  and measure_Ioi: "measure p {y. R x y \<and> \<not> R y x} \<le> measure q {y. R x y \<and> \<not> R y x}" (is ?thesis2)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1627
proof -
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1628
  from * obtain pq
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1629
    where pq: "\<And>x y. (x, y) \<in> set_pmf pq \<Longrightarrow> R x y"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1630
    and p: "p = map_pmf fst pq"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1631
    and q: "q = map_pmf snd pq"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1632
    by cases auto
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1633
  show ?thesis1 ?thesis2 unfolding p q map_pmf_rep_eq using refl trans
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1634
    by(auto 4 3 simp add: measure_distr reflpD AE_measure_pmf_iff intro!: measure_pmf.finite_measure_mono_AE dest!: pq elim: transpE)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1635
qed
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1636
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1637
lemma rel_pmf_inf:
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1638
  fixes p q :: "'a pmf"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1639
  assumes 1: "rel_pmf R p q"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1640
  assumes 2: "rel_pmf R q p"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1641
  and refl: "reflp R" and trans: "transp R"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1642
  shows "rel_pmf (inf R R\<inverse>\<inverse>) p q"
59681
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1643
proof (subst rel_pmf_iff_equivp, safe)
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1644
  show "equivp (inf R R\<inverse>\<inverse>)"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1645
    using trans refl by (auto simp: equivp_reflp_symp_transp intro: sympI transpI reflpI dest: transpD reflpD)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1646
59681
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1647
  fix C assume "C \<in> UNIV // {(x, y). inf R R\<inverse>\<inverse> x y}"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1648
  then obtain x where C: "C = {y. R x y \<and> R y x}"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1649
    by (auto elim: quotientE)
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1650
59670
dee043d19729 generalized bind_cond_pmf_cancel
hoelzl
parents: 59667
diff changeset
  1651
  let ?R = "\<lambda>x y. R x y \<and> R y x"
dee043d19729 generalized bind_cond_pmf_cancel
hoelzl
parents: 59667
diff changeset
  1652
  let ?\<mu>R = "\<lambda>y. measure q {x. ?R x y}"
59681
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1653
  have "measure p {y. ?R x y} = measure p ({y. R x y} - {y. R x y \<and> \<not> R y x})"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1654
    by(auto intro!: arg_cong[where f="measure p"])
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1655
  also have "\<dots> = measure p {y. R x y} - measure p {y. R x y \<and> \<not> R y x}"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1656
    by (rule measure_pmf.finite_measure_Diff) auto
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1657
  also have "measure p {y. R x y \<and> \<not> R y x} = measure q {y. R x y \<and> \<not> R y x}"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1658
    using 1 2 refl trans by(auto intro!: Orderings.antisym measure_Ioi)
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1659
  also have "measure p {y. R x y} = measure q {y. R x y}"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1660
    using 1 2 refl trans by(auto intro!: Orderings.antisym measure_Ici)
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1661
  also have "measure q {y. R x y} - measure q {y. R x y \<and> \<not> R y x} =
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1662
    measure q ({y. R x y} - {y. R x y \<and> \<not> R y x})"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1663
    by(rule measure_pmf.finite_measure_Diff[symmetric]) auto
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1664
  also have "\<dots> = ?\<mu>R x"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1665
    by(auto intro!: arg_cong[where f="measure q"])
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1666
  finally show "measure p C = measure q C"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1667
    by (simp add: C conj_commute)
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1668
qed
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1669
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1670
lemma rel_pmf_antisym:
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1671
  fixes p q :: "'a pmf"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1672
  assumes 1: "rel_pmf R p q"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1673
  assumes 2: "rel_pmf R q p"
64634
5bd30359e46e proper logical constants
haftmann
parents: 64267
diff changeset
  1674
  and refl: "reflp R" and trans: "transp R" and antisym: "antisymp R"
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1675
  shows "p = q"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1676
proof -
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1677
  from 1 2 refl trans have "rel_pmf (inf R R\<inverse>\<inverse>) p q" by(rule rel_pmf_inf)
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67226
diff changeset
  1678
  also have "inf R R\<inverse>\<inverse> = (=)"
64634
5bd30359e46e proper logical constants
haftmann
parents: 64267
diff changeset
  1679
    using refl antisym by (auto intro!: ext simp add: reflpD dest: antisympD)
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1680
  finally show ?thesis unfolding pmf.rel_eq .
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1681
qed
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1682
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1683
lemma reflp_rel_pmf: "reflp R \<Longrightarrow> reflp (rel_pmf R)"
64634
5bd30359e46e proper logical constants
haftmann
parents: 64267
diff changeset
  1684
  by (fact pmf.rel_reflp)
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1685
64634
5bd30359e46e proper logical constants
haftmann
parents: 64267
diff changeset
  1686
lemma antisymp_rel_pmf:
5bd30359e46e proper logical constants
haftmann
parents: 64267
diff changeset
  1687
  "\<lbrakk> reflp R; transp R; antisymp R \<rbrakk>
5bd30359e46e proper logical constants
haftmann
parents: 64267
diff changeset
  1688
  \<Longrightarrow> antisymp (rel_pmf R)"
5bd30359e46e proper logical constants
haftmann
parents: 64267
diff changeset
  1689
by(rule antisympI)(blast intro: rel_pmf_antisym)
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1690
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1691
lemma transp_rel_pmf:
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1692
  assumes "transp R"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1693
  shows "transp (rel_pmf R)"
64634
5bd30359e46e proper logical constants
haftmann
parents: 64267
diff changeset
  1694
  using assms by (fact pmf.rel_transp)
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1695
67486
Lars Hupel <lars.hupel@mytum.de>
parents: 67399
diff changeset
  1696
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1697
subsection \<open> Distributions \<close>
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1698
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1699
context
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1700
begin
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1701
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1702
interpretation pmf_as_function .
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1703
59093
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  1704
subsubsection \<open> Bernoulli Distribution \<close>
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  1705
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1706
lift_definition bernoulli_pmf :: "real \<Rightarrow> bool pmf" is
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1707
  "\<lambda>p b. ((\<lambda>p. if b then p else 1 - p) \<circ> min 1 \<circ> max 0) p"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1708
  by (auto simp: nn_integral_count_space_finite[where A="{False, True}"] UNIV_bool
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1709
           split: split_max split_min)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1710
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1711
lemma pmf_bernoulli_True[simp]: "0 \<le> p \<Longrightarrow> p \<le> 1 \<Longrightarrow> pmf (bernoulli_pmf p) True = p"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1712
  by transfer simp
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1713
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1714
lemma pmf_bernoulli_False[simp]: "0 \<le> p \<Longrightarrow> p \<le> 1 \<Longrightarrow> pmf (bernoulli_pmf p) False = 1 - p"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1715
  by transfer simp
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1716
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1717
lemma set_pmf_bernoulli[simp]: "0 < p \<Longrightarrow> p < 1 \<Longrightarrow> set_pmf (bernoulli_pmf p) = UNIV"
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1718
  by (auto simp add: set_pmf_iff UNIV_bool)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1719
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59665
diff changeset
  1720
lemma nn_integral_bernoulli_pmf[simp]:
59002
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
  1721
  assumes [simp]: "0 \<le> p" "p \<le> 1" "\<And>x. 0 \<le> f x"
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
  1722
  shows "(\<integral>\<^sup>+x. f x \<partial>bernoulli_pmf p) = f True * p + f False * (1 - p)"
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
  1723
  by (subst nn_integral_measure_pmf_support[of UNIV])
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
  1724
     (auto simp: UNIV_bool field_simps)
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
  1725
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59665
diff changeset
  1726
lemma integral_bernoulli_pmf[simp]:
59002
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
  1727
  assumes [simp]: "0 \<le> p" "p \<le> 1"
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
  1728
  shows "(\<integral>x. f x \<partial>bernoulli_pmf p) = f True * p + f False * (1 - p)"
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
  1729
  by (subst integral_measure_pmf[of UNIV]) (auto simp: UNIV_bool)
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
  1730
59525
dfe6449aecd8 more lemmas
Andreas Lochbihler
parents: 59496
diff changeset
  1731
lemma pmf_bernoulli_half [simp]: "pmf (bernoulli_pmf (1 / 2)) x = 1 / 2"
dfe6449aecd8 more lemmas
Andreas Lochbihler
parents: 59496
diff changeset
  1732
by(cases x) simp_all
dfe6449aecd8 more lemmas
Andreas Lochbihler
parents: 59496
diff changeset
  1733
dfe6449aecd8 more lemmas
Andreas Lochbihler
parents: 59496
diff changeset
  1734
lemma measure_pmf_bernoulli_half: "measure_pmf (bernoulli_pmf (1 / 2)) = uniform_count_measure UNIV"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1735
  by (rule measure_eqI)
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1736
     (simp_all add: nn_integral_pmf[symmetric] emeasure_uniform_count_measure ennreal_divide_numeral[symmetric]
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1737
                    nn_integral_count_space_finite sets_uniform_count_measure divide_ennreal_def mult_ac
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1738
                    ennreal_of_nat_eq_real_of_nat)
59525
dfe6449aecd8 more lemmas
Andreas Lochbihler
parents: 59496
diff changeset
  1739
59093
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  1740
subsubsection \<open> Geometric Distribution \<close>
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  1741
60602
37588fbe39f9 generalized geometric distribution
hoelzl
parents: 60595
diff changeset
  1742
context
37588fbe39f9 generalized geometric distribution
hoelzl
parents: 60595
diff changeset
  1743
  fixes p :: real assumes p[arith]: "0 < p" "p \<le> 1"
37588fbe39f9 generalized geometric distribution
hoelzl
parents: 60595
diff changeset
  1744
begin
37588fbe39f9 generalized geometric distribution
hoelzl
parents: 60595
diff changeset
  1745
37588fbe39f9 generalized geometric distribution
hoelzl
parents: 60595
diff changeset
  1746
lift_definition geometric_pmf :: "nat pmf" is "\<lambda>n. (1 - p)^n * p"
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1747
proof
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1748
  have "(\<Sum>i. ennreal (p * (1 - p) ^ i)) = ennreal (p * (1 / (1 - (1 - p))))"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1749
    by (intro suminf_ennreal_eq sums_mult geometric_sums) auto
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1750
  then show "(\<integral>\<^sup>+ x. ennreal ((1 - p)^x * p) \<partial>count_space UNIV) = 1"
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1751
    by (simp add: nn_integral_count_space_nat field_simps)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1752
qed simp
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1753
60602
37588fbe39f9 generalized geometric distribution
hoelzl
parents: 60595
diff changeset
  1754
lemma pmf_geometric[simp]: "pmf geometric_pmf n = (1 - p)^n * p"
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1755
  by transfer rule
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1756
60602
37588fbe39f9 generalized geometric distribution
hoelzl
parents: 60595
diff changeset
  1757
end
37588fbe39f9 generalized geometric distribution
hoelzl
parents: 60595
diff changeset
  1758
73253
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1759
lemma geometric_pmf_1 [simp]: "geometric_pmf 1 = return_pmf 0"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1760
  by (intro pmf_eqI) (auto simp: indicator_def)
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1761
60602
37588fbe39f9 generalized geometric distribution
hoelzl
parents: 60595
diff changeset
  1762
lemma set_pmf_geometric: "0 < p \<Longrightarrow> p < 1 \<Longrightarrow> set_pmf (geometric_pmf p) = UNIV"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1763
  by (auto simp: set_pmf_iff)
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1764
73253
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1765
lemma geometric_sums_times_n:
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1766
  fixes c::"'a::{banach,real_normed_field}"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1767
  assumes "norm c < 1"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1768
  shows "(\<lambda>n. c^n * of_nat n) sums (c / (1 - c)\<^sup>2)"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1769
proof -
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1770
  have "(\<lambda>n. c * z ^ n) sums (c / (1 - z))" if "norm z < 1" for z
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1771
    using geometric_sums sums_mult that by fastforce
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1772
  moreover have "((\<lambda>z. c / (1 - z)) has_field_derivative (c / (1 - c)\<^sup>2)) (at c)"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1773
      using assms by (auto intro!: derivative_eq_intros simp add: semiring_normalization_rules)
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1774
  ultimately have "(\<lambda>n. diffs (\<lambda>n. c) n * c ^ n) sums (c / (1 - c)\<^sup>2)"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1775
    using assms by (intro termdiffs_sums_strong)
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1776
  then have "(\<lambda>n. of_nat (Suc n) * c ^ (Suc n)) sums (c / (1 - c)\<^sup>2)"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1777
    unfolding diffs_def by (simp add: power_eq_if mult.assoc)
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1778
  then show ?thesis
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1779
    by (subst (asm) sums_Suc_iff) (auto simp add: mult.commute)
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1780
qed
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1781
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1782
lemma geometric_sums_times_norm:
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1783
  fixes c::"'a::{banach,real_normed_field}"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1784
  assumes "norm c < 1"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1785
  shows "(\<lambda>n. norm (c^n * of_nat n)) sums (norm c / (1 - norm c)\<^sup>2)"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1786
proof -
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1787
  have "norm (c^n * of_nat n) = (norm c) ^ n * of_nat n" for n::nat
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1788
    by (simp add: norm_power norm_mult)
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1789
  then show ?thesis
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1790
    using geometric_sums_times_n[of "norm c"] assms
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1791
    by force
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1792
qed
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1793
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1794
lemma integrable_real_geometric_pmf:
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1795
  assumes "p \<in> {0<..1}"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1796
  shows   "integrable (geometric_pmf p) real"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1797
proof -
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1798
  have "summable (\<lambda>x. p * ((1 - p) ^ x * real x))"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1799
    using geometric_sums_times_norm[of "1 - p"] assms
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1800
    by (intro summable_mult) (auto simp: sums_iff)
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1801
  hence "summable (\<lambda>x. (1 - p) ^ x * real x)"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1802
    by (rule summable_mult_D) (use assms in auto)
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1803
  thus ?thesis
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1804
    unfolding measure_pmf_eq_density using assms
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1805
    by (subst integrable_density)
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1806
       (auto simp: integrable_count_space_nat_iff mult_ac)
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1807
qed
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1808
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1809
lemma expectation_geometric_pmf:
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1810
  assumes "p \<in> {0<..1}"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1811
  shows   "measure_pmf.expectation (geometric_pmf p) real = (1 - p) / p"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1812
proof -
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1813
  have "(\<lambda>n. p * ((1 - p) ^ n * n)) sums (p * ((1 - p) / p^2))"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1814
    using assms geometric_sums_times_n[of "1-p"] by (intro sums_mult) auto
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1815
  moreover have "(\<lambda>n. p * ((1 - p) ^ n * n)) = (\<lambda>n. (1 - p) ^ n * p  * real n)"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1816
    by auto
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1817
  ultimately have *: "(\<lambda>n. (1 - p) ^ n * p  * real n) sums ((1 - p) / p)"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1818
    using assms sums_subst by (auto simp add: power2_eq_square)
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1819
  have "measure_pmf.expectation (geometric_pmf p) real =
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1820
        (\<integral>n. pmf (geometric_pmf p) n * real n \<partial>count_space UNIV)"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1821
    unfolding measure_pmf_eq_density by (subst integral_density) auto
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1822
  also have "integrable (count_space UNIV) (\<lambda>n. pmf (geometric_pmf p) n * real n)"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1823
    using * assms unfolding integrable_count_space_nat_iff by (simp add: sums_iff)
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1824
  hence "(\<integral>n. pmf (geometric_pmf p) n * real n \<partial>count_space UNIV) = (1 - p) / p"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1825
    using * assms by (subst integral_count_space_nat) (simp_all add: sums_iff)
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1826
  finally show ?thesis by auto
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1827
qed
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1828
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1829
lemma geometric_bind_pmf_unfold:
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1830
  assumes "p \<in> {0<..1}"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1831
  shows "geometric_pmf p =
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1832
     do {b \<leftarrow> bernoulli_pmf p;
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1833
         if b then return_pmf 0 else map_pmf Suc (geometric_pmf p)}"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1834
proof -
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1835
  have *: "(Suc -` {i}) = (if i = 0 then {} else {i - 1})" for i
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1836
    by force
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1837
  have "pmf (geometric_pmf p) i =
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1838
        pmf (bernoulli_pmf p \<bind>
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1839
            (\<lambda>b. if b then return_pmf 0 else map_pmf Suc (geometric_pmf p)))
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1840
            i" for i
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1841
  proof -
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1842
    have "pmf (geometric_pmf p) i =
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1843
          (if i = 0 then p else (1 - p) * pmf (geometric_pmf p) (i - 1))"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1844
      using assms by (simp add: power_eq_if)
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1845
    also have "\<dots> = (if i = 0  then p else (1 - p) * pmf (map_pmf Suc (geometric_pmf p)) i)"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1846
      by (simp add: pmf_map indicator_def measure_pmf_single *)
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1847
    also have "\<dots> = measure_pmf.expectation (bernoulli_pmf p)
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1848
          (\<lambda>x. pmf (if x then return_pmf 0 else map_pmf Suc (geometric_pmf p)) i)"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1849
      using assms by (auto simp add: pmf_map *)
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1850
    also have "\<dots> = pmf (bernoulli_pmf p \<bind>
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1851
                   (\<lambda>b. if b then return_pmf 0 else map_pmf Suc (geometric_pmf p)))
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1852
                   i"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1853
      by (auto simp add: pmf_bind)
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1854
    finally show ?thesis .
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1855
  qed
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1856
  then show ?thesis
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1857
    using pmf_eqI by blast
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1858
qed
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1859
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  1860
59093
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  1861
subsubsection \<open> Uniform Multiset Distribution \<close>
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  1862
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1863
context
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1864
  fixes M :: "'a multiset" assumes M_not_empty: "M \<noteq> {#}"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1865
begin
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1866
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1867
lift_definition pmf_of_multiset :: "'a pmf" is "\<lambda>x. count M x / size M"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1868
proof
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1869
  show "(\<integral>\<^sup>+ x. ennreal (real (count M x) / real (size M)) \<partial>count_space UNIV) = 1"
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1870
    using M_not_empty
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1871
    by (simp add: zero_less_divide_iff nn_integral_count_space nonempty_has_size
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64008
diff changeset
  1872
                  sum_divide_distrib[symmetric])
b9a1486e79be setsum -> sum
nipkow
parents: 64008
diff changeset
  1873
       (auto simp: size_multiset_overloaded_eq intro!: sum.cong)
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1874
qed simp
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1875
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1876
lemma pmf_of_multiset[simp]: "pmf pmf_of_multiset x = count M x / size M"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1877
  by transfer rule
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1878
60495
d7ff0a1df90a renamed Multiset.set_of to the canonical set_mset
nipkow
parents: 60068
diff changeset
  1879
lemma set_pmf_of_multiset[simp]: "set_pmf pmf_of_multiset = set_mset M"
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1880
  by (auto simp: set_pmf_iff)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1881
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1882
end
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1883
59093
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  1884
subsubsection \<open> Uniform Distribution \<close>
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  1885
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1886
context
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1887
  fixes S :: "'a set" assumes S_not_empty: "S \<noteq> {}" and S_finite: "finite S"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1888
begin
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1889
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1890
lift_definition pmf_of_set :: "'a pmf" is "\<lambda>x. indicator S x / card S"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1891
proof
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1892
  show "(\<integral>\<^sup>+ x. ennreal (indicator S x / real (card S)) \<partial>count_space UNIV) = 1"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1893
    using S_not_empty S_finite
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1894
    by (subst nn_integral_count_space'[of S])
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1895
       (auto simp: ennreal_of_nat_eq_real_of_nat ennreal_mult[symmetric])
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1896
qed simp
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1897
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1898
lemma pmf_of_set[simp]: "pmf pmf_of_set x = indicator S x / card S"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1899
  by transfer rule
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1900
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1901
lemma set_pmf_of_set[simp]: "set_pmf pmf_of_set = S"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1902
  using S_finite S_not_empty by (auto simp: set_pmf_iff)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1903
61634
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1904
lemma emeasure_pmf_of_set_space[simp]: "emeasure pmf_of_set S = 1"
59002
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
  1905
  by (rule measure_pmf.emeasure_eq_1_AE) (auto simp: AE_measure_pmf_iff)
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
  1906
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64008
diff changeset
  1907
lemma nn_integral_pmf_of_set: "nn_integral (measure_pmf pmf_of_set) f = sum f S / card S"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1908
  by (subst nn_integral_measure_pmf_finite)
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64008
diff changeset
  1909
     (simp_all add: sum_distrib_right[symmetric] card_gt_0_iff S_not_empty S_finite divide_ennreal_def
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1910
                divide_ennreal[symmetric] ennreal_of_nat_eq_real_of_nat[symmetric] ennreal_times_divide)
60068
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
  1911
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64008
diff changeset
  1912
lemma integral_pmf_of_set: "integral\<^sup>L (measure_pmf pmf_of_set) f = sum f S / card S"
b9a1486e79be setsum -> sum
nipkow
parents: 64008
diff changeset
  1913
  by (subst integral_measure_pmf[of S]) (auto simp: S_finite sum_divide_distrib)
60068
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
  1914
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1915
lemma emeasure_pmf_of_set: "emeasure (measure_pmf pmf_of_set) A = card (S \<inter> A) / card S"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1916
  by (subst nn_integral_indicator[symmetric], simp)
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64008
diff changeset
  1917
     (simp add: S_finite S_not_empty card_gt_0_iff indicator_def sum.If_cases divide_ennreal
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1918
                ennreal_of_nat_eq_real_of_nat nn_integral_pmf_of_set)
60068
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
  1919
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1920
lemma measure_pmf_of_set: "measure (measure_pmf pmf_of_set) A = card (S \<inter> A) / card S"
63092
a949b2a5f51d eliminated use of empty "assms";
wenzelm
parents: 63040
diff changeset
  1921
  using emeasure_pmf_of_set[of A]
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1922
  by (simp add: measure_nonneg measure_pmf.emeasure_eq_measure)
61634
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  1923
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1924
end
67486
Lars Hupel <lars.hupel@mytum.de>
parents: 67399
diff changeset
  1925
65395
7504569a73c7 moved material from AFP to distribution
eberlm <eberlm@in.tum.de>
parents: 64634
diff changeset
  1926
lemma pmf_expectation_bind_pmf_of_set:
7504569a73c7 moved material from AFP to distribution
eberlm <eberlm@in.tum.de>
parents: 64634
diff changeset
  1927
  fixes A :: "'a set" and f :: "'a \<Rightarrow> 'b pmf"
7504569a73c7 moved material from AFP to distribution
eberlm <eberlm@in.tum.de>
parents: 64634
diff changeset
  1928
    and  h :: "'b \<Rightarrow> 'c::{banach, second_countable_topology}"
7504569a73c7 moved material from AFP to distribution
eberlm <eberlm@in.tum.de>
parents: 64634
diff changeset
  1929
  assumes "A \<noteq> {}" "finite A" "\<And>x. x \<in> A \<Longrightarrow> finite (set_pmf (f x))"
7504569a73c7 moved material from AFP to distribution
eberlm <eberlm@in.tum.de>
parents: 64634
diff changeset
  1930
  shows "measure_pmf.expectation (pmf_of_set A \<bind> f) h =
7504569a73c7 moved material from AFP to distribution
eberlm <eberlm@in.tum.de>
parents: 64634
diff changeset
  1931
           (\<Sum>a\<in>A. measure_pmf.expectation (f a) h /\<^sub>R real (card A))"
70817
dd675800469d dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents: 70532
diff changeset
  1932
  using assms by (subst pmf_expectation_bind[of A]) (auto simp: field_split_simps)
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1933
63099
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  1934
lemma map_pmf_of_set:
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  1935
  assumes "finite A" "A \<noteq> {}"
63886
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63882
diff changeset
  1936
  shows   "map_pmf f (pmf_of_set A) = pmf_of_multiset (image_mset f (mset_set A))"
63099
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  1937
    (is "?lhs = ?rhs")
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  1938
proof (intro pmf_eqI)
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  1939
  fix x
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  1940
  from assms have "ennreal (pmf ?lhs x) = ennreal (pmf ?rhs x)"
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  1941
    by (subst ennreal_pmf_map)
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  1942
       (simp_all add: emeasure_pmf_of_set mset_set_empty_iff count_image_mset Int_commute)
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  1943
  thus "pmf ?lhs x = pmf ?rhs x" by simp
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  1944
qed
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  1945
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  1946
lemma pmf_bind_pmf_of_set:
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  1947
  assumes "A \<noteq> {}" "finite A"
63886
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63882
diff changeset
  1948
  shows   "pmf (bind_pmf (pmf_of_set A) f) x =
63099
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  1949
             (\<Sum>xa\<in>A. pmf (f xa) x) / real_of_nat (card A)" (is "?lhs = ?rhs")
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  1950
proof -
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  1951
  from assms have "card A > 0" by auto
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  1952
  with assms have "ennreal ?lhs = ennreal ?rhs"
63886
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63882
diff changeset
  1953
    by (subst ennreal_pmf_bind)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63882
diff changeset
  1954
       (simp_all add: nn_integral_pmf_of_set max_def pmf_nonneg divide_ennreal [symmetric]
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64008
diff changeset
  1955
        sum_nonneg ennreal_of_nat_eq_real_of_nat)
b9a1486e79be setsum -> sum
nipkow
parents: 64008
diff changeset
  1956
  thus ?thesis by (subst (asm) ennreal_inj) (auto intro!: sum_nonneg divide_nonneg_nonneg)
63099
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  1957
qed
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  1958
60068
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
  1959
lemma pmf_of_set_singleton: "pmf_of_set {x} = return_pmf x"
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
  1960
by(rule pmf_eqI)(simp add: indicator_def)
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
  1961
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1962
lemma map_pmf_of_set_inj:
60068
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
  1963
  assumes f: "inj_on f A"
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
  1964
  and [simp]: "A \<noteq> {}" "finite A"
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
  1965
  shows "map_pmf f (pmf_of_set A) = pmf_of_set (f ` A)" (is "?lhs = ?rhs")
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
  1966
proof(rule pmf_eqI)
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
  1967
  fix i
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
  1968
  show "pmf ?lhs i = pmf ?rhs i"
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
  1969
  proof(cases "i \<in> f ` A")
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
  1970
    case True
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
  1971
    then obtain i' where "i = f i'" "i' \<in> A" by auto
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
  1972
    thus ?thesis using f by(simp add: card_image pmf_map_inj)
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
  1973
  next
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
  1974
    case False
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
  1975
    hence "pmf ?lhs i = 0" by(simp add: pmf_eq_0_set_pmf set_map_pmf)
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
  1976
    moreover have "pmf ?rhs i = 0" using False by simp
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
  1977
    ultimately show ?thesis by simp
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
  1978
  qed
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
  1979
qed
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
  1980
65395
7504569a73c7 moved material from AFP to distribution
eberlm <eberlm@in.tum.de>
parents: 64634
diff changeset
  1981
lemma map_pmf_of_set_bij_betw:
7504569a73c7 moved material from AFP to distribution
eberlm <eberlm@in.tum.de>
parents: 64634
diff changeset
  1982
  assumes "bij_betw f A B" "A \<noteq> {}" "finite A"
7504569a73c7 moved material from AFP to distribution
eberlm <eberlm@in.tum.de>
parents: 64634
diff changeset
  1983
  shows   "map_pmf f (pmf_of_set A) = pmf_of_set B"
7504569a73c7 moved material from AFP to distribution
eberlm <eberlm@in.tum.de>
parents: 64634
diff changeset
  1984
proof -
7504569a73c7 moved material from AFP to distribution
eberlm <eberlm@in.tum.de>
parents: 64634
diff changeset
  1985
  have "map_pmf f (pmf_of_set A) = pmf_of_set (f ` A)"
7504569a73c7 moved material from AFP to distribution
eberlm <eberlm@in.tum.de>
parents: 64634
diff changeset
  1986
    by (intro map_pmf_of_set_inj assms bij_betw_imp_inj_on[OF assms(1)])
7504569a73c7 moved material from AFP to distribution
eberlm <eberlm@in.tum.de>
parents: 64634
diff changeset
  1987
  also from assms have "f ` A = B" by (simp add: bij_betw_def)
7504569a73c7 moved material from AFP to distribution
eberlm <eberlm@in.tum.de>
parents: 64634
diff changeset
  1988
  finally show ?thesis .
7504569a73c7 moved material from AFP to distribution
eberlm <eberlm@in.tum.de>
parents: 64634
diff changeset
  1989
qed
7504569a73c7 moved material from AFP to distribution
eberlm <eberlm@in.tum.de>
parents: 64634
diff changeset
  1990
63099
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  1991
text \<open>
63886
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63882
diff changeset
  1992
  Choosing an element uniformly at random from the union of a disjoint family
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63882
diff changeset
  1993
  of finite non-empty sets with the same size is the same as first choosing a set
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63882
diff changeset
  1994
  from the family uniformly at random and then choosing an element from the chosen set
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63882
diff changeset
  1995
  uniformly at random.
63099
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  1996
\<close>
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  1997
lemma pmf_of_set_UN:
69313
b021008c5397 removed legacy input syntax
haftmann
parents: 68386
diff changeset
  1998
  assumes "finite (\<Union>(f ` A))" "A \<noteq> {}" "\<And>x. x \<in> A \<Longrightarrow> f x \<noteq> {}"
63099
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  1999
          "\<And>x. x \<in> A \<Longrightarrow> card (f x) = n" "disjoint_family_on f A"
69313
b021008c5397 removed legacy input syntax
haftmann
parents: 68386
diff changeset
  2000
  shows   "pmf_of_set (\<Union>(f ` A)) = do {x \<leftarrow> pmf_of_set A; pmf_of_set (f x)}"
63099
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2001
            (is "?lhs = ?rhs")
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2002
proof (intro pmf_eqI)
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2003
  fix x
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2004
  from assms have [simp]: "finite A"
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2005
    using infinite_disjoint_family_imp_infinite_UNION[of A f] by blast
69313
b021008c5397 removed legacy input syntax
haftmann
parents: 68386
diff changeset
  2006
  from assms have "ereal (pmf (pmf_of_set (\<Union>(f ` A))) x) =
63099
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2007
    ereal (indicator (\<Union>x\<in>A. f x) x / real (card (\<Union>x\<in>A. f x)))"
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2008
    by (subst pmf_of_set) auto
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2009
  also from assms have "card (\<Union>x\<in>A. f x) = card A * n"
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2010
    by (subst card_UN_disjoint) (auto simp: disjoint_family_on_def)
63886
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63882
diff changeset
  2011
  also from assms
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63882
diff changeset
  2012
    have "indicator (\<Union>x\<in>A. f x) x / real \<dots> =
63099
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2013
              indicator (\<Union>x\<in>A. f x) x / (n * real (card A))"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64008
diff changeset
  2014
      by (simp add: sum_divide_distrib [symmetric] mult_ac)
63099
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2015
  also from assms have "indicator (\<Union>x\<in>A. f x) x = (\<Sum>y\<in>A. indicator (f y) x)"
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2016
    by (intro indicator_UN_disjoint) simp_all
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2017
  also from assms have "ereal ((\<Sum>y\<in>A. indicator (f y) x) / (real n * real (card A))) =
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2018
                          ereal (pmf ?rhs x)"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64008
diff changeset
  2019
    by (subst pmf_bind_pmf_of_set) (simp_all add: sum_divide_distrib)
63099
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2020
  finally show "pmf ?lhs x = pmf ?rhs x" by simp
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2021
qed
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2022
60068
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
  2023
lemma bernoulli_pmf_half_conv_pmf_of_set: "bernoulli_pmf (1 / 2) = pmf_of_set UNIV"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2024
  by (rule pmf_eqI) simp_all
61634
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  2025
59093
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  2026
subsubsection \<open> Poisson Distribution \<close>
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  2027
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  2028
context
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  2029
  fixes rate :: real assumes rate_pos: "0 < rate"
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  2030
begin
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  2031
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  2032
lift_definition poisson_pmf :: "nat pmf" is "\<lambda>k. rate ^ k / fact k * exp (-rate)"
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 59667
diff changeset
  2033
proof  (* by Manuel Eberl *)
59093
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  2034
  have summable: "summable (\<lambda>x::nat. rate ^ x / fact x)" using summable_exp
59557
ebd8ecacfba6 establish unique preferred fact names
haftmann
parents: 59527
diff changeset
  2035
    by (simp add: field_simps divide_inverse [symmetric])
59093
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  2036
  have "(\<integral>\<^sup>+(x::nat). rate ^ x / fact x * exp (-rate) \<partial>count_space UNIV) =
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  2037
          exp (-rate) * (\<integral>\<^sup>+(x::nat). rate ^ x / fact x \<partial>count_space UNIV)"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2038
    by (simp add: field_simps nn_integral_cmult[symmetric] ennreal_mult'[symmetric])
59093
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  2039
  also from rate_pos have "(\<integral>\<^sup>+(x::nat). rate ^ x / fact x \<partial>count_space UNIV) = (\<Sum>x. rate ^ x / fact x)"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2040
    by (simp_all add: nn_integral_count_space_nat suminf_ennreal summable ennreal_suminf_neq_top)
59093
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  2041
  also have "... = exp rate" unfolding exp_def
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 59667
diff changeset
  2042
    by (simp add: field_simps divide_inverse [symmetric])
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2043
  also have "ennreal (exp (-rate)) * ennreal (exp rate) = 1"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2044
    by (simp add: mult_exp_exp ennreal_mult[symmetric])
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2045
  finally show "(\<integral>\<^sup>+ x. ennreal (rate ^ x / (fact x) * exp (- rate)) \<partial>count_space UNIV) = 1" .
59093
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  2046
qed (simp add: rate_pos[THEN less_imp_le])
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  2047
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  2048
lemma pmf_poisson[simp]: "pmf poisson_pmf k = rate ^ k / fact k * exp (-rate)"
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  2049
  by transfer rule
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  2050
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  2051
lemma set_pmf_poisson[simp]: "set_pmf poisson_pmf = UNIV"
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  2052
  using rate_pos by (auto simp: set_pmf_iff)
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  2053
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  2054
end
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  2055
59093
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  2056
subsubsection \<open> Binomial Distribution \<close>
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  2057
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  2058
context
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  2059
  fixes n :: nat and p :: real assumes p_nonneg: "0 \<le> p" and p_le_1: "p \<le> 1"
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  2060
begin
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  2061
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  2062
lift_definition binomial_pmf :: "nat pmf" is "\<lambda>k. (n choose k) * p^k * (1 - p)^(n - k)"
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  2063
proof
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2064
  have "(\<integral>\<^sup>+k. ennreal (real (n choose k) * p ^ k * (1 - p) ^ (n - k)) \<partial>count_space UNIV) =
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2065
    ennreal (\<Sum>k\<le>n. real (n choose k) * p ^ k * (1 - p) ^ (n - k))"
59093
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  2066
    using p_le_1 p_nonneg by (subst nn_integral_count_space') auto
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  2067
  also have "(\<Sum>k\<le>n. real (n choose k) * p ^ k * (1 - p) ^ (n - k)) = (p + (1 - p)) ^ n"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  2068
    by (subst binomial_ring) (simp add: atLeast0AtMost)
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  2069
  finally show "(\<integral>\<^sup>+ x. ennreal (real (n choose x) * p ^ x * (1 - p) ^ (n - x)) \<partial>count_space UNIV) = 1"
59093
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  2070
    by simp
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  2071
qed (insert p_nonneg p_le_1, simp)
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  2072
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  2073
lemma pmf_binomial[simp]: "pmf binomial_pmf k = (n choose k) * p^k * (1 - p)^(n - k)"
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  2074
  by transfer rule
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  2075
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  2076
lemma set_pmf_binomial_eq: "set_pmf binomial_pmf = (if p = 0 then {0} else if p = 1 then {n} else {.. n})"
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  2077
  using p_nonneg p_le_1 unfolding set_eq_iff set_pmf_iff pmf_binomial by (auto simp: set_pmf_iff)
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  2078
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  2079
end
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  2080
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  2081
end
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  2082
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  2083
lemma set_pmf_binomial_0[simp]: "set_pmf (binomial_pmf n 0) = {0}"
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  2084
  by (simp add: set_pmf_binomial_eq)
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  2085
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  2086
lemma set_pmf_binomial_1[simp]: "set_pmf (binomial_pmf n 1) = {n}"
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  2087
  by (simp add: set_pmf_binomial_eq)
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  2088
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  2089
lemma set_pmf_binomial[simp]: "0 < p \<Longrightarrow> p < 1 \<Longrightarrow> set_pmf (binomial_pmf n p) = {..n}"
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  2090
  by (simp add: set_pmf_binomial_eq)
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  2091
73253
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2092
lemma finite_set_pmf_binomial_pmf [intro]: "p \<in> {0..1} \<Longrightarrow> finite (set_pmf (binomial_pmf n p))"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2093
  by (subst set_pmf_binomial_eq) auto
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2094
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2095
lemma expectation_binomial_pmf':
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2096
  fixes f :: "nat \<Rightarrow> 'a :: {banach, second_countable_topology}"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2097
  assumes p: "p \<in> {0..1}"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2098
  shows   "measure_pmf.expectation (binomial_pmf n p) f =
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2099
             (\<Sum>k\<le>n. (real (n choose k) * p ^ k * (1 - p) ^ (n - k)) *\<^sub>R f k)"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2100
  using p by (subst integral_measure_pmf[where A = "{..n}"])
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2101
             (auto simp: set_pmf_binomial_eq split: if_splits)
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2102
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2103
lemma integrable_binomial_pmf [simp, intro]:
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2104
  fixes f :: "nat \<Rightarrow> 'a :: {banach, second_countable_topology}"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2105
  assumes p: "p \<in> {0..1}"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2106
  shows "integrable (binomial_pmf n p) f"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2107
  by (rule integrable_measure_pmf_finite) (use assms in auto)
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2108
63343
fb5d8a50c641 bundle lifting_syntax;
wenzelm
parents: 63194
diff changeset
  2109
context includes lifting_syntax
fb5d8a50c641 bundle lifting_syntax;
wenzelm
parents: 63194
diff changeset
  2110
begin
61634
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  2111
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  2112
lemma bind_pmf_parametric [transfer_rule]:
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  2113
  "(rel_pmf A ===> (A ===> rel_pmf B) ===> rel_pmf B) bind_pmf bind_pmf"
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  2114
by(blast intro: rel_pmf_bindI dest: rel_funD)
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  2115
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  2116
lemma return_pmf_parametric [transfer_rule]: "(A ===> rel_pmf A) return_pmf return_pmf"
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  2117
by(rule rel_funI) simp
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  2118
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  2119
end
61634
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  2120
63099
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2121
63194
0b7bdb75f451 Added code generation for PMFs
eberlm
parents: 63101
diff changeset
  2122
primrec replicate_pmf :: "nat \<Rightarrow> 'a pmf \<Rightarrow> 'a list pmf" where
0b7bdb75f451 Added code generation for PMFs
eberlm
parents: 63101
diff changeset
  2123
  "replicate_pmf 0 _ = return_pmf []"
0b7bdb75f451 Added code generation for PMFs
eberlm
parents: 63101
diff changeset
  2124
| "replicate_pmf (Suc n) p = do {x \<leftarrow> p; xs \<leftarrow> replicate_pmf n p; return_pmf (x#xs)}"
0b7bdb75f451 Added code generation for PMFs
eberlm
parents: 63101
diff changeset
  2125
0b7bdb75f451 Added code generation for PMFs
eberlm
parents: 63101
diff changeset
  2126
lemma replicate_pmf_1: "replicate_pmf 1 p = map_pmf (\<lambda>x. [x]) p"
0b7bdb75f451 Added code generation for PMFs
eberlm
parents: 63101
diff changeset
  2127
  by (simp add: map_pmf_def bind_return_pmf)
63886
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63882
diff changeset
  2128
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63882
diff changeset
  2129
lemma set_replicate_pmf:
63194
0b7bdb75f451 Added code generation for PMFs
eberlm
parents: 63101
diff changeset
  2130
  "set_pmf (replicate_pmf n p) = {xs\<in>lists (set_pmf p). length xs = n}"
0b7bdb75f451 Added code generation for PMFs
eberlm
parents: 63101
diff changeset
  2131
  by (induction n) (auto simp: length_Suc_conv)
0b7bdb75f451 Added code generation for PMFs
eberlm
parents: 63101
diff changeset
  2132
0b7bdb75f451 Added code generation for PMFs
eberlm
parents: 63101
diff changeset
  2133
lemma replicate_pmf_distrib:
63886
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63882
diff changeset
  2134
  "replicate_pmf (m + n) p =
63194
0b7bdb75f451 Added code generation for PMFs
eberlm
parents: 63101
diff changeset
  2135
     do {xs \<leftarrow> replicate_pmf m p; ys \<leftarrow> replicate_pmf n p; return_pmf (xs @ ys)}"
0b7bdb75f451 Added code generation for PMFs
eberlm
parents: 63101
diff changeset
  2136
  by (induction m) (simp_all add: bind_return_pmf bind_return_pmf' bind_assoc_pmf)
0b7bdb75f451 Added code generation for PMFs
eberlm
parents: 63101
diff changeset
  2137
63886
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63882
diff changeset
  2138
lemma power_diff':
63194
0b7bdb75f451 Added code generation for PMFs
eberlm
parents: 63101
diff changeset
  2139
  assumes "b \<le> a"
0b7bdb75f451 Added code generation for PMFs
eberlm
parents: 63101
diff changeset
  2140
  shows   "x ^ (a - b) = (if x = 0 \<and> a = b then 1 else x ^ a / (x::'a::field) ^ b)"
0b7bdb75f451 Added code generation for PMFs
eberlm
parents: 63101
diff changeset
  2141
proof (cases "x = 0")
0b7bdb75f451 Added code generation for PMFs
eberlm
parents: 63101
diff changeset
  2142
  case True
0b7bdb75f451 Added code generation for PMFs
eberlm
parents: 63101
diff changeset
  2143
  with assms show ?thesis by (cases "a - b") simp_all
0b7bdb75f451 Added code generation for PMFs
eberlm
parents: 63101
diff changeset
  2144
qed (insert assms, simp_all add: power_diff)
0b7bdb75f451 Added code generation for PMFs
eberlm
parents: 63101
diff changeset
  2145
63886
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63882
diff changeset
  2146
63194
0b7bdb75f451 Added code generation for PMFs
eberlm
parents: 63101
diff changeset
  2147
lemma binomial_pmf_Suc:
0b7bdb75f451 Added code generation for PMFs
eberlm
parents: 63101
diff changeset
  2148
  assumes "p \<in> {0..1}"
63886
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63882
diff changeset
  2149
  shows   "binomial_pmf (Suc n) p =
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63882
diff changeset
  2150
             do {b \<leftarrow> bernoulli_pmf p;
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63882
diff changeset
  2151
                 k \<leftarrow> binomial_pmf n p;
63194
0b7bdb75f451 Added code generation for PMFs
eberlm
parents: 63101
diff changeset
  2152
                 return_pmf ((if b then 1 else 0) + k)}" (is "_ = ?rhs")
0b7bdb75f451 Added code generation for PMFs
eberlm
parents: 63101
diff changeset
  2153
proof (intro pmf_eqI)
0b7bdb75f451 Added code generation for PMFs
eberlm
parents: 63101
diff changeset
  2154
  fix k
0b7bdb75f451 Added code generation for PMFs
eberlm
parents: 63101
diff changeset
  2155
  have A: "indicator {Suc a} (Suc b) = indicator {a} b" for a b
0b7bdb75f451 Added code generation for PMFs
eberlm
parents: 63101
diff changeset
  2156
    by (simp add: indicator_def)
0b7bdb75f451 Added code generation for PMFs
eberlm
parents: 63101
diff changeset
  2157
  show "pmf (binomial_pmf (Suc n) p) k = pmf ?rhs k"
0b7bdb75f451 Added code generation for PMFs
eberlm
parents: 63101
diff changeset
  2158
    by (cases k; cases "k > n")
70817
dd675800469d dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents: 70532
diff changeset
  2159
       (insert assms, auto simp: pmf_bind measure_pmf_single A field_split_simps algebra_simps
63194
0b7bdb75f451 Added code generation for PMFs
eberlm
parents: 63101
diff changeset
  2160
          not_less less_eq_Suc_le [symmetric] power_diff')
0b7bdb75f451 Added code generation for PMFs
eberlm
parents: 63101
diff changeset
  2161
qed
0b7bdb75f451 Added code generation for PMFs
eberlm
parents: 63101
diff changeset
  2162
0b7bdb75f451 Added code generation for PMFs
eberlm
parents: 63101
diff changeset
  2163
lemma binomial_pmf_0: "p \<in> {0..1} \<Longrightarrow> binomial_pmf 0 p = return_pmf 0"
0b7bdb75f451 Added code generation for PMFs
eberlm
parents: 63101
diff changeset
  2164
  by (rule pmf_eqI) (simp_all add: indicator_def)
0b7bdb75f451 Added code generation for PMFs
eberlm
parents: 63101
diff changeset
  2165
0b7bdb75f451 Added code generation for PMFs
eberlm
parents: 63101
diff changeset
  2166
lemma binomial_pmf_altdef:
0b7bdb75f451 Added code generation for PMFs
eberlm
parents: 63101
diff changeset
  2167
  assumes "p \<in> {0..1}"
0b7bdb75f451 Added code generation for PMFs
eberlm
parents: 63101
diff changeset
  2168
  shows   "binomial_pmf n p = map_pmf (length \<circ> filter id) (replicate_pmf n (bernoulli_pmf p))"
63886
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63882
diff changeset
  2169
  by (induction n)
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63882
diff changeset
  2170
     (insert assms, auto simp: binomial_pmf_Suc map_pmf_def bind_return_pmf bind_assoc_pmf
63194
0b7bdb75f451 Added code generation for PMFs
eberlm
parents: 63101
diff changeset
  2171
        bind_return_pmf' binomial_pmf_0 intro!: bind_pmf_cong)
0b7bdb75f451 Added code generation for PMFs
eberlm
parents: 63101
diff changeset
  2172
0b7bdb75f451 Added code generation for PMFs
eberlm
parents: 63101
diff changeset
  2173
73253
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2174
subsection \<open>Negative Binomial distribution\<close>
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2175
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2176
text \<open>
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2177
  The negative binomial distribution counts the number of times a weighted coin comes up
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2178
  tails before having come up heads \<open>n\<close> times. In other words: how many failures do we see before
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2179
  seeing the \<open>n\<close>-th success?
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2180
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2181
  An alternative view is that the negative binomial distribution is the sum of \<open>n\<close> i.i.d.
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2182
  geometric variables (this is the definition that we use).
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2183
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2184
  Note that there are sometimes different conventions for this distributions in the literature;
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2185
  for instance, sometimes the number of \<^emph>\<open>attempts\<close> is counted instead of the number of failures.
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2186
  This only shifts the entire distribution by a constant number and is thus not a big difference.
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2187
  I think that the convention we use is the most natural one since the support of the distribution
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2188
  starts at 0, whereas for the other convention it starts at \<open>n\<close>.
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2189
\<close>
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2190
primrec neg_binomial_pmf :: "nat \<Rightarrow> real \<Rightarrow> nat pmf" where
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2191
  "neg_binomial_pmf 0 p = return_pmf 0"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2192
| "neg_binomial_pmf (Suc n) p =
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2193
     map_pmf (\<lambda>(x,y). (x + y)) (pair_pmf (geometric_pmf p) (neg_binomial_pmf n p))"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2194
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2195
lemma neg_binomial_pmf_Suc_0 [simp]: "neg_binomial_pmf (Suc 0) p = geometric_pmf p"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2196
  by (auto simp: pair_pmf_def bind_return_pmf map_pmf_def bind_assoc_pmf bind_return_pmf')
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2197
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2198
lemmas neg_binomial_pmf_Suc [simp del] = neg_binomial_pmf.simps(2)
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2199
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2200
lemma neg_binomial_prob_1 [simp]: "neg_binomial_pmf n 1 = return_pmf 0"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2201
  by (induction n) (simp_all add: neg_binomial_pmf_Suc)
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2202
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2203
text \<open>
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2204
  We can now show the aforementioned intuition about counting the failures before the
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2205
  \<open>n\<close>-th success with the following recurrence:
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2206
\<close>
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2207
lemma neg_binomial_pmf_unfold:
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2208
  assumes p: "p \<in> {0<..1}"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2209
  shows "neg_binomial_pmf (Suc n) p =
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2210
           do {b \<leftarrow> bernoulli_pmf p;
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2211
               if b then neg_binomial_pmf n p else map_pmf Suc (neg_binomial_pmf (Suc n) p)}"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2212
  (is "_ = ?rhs")
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2213
  unfolding neg_binomial_pmf_Suc
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2214
  by (subst geometric_bind_pmf_unfold[OF p])
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2215
     (auto simp: map_pmf_def pair_pmf_def bind_assoc_pmf bind_return_pmf bind_return_pmf'
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2216
           intro!: bind_pmf_cong)
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2217
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2218
text \<open>
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2219
  Next, we show an explicit formula for the probability mass function of the negative
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2220
  binomial distribution:
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2221
\<close>
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2222
lemma pmf_neg_binomial:
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2223
  assumes p: "p \<in> {0<..1}"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2224
  shows   "pmf (neg_binomial_pmf n p) k = real ((k + n - 1) choose k) * p ^ n * (1 - p) ^ k"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2225
proof (induction n arbitrary: k)
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2226
  case 0
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2227
  thus ?case using assms by (auto simp: indicator_def)
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2228
next
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2229
  case (Suc n)
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2230
  show ?case
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2231
  proof (cases "n = 0")
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2232
    case True
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2233
    thus ?thesis using assms by auto
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2234
  next
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2235
    case False
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2236
    let ?f = "pmf (neg_binomial_pmf n p)"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2237
    have "pmf (neg_binomial_pmf (Suc n) p) k =
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2238
            pmf (geometric_pmf p \<bind> (\<lambda>x. map_pmf ((+) x) (neg_binomial_pmf n p))) k"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2239
      by (auto simp: pair_pmf_def bind_return_pmf map_pmf_def bind_assoc_pmf neg_binomial_pmf_Suc)
75625
0dd3ac5fdbaa tuned BNF bounds for function space and bounded sets; NEWS and CONTRIBUTORS
traytel
parents: 75624
diff changeset
  2240
    also have "\<dots> = measure_pmf.expectation (geometric_pmf p)
73253
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2241
                      (\<lambda>x. measure_pmf.prob (neg_binomial_pmf n p) ((+) x -` {k}))"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2242
      by (simp add: pmf_bind pmf_map)
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2243
    also have "(\<lambda>x. (+) x -` {k}) = (\<lambda>x. if x \<le> k then {k - x} else {})"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2244
      by (auto simp: fun_eq_iff)
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2245
    also have "(\<lambda>x. measure_pmf.prob (neg_binomial_pmf n p) (\<dots> x)) =
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2246
               (\<lambda>x. if x \<le> k then ?f(k - x) else 0)"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2247
      by (auto simp: fun_eq_iff measure_pmf_single)
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2248
    also have "measure_pmf.expectation (geometric_pmf p) \<dots> =
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2249
                 (\<Sum>i\<le>k. pmf (neg_binomial_pmf n p) (k - i) * pmf (geometric_pmf p) i)"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2250
      by (subst integral_measure_pmf_real[where A = "{..k}"]) (auto split: if_splits)
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2251
    also have "\<dots> = p^(n+1) * (1-p)^k * real (\<Sum>i\<le>k. (k - i + n - 1) choose (k - i))"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2252
      unfolding sum_distrib_left of_nat_sum
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2253
    proof (intro sum.cong refl, goal_cases)
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2254
      case (1 i)
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2255
      have "pmf (neg_binomial_pmf n p) (k - i) * pmf (geometric_pmf p) i =
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2256
              real ((k - i + n - 1) choose (k - i)) * p^(n+1) * ((1-p)^(k-i) * (1-p)^i)"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2257
        using assms Suc.IH by (simp add: mult_ac)
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2258
      also have "(1-p)^(k-i) * (1-p)^i = (1-p)^k"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2259
        using 1 by (subst power_add [symmetric]) auto
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2260
      finally show ?case by simp
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2261
    qed
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2262
    also have "(\<Sum>i\<le>k. (k - i + n - 1) choose (k - i)) = (\<Sum>i\<le>k. (n - 1 + i) choose i)"
75625
0dd3ac5fdbaa tuned BNF bounds for function space and bounded sets; NEWS and CONTRIBUTORS
traytel
parents: 75624
diff changeset
  2263
      by (intro sum.reindex_bij_witness[of _ "\<lambda>i. k - i" "\<lambda>i. k - i"])
73253
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2264
         (use \<open>n \<noteq> 0\<close> in \<open>auto simp: algebra_simps\<close>)
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2265
    also have "\<dots> = (n + k) choose k"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2266
      by (subst sum_choose_lower) (use \<open>n \<noteq> 0\<close> in auto)
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2267
    finally show ?thesis
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2268
      by (simp add: add_ac)
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2269
  qed
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2270
qed
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2271
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2272
(* TODO: Move? *)
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2273
lemma gbinomial_0_left: "0 gchoose k = (if k = 0 then 1 else 0)"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2274
  by (cases k) auto
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2275
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2276
text \<open>
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2277
  The following alternative formula highlights why it is called `negative binomial distribution':
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2278
\<close>
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2279
lemma pmf_neg_binomial':
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2280
  assumes p: "p \<in> {0<..1}"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2281
  shows   "pmf (neg_binomial_pmf n p) k = (-1) ^ k * ((-real n) gchoose k) * p ^ n * (1 - p) ^ k"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2282
proof (cases "n > 0")
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2283
  case n: True
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2284
  have "pmf (neg_binomial_pmf n p) k = real ((k + n - 1) choose k) * p ^ n * (1 - p) ^ k"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2285
    by (rule pmf_neg_binomial) fact+
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2286
  also have "real ((k + n - 1) choose k) = ((real k + real n - 1) gchoose k)"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2287
    using n by (subst binomial_gbinomial) (auto simp: of_nat_diff)
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2288
  also have "\<dots> = (-1) ^ k * ((-real n) gchoose k)"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2289
    by (subst gbinomial_negated_upper) auto
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2290
  finally show ?thesis by simp
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2291
qed (auto simp: indicator_def gbinomial_0_left)
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2292
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2293
text \<open>
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2294
  The cumulative distribution function of the negative binomial distribution can be
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2295
  expressed in terms of that of the `normal' binomial distribution.
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2296
\<close>
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2297
lemma prob_neg_binomial_pmf_atMost:
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2298
  assumes p: "p \<in> {0<..1}"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2299
  shows "measure_pmf.prob (neg_binomial_pmf n p) {..k} =
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2300
         measure_pmf.prob (binomial_pmf (n + k) (1 - p)) {..k}"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2301
proof (cases "n = 0")
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2302
  case [simp]: True
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2303
  have "set_pmf (binomial_pmf (n + k) (1 - p)) \<subseteq> {..n+k}"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2304
    using p by (subst set_pmf_binomial_eq) auto
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2305
  hence "measure_pmf.prob (binomial_pmf (n + k) (1 - p)) {..k} = 1"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2306
    by (subst measure_pmf.prob_eq_1) (auto intro!: AE_pmfI)
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2307
  thus ?thesis by simp
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2308
next
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2309
  case False
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2310
  hence n: "n > 0" by auto
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2311
  have "measure_pmf.prob (binomial_pmf (n + k) (1 - p)) {..k} = (\<Sum>i\<le>k. pmf (binomial_pmf (n + k) (1 - p)) i)"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2312
    by (intro measure_measure_pmf_finite) auto
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2313
  also have "\<dots> = (\<Sum>i\<le>k. real ((n + k) choose i) * p ^ (n + k - i) * (1 - p) ^ i)"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2314
    using p by (simp add: mult_ac)
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2315
  also have "\<dots> = p ^ n * (\<Sum>i\<le>k. real ((n + k) choose i) * (1 - p) ^ i * p ^ (k - i))"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2316
    unfolding sum_distrib_left by (intro sum.cong) (auto simp: algebra_simps simp flip: power_add)
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2317
  also have "(\<Sum>i\<le>k. real ((n + k) choose i) * (1 - p) ^ i * p ^ (k - i)) =
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2318
             (\<Sum>i\<le>k. ((n + i - 1) choose i) * (1 - p) ^ i)"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2319
    using gbinomial_partial_sum_poly_xpos[of k "real n" "1 - p" p] n
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2320
    by (simp add: binomial_gbinomial add_ac of_nat_diff)
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2321
  also have "p ^ n * \<dots> = (\<Sum>i\<le>k. pmf (neg_binomial_pmf n p) i)"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2322
    using p unfolding sum_distrib_left by (simp add: pmf_neg_binomial algebra_simps)
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2323
  also have "\<dots> = measure_pmf.prob (neg_binomial_pmf n p) {..k}"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2324
    by (intro measure_measure_pmf_finite [symmetric]) auto
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2325
  finally show ?thesis ..
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2326
qed
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2327
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2328
lemma prob_neg_binomial_pmf_lessThan:
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2329
  assumes p: "p \<in> {0<..1}"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2330
  shows "measure_pmf.prob (neg_binomial_pmf n p) {..<k} =
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2331
         measure_pmf.prob (binomial_pmf (n + k - 1) (1 - p)) {..<k}"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2332
proof (cases "k = 0")
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2333
  case False
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2334
  hence "{..<k} = {..k-1}"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2335
    by auto
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2336
  thus ?thesis
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2337
    using prob_neg_binomial_pmf_atMost[OF p, of n "k - 1"] False by simp
75625
0dd3ac5fdbaa tuned BNF bounds for function space and bounded sets; NEWS and CONTRIBUTORS
traytel
parents: 75624
diff changeset
  2338
qed auto
73253
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2339
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2340
text \<open>
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2341
  The expected value of the negative binomial distribution is $n(1-p)/p$:
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2342
\<close>
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2343
lemma nn_integral_neg_binomial_pmf_real:
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2344
  assumes p: "p \<in> {0<..1}"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2345
  shows "nn_integral (measure_pmf (neg_binomial_pmf n p)) of_nat = ennreal (n * (1 - p) / p)"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2346
proof (induction n)
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2347
  case 0
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2348
  thus ?case by auto
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2349
next
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2350
  case (Suc n)
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2351
  have "nn_integral (measure_pmf (neg_binomial_pmf (Suc n) p)) of_nat =
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2352
        nn_integral (measure_pmf (geometric_pmf p)) of_nat +
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2353
        nn_integral (measure_pmf (neg_binomial_pmf n p)) of_nat"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2354
    by (simp add: neg_binomial_pmf_Suc case_prod_unfold nn_integral_add nn_integral_pair_pmf')
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2355
  also have "nn_integral (measure_pmf (geometric_pmf p)) of_nat = ennreal ((1-p) / p)"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2356
    unfolding ennreal_of_nat_eq_real_of_nat
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2357
    using expectation_geometric_pmf[OF p] integrable_real_geometric_pmf[OF p]
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2358
    by (subst nn_integral_eq_integral) auto
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2359
  also have "nn_integral (measure_pmf (neg_binomial_pmf n p)) of_nat = n * (1 - p) / p" using p
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2360
    by (subst Suc.IH)
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2361
       (auto simp: ennreal_of_nat_eq_real_of_nat ennreal_mult simp flip: divide_ennreal ennreal_minus)
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2362
  also have "ennreal ((1 - p) / p) + ennreal (real n * (1 - p) / p) =
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2363
             ennreal ((1-p) / p + real n * (1 - p) / p)"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2364
    by (intro ennreal_plus [symmetric] divide_nonneg_pos mult_nonneg_nonneg) (use p in auto)
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2365
  also have "(1-p) / p + real n * (1 - p) / p = real (Suc n) * (1 - p) / p"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2366
    using p by (auto simp: field_simps)
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2367
  finally show ?case
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2368
    by (simp add: ennreal_of_nat_eq_real_of_nat)
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2369
qed
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2370
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2371
lemma integrable_neg_binomial_pmf_real:
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2372
  assumes p: "p \<in> {0<..1}"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2373
  shows "integrable (measure_pmf (neg_binomial_pmf n p)) real"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2374
  using nn_integral_neg_binomial_pmf_real[OF p, of n]
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2375
  by (subst integrable_iff_bounded) (auto simp flip: ennreal_of_nat_eq_real_of_nat)
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2376
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2377
lemma expectation_neg_binomial_pmf:
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2378
  assumes p: "p \<in> {0<..1}"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2379
  shows "measure_pmf.expectation (neg_binomial_pmf n p) real = n * (1 - p) / p"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2380
proof -
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2381
  have "nn_integral (measure_pmf (neg_binomial_pmf n p)) of_nat = ennreal (n * (1 - p) / p)"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2382
    by (intro nn_integral_neg_binomial_pmf_real p)
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2383
  also have "of_nat = (\<lambda>x. ennreal (real x))"
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2384
    by (simp add: ennreal_of_nat_eq_real_of_nat fun_eq_iff)
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2385
  finally show ?thesis
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2386
    using p by (subst (asm) nn_integral_eq_integrable) auto
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2387
qed
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2388
f6bb31879698 HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
Manuel Eberl <eberlm@in.tum.de>
parents: 70817
diff changeset
  2389
67486
Lars Hupel <lars.hupel@mytum.de>
parents: 67399
diff changeset
  2390
subsection \<open>PMFs from association lists\<close>
63099
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2391
63886
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63882
diff changeset
  2392
definition pmf_of_list ::" ('a \<times> real) list \<Rightarrow> 'a pmf" where
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63680
diff changeset
  2393
  "pmf_of_list xs = embed_pmf (\<lambda>x. sum_list (map snd (filter (\<lambda>z. fst z = x) xs)))"
63099
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2394
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2395
definition pmf_of_list_wf where
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63680
diff changeset
  2396
  "pmf_of_list_wf xs \<longleftrightarrow> (\<forall>x\<in>set (map snd xs) . x \<ge> 0) \<and> sum_list (map snd xs) = 1"
63099
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2397
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2398
lemma pmf_of_list_wfI:
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63680
diff changeset
  2399
  "(\<And>x. x \<in> set (map snd xs) \<Longrightarrow> x \<ge> 0) \<Longrightarrow> sum_list (map snd xs) = 1 \<Longrightarrow> pmf_of_list_wf xs"
63099
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2400
  unfolding pmf_of_list_wf_def by simp
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2401
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2402
context
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2403
begin
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2404
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2405
private lemma pmf_of_list_aux:
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2406
  assumes "\<And>x. x \<in> set (map snd xs) \<Longrightarrow> x \<ge> 0"
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63680
diff changeset
  2407
  assumes "sum_list (map snd xs) = 1"
68386
98cf1c823c48 Keep filter input syntax
nipkow
parents: 68249
diff changeset
  2408
  shows "(\<integral>\<^sup>+ x. ennreal (sum_list (map snd [z\<leftarrow>xs . fst z = x])) \<partial>count_space UNIV) = 1"
63099
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2409
proof -
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63680
diff changeset
  2410
  have "(\<integral>\<^sup>+ x. ennreal (sum_list (map snd (filter (\<lambda>z. fst z = x) xs))) \<partial>count_space UNIV) =
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63680
diff changeset
  2411
            (\<integral>\<^sup>+ x. ennreal (sum_list (map (\<lambda>(x',p). indicator {x'} x * p) xs)) \<partial>count_space UNIV)"
67486
Lars Hupel <lars.hupel@mytum.de>
parents: 67399
diff changeset
  2412
    apply (intro nn_integral_cong ennreal_cong, subst sum_list_map_filter')
Lars Hupel <lars.hupel@mytum.de>
parents: 67399
diff changeset
  2413
    apply (rule arg_cong[where f = sum_list])
Lars Hupel <lars.hupel@mytum.de>
parents: 67399
diff changeset
  2414
    apply (auto cong: map_cong)
Lars Hupel <lars.hupel@mytum.de>
parents: 67399
diff changeset
  2415
    done
63099
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2416
  also have "\<dots> = (\<Sum>(x',p)\<leftarrow>xs. (\<integral>\<^sup>+ x. ennreal (indicator {x'} x * p) \<partial>count_space UNIV))"
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2417
    using assms(1)
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2418
  proof (induction xs)
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2419
    case (Cons x xs)
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2420
    from Cons.prems have "snd x \<ge> 0" by simp
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2421
    moreover have "b \<ge> 0" if "(a,b) \<in> set xs" for a b
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2422
      using Cons.prems[of b] that by force
63886
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63882
diff changeset
  2423
    ultimately have "(\<integral>\<^sup>+ y. ennreal (\<Sum>(x', p)\<leftarrow>x # xs. indicator {x'} y * p) \<partial>count_space UNIV) =
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63882
diff changeset
  2424
            (\<integral>\<^sup>+ y. ennreal (indicator {fst x} y * snd x) +
63099
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2425
            ennreal (\<Sum>(x', p)\<leftarrow>xs. indicator {x'} y * p) \<partial>count_space UNIV)"
63886
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63882
diff changeset
  2426
      by (intro nn_integral_cong, subst ennreal_plus [symmetric])
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63680
diff changeset
  2427
         (auto simp: case_prod_unfold indicator_def intro!: sum_list_nonneg)
63886
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63882
diff changeset
  2428
    also have "\<dots> = (\<integral>\<^sup>+ y. ennreal (indicator {fst x} y * snd x) \<partial>count_space UNIV) +
63099
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2429
                      (\<integral>\<^sup>+ y. ennreal (\<Sum>(x', p)\<leftarrow>xs. indicator {x'} y * p) \<partial>count_space UNIV)"
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2430
      by (intro nn_integral_add)
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63680
diff changeset
  2431
         (force intro!: sum_list_nonneg AE_I2 intro: Cons simp: indicator_def)+
63099
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2432
    also have "(\<integral>\<^sup>+ y. ennreal (\<Sum>(x', p)\<leftarrow>xs. indicator {x'} y * p) \<partial>count_space UNIV) =
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2433
               (\<Sum>(x', p)\<leftarrow>xs. (\<integral>\<^sup>+ y. ennreal (indicator {x'} y * p) \<partial>count_space UNIV))"
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2434
      using Cons(1) by (intro Cons) simp_all
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2435
    finally show ?case by (simp add: case_prod_unfold)
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2436
  qed simp
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2437
  also have "\<dots> = (\<Sum>(x',p)\<leftarrow>xs. ennreal p * (\<integral>\<^sup>+ x. indicator {x'} x \<partial>count_space UNIV))"
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2438
    using assms(1)
67489
f1ba59ddd9a6 drop redundant cong rules
Lars Hupel <lars.hupel@mytum.de>
parents: 67486
diff changeset
  2439
    by (simp cong: map_cong only: case_prod_unfold, subst nn_integral_cmult [symmetric])
63099
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2440
       (auto intro!: assms(1) simp: max_def times_ereal.simps [symmetric] mult_ac ereal_indicator
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2441
             simp del: times_ereal.simps)+
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63680
diff changeset
  2442
  also from assms have "\<dots> = sum_list (map snd xs)" by (simp add: case_prod_unfold sum_list_ennreal)
63099
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2443
  also have "\<dots> = 1" using assms(2) by simp
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2444
  finally show ?thesis .
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2445
qed
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2446
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2447
lemma pmf_pmf_of_list:
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2448
  assumes "pmf_of_list_wf xs"
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63680
diff changeset
  2449
  shows   "pmf (pmf_of_list xs) x = sum_list (map snd (filter (\<lambda>z. fst z = x) xs))"
63099
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2450
  using assms pmf_of_list_aux[of xs] unfolding pmf_of_list_def pmf_of_list_wf_def
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63680
diff changeset
  2451
  by (subst pmf_embed_pmf) (auto intro!: sum_list_nonneg)
63099
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2452
61634
48e2de1b1df5 add various lemmas
Andreas Lochbihler
parents: 61610
diff changeset
  2453
end
63099
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2454
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2455
lemma set_pmf_of_list:
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2456
  assumes "pmf_of_list_wf xs"
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2457
  shows   "set_pmf (pmf_of_list xs) \<subseteq> set (map fst xs)"
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2458
proof clarify
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2459
  fix x assume A: "x \<in> set_pmf (pmf_of_list xs)"
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2460
  show "x \<in> set (map fst xs)"
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2461
  proof (rule ccontr)
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2462
    assume "x \<notin> set (map fst xs)"
68386
98cf1c823c48 Keep filter input syntax
nipkow
parents: 68249
diff changeset
  2463
    hence "[z\<leftarrow>xs . fst z = x] = []" by (auto simp: filter_empty_conv)
63099
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2464
    with A assms show False by (simp add: pmf_pmf_of_list set_pmf_eq)
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2465
  qed
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2466
qed
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2467
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2468
lemma finite_set_pmf_of_list:
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2469
  assumes "pmf_of_list_wf xs"
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2470
  shows   "finite (set_pmf (pmf_of_list xs))"
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2471
  using assms by (rule finite_subset[OF set_pmf_of_list]) simp_all
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2472
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2473
lemma emeasure_Int_set_pmf:
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2474
  "emeasure (measure_pmf p) (A \<inter> set_pmf p) = emeasure (measure_pmf p) A"
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2475
  by (rule emeasure_eq_AE) (auto simp: AE_measure_pmf_iff)
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2476
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2477
lemma measure_Int_set_pmf:
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2478
  "measure (measure_pmf p) (A \<inter> set_pmf p) = measure (measure_pmf p) A"
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2479
  using emeasure_Int_set_pmf[of p A] by (simp add: Sigma_Algebra.measure_def)
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2480
66568
52b5cf533fd6 Connecting PMFs to infinite sums
eberlm <eberlm@in.tum.de>
parents: 66453
diff changeset
  2481
lemma measure_prob_cong_0:
52b5cf533fd6 Connecting PMFs to infinite sums
eberlm <eberlm@in.tum.de>
parents: 66453
diff changeset
  2482
  assumes "\<And>x. x \<in> A - B \<Longrightarrow> pmf p x = 0"
52b5cf533fd6 Connecting PMFs to infinite sums
eberlm <eberlm@in.tum.de>
parents: 66453
diff changeset
  2483
  assumes "\<And>x. x \<in> B - A \<Longrightarrow> pmf p x = 0"
52b5cf533fd6 Connecting PMFs to infinite sums
eberlm <eberlm@in.tum.de>
parents: 66453
diff changeset
  2484
  shows   "measure (measure_pmf p) A = measure (measure_pmf p) B"
52b5cf533fd6 Connecting PMFs to infinite sums
eberlm <eberlm@in.tum.de>
parents: 66453
diff changeset
  2485
proof -
52b5cf533fd6 Connecting PMFs to infinite sums
eberlm <eberlm@in.tum.de>
parents: 66453
diff changeset
  2486
  have "measure_pmf.prob p A = measure_pmf.prob p (A \<inter> set_pmf p)"
52b5cf533fd6 Connecting PMFs to infinite sums
eberlm <eberlm@in.tum.de>
parents: 66453
diff changeset
  2487
    by (simp add: measure_Int_set_pmf)
52b5cf533fd6 Connecting PMFs to infinite sums
eberlm <eberlm@in.tum.de>
parents: 66453
diff changeset
  2488
  also have "A \<inter> set_pmf p = B \<inter> set_pmf p"
52b5cf533fd6 Connecting PMFs to infinite sums
eberlm <eberlm@in.tum.de>
parents: 66453
diff changeset
  2489
    using assms by (auto simp: set_pmf_eq)
52b5cf533fd6 Connecting PMFs to infinite sums
eberlm <eberlm@in.tum.de>
parents: 66453
diff changeset
  2490
  also have "measure_pmf.prob p \<dots> = measure_pmf.prob p B"
52b5cf533fd6 Connecting PMFs to infinite sums
eberlm <eberlm@in.tum.de>
parents: 66453
diff changeset
  2491
    by (simp add: measure_Int_set_pmf)
52b5cf533fd6 Connecting PMFs to infinite sums
eberlm <eberlm@in.tum.de>
parents: 66453
diff changeset
  2492
  finally show ?thesis .
52b5cf533fd6 Connecting PMFs to infinite sums
eberlm <eberlm@in.tum.de>
parents: 66453
diff changeset
  2493
qed
52b5cf533fd6 Connecting PMFs to infinite sums
eberlm <eberlm@in.tum.de>
parents: 66453
diff changeset
  2494
63099
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2495
lemma emeasure_pmf_of_list:
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2496
  assumes "pmf_of_list_wf xs"
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63680
diff changeset
  2497
  shows   "emeasure (pmf_of_list xs) A = ennreal (sum_list (map snd (filter (\<lambda>x. fst x \<in> A) xs)))"
63099
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2498
proof -
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2499
  have "emeasure (pmf_of_list xs) A = nn_integral (measure_pmf (pmf_of_list xs)) (indicator A)"
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2500
    by simp
63886
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63882
diff changeset
  2501
  also from assms
68386
98cf1c823c48 Keep filter input syntax
nipkow
parents: 68249
diff changeset
  2502
    have "\<dots> = (\<Sum>x\<in>set_pmf (pmf_of_list xs) \<inter> A. ennreal (sum_list (map snd [z\<leftarrow>xs . fst z = x])))"
63973
b09f16aeb694 Probability: fix proof
hoelzl
parents: 63918
diff changeset
  2503
    by (subst nn_integral_measure_pmf_finite) (simp_all add: finite_set_pmf_of_list pmf_pmf_of_list Int_def)
63886
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63882
diff changeset
  2504
  also from assms
68386
98cf1c823c48 Keep filter input syntax
nipkow
parents: 68249
diff changeset
  2505
    have "\<dots> = ennreal (\<Sum>x\<in>set_pmf (pmf_of_list xs) \<inter> A. sum_list (map snd [z\<leftarrow>xs . fst z = x]))"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64008
diff changeset
  2506
    by (subst sum_ennreal) (auto simp: pmf_of_list_wf_def intro!: sum_list_nonneg)
63886
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63882
diff changeset
  2507
  also have "\<dots> = ennreal (\<Sum>x\<in>set_pmf (pmf_of_list xs) \<inter> A.
63099
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2508
      indicator A x * pmf (pmf_of_list xs) x)" (is "_ = ennreal ?S")
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64008
diff changeset
  2509
    using assms by (intro ennreal_cong sum.cong) (auto simp: pmf_pmf_of_list)
63099
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2510
  also have "?S = (\<Sum>x\<in>set_pmf (pmf_of_list xs). indicator A x * pmf (pmf_of_list xs) x)"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64008
diff changeset
  2511
    using assms by (intro sum.mono_neutral_left set_pmf_of_list finite_set_pmf_of_list) auto
63099
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2512
  also have "\<dots> = (\<Sum>x\<in>set (map fst xs). indicator A x * pmf (pmf_of_list xs) x)"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64008
diff changeset
  2513
    using assms by (intro sum.mono_neutral_left set_pmf_of_list) (auto simp: set_pmf_eq)
63886
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63882
diff changeset
  2514
  also have "\<dots> = (\<Sum>x\<in>set (map fst xs). indicator A x *
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63680
diff changeset
  2515
                      sum_list (map snd (filter (\<lambda>z. fst z = x) xs)))"
63099
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2516
    using assms by (simp add: pmf_pmf_of_list)
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63680
diff changeset
  2517
  also have "\<dots> = (\<Sum>x\<in>set (map fst xs). sum_list (map snd (filter (\<lambda>z. fst z = x \<and> x \<in> A) xs)))"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64008
diff changeset
  2518
    by (intro sum.cong) (auto simp: indicator_def)
63099
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2519
  also have "\<dots> = (\<Sum>x\<in>set (map fst xs). (\<Sum>xa = 0..<length xs.
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2520
                     if fst (xs ! xa) = x \<and> x \<in> A then snd (xs ! xa) else 0))"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64008
diff changeset
  2521
    by (intro sum.cong refl, subst sum_list_map_filter', subst sum_list_sum_nth) simp
63886
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63882
diff changeset
  2522
  also have "\<dots> = (\<Sum>xa = 0..<length xs. (\<Sum>x\<in>set (map fst xs).
63099
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2523
                     if fst (xs ! xa) = x \<and> x \<in> A then snd (xs ! xa) else 0))"
66804
3f9bb52082c4 avoid name clashes on interpretation of abstract locales
haftmann
parents: 66568
diff changeset
  2524
    by (rule sum.swap)
63886
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63882
diff changeset
  2525
  also have "\<dots> = (\<Sum>xa = 0..<length xs. if fst (xs ! xa) \<in> A then
63099
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2526
                     (\<Sum>x\<in>set (map fst xs). if x = fst (xs ! xa) then snd (xs ! xa) else 0) else 0)"
66089
def95e0bc529 Some new material. SIMPRULE STATUS for sum/prod.delta rules!
paulson <lp15@cam.ac.uk>
parents: 65395
diff changeset
  2527
    by (auto intro!: sum.cong sum.neutral simp del: sum.delta)
63099
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2528
  also have "\<dots> = (\<Sum>xa = 0..<length xs. if fst (xs ! xa) \<in> A then snd (xs ! xa) else 0)"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64008
diff changeset
  2529
    by (intro sum.cong refl) (simp_all add: sum.delta)
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63680
diff changeset
  2530
  also have "\<dots> = sum_list (map snd (filter (\<lambda>x. fst x \<in> A) xs))"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64008
diff changeset
  2531
    by (subst sum_list_map_filter', subst sum_list_sum_nth) simp_all
63886
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63882
diff changeset
  2532
  finally show ?thesis .
63099
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2533
qed
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2534
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2535
lemma measure_pmf_of_list:
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2536
  assumes "pmf_of_list_wf xs"
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63680
diff changeset
  2537
  shows   "measure (pmf_of_list xs) A = sum_list (map snd (filter (\<lambda>x. fst x \<in> A) xs))"
63099
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2538
  using assms unfolding pmf_of_list_wf_def Sigma_Algebra.measure_def
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63680
diff changeset
  2539
  by (subst emeasure_pmf_of_list [OF assms], subst enn2real_ennreal) (auto intro!: sum_list_nonneg)
63099
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2540
63194
0b7bdb75f451 Added code generation for PMFs
eberlm
parents: 63101
diff changeset
  2541
(* TODO Move? *)
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63680
diff changeset
  2542
lemma sum_list_nonneg_eq_zero_iff:
63194
0b7bdb75f451 Added code generation for PMFs
eberlm
parents: 63101
diff changeset
  2543
  fixes xs :: "'a :: linordered_ab_group_add list"
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63680
diff changeset
  2544
  shows "(\<And>x. x \<in> set xs \<Longrightarrow> x \<ge> 0) \<Longrightarrow> sum_list xs = 0 \<longleftrightarrow> set xs \<subseteq> {0}"
63194
0b7bdb75f451 Added code generation for PMFs
eberlm
parents: 63101
diff changeset
  2545
proof (induction xs)
0b7bdb75f451 Added code generation for PMFs
eberlm
parents: 63101
diff changeset
  2546
  case (Cons x xs)
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63680
diff changeset
  2547
  from Cons.prems have "sum_list (x#xs) = 0 \<longleftrightarrow> x = 0 \<and> sum_list xs = 0"
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63680
diff changeset
  2548
    unfolding sum_list_simps by (subst add_nonneg_eq_0_iff) (auto intro: sum_list_nonneg)
63194
0b7bdb75f451 Added code generation for PMFs
eberlm
parents: 63101
diff changeset
  2549
  with Cons.IH Cons.prems show ?case by simp
0b7bdb75f451 Added code generation for PMFs
eberlm
parents: 63101
diff changeset
  2550
qed simp_all
0b7bdb75f451 Added code generation for PMFs
eberlm
parents: 63101
diff changeset
  2551
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63680
diff changeset
  2552
lemma sum_list_filter_nonzero:
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63680
diff changeset
  2553
  "sum_list (filter (\<lambda>x. x \<noteq> 0) xs) = sum_list xs"
63194
0b7bdb75f451 Added code generation for PMFs
eberlm
parents: 63101
diff changeset
  2554
  by (induction xs) simp_all
0b7bdb75f451 Added code generation for PMFs
eberlm
parents: 63101
diff changeset
  2555
(* END MOVE *)
63886
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63882
diff changeset
  2556
63194
0b7bdb75f451 Added code generation for PMFs
eberlm
parents: 63101
diff changeset
  2557
lemma set_pmf_of_list_eq:
0b7bdb75f451 Added code generation for PMFs
eberlm
parents: 63101
diff changeset
  2558
  assumes "pmf_of_list_wf xs" "\<And>x. x \<in> snd ` set xs \<Longrightarrow> x > 0"
0b7bdb75f451 Added code generation for PMFs
eberlm
parents: 63101
diff changeset
  2559
  shows   "set_pmf (pmf_of_list xs) = fst ` set xs"
0b7bdb75f451 Added code generation for PMFs
eberlm
parents: 63101
diff changeset
  2560
proof
0b7bdb75f451 Added code generation for PMFs
eberlm
parents: 63101
diff changeset
  2561
  {
0b7bdb75f451 Added code generation for PMFs
eberlm
parents: 63101
diff changeset
  2562
    fix x assume A: "x \<in> fst ` set xs" and B: "x \<notin> set_pmf (pmf_of_list xs)"
0b7bdb75f451 Added code generation for PMFs
eberlm
parents: 63101
diff changeset
  2563
    then obtain y where y: "(x, y) \<in> set xs" by auto
68386
98cf1c823c48 Keep filter input syntax
nipkow
parents: 68249
diff changeset
  2564
    from B have "sum_list (map snd [z\<leftarrow>xs. fst z = x]) = 0"
63194
0b7bdb75f451 Added code generation for PMFs
eberlm
parents: 63101
diff changeset
  2565
      by (simp add: pmf_pmf_of_list[OF assms(1)] set_pmf_eq)
0b7bdb75f451 Added code generation for PMFs
eberlm
parents: 63101
diff changeset
  2566
    moreover from y have "y \<in> snd ` {xa \<in> set xs. fst xa = x}" by force
63886
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63882
diff changeset
  2567
    ultimately have "y = 0" using assms(1)
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63680
diff changeset
  2568
      by (subst (asm) sum_list_nonneg_eq_zero_iff) (auto simp: pmf_of_list_wf_def)
63194
0b7bdb75f451 Added code generation for PMFs
eberlm
parents: 63101
diff changeset
  2569
    with assms(2) y have False by force
0b7bdb75f451 Added code generation for PMFs
eberlm
parents: 63101
diff changeset
  2570
  }
0b7bdb75f451 Added code generation for PMFs
eberlm
parents: 63101
diff changeset
  2571
  thus "fst ` set xs \<subseteq> set_pmf (pmf_of_list xs)" by blast
0b7bdb75f451 Added code generation for PMFs
eberlm
parents: 63101
diff changeset
  2572
qed (insert set_pmf_of_list[OF assms(1)], simp_all)
63886
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63882
diff changeset
  2573
63194
0b7bdb75f451 Added code generation for PMFs
eberlm
parents: 63101
diff changeset
  2574
lemma pmf_of_list_remove_zeros:
0b7bdb75f451 Added code generation for PMFs
eberlm
parents: 63101
diff changeset
  2575
  assumes "pmf_of_list_wf xs"
0b7bdb75f451 Added code generation for PMFs
eberlm
parents: 63101
diff changeset
  2576
  defines "xs' \<equiv> filter (\<lambda>z. snd z \<noteq> 0) xs"
0b7bdb75f451 Added code generation for PMFs
eberlm
parents: 63101
diff changeset
  2577
  shows   "pmf_of_list_wf xs'" "pmf_of_list xs' = pmf_of_list xs"
0b7bdb75f451 Added code generation for PMFs
eberlm
parents: 63101
diff changeset
  2578
proof -
68386
98cf1c823c48 Keep filter input syntax
nipkow
parents: 68249
diff changeset
  2579
  have "map snd [z\<leftarrow>xs . snd z \<noteq> 0] = filter (\<lambda>x. x \<noteq> 0) (map snd xs)"
63194
0b7bdb75f451 Added code generation for PMFs
eberlm
parents: 63101
diff changeset
  2580
    by (induction xs) simp_all
0b7bdb75f451 Added code generation for PMFs
eberlm
parents: 63101
diff changeset
  2581
  with assms(1) show wf: "pmf_of_list_wf xs'"
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63680
diff changeset
  2582
    by (auto simp: pmf_of_list_wf_def xs'_def sum_list_filter_nonzero)
68386
98cf1c823c48 Keep filter input syntax
nipkow
parents: 68249
diff changeset
  2583
  have "sum_list (map snd [z\<leftarrow>xs' . fst z = i]) = sum_list (map snd [z\<leftarrow>xs . fst z = i])" for i
63194
0b7bdb75f451 Added code generation for PMFs
eberlm
parents: 63101
diff changeset
  2584
    unfolding xs'_def by (induction xs) simp_all
0b7bdb75f451 Added code generation for PMFs
eberlm
parents: 63101
diff changeset
  2585
  with assms(1) wf show "pmf_of_list xs' = pmf_of_list xs"
0b7bdb75f451 Added code generation for PMFs
eberlm
parents: 63101
diff changeset
  2586
    by (intro pmf_eqI) (simp_all add: pmf_pmf_of_list)
0b7bdb75f451 Added code generation for PMFs
eberlm
parents: 63101
diff changeset
  2587
qed
0b7bdb75f451 Added code generation for PMFs
eberlm
parents: 63101
diff changeset
  2588
63099
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63092
diff changeset
  2589
end