diff -r 3e3097ac37d1 -r e5abbdee461a src/HOL/Probability/Borel_Space.thy --- a/src/HOL/Probability/Borel_Space.thy Fri Jul 29 08:22:59 2016 +0200 +++ b/src/HOL/Probability/Borel_Space.thy Fri Jul 29 20:34:07 2016 +0200 @@ -414,7 +414,7 @@ shows "f \ borel_measurable (restrict_space M \) \ (\x. f x * indicator \ x) \ borel_measurable M" by (subst measurable_restrict_space_iff) - (auto simp: indicator_def if_distrib[where f="\x. a * x" for a] cong del: if_cong) + (auto simp: indicator_def if_distrib[where f="\x. a * x" for a] cong del: if_weak_cong) lemma borel_measurable_restrict_space_iff_ennreal: fixes f :: "'a \ ennreal" @@ -422,7 +422,7 @@ shows "f \ borel_measurable (restrict_space M \) \ (\x. f x * indicator \ x) \ borel_measurable M" by (subst measurable_restrict_space_iff) - (auto simp: indicator_def if_distrib[where f="\x. a * x" for a] cong del: if_cong) + (auto simp: indicator_def if_distrib[where f="\x. a * x" for a] cong del: if_weak_cong) lemma borel_measurable_restrict_space_iff: fixes f :: "'a \ 'b::real_normed_vector" @@ -430,7 +430,8 @@ shows "f \ borel_measurable (restrict_space M \) \ (\x. indicator \ x *\<^sub>R f x) \ borel_measurable M" by (subst measurable_restrict_space_iff) - (auto simp: indicator_def if_distrib[where f="\x. x *\<^sub>R a" for a] ac_simps cong del: if_cong) + (auto simp: indicator_def if_distrib[where f="\x. x *\<^sub>R a" for a] ac_simps + cong del: if_weak_cong) lemma cbox_borel[measurable]: "cbox a b \ sets borel" by (auto intro: borel_closed)