src/HOL/Probability/Probability_Mass_Function.thy
author paulson <lp15@cam.ac.uk>
Tue, 10 Nov 2015 14:43:29 +0000
changeset 61610 4f54d2759a0b
parent 61609 77b453bd616f
parent 61605 1bf7b186542e
child 61634 48e2de1b1df5
permissions -rw-r--r--
Merge
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
58606
9c66f7c541fb add Giry monad
hoelzl
parents: 58587
diff changeset
     1
(*  Title:      HOL/Probability/Probability_Mass_Function.thy
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59665
diff changeset
     2
    Author:     Johannes Hölzl, TU München
59023
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
     3
    Author:     Andreas Lochbihler, ETH Zurich
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
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    11
  "~~/src/HOL/Library/Multiset"
58587
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
    12
begin
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
    13
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
    14
lemma AE_emeasure_singleton:
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
    15
  assumes x: "emeasure M {x} \<noteq> 0" and ae: "AE x in M. P x" shows "P x"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
    16
proof -
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
    17
  from x have x_M: "{x} \<in> sets M"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
    18
    by (auto intro: emeasure_notin_sets)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
    19
  from ae obtain N where N: "{x\<in>space M. \<not> P x} \<subseteq> N" "emeasure M N = 0" "N \<in> sets M"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
    20
    by (auto elim: AE_E)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
    21
  { assume "\<not> P x"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
    22
    with x_M[THEN sets.sets_into_space] N have "emeasure M {x} \<le> emeasure M N"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
    23
      by (intro emeasure_mono) auto
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
    24
    with x N have False
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
    25
      by (auto simp: emeasure_le_0_iff) }
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
    26
  then show "P x" by auto
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
    27
qed
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
    28
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
    29
lemma AE_measure_singleton: "measure M {x} \<noteq> 0 \<Longrightarrow> AE x in M. P x \<Longrightarrow> P x"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
    30
  by (metis AE_emeasure_singleton measure_def emeasure_empty measure_empty)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
    31
59494
054f9c9d73ea add cond_map_pmf
hoelzl
parents: 59493
diff changeset
    32
lemma ereal_divide': "b \<noteq> 0 \<Longrightarrow> ereal (a / b) = ereal a / ereal b"
054f9c9d73ea add cond_map_pmf
hoelzl
parents: 59493
diff changeset
    33
  using ereal_divide[of a b] by simp
054f9c9d73ea add cond_map_pmf
hoelzl
parents: 59493
diff changeset
    34
59052
a05c8305781e projections of pair_pmf (by D. Traytel)
hoelzl
parents: 59048
diff changeset
    35
lemma (in finite_measure) countable_support:
58587
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
    36
  "countable {x. measure M {x} \<noteq> 0}"
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    37
proof cases
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    38
  assume "measure M (space M) = 0"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    39
  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
    40
    by auto
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    41
  then show ?thesis
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    42
    by simp
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    43
next
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    44
  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
    45
  assume "?M \<noteq> 0"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    46
  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
    47
    using reals_Archimedean[of "?m x / ?M" for x]
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    48
    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
    49
  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
    50
  proof (rule ccontr)
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    51
    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
    52
    then obtain X where "finite X" "card X = Suc (Suc n)" "X \<subseteq> ?X"
58606
9c66f7c541fb add Giry monad
hoelzl
parents: 58587
diff changeset
    53
      by (metis infinite_arbitrarily_large)
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59665
diff changeset
    54
    from this(3) have *: "\<And>x. x \<in> X \<Longrightarrow> ?M / Suc n \<le> ?m x"
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    55
      by auto
58587
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
    56
    { fix x assume "x \<in> X"
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    57
      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
    58
      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
    59
    note singleton_sets = this
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    60
    have "?M < (\<Sum>x\<in>X. ?M / Suc n)"
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59665
diff changeset
    61
      using `?M \<noteq> 0`
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
    62
      by (simp add: `card X = Suc (Suc n)` 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
    63
    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
    64
      by (rule setsum_mono) fact
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
    65
    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
    66
      using singleton_sets `finite X`
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
    67
      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
    68
    finally have "?M < measure M (\<Union>x\<in>X. {x})" .
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    69
    moreover have "measure M (\<Union>x\<in>X. {x}) \<le> ?M"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    70
      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
    71
    ultimately show False by simp
58587
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
    72
  qed
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
    73
  show ?thesis
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
    74
    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
    75
qed
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
    76
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    77
lemma (in finite_measure) AE_support_countable:
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    78
  assumes [simp]: "sets M = UNIV"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    79
  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
    80
proof
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    81
  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
    82
  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
    83
    by auto
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59665
diff changeset
    84
  then have "emeasure M (\<Union>x\<in>{x\<in>S. emeasure M {x} \<noteq> 0}. {x}) =
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    85
    (\<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
    86
    by (subst emeasure_UN_countable)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    87
       (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
    88
  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
    89
    by (auto intro!: nn_integral_cong split: split_indicator)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    90
  also have "\<dots> = emeasure M (\<Union>x\<in>S. {x})"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    91
    by (subst emeasure_UN_countable)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    92
       (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
    93
  also have "\<dots> = emeasure M (space M)"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    94
    using ae by (intro emeasure_eq_AE) auto
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    95
  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
    96
    by (simp add: emeasure_single_in_space cong: rev_conj_cong)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
    97
  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
    98
  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
    99
    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
   100
  then show "AE x in M. measure M {x} \<noteq> 0"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   101
    by (auto simp: emeasure_eq_measure)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   102
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
   103
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   104
subsection \<open> PMF as measure \<close>
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   105
58587
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   106
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
   107
  morphisms measure_pmf Abs_pmf
58606
9c66f7c541fb add Giry monad
hoelzl
parents: 58587
diff changeset
   108
  by (intro exI[of _ "uniform_measure (count_space UNIV) {undefined}"])
9c66f7c541fb add Giry monad
hoelzl
parents: 58587
diff changeset
   109
     (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
   110
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   111
declare [[coercion measure_pmf]]
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   112
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   113
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
   114
  using pmf.measure_pmf[of p] by auto
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   115
61605
1bf7b186542e qualifier is mandatory by default;
wenzelm
parents: 61424
diff changeset
   116
interpretation measure_pmf: prob_space "measure_pmf M" for M
58587
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   117
  by (rule prob_space_measure_pmf)
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   118
61605
1bf7b186542e qualifier is mandatory by default;
wenzelm
parents: 61424
diff changeset
   119
interpretation measure_pmf: subprob_space "measure_pmf M" for M
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   120
  by (rule prob_space_imp_subprob_space) unfold_locales
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   121
59048
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59024
diff changeset
   122
lemma subprob_space_measure_pmf: "subprob_space (measure_pmf x)"
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59024
diff changeset
   123
  by unfold_locales
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59024
diff changeset
   124
58587
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   125
locale pmf_as_measure
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   126
begin
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   127
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   128
setup_lifting type_definition_pmf
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   129
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   130
end
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   131
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   132
context
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   133
begin
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   134
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   135
interpretation pmf_as_measure .
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   136
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   137
lemma sets_measure_pmf[simp]: "sets (measure_pmf p) = UNIV"
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59665
diff changeset
   138
  by transfer blast
58587
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   139
59048
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59024
diff changeset
   140
lemma sets_measure_pmf_count_space[measurable_cong]:
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59024
diff changeset
   141
  "sets (measure_pmf M) = sets (count_space UNIV)"
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   142
  by simp
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   143
58587
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   144
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
   145
  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
   146
59048
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59024
diff changeset
   147
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
   148
  by (simp add: space_subprob_algebra subprob_space_measure_pmf)
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59024
diff changeset
   149
58587
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   150
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
   151
  by (auto simp: measurable_def)
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   152
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   153
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
   154
  by (intro measurable_cong_sets) simp_all
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   155
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   156
lemma measurable_pair_restrict_pmf2:
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   157
  assumes "countable A"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   158
  assumes [measurable]: "\<And>y. y \<in> A \<Longrightarrow> (\<lambda>x. f (x, y)) \<in> measurable M L"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   159
  shows "f \<in> measurable (M \<Otimes>\<^sub>M restrict_space (measure_pmf N) A) L" (is "f \<in> measurable ?M _")
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   160
proof -
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   161
  have [measurable_cong]: "sets (restrict_space (count_space UNIV) A) = sets (count_space A)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   162
    by (simp add: restrict_count_space)
58587
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   163
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   164
  show ?thesis
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   165
    by (intro measurable_compose_countable'[where f="\<lambda>a b. f (fst b, a)" and g=snd and I=A,
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61169
diff changeset
   166
                                            unfolded prod.collapse] assms)
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   167
        measurable
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   168
qed
58587
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   169
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   170
lemma measurable_pair_restrict_pmf1:
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   171
  assumes "countable A"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   172
  assumes [measurable]: "\<And>x. x \<in> A \<Longrightarrow> (\<lambda>y. f (x, y)) \<in> measurable N L"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   173
  shows "f \<in> measurable (restrict_space (measure_pmf M) A \<Otimes>\<^sub>M N) L"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   174
proof -
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   175
  have [measurable_cong]: "sets (restrict_space (count_space UNIV) A) = sets (count_space A)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   176
    by (simp add: restrict_count_space)
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   177
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   178
  show ?thesis
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   179
    by (intro measurable_compose_countable'[where f="\<lambda>a b. f (a, snd b)" and g=fst and I=A,
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61169
diff changeset
   180
                                            unfolded prod.collapse] assms)
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   181
        measurable
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   182
qed
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   183
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   184
lift_definition pmf :: "'a pmf \<Rightarrow> 'a \<Rightarrow> real" is "\<lambda>M x. measure M {x}" .
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   185
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   186
lift_definition set_pmf :: "'a pmf \<Rightarrow> 'a set" is "\<lambda>M. {x. measure M {x} \<noteq> 0}" .
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   187
declare [[coercion set_pmf]]
58587
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   188
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   189
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
   190
  by transfer simp
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   191
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   192
lemma emeasure_pmf_single_eq_zero_iff:
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   193
  fixes M :: "'a pmf"
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   194
  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
   195
  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
   196
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   197
lemma AE_measure_pmf_iff: "(AE x in measure_pmf M. P x) \<longleftrightarrow> (\<forall>y\<in>M. P y)"
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   198
  using AE_measure_singleton[of M] AE_measure_pmf[of M]
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   199
  by (auto simp: set_pmf.rep_eq)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   200
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   201
lemma countable_set_pmf [simp]: "countable (set_pmf p)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   202
  by transfer (metis prob_space.finite_measure finite_measure.countable_support)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   203
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   204
lemma pmf_positive: "x \<in> set_pmf p \<Longrightarrow> 0 < pmf p x"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   205
  by transfer (simp add: less_le measure_nonneg)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   206
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   207
lemma pmf_nonneg: "0 \<le> pmf p x"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   208
  by transfer (simp add: measure_nonneg)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   209
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   210
lemma pmf_le_1: "pmf p x \<le> 1"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   211
  by (simp add: pmf.rep_eq)
58587
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   212
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   213
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
   214
  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
   215
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   216
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
   217
  by transfer simp
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   218
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   219
lemma set_pmf_eq: "set_pmf M = {x. pmf M x \<noteq> 0}"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   220
  by (auto simp: set_pmf_iff)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   221
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   222
lemma emeasure_pmf_single:
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   223
  fixes M :: "'a pmf"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   224
  shows "emeasure M {x} = pmf M x"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   225
  by transfer (simp add: finite_measure.emeasure_eq_measure[OF prob_space.finite_measure])
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   226
60068
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
   227
lemma measure_pmf_single: "measure (measure_pmf M) {x} = pmf M x"
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
   228
using emeasure_pmf_single[of M x] by(simp add: measure_pmf.emeasure_eq_measure)
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
   229
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   230
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
   231
  by (subst emeasure_eq_setsum_singleton) (auto simp: emeasure_pmf_single)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   232
59023
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
   233
lemma measure_measure_pmf_finite: "finite S \<Longrightarrow> measure (measure_pmf M) S = setsum (pmf M) S"
59425
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59327
diff changeset
   234
  using emeasure_measure_pmf_finite[of S M] by(simp add: measure_pmf.emeasure_eq_measure)
59023
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
   235
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   236
lemma nn_integral_measure_pmf_support:
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   237
  fixes f :: "'a \<Rightarrow> ereal"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   238
  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
   239
  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
   240
proof -
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   241
  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
   242
    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
   243
  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
   244
    using assms by (intro nn_integral_indicator_finite) auto
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   245
  finally show ?thesis
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   246
    by (simp add: emeasure_measure_pmf_finite)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   247
qed
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   248
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   249
lemma nn_integral_measure_pmf_finite:
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   250
  fixes f :: "'a \<Rightarrow> ereal"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   251
  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
   252
  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
   253
  using assms by (intro nn_integral_measure_pmf_support) auto
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   254
lemma integrable_measure_pmf_finite:
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   255
  fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   256
  shows "finite (set_pmf M) \<Longrightarrow> integrable M f"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   257
  by (auto intro!: integrableI_bounded simp: nn_integral_measure_pmf_finite)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   258
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   259
lemma integral_measure_pmf:
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   260
  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
   261
  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
   262
proof -
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   263
  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
   264
    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
   265
  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
   266
    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
   267
  finally show ?thesis .
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   268
qed
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   269
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   270
lemma integrable_pmf: "integrable (count_space X) (pmf M)"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   271
proof -
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   272
  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
   273
    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
   274
  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
   275
    by (simp add: integrable_iff_bounded pmf_nonneg)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   276
  then show ?thesis
59023
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
   277
    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
   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: "(\<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
   281
proof -
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   282
  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
   283
    by (simp add: pmf_nonneg integrable_pmf nn_integral_eq_integral)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   284
  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
   285
    by (auto intro!: nn_integral_cong_AE split: split_indicator
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   286
             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
   287
                   AE_count_space set_pmf_iff)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   288
  also have "\<dots> = emeasure M (X \<inter> M)"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   289
    by (rule emeasure_countable_singleton[symmetric]) (auto intro: countable_set_pmf)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   290
  also have "\<dots> = emeasure M X"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   291
    by (auto intro!: emeasure_eq_AE simp: AE_measure_pmf_iff)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   292
  finally show ?thesis
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   293
    by (simp add: measure_pmf.emeasure_eq_measure)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   294
qed
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   295
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   296
lemma integral_pmf_restrict:
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   297
  "(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
   298
    (\<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
   299
  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
   300
58587
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   301
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
   302
proof -
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   303
  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
   304
    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
   305
  then show ?thesis
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   306
    using measure_pmf.emeasure_space_1 by simp
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   307
qed
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   308
59490
f71732294f29 tune proof
Andreas Lochbihler
parents: 59475
diff changeset
   309
lemma emeasure_pmf_UNIV [simp]: "emeasure (measure_pmf M) UNIV = 1"
f71732294f29 tune proof
Andreas Lochbihler
parents: 59475
diff changeset
   310
using measure_pmf.emeasure_space_1[of M] by simp
f71732294f29 tune proof
Andreas Lochbihler
parents: 59475
diff changeset
   311
59023
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
   312
lemma in_null_sets_measure_pmfI:
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
   313
  "A \<inter> set_pmf p = {} \<Longrightarrow> A \<in> null_sets (measure_pmf p)"
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
   314
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
   315
by(auto simp add: null_sets_def AE_measure_pmf_iff)
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
   316
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   317
lemma measure_subprob: "measure_pmf M \<in> space (subprob_algebra (count_space UNIV))"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   318
  by (simp add: space_subprob_algebra subprob_space_measure_pmf)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   319
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   320
subsection \<open> Monad Interpretation \<close>
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   321
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   322
lemma measurable_measure_pmf[measurable]:
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   323
  "(\<lambda>x. measure_pmf (M x)) \<in> measurable (count_space UNIV) (subprob_algebra (count_space UNIV))"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   324
  by (auto simp: space_subprob_algebra intro!: prob_space_imp_subprob_space) unfold_locales
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   325
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   326
lemma bind_measure_pmf_cong:
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   327
  assumes "\<And>x. A x \<in> space (subprob_algebra N)" "\<And>x. B x \<in> space (subprob_algebra N)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   328
  assumes "\<And>i. i \<in> set_pmf x \<Longrightarrow> A i = B i"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   329
  shows "bind (measure_pmf x) A = bind (measure_pmf x) B"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   330
proof (rule measure_eqI)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   331
  show "sets (measure_pmf x \<guillemotright>= A) = sets (measure_pmf x \<guillemotright>= B)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   332
    using assms by (subst (1 2) sets_bind) (auto simp: space_subprob_algebra)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   333
next
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   334
  fix X assume "X \<in> sets (measure_pmf x \<guillemotright>= A)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   335
  then have X: "X \<in> sets N"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   336
    using assms by (subst (asm) sets_bind) (auto simp: space_subprob_algebra)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   337
  show "emeasure (measure_pmf x \<guillemotright>= A) X = emeasure (measure_pmf x \<guillemotright>= B) X"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   338
    using assms
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   339
    by (subst (1 2) emeasure_bind[where N=N, OF _ _ X])
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   340
       (auto intro!: nn_integral_cong_AE simp: AE_measure_pmf_iff)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   341
qed
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   342
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   343
lift_definition bind_pmf :: "'a pmf \<Rightarrow> ('a \<Rightarrow> 'b pmf ) \<Rightarrow> 'b pmf" is bind
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   344
proof (clarify, intro conjI)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   345
  fix f :: "'a measure" and g :: "'a \<Rightarrow> 'b measure"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   346
  assume "prob_space f"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   347
  then interpret f: prob_space f .
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   348
  assume "sets f = UNIV" and ae_f: "AE x in f. measure f {x} \<noteq> 0"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   349
  then have s_f[simp]: "sets f = sets (count_space UNIV)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   350
    by simp
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   351
  assume g: "\<And>x. prob_space (g x) \<and> sets (g x) = UNIV \<and> (AE y in g x. measure (g x) {y} \<noteq> 0)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   352
  then have g: "\<And>x. prob_space (g x)" and s_g[simp]: "\<And>x. sets (g x) = sets (count_space UNIV)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   353
    and ae_g: "\<And>x. AE y in g x. measure (g x) {y} \<noteq> 0"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   354
    by auto
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   355
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   356
  have [measurable]: "g \<in> measurable f (subprob_algebra (count_space UNIV))"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   357
    by (auto simp: measurable_def space_subprob_algebra prob_space_imp_subprob_space g)
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59665
diff changeset
   358
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   359
  show "prob_space (f \<guillemotright>= g)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   360
    using g by (intro f.prob_space_bind[where S="count_space UNIV"]) auto
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59665
diff changeset
   361
  then interpret fg: prob_space "f \<guillemotright>= g" .
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   362
  show [simp]: "sets (f \<guillemotright>= g) = UNIV"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   363
    using sets_eq_imp_space_eq[OF s_f]
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   364
    by (subst sets_bind[where N="count_space UNIV"]) auto
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   365
  show "AE x in f \<guillemotright>= g. measure (f \<guillemotright>= g) {x} \<noteq> 0"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   366
    apply (simp add: fg.prob_eq_0 AE_bind[where B="count_space UNIV"])
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   367
    using ae_f
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   368
    apply eventually_elim
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   369
    using ae_g
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   370
    apply eventually_elim
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   371
    apply (auto dest: AE_measure_singleton)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   372
    done
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   373
qed
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   374
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   375
lemma ereal_pmf_bind: "pmf (bind_pmf N f) i = (\<integral>\<^sup>+x. pmf (f x) i \<partial>measure_pmf N)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   376
  unfolding pmf.rep_eq bind_pmf.rep_eq
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   377
  by (auto simp: measure_pmf.measure_bind[where N="count_space UNIV"] measure_subprob measure_nonneg
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   378
           intro!: nn_integral_eq_integral[symmetric] measure_pmf.integrable_const_bound[where B=1])
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   379
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   380
lemma pmf_bind: "pmf (bind_pmf N f) i = (\<integral>x. pmf (f x) i \<partial>measure_pmf N)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   381
  using ereal_pmf_bind[of N f i]
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   382
  by (subst (asm) nn_integral_eq_integral)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   383
     (auto simp: pmf_nonneg pmf_le_1
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   384
           intro!: nn_integral_eq_integral[symmetric] measure_pmf.integrable_const_bound[where B=1])
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   385
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   386
lemma bind_pmf_const[simp]: "bind_pmf M (\<lambda>x. c) = c"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   387
  by transfer (simp add: bind_const' prob_space_imp_subprob_space)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   388
59665
37adca7fd48f add set_pmf lemmas to simpset
hoelzl
parents: 59664
diff changeset
   389
lemma set_bind_pmf[simp]: "set_pmf (bind_pmf M N) = (\<Union>M\<in>set_pmf M. set_pmf (N M))"
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59665
diff changeset
   390
  unfolding set_pmf_eq ereal_eq_0(1)[symmetric] ereal_pmf_bind
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   391
  by (auto simp add: nn_integral_0_iff_AE AE_measure_pmf_iff set_pmf_eq not_le less_le pmf_nonneg)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   392
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   393
lemma bind_pmf_cong:
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   394
  assumes "p = q"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   395
  shows "(\<And>x. x \<in> set_pmf q \<Longrightarrow> f x = g x) \<Longrightarrow> bind_pmf p f = bind_pmf q g"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   396
  unfolding `p = q`[symmetric] measure_pmf_inject[symmetric] bind_pmf.rep_eq
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   397
  by (auto simp: AE_measure_pmf_iff Pi_iff space_subprob_algebra subprob_space_measure_pmf
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   398
                 sets_bind[where N="count_space UNIV"] emeasure_bind[where N="count_space UNIV"]
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   399
           intro!: nn_integral_cong_AE measure_eqI)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   400
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   401
lemma bind_pmf_cong_simp:
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   402
  "p = q \<Longrightarrow> (\<And>x. x \<in> set_pmf q =simp=> f x = g x) \<Longrightarrow> bind_pmf p f = bind_pmf q g"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   403
  by (simp add: simp_implies_def cong: bind_pmf_cong)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   404
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   405
lemma measure_pmf_bind: "measure_pmf (bind_pmf M f) = (measure_pmf M \<guillemotright>= (\<lambda>x. measure_pmf (f x)))"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   406
  by transfer simp
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   407
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   408
lemma nn_integral_bind_pmf[simp]: "(\<integral>\<^sup>+x. f x \<partial>bind_pmf M N) = (\<integral>\<^sup>+x. \<integral>\<^sup>+y. f y \<partial>N x \<partial>M)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   409
  using measurable_measure_pmf[of N]
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   410
  unfolding measure_pmf_bind
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   411
  apply (subst (1 3) nn_integral_max_0[symmetric])
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   412
  apply (intro nn_integral_bind[where B="count_space UNIV"])
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   413
  apply auto
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   414
  done
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   415
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   416
lemma emeasure_bind_pmf[simp]: "emeasure (bind_pmf M N) X = (\<integral>\<^sup>+x. emeasure (N x) X \<partial>M)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   417
  using measurable_measure_pmf[of N]
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   418
  unfolding measure_pmf_bind
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   419
  by (subst emeasure_bind[where N="count_space UNIV"]) auto
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59665
diff changeset
   420
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   421
lift_definition return_pmf :: "'a \<Rightarrow> 'a pmf" is "return (count_space UNIV)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   422
  by (auto intro!: prob_space_return simp: AE_return measure_return)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   423
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   424
lemma bind_return_pmf: "bind_pmf (return_pmf x) f = f x"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   425
  by transfer
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   426
     (auto intro!: prob_space_imp_subprob_space bind_return[where N="count_space UNIV"]
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   427
           simp: space_subprob_algebra)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   428
59665
37adca7fd48f add set_pmf lemmas to simpset
hoelzl
parents: 59664
diff changeset
   429
lemma set_return_pmf[simp]: "set_pmf (return_pmf x) = {x}"
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   430
  by transfer (auto simp add: measure_return split: split_indicator)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   431
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   432
lemma bind_return_pmf': "bind_pmf N return_pmf = N"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   433
proof (transfer, clarify)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   434
  fix N :: "'a measure" assume "sets N = UNIV" then show "N \<guillemotright>= return (count_space UNIV) = N"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   435
    by (subst return_sets_cong[where N=N]) (simp_all add: bind_return')
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   436
qed
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   437
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   438
lemma bind_assoc_pmf: "bind_pmf (bind_pmf A B) C = bind_pmf A (\<lambda>x. bind_pmf (B x) C)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   439
  by transfer
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   440
     (auto intro!: bind_assoc[where N="count_space UNIV" and R="count_space UNIV"]
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   441
           simp: measurable_def space_subprob_algebra prob_space_imp_subprob_space)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   442
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   443
definition "map_pmf f M = bind_pmf M (\<lambda>x. return_pmf (f x))"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   444
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   445
lemma map_bind_pmf: "map_pmf f (bind_pmf M g) = bind_pmf M (\<lambda>x. map_pmf f (g x))"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   446
  by (simp add: map_pmf_def bind_assoc_pmf)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   447
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   448
lemma bind_map_pmf: "bind_pmf (map_pmf f M) g = bind_pmf M (\<lambda>x. g (f x))"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   449
  by (simp add: map_pmf_def bind_assoc_pmf bind_return_pmf)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   450
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   451
lemma map_pmf_transfer[transfer_rule]:
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   452
  "rel_fun op = (rel_fun cr_pmf cr_pmf) (\<lambda>f M. distr M (count_space UNIV) f) map_pmf"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   453
proof -
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   454
  have "rel_fun op = (rel_fun pmf_as_measure.cr_pmf pmf_as_measure.cr_pmf)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   455
     (\<lambda>f M. M \<guillemotright>= (return (count_space UNIV) o f)) map_pmf"
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59665
diff changeset
   456
    unfolding map_pmf_def[abs_def] comp_def by transfer_prover
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   457
  then show ?thesis
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   458
    by (force simp: rel_fun_def cr_pmf_def bind_return_distr)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   459
qed
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   460
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   461
lemma map_pmf_rep_eq:
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   462
  "measure_pmf (map_pmf f M) = distr (measure_pmf M) (count_space UNIV) f"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   463
  unfolding map_pmf_def bind_pmf.rep_eq comp_def return_pmf.rep_eq
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   464
  using bind_return_distr[of M f "count_space UNIV"] by (simp add: comp_def)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   465
58587
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   466
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
   467
  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
   468
59053
43e07797269b tuned proof that pmfs are bnfs
hoelzl
parents: 59052
diff changeset
   469
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
   470
  using map_pmf_id unfolding id_def .
43e07797269b tuned proof that pmfs are bnfs
hoelzl
parents: 59052
diff changeset
   471
58587
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   472
lemma map_pmf_compose: "map_pmf (f \<circ> g) = map_pmf f \<circ> map_pmf g"
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59665
diff changeset
   473
  by (rule, transfer) (simp add: distr_distr[symmetric, where N="count_space UNIV"] measurable_def)
58587
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   474
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   475
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
   476
  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
   477
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   478
lemma map_pmf_cong: "p = q \<Longrightarrow> (\<And>x. x \<in> set_pmf q \<Longrightarrow> f x = g x) \<Longrightarrow> map_pmf f p = map_pmf g q"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   479
  unfolding map_pmf_def by (rule bind_pmf_cong) auto
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   480
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   481
lemma pmf_set_map: "set_pmf \<circ> map_pmf f = op ` f \<circ> set_pmf"
59665
37adca7fd48f add set_pmf lemmas to simpset
hoelzl
parents: 59664
diff changeset
   482
  by (auto simp add: comp_def fun_eq_iff map_pmf_def)
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   483
59665
37adca7fd48f add set_pmf lemmas to simpset
hoelzl
parents: 59664
diff changeset
   484
lemma set_map_pmf[simp]: "set_pmf (map_pmf f M) = f`set_pmf M"
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   485
  using pmf_set_map[of f] by (auto simp: comp_def fun_eq_iff)
58587
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   486
59002
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   487
lemma emeasure_map_pmf[simp]: "emeasure (map_pmf f M) X = emeasure M (f -` X)"
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   488
  unfolding map_pmf_rep_eq by (subst emeasure_distr) auto
59002
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   489
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   490
lemma nn_integral_map_pmf[simp]: "(\<integral>\<^sup>+x. f x \<partial>map_pmf g M) = (\<integral>\<^sup>+x. f (g x) \<partial>M)"
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   491
  unfolding map_pmf_rep_eq by (intro nn_integral_distr) auto
59002
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
   492
59023
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
   493
lemma ereal_pmf_map: "pmf (map_pmf f p) x = (\<integral>\<^sup>+ y. indicator (f -` {x}) y \<partial>measure_pmf p)"
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   494
proof (transfer fixing: f x)
59023
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
   495
  fix p :: "'b measure"
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
   496
  presume "prob_space p"
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
   497
  then interpret prob_space p .
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
   498
  presume "sets p = UNIV"
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
   499
  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
   500
    by(simp add: measure_distr measurable_def emeasure_eq_measure)
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
   501
qed simp_all
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
   502
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
   503
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
   504
proof -
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
   505
  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
   506
    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
   507
  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
   508
    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
   509
  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
   510
    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
   511
  also have "\<dots> = emeasure (measure_pmf p) A"
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
   512
    by(auto intro: arg_cong2[where f=emeasure])
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
   513
  finally show ?thesis .
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
   514
qed
4999a616336c register pmf as BNF
Andreas Lochbihler
parents: 59002
diff changeset
   515
60068
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
   516
lemma map_return_pmf [simp]: "map_pmf f (return_pmf x) = return_pmf (f x)"
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   517
  by transfer (simp add: distr_return)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   518
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   519
lemma map_pmf_const[simp]: "map_pmf (\<lambda>_. c) M = return_pmf c"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   520
  by transfer (auto simp: prob_space.distr_const)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   521
60068
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
   522
lemma pmf_return [simp]: "pmf (return_pmf x) y = indicator {y} x"
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   523
  by transfer (simp add: measure_return)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   524
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   525
lemma nn_integral_return_pmf[simp]: "0 \<le> f x \<Longrightarrow> (\<integral>\<^sup>+x. f x \<partial>return_pmf x) = f x"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   526
  unfolding return_pmf.rep_eq by (intro nn_integral_return) auto
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   527
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   528
lemma emeasure_return_pmf[simp]: "emeasure (return_pmf x) X = indicator X x"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   529
  unfolding return_pmf.rep_eq by (intro emeasure_return) auto
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   530
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   531
lemma return_pmf_inj[simp]: "return_pmf x = return_pmf y \<longleftrightarrow> x = y"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   532
  by (metis insertI1 set_return_pmf singletonD)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   533
59665
37adca7fd48f add set_pmf lemmas to simpset
hoelzl
parents: 59664
diff changeset
   534
lemma map_pmf_eq_return_pmf_iff:
37adca7fd48f add set_pmf lemmas to simpset
hoelzl
parents: 59664
diff changeset
   535
  "map_pmf f p = return_pmf x \<longleftrightarrow> (\<forall>y \<in> set_pmf p. f y = x)"
37adca7fd48f add set_pmf lemmas to simpset
hoelzl
parents: 59664
diff changeset
   536
proof
37adca7fd48f add set_pmf lemmas to simpset
hoelzl
parents: 59664
diff changeset
   537
  assume "map_pmf f p = return_pmf x"
37adca7fd48f add set_pmf lemmas to simpset
hoelzl
parents: 59664
diff changeset
   538
  then have "set_pmf (map_pmf f p) = set_pmf (return_pmf x)" by simp
37adca7fd48f add set_pmf lemmas to simpset
hoelzl
parents: 59664
diff changeset
   539
  then show "\<forall>y \<in> set_pmf p. f y = x" by auto
37adca7fd48f add set_pmf lemmas to simpset
hoelzl
parents: 59664
diff changeset
   540
next
37adca7fd48f add set_pmf lemmas to simpset
hoelzl
parents: 59664
diff changeset
   541
  assume "\<forall>y \<in> set_pmf p. f y = x"
37adca7fd48f add set_pmf lemmas to simpset
hoelzl
parents: 59664
diff changeset
   542
  then show "map_pmf f p = return_pmf x"
37adca7fd48f add set_pmf lemmas to simpset
hoelzl
parents: 59664
diff changeset
   543
    unfolding map_pmf_const[symmetric, of _ p] by (intro map_pmf_cong) auto
37adca7fd48f add set_pmf lemmas to simpset
hoelzl
parents: 59664
diff changeset
   544
qed
37adca7fd48f add set_pmf lemmas to simpset
hoelzl
parents: 59664
diff changeset
   545
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   546
definition "pair_pmf A B = bind_pmf A (\<lambda>x. bind_pmf B (\<lambda>y. return_pmf (x, y)))"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   547
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   548
lemma pmf_pair: "pmf (pair_pmf M N) (a, b) = pmf M a * pmf N b"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   549
  unfolding pair_pmf_def pmf_bind pmf_return
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   550
  apply (subst integral_measure_pmf[where A="{b}"])
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   551
  apply (auto simp: indicator_eq_0_iff)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   552
  apply (subst integral_measure_pmf[where A="{a}"])
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   553
  apply (auto simp: indicator_eq_0_iff setsum_nonneg_eq_0_iff pmf_nonneg)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   554
  done
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   555
59665
37adca7fd48f add set_pmf lemmas to simpset
hoelzl
parents: 59664
diff changeset
   556
lemma set_pair_pmf[simp]: "set_pmf (pair_pmf A B) = set_pmf A \<times> set_pmf B"
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   557
  unfolding pair_pmf_def set_bind_pmf set_return_pmf by auto
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   558
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   559
lemma measure_pmf_in_subprob_space[measurable (raw)]:
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   560
  "measure_pmf M \<in> space (subprob_algebra (count_space UNIV))"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   561
  by (simp add: space_subprob_algebra) intro_locales
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   562
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   563
lemma nn_integral_pair_pmf': "(\<integral>\<^sup>+x. f x \<partial>pair_pmf A B) = (\<integral>\<^sup>+a. \<integral>\<^sup>+b. f (a, b) \<partial>B \<partial>A)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   564
proof -
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   565
  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)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   566
    by (subst nn_integral_max_0[symmetric])
59665
37adca7fd48f add set_pmf lemmas to simpset
hoelzl
parents: 59664
diff changeset
   567
       (auto simp: AE_measure_pmf_iff intro!: nn_integral_cong_AE)
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   568
  also have "\<dots> = (\<integral>\<^sup>+a. \<integral>\<^sup>+b. max 0 (f (a, b)) * indicator (A \<times> B) (a, b) \<partial>B \<partial>A)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   569
    by (simp add: pair_pmf_def)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   570
  also have "\<dots> = (\<integral>\<^sup>+a. \<integral>\<^sup>+b. max 0 (f (a, b)) \<partial>B \<partial>A)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   571
    by (auto intro!: nn_integral_cong_AE simp: AE_measure_pmf_iff)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   572
  finally show ?thesis
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   573
    unfolding nn_integral_max_0 .
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   574
qed
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   575
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   576
lemma bind_pair_pmf:
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   577
  assumes M[measurable]: "M \<in> measurable (count_space UNIV \<Otimes>\<^sub>M count_space UNIV) (subprob_algebra N)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   578
  shows "measure_pmf (pair_pmf A B) \<guillemotright>= M = (measure_pmf A \<guillemotright>= (\<lambda>x. measure_pmf B \<guillemotright>= (\<lambda>y. M (x, y))))"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   579
    (is "?L = ?R")
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   580
proof (rule measure_eqI)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   581
  have M'[measurable]: "M \<in> measurable (pair_pmf A B) (subprob_algebra N)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   582
    using M[THEN measurable_space] by (simp_all add: space_pair_measure)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   583
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   584
  note measurable_bind[where N="count_space UNIV", measurable]
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   585
  note measure_pmf_in_subprob_space[simp]
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   586
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   587
  have sets_eq_N: "sets ?L = N"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   588
    by (subst sets_bind[OF sets_kernel[OF M']]) auto
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   589
  show "sets ?L = sets ?R"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   590
    using measurable_space[OF M]
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   591
    by (simp add: sets_eq_N space_pair_measure space_subprob_algebra)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   592
  fix X assume "X \<in> sets ?L"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   593
  then have X[measurable]: "X \<in> sets N"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   594
    unfolding sets_eq_N .
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   595
  then show "emeasure ?L X = emeasure ?R X"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   596
    apply (simp add: emeasure_bind[OF _ M' X])
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   597
    apply (simp add: nn_integral_bind[where B="count_space UNIV"] pair_pmf_def measure_pmf_bind[of A]
60068
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
   598
                     nn_integral_measure_pmf_finite emeasure_nonneg one_ereal_def[symmetric])
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   599
    apply (subst emeasure_bind[OF _ _ X])
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   600
    apply measurable
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   601
    apply (subst emeasure_bind[OF _ _ X])
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   602
    apply measurable
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   603
    done
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   604
qed
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   605
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   606
lemma map_fst_pair_pmf: "map_pmf fst (pair_pmf A B) = A"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   607
  by (simp add: pair_pmf_def map_pmf_def bind_assoc_pmf bind_return_pmf bind_return_pmf')
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   608
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   609
lemma map_snd_pair_pmf: "map_pmf snd (pair_pmf A B) = B"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   610
  by (simp add: pair_pmf_def map_pmf_def bind_assoc_pmf bind_return_pmf bind_return_pmf')
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   611
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   612
lemma nn_integral_pmf':
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   613
  "inj_on f A \<Longrightarrow> (\<integral>\<^sup>+x. pmf p (f x) \<partial>count_space A) = emeasure p (f ` A)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   614
  by (subst nn_integral_bij_count_space[where g=f and B="f`A"])
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   615
     (auto simp: bij_betw_def nn_integral_pmf)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   616
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   617
lemma pmf_le_0_iff[simp]: "pmf M p \<le> 0 \<longleftrightarrow> pmf M p = 0"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   618
  using pmf_nonneg[of M p] by simp
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   619
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   620
lemma min_pmf_0[simp]: "min (pmf M p) 0 = 0" "min 0 (pmf M p) = 0"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   621
  using pmf_nonneg[of M p] by simp_all
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   622
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   623
lemma pmf_eq_0_set_pmf: "pmf M p = 0 \<longleftrightarrow> p \<notin> set_pmf M"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   624
  unfolding set_pmf_iff by simp
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   625
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   626
lemma pmf_map_inj: "inj_on f (set_pmf M) \<Longrightarrow> x \<in> set_pmf M \<Longrightarrow> pmf (map_pmf f M) (f x) = pmf M x"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   627
  by (auto simp: pmf.rep_eq map_pmf_rep_eq measure_distr AE_measure_pmf_iff inj_onD
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   628
           intro!: measure_pmf.finite_measure_eq_AE)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   629
60068
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
   630
lemma pmf_map_inj': "inj f \<Longrightarrow> pmf (map_pmf f M) (f x) = pmf M x"
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
   631
apply(cases "x \<in> set_pmf M")
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
   632
 apply(simp add: pmf_map_inj[OF subset_inj_on])
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
   633
apply(simp add: pmf_eq_0_set_pmf[symmetric])
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
   634
apply(auto simp add: pmf_eq_0_set_pmf dest: injD)
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
   635
done
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
   636
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
   637
lemma pmf_map_outside: "x \<notin> f ` set_pmf M \<Longrightarrow> pmf (map_pmf f M) x = 0"
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
   638
unfolding pmf_eq_0_set_pmf by simp
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
   639
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   640
subsection \<open> PMFs as function \<close>
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   641
58587
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   642
context
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   643
  fixes f :: "'a \<Rightarrow> real"
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   644
  assumes nonneg: "\<And>x. 0 \<le> f x"
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   645
  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
   646
begin
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   647
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   648
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
   649
proof (intro conjI)
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   650
  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
   651
    by (simp split: split_indicator)
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   652
  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
   653
    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
   654
    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
   655
  show "prob_space (density (count_space UNIV) (ereal \<circ> f))"
61169
4de9ff3ea29a tuned proofs -- less legacy;
wenzelm
parents: 60602
diff changeset
   656
    by standard (simp add: emeasure_density prob)
58587
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   657
qed simp
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   658
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   659
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
   660
proof transfer
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   661
  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
   662
    by (simp split: split_indicator)
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   663
  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
   664
    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
   665
qed
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   666
60068
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
   667
lemma set_embed_pmf: "set_pmf embed_pmf = {x. f x \<noteq> 0}"
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
   668
by(auto simp add: set_pmf_eq assms pmf_embed_pmf)
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
   669
58587
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   670
end
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   671
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   672
lemma embed_pmf_transfer:
58730
b3fd0628f849 add transfer rule for set_pmf
hoelzl
parents: 58606
diff changeset
   673
  "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
   674
  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
   675
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   676
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
   677
proof (transfer, elim conjE)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   678
  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
   679
  assume "prob_space M" then interpret prob_space M .
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   680
  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
   681
  proof (rule measure_eqI)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   682
    fix A :: "'a set"
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59665
diff changeset
   683
    have "(\<integral>\<^sup>+ x. ereal (measure M {x}) * indicator A x \<partial>count_space UNIV) =
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   684
      (\<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
   685
      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
   686
    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
   687
      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
   688
    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
   689
      by (intro emeasure_UN_countable[symmetric] countable_Int2 countable_support)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   690
         (auto simp: disjoint_family_on_def)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   691
    also have "\<dots> = emeasure M A"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   692
      using ae by (intro emeasure_eq_AE) auto
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   693
    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
   694
      using emeasure_space_1 by (simp add: emeasure_density)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   695
  qed simp
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   696
qed
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   697
58587
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   698
lemma td_pmf_embed_pmf:
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   699
  "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
   700
  unfolding type_definition_def
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   701
proof safe
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   702
  fix p :: "'a pmf"
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   703
  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
   704
    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
   705
  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
   706
    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
   707
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   708
  show "embed_pmf (pmf p) = p"
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   709
    by (intro measure_pmf_inject[THEN iffD1])
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   710
       (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
   711
next
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   712
  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
   713
  then show "pmf (embed_pmf f) = f"
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   714
    by (auto intro!: pmf_embed_pmf)
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   715
qed (rule pmf_nonneg)
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   716
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   717
end
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   718
60068
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
   719
lemma nn_integral_measure_pmf: "(\<integral>\<^sup>+ x. f x \<partial>measure_pmf p) = \<integral>\<^sup>+ x. ereal (pmf p x) * f x \<partial>count_space UNIV"
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
   720
by(simp add: measure_pmf_eq_density nn_integral_density pmf_nonneg)
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
   721
58587
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   722
locale pmf_as_function
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   723
begin
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   724
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   725
setup_lifting td_pmf_embed_pmf
5484f6079bcd add type for probability mass functions, i.e. discrete probability distribution
hoelzl
parents:
diff changeset
   726
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59665
diff changeset
   727
lemma set_pmf_transfer[transfer_rule]:
58730
b3fd0628f849 add transfer rule for set_pmf
hoelzl
parents: 58606
diff changeset
   728
  assumes "bi_total A"
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59665
diff changeset
   729
  shows "rel_fun (pcr_pmf A) (rel_set A) (\<lambda>f. {x. f x \<noteq> 0}) set_pmf"
58730
b3fd0628f849 add transfer rule for set_pmf
hoelzl
parents: 58606
diff changeset
   730
  using `bi_total A`
b3fd0628f849 add transfer rule for set_pmf
hoelzl
parents: 58606
diff changeset
   731
  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
   732
     metis+
b3fd0628f849 add transfer rule for set_pmf
hoelzl
parents: 58606
diff changeset
   733
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   734
end
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   735
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   736
context
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   737
begin
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   738
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   739
interpretation pmf_as_function .
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   740
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   741
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
   742
  by transfer auto
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   743
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   744
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
   745
  by (auto intro: pmf_eqI)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   746
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   747
lemma bind_commute_pmf: "bind_pmf A (\<lambda>x. bind_pmf B (C x)) = bind_pmf B (\<lambda>y. bind_pmf A (\<lambda>x. C x y))"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   748
  unfolding pmf_eq_iff pmf_bind
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   749
proof
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   750
  fix i
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   751
  interpret B: prob_space "restrict_space B B"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   752
    by (intro prob_space_restrict_space measure_pmf.emeasure_eq_1_AE)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   753
       (auto simp: AE_measure_pmf_iff)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   754
  interpret A: prob_space "restrict_space A A"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   755
    by (intro prob_space_restrict_space measure_pmf.emeasure_eq_1_AE)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   756
       (auto simp: AE_measure_pmf_iff)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   757
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   758
  interpret AB: pair_prob_space "restrict_space A A" "restrict_space B B"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   759
    by unfold_locales
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   760
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   761
  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)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   762
    by (rule integral_cong) (auto intro!: integral_pmf_restrict)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   763
  also have "\<dots> = (\<integral> x. (\<integral> y. pmf (C x y) i \<partial>restrict_space B B) \<partial>restrict_space A A)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   764
    by (intro integral_pmf_restrict B.borel_measurable_lebesgue_integral measurable_pair_restrict_pmf2
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   765
              countable_set_pmf borel_measurable_count_space)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   766
  also have "\<dots> = (\<integral> y. \<integral> x. pmf (C x y) i \<partial>restrict_space A A \<partial>restrict_space B B)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   767
    by (rule AB.Fubini_integral[symmetric])
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   768
       (auto intro!: AB.integrable_const_bound[where B=1] measurable_pair_restrict_pmf2
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   769
             simp: pmf_nonneg pmf_le_1 measurable_restrict_space1)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   770
  also have "\<dots> = (\<integral> y. \<integral> x. pmf (C x y) i \<partial>restrict_space A A \<partial>B)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   771
    by (intro integral_pmf_restrict[symmetric] A.borel_measurable_lebesgue_integral measurable_pair_restrict_pmf2
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   772
              countable_set_pmf borel_measurable_count_space)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   773
  also have "\<dots> = (\<integral> y. \<integral> x. pmf (C x y) i \<partial>A \<partial>B)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   774
    by (rule integral_cong) (auto intro!: integral_pmf_restrict[symmetric])
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   775
  finally show "(\<integral> x. \<integral> y. pmf (C x y) i \<partial>B \<partial>A) = (\<integral> y. \<integral> x. pmf (C x y) i \<partial>A \<partial>B)" .
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   776
qed
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   777
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   778
lemma pair_map_pmf1: "pair_pmf (map_pmf f A) B = map_pmf (apfst f) (pair_pmf A B)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   779
proof (safe intro!: pmf_eqI)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   780
  fix a :: "'a" and b :: "'b"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   781
  have [simp]: "\<And>c d. indicator (apfst f -` {(a, b)}) (c, d) = indicator (f -` {a}) c * (indicator {b} d::ereal)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   782
    by (auto split: split_indicator)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   783
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   784
  have "ereal (pmf (pair_pmf (map_pmf f A) B) (a, b)) =
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   785
         ereal (pmf (map_pmf (apfst f) (pair_pmf A B)) (a, b))"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   786
    unfolding pmf_pair ereal_pmf_map
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   787
    by (simp add: nn_integral_pair_pmf' max_def emeasure_pmf_single nn_integral_multc pmf_nonneg
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   788
                  emeasure_map_pmf[symmetric] del: emeasure_map_pmf)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   789
  then show "pmf (pair_pmf (map_pmf f A) B) (a, b) = pmf (map_pmf (apfst f) (pair_pmf A B)) (a, b)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   790
    by simp
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   791
qed
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   792
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   793
lemma pair_map_pmf2: "pair_pmf A (map_pmf f B) = map_pmf (apsnd f) (pair_pmf A B)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   794
proof (safe intro!: pmf_eqI)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   795
  fix a :: "'a" and b :: "'b"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   796
  have [simp]: "\<And>c d. indicator (apsnd f -` {(a, b)}) (c, d) = indicator {a} c * (indicator (f -` {b}) d::ereal)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   797
    by (auto split: split_indicator)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   798
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   799
  have "ereal (pmf (pair_pmf A (map_pmf f B)) (a, b)) =
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   800
         ereal (pmf (map_pmf (apsnd f) (pair_pmf A B)) (a, b))"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   801
    unfolding pmf_pair ereal_pmf_map
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   802
    by (simp add: nn_integral_pair_pmf' max_def emeasure_pmf_single nn_integral_cmult nn_integral_multc pmf_nonneg
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   803
                  emeasure_map_pmf[symmetric] del: emeasure_map_pmf)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   804
  then show "pmf (pair_pmf A (map_pmf f B)) (a, b) = pmf (map_pmf (apsnd f) (pair_pmf A B)) (a, b)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   805
    by simp
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   806
qed
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   807
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   808
lemma map_pair: "map_pmf (\<lambda>(a, b). (f a, g b)) (pair_pmf A B) = pair_pmf (map_pmf f A) (map_pmf g B)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   809
  by (simp add: pair_map_pmf2 pair_map_pmf1 map_pmf_comp split_beta')
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   810
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   811
end
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
   812
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   813
subsection \<open> Conditional Probabilities \<close>
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   814
59670
dee043d19729 generalized bind_cond_pmf_cancel
hoelzl
parents: 59667
diff changeset
   815
lemma measure_pmf_zero_iff: "measure (measure_pmf p) s = 0 \<longleftrightarrow> set_pmf p \<inter> s = {}"
dee043d19729 generalized bind_cond_pmf_cancel
hoelzl
parents: 59667
diff changeset
   816
  by (subst measure_pmf.prob_eq_0) (auto simp: AE_measure_pmf_iff)
dee043d19729 generalized bind_cond_pmf_cancel
hoelzl
parents: 59667
diff changeset
   817
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   818
context
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   819
  fixes p :: "'a pmf" and s :: "'a set"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   820
  assumes not_empty: "set_pmf p \<inter> s \<noteq> {}"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   821
begin
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   822
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   823
interpretation pmf_as_measure .
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   824
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   825
lemma emeasure_measure_pmf_not_zero: "emeasure (measure_pmf p) s \<noteq> 0"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   826
proof
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   827
  assume "emeasure (measure_pmf p) s = 0"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   828
  then have "AE x in measure_pmf p. x \<notin> s"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   829
    by (rule AE_I[rotated]) auto
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   830
  with not_empty show False
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   831
    by (auto simp: AE_measure_pmf_iff)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   832
qed
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   833
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   834
lemma measure_measure_pmf_not_zero: "measure (measure_pmf p) s \<noteq> 0"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   835
  using emeasure_measure_pmf_not_zero unfolding measure_pmf.emeasure_eq_measure by simp
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   836
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   837
lift_definition cond_pmf :: "'a pmf" is
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   838
  "uniform_measure (measure_pmf p) s"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   839
proof (intro conjI)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   840
  show "prob_space (uniform_measure (measure_pmf p) s)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   841
    by (intro prob_space_uniform_measure) (auto simp: emeasure_measure_pmf_not_zero)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   842
  show "AE x in uniform_measure (measure_pmf p) s. measure (uniform_measure (measure_pmf p) s) {x} \<noteq> 0"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   843
    by (simp add: emeasure_measure_pmf_not_zero measure_measure_pmf_not_zero AE_uniform_measure
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   844
                  AE_measure_pmf_iff set_pmf.rep_eq)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   845
qed simp
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   846
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   847
lemma pmf_cond: "pmf cond_pmf x = (if x \<in> s then pmf p x / measure p s else 0)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   848
  by transfer (simp add: emeasure_measure_pmf_not_zero pmf.rep_eq)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   849
59665
37adca7fd48f add set_pmf lemmas to simpset
hoelzl
parents: 59664
diff changeset
   850
lemma set_cond_pmf[simp]: "set_pmf cond_pmf = set_pmf p \<inter> s"
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   851
  by (auto simp add: set_pmf_iff pmf_cond measure_measure_pmf_not_zero split: split_if_asm)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   852
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   853
end
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   854
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   855
lemma cond_map_pmf:
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   856
  assumes "set_pmf p \<inter> f -` s \<noteq> {}"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   857
  shows "cond_pmf (map_pmf f p) s = map_pmf f (cond_pmf p (f -` s))"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   858
proof -
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   859
  have *: "set_pmf (map_pmf f p) \<inter> s \<noteq> {}"
59665
37adca7fd48f add set_pmf lemmas to simpset
hoelzl
parents: 59664
diff changeset
   860
    using assms by auto
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   861
  { fix x
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   862
    have "ereal (pmf (map_pmf f (cond_pmf p (f -` s))) x) =
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   863
      emeasure p (f -` s \<inter> f -` {x}) / emeasure p (f -` s)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   864
      unfolding ereal_pmf_map cond_pmf.rep_eq[OF assms] by (simp add: nn_integral_uniform_measure)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   865
    also have "f -` s \<inter> f -` {x} = (if x \<in> s then f -` {x} else {})"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   866
      by auto
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   867
    also have "emeasure p (if x \<in> s then f -` {x} else {}) / emeasure p (f -` s) =
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   868
      ereal (pmf (cond_pmf (map_pmf f p) s) x)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   869
      using measure_measure_pmf_not_zero[OF *]
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   870
      by (simp add: pmf_cond[OF *] ereal_divide' ereal_pmf_map measure_pmf.emeasure_eq_measure[symmetric]
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   871
               del: ereal_divide)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   872
    finally have "ereal (pmf (cond_pmf (map_pmf f p) s) x) = ereal (pmf (map_pmf f (cond_pmf p (f -` s))) x)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   873
      by simp }
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   874
  then show ?thesis
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   875
    by (intro pmf_eqI) simp
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   876
qed
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   877
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   878
lemma bind_cond_pmf_cancel:
59670
dee043d19729 generalized bind_cond_pmf_cancel
hoelzl
parents: 59667
diff changeset
   879
  assumes [simp]: "\<And>x. x \<in> set_pmf p \<Longrightarrow> set_pmf q \<inter> {y. R x y} \<noteq> {}"
dee043d19729 generalized bind_cond_pmf_cancel
hoelzl
parents: 59667
diff changeset
   880
  assumes [simp]: "\<And>y. y \<in> set_pmf q \<Longrightarrow> set_pmf p \<inter> {x. R x y} \<noteq> {}"
dee043d19729 generalized bind_cond_pmf_cancel
hoelzl
parents: 59667
diff changeset
   881
  assumes [simp]: "\<And>x y. x \<in> set_pmf p \<Longrightarrow> y \<in> set_pmf q \<Longrightarrow> R x y \<Longrightarrow> measure q {y. R x y} = measure p {x. R x y}"
dee043d19729 generalized bind_cond_pmf_cancel
hoelzl
parents: 59667
diff changeset
   882
  shows "bind_pmf p (\<lambda>x. cond_pmf q {y. R x y}) = q"
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   883
proof (rule pmf_eqI)
59670
dee043d19729 generalized bind_cond_pmf_cancel
hoelzl
parents: 59667
diff changeset
   884
  fix i
dee043d19729 generalized bind_cond_pmf_cancel
hoelzl
parents: 59667
diff changeset
   885
  have "ereal (pmf (bind_pmf p (\<lambda>x. cond_pmf q {y. R x y})) i) =
dee043d19729 generalized bind_cond_pmf_cancel
hoelzl
parents: 59667
diff changeset
   886
    (\<integral>\<^sup>+x. ereal (pmf q i / measure p {x. R x i}) * ereal (indicator {x. R x i} x) \<partial>p)"
dee043d19729 generalized bind_cond_pmf_cancel
hoelzl
parents: 59667
diff changeset
   887
    by (auto simp add: ereal_pmf_bind AE_measure_pmf_iff pmf_cond pmf_eq_0_set_pmf intro!: nn_integral_cong_AE)
dee043d19729 generalized bind_cond_pmf_cancel
hoelzl
parents: 59667
diff changeset
   888
  also have "\<dots> = (pmf q i * measure p {x. R x i}) / measure p {x. R x i}"
dee043d19729 generalized bind_cond_pmf_cancel
hoelzl
parents: 59667
diff changeset
   889
    by (simp add: pmf_nonneg measure_nonneg zero_ereal_def[symmetric] ereal_indicator
dee043d19729 generalized bind_cond_pmf_cancel
hoelzl
parents: 59667
diff changeset
   890
                  nn_integral_cmult measure_pmf.emeasure_eq_measure)
dee043d19729 generalized bind_cond_pmf_cancel
hoelzl
parents: 59667
diff changeset
   891
  also have "\<dots> = pmf q i"
dee043d19729 generalized bind_cond_pmf_cancel
hoelzl
parents: 59667
diff changeset
   892
    by (cases "pmf q i = 0") (simp_all add: pmf_eq_0_set_pmf measure_measure_pmf_not_zero)
dee043d19729 generalized bind_cond_pmf_cancel
hoelzl
parents: 59667
diff changeset
   893
  finally show "pmf (bind_pmf p (\<lambda>x. cond_pmf q {y. R x y})) i = pmf q i"
dee043d19729 generalized bind_cond_pmf_cancel
hoelzl
parents: 59667
diff changeset
   894
    by simp
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   895
qed
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   896
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   897
subsection \<open> Relator \<close>
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   898
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   899
inductive rel_pmf :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'a pmf \<Rightarrow> 'b pmf \<Rightarrow> bool"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   900
for R p q
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   901
where
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59665
diff changeset
   902
  "\<lbrakk> \<And>x y. (x, y) \<in> set_pmf pq \<Longrightarrow> R x y;
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   903
     map_pmf fst pq = p; map_pmf snd pq = q \<rbrakk>
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   904
  \<Longrightarrow> rel_pmf R p q"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
   905
59681
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   906
lemma rel_pmfI:
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   907
  assumes R: "rel_set R (set_pmf p) (set_pmf q)"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   908
  assumes eq: "\<And>x y. x \<in> set_pmf p \<Longrightarrow> y \<in> set_pmf q \<Longrightarrow> R x y \<Longrightarrow>
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   909
    measure p {x. R x y} = measure q {y. R x y}"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   910
  shows "rel_pmf R p q"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   911
proof
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   912
  let ?pq = "bind_pmf p (\<lambda>x. bind_pmf (cond_pmf q {y. R x y}) (\<lambda>y. return_pmf (x, y)))"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   913
  have "\<And>x. x \<in> set_pmf p \<Longrightarrow> set_pmf q \<inter> {y. R x y} \<noteq> {}"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   914
    using R by (auto simp: rel_set_def)
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   915
  then show "\<And>x y. (x, y) \<in> set_pmf ?pq \<Longrightarrow> R x y"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   916
    by auto
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   917
  show "map_pmf fst ?pq = p"
60068
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
   918
    by (simp add: map_bind_pmf bind_return_pmf')
59681
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   919
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   920
  show "map_pmf snd ?pq = q"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   921
    using R eq
60068
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
   922
    apply (simp add: bind_cond_pmf_cancel map_bind_pmf bind_return_pmf')
59681
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   923
    apply (rule bind_cond_pmf_cancel)
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   924
    apply (auto simp: rel_set_def)
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   925
    done
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   926
qed
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   927
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   928
lemma rel_pmf_imp_rel_set: "rel_pmf R p q \<Longrightarrow> rel_set R (set_pmf p) (set_pmf q)"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   929
  by (force simp add: rel_pmf.simps rel_set_def)
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   930
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   931
lemma rel_pmfD_measure:
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   932
  assumes rel_R: "rel_pmf R p q" and R: "\<And>a b. R a b \<Longrightarrow> R a y \<longleftrightarrow> R x b"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   933
  assumes "x \<in> set_pmf p" "y \<in> set_pmf q"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   934
  shows "measure p {x. R x y} = measure q {y. R x y}"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   935
proof -
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   936
  from rel_R obtain pq where pq: "\<And>x y. (x, y) \<in> set_pmf pq \<Longrightarrow> R x y"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   937
    and eq: "p = map_pmf fst pq" "q = map_pmf snd pq"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   938
    by (auto elim: rel_pmf.cases)
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   939
  have "measure p {x. R x y} = measure pq {x. R (fst x) y}"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   940
    by (simp add: eq map_pmf_rep_eq measure_distr)
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   941
  also have "\<dots> = measure pq {y. R x (snd y)}"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   942
    by (intro measure_pmf.finite_measure_eq_AE)
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   943
       (auto simp: AE_measure_pmf_iff R dest!: pq)
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   944
  also have "\<dots> = measure q {y. R x y}"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   945
    by (simp add: eq map_pmf_rep_eq measure_distr)
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   946
  finally show "measure p {x. R x y} = measure q {y. R x y}" .
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   947
qed
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   948
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   949
lemma rel_pmf_iff_measure:
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   950
  assumes "symp R" "transp R"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   951
  shows "rel_pmf R p q \<longleftrightarrow>
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   952
    rel_set R (set_pmf p) (set_pmf q) \<and>
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   953
    (\<forall>x\<in>set_pmf p. \<forall>y\<in>set_pmf q. R x y \<longrightarrow> measure p {x. R x y} = measure q {y. R x y})"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   954
  by (safe intro!: rel_pmf_imp_rel_set rel_pmfI)
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   955
     (auto intro!: rel_pmfD_measure dest: sympD[OF \<open>symp R\<close>] transpD[OF \<open>transp R\<close>])
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   956
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   957
lemma quotient_rel_set_disjoint:
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   958
  "equivp R \<Longrightarrow> C \<in> UNIV // {(x, y). R x y} \<Longrightarrow> rel_set R A B \<Longrightarrow> A \<inter> C = {} \<longleftrightarrow> B \<inter> C = {}"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
   959
  using in_quotient_imp_closed[of UNIV "{(x, y). R x y}" C]
59681
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   960
  by (auto 0 0 simp: equivp_equiv rel_set_def set_eq_iff elim: equivpE)
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   961
     (blast dest: equivp_symp)+
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   962
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   963
lemma quotientD: "equiv X R \<Longrightarrow> A \<in> X // R \<Longrightarrow> x \<in> A \<Longrightarrow> A = R `` {x}"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   964
  by (metis Image_singleton_iff equiv_class_eq_iff quotientE)
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   965
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   966
lemma rel_pmf_iff_equivp:
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   967
  assumes "equivp R"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   968
  shows "rel_pmf R p q \<longleftrightarrow> (\<forall>C\<in>UNIV // {(x, y). R x y}. measure p C = measure q C)"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   969
    (is "_ \<longleftrightarrow>   (\<forall>C\<in>_//?R. _)")
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   970
proof (subst rel_pmf_iff_measure, safe)
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   971
  show "symp R" "transp R"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   972
    using assms by (auto simp: equivp_reflp_symp_transp)
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   973
next
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   974
  fix C assume C: "C \<in> UNIV // ?R" and R: "rel_set R (set_pmf p) (set_pmf q)"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   975
  assume eq: "\<forall>x\<in>set_pmf p. \<forall>y\<in>set_pmf q. R x y \<longrightarrow> measure p {x. R x y} = measure q {y. R x y}"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
   976
59681
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   977
  show "measure p C = measure q C"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   978
  proof cases
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   979
    assume "p \<inter> C = {}"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
   980
    moreover then have "q \<inter> C = {}"
59681
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   981
      using quotient_rel_set_disjoint[OF assms C R] by simp
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   982
    ultimately show ?thesis
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   983
      unfolding measure_pmf_zero_iff[symmetric] by simp
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   984
  next
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   985
    assume "p \<inter> C \<noteq> {}"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
   986
    moreover then have "q \<inter> C \<noteq> {}"
59681
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   987
      using quotient_rel_set_disjoint[OF assms C R] by simp
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   988
    ultimately obtain x y where in_set: "x \<in> set_pmf p" "y \<in> set_pmf q" and in_C: "x \<in> C" "y \<in> C"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   989
      by auto
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   990
    then have "R x y"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   991
      using in_quotient_imp_in_rel[of UNIV ?R C x y] C assms
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   992
      by (simp add: equivp_equiv)
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   993
    with in_set eq have "measure p {x. R x y} = measure q {y. R x y}"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   994
      by auto
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   995
    moreover have "{y. R x y} = C"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   996
      using assms `x \<in> C` C quotientD[of UNIV ?R C x] by (simp add: equivp_equiv)
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   997
    moreover have "{x. R x y} = C"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   998
      using assms `y \<in> C` C quotientD[of UNIV "?R" C y] sympD[of R]
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
   999
      by (auto simp add: equivp_equiv elim: equivpE)
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1000
    ultimately show ?thesis
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1001
      by auto
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1002
  qed
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1003
next
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1004
  assume eq: "\<forall>C\<in>UNIV // ?R. measure p C = measure q C"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1005
  show "rel_set R (set_pmf p) (set_pmf q)"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1006
    unfolding rel_set_def
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1007
  proof safe
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1008
    fix x assume x: "x \<in> set_pmf p"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1009
    have "{y. R x y} \<in> UNIV // ?R"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1010
      by (auto simp: quotient_def)
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1011
    with eq have *: "measure q {y. R x y} = measure p {y. R x y}"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1012
      by auto
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1013
    have "measure q {y. R x y} \<noteq> 0"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1014
      using x assms unfolding * by (auto simp: measure_pmf_zero_iff set_eq_iff dest: equivp_reflp)
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1015
    then show "\<exists>y\<in>set_pmf q. R x y"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1016
      unfolding measure_pmf_zero_iff by auto
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1017
  next
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1018
    fix y assume y: "y \<in> set_pmf q"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1019
    have "{x. R x y} \<in> UNIV // ?R"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1020
      using assms by (auto simp: quotient_def dest: equivp_symp)
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1021
    with eq have *: "measure p {x. R x y} = measure q {x. R x y}"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1022
      by auto
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1023
    have "measure p {x. R x y} \<noteq> 0"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1024
      using y assms unfolding * by (auto simp: measure_pmf_zero_iff set_eq_iff dest: equivp_reflp)
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1025
    then show "\<exists>x\<in>set_pmf p. R x y"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1026
      unfolding measure_pmf_zero_iff by auto
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1027
  qed
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1028
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1029
  fix x y assume "x \<in> set_pmf p" "y \<in> set_pmf q" "R x y"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1030
  have "{y. R x y} \<in> UNIV // ?R" "{x. R x y} = {y. R x y}"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1031
    using assms `R x y` by (auto simp: quotient_def dest: equivp_symp equivp_transp)
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1032
  with eq show "measure p {x. R x y} = measure q {y. R x y}"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1033
    by auto
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1034
qed
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1035
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1036
bnf pmf: "'a pmf" map: map_pmf sets: set_pmf bd : "natLeq" rel: rel_pmf
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1037
proof -
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1038
  show "map_pmf id = id" by (rule map_pmf_id)
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59665
diff changeset
  1039
  show "\<And>f g. map_pmf (f \<circ> g) = map_pmf f \<circ> map_pmf g" by (rule map_pmf_compose)
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1040
  show "\<And>f g::'a \<Rightarrow> 'b. \<And>p. (\<And>x. x \<in> set_pmf p \<Longrightarrow> f x = g x) \<Longrightarrow> map_pmf f p = map_pmf g p"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1041
    by (intro map_pmf_cong refl)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1042
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1043
  show "\<And>f::'a \<Rightarrow> 'b. set_pmf \<circ> map_pmf f = op ` f \<circ> set_pmf"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1044
    by (rule pmf_set_map)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1045
60595
804dfdc82835 premises in 'show' are treated like 'assume';
wenzelm
parents: 60495
diff changeset
  1046
  show "(card_of (set_pmf p), natLeq) \<in> ordLeq" for p :: "'s pmf"
804dfdc82835 premises in 'show' are treated like 'assume';
wenzelm
parents: 60495
diff changeset
  1047
  proof -
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1048
    have "(card_of (set_pmf p), card_of (UNIV :: nat set)) \<in> ordLeq"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1049
      by (rule card_of_ordLeqI[where f="to_nat_on (set_pmf p)"])
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1050
         (auto intro: countable_set_pmf)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1051
    also have "(card_of (UNIV :: nat set), natLeq) \<in> ordLeq"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1052
      by (metis Field_natLeq card_of_least natLeq_Well_order)
60595
804dfdc82835 premises in 'show' are treated like 'assume';
wenzelm
parents: 60495
diff changeset
  1053
    finally show ?thesis .
804dfdc82835 premises in 'show' are treated like 'assume';
wenzelm
parents: 60495
diff changeset
  1054
  qed
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1055
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1056
  show "\<And>R. rel_pmf R =
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1057
         (BNF_Def.Grp {x. set_pmf x \<subseteq> {(x, y). R x y}} (map_pmf fst))\<inverse>\<inverse> OO
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1058
         BNF_Def.Grp {x. set_pmf x \<subseteq> {(x, y). R x y}} (map_pmf snd)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1059
     by (auto simp add: fun_eq_iff BNF_Def.Grp_def OO_def rel_pmf.simps)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1060
60595
804dfdc82835 premises in 'show' are treated like 'assume';
wenzelm
parents: 60495
diff changeset
  1061
  show "rel_pmf R OO rel_pmf S \<le> rel_pmf (R OO S)"
804dfdc82835 premises in 'show' are treated like 'assume';
wenzelm
parents: 60495
diff changeset
  1062
    for R :: "'a \<Rightarrow> 'b \<Rightarrow> bool" and S :: "'b \<Rightarrow> 'c \<Rightarrow> bool"
804dfdc82835 premises in 'show' are treated like 'assume';
wenzelm
parents: 60495
diff changeset
  1063
  proof -
804dfdc82835 premises in 'show' are treated like 'assume';
wenzelm
parents: 60495
diff changeset
  1064
    { fix p q r
804dfdc82835 premises in 'show' are treated like 'assume';
wenzelm
parents: 60495
diff changeset
  1065
      assume pq: "rel_pmf R p q"
804dfdc82835 premises in 'show' are treated like 'assume';
wenzelm
parents: 60495
diff changeset
  1066
        and qr:"rel_pmf S q r"
804dfdc82835 premises in 'show' are treated like 'assume';
wenzelm
parents: 60495
diff changeset
  1067
      from pq obtain pq where pq: "\<And>x y. (x, y) \<in> set_pmf pq \<Longrightarrow> R x y"
804dfdc82835 premises in 'show' are treated like 'assume';
wenzelm
parents: 60495
diff changeset
  1068
        and p: "p = map_pmf fst pq" and q: "q = map_pmf snd pq" by cases auto
804dfdc82835 premises in 'show' are treated like 'assume';
wenzelm
parents: 60495
diff changeset
  1069
      from qr obtain qr where qr: "\<And>y z. (y, z) \<in> set_pmf qr \<Longrightarrow> S y z"
804dfdc82835 premises in 'show' are treated like 'assume';
wenzelm
parents: 60495
diff changeset
  1070
        and q': "q = map_pmf fst qr" and r: "r = map_pmf snd qr" by cases auto
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1071
60595
804dfdc82835 premises in 'show' are treated like 'assume';
wenzelm
parents: 60495
diff changeset
  1072
      def pr \<equiv> "bind_pmf pq (\<lambda>xy. bind_pmf (cond_pmf qr {yz. fst yz = snd xy}) (\<lambda>yz. return_pmf (fst xy, snd yz)))"
804dfdc82835 premises in 'show' are treated like 'assume';
wenzelm
parents: 60495
diff changeset
  1073
      have pr_welldefined: "\<And>y. y \<in> q \<Longrightarrow> qr \<inter> {yz. fst yz = y} \<noteq> {}"
804dfdc82835 premises in 'show' are treated like 'assume';
wenzelm
parents: 60495
diff changeset
  1074
        by (force simp: q')
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1075
60595
804dfdc82835 premises in 'show' are treated like 'assume';
wenzelm
parents: 60495
diff changeset
  1076
      have "rel_pmf (R OO S) p r"
804dfdc82835 premises in 'show' are treated like 'assume';
wenzelm
parents: 60495
diff changeset
  1077
      proof (rule rel_pmf.intros)
804dfdc82835 premises in 'show' are treated like 'assume';
wenzelm
parents: 60495
diff changeset
  1078
        fix x z assume "(x, z) \<in> pr"
804dfdc82835 premises in 'show' are treated like 'assume';
wenzelm
parents: 60495
diff changeset
  1079
        then have "\<exists>y. (x, y) \<in> pq \<and> (y, z) \<in> qr"
804dfdc82835 premises in 'show' are treated like 'assume';
wenzelm
parents: 60495
diff changeset
  1080
          by (auto simp: q pr_welldefined pr_def split_beta)
804dfdc82835 premises in 'show' are treated like 'assume';
wenzelm
parents: 60495
diff changeset
  1081
        with pq qr show "(R OO S) x z"
804dfdc82835 premises in 'show' are treated like 'assume';
wenzelm
parents: 60495
diff changeset
  1082
          by blast
804dfdc82835 premises in 'show' are treated like 'assume';
wenzelm
parents: 60495
diff changeset
  1083
      next
804dfdc82835 premises in 'show' are treated like 'assume';
wenzelm
parents: 60495
diff changeset
  1084
        have "map_pmf snd pr = map_pmf snd (bind_pmf q (\<lambda>y. cond_pmf qr {yz. fst yz = y}))"
804dfdc82835 premises in 'show' are treated like 'assume';
wenzelm
parents: 60495
diff changeset
  1085
          by (simp add: pr_def q split_beta bind_map_pmf map_pmf_def[symmetric] map_bind_pmf map_pmf_comp)
804dfdc82835 premises in 'show' are treated like 'assume';
wenzelm
parents: 60495
diff changeset
  1086
        then show "map_pmf snd pr = r"
804dfdc82835 premises in 'show' are treated like 'assume';
wenzelm
parents: 60495
diff changeset
  1087
          unfolding r q' bind_map_pmf by (subst (asm) bind_cond_pmf_cancel) (auto simp: eq_commute)
804dfdc82835 premises in 'show' are treated like 'assume';
wenzelm
parents: 60495
diff changeset
  1088
      qed (simp add: pr_def map_bind_pmf split_beta map_pmf_def[symmetric] p map_pmf_comp)
804dfdc82835 premises in 'show' are treated like 'assume';
wenzelm
parents: 60495
diff changeset
  1089
    }
804dfdc82835 premises in 'show' are treated like 'assume';
wenzelm
parents: 60495
diff changeset
  1090
    then show ?thesis
804dfdc82835 premises in 'show' are treated like 'assume';
wenzelm
parents: 60495
diff changeset
  1091
      by(auto simp add: le_fun_def)
804dfdc82835 premises in 'show' are treated like 'assume';
wenzelm
parents: 60495
diff changeset
  1092
  qed
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1093
qed (fact natLeq_card_order natLeq_cinfinite)+
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1094
59665
37adca7fd48f add set_pmf lemmas to simpset
hoelzl
parents: 59664
diff changeset
  1095
lemma rel_pmf_conj[simp]:
37adca7fd48f add set_pmf lemmas to simpset
hoelzl
parents: 59664
diff changeset
  1096
  "rel_pmf (\<lambda>x y. P \<and> Q x y) x y \<longleftrightarrow> P \<and> rel_pmf Q x y"
37adca7fd48f add set_pmf lemmas to simpset
hoelzl
parents: 59664
diff changeset
  1097
  "rel_pmf (\<lambda>x y. Q x y \<and> P) x y \<longleftrightarrow> P \<and> rel_pmf Q x y"
37adca7fd48f add set_pmf lemmas to simpset
hoelzl
parents: 59664
diff changeset
  1098
  using set_pmf_not_empty by (fastforce simp: pmf.in_rel subset_eq)+
37adca7fd48f add set_pmf lemmas to simpset
hoelzl
parents: 59664
diff changeset
  1099
37adca7fd48f add set_pmf lemmas to simpset
hoelzl
parents: 59664
diff changeset
  1100
lemma rel_pmf_top[simp]: "rel_pmf top = top"
37adca7fd48f add set_pmf lemmas to simpset
hoelzl
parents: 59664
diff changeset
  1101
  by (auto simp: pmf.in_rel[abs_def] fun_eq_iff map_fst_pair_pmf map_snd_pair_pmf
37adca7fd48f add set_pmf lemmas to simpset
hoelzl
parents: 59664
diff changeset
  1102
           intro: exI[of _ "pair_pmf x y" for x y])
37adca7fd48f add set_pmf lemmas to simpset
hoelzl
parents: 59664
diff changeset
  1103
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1104
lemma rel_pmf_return_pmf1: "rel_pmf R (return_pmf x) M \<longleftrightarrow> (\<forall>a\<in>M. R x a)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1105
proof safe
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1106
  fix a assume "a \<in> M" "rel_pmf R (return_pmf x) M"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1107
  then obtain pq where *: "\<And>a b. (a, b) \<in> set_pmf pq \<Longrightarrow> R a b"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1108
    and eq: "return_pmf x = map_pmf fst pq" "M = map_pmf snd pq"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1109
    by (force elim: rel_pmf.cases)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1110
  moreover have "set_pmf (return_pmf x) = {x}"
59665
37adca7fd48f add set_pmf lemmas to simpset
hoelzl
parents: 59664
diff changeset
  1111
    by simp
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1112
  with `a \<in> M` have "(x, a) \<in> pq"
59665
37adca7fd48f add set_pmf lemmas to simpset
hoelzl
parents: 59664
diff changeset
  1113
    by (force simp: eq)
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1114
  with * show "R x a"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1115
    by auto
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1116
qed (auto intro!: rel_pmf.intros[where pq="pair_pmf (return_pmf x) M"]
59665
37adca7fd48f add set_pmf lemmas to simpset
hoelzl
parents: 59664
diff changeset
  1117
          simp: map_fst_pair_pmf map_snd_pair_pmf)
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1118
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1119
lemma rel_pmf_return_pmf2: "rel_pmf R M (return_pmf x) \<longleftrightarrow> (\<forall>a\<in>M. R a x)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1120
  by (subst pmf.rel_flip[symmetric]) (simp add: rel_pmf_return_pmf1)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1121
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1122
lemma rel_return_pmf[simp]: "rel_pmf R (return_pmf x1) (return_pmf x2) = R x1 x2"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1123
  unfolding rel_pmf_return_pmf2 set_return_pmf by simp
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1124
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1125
lemma rel_pmf_False[simp]: "rel_pmf (\<lambda>x y. False) x y = False"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1126
  unfolding pmf.in_rel fun_eq_iff using set_pmf_not_empty by fastforce
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1127
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1128
lemma rel_pmf_rel_prod:
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1129
  "rel_pmf (rel_prod R S) (pair_pmf A A') (pair_pmf B B') \<longleftrightarrow> rel_pmf R A B \<and> rel_pmf S A' B'"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1130
proof safe
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1131
  assume "rel_pmf (rel_prod R S) (pair_pmf A A') (pair_pmf B B')"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1132
  then obtain pq where pq: "\<And>a b c d. ((a, c), (b, d)) \<in> set_pmf pq \<Longrightarrow> R a b \<and> S c d"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1133
    and eq: "map_pmf fst pq = pair_pmf A A'" "map_pmf snd pq = pair_pmf B B'"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1134
    by (force elim: rel_pmf.cases)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1135
  show "rel_pmf R A B"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1136
  proof (rule rel_pmf.intros)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1137
    let ?f = "\<lambda>(a, b). (fst a, fst b)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1138
    have [simp]: "(\<lambda>x. fst (?f x)) = fst o fst" "(\<lambda>x. snd (?f x)) = fst o snd"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1139
      by auto
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1140
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1141
    show "map_pmf fst (map_pmf ?f pq) = A"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1142
      by (simp add: map_pmf_comp pmf.map_comp[symmetric] eq map_fst_pair_pmf)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1143
    show "map_pmf snd (map_pmf ?f pq) = B"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1144
      by (simp add: map_pmf_comp pmf.map_comp[symmetric] eq map_fst_pair_pmf)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1145
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1146
    fix a b assume "(a, b) \<in> set_pmf (map_pmf ?f pq)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1147
    then obtain c d where "((a, c), (b, d)) \<in> set_pmf pq"
59665
37adca7fd48f add set_pmf lemmas to simpset
hoelzl
parents: 59664
diff changeset
  1148
      by auto
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1149
    from pq[OF this] show "R a b" ..
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1150
  qed
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1151
  show "rel_pmf S A' B'"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1152
  proof (rule rel_pmf.intros)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1153
    let ?f = "\<lambda>(a, b). (snd a, snd b)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1154
    have [simp]: "(\<lambda>x. fst (?f x)) = snd o fst" "(\<lambda>x. snd (?f x)) = snd o snd"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1155
      by auto
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1156
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1157
    show "map_pmf fst (map_pmf ?f pq) = A'"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1158
      by (simp add: map_pmf_comp pmf.map_comp[symmetric] eq map_snd_pair_pmf)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1159
    show "map_pmf snd (map_pmf ?f pq) = B'"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1160
      by (simp add: map_pmf_comp pmf.map_comp[symmetric] eq map_snd_pair_pmf)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1161
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1162
    fix c d assume "(c, d) \<in> set_pmf (map_pmf ?f pq)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1163
    then obtain a b where "((a, c), (b, d)) \<in> set_pmf pq"
59665
37adca7fd48f add set_pmf lemmas to simpset
hoelzl
parents: 59664
diff changeset
  1164
      by auto
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1165
    from pq[OF this] show "S c d" ..
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1166
  qed
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1167
next
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1168
  assume "rel_pmf R A B" "rel_pmf S A' B'"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1169
  then obtain Rpq Spq
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1170
    where Rpq: "\<And>a b. (a, b) \<in> set_pmf Rpq \<Longrightarrow> R a b"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1171
        "map_pmf fst Rpq = A" "map_pmf snd Rpq = B"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1172
      and Spq: "\<And>a b. (a, b) \<in> set_pmf Spq \<Longrightarrow> S a b"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1173
        "map_pmf fst Spq = A'" "map_pmf snd Spq = B'"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1174
    by (force elim: rel_pmf.cases)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1175
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1176
  let ?f = "(\<lambda>((a, c), (b, d)). ((a, b), (c, d)))"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1177
  let ?pq = "map_pmf ?f (pair_pmf Rpq Spq)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1178
  have [simp]: "(\<lambda>x. fst (?f x)) = (\<lambda>(a, b). (fst a, fst b))" "(\<lambda>x. snd (?f x)) = (\<lambda>(a, b). (snd a, snd b))"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1179
    by auto
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1180
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1181
  show "rel_pmf (rel_prod R S) (pair_pmf A A') (pair_pmf B B')"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1182
    by (rule rel_pmf.intros[where pq="?pq"])
59665
37adca7fd48f add set_pmf lemmas to simpset
hoelzl
parents: 59664
diff changeset
  1183
       (auto simp: map_snd_pair_pmf map_fst_pair_pmf map_pmf_comp Rpq Spq
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1184
                   map_pair)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1185
qed
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1186
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59665
diff changeset
  1187
lemma rel_pmf_reflI:
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1188
  assumes "\<And>x. x \<in> set_pmf p \<Longrightarrow> P x x"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1189
  shows "rel_pmf P p p"
59665
37adca7fd48f add set_pmf lemmas to simpset
hoelzl
parents: 59664
diff changeset
  1190
  by (rule rel_pmf.intros[where pq="map_pmf (\<lambda>x. (x, x)) p"])
37adca7fd48f add set_pmf lemmas to simpset
hoelzl
parents: 59664
diff changeset
  1191
     (auto simp add: pmf.map_comp o_def assms)
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1192
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1193
context
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1194
begin
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1195
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1196
interpretation pmf_as_measure .
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1197
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1198
definition "join_pmf M = bind_pmf M (\<lambda>x. x)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1199
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1200
lemma bind_eq_join_pmf: "bind_pmf M f = join_pmf (map_pmf f M)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1201
  unfolding join_pmf_def bind_map_pmf ..
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1202
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1203
lemma join_eq_bind_pmf: "join_pmf M = bind_pmf M id"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1204
  by (simp add: join_pmf_def id_def)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1205
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1206
lemma pmf_join: "pmf (join_pmf N) i = (\<integral>M. pmf M i \<partial>measure_pmf N)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1207
  unfolding join_pmf_def pmf_bind ..
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1208
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1209
lemma ereal_pmf_join: "ereal (pmf (join_pmf N) i) = (\<integral>\<^sup>+M. pmf M i \<partial>measure_pmf N)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1210
  unfolding join_pmf_def ereal_pmf_bind ..
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1211
59665
37adca7fd48f add set_pmf lemmas to simpset
hoelzl
parents: 59664
diff changeset
  1212
lemma set_pmf_join_pmf[simp]: "set_pmf (join_pmf f) = (\<Union>p\<in>set_pmf f. set_pmf p)"
37adca7fd48f add set_pmf lemmas to simpset
hoelzl
parents: 59664
diff changeset
  1213
  by (simp add: join_pmf_def)
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1214
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1215
lemma join_return_pmf: "join_pmf (return_pmf M) = M"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1216
  by (simp add: integral_return pmf_eq_iff pmf_join return_pmf.rep_eq)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1217
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1218
lemma map_join_pmf: "map_pmf f (join_pmf AA) = join_pmf (map_pmf (map_pmf f) AA)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1219
  by (simp add: join_pmf_def map_pmf_def bind_assoc_pmf bind_return_pmf)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1220
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1221
lemma join_map_return_pmf: "join_pmf (map_pmf return_pmf A) = A"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1222
  by (simp add: join_pmf_def map_pmf_def bind_assoc_pmf bind_return_pmf bind_return_pmf')
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1223
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1224
end
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1225
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1226
lemma rel_pmf_joinI:
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1227
  assumes "rel_pmf (rel_pmf P) p q"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1228
  shows "rel_pmf P (join_pmf p) (join_pmf q)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1229
proof -
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1230
  from assms obtain pq where p: "p = map_pmf fst pq"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1231
    and q: "q = map_pmf snd pq"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1232
    and P: "\<And>x y. (x, y) \<in> set_pmf pq \<Longrightarrow> rel_pmf P x y"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1233
    by cases auto
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59665
diff changeset
  1234
  from P obtain PQ
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1235
    where PQ: "\<And>x y a b. \<lbrakk> (x, y) \<in> set_pmf pq; (a, b) \<in> set_pmf (PQ x y) \<rbrakk> \<Longrightarrow> P a b"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1236
    and x: "\<And>x y. (x, y) \<in> set_pmf pq \<Longrightarrow> map_pmf fst (PQ x y) = x"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1237
    and y: "\<And>x y. (x, y) \<in> set_pmf pq \<Longrightarrow> map_pmf snd (PQ x y) = y"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1238
    by(metis rel_pmf.simps)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1239
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1240
  let ?r = "bind_pmf pq (\<lambda>(x, y). PQ x y)"
59665
37adca7fd48f add set_pmf lemmas to simpset
hoelzl
parents: 59664
diff changeset
  1241
  have "\<And>a b. (a, b) \<in> set_pmf ?r \<Longrightarrow> P a b" by (auto intro: PQ)
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1242
  moreover have "map_pmf fst ?r = join_pmf p" "map_pmf snd ?r = join_pmf q"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1243
    by (simp_all add: p q x y join_pmf_def map_bind_pmf bind_map_pmf split_def cong: bind_pmf_cong)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1244
  ultimately show ?thesis ..
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1245
qed
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1246
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1247
lemma rel_pmf_bindI:
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1248
  assumes pq: "rel_pmf R p q"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1249
  and fg: "\<And>x y. R x y \<Longrightarrow> rel_pmf P (f x) (g y)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1250
  shows "rel_pmf P (bind_pmf p f) (bind_pmf q g)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1251
  unfolding bind_eq_join_pmf
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1252
  by (rule rel_pmf_joinI)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1253
     (auto simp add: pmf.rel_map intro: pmf.rel_mono[THEN le_funD, THEN le_funD, THEN le_boolD, THEN mp, OF _ pq] fg)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1254
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1255
text {*
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1256
  Proof that @{const rel_pmf} preserves orders.
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59665
diff changeset
  1257
  Antisymmetry proof follows Thm. 1 in N. Saheb-Djahromi, Cpo's of measures for nondeterminism,
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59665
diff changeset
  1258
  Theoretical Computer Science 12(1):19--37, 1980,
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1259
  @{url "http://dx.doi.org/10.1016/0304-3975(80)90003-1"}
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1260
*}
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1261
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59665
diff changeset
  1262
lemma
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1263
  assumes *: "rel_pmf R p q"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1264
  and refl: "reflp R" and trans: "transp R"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1265
  shows measure_Ici: "measure p {y. R x y} \<le> measure q {y. R x y}" (is ?thesis1)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1266
  and measure_Ioi: "measure p {y. R x y \<and> \<not> R y x} \<le> measure q {y. R x y \<and> \<not> R y x}" (is ?thesis2)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1267
proof -
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1268
  from * obtain pq
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1269
    where pq: "\<And>x y. (x, y) \<in> set_pmf pq \<Longrightarrow> R x y"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1270
    and p: "p = map_pmf fst pq"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1271
    and q: "q = map_pmf snd pq"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1272
    by cases auto
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1273
  show ?thesis1 ?thesis2 unfolding p q map_pmf_rep_eq using refl trans
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1274
    by(auto 4 3 simp add: measure_distr reflpD AE_measure_pmf_iff intro!: measure_pmf.finite_measure_mono_AE dest!: pq elim: transpE)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1275
qed
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1276
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1277
lemma rel_pmf_inf:
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1278
  fixes p q :: "'a pmf"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1279
  assumes 1: "rel_pmf R p q"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1280
  assumes 2: "rel_pmf R q p"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1281
  and refl: "reflp R" and trans: "transp R"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1282
  shows "rel_pmf (inf R R\<inverse>\<inverse>) p q"
59681
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1283
proof (subst rel_pmf_iff_equivp, safe)
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1284
  show "equivp (inf R R\<inverse>\<inverse>)"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1285
    using trans refl by (auto simp: equivp_reflp_symp_transp intro: sympI transpI reflpI dest: transpD reflpD)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1286
59681
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1287
  fix C assume "C \<in> UNIV // {(x, y). inf R R\<inverse>\<inverse> x y}"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1288
  then obtain x where C: "C = {y. R x y \<and> R y x}"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1289
    by (auto elim: quotientE)
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1290
59670
dee043d19729 generalized bind_cond_pmf_cancel
hoelzl
parents: 59667
diff changeset
  1291
  let ?R = "\<lambda>x y. R x y \<and> R y x"
dee043d19729 generalized bind_cond_pmf_cancel
hoelzl
parents: 59667
diff changeset
  1292
  let ?\<mu>R = "\<lambda>y. measure q {x. ?R x y}"
59681
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1293
  have "measure p {y. ?R x y} = measure p ({y. R x y} - {y. R x y \<and> \<not> R y x})"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1294
    by(auto intro!: arg_cong[where f="measure p"])
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1295
  also have "\<dots> = measure p {y. R x y} - measure p {y. R x y \<and> \<not> R y x}"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1296
    by (rule measure_pmf.finite_measure_Diff) auto
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1297
  also have "measure p {y. R x y \<and> \<not> R y x} = measure q {y. R x y \<and> \<not> R y x}"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1298
    using 1 2 refl trans by(auto intro!: Orderings.antisym measure_Ioi)
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1299
  also have "measure p {y. R x y} = measure q {y. R x y}"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1300
    using 1 2 refl trans by(auto intro!: Orderings.antisym measure_Ici)
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1301
  also have "measure q {y. R x y} - measure q {y. R x y \<and> \<not> R y x} =
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1302
    measure q ({y. R x y} - {y. R x y \<and> \<not> R y x})"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1303
    by(rule measure_pmf.finite_measure_Diff[symmetric]) auto
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1304
  also have "\<dots> = ?\<mu>R x"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1305
    by(auto intro!: arg_cong[where f="measure q"])
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1306
  finally show "measure p C = measure q C"
f24ab09e4c37 rel_pmf on equivalence relation
hoelzl
parents: 59670
diff changeset
  1307
    by (simp add: C conj_commute)
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1308
qed
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1309
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1310
lemma rel_pmf_antisym:
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1311
  fixes p q :: "'a pmf"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1312
  assumes 1: "rel_pmf R p q"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1313
  assumes 2: "rel_pmf R q p"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1314
  and refl: "reflp R" and trans: "transp R" and antisym: "antisymP R"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1315
  shows "p = q"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1316
proof -
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1317
  from 1 2 refl trans have "rel_pmf (inf R R\<inverse>\<inverse>) p q" by(rule rel_pmf_inf)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1318
  also have "inf R R\<inverse>\<inverse> = op ="
59665
37adca7fd48f add set_pmf lemmas to simpset
hoelzl
parents: 59664
diff changeset
  1319
    using refl antisym by (auto intro!: ext simp add: reflpD dest: antisymD)
59664
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1320
  finally show ?thesis unfolding pmf.rel_eq .
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1321
qed
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1322
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1323
lemma reflp_rel_pmf: "reflp R \<Longrightarrow> reflp (rel_pmf R)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1324
by(blast intro: reflpI rel_pmf_reflI reflpD)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1325
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1326
lemma antisymP_rel_pmf:
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1327
  "\<lbrakk> reflp R; transp R; antisymP R \<rbrakk>
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1328
  \<Longrightarrow> antisymP (rel_pmf R)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1329
by(rule antisymI)(blast intro: rel_pmf_antisym)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1330
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1331
lemma transp_rel_pmf:
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1332
  assumes "transp R"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1333
  shows "transp (rel_pmf R)"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1334
proof (rule transpI)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1335
  fix x y z
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1336
  assume "rel_pmf R x y" and "rel_pmf R y z"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1337
  hence "rel_pmf (R OO R) x z" by (simp add: pmf.rel_compp relcompp.relcompI)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1338
  thus "rel_pmf R x z"
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1339
    using assms by (metis (no_types) pmf.rel_mono rev_predicate2D transp_relcompp_less_eq)
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1340
qed
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1341
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1342
subsection \<open> Distributions \<close>
224741ede5ae build pmf's on bind
hoelzl
parents: 59557
diff changeset
  1343
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1344
context
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1345
begin
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1346
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1347
interpretation pmf_as_function .
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1348
59093
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  1349
subsubsection \<open> Bernoulli Distribution \<close>
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  1350
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1351
lift_definition bernoulli_pmf :: "real \<Rightarrow> bool pmf" is
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1352
  "\<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
  1353
  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
  1354
           split: split_max split_min)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1355
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1356
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
  1357
  by transfer simp
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1358
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1359
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
  1360
  by transfer simp
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1361
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1362
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
  1363
  by (auto simp add: set_pmf_iff UNIV_bool)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1364
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59665
diff changeset
  1365
lemma nn_integral_bernoulli_pmf[simp]:
59002
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
  1366
  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
  1367
  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
  1368
  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
  1369
     (auto simp: UNIV_bool field_simps)
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
  1370
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59665
diff changeset
  1371
lemma integral_bernoulli_pmf[simp]:
59002
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
  1372
  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
  1373
  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
  1374
  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
  1375
59525
dfe6449aecd8 more lemmas
Andreas Lochbihler
parents: 59496
diff changeset
  1376
lemma pmf_bernoulli_half [simp]: "pmf (bernoulli_pmf (1 / 2)) x = 1 / 2"
dfe6449aecd8 more lemmas
Andreas Lochbihler
parents: 59496
diff changeset
  1377
by(cases x) simp_all
dfe6449aecd8 more lemmas
Andreas Lochbihler
parents: 59496
diff changeset
  1378
dfe6449aecd8 more lemmas
Andreas Lochbihler
parents: 59496
diff changeset
  1379
lemma measure_pmf_bernoulli_half: "measure_pmf (bernoulli_pmf (1 / 2)) = uniform_count_measure UNIV"
dfe6449aecd8 more lemmas
Andreas Lochbihler
parents: 59496
diff changeset
  1380
by(rule measure_eqI)(simp_all add: nn_integral_pmf[symmetric] emeasure_uniform_count_measure nn_integral_count_space_finite sets_uniform_count_measure)
dfe6449aecd8 more lemmas
Andreas Lochbihler
parents: 59496
diff changeset
  1381
59093
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  1382
subsubsection \<open> Geometric Distribution \<close>
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  1383
60602
37588fbe39f9 generalized geometric distribution
hoelzl
parents: 60595
diff changeset
  1384
context
37588fbe39f9 generalized geometric distribution
hoelzl
parents: 60595
diff changeset
  1385
  fixes p :: real assumes p[arith]: "0 < p" "p \<le> 1"
37588fbe39f9 generalized geometric distribution
hoelzl
parents: 60595
diff changeset
  1386
begin
37588fbe39f9 generalized geometric distribution
hoelzl
parents: 60595
diff changeset
  1387
37588fbe39f9 generalized geometric distribution
hoelzl
parents: 60595
diff changeset
  1388
lift_definition geometric_pmf :: "nat pmf" is "\<lambda>n. (1 - p)^n * p"
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1389
proof
60602
37588fbe39f9 generalized geometric distribution
hoelzl
parents: 60595
diff changeset
  1390
  have "(\<Sum>i. ereal (p * (1 - p) ^ i)) = ereal (p * (1 / (1 - (1 - p))))"
37588fbe39f9 generalized geometric distribution
hoelzl
parents: 60595
diff changeset
  1391
    by (intro sums_suminf_ereal sums_mult geometric_sums) auto
37588fbe39f9 generalized geometric distribution
hoelzl
parents: 60595
diff changeset
  1392
  then show "(\<integral>\<^sup>+ x. ereal ((1 - p)^x * p) \<partial>count_space UNIV) = 1"
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1393
    by (simp add: nn_integral_count_space_nat field_simps)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1394
qed simp
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1395
60602
37588fbe39f9 generalized geometric distribution
hoelzl
parents: 60595
diff changeset
  1396
lemma pmf_geometric[simp]: "pmf geometric_pmf n = (1 - p)^n * p"
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1397
  by transfer rule
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1398
60602
37588fbe39f9 generalized geometric distribution
hoelzl
parents: 60595
diff changeset
  1399
end
37588fbe39f9 generalized geometric distribution
hoelzl
parents: 60595
diff changeset
  1400
37588fbe39f9 generalized geometric distribution
hoelzl
parents: 60595
diff changeset
  1401
lemma set_pmf_geometric: "0 < p \<Longrightarrow> p < 1 \<Longrightarrow> set_pmf (geometric_pmf p) = UNIV"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1402
  by (auto simp: set_pmf_iff)
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1403
59093
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  1404
subsubsection \<open> Uniform Multiset Distribution \<close>
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  1405
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1406
context
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1407
  fixes M :: "'a multiset" assumes M_not_empty: "M \<noteq> {#}"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1408
begin
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1409
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1410
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
  1411
proof
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59665
diff changeset
  1412
  show "(\<integral>\<^sup>+ x. ereal (real (count M x) / real (size M)) \<partial>count_space UNIV) = 1"
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1413
    using M_not_empty
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1414
    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
  1415
                  setsum_divide_distrib[symmetric])
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1416
       (auto simp: size_multiset_overloaded_eq intro!: setsum.cong)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1417
qed simp
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1418
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1419
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
  1420
  by transfer rule
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1421
60495
d7ff0a1df90a renamed Multiset.set_of to the canonical set_mset
nipkow
parents: 60068
diff changeset
  1422
lemma set_pmf_of_multiset[simp]: "set_pmf pmf_of_multiset = set_mset M"
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1423
  by (auto simp: set_pmf_iff)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1424
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1425
end
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1426
59093
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  1427
subsubsection \<open> Uniform Distribution \<close>
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  1428
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1429
context
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1430
  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
  1431
begin
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1432
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1433
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
  1434
proof
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59665
diff changeset
  1435
  show "(\<integral>\<^sup>+ x. ereal (indicator S x / real (card S)) \<partial>count_space UNIV) = 1"
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1436
    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
  1437
qed simp
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1438
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1439
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
  1440
  by transfer rule
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1441
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1442
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
  1443
  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
  1444
59002
2c8b2fb54b88 cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
hoelzl
parents: 59000
diff changeset
  1445
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
  1446
  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
  1447
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1448
lemma nn_integral_pmf_of_set':
60068
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
  1449
  "(\<And>x. x \<in> S \<Longrightarrow> f x \<ge> 0) \<Longrightarrow> nn_integral (measure_pmf pmf_of_set) f = setsum f S / card S"
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
  1450
apply(subst nn_integral_measure_pmf_finite, simp_all add: S_finite)
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
  1451
apply(simp add: setsum_ereal_left_distrib[symmetric])
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
  1452
apply(subst ereal_divide', simp add: S_not_empty S_finite)
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
  1453
apply(simp add: ereal_times_divide_eq one_ereal_def[symmetric])
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
  1454
done
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
  1455
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1456
lemma nn_integral_pmf_of_set:
60068
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
  1457
  "nn_integral (measure_pmf pmf_of_set) f = setsum (max 0 \<circ> f) S / card S"
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
  1458
apply(subst nn_integral_max_0[symmetric])
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
  1459
apply(subst nn_integral_pmf_of_set')
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
  1460
apply simp_all
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
  1461
done
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
  1462
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
  1463
lemma integral_pmf_of_set:
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
  1464
  "integral\<^sup>L (measure_pmf pmf_of_set) f = setsum f S / card S"
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
  1465
apply(simp add: real_lebesgue_integral_def integrable_measure_pmf_finite nn_integral_pmf_of_set S_finite)
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
  1466
apply(subst real_of_ereal_minus')
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
  1467
 apply(simp add: ereal_max_0 S_finite del: ereal_max)
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
  1468
apply(simp add: ereal_max_0 S_finite S_not_empty del: ereal_max)
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
  1469
apply(simp add: field_simps S_finite S_not_empty)
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
  1470
apply(subst setsum.distrib[symmetric])
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
  1471
apply(rule setsum.cong; simp_all)
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
  1472
done
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
  1473
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1474
end
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1475
60068
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
  1476
lemma pmf_of_set_singleton: "pmf_of_set {x} = return_pmf x"
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
  1477
by(rule pmf_eqI)(simp add: indicator_def)
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
  1478
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1479
lemma map_pmf_of_set_inj:
60068
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
  1480
  assumes f: "inj_on f A"
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
  1481
  and [simp]: "A \<noteq> {}" "finite A"
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
  1482
  shows "map_pmf f (pmf_of_set A) = pmf_of_set (f ` A)" (is "?lhs = ?rhs")
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
  1483
proof(rule pmf_eqI)
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
  1484
  fix i
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
  1485
  show "pmf ?lhs i = pmf ?rhs i"
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
  1486
  proof(cases "i \<in> f ` A")
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
  1487
    case True
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
  1488
    then obtain i' where "i = f i'" "i' \<in> A" by auto
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
  1489
    thus ?thesis using f by(simp add: card_image pmf_map_inj)
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
  1490
  next
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
  1491
    case False
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
  1492
    hence "pmf ?lhs i = 0" by(simp add: pmf_eq_0_set_pmf set_map_pmf)
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
  1493
    moreover have "pmf ?rhs i = 0" using False by simp
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
  1494
    ultimately show ?thesis by simp
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
  1495
  qed
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
  1496
qed
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
  1497
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
  1498
lemma bernoulli_pmf_half_conv_pmf_of_set: "bernoulli_pmf (1 / 2) = pmf_of_set UNIV"
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
  1499
by(rule pmf_eqI) simp_all
ef2123db900c add various lemmas about pmfs
Andreas Lochbihler
parents: 59731
diff changeset
  1500
59093
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  1501
subsubsection \<open> Poisson Distribution \<close>
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  1502
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  1503
context
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  1504
  fixes rate :: real assumes rate_pos: "0 < rate"
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  1505
begin
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  1506
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  1507
lift_definition poisson_pmf :: "nat pmf" is "\<lambda>k. rate ^ k / fact k * exp (-rate)"
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 59667
diff changeset
  1508
proof  (* by Manuel Eberl *)
59093
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  1509
  have summable: "summable (\<lambda>x::nat. rate ^ x / fact x)" using summable_exp
59557
ebd8ecacfba6 establish unique preferred fact names
haftmann
parents: 59527
diff changeset
  1510
    by (simp add: field_simps divide_inverse [symmetric])
59093
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  1511
  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
  1512
          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
  1513
    by (simp add: field_simps nn_integral_cmult[symmetric])
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  1514
  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
  1515
    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
  1516
  also have "... = exp rate" unfolding exp_def
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 59667
diff changeset
  1517
    by (simp add: field_simps divide_inverse [symmetric])
59093
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  1518
  also have "ereal (exp (-rate)) * ereal (exp rate) = 1"
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  1519
    by (simp add: mult_exp_exp)
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 59667
diff changeset
  1520
  finally show "(\<integral>\<^sup>+ x. ereal (rate ^ x / (fact x) * exp (- rate)) \<partial>count_space UNIV) = 1" .
59093
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  1521
qed (simp add: rate_pos[THEN less_imp_le])
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  1522
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  1523
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
  1524
  by transfer rule
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  1525
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  1526
lemma set_pmf_poisson[simp]: "set_pmf poisson_pmf = UNIV"
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  1527
  using rate_pos by (auto simp: set_pmf_iff)
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  1528
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1529
end
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1530
59093
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  1531
subsubsection \<open> Binomial Distribution \<close>
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  1532
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  1533
context
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  1534
  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
  1535
begin
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  1536
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  1537
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
  1538
proof
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  1539
  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
  1540
    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
  1541
    using p_le_1 p_nonneg by (subst nn_integral_count_space') auto
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  1542
  also have "(\<Sum>k\<le>n. real (n choose k) * p ^ k * (1 - p) ^ (n - k)) = (p + (1 - p)) ^ n"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61424
diff changeset
  1543
    by (subst binomial_ring) (simp add: atLeast0AtMost)
59093
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  1544
  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
  1545
    by simp
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  1546
qed (insert p_nonneg p_le_1, simp)
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  1547
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  1548
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
  1549
  by transfer rule
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  1550
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  1551
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
  1552
  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
  1553
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  1554
end
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  1555
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  1556
end
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  1557
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  1558
lemma set_pmf_binomial_0[simp]: "set_pmf (binomial_pmf n 0) = {0}"
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  1559
  by (simp add: set_pmf_binomial_eq)
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  1560
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  1561
lemma set_pmf_binomial_1[simp]: "set_pmf (binomial_pmf n 1) = {n}"
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  1562
  by (simp add: set_pmf_binomial_eq)
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  1563
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  1564
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
  1565
  by (simp add: set_pmf_binomial_eq)
2b106e58a177 add Poisson and Binomial distribution
hoelzl
parents: 59092
diff changeset
  1566
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58730
diff changeset
  1567
end