src/HOL/Probability/Probability_Mass_Function.thy
author wenzelm
Mon, 19 Jan 2015 20:39:01 +0100
changeset 59409 b7cfe12acf2e
parent 59327 8a779359df67
child 59425 c5e79df8cc21
permissions -rw-r--r--
tuned;
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
59023
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
     2
    Author:     Johannes Hölzl, TU München 
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
     3
    Author:     Andreas Lochbihler, ETH Zurich
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
     4
*)
58606
9c66f7c541fb add Giry monad
hoelzl
parents: 58587
diff changeset
     5
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
     6
section \<open> Probability mass function \<close>
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
     7
58587
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
     8
theory Probability_Mass_Function
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
     9
imports
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    10
  Giry_Monad
59093
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
    11
  "~~/src/HOL/Number_Theory/Binomial"
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    12
  "~~/src/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
59052
a05c8305781e projections of pair_pmf (by D. Traytel)
hoelzl
parents: 59048
diff changeset
    15
lemma bind_return'': "sets M = sets N \<Longrightarrow> M \<guillemotright>= return N = M"
a05c8305781e projections of pair_pmf (by D. Traytel)
hoelzl
parents: 59048
diff changeset
    16
   by (cases "space M = {}")
a05c8305781e projections of pair_pmf (by D. Traytel)
hoelzl
parents: 59048
diff changeset
    17
      (simp_all add: bind_empty space_empty[symmetric] bind_nonempty join_return'
a05c8305781e projections of pair_pmf (by D. Traytel)
hoelzl
parents: 59048
diff changeset
    18
                cong: subprob_algebra_cong)
a05c8305781e projections of pair_pmf (by D. Traytel)
hoelzl
parents: 59048
diff changeset
    19
a05c8305781e projections of pair_pmf (by D. Traytel)
hoelzl
parents: 59048
diff changeset
    20
a05c8305781e projections of pair_pmf (by D. Traytel)
hoelzl
parents: 59048
diff changeset
    21
lemma (in prob_space) distr_const[simp]:
a05c8305781e projections of pair_pmf (by D. Traytel)
hoelzl
parents: 59048
diff changeset
    22
  "c \<in> space N \<Longrightarrow> distr M N (\<lambda>x. c) = return N c"
a05c8305781e projections of pair_pmf (by D. Traytel)
hoelzl
parents: 59048
diff changeset
    23
  by (rule measure_eqI) (auto simp: emeasure_distr emeasure_space_1)
a05c8305781e projections of pair_pmf (by D. Traytel)
hoelzl
parents: 59048
diff changeset
    24
a05c8305781e projections of pair_pmf (by D. Traytel)
hoelzl
parents: 59048
diff changeset
    25
lemma (in finite_measure) countable_support:
58587
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
    26
  "countable {x. measure M {x} \<noteq> 0}"
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    27
proof cases
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    28
  assume "measure M (space M) = 0"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    29
  with bounded_measure measure_le_0_iff have "{x. measure M {x} \<noteq> 0} = {}"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    30
    by auto
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    31
  then show ?thesis
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    32
    by simp
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    33
next
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    34
  let ?M = "measure M (space M)" and ?m = "\<lambda>x. measure M {x}"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    35
  assume "?M \<noteq> 0"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    36
  then have *: "{x. ?m x \<noteq> 0} = (\<Union>n. {x. ?M / Suc n < ?m x})"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    37
    using reals_Archimedean[of "?m x / ?M" for x]
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    38
    by (auto simp: field_simps not_le[symmetric] measure_nonneg divide_le_0_iff measure_le_0_iff)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    39
  have **: "\<And>n. finite {x. ?M / Suc n < ?m x}"
58587
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
    40
  proof (rule ccontr)
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    41
    fix n assume "infinite {x. ?M / Suc n < ?m x}" (is "infinite ?X")
58587
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
    42
    then obtain X where "finite X" "card X = Suc (Suc n)" "X \<subseteq> ?X"
58606
9c66f7c541fb add Giry monad
hoelzl
parents: 58587
diff changeset
    43
      by (metis infinite_arbitrarily_large)
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    44
    from this(3) have *: "\<And>x. x \<in> X \<Longrightarrow> ?M / Suc n \<le> ?m x" 
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    45
      by auto
58587
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
    46
    { fix x assume "x \<in> X"
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    47
      from `?M \<noteq> 0` *[OF this] have "?m x \<noteq> 0" by (auto simp: field_simps measure_le_0_iff)
58587
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
    48
      then have "{x} \<in> sets M" by (auto dest: measure_notin_sets) }
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
    49
    note singleton_sets = this
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    50
    have "?M < (\<Sum>x\<in>X. ?M / Suc n)"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    51
      using `?M \<noteq> 0` 
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    52
      by (simp add: `card X = Suc (Suc n)` real_eq_of_nat[symmetric] real_of_nat_Suc field_simps less_le measure_nonneg)
58587
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
    53
    also have "\<dots> \<le> (\<Sum>x\<in>X. ?m x)"
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
    54
      by (rule setsum_mono) fact
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
    55
    also have "\<dots> = measure M (\<Union>x\<in>X. {x})"
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
    56
      using singleton_sets `finite X`
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
    57
      by (intro finite_measure_finite_Union[symmetric]) (auto simp: disjoint_family_on_def)
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    58
    finally have "?M < measure M (\<Union>x\<in>X. {x})" .
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    59
    moreover have "measure M (\<Union>x\<in>X. {x}) \<le> ?M"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    60
      using singleton_sets[THEN sets.sets_into_space] by (intro finite_measure_mono) auto
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    61
    ultimately show False by simp
58587
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
    62
  qed
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
    63
  show ?thesis
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
    64
    unfolding * by (intro countable_UN countableI_type countable_finite[OF **])
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
    65
qed
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
    66
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    67
lemma (in finite_measure) AE_support_countable:
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    68
  assumes [simp]: "sets M = UNIV"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    69
  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
    70
