diff -r d45b78cb86cf -r 6aba668aea78 src/HOL/Analysis/Bochner_Integration.thy --- a/src/HOL/Analysis/Bochner_Integration.thy Wed Apr 25 21:29:02 2018 +0100 +++ b/src/HOL/Analysis/Bochner_Integration.thy Thu Apr 26 19:51:32 2018 +0200 @@ -474,8 +474,8 @@ by (subst simple_bochner_integral_partition[OF f(1), where g="\x. ennreal (f x)" and v=enn2real]) (auto intro: f simple_function_compose1 elim: simple_bochner_integrable.cases intro!: sum.cong ennreal_cong_mult - simp: sum_ennreal[symmetric] ac_simps ennreal_mult - simp del: sum_ennreal) + simp: ac_simps ennreal_mult + reorient: sum_ennreal) also have "\ = (\\<^sup>+x. f x \M)" using f by (intro nn_integral_eq_simple_integral[symmetric]) @@ -503,7 +503,7 @@ using simple_bochner_integrable_compose2[of "\x y. norm (x - y)" M "s" "t"] s t by (auto intro!: simple_bochner_integral_eq_nn_integral) also have "\ \ (\\<^sup>+x. ennreal (norm (f x - s x)) + ennreal (norm (f x - t x)) \M)" - by (auto intro!: nn_integral_mono simp: ennreal_plus[symmetric] simp del: ennreal_plus) + by (auto intro!: nn_integral_mono reorient: ennreal_plus) (metis (erased, hide_lams) add_diff_cancel_left add_diff_eq diff_add_eq order_trans norm_minus_commute norm_triangle_ineq4 order_refl) also have "\ = ?S + ?T" @@ -593,7 +593,7 @@ proof (intro always_eventually allI) fix i have "?f i \ (\\<^sup>+ x. (norm (f x - sf i x)) + ennreal (norm (g x - sg i x)) \M)" by (auto intro!: nn_integral_mono norm_diff_triangle_ineq - simp del: ennreal_plus simp add: ennreal_plus[symmetric]) + reorient: ennreal_plus) also have "\ = ?g i" by (intro nn_integral_add) auto finally show "?f i \ ?g i" . @@ -746,7 +746,7 @@ finally have s_fin: "(\\<^sup>+x. norm (s i x) \M) < \" . have "(\\<^sup>+ x. norm (f x) \M) \ (\\<^sup>+ x. ennreal (norm (f x - s i x)) + ennreal (norm (s i x)) \M)" - by (auto intro!: nn_integral_mono simp del: ennreal_plus simp add: ennreal_plus[symmetric]) + by (auto intro!: nn_integral_mono reorient: ennreal_plus) (metis add.commute norm_triangle_sub) also have "\ = (\\<^sup>+x. norm (f x - s i x) \M) + (\\<^sup>+x. norm (s i x) \M)" by (rule nn_integral_add) auto @@ -782,7 +782,7 @@ by (intro simple_bochner_integral_eq_nn_integral) (auto intro: s simple_bochner_integrable_compose2) also have "\ \ (\\<^sup>+x. ennreal (norm (f x - s n x)) + norm (f x) \M)" - by (auto intro!: nn_integral_mono simp del: ennreal_plus simp add: ennreal_plus[symmetric]) + by (auto intro!: nn_integral_mono reorient: ennreal_plus) (metis add.commute norm_minus_commute norm_triangle_sub) also have "\ = ?t n" by (rule nn_integral_add) auto @@ -827,7 +827,7 @@ using tendsto_add[OF \?S \ 0\ \?T \ 0\] by simp qed then have "(\i. norm (?s i - ?t i)) \ 0" - by (simp add: ennreal_0[symmetric] del: ennreal_0) + by (simp reorient: ennreal_0) ultimately have "norm (x - y) = 0" by (rule LIMSEQ_unique) then show "x = y" by simp @@ -1173,7 +1173,7 @@ by (intro simple_bochner_integral_bounded s f) also have "\ < ennreal (e / 2) + e / 2" by (intro add_strict_mono M n m) - also have "\ = e" using \0 by (simp del: ennreal_plus add: ennreal_plus[symmetric]) + also have "\ = e" using \0 by (simp reorient: ennreal_plus) finally show "dist (?s n) (?s m) < e" using \0 by (simp add: dist_norm ennreal_less_iff) qed @@ -1218,7 +1218,7 @@ fix x assume "(\i. u i x) \ u' x" from tendsto_diff[OF tendsto_const[of "u' x"] this] show "(\i. ennreal (norm (u' x - u i x))) \ 0" - by (simp add: tendsto_norm_zero_iff ennreal_0[symmetric] del: ennreal_0) + by (simp add: tendsto_norm_zero_iff reorient: ennreal_0) qed qed (insert bnd w_nonneg, auto) then show ?thesis by simp @@ -2116,7 +2116,7 @@ by auto qed then have "((\n. norm((\x. u n x \M) - (\x. f x \M))) \ 0) F" - by (simp add: ennreal_0[symmetric] del: ennreal_0) + by (simp reorient: ennreal_0) then have "((\n. ((\x. u n x \M) - (\x. f x \M))) \ 0) F" using tendsto_norm_zero_iff by blast then show ?thesis using Lim_null by auto qed @@ -2214,7 +2214,7 @@ ultimately have "(\n. ennreal (norm(u (r n) x))) \ 0" using tendsto_Limsup[of sequentially "\n. ennreal (norm(u (r n) x))"] by auto then have "(\n. norm(u (r n) x)) \ 0" - by (simp add: ennreal_0[symmetric] del: ennreal_0) + by (simp reorient: ennreal_0) then have "(\n. u (r n) x) \ 0" by (simp add: tendsto_norm_zero_iff) }