diff -r 933eb9e6a1cc -r 77b453bd616f src/HOL/Probability/Probability_Mass_Function.thy --- a/src/HOL/Probability/Probability_Mass_Function.thy Tue Nov 03 11:20:21 2015 +0100 +++ b/src/HOL/Probability/Probability_Mass_Function.thy Tue Nov 10 14:18:41 2015 +0000 @@ -59,7 +59,7 @@ note singleton_sets = this have "?M < (\x\X. ?M / Suc n)" using `?M \ 0` - by (simp add: `card X = Suc (Suc n)` real_eq_of_nat[symmetric] real_of_nat_Suc field_simps less_le measure_nonneg) + by (simp add: `card X = Suc (Suc n)` of_nat_Suc field_simps less_le measure_nonneg) also have "\ \ (\x\X. ?m x)" by (rule setsum_mono) fact also have "\ = measure M (\x\X. {x})" @@ -956,7 +956,7 @@ lemma quotient_rel_set_disjoint: "equivp R \ C \ UNIV // {(x, y). R x y} \ rel_set R A B \ A \ C = {} \ B \ C = {}" - using in_quotient_imp_closed[of UNIV "{(x, y). R x y}" C] + using in_quotient_imp_closed[of UNIV "{(x, y). R x y}" C] by (auto 0 0 simp: equivp_equiv rel_set_def set_eq_iff elim: equivpE) (blast dest: equivp_symp)+ @@ -973,17 +973,17 @@ next fix C assume C: "C \ UNIV // ?R" and R: "rel_set R (set_pmf p) (set_pmf q)" assume eq: "\x\set_pmf p. \y\set_pmf q. R x y \ measure p {x. R x y} = measure q {y. R x y}" - + show "measure p C = measure q C" proof cases assume "p \ C = {}" - moreover then have "q \ C = {}" + moreover then have "q \ C = {}" using quotient_rel_set_disjoint[OF assms C R] by simp ultimately show ?thesis unfolding measure_pmf_zero_iff[symmetric] by simp next assume "p \ C \ {}" - moreover then have "q \ C \ {}" + moreover then have "q \ C \ {}" using quotient_rel_set_disjoint[OF assms C R] by simp ultimately obtain x y where in_set: "x \ set_pmf p" "y \ set_pmf q" and in_C: "x \ C" "y \ C" by auto @@ -1068,11 +1068,11 @@ and p: "p = map_pmf fst pq" and q: "q = map_pmf snd pq" by cases auto from qr obtain qr where qr: "\y z. (y, z) \ set_pmf qr \ S y z" and q': "q = map_pmf fst qr" and r: "r = map_pmf snd qr" by cases auto - + def pr \ "bind_pmf pq (\xy. bind_pmf (cond_pmf qr {yz. fst yz = snd xy}) (\yz. return_pmf (fst xy, snd yz)))" have pr_welldefined: "\y. y \ q \ qr \ {yz. fst yz = y} \ {}" by (force simp: q') - + have "rel_pmf (R OO S) p r" proof (rule rel_pmf.intros) fix x z assume "(x, z) \ pr" @@ -1283,7 +1283,7 @@ proof (subst rel_pmf_iff_equivp, safe) show "equivp (inf R R\\)" using trans refl by (auto simp: equivp_reflp_symp_transp intro: sympI transpI reflpI dest: transpD reflpD) - + fix C assume "C \ UNIV // {(x, y). inf R R\\ x y}" then obtain x where C: "C = {y. R x y \ R y x}" by (auto elim: quotientE) @@ -1399,7 +1399,7 @@ end lemma set_pmf_geometric: "0 < p \ p < 1 \ set_pmf (geometric_pmf p) = UNIV" - by (auto simp: set_pmf_iff) + by (auto simp: set_pmf_iff) subsubsection \ Uniform Multiset Distribution \ @@ -1445,7 +1445,7 @@ lemma emeasure_pmf_of_set[simp]: "emeasure pmf_of_set S = 1" by (rule measure_pmf.emeasure_eq_1_AE) (auto simp: AE_measure_pmf_iff) -lemma nn_integral_pmf_of_set': +lemma nn_integral_pmf_of_set': "(\x. x \ S \ f x \ 0) \ nn_integral (measure_pmf pmf_of_set) f = setsum f S / card S" apply(subst nn_integral_measure_pmf_finite, simp_all add: S_finite) apply(simp add: setsum_ereal_left_distrib[symmetric]) @@ -1453,7 +1453,7 @@ apply(simp add: ereal_times_divide_eq one_ereal_def[symmetric]) done -lemma nn_integral_pmf_of_set: +lemma nn_integral_pmf_of_set: "nn_integral (measure_pmf pmf_of_set) f = setsum (max 0 \ f) S / card S" apply(subst nn_integral_max_0[symmetric]) apply(subst nn_integral_pmf_of_set') @@ -1476,7 +1476,7 @@ lemma pmf_of_set_singleton: "pmf_of_set {x} = return_pmf x" by(rule pmf_eqI)(simp add: indicator_def) -lemma map_pmf_of_set_inj: +lemma map_pmf_of_set_inj: assumes f: "inj_on f A" and [simp]: "A \ {}" "finite A" shows "map_pmf f (pmf_of_set A) = pmf_of_set (f ` A)" (is "?lhs = ?rhs") @@ -1540,7 +1540,7 @@ ereal (\k\n. real (n choose k) * p ^ k * (1 - p) ^ (n - k))" using p_le_1 p_nonneg by (subst nn_integral_count_space') auto also have "(\k\n. real (n choose k) * p ^ k * (1 - p) ^ (n - k)) = (p + (1 - p)) ^ n" - by (subst binomial_ring) (simp add: atLeast0AtMost real_of_nat_def) + by (subst binomial_ring) (simp add: atLeast0AtMost) finally show "(\\<^sup>+ x. ereal (real (n choose x) * p ^ x * (1 - p) ^ (n - x)) \count_space UNIV) = 1" by simp qed (insert p_nonneg p_le_1, simp)