# HG changeset patch # User hoelzl # Date 1412605627 -7200 # Node ID 5484f6079bcd52185ca0e67468894aceacb03af7 # Parent 1b11669a5c664beddfbf8c3ea408b75217bd1128 add type for probability mass functions, i.e. discrete probability distribution diff -r 1b11669a5c66 -r 5484f6079bcd src/HOL/Probability/Complete_Measure.thy --- a/src/HOL/Probability/Complete_Measure.thy Mon Oct 06 13:42:48 2014 +0200 +++ b/src/HOL/Probability/Complete_Measure.thy Mon Oct 06 16:27:07 2014 +0200 @@ -3,7 +3,7 @@ *) theory Complete_Measure -imports Bochner_Integration + imports Bochner_Integration Probability_Measure begin definition @@ -314,4 +314,19 @@ qed auto qed +lemma (in prob_space) prob_space_completion: "prob_space (completion M)" + by (rule prob_spaceI) (simp add: emeasure_space_1) + +lemma null_sets_completionI: "N \ null_sets M \ N \ null_sets (completion M)" + by (auto simp: null_sets_def) + +lemma AE_completion: "(AE x in M. P x) \ (AE x in completion M. P x)" + unfolding eventually_ae_filter by (auto intro: null_sets_completionI) + +lemma null_sets_completion_iff: "N \ sets M \ N \ null_sets (completion M) \ N \ null_sets M" + by (auto simp: null_sets_def) + +lemma AE_completion_iff: "{x\space M. P x} \ sets M \ (AE x in M. P x) \ (AE x in completion M. P x)" + by (simp add: AE_iff_null null_sets_completion_iff) + end diff -r 1b11669a5c66 -r 5484f6079bcd src/HOL/Probability/Probability.thy --- a/src/HOL/Probability/Probability.thy Mon Oct 06 13:42:48 2014 +0200 +++ b/src/HOL/Probability/Probability.thy Mon Oct 06 16:27:07 2014 +0200 @@ -7,6 +7,7 @@ Projective_Limit Independent_Family Distributions + Probability_Mass_Function begin end diff -r 1b11669a5c66 -r 5484f6079bcd src/HOL/Probability/Probability_Mass_Function.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Probability/Probability_Mass_Function.thy Mon Oct 06 16:27:07 2014 +0200 @@ -0,0 +1,358 @@ +theory Probability_Mass_Function + imports Probability_Measure +begin + +lemma sets_Pair: "{x} \ sets M1 \ {y} \ sets M2 \ {(x, y)} \ sets (M1 \\<^sub>M M2)" + using pair_measureI[of "{x}" M1 "{y}" M2] by simp + +lemma finite_subset_card: + assumes X: "infinite X" shows "\A\X. finite A \ card A = n" +proof (induct n) + case (Suc n) then guess A .. note A = this + with X obtain x where "x \ X" "x \ A" + by (metis subset_antisym subset_eq) + with A show ?case + by (intro exI[of _ "insert x A"]) auto +qed (simp cong: conj_cong) + +lemma (in prob_space) countable_support: + "countable {x. measure M {x} \ 0}" +proof - + let ?m = "\x. measure M {x}" + have *: "{x. ?m x \ 0} = (\n. {x. inverse (real (Suc n)) < ?m x})" + by (auto intro!: measure_nonneg reals_Archimedean order_le_neq_trans) + have **: "\n. finite {x. inverse (Suc n) < ?m x}" + proof (rule ccontr) + fix n assume "infinite {x. inverse (Suc n) < ?m x}" (is "infinite ?X") + then obtain X where "finite X" "card X = Suc (Suc n)" "X \ ?X" + by (metis finite_subset_card) + from this(3) have *: "\x. x \ X \ 1 / Suc n \ ?m x" + by (auto simp: inverse_eq_divide) + { fix x assume "x \ X" + from *[OF this] have "?m x \ 0" by auto + then have "{x} \ sets M" by (auto dest: measure_notin_sets) } + note singleton_sets = this + have "1 < (\x\X. 1 / Suc n)" + by (simp add: `card X = Suc (Suc n)` real_eq_of_nat[symmetric] real_of_nat_Suc) + also have "\ \ (\x\X. ?m x)" + by (rule setsum_mono) fact + also have "\ = measure M (\x\X. {x})" + using singleton_sets `finite X` + by (intro finite_measure_finite_Union[symmetric]) (auto simp: disjoint_family_on_def) + finally show False + using prob_le_1[of "\x\X. {x}"] by arith + qed + show ?thesis + unfolding * by (intro countable_UN countableI_type countable_finite[OF **]) +qed + +lemma measure_count_space: "measure (count_space A) X = (if X \ A then card X else 0)" + unfolding measure_def + by (cases "finite X") (simp_all add: emeasure_notin_sets) + +typedef 'a pmf = "{M :: 'a measure. prob_space M \ sets M = UNIV \ (AE x in M. measure M {x} \ 0)}" + morphisms measure_pmf Abs_pmf + apply (intro exI[of _ "uniform_measure (count_space UNIV) {undefined}"]) + apply (auto intro!: prob_space_uniform_measure simp: measure_count_space) + apply (subst uniform_measure_def) + apply (simp add: AE_density AE_count_space split: split_indicator) + done + +declare [[coercion measure_pmf]] + +lemma prob_space_measure_pmf: "prob_space (measure_pmf p)" + using pmf.measure_pmf[of p] by auto + +interpretation measure_pmf!: prob_space "measure_pmf M" for M + by (rule prob_space_measure_pmf) + +locale pmf_as_measure +begin + +setup_lifting type_definition_pmf + +end + +context +begin + +interpretation pmf_as_measure . + +lift_definition pmf :: "'a pmf \ 'a \ real" is "\M x. measure M {x}" . + +lift_definition set_pmf :: "'a pmf \ 'a set" is "\M. {x. measure M {x} \ 0}" . + +lift_definition map_pmf :: "('a \ 'b) \ 'a pmf \ 'b pmf" is + "\f M. distr M (count_space UNIV) f" +proof safe + fix M and f :: "'a \ 'b" + let ?D = "distr M (count_space UNIV) f" + assume "prob_space M" and [simp]: "sets M = UNIV" and ae: "AE x in M. measure M {x} \ 0" + interpret prob_space M by fact + from ae have "AE x in M. measure M (f -` {f x}) \ 0" + proof eventually_elim + fix x + have "measure M {x} \ measure M (f -` {f x})" + by (intro finite_measure_mono) auto + then show "measure M {x} \ 0 \ measure M (f -` {f x}) \ 0" + using measure_nonneg[of M "{x}"] by auto + qed + then show "AE x in ?D. measure ?D {x} \ 0" + by (simp add: AE_distr_iff measure_distr measurable_def) +qed (auto simp: measurable_def prob_space.prob_space_distr) + +declare [[coercion set_pmf]] + +lemma countable_set_pmf: "countable (set_pmf p)" + by transfer (metis prob_space.countable_support) + +lemma sets_measure_pmf[simp]: "sets (measure_pmf p) = UNIV" + by transfer metis + +lemma space_measure_pmf[simp]: "space (measure_pmf p) = UNIV" + using sets_eq_imp_space_eq[of "measure_pmf p" "count_space UNIV"] by simp + +lemma measurable_pmf_measure1[simp]: "measurable (M :: 'a pmf) N = UNIV \ space N" + by (auto simp: measurable_def) + +lemma measurable_pmf_measure2[simp]: "measurable N (M :: 'a pmf) = measurable N (count_space UNIV)" + by (intro measurable_cong_sets) simp_all + +lemma pmf_positive: "x \ set_pmf p \ 0 < pmf p x" + by transfer (simp add: less_le measure_nonneg) + +lemma pmf_nonneg: "0 \ pmf p x" + by transfer (simp add: measure_nonneg) + +lemma emeasure_pmf_single: + fixes M :: "'a pmf" + shows "emeasure M {x} = pmf M x" + by transfer (simp add: finite_measure.emeasure_eq_measure[OF prob_space.finite_measure]) + +lemma AE_measure_pmf: "AE x in (M::'a pmf). x \ M" + by transfer simp + +lemma emeasure_pmf_single_eq_zero_iff: + fixes M :: "'a pmf" + shows "emeasure M {y} = 0 \ y \ M" + by transfer (simp add: finite_measure.emeasure_eq_measure[OF prob_space.finite_measure]) + +lemma AE_measure_pmf_iff: "(AE x in measure_pmf M. P x) \ (\y\M. P y)" +proof - + { fix y assume y: "y \ M" and P: "AE x in M. P x" "\ P y" + with P have "AE x in M. x \ y" + by auto + with y have False + by (simp add: emeasure_pmf_single_eq_zero_iff AE_iff_measurable[OF _ refl]) } + then show ?thesis + using AE_measure_pmf[of M] by auto +qed + +lemma measure_pmf_eq_density: "measure_pmf p = density (count_space UNIV) (pmf p)" +proof (transfer, elim conjE) + fix M :: "'a measure" assume [simp]: "sets M = UNIV" and ae: "AE x in M. measure M {x} \ 0" + assume "prob_space M" then interpret prob_space M . + show "M = density (count_space UNIV) (\x. ereal (measure M {x}))" + proof (rule measure_eqI) + fix A :: "'a set" + have "(\\<^sup>+ x. ereal (measure M {x}) * indicator A x \count_space UNIV) = + (\\<^sup>+ x. emeasure M {x} * indicator (A \ {x. measure M {x} \ 0}) x \count_space UNIV)" + by (auto intro!: nn_integral_cong simp: emeasure_eq_measure split: split_indicator) + also have "\ = (\\<^sup>+ x. emeasure M {x} \count_space (A \ {x. measure M {x} \ 0}))" + by (subst nn_integral_restrict_space[symmetric]) (auto simp: restrict_count_space) + also have "\ = emeasure M (\x\(A \ {x. measure M {x} \ 0}). {x})" + by (intro emeasure_UN_countable[symmetric] countable_Int2 countable_support) + (auto simp: disjoint_family_on_def) + also have "\ = emeasure M A" + using ae by (intro emeasure_eq_AE) auto + finally show " emeasure M A = emeasure (density (count_space UNIV) (\x. ereal (measure M {x}))) A" + using emeasure_space_1 by (simp add: emeasure_density) + qed simp +qed + +lemma set_pmf_not_empty: "set_pmf M \ {}" + using AE_measure_pmf[of M] by (intro notI) simp + +lemma set_pmf_iff: "x \ set_pmf M \ pmf M x \ 0" + by transfer simp + +lemma emeasure_pmf: "emeasure (M::'a pmf) M = 1" +proof - + have "emeasure (M::'a pmf) M = emeasure (M::'a pmf) (space M)" + by (intro emeasure_eq_AE) (simp_all add: AE_measure_pmf) + then show ?thesis + using measure_pmf.emeasure_space_1 by simp +qed + +lemma map_pmf_id[simp]: "map_pmf id = id" + by (rule, transfer) (auto simp: emeasure_distr measurable_def intro!: measure_eqI) + +lemma map_pmf_compose: "map_pmf (f \ g) = map_pmf f \ map_pmf g" + by (rule, transfer) (simp add: distr_distr[symmetric, where N="count_space UNIV"] measurable_def) + +lemma map_pmf_cong: + assumes "p = q" + shows "(\x. x \ set_pmf q \ f x = g x) \ map_pmf f p = map_pmf g q" + unfolding `p = q`[symmetric] measure_pmf_inject[symmetric] map_pmf.rep_eq + by (auto simp add: emeasure_distr AE_measure_pmf_iff intro!: emeasure_eq_AE measure_eqI) + +lemma pmf_set_map: + fixes f :: "'a \ 'b" + shows "set_pmf \ map_pmf f = op ` f \ set_pmf" +proof (rule, transfer, clarsimp simp add: measure_distr measurable_def) + fix f :: "'a \ 'b" and M :: "'a measure" + assume "prob_space M" and ae: "AE x in M. measure M {x} \ 0" and [simp]: "sets M = UNIV" + interpret prob_space M by fact + show "{x. measure M (f -` {x}) \ 0} = f ` {x. measure M {x} \ 0}" + proof safe + fix x assume "measure M (f -` {x}) \ 0" + moreover have "measure M (f -` {x}) = measure M {y. f y = x \ measure M {y} \ 0}" + using ae by (intro finite_measure_eq_AE) auto + ultimately have "{y. f y = x \ measure M {y} \ 0} \ {}" + by (metis measure_empty) + then show "x \ f ` {x. measure M {x} \ 0}" + by auto + next + fix x assume "measure M {x} \ 0" + then have "0 < measure M {x}" + using measure_nonneg[of M "{x}"] by auto + also have "measure M {x} \ measure M (f -` {f x})" + by (intro finite_measure_mono) auto + finally show "measure M (f -` {f x}) = 0 \ False" + by simp + qed +qed + +context + fixes f :: "'a \ real" + assumes nonneg: "\x. 0 \ f x" + assumes prob: "(\\<^sup>+x. f x \count_space UNIV) = 1" +begin + +lift_definition embed_pmf :: "'a pmf" is "density (count_space UNIV) (ereal \ f)" +proof (intro conjI) + have *[simp]: "\x y. ereal (f y) * indicator {x} y = ereal (f x) * indicator {x} y" + by (simp split: split_indicator) + show "AE x in density (count_space UNIV) (ereal \ f). + measure (density (count_space UNIV) (ereal \ f)) {x} \ 0" + by (simp add: AE_density nonneg emeasure_density measure_def nn_integral_cmult_indicator) + show "prob_space (density (count_space UNIV) (ereal \ f))" + by default (simp add: emeasure_density prob) +qed simp + +lemma pmf_embed_pmf: "pmf embed_pmf x = f x" +proof transfer + have *[simp]: "\x y. ereal (f y) * indicator {x} y = ereal (f x) * indicator {x} y" + by (simp split: split_indicator) + fix x show "measure (density (count_space UNIV) (ereal \ f)) {x} = f x" + by transfer (simp add: measure_def emeasure_density nn_integral_cmult_indicator nonneg) +qed + +end + +lemma embed_pmf_transfer: + "rel_fun (eq_onp (\f::'a \ real. (\x. 0 \ f x) \ (\\<^sup>+x. ereal (f x) \count_space UNIV) = 1)) pmf_as_measure.cr_pmf (\f. density (count_space UNIV) (ereal \ f)) embed_pmf" + by (auto simp: rel_fun_def eq_onp_def embed_pmf.transfer) + +lemma td_pmf_embed_pmf: + "type_definition pmf embed_pmf {f::'a \ real. (\x. 0 \ f x) \ (\\<^sup>+x. ereal (f x) \count_space UNIV) = 1}" + unfolding type_definition_def +proof safe + fix p :: "'a pmf" + have "(\\<^sup>+ x. 1 \measure_pmf p) = 1" + using measure_pmf.emeasure_space_1[of p] by simp + then show *: "(\\<^sup>+ x. ereal (pmf p x) \count_space UNIV) = 1" + by (simp add: measure_pmf_eq_density nn_integral_density pmf_nonneg del: nn_integral_const) + + show "embed_pmf (pmf p) = p" + by (intro measure_pmf_inject[THEN iffD1]) + (simp add: * embed_pmf.rep_eq pmf_nonneg measure_pmf_eq_density[of p] comp_def) +next + fix f :: "'a \ real" assume "\x. 0 \ f x" "(\\<^sup>+x. f x \count_space UNIV) = 1" + then show "pmf (embed_pmf f) = f" + by (auto intro!: pmf_embed_pmf) +qed (rule pmf_nonneg) + +end + +locale pmf_as_function +begin + +setup_lifting td_pmf_embed_pmf + +end + +(* + +definition + "rel_pmf P d1 d2 \ (\p3. (\(x, y) \ set_pmf p3. P x y) \ map_pmf fst p3 = d1 \ map_pmf snd p3 = d2)" + +lift_definition pmf_join :: "real \ 'a pmf \ 'a pmf \ 'a pmf" is + "\p M1 M2. density (count_space UNIV) (\x. p * measure M1 {x} + (1 - p) * measure M2 {x})" +sorry + +lift_definition pmf_single :: "'a \ 'a pmf" is + "\x. uniform_measure (count_space UNIV) {x}" +sorry + +bnf pmf: "'a pmf" map: map_pmf sets: set_pmf bd : "natLeq" rel: pmf_rel +proof - + show "map_pmf id = id" by (rule map_pmf_id) + show "\f g. map_pmf (f \ g) = map_pmf f \ map_pmf g" by (rule map_pmf_compose) + show "\f g::'a \ 'b. \p. (\x. x \ set_pmf p \ f x = g x) \ map_pmf f p = map_pmf g p" + by (intro map_pmg_cong refl) + + show "\f::'a \ 'b. set_pmf \ map_pmf f = op ` f \ set_pmf" + by (rule pmf_set_map) + + { fix p :: "'s pmf" + have "(card_of (set_pmf p), card_of (UNIV :: nat set)) \ ordLeq" + by (rule card_of_ordLeqI[where f="to_nat_on (set_pmf p)"]) + (auto intro: countable_set_pmf inj_on_to_nat_on) + also have "(card_of (UNIV :: nat set), natLeq) \ ordLeq" + by (metis Field_natLeq card_of_least natLeq_Well_order) + finally show "(card_of (set_pmf p), natLeq) \ ordLeq" . } + + show "\R. pmf_rel R = + (BNF_Util.Grp {x. set_pmf x \ {(x, y). R x y}} (map_pmf fst))\\ OO + BNF_Util.Grp {x. set_pmf x \ {(x, y). R x y}} (map_pmf snd)" + by (auto simp add: fun_eq_iff pmf_rel_def BNF_Util.Grp_def OO_def) + + { let ?f = "map_pmf fst" and ?s = "map_pmf snd" + fix R :: "'a \ 'b \ bool" and A assume "\x y. (x, y) \ set_pmf A \ R x y" + fix S :: "'b \ 'c \ bool" and B assume "\y z. (y, z) \ set_pmf B \ S y z" + assume "?f B = ?s A" + have "\C. (\(x, z)\set_pmf C. \y. R x y \ S y z) \ ?f C = ?f A \ ?s C = ?s B" + sorry } +oops + then show "\R::'a \ 'b \ bool. \S::'b \ 'c \ bool. pmf_rel R OO pmf_rel S \ pmf_rel (R OO S)" + by (auto simp add: subset_eq pmf_rel_def fun_eq_iff OO_def Ball_def) +qed (fact natLeq_card_order natLeq_cinfinite)+ + +notepad +begin + fix x y :: "nat \ real" + def IJz \ "rec_nat ((0, 0), \_. 0) (\n ((I, J), z). + let a = x I - (\ji b then I + 1 else I, if b \ a then J + 1 else J), z((I, J) := min a b)))" + def I == "fst \ fst \ IJz" def J == "snd \ fst \ IJz" def z == "snd \ IJz" + let ?a = "\n. x (I n) - (\jp. z 0 p = 0" "I 0 = 0" "J 0 = 0" + by (simp_all add: I_def J_def z_def IJz_def) + have z_Suc[simp]: "\n. z (Suc n) = (z n)((I n, J n) := min (?a n) (?b n))" + by (simp add: z_def I_def J_def IJz_def Let_def split_beta) + have I_Suc[simp]: "\n. I (Suc n) = (if ?a n \ ?b n then I n + 1 else I n)" + by (simp add: z_def I_def J_def IJz_def Let_def split_beta) + have J_Suc[simp]: "\n. J (Suc n) = (if ?b n \ ?a n then J n + 1 else J n)" + by (simp add: z_def I_def J_def IJz_def Let_def split_beta) + + { fix N have "\p. z N p \ 0 \ \nj. z n (i, j)) = x i" + oops +*) + +end +