proof
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    71
  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
    72
  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
    73
    by auto
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    74
  then have "emeasure M (\<Union>x\<in>{x\<in>S. emeasure M {x} \<noteq> 0}. {x}) = 
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    75
    (\<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
    76
    by (subst emeasure_UN_countable)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    77
       (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
    78
  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
    79
    by (auto intro!: nn_integral_cong split: split_indicator)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    80
  also have "\<dots> = emeasure M (\<Union>x\<in>S. {x})"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    81
    by (subst emeasure_UN_countable)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    82
       (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
    83
  also have "\<dots> = emeasure M (space M)"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    84
    using ae by (intro emeasure_eq_AE) auto
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    85
  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
    86
    by (simp add: emeasure_single_in_space cong: rev_conj_cong)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    87
  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
    88
  have "AE x in M. x \<in> S \<and> emeasure M {x} \<noteq> 0"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    89
    by (intro AE_I[OF order_refl]) (auto simp: emeasure_eq_measure set_diff_eq cong: conj_cong)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    90
  then show "AE x in M. measure M {x} \<noteq> 0"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    91
    by (auto simp: emeasure_eq_measure)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    92
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
    93
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    94
subsection {* PMF as measure *}
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    95
58587
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
    96
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
    97
  morphisms measure_pmf Abs_pmf
58606
9c66f7c541fb add Giry monad
hoelzl
parents: 58587
diff changeset
    98
  by (intro exI[of _ "uniform_measure (count_space UNIV) {undefined}"])
9c66f7c541fb add Giry monad
hoelzl
parents: 58587
diff changeset
    99
     (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
   100
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   101
declare [[coercion measure_pmf]]
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   102
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   103
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
   104
  using pmf.measure_pmf[of p] by auto
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   105
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   106
interpretation measure_pmf!: prob_space "measure_pmf M" for M
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   107
  by (rule prob_space_measure_pmf)
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   108
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   109
interpretation measure_pmf!: subprob_space "measure_pmf M" for M
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   110
  by (rule prob_space_imp_subprob_space) unfold_locales
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   111
59048
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59024
diff changeset
   112
lemma subprob_space_measure_pmf: "subprob_space (measure_pmf x)"
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59024
diff changeset
   113
  by unfold_locales
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59024
diff changeset
   114
58587
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   115
locale pmf_as_measure
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   116
begin
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   117
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   118
setup_lifting type_definition_pmf
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   119
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   120
end
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   121
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   122
context
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   123
begin
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   124
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   125
interpretation pmf_as_measure .
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   126
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   127
lift_definition pmf :: "'a pmf \<Rightarrow> 'a \<Rightarrow> real" is "\<lambda>M x. measure M {x}" .
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   128
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   129
lift_definition set_pmf :: "'a pmf \<Rightarrow> 'a set" is "\<lambda>M. {x. measure M {x} \<noteq> 0}" .
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   130
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   131
lift_definition map_pmf :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a pmf \<Rightarrow> 'b pmf" is
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   132
  "\<lambda>f M. distr M (count_space UNIV) f"
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   133
proof safe
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   134
  fix M and f :: "'a \<Rightarrow> 'b"
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   135
  let ?D = "distr M (count_space UNIV) f"
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   136
  assume "prob_space M" and [simp]: "sets M = UNIV" and ae: "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
   137
  interpret prob_space M by fact
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   138
  from ae have "AE x in M. measure M (f -` {f x}) \<noteq> 0"
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   139
  proof eventually_elim
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   140
    fix x
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   141
    have "measure M {x} \<le> measure M (f -` {f x})"
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   142
      by (intro finite_measure_mono) auto
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   143
    then show "measure M {x} \<noteq> 0 \<Longrightarrow> measure M (f -` {f x}) \<noteq> 0"
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   144
      using measure_nonneg[of M "{x}"] by auto
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   145
  qed
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   146
  then show "AE x in ?D. measure ?D {x} \<noteq> 0"
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   147
    by (simp add: AE_distr_iff measure_distr measurable_def)
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   148
qed (auto simp: measurable_def prob_space.prob_space_distr)
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   149
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   150
declare [[coercion set_pmf]]
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   151
59023
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
   152
lemma countable_set_pmf [simp]: "countable (set_pmf p)"
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   153
  by transfer (metis prob_space.finite_measure finite_measure.countable_support)
58587
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   154
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   155
lemma sets_measure_pmf[simp]: "sets (measure_pmf p) = UNIV"
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   156
  by transfer metis
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   157
59048
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59024
diff changeset
   158
lemma sets_measure_pmf_count_space[measurable_cong]:
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59024
diff changeset
   159
  "sets (measure_pmf M) = sets (count_space UNIV)"
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   160
  by simp
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   161
58587
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   162
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
   163
  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
   164
59048
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59024
diff changeset
   165
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
   166
  by (simp add: space_subprob_algebra subprob_space_measure_pmf)
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59024
diff changeset
   167
58587
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   168
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
   169
  by (auto simp: measurable_def)
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   170
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   171
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
   172
  by (intro measurable_cong_sets) simp_all
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   173
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   174
lemma pmf_positive: "x \<in> set_pmf p \<Longrightarrow> 0 < pmf p x"
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   175
  by transfer (simp add: less_le measure_nonneg)
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   176
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   177
lemma pmf_nonneg: "0 \<le> pmf p x"
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   178
  by transfer (simp add: measure_nonneg)
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   179
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   180
lemma pmf_le_1: "pmf p x \<le> 1"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   181
  by (simp add: pmf.rep_eq)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   182
58587
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   183
lemma emeasure_pmf_single:
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   184
  fixes M :: "'a pmf"
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   185
  shows "emeasure M {x} = pmf M x"
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   186
  by transfer (simp add: finite_measure.emeasure_eq_measure[OF prob_space.finite_measure])
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   187
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   188
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
   189
  by transfer simp
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   190
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   191
lemma emeasure_pmf_single_eq_zero_iff:
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   192
  fixes M :: "'a pmf"
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   193
  shows "emeasure M {y} = 0 \<longleftrightarrow> y \<notin> M"
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   194
  by transfer (simp add: finite_measure.emeasure_eq_measure[OF prob_space.finite_measure])
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   195
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   196
lemma AE_measure_pmf_iff: "(AE x in measure_pmf M. P x) \<longleftrightarrow> (\<forall>y\<in>M. P y)"
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   197
proof -
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   198
  { fix y assume y: "y \<in> M" and P: "AE x in M. P x" "\<not> P y"
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   199
    with P have "AE x in M. x \<noteq> y"
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   200
      by auto
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   201
    with y have False
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   202
      by (simp add: emeasure_pmf_single_eq_zero_iff AE_iff_measurable[OF _ refl]) }
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   203
  then show ?thesis
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   204
    using AE_measure_pmf[of M] by auto
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   205
qed
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   206
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   207
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
   208
  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
   209
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   210
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
   211
  by transfer simp
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   212
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   213
lemma emeasure_measure_pmf_finite: "finite S \<Longrightarrow> emeasure (measure_pmf M) S = (\<Sum>s\<in>S. pmf M s)"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   214
  by (subst emeasure_eq_setsum_singleton) (auto simp: emeasure_pmf_single)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   215
59023
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
   216
lemma measure_measure_pmf_finite: "finite S \<Longrightarrow> measure (measure_pmf M) S = setsum (pmf M) S"
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
   217
using emeasure_measure_pmf_finite[of S M]
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
   218
by(simp add: measure_pmf.emeasure_eq_measure)
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
   219
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   220
lemma nn_integral_measure_pmf_support:
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   221
  fixes f :: "'a \<Rightarrow> ereal"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   222
  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
   223
  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
   224
proof -
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   225
  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
   226
    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
   227
  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
   228
    using assms by (intro nn_integral_indicator_finite) auto
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   229
  finally show ?thesis
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   230
    by (simp add: emeasure_measure_pmf_finite)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   231
qed
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   232
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   233
lemma nn_integral_measure_pmf_finite:
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   234
  fixes f :: "'a \<Rightarrow> ereal"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   235
  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
   236
  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
   237
  using assms by (intro nn_integral_measure_pmf_support) auto
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   238
lemma integrable_measure_pmf_finite:
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   239
  fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   240
  shows "finite (set_pmf M) \<Longrightarrow> integrable M f"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   241
  by (auto intro!: integrableI_bounded simp: nn_integral_measure_pmf_finite)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   242
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   243
lemma integral_measure_pmf:
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   244
  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
   245
  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
   246
proof -
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   247
  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
   248
    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
   249
  also have "\<dots> = (\<Sum>a\<in>A. f a * pmf M a)"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   250
    by (subst integral_indicator_finite_real) (auto simp: measure_def emeasure_measure_pmf_finite)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   251
  finally show ?thesis .
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   252
qed
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   253
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   254
lemma integrable_pmf: "integrable (count_space X) (pmf M)"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   255
proof -
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   256
  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
   257
    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
   258
  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
   259
    by (simp add: integrable_iff_bounded pmf_nonneg)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   260
  then show ?thesis
59023
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
   261
    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
   262
qed
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   263
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   264
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
   265
proof -
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   266
  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
   267
    by (simp add: pmf_nonneg integrable_pmf nn_integral_eq_integral)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   268
  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
   269
    by (auto intro!: nn_integral_cong_AE split: split_indicator
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   270
             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
   271
                   AE_count_space set_pmf_iff)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   272
  also have "\<dots> = emeasure M (X \<inter> M)"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   273
    by (rule emeasure_countable_singleton[symmetric]) (auto intro: countable_set_pmf)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   274
  also have "\<dots> = emeasure M X"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   275
    by (auto intro!: emeasure_eq_AE simp: AE_measure_pmf_iff)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   276
  finally show ?thesis
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   277
    by (simp add: measure_pmf.emeasure_eq_measure)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   278
qed
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   279
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   280
lemma integral_pmf_restrict:
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   281
  "(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
   282
    (\<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
   283
  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
   284
58587
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   285
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
   286
proof -
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   287
  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
   288
    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
   289
  then show ?thesis
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   290
    using measure_pmf.emeasure_space_1 by simp
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   291
qed
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   292
59023
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
   293
lemma in_null_sets_measure_pmfI:
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
   294
  "A \<inter> set_pmf p = {} \<Longrightarrow> A \<in> null_sets (measure_pmf p)"
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
   295
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
   296
by(auto simp add: null_sets_def AE_measure_pmf_iff)
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
   297
58587
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   298
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
   299
  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
   300
59053
43e07797269b tuned proof that pmfs are bnfs
hoelzl
parents: 59052
diff changeset
   301
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
   302
  using map_pmf_id unfolding id_def .
43e07797269b tuned proof that pmfs are bnfs
hoelzl
parents: 59052
diff changeset
   303
58587
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   304
lemma map_pmf_compose: "map_pmf (f \<circ> g) = map_pmf f \<circ> map_pmf g"
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   305
  by (rule, transfer) (simp add: distr_distr[symmetric, where N="count_space UNIV"] measurable_def) 
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   306
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   307
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
   308
  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
   309
58587
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   310
lemma map_pmf_cong:
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   311
  assumes "p = q"
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   312
  shows "(\<And>x. x \<in> set_pmf q \<Longrightarrow> f x = g x) \<Longrightarrow> map_pmf f p = map_pmf g q"
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   313
  unfolding `p = q`[symmetric] measure_pmf_inject[symmetric] map_pmf.rep_eq
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   314
  by (auto simp add: emeasure_distr AE_measure_pmf_iff intro!: emeasure_eq_AE measure_eqI)
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   315
59002
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   316
lemma emeasure_map_pmf[simp]: "emeasure (map_pmf f M) X = emeasure M (f -` X)"
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   317
  unfolding map_pmf.rep_eq by (subst emeasure_distr) auto
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   318
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   319
lemma nn_integral_map_pmf[simp]: "(\<integral>\<^sup>+x. f x \<partial>map_pmf g M) = (\<integral>\<^sup>+x. f (g x) \<partial>M)"
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   320
  unfolding map_pmf.rep_eq by (intro nn_integral_distr) auto
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   321
59023
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
   322
lemma ereal_pmf_map: "pmf (map_pmf f p) x = (\<integral>\<^sup>+ y. indicator (f -` {x}) y \<partial>measure_pmf p)"
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
   323
proof(transfer fixing: f x)
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
   324
  fix p :: "'b measure"
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
   325
  presume "prob_space p"
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
   326
  then interpret prob_space p .
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
   327
  presume "sets p = UNIV"
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
   328
  then show "ereal (measure (distr p (count_space UNIV) f) {x}) = integral\<^sup>N p (indicator (f -` {x}))"
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
   329
    by(simp add: measure_distr measurable_def emeasure_eq_measure)
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
   330
qed simp_all
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
   331
58587
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   332
lemma pmf_set_map: 
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   333
  fixes f :: "'a \<Rightarrow> 'b"
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   334
  shows "set_pmf \<circ> map_pmf f = op ` f \<circ> set_pmf"
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   335
proof (rule, transfer, clarsimp simp add: measure_distr measurable_def)
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   336
  fix f :: "'a \<Rightarrow> 'b" and M :: "'a measure"
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   337
  assume "prob_space M" and ae: "AE x in M. measure M {x} \<noteq> 0" and [simp]: "sets M = UNIV"
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   338
  interpret prob_space M by fact
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   339
  show "{x. measure M (f -` {x}) \<noteq> 0} = f ` {x. measure M {x} \<noteq> 0}"
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   340
  proof safe
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   341
    fix x assume "measure M (f -` {x}) \<noteq> 0"
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   342
    moreover have "measure M (f -` {x}) = measure M {y. f y = x \<and> measure M {y} \<noteq> 0}"
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   343
      using ae by (intro finite_measure_eq_AE) auto
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   344
    ultimately have "{y. f y = x \<and> measure M {y} \<noteq> 0} \<noteq> {}"
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   345
      by (metis measure_empty)
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   346
    then show "x \<in> f ` {x. measure M {x} \<noteq> 0}"
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   347
      by auto
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   348
  next
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   349
    fix x assume "measure M {x} \<noteq> 0"
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   350
    then have "0 < measure M {x}"
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   351
      using measure_nonneg[of M "{x}"] by auto
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   352
    also have "measure M {x} \<le> measure M (f -` {f x})"
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   353
      by (intro finite_measure_mono) auto
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   354
    finally show "measure M (f -` {f x}) = 0 \<Longrightarrow> False"
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   355
      by simp
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   356
  qed
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   357
qed
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   358
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   359
lemma set_map_pmf: "set_pmf (map_pmf f M) = f`set_pmf M"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   360
  using pmf_set_map[of f] by (auto simp: comp_def fun_eq_iff)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   361
59023
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
   362
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
   363
proof -
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
   364
  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
   365
    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
   366
  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
   367
    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
   368
  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
   369
    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
   370
  also have "\<dots> = emeasure (measure_pmf p) A"
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
   371
    by(auto intro: arg_cong2[where f=emeasure])
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
   372
  finally show ?thesis .
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
   373
qed
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
   374
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   375
subsection {* PMFs as function *}
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   376
58587
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   377
context
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   378
  fixes f :: "'a \<Rightarrow> real"
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   379
  assumes nonneg: "\<And>x. 0 \<le> f x"
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   380
  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
   381
begin
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   382
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   383
lift_definition embed_pmf :: "'a pmf" is "density (count_space UNIV) (ereal \<circ> f)"
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   384
proof (intro conjI)
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   385
  have *[simp]: "\<And>x y. ereal (f y) * indicator {x} y = ereal (f x) * indicator {x} y"
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   386
    by (simp split: split_indicator)
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   387
  show "AE x in density (count_space UNIV) (ereal \<circ> f).
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   388
    measure (density (count_space UNIV) (ereal \<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
   389
    by (simp add: AE_density nonneg measure_def emeasure_density max_def)
58587
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   390
  show "prob_space (density (count_space UNIV) (ereal \<circ> f))"
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   391
    by default (simp add: emeasure_density prob)
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   392
qed simp
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   393
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   394
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
   395
proof transfer
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   396
  have *[simp]: "\<And>x y. ereal (f y) * indicator {x} y = ereal (f x) * indicator {x} y"
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   397
    by (simp split: split_indicator)
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   398
  fix x show "measure (density (count_space UNIV) (ereal \<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
   399
    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
   400
qed
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   401
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   402
end
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   403
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   404
lemma embed_pmf_transfer:
58730
b3fd0628f849 add transfer rule for set_pmf
hoelzl
parents: 58606
diff changeset
   405
  "rel_fun (eq_onp (\<lambda>f. (\<forall>x. 0 \<le> f x) \<and> (\<integral>\<^sup>+x. ereal (f x) \<partial>count_space UNIV) = 1)) pmf_as_measure.cr_pmf (\<lambda>f. density (count_space UNIV) (ereal \<circ> f)) embed_pmf"
58587
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   406
  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
   407
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   408
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
   409
proof (transfer, elim conjE)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   410
  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
   411
  assume "prob_space M" then interpret prob_space M .
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   412
  show "M = density (count_space UNIV) (\<lambda>x. ereal (measure M {x}))"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   413
  proof (rule measure_eqI)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   414
    fix A :: "'a set"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   415
    have "(\<integral>\<^sup>+ x. ereal (measure M {x}) * indicator A x \<partial>count_space UNIV) = 
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   416
      (\<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
   417
      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
   418
    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
   419
      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
   420
    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
   421
      by (intro emeasure_UN_countable[symmetric] countable_Int2 countable_support)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   422
         (auto simp: disjoint_family_on_def)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   423
    also have "\<dots> = emeasure M A"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   424
      using ae by (intro emeasure_eq_AE) auto
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   425
    finally show " emeasure M A = emeasure (density (count_space UNIV) (\<lambda>x. ereal (measure M {x}))) A"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   426
      using emeasure_space_1 by (simp add: emeasure_density)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   427
  qed simp
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   428
qed
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   429
58587
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   430
lemma td_pmf_embed_pmf:
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   431
  "type_definition pmf embed_pmf {f::'a \<Rightarrow> real. (\<forall>x. 0 \<le> f x) \<and> (\<integral>\<^sup>+x. ereal (f x) \<partial>count_space UNIV) = 1}"
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   432
  unfolding type_definition_def
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   433
proof safe
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   434
  fix p :: "'a pmf"
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   435
  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
   436
    using measure_pmf.emeasure_space_1[of p] by simp
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   437
  then show *: "(\<integral>\<^sup>+ x. ereal (pmf p x) \<partial>count_space UNIV) = 1"
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   438
    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
   439
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   440
  show "embed_pmf (pmf p) = p"
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   441
    by (intro measure_pmf_inject[THEN iffD1])
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   442
       (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
   443
next
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   444
  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
   445
  then show "pmf (embed_pmf f) = f"
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   446
    by (auto intro!: pmf_embed_pmf)
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   447
qed (rule pmf_nonneg)
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   448
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   449
end
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   450
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   451
locale pmf_as_function
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   452
begin
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   453
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   454
setup_lifting td_pmf_embed_pmf
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   455
58730
b3fd0628f849 add transfer rule for set_pmf
hoelzl
parents: 58606
diff changeset
   456
lemma set_pmf_transfer[transfer_rule]: 
b3fd0628f849 add transfer rule for set_pmf
hoelzl
parents: 58606
diff changeset
   457
  assumes "bi_total A"
b3fd0628f849 add transfer rule for set_pmf
hoelzl
parents: 58606
diff changeset
   458
  shows "rel_fun (pcr_pmf A) (rel_set A) (\<lambda>f. {x. f x \<noteq> 0}) set_pmf"  
b3fd0628f849 add transfer rule for set_pmf
hoelzl
parents: 58606
diff changeset
   459
  using `bi_total A`
b3fd0628f849 add transfer rule for set_pmf
hoelzl
parents: 58606
diff changeset
   460
  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
   461
     metis+
b3fd0628f849 add transfer rule for set_pmf
hoelzl
parents: 58606
diff changeset
   462
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   463
end
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   464
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   465
context
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   466
begin
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   467
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   468
interpretation pmf_as_function .
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   469
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   470
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
   471
  by transfer auto
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   472
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   473
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
   474
  by (auto intro: pmf_eqI)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   475
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   476
end
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   477
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   478
context
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   479
begin
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   480
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   481
interpretation pmf_as_function .
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   482
59093
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
   483
subsubsection \<open> Bernoulli Distribution \<close>
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
   484
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   485
lift_definition bernoulli_pmf :: "real \<Rightarrow> bool pmf" is
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   486
  "\<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
   487
  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
   488
           split: split_max split_min)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   489
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   490
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
   491
  by transfer simp
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   492
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   493
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
   494
  by transfer simp
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   495
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   496
lemma set_pmf_bernoulli: "0 < p \<Longrightarrow> p < 1 \<Longrightarrow> set_pmf (bernoulli_pmf p) = UNIV"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   497
  by (auto simp add: set_pmf_iff UNIV_bool)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   498
59002
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   499
lemma nn_integral_bernoulli_pmf[simp]: 
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   500
  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
   501
  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
   502
  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
   503
     (auto simp: UNIV_bool field_simps)
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   504
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   505
lemma integral_bernoulli_pmf[simp]: 
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   506
  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
   507
  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
   508
  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
   509
59093
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
   510
subsubsection \<open> Geometric Distribution \<close>
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
   511
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   512
lift_definition geometric_pmf :: "nat pmf" is "\<lambda>n. 1 / 2^Suc n"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   513
proof
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   514
  note geometric_sums[of "1 / 2"]
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   515
  note sums_mult[OF this, of "1 / 2"]
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   516
  from sums_suminf_ereal[OF this]
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   517
  show "(\<integral>\<^sup>+ x. ereal (1 / 2 ^ Suc x) \<partial>count_space UNIV) = 1"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   518
    by (simp add: nn_integral_count_space_nat field_simps)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   519
qed simp
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   520
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   521
lemma pmf_geometric[simp]: "pmf geometric_pmf n = 1 / 2^Suc n"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   522
  by transfer rule
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   523
59002
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   524
lemma set_pmf_geometric[simp]: "set_pmf geometric_pmf = UNIV"
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   525
  by (auto simp: set_pmf_iff)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   526
59093
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
   527
subsubsection \<open> Uniform Multiset Distribution \<close>
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
   528
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   529
context
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   530
  fixes M :: "'a multiset" assumes M_not_empty: "M \<noteq> {#}"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   531
begin
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   532
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   533
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
   534
proof
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   535
  show "(\<integral>\<^sup>+ x. ereal (real (count M x) / real (size M)) \<partial>count_space UNIV) = 1"  
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   536
    using M_not_empty
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   537
    by (simp add: zero_less_divide_iff nn_integral_count_space nonempty_has_size
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   538
                  setsum_divide_distrib[symmetric])
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   539
       (auto simp: size_multiset_overloaded_eq intro!: setsum.cong)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   540
qed simp
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   541
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   542
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
   543
  by transfer rule
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   544
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   545
lemma set_pmf_of_multiset[simp]: "set_pmf pmf_of_multiset = set_of M"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   546
  by (auto simp: set_pmf_iff)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   547
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   548
end
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   549
59093
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
   550
subsubsection \<open> Uniform Distribution \<close>
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
   551
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   552
context
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   553
  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
   554
begin
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   555
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   556
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
   557
proof
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   558
  show "(\<integral>\<^sup>+ x. ereal (indicator S x / real (card S)) \<partial>count_space UNIV) = 1"  
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   559
    using S_not_empty S_finite by (subst nn_integral_count_space'[of S]) auto
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   560
qed simp
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   561
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   562
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
   563
  by transfer rule
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   564
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   565
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
   566
  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
   567
59002
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   568
lemma emeasure_pmf_of_set[simp]: "emeasure pmf_of_set S = 1"
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   569
  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
   570
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   571
end
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   572
59093
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
   573
subsubsection \<open> Poisson Distribution \<close>
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
   574
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
   575
context
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
   576
  fixes rate :: real assumes rate_pos: "0 < rate"
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
   577
begin
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
   578
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
   579
lift_definition poisson_pmf :: "nat pmf" is "\<lambda>k. rate ^ k / fact k * exp (-rate)"
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
   580
proof
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
   581
  (* Proof by Manuel Eberl *)
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
   582
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
   583
  have summable: "summable (\<lambda>x::nat. rate ^ x / fact x)" using summable_exp
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
   584
    by (simp add: field_simps field_divide_inverse[symmetric])
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
   585
  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
   586
          exp (-rate) * (\<integral>\<^sup>+(x::nat). rate ^ x / fact x \<partial>count_space UNIV)"
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
   587
    by (simp add: field_simps nn_integral_cmult[symmetric])
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
   588
  also from rate_pos have "(\<integral>\<^sup>+(x::nat). rate ^ x / fact x \<partial>count_space UNIV) = (\<Sum>x. rate ^ x / fact x)"
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
   589
    by (simp_all add: nn_integral_count_space_nat suminf_ereal summable suminf_ereal_finite)
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
   590
  also have "... = exp rate" unfolding exp_def
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
   591
    by (simp add: field_simps field_divide_inverse[symmetric] transfer_int_nat_factorial)
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
   592
  also have "ereal (exp (-rate)) * ereal (exp rate) = 1"
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
   593
    by (simp add: mult_exp_exp)
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
   594
  finally show "(\<integral>\<^sup>+ x. ereal (rate ^ x / real (fact x) * exp (- rate)) \<partial>count_space UNIV) = 1" .
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
   595
qed (simp add: rate_pos[THEN less_imp_le])
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
   596
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
   597
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
   598
  by transfer rule
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
   599
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
   600
lemma set_pmf_poisson[simp]: "set_pmf poisson_pmf = UNIV"
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
   601
  using rate_pos by (auto simp: set_pmf_iff)
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
   602
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   603
end
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   604
59093
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
   605
subsubsection \<open> Binomial Distribution \<close>
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
   606
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
   607
context
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
   608
  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
   609
begin
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
   610
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
   611
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
   612
proof
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
   613
  have "(\<integral>\<^sup>+k. ereal (real (n choose k) * p ^ k * (1 - p) ^ (n - k)) \<partial>count_space UNIV) =
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
   614
    ereal (\<Sum>k\<le>n. real (n choose k) * p ^ k * (1 - p) ^ (n - k))"
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
   615
    using p_le_1 p_nonneg by (subst nn_integral_count_space') auto
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
   616
  also have "(\<Sum>k\<le>n. real (n choose k) * p ^ k * (1 - p) ^ (n - k)) = (p + (1 - p)) ^ n"
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
   617
    by (subst binomial_ring) (simp add: atLeast0AtMost real_of_nat_def)
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
   618
  finally show "(\<integral>\<^sup>+ x. ereal (real (n choose x) * p ^ x * (1 - p) ^ (n - x)) \<partial>count_space UNIV) = 1"
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
   619
    by simp
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
   620
qed (insert p_nonneg p_le_1, simp)
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
   621
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
   622
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
   623
  by transfer rule
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
   624
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
   625
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
   626
  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
   627
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
   628
end
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
   629
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
   630
end
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
   631
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
   632
lemma set_pmf_binomial_0[simp]: "set_pmf (binomial_pmf n 0) = {0}"
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
   633
  by (simp add: set_pmf_binomial_eq)
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
   634
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
   635
lemma set_pmf_binomial_1[simp]: "set_pmf (binomial_pmf n 1) = {n}"
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
   636
  by (simp add: set_pmf_binomial_eq)
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
   637
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
   638
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
   639
  by (simp add: set_pmf_binomial_eq)
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
   640
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
   641
subsection \<open> Monad Interpretation \<close>
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   642
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   643
lemma measurable_measure_pmf[measurable]:
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   644
  "(\<lambda>x. measure_pmf (M x)) \<in> measurable (count_space UNIV) (subprob_algebra (count_space UNIV))"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   645
  by (auto simp: space_subprob_algebra intro!: prob_space_imp_subprob_space) unfold_locales
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   646
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   647
lemma bind_pmf_cong:
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   648
  assumes "\<And>x. A x \<in> space (subprob_algebra N)" "\<And>x. B x \<in> space (subprob_algebra N)"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   649
  assumes "\<And>i. i \<in> set_pmf x \<Longrightarrow> A i = B i"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   650
  shows "bind (measure_pmf x) A = bind (measure_pmf x) B"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   651
proof (rule measure_eqI)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   652
  show "sets (measure_pmf x \<guillemotright>= A) = sets (measure_pmf x \<guillemotright>= B)"
59048
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59024
diff changeset
   653
    using assms by (subst (1 2) sets_bind) (auto simp: space_subprob_algebra)
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   654
next
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   655
  fix X assume "X \<in> sets (measure_pmf x \<guillemotright>= A)"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   656
  then have X: "X \<in> sets N"
59048
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59024
diff changeset
   657
    using assms by (subst (asm) sets_bind) (auto simp: space_subprob_algebra)
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   658
  show "emeasure (measure_pmf x \<guillemotright>= A) X = emeasure (measure_pmf x \<guillemotright>= B) X"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   659
    using assms
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   660
    by (subst (1 2) emeasure_bind[where N=N, OF _ _ X])
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   661
       (auto intro!: nn_integral_cong_AE simp: AE_measure_pmf_iff)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   662
qed
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   663
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   664
context
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   665
begin
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   666
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   667
interpretation pmf_as_measure .
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   668
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   669
lift_definition join_pmf :: "'a pmf pmf \<Rightarrow> 'a pmf" is "\<lambda>M. measure_pmf M \<guillemotright>= measure_pmf"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   670
proof (intro conjI)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   671
  fix M :: "'a pmf pmf"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   672
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   673
  interpret bind: prob_space "measure_pmf M \<guillemotright>= measure_pmf"
59048
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59024
diff changeset
   674
    apply (intro measure_pmf.prob_space_bind[where S="count_space UNIV"] AE_I2)
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59024
diff changeset
   675
    apply (auto intro!: subprob_space_measure_pmf simp: space_subprob_algebra)
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   676
    apply unfold_locales
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   677
    done
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   678
  show "prob_space (measure_pmf M \<guillemotright>= measure_pmf)"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   679
    by intro_locales
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   680
  show "sets (measure_pmf M \<guillemotright>= measure_pmf) = UNIV"
59048
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59024
diff changeset
   681
    by (subst sets_bind) auto
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   682
  have "AE x in measure_pmf M \<guillemotright>= measure_pmf. emeasure (measure_pmf M \<guillemotright>= measure_pmf) {x} \<noteq> 0"
59048
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59024
diff changeset
   683
    by (auto simp: AE_bind[where B="count_space UNIV"] measure_pmf_in_subprob_algebra
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59024
diff changeset
   684
                   emeasure_bind[where N="count_space UNIV"] AE_measure_pmf_iff nn_integral_0_iff_AE
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59024
diff changeset
   685
                   measure_pmf.emeasure_eq_measure measure_le_0_iff set_pmf_iff pmf.rep_eq)
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   686
  then show "AE x in measure_pmf M \<guillemotright>= measure_pmf. measure (measure_pmf M \<guillemotright>= measure_pmf) {x} \<noteq> 0"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   687
    unfolding bind.emeasure_eq_measure by simp
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   688
qed
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   689
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   690
lemma pmf_join: "pmf (join_pmf N) i = (\<integral>M. pmf M i \<partial>measure_pmf N)"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   691
proof (transfer fixing: N i)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   692
  have N: "subprob_space (measure_pmf N)"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   693
    by (rule prob_space_imp_subprob_space) intro_locales
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   694
  show "measure (measure_pmf N \<guillemotright>= measure_pmf) {i} = integral\<^sup>L (measure_pmf N) (\<lambda>M. measure M {i})"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   695
    using measurable_measure_pmf[of "\<lambda>x. x"]
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   696
    by (intro subprob_space.measure_bind[where N="count_space UNIV", OF N]) auto
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   697
qed (auto simp: Transfer.Rel_def rel_fun_def cr_pmf_def)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   698
59024
5fcfeae84b96 add lemma
Andreas Lochbihler
parents: 59023
diff changeset
   699
lemma set_pmf_join_pmf: "set_pmf (join_pmf f) = (\<Union>p\<in>set_pmf f. set_pmf p)"
5fcfeae84b96 add lemma
Andreas Lochbihler
parents: 59023
diff changeset
   700
apply(simp add: set_eq_iff set_pmf_iff pmf_join)
5fcfeae84b96 add lemma
Andreas Lochbihler
parents: 59023
diff changeset
   701
apply(subst integral_nonneg_eq_0_iff_AE)
5fcfeae84b96 add lemma
Andreas Lochbihler
parents: 59023
diff changeset
   702
apply(auto simp add: pmf_le_1 pmf_nonneg AE_measure_pmf_iff intro!: measure_pmf.integrable_const_bound[where B=1])
5fcfeae84b96 add lemma
Andreas Lochbihler
parents: 59023
diff changeset
   703
done
5fcfeae84b96 add lemma
Andreas Lochbihler
parents: 59023
diff changeset
   704
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   705
lift_definition return_pmf :: "'a \<Rightarrow> 'a pmf" is "return (count_space UNIV)"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   706
  by (auto intro!: prob_space_return simp: AE_return measure_return)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   707
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   708
lemma join_return_pmf: "join_pmf (return_pmf M) = M"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   709
  by (simp add: integral_return pmf_eq_iff pmf_join return_pmf.rep_eq)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   710
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   711
lemma map_return_pmf: "map_pmf f (return_pmf x) = return_pmf (f x)"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   712
  by transfer (simp add: distr_return)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   713
59052
a05c8305781e projections of pair_pmf (by D. Traytel)
hoelzl
parents: 59048
diff changeset
   714
lemma map_pmf_const[simp]: "map_pmf (\<lambda>_. c) M = return_pmf c"
a05c8305781e projections of pair_pmf (by D. Traytel)
hoelzl
parents: 59048
diff changeset
   715
  by transfer (auto simp: prob_space.distr_const)
a05c8305781e projections of pair_pmf (by D. Traytel)
hoelzl
parents: 59048
diff changeset
   716
59002
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   717
lemma set_return_pmf: "set_pmf (return_pmf x) = {x}"
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   718
  by transfer (auto simp add: measure_return split: split_indicator)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   719
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   720
lemma pmf_return: "pmf (return_pmf x) y = indicator {y} x"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   721
  by transfer (simp add: measure_return)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   722
59002
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   723
lemma nn_integral_return_pmf[simp]: "0 \<le> f x \<Longrightarrow> (\<integral>\<^sup>+x. f x \<partial>return_pmf x) = f x"
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   724
  unfolding return_pmf.rep_eq by (intro nn_integral_return) auto
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   725
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   726
lemma emeasure_return_pmf[simp]: "emeasure (return_pmf x) X = indicator X x"
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   727
  unfolding return_pmf.rep_eq by (intro emeasure_return) auto
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   728
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   729
end
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   730
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   731
definition "bind_pmf M f = join_pmf (map_pmf f M)"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   732
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   733
lemma (in pmf_as_measure) bind_transfer[transfer_rule]:
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   734
  "rel_fun pmf_as_measure.cr_pmf (rel_fun (rel_fun op = pmf_as_measure.cr_pmf) pmf_as_measure.cr_pmf) op \<guillemotright>= bind_pmf"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   735
proof (auto simp: pmf_as_measure.cr_pmf_def rel_fun_def bind_pmf_def join_pmf.rep_eq map_pmf.rep_eq)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   736
  fix M f and g :: "'a \<Rightarrow> 'b pmf" assume "\<forall>x. f x = measure_pmf (g x)"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   737
  then have f: "f = (\<lambda>x. measure_pmf (g x))"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   738
    by auto
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   739
  show "measure_pmf M \<guillemotright>= f = distr (measure_pmf M) (count_space UNIV) g \<guillemotright>= measure_pmf"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   740
    unfolding f by (subst bind_distr[OF _ measurable_measure_pmf]) auto
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   741
qed
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   742
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   743
lemma pmf_bind: "pmf (bind_pmf N f) i = (\<integral>x. pmf (f x) i \<partial>measure_pmf N)"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   744
  by (auto intro!: integral_distr simp: bind_pmf_def pmf_join map_pmf.rep_eq)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   745
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   746
lemma bind_return_pmf: "bind_pmf (return_pmf x) f = f x"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   747
  unfolding bind_pmf_def map_return_pmf join_return_pmf ..
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   748
59052
a05c8305781e projections of pair_pmf (by D. Traytel)
hoelzl
parents: 59048
diff changeset
   749
lemma join_eq_bind_pmf: "join_pmf M = bind_pmf M id"
a05c8305781e projections of pair_pmf (by D. Traytel)
hoelzl
parents: 59048
diff changeset
   750
  by (simp add: bind_pmf_def)
a05c8305781e projections of pair_pmf (by D. Traytel)
hoelzl
parents: 59048
diff changeset
   751
a05c8305781e projections of pair_pmf (by D. Traytel)
hoelzl
parents: 59048
diff changeset
   752
lemma bind_pmf_const[simp]: "bind_pmf M (\<lambda>x. c) = c"
a05c8305781e projections of pair_pmf (by D. Traytel)
hoelzl
parents: 59048
diff changeset
   753
  unfolding bind_pmf_def map_pmf_const join_return_pmf ..
a05c8305781e projections of pair_pmf (by D. Traytel)
hoelzl
parents: 59048
diff changeset
   754
59002
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   755
lemma set_bind_pmf: "set_pmf (bind_pmf M N) = (\<Union>M\<in>set_pmf M. set_pmf (N M))"
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   756
  apply (simp add: set_eq_iff set_pmf_iff pmf_bind)
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   757
  apply (subst integral_nonneg_eq_0_iff_AE)
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   758
  apply (auto simp: pmf_nonneg pmf_le_1 AE_measure_pmf_iff
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   759
              intro!: measure_pmf.integrable_const_bound[where B=1])
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   760
  done
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   761
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   762
lemma measurable_pair_restrict_pmf2:
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   763
  assumes "countable A"
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   764
  assumes "\<And>y. y \<in> A \<Longrightarrow> (\<lambda>x. f (x, y)) \<in> measurable M L"
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   765
  shows "f \<in> measurable (M \<Otimes>\<^sub>M restrict_space (measure_pmf N) A) L"
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   766
  apply (subst measurable_cong_sets)
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   767
  apply (rule sets_pair_measure_cong sets_restrict_space_cong sets_measure_pmf_count_space refl)+
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   768
  apply (simp_all add: restrict_count_space)
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   769
  apply (subst split_eta[symmetric])
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   770
  unfolding measurable_split_conv
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   771
  apply (rule measurable_compose_countable'[OF _ measurable_snd `countable A`])
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   772
  apply (rule measurable_compose[OF measurable_fst])
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   773
  apply fact
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   774
  done
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   775
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   776
lemma measurable_pair_restrict_pmf1:
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   777
  assumes "countable A"
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   778
  assumes "\<And>x. x \<in> A \<Longrightarrow> (\<lambda>y. f (x, y)) \<in> measurable N L"
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   779
  shows "f \<in> measurable (restrict_space (measure_pmf M) A \<Otimes>\<^sub>M N) L"
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   780
  apply (subst measurable_cong_sets)
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   781
  apply (rule sets_pair_measure_cong sets_restrict_space_cong sets_measure_pmf_count_space refl)+
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   782
  apply (simp_all add: restrict_count_space)
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   783
  apply (subst split_eta[symmetric])
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   784
  unfolding measurable_split_conv
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   785
  apply (rule measurable_compose_countable'[OF _ measurable_fst `countable A`])
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   786
  apply (rule measurable_compose[OF measurable_snd])
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   787
  apply fact
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   788
  done
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   789
                                
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   790
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))"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   791
  unfolding pmf_eq_iff pmf_bind
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   792
proof
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   793
  fix i
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   794
  interpret B: prob_space "restrict_space B B"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   795
    by (intro prob_space_restrict_space measure_pmf.emeasure_eq_1_AE)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   796
       (auto simp: AE_measure_pmf_iff)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   797
  interpret A: prob_space "restrict_space A A"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   798
    by (intro prob_space_restrict_space measure_pmf.emeasure_eq_1_AE)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   799
       (auto simp: AE_measure_pmf_iff)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   800
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   801
  interpret AB: pair_prob_space "restrict_space A A" "restrict_space B B"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   802
    by unfold_locales
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   803
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   804
  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)"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   805
    by (rule integral_cong) (auto intro!: integral_pmf_restrict)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   806
  also have "\<dots> = (\<integral> x. (\<integral> y. pmf (C x y) i \<partial>restrict_space B B) \<partial>restrict_space A A)"
59002
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   807
    by (intro integral_pmf_restrict B.borel_measurable_lebesgue_integral measurable_pair_restrict_pmf2
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   808
              countable_set_pmf borel_measurable_count_space)
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   809
  also have "\<dots> = (\<integral> y. \<integral> x. pmf (C x y) i \<partial>restrict_space A A \<partial>restrict_space B B)"
59002
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   810
    by (rule AB.Fubini_integral[symmetric])
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   811
       (auto intro!: AB.integrable_const_bound[where B=1] measurable_pair_restrict_pmf2
59023
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
   812
             simp: pmf_nonneg pmf_le_1 measurable_restrict_space1)
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   813
  also have "\<dots> = (\<integral> y. \<integral> x. pmf (C x y) i \<partial>restrict_space A A \<partial>B)"
59002
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   814
    by (intro integral_pmf_restrict[symmetric] A.borel_measurable_lebesgue_integral measurable_pair_restrict_pmf2
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   815
              countable_set_pmf borel_measurable_count_space)
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   816
  also have "\<dots> = (\<integral> y. \<integral> x. pmf (C x y) i \<partial>A \<partial>B)"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   817
    by (rule integral_cong) (auto intro!: integral_pmf_restrict[symmetric])
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   818
  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)" .
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   819
qed
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   820
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   821
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   822
context
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   823
begin
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   824
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   825
interpretation pmf_as_measure .
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   826
59002
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   827
lemma measure_pmf_bind: "measure_pmf (bind_pmf M f) = (measure_pmf M \<guillemotright>= (\<lambda>x. measure_pmf (f x)))"
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   828
  by transfer simp
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   829
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   830
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)"
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   831
  using measurable_measure_pmf[of N]
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   832
  unfolding measure_pmf_bind
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   833
  apply (subst (1 3) nn_integral_max_0[symmetric])
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   834
  apply (intro nn_integral_bind[where B="count_space UNIV"])
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   835
  apply auto
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   836
  done
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   837
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   838
lemma emeasure_bind_pmf[simp]: "emeasure (bind_pmf M N) X = (\<integral>\<^sup>+x. emeasure (N x) X \<partial>M)"
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   839
  using measurable_measure_pmf[of N]
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   840
  unfolding measure_pmf_bind
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   841
  by (subst emeasure_bind[where N="count_space UNIV"]) auto
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   842
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   843
lemma bind_return_pmf': "bind_pmf N return_pmf = N"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   844
proof (transfer, clarify)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   845
  fix N :: "'a measure" assume "sets N = UNIV" then show "N \<guillemotright>= return (count_space UNIV) = N"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   846
    by (subst return_sets_cong[where N=N]) (simp_all add: bind_return')
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   847
qed
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   848
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   849
lemma bind_return_pmf'': "bind_pmf N (\<lambda>x. return_pmf (f x)) = map_pmf f N"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   850
proof (transfer, clarify)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   851
  fix N :: "'b measure" and f :: "'b \<Rightarrow> 'a" assume "prob_space N" "sets N = UNIV"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   852
  then show "N \<guillemotright>= (\<lambda>x. return (count_space UNIV) (f x)) = distr N (count_space UNIV) f"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   853
    by (subst bind_return_distr[symmetric])
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   854
       (auto simp: prob_space.not_empty measurable_def comp_def)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   855
qed
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   856
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   857
lemma bind_assoc_pmf: "bind_pmf (bind_pmf A B) C = bind_pmf A (\<lambda>x. bind_pmf (B x) C)"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   858
  by transfer
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   859
     (auto intro!: bind_assoc[where N="count_space UNIV" and R="count_space UNIV"]
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   860
           simp: measurable_def space_subprob_algebra prob_space_imp_subprob_space)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   861
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   862
end
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   863
59052
a05c8305781e projections of pair_pmf (by D. Traytel)
hoelzl
parents: 59048
diff changeset
   864
lemma map_join_pmf: "map_pmf f (join_pmf AA) = join_pmf (map_pmf (map_pmf f) AA)"
a05c8305781e projections of pair_pmf (by D. Traytel)
hoelzl
parents: 59048
diff changeset
   865
  unfolding bind_pmf_def[symmetric]
a05c8305781e projections of pair_pmf (by D. Traytel)
hoelzl
parents: 59048
diff changeset
   866
  unfolding bind_return_pmf''[symmetric] join_eq_bind_pmf bind_assoc_pmf
a05c8305781e projections of pair_pmf (by D. Traytel)
hoelzl
parents: 59048
diff changeset
   867
  by (simp add: bind_return_pmf'')
a05c8305781e projections of pair_pmf (by D. Traytel)
hoelzl
parents: 59048
diff changeset
   868
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   869
definition "pair_pmf A B = bind_pmf A (\<lambda>x. bind_pmf B (\<lambda>y. return_pmf (x, y)))"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   870
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   871
lemma pmf_pair: "pmf (pair_pmf M N) (a, b) = pmf M a * pmf N b"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   872
  unfolding pair_pmf_def pmf_bind pmf_return
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   873
  apply (subst integral_measure_pmf[where A="{b}"])
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   874
  apply (auto simp: indicator_eq_0_iff)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   875
  apply (subst integral_measure_pmf[where A="{a}"])
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   876
  apply (auto simp: indicator_eq_0_iff setsum_nonneg_eq_0_iff pmf_nonneg)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   877
  done
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   878
59002
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   879
lemma set_pair_pmf: "set_pmf (pair_pmf A B) = set_pmf A \<times> set_pmf B"
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   880
  unfolding pair_pmf_def set_bind_pmf set_return_pmf by auto
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   881
59048
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59024
diff changeset
   882
lemma measure_pmf_in_subprob_space[measurable (raw)]:
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59024
diff changeset
   883
  "measure_pmf M \<in> space (subprob_algebra (count_space UNIV))"
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59024
diff changeset
   884
  by (simp add: space_subprob_algebra) intro_locales
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59024
diff changeset
   885
59134
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
   886
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)"
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
   887
proof -
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
   888
  have "(\<integral>\<^sup>+x. f x \<partial>pair_pmf A B) = (\<integral>\<^sup>+x. max 0 (f x) * indicator (A \<times> B) x \<partial>pair_pmf A B)"
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
   889
    by (subst nn_integral_max_0[symmetric])
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
   890
       (auto simp: AE_measure_pmf_iff set_pair_pmf intro!: nn_integral_cong_AE)
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
   891
  also have "\<dots> = (\<integral>\<^sup>+a. \<integral>\<^sup>+b. max 0 (f (a, b)) * indicator (A \<times> B) (a, b) \<partial>B \<partial>A)"
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
   892
    by (simp add: pair_pmf_def)
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
   893
  also have "\<dots> = (\<integral>\<^sup>+a. \<integral>\<^sup>+b. max 0 (f (a, b)) \<partial>B \<partial>A)"
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
   894
    by (auto intro!: nn_integral_cong_AE simp: AE_measure_pmf_iff)
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
   895
  finally show ?thesis
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
   896
    unfolding nn_integral_max_0 .
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
   897
qed
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
   898
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
   899
lemma pair_map_pmf1: "pair_pmf (map_pmf f A) B = map_pmf (apfst f) (pair_pmf A B)"
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
   900
proof (safe intro!: pmf_eqI)
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
   901
  fix a :: "'a" and b :: "'b"
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
   902
  have [simp]: "\<And>c d. indicator (apfst f -` {(a, b)}) (c, d) = indicator (f -` {a}) c * (indicator {b} d::ereal)"
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
   903
    by (auto split: split_indicator)
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
   904
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
   905
  have "ereal (pmf (pair_pmf (map_pmf f A) B) (a, b)) =
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
   906
         ereal (pmf (map_pmf (apfst f) (pair_pmf A B)) (a, b))"
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
   907
    unfolding pmf_pair ereal_pmf_map
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
   908
    by (simp add: nn_integral_pair_pmf' max_def emeasure_pmf_single nn_integral_multc pmf_nonneg
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
   909
                  emeasure_map_pmf[symmetric] del: emeasure_map_pmf)
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
   910
  then show "pmf (pair_pmf (map_pmf f A) B) (a, b) = pmf (map_pmf (apfst f) (pair_pmf A B)) (a, b)"
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
   911
    by simp
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
   912
qed
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
   913
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
   914
lemma pair_map_pmf2: "pair_pmf A (map_pmf f B) = map_pmf (apsnd f) (pair_pmf A B)"
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
   915
proof (safe intro!: pmf_eqI)
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
   916
  fix a :: "'a" and b :: "'b"
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
   917
  have [simp]: "\<And>c d. indicator (apsnd f -` {(a, b)}) (c, d) = indicator {a} c * (indicator (f -` {b}) d::ereal)"
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
   918
    by (auto split: split_indicator)
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
   919
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
   920
  have "ereal (pmf (pair_pmf A (map_pmf f B)) (a, b)) =
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
   921
         ereal (pmf (map_pmf (apsnd f) (pair_pmf A B)) (a, b))"
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
   922
    unfolding pmf_pair ereal_pmf_map
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
   923
    by (simp add: nn_integral_pair_pmf' max_def emeasure_pmf_single nn_integral_cmult nn_integral_multc pmf_nonneg
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
   924
                  emeasure_map_pmf[symmetric] del: emeasure_map_pmf)
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
   925
  then show "pmf (pair_pmf A (map_pmf f B)) (a, b) = pmf (map_pmf (apsnd f) (pair_pmf A B)) (a, b)"
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
   926
    by simp
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
   927
qed
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
   928
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
   929
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)"
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
   930
  by (simp add: pair_map_pmf2 pair_map_pmf1 map_pmf_comp split_beta')
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
   931
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   932
lemma bind_pair_pmf:
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   933
  assumes M[measurable]: "M \<in> measurable (count_space UNIV \<Otimes>\<^sub>M count_space UNIV) (subprob_algebra N)"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   934
  shows "measure_pmf (pair_pmf A B) \<guillemotright>= M = (measure_pmf A \<guillemotright>= (\<lambda>x. measure_pmf B \<guillemotright>= (\<lambda>y. M (x, y))))"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   935
    (is "?L = ?R")
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   936
proof (rule measure_eqI)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   937
  have M'[measurable]: "M \<in> measurable (pair_pmf A B) (subprob_algebra N)"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   938
    using M[THEN measurable_space] by (simp_all add: space_pair_measure)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   939
59048
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59024
diff changeset
   940
  note measurable_bind[where N="count_space UNIV", measurable]
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59024
diff changeset
   941
  note measure_pmf_in_subprob_space[simp]
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59024
diff changeset
   942
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   943
  have sets_eq_N: "sets ?L = N"
59048
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59024
diff changeset
   944
    by (subst sets_bind[OF sets_kernel[OF M']]) auto
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   945
  show "sets ?L = sets ?R"
59048
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59024
diff changeset
   946
    using measurable_space[OF M]
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59024
diff changeset
   947
    by (simp add: sets_eq_N space_pair_measure space_subprob_algebra)
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   948
  fix X assume "X \<in> sets ?L"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   949
  then have X[measurable]: "X \<in> sets N"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   950
    unfolding sets_eq_N .
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   951
  then show "emeasure ?L X = emeasure ?R X"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   952
    apply (simp add: emeasure_bind[OF _ M' X])
59048
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59024
diff changeset
   953
    apply (simp add: nn_integral_bind[where B="count_space UNIV"] pair_pmf_def measure_pmf_bind[of A]
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59024
diff changeset
   954
      nn_integral_measure_pmf_finite set_return_pmf emeasure_nonneg pmf_return one_ereal_def[symmetric])
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59024
diff changeset
   955
    apply (subst emeasure_bind[OF _ _ X])
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   956
    apply measurable
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   957
    apply (subst emeasure_bind[OF _ _ X])
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   958
    apply measurable
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   959
    done
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   960
qed
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   961
59052
a05c8305781e projections of pair_pmf (by D. Traytel)
hoelzl
parents: 59048
diff changeset
   962
lemma join_map_return_pmf: "join_pmf (map_pmf return_pmf A) = A"
a05c8305781e projections of pair_pmf (by D. Traytel)
hoelzl
parents: 59048
diff changeset
   963
  unfolding bind_pmf_def[symmetric] bind_return_pmf' ..
a05c8305781e projections of pair_pmf (by D. Traytel)
hoelzl
parents: 59048
diff changeset
   964
a05c8305781e projections of pair_pmf (by D. Traytel)
hoelzl
parents: 59048
diff changeset
   965
lemma map_fst_pair_pmf: "map_pmf fst (pair_pmf A B) = A"
a05c8305781e projections of pair_pmf (by D. Traytel)
hoelzl
parents: 59048
diff changeset
   966
  by (simp add: pair_pmf_def bind_return_pmf''[symmetric] bind_assoc_pmf bind_return_pmf bind_return_pmf')
a05c8305781e projections of pair_pmf (by D. Traytel)
hoelzl
parents: 59048
diff changeset
   967
a05c8305781e projections of pair_pmf (by D. Traytel)
hoelzl
parents: 59048
diff changeset
   968
lemma map_snd_pair_pmf: "map_pmf snd (pair_pmf A B) = B"
a05c8305781e projections of pair_pmf (by D. Traytel)
hoelzl
parents: 59048
diff changeset
   969
  by (simp add: pair_pmf_def bind_return_pmf''[symmetric] bind_assoc_pmf bind_return_pmf bind_return_pmf')
a05c8305781e projections of pair_pmf (by D. Traytel)
hoelzl
parents: 59048
diff changeset
   970
59053
43e07797269b tuned proof that pmfs are bnfs
hoelzl
parents: 59052
diff changeset
   971
lemma nn_integral_pmf':
43e07797269b tuned proof that pmfs are bnfs
hoelzl
parents: 59052
diff changeset
   972
  "inj_on f A \<Longrightarrow> (\<integral>\<^sup>+x. pmf p (f x) \<partial>count_space A) = emeasure p (f ` A)"
43e07797269b tuned proof that pmfs are bnfs
hoelzl
parents: 59052
diff changeset
   973
  by (subst nn_integral_bij_count_space[where g=f and B="f`A"])
43e07797269b tuned proof that pmfs are bnfs
hoelzl
parents: 59052
diff changeset
   974
     (auto simp: bij_betw_def nn_integral_pmf)
43e07797269b tuned proof that pmfs are bnfs
hoelzl
parents: 59052
diff changeset
   975
43e07797269b tuned proof that pmfs are bnfs
hoelzl
parents: 59052
diff changeset
   976
lemma pmf_le_0_iff[simp]: "pmf M p \<le> 0 \<longleftrightarrow> pmf M p = 0"
43e07797269b tuned proof that pmfs are bnfs
hoelzl
parents: 59052
diff changeset
   977
  using pmf_nonneg[of M p] by simp
43e07797269b tuned proof that pmfs are bnfs
hoelzl
parents: 59052
diff changeset
   978
43e07797269b tuned proof that pmfs are bnfs
hoelzl
parents: 59052
diff changeset
   979
lemma min_pmf_0[simp]: "min (pmf M p) 0 = 0" "min 0 (pmf M p) = 0"
43e07797269b tuned proof that pmfs are bnfs
hoelzl
parents: 59052
diff changeset
   980
  using pmf_nonneg[of M p] by simp_all
43e07797269b tuned proof that pmfs are bnfs
hoelzl
parents: 59052
diff changeset
   981
43e07797269b tuned proof that pmfs are bnfs
hoelzl
parents: 59052
diff changeset
   982
lemma pmf_eq_0_set_pmf: "pmf M p = 0 \<longleftrightarrow> p \<notin> set_pmf M"
43e07797269b tuned proof that pmfs are bnfs
hoelzl
parents: 59052
diff changeset
   983
  unfolding set_pmf_iff by simp
43e07797269b tuned proof that pmfs are bnfs
hoelzl
parents: 59052
diff changeset
   984
43e07797269b tuned proof that pmfs are bnfs
hoelzl
parents: 59052
diff changeset
   985
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"
43e07797269b tuned proof that pmfs are bnfs
hoelzl
parents: 59052
diff changeset
   986
  by (auto simp: pmf.rep_eq map_pmf.rep_eq measure_distr AE_measure_pmf_iff inj_onD
43e07797269b tuned proof that pmfs are bnfs
hoelzl
parents: 59052
diff changeset
   987
           intro!: measure_pmf.finite_measure_eq_AE)
43e07797269b tuned proof that pmfs are bnfs
hoelzl
parents: 59052
diff changeset
   988
59023
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
   989
inductive rel_pmf :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'a pmf \<Rightarrow> 'b pmf \<Rightarrow> bool"
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
   990
for R p q
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
   991
where
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
   992
  "\<lbrakk> \<And>x y. (x, y) \<in> set_pmf pq \<Longrightarrow> R x y; 
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
   993
     map_pmf fst pq = p; map_pmf snd pq = q \<rbrakk>
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
   994
  \<Longrightarrow> rel_pmf R p q"
58587
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   995
59327
8a779359df67 rel_pmf OO: conversion to nat is not necessary
hoelzl
parents: 59325
diff changeset
   996
lemma nn_integral_count_space_eq:
8a779359df67 rel_pmf OO: conversion to nat is not necessary
hoelzl
parents: 59325
diff changeset
   997
  "(\<And>x. x \<in> A - B \<Longrightarrow> f x = 0) \<Longrightarrow> (\<And>x. x \<in> B - A \<Longrightarrow> f x = 0) \<Longrightarrow>
8a779359df67 rel_pmf OO: conversion to nat is not necessary
hoelzl
parents: 59325
diff changeset
   998
    (\<integral>\<^sup>+x. f x \<partial>count_space A) = (\<integral>\<^sup>+x. f x \<partial>count_space B)"
8a779359df67 rel_pmf OO: conversion to nat is not necessary
hoelzl
parents: 59325
diff changeset
   999
  by (auto simp: nn_integral_count_space_indicator intro!: nn_integral_cong split: split_indicator)
8a779359df67 rel_pmf OO: conversion to nat is not necessary
hoelzl
parents: 59325
diff changeset
  1000
59023
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
  1001
bnf pmf: "'a pmf" map: map_pmf sets: set_pmf bd : "natLeq" rel: rel_pmf
58587
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
  1002
proof -
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
  1003
  show "map_pmf id = id" by (rule map_pmf_id)
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
  1004
  show "\<And>f g. map_pmf (f \<circ> g) = map_pmf f \<circ> map_pmf g" by (rule map_pmf_compose) 
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
  1005
  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"
59023
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
  1006
    by (intro map_pmf_cong refl)
58587
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
  1007
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
  1008
  show "\<And>f::'a \<Rightarrow> 'b. set_pmf \<circ> map_pmf f = op ` f \<circ> set_pmf"
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
  1009
    by (rule pmf_set_map)
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
  1010
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
  1011
  { fix p :: "'s pmf"
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
  1012
    have "(card_of (set_pmf p), card_of (UNIV :: nat set)) \<in> ordLeq"
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
  1013
      by (rule card_of_ordLeqI[where f="to_nat_on (set_pmf p)"])
59053
43e07797269b tuned proof that pmfs are bnfs
hoelzl
parents: 59052
diff changeset
  1014
         (auto intro: countable_set_pmf)
58587
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
  1015
    also have "(card_of (UNIV :: nat set), natLeq) \<in> ordLeq"
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
  1016
      by (metis Field_natLeq card_of_least natLeq_Well_order)
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
  1017
    finally show "(card_of (set_pmf p), natLeq) \<in> ordLeq" . }
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
  1018
59023
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
  1019
  show "\<And>R. rel_pmf R =
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
  1020
         (BNF_Def.Grp {x. set_pmf x \<subseteq> {(x, y). R x y}} (map_pmf fst))\<inverse>\<inverse> OO
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
  1021
         BNF_Def.Grp {x. set_pmf x \<subseteq> {(x, y). R x y}} (map_pmf snd)"
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
  1022
     by (auto simp add: fun_eq_iff BNF_Def.Grp_def OO_def rel_pmf.simps)
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
  1023
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
  1024
  { fix p :: "'a pmf" and f :: "'a \<Rightarrow> 'b" and g x
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
  1025
    assume p: "\<And>z. z \<in> set_pmf p \<Longrightarrow> f z = g z"
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
  1026
      and x: "x \<in> set_pmf p"
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
  1027
    thus "f x = g x" by simp }
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
  1028
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
  1029
  fix R :: "'a => 'b \<Rightarrow> bool" and S :: "'b \<Rightarrow> 'c \<Rightarrow> bool"
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
  1030
  { fix p q r
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
  1031
    assume pq: "rel_pmf R p q"
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
  1032
      and qr:"rel_pmf S q r"
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
  1033
    from pq obtain pq where pq: "\<And>x y. (x, y) \<in> set_pmf pq \<Longrightarrow> R x y"
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
  1034
      and p: "p = map_pmf fst pq" and q: "q = map_pmf snd pq" by cases auto
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
  1035
    from qr obtain qr where qr: "\<And>y z. (y, z) \<in> set_pmf qr \<Longrightarrow> S y z"
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
  1036
      and q': "q = map_pmf fst qr" and r: "r = map_pmf snd qr" by cases auto
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
  1037
59053
43e07797269b tuned proof that pmfs are bnfs
hoelzl
parents: 59052
diff changeset
  1038
    note pmf_nonneg[intro, simp]
59327
8a779359df67 rel_pmf OO: conversion to nat is not necessary
hoelzl
parents: 59325
diff changeset
  1039
    let ?pq = "\<lambda>y x. pmf pq (x, y)"
8a779359df67 rel_pmf OO: conversion to nat is not necessary
hoelzl
parents: 59325
diff changeset
  1040
    let ?qr = "\<lambda>y z. pmf qr (y, z)"
59023
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
  1041
59327
8a779359df67 rel_pmf OO: conversion to nat is not necessary
hoelzl
parents: 59325
diff changeset
  1042
    have nn_integral_pp2: "\<And>y. (\<integral>\<^sup>+ x. ?pq y x \<partial>count_space UNIV) = pmf q y"
8a779359df67 rel_pmf OO: conversion to nat is not necessary
hoelzl
parents: 59325
diff changeset
  1043
      by (simp add: nn_integral_pmf' inj_on_def q)
59325
922d31f5c3f5 simplify construction for distribution of rel_pmf over op OO
Andreas Lochbihler
parents: 59134
diff changeset
  1044
         (auto simp add: ereal_pmf_map intro!: arg_cong2[where f=emeasure])
59327
8a779359df67 rel_pmf OO: conversion to nat is not necessary
hoelzl
parents: 59325
diff changeset
  1045
    have nn_integral_rr1: "\<And>y. (\<integral>\<^sup>+ x. ?qr y x \<partial>count_space UNIV) = pmf q y"
8a779359df67 rel_pmf OO: conversion to nat is not necessary
hoelzl
parents: 59325
diff changeset
  1046
      by (simp add: nn_integral_pmf' inj_on_def q')
59325
922d31f5c3f5 simplify construction for distribution of rel_pmf over op OO
Andreas Lochbihler
parents: 59134
diff changeset
  1047
         (auto simp add: ereal_pmf_map intro!: arg_cong2[where f=emeasure])
59327
8a779359df67 rel_pmf OO: conversion to nat is not necessary
hoelzl
parents: 59325
diff changeset
  1048
    have eq: "\<And>y. (\<integral>\<^sup>+ x. ?pq y x \<partial>count_space UNIV) = (\<integral>\<^sup>+ z. ?qr y z \<partial>count_space UNIV)"
59325
922d31f5c3f5 simplify construction for distribution of rel_pmf over op OO
Andreas Lochbihler
parents: 59134
diff changeset
  1049
      by(simp add: nn_integral_pp2 nn_integral_rr1)
59023
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
  1050
59327
8a779359df67 rel_pmf OO: conversion to nat is not necessary
hoelzl
parents: 59325
diff changeset
  1051
    def assign \<equiv> "\<lambda>y x z. ?pq y x * ?qr y z / pmf q y"
59325
922d31f5c3f5 simplify construction for distribution of rel_pmf over op OO
Andreas Lochbihler
parents: 59134
diff changeset
  1052
    have assign_nonneg [simp]: "\<And>y x z. 0 \<le> assign y x z" by(simp add: assign_def)
59327
8a779359df67 rel_pmf OO: conversion to nat is not necessary
hoelzl
parents: 59325
diff changeset
  1053
    have assign_eq_0_outside: "\<And>y x z. \<lbrakk> ?pq y x = 0 \<or> ?qr y z = 0 \<rbrakk> \<Longrightarrow> assign y x z = 0"
59325
922d31f5c3f5 simplify construction for distribution of rel_pmf over op OO
Andreas Lochbihler
parents: 59134
diff changeset
  1054
      by(auto simp add: assign_def)
59327
8a779359df67 rel_pmf OO: conversion to nat is not necessary
hoelzl
parents: 59325
diff changeset
  1055
    have nn_integral_assign1: "\<And>y z. (\<integral>\<^sup>+ x. assign y x z \<partial>count_space UNIV) = ?qr y z"
59023
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
  1056
    proof -
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
  1057
      fix y z
59325
922d31f5c3f5 simplify construction for distribution of rel_pmf over op OO
Andreas Lochbihler
parents: 59134
diff changeset
  1058
      have "(\<integral>\<^sup>+ x. assign y x z \<partial>count_space UNIV) = 
59327
8a779359df67 rel_pmf OO: conversion to nat is not necessary
hoelzl
parents: 59325
diff changeset
  1059
            (\<integral>\<^sup>+ x. ?pq y x \<partial>count_space UNIV) * (?qr y z / pmf q y)"
59325
922d31f5c3f5 simplify construction for distribution of rel_pmf over op OO
Andreas Lochbihler
parents: 59134
diff changeset
  1060
        by(simp add: assign_def nn_integral_multc times_ereal.simps(1)[symmetric] divide_real_def mult.assoc del: times_ereal.simps(1))
59327
8a779359df67 rel_pmf OO: conversion to nat is not necessary
hoelzl
parents: 59325
diff changeset
  1061
      also have "\<dots> = ?qr y z" by(auto simp add: image_iff q' pmf_eq_0_set_pmf set_map_pmf nn_integral_pp2)
59023
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
  1062
      finally show "?thesis y z" .
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
  1063
    qed
59327
8a779359df67 rel_pmf OO: conversion to nat is not necessary
hoelzl
parents: 59325
diff changeset
  1064
    have nn_integral_assign2: "\<And>y x. (\<integral>\<^sup>+ z. assign y x z \<partial>count_space UNIV) = ?pq y x"
59325
922d31f5c3f5 simplify construction for distribution of rel_pmf over op OO
Andreas Lochbihler
parents: 59134
diff changeset
  1065
    proof -
922d31f5c3f5 simplify construction for distribution of rel_pmf over op OO
Andreas Lochbihler
parents: 59134
diff changeset
  1066
      fix x y
59327
8a779359df67 rel_pmf OO: conversion to nat is not necessary
hoelzl
parents: 59325
diff changeset
  1067
      have "(\<integral>\<^sup>+ z. assign y x z \<partial>count_space UNIV) = (\<integral>\<^sup>+ z. ?qr y z \<partial>count_space UNIV) * (?pq y x / pmf q y)"
8a779359df67 rel_pmf OO: conversion to nat is not necessary
hoelzl
parents: 59325
diff changeset
  1068
        by(simp add: assign_def divide_real_def mult.commute[where a="?pq y x"] mult.assoc nn_integral_multc times_ereal.simps(1)[symmetric] del: times_ereal.simps(1))
8a779359df67 rel_pmf OO: conversion to nat is not necessary
hoelzl
parents: 59325
diff changeset
  1069
      also have "\<dots> = ?pq y x" by(auto simp add: image_iff pmf_eq_0_set_pmf set_map_pmf q nn_integral_rr1)
59325
922d31f5c3f5 simplify construction for distribution of rel_pmf over op OO
Andreas Lochbihler
parents: 59134
diff changeset
  1070
      finally show "?thesis y x" .
922d31f5c3f5 simplify construction for distribution of rel_pmf over op OO
Andreas Lochbihler
parents: 59134
diff changeset
  1071
    qed
59023
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
  1072
59327
8a779359df67 rel_pmf OO: conversion to nat is not necessary
hoelzl
parents: 59325
diff changeset
  1073
    def pqr \<equiv> "embed_pmf (\<lambda>(y, x, z). assign y x z)"
59023
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
  1074
    { fix y x z
59327
8a779359df67 rel_pmf OO: conversion to nat is not necessary
hoelzl
parents: 59325
diff changeset
  1075
      have "assign y x z = pmf pqr (y, x, z)"
8a779359df67 rel_pmf OO: conversion to nat is not necessary
hoelzl
parents: 59325
diff changeset
  1076
        unfolding pqr_def
59053
43e07797269b tuned proof that pmfs are bnfs
hoelzl
parents: 59052
diff changeset
  1077
      proof (subst pmf_embed_pmf)
43e07797269b tuned proof that pmfs are bnfs
hoelzl
parents: 59052
diff changeset
  1078
        have "(\<integral>\<^sup>+ x. ereal ((\<lambda>(y, x, z). assign y x z) x) \<partial>count_space UNIV) =
59327
8a779359df67 rel_pmf OO: conversion to nat is not necessary
hoelzl
parents: 59325
diff changeset
  1079
          (\<integral>\<^sup>+ x. ereal ((\<lambda>(y, x, z). assign y x z) x) \<partial>(count_space ((\<lambda>((x, y), z). (y, x, z)) ` (pq \<times> r))))"
8a779359df67 rel_pmf OO: conversion to nat is not necessary
hoelzl
parents: 59325
diff changeset
  1080
          by (force simp add: pmf_eq_0_set_pmf r set_map_pmf split: split_indicator
8a779359df67 rel_pmf OO: conversion to nat is not necessary
hoelzl
parents: 59325
diff changeset
  1081
                    intro!: nn_integral_count_space_eq assign_eq_0_outside)
8a779359df67 rel_pmf OO: conversion to nat is not necessary
hoelzl
parents: 59325
diff changeset
  1082
        also have "\<dots> = (\<integral>\<^sup>+ x. ereal ((\<lambda>((x, y), z). assign y x z) x) \<partial>(count_space (pq \<times> r)))"
59053
43e07797269b tuned proof that pmfs are bnfs
hoelzl
parents: 59052
diff changeset
  1083
          by (subst nn_integral_bij_count_space[OF inj_on_imp_bij_betw, symmetric])
43e07797269b tuned proof that pmfs are bnfs
hoelzl
parents: 59052
diff changeset
  1084
             (auto simp: inj_on_def intro!: nn_integral_cong)
59327
8a779359df67 rel_pmf OO: conversion to nat is not necessary
hoelzl
parents: 59325
diff changeset
  1085
        also have "\<dots> = (\<integral>\<^sup>+ xy. \<integral>\<^sup>+z. ereal ((\<lambda>((x, y), z). assign y x z) (xy, z)) \<partial>count_space r \<partial>count_space pq)"
59053
43e07797269b tuned proof that pmfs are bnfs
hoelzl
parents: 59052
diff changeset
  1086
          by (subst sigma_finite_measure.nn_integral_fst)
43e07797269b tuned proof that pmfs are bnfs
hoelzl
parents: 59052
diff changeset
  1087
             (auto simp: pair_measure_countable sigma_finite_measure_count_space_countable)
59327
8a779359df67 rel_pmf OO: conversion to nat is not necessary
hoelzl
parents: 59325
diff changeset
  1088
        also have "\<dots> = (\<integral>\<^sup>+ xy. \<integral>\<^sup>+z. ereal ((\<lambda>((x, y), z). assign y x z) (xy, z)) \<partial>count_space UNIV \<partial>count_space pq)"
8a779359df67 rel_pmf OO: conversion to nat is not necessary
hoelzl
parents: 59325
diff changeset
  1089
          by (intro nn_integral_cong nn_integral_count_space_eq)
8a779359df67 rel_pmf OO: conversion to nat is not necessary
hoelzl
parents: 59325
diff changeset
  1090
             (force simp: r set_map_pmf pmf_eq_0_set_pmf intro!: assign_eq_0_outside)+
8a779359df67 rel_pmf OO: conversion to nat is not necessary
hoelzl
parents: 59325
diff changeset
  1091
        also have "\<dots> = (\<integral>\<^sup>+ z. ?pq (snd z) (fst z) \<partial>count_space pq)"
59053
43e07797269b tuned proof that pmfs are bnfs
hoelzl
parents: 59052
diff changeset
  1092
          by (subst nn_integral_assign2[symmetric]) (auto intro!: nn_integral_cong)
43e07797269b tuned proof that pmfs are bnfs
hoelzl
parents: 59052
diff changeset
  1093
        finally show "(\<integral>\<^sup>+ x. ereal ((\<lambda>(y, x, z). assign y x z) x) \<partial>count_space UNIV) = 1"
43e07797269b tuned proof that pmfs are bnfs
hoelzl
parents: 59052
diff changeset
  1094
          by (simp add: nn_integral_pmf emeasure_pmf)
43e07797269b tuned proof that pmfs are bnfs
hoelzl
parents: 59052
diff changeset
  1095
      qed auto }
43e07797269b tuned proof that pmfs are bnfs
hoelzl
parents: 59052
diff changeset
  1096
    note a = this
59023
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
  1097
59327
8a779359df67 rel_pmf OO: conversion to nat is not necessary
hoelzl
parents: 59325
diff changeset
  1098
    def pr \<equiv> "map_pmf (\<lambda>(y, x, z). (x, z)) pqr"
59023
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
  1099
59053
43e07797269b tuned proof that pmfs are bnfs
hoelzl
parents: 59052
diff changeset
  1100
    have "rel_pmf (R OO S) p r"
59023
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
  1101
    proof
59327
8a779359df67 rel_pmf OO: conversion to nat is not necessary
hoelzl
parents: 59325
diff changeset
  1102
      have pq_eq: "pq = map_pmf (\<lambda>(y, x, z). (x, y)) pqr"
59053
43e07797269b tuned proof that pmfs are bnfs
hoelzl
parents: 59052
diff changeset
  1103
      proof (rule pmf_eqI)
43e07797269b tuned proof that pmfs are bnfs
hoelzl
parents: 59052
diff changeset
  1104
        fix i
59327
8a779359df67 rel_pmf OO: conversion to nat is not necessary
hoelzl
parents: 59325
diff changeset
  1105
        show "pmf pq i = pmf (map_pmf (\<lambda>(y, x, z). (x, y)) pqr) i"
8a779359df67 rel_pmf OO: conversion to nat is not necessary
hoelzl
parents: 59325
diff changeset
  1106
          using nn_integral_assign2[of "snd i" "fst i", symmetric]
59053
43e07797269b tuned proof that pmfs are bnfs
hoelzl
parents: 59052
diff changeset
  1107
          by (auto simp add: a nn_integral_pmf' inj_on_def ereal.inject[symmetric] ereal_pmf_map 
43e07797269b tuned proof that pmfs are bnfs
hoelzl
parents: 59052
diff changeset
  1108
                   simp del: ereal.inject intro!: arg_cong2[where f=emeasure])
43e07797269b tuned proof that pmfs are bnfs
hoelzl
parents: 59052
diff changeset
  1109
      qed
59327
8a779359df67 rel_pmf OO: conversion to nat is not necessary
hoelzl
parents: 59325
diff changeset
  1110
      then show "map_pmf fst pr = p"
59053
43e07797269b tuned proof that pmfs are bnfs
hoelzl
parents: 59052
diff changeset
  1111
        unfolding p pr_def by (simp add: map_pmf_comp split_beta)
43e07797269b tuned proof that pmfs are bnfs
hoelzl
parents: 59052
diff changeset
  1112
59327
8a779359df67 rel_pmf OO: conversion to nat is not necessary
hoelzl
parents: 59325
diff changeset
  1113
      have qr_eq: "qr = map_pmf (\<lambda>(y, x, z). (y, z)) pqr"
59053
43e07797269b tuned proof that pmfs are bnfs
hoelzl
parents: 59052
diff changeset
  1114
      proof (rule pmf_eqI)
59327
8a779359df67 rel_pmf OO: conversion to nat is not necessary
hoelzl
parents: 59325
diff changeset
  1115
        fix i show "pmf qr i = pmf (map_pmf (\<lambda>(y, x, z). (y, z)) pqr) i"
59053
43e07797269b tuned proof that pmfs are bnfs
hoelzl
parents: 59052
diff changeset
  1116
          using nn_integral_assign1[of "fst i" "snd i", symmetric]
43e07797269b tuned proof that pmfs are bnfs
hoelzl
parents: 59052
diff changeset
  1117
          by (auto simp add: a nn_integral_pmf' inj_on_def ereal.inject[symmetric] ereal_pmf_map 
43e07797269b tuned proof that pmfs are bnfs
hoelzl
parents: 59052
diff changeset
  1118
                   simp del: ereal.inject intro!: arg_cong2[where f=emeasure])
43e07797269b tuned proof that pmfs are bnfs
hoelzl
parents: 59052
diff changeset
  1119
      qed
59327
8a779359df67 rel_pmf OO: conversion to nat is not necessary
hoelzl
parents: 59325
diff changeset
  1120
      then show "map_pmf snd pr = r"
59053
43e07797269b tuned proof that pmfs are bnfs
hoelzl
parents: 59052
diff changeset
  1121
        unfolding r pr_def by (simp add: map_pmf_comp split_beta)
43e07797269b tuned proof that pmfs are bnfs
hoelzl
parents: 59052
diff changeset
  1122
43e07797269b tuned proof that pmfs are bnfs
hoelzl
parents: 59052
diff changeset
  1123
      fix x z assume "(x, z) \<in> set_pmf pr"
43e07797269b tuned proof that pmfs are bnfs
hoelzl
parents: 59052
diff changeset
  1124
      then have "\<exists>y. (x, y) \<in> set_pmf pq \<and> (y, z) \<in> set_pmf qr"
59327
8a779359df67 rel_pmf OO: conversion to nat is not necessary
hoelzl
parents: 59325
diff changeset
  1125
        unfolding pr_def pq_eq qr_eq by (force simp: set_map_pmf)
59053
43e07797269b tuned proof that pmfs are bnfs
hoelzl
parents: 59052
diff changeset
  1126
      with pq qr show "(R OO S) x z"
43e07797269b tuned proof that pmfs are bnfs
hoelzl
parents: 59052
diff changeset
  1127
        by blast
43e07797269b tuned proof that pmfs are bnfs
hoelzl
parents: 59052
diff changeset
  1128
    qed }
59023
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
  1129
  then show "rel_pmf R OO rel_pmf S \<le> rel_pmf (R OO S)"
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
  1130
    by(auto simp add: le_fun_def)
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
  1131
qed (fact natLeq_card_order natLeq_cinfinite)+
58587
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
  1132
59134
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
  1133
lemma rel_pmf_return_pmf1: "rel_pmf R (return_pmf x) M \<longleftrightarrow> (\<forall>a\<in>M. R x a)"
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
  1134
proof safe
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
  1135
  fix a assume "a \<in> M" "rel_pmf R (return_pmf x) M"
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
  1136
  then obtain pq where *: "\<And>a b. (a, b) \<in> set_pmf pq \<Longrightarrow> R a b"
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
  1137
    and eq: "return_pmf x = map_pmf fst pq" "M = map_pmf snd pq"
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
  1138
    by (force elim: rel_pmf.cases)
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
  1139
  moreover have "set_pmf (return_pmf x) = {x}"
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
  1140
    by (simp add: set_return_pmf)
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
  1141
  with `a \<in> M` have "(x, a) \<in> pq"
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
  1142
    by (force simp: eq set_map_pmf)
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
  1143
  with * show "R x a"
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
  1144
    by auto
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
  1145
qed (auto intro!: rel_pmf.intros[where pq="pair_pmf (return_pmf x) M"]
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
  1146
          simp: map_fst_pair_pmf map_snd_pair_pmf set_pair_pmf set_return_pmf)
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
  1147
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
  1148
lemma rel_pmf_return_pmf2: "rel_pmf R M (return_pmf x) \<longleftrightarrow> (\<forall>a\<in>M. R a x)"
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
  1149
  by (subst pmf.rel_flip[symmetric]) (simp add: rel_pmf_return_pmf1)
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
  1150
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
  1151
lemma rel_pmf_rel_prod:
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
  1152
  "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'"
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
  1153
proof safe
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
  1154
  assume "rel_pmf (rel_prod R S) (pair_pmf A A') (pair_pmf B B')"
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
  1155
  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"
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
  1156
    and eq: "map_pmf fst pq = pair_pmf A A'" "map_pmf snd pq = pair_pmf B B'"
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
  1157
    by (force elim: rel_pmf.cases)
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
  1158
  show "rel_pmf R A B"
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
  1159
  proof (rule rel_pmf.intros)
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
  1160
    let ?f = "\<lambda>(a, b). (fst a, fst b)"
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
  1161
    have [simp]: "(\<lambda>x. fst (?f x)) = fst o fst" "(\<lambda>x. snd (?f x)) = fst o snd"
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
  1162
      by auto
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
  1163
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
  1164
    show "map_pmf fst (map_pmf ?f pq) = A"
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
  1165
      by (simp add: map_pmf_comp pmf.map_comp[symmetric] eq map_fst_pair_pmf)
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
  1166
    show "map_pmf snd (map_pmf ?f pq) = B"
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
  1167
      by (simp add: map_pmf_comp pmf.map_comp[symmetric] eq map_fst_pair_pmf)
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
  1168
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
  1169
    fix a b assume "(a, b) \<in> set_pmf (map_pmf ?f pq)"
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
  1170
    then obtain c d where "((a, c), (b, d)) \<in> set_pmf pq"
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
  1171
      by (auto simp: set_map_pmf)
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
  1172
    from pq[OF this] show "R a b" ..
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
  1173
  qed
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
  1174
  show "rel_pmf S A' B'"
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
  1175
  proof (rule rel_pmf.intros)
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
  1176
    let ?f = "\<lambda>(a, b). (snd a, snd b)"
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
  1177
    have [simp]: "(\<lambda>x. fst (?f x)) = snd o fst" "(\<lambda>x. snd (?f x)) = snd o snd"
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
  1178
      by auto
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
  1179
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
  1180
    show "map_pmf fst (map_pmf ?f pq) = A'"
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
  1181
      by (simp add: map_pmf_comp pmf.map_comp[symmetric] eq map_snd_pair_pmf)
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
  1182
    show "map_pmf snd (map_pmf ?f pq) = B'"
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
  1183
      by (simp add: map_pmf_comp pmf.map_comp[symmetric] eq map_snd_pair_pmf)
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
  1184
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
  1185
    fix c d assume "(c, d) \<in> set_pmf (map_pmf ?f pq)"
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
  1186
    then obtain a b where "((a, c), (b, d)) \<in> set_pmf pq"
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
  1187
      by (auto simp: set_map_pmf)
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
  1188
    from pq[OF this] show "S c d" ..
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
  1189
  qed
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
  1190
next
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
  1191
  assume "rel_pmf R A B" "rel_pmf S A' B'"
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
  1192
  then obtain Rpq Spq
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
  1193
    where Rpq: "\<And>a b. (a, b) \<in> set_pmf Rpq \<Longrightarrow> R a b"
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
  1194
        "map_pmf fst Rpq = A" "map_pmf snd Rpq = B"
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
  1195
      and Spq: "\<And>a b. (a, b) \<in> set_pmf Spq \<Longrightarrow> S a b"
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
  1196
        "map_pmf fst Spq = A'" "map_pmf snd Spq = B'"
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
  1197
    by (force elim: rel_pmf.cases)
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
  1198
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
  1199
  let ?f = "(\<lambda>((a, c), (b, d)). ((a, b), (c, d)))"
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
  1200
  let ?pq = "map_pmf ?f (pair_pmf Rpq Spq)"
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
  1201
  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))"
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
  1202
    by auto
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
  1203
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
  1204
  show "rel_pmf (rel_prod R S) (pair_pmf A A') (pair_pmf B B')"
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
  1205
    by (rule rel_pmf.intros[where pq="?pq"])
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
  1206
       (auto simp: map_snd_pair_pmf map_fst_pair_pmf set_pair_pmf set_map_pmf map_pmf_comp Rpq Spq
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
  1207
                   map_pair)
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
  1208
qed
a71f2e256ee2 rel_pmf commutes with rel_prod
hoelzl
parents: 59093
diff changeset
  1209
58587
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
  1210
end
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
  1211