# HG changeset patch # User paulson # Date 1691082643 -7200 # Node ID 032a4344903e867e2b0dd0cf8979d41e4b2e7aac # Parent ba2afdd29e1d5f1f22b79d5036d73284f776fd6c# Parent a5f6d2fc1b1f3b87d4612957e979d6f57cd09a9d merged diff -r ba2afdd29e1d -r 032a4344903e src/HOL/Analysis/Bochner_Integration.thy --- a/src/HOL/Analysis/Bochner_Integration.thy Tue Aug 01 11:27:55 2023 +0200 +++ b/src/HOL/Analysis/Bochner_Integration.thy Thu Aug 03 19:10:43 2023 +0200 @@ -99,22 +99,19 @@ unfolding A_def by simp qed ultimately have "dist (f x) (F N x) < 1 / Suc (m N x)" - by (auto simp add: F_def L_def) } + by (auto simp: F_def L_def) } note * = this fix x assume "x \ space M" show "(\i. F i x) \ f x" - proof cases - assume "f x = z" + proof (cases "f x = z") + case True then have "\i n. x \ A i n" unfolding A_def by auto - then have "\i. F i x = z" - by (auto simp: F_def) then show ?thesis - using \f x = z\ by auto + by (metis B_imp_A F_def LIMSEQ_def True dist_self) next - assume "f x \ z" - + case False show ?thesis proof (rule tendstoI) fix e :: real assume "0 < e" @@ -148,10 +145,7 @@ finally (xtrans) have "dist (e m) z \ 2 * dist (f x) z" . } then show "dist (F i x) z \ 2 * dist (f x) z" unfolding F_def - apply auto - apply (rule LeastI2) - apply auto - done + by (smt (verit, ccfv_threshold) LeastI2 B_imp_A dist_eq_0_iff zero_le_dist) qed qed @@ -184,10 +178,7 @@ show "P u" proof (rule seq) show U': "U' i \ borel_measurable M" "\x. 0 \ U' i x" for i - using U by (auto - intro: borel_measurable_simple_function - intro!: borel_measurable_enn2real borel_measurable_times - simp: U'_def zero_le_mult_iff) + using U'_sf by (auto simp: U'_def borel_measurable_simple_function) show "incseq U'" using U(2,3) by (auto simp: incseq_def le_fun_def image_iff eq_commute U'_def indicator_def enn2real_mono) @@ -238,7 +229,7 @@ lemma scaleR_cong_right: fixes x :: "'a :: real_vector" shows "(x \ 0 \ r = p) \ r *\<^sub>R x = p *\<^sub>R x" - by (cases "x = 0") auto + by auto inductive simple_bochner_integrable :: "'a measure \ ('a \ 'b::real_vector) \ bool" for M f where "simple_function M f \ emeasure M {y\space M. f y \ 0} \ \ \ @@ -309,11 +300,7 @@ shows "simple_bochner_integrable M f" proof have "emeasure M {y \ space M. ennreal (norm (f y)) \ 0} \ \" - proof (rule simple_function_finite_support) - show "simple_function M (\x. ennreal (norm (f x)))" - using f by (rule simple_function_compose1) - show "(\\<^sup>+ y. ennreal (norm (f y)) \M) < \" by fact - qed simp + using simple_function_finite_support simple_function_compose1 f fin by force then show "emeasure M {y \ space M. f y \ 0} \ \" by simp qed fact @@ -453,18 +440,17 @@ lemma simple_bochner_integral_nonneg[simp]: fixes f :: "'a \ real" shows "(\x. 0 \ f x) \ 0 \ simple_bochner_integral M f" - by (force simp add: simple_bochner_integral_def intro: sum_nonneg) + by (force simp: simple_bochner_integral_def intro: sum_nonneg) lemma simple_bochner_integral_eq_nn_integral: assumes f: "simple_bochner_integrable M f" "\x. 0 \ f x" shows "simple_bochner_integral M f = (\\<^sup>+x. f x \M)" proof - - { fix x y z have "(x \ 0 \ y = z) \ ennreal x * y = ennreal x * z" - by (cases "x = 0") (auto simp: zero_ennreal_def[symmetric]) } - note ennreal_cong_mult = this + have ennreal_cong_mult: "(x \ 0 \ y = z) \ ennreal x * y = ennreal x * z" for x y z + by fastforce have [measurable]: "f \ borel_measurable M" - using f(1) by (auto intro: borel_measurable_simple_function elim: simple_bochner_integrable.cases) + by (meson borel_measurable_simple_function f(1) simple_bochner_integrable.cases) { fix y assume y: "y \ space M" "f y \ 0" have "ennreal (measure M {x \ space M. f x = f y}) = emeasure M {x \ space M. f x = f y}" @@ -487,8 +473,7 @@ simp flip: sum_ennreal) also have "\ = (\\<^sup>+x. f x \M)" using f - by (intro nn_integral_eq_simple_integral[symmetric]) - (auto simp: simple_function_compose1 simple_bochner_integrable.simps) + by (metis nn_integral_eq_simple_integral simple_bochner_integrable.cases simple_function_compose1) finally show ?thesis . qed @@ -512,9 +497,13 @@ 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 flip: ennreal_plus) - (metis (erased, opaque_lifting) add_diff_cancel_left add_diff_eq diff_add_eq order_trans - norm_minus_commute norm_triangle_ineq4 order_refl) + proof - + have "\x. x \ space M \ + norm (s x - t x) \ norm (f x - s x) + norm (f x - t x)" + by (metis dual_order.refl norm_diff_triangle_le norm_minus_commute) + then show ?thesis + by (metis (mono_tags, lifting) ennreal_leI ennreal_plus nn_integral_mono norm_ge_zero) + qed also have "\ = ?S + ?T" by (rule nn_integral_add) auto finally show ?thesis . @@ -555,7 +544,7 @@ by (rule has_bochner_integral.intros[where s="\_. f"]) (auto intro: borel_measurable_simple_function elim: simple_bochner_integrable.cases - simp: zero_ennreal_def[symmetric]) + simp flip: zero_ennreal_def) lemma has_bochner_integral_real_indicator: assumes [measurable]: "A \ sets M" and A: "emeasure M A < \" @@ -665,21 +654,21 @@ lemma has_bochner_integral_scaleR_left[intro]: "(c \ 0 \ has_bochner_integral M f x) \ has_bochner_integral M (\x. f x *\<^sub>R c) (x *\<^sub>R c)" - by (cases "c = 0") (auto simp add: has_bochner_integral_bounded_linear[OF bounded_linear_scaleR_left]) + by (cases "c = 0") (auto simp: has_bochner_integral_bounded_linear[OF bounded_linear_scaleR_left]) lemma has_bochner_integral_scaleR_right[intro]: "(c \ 0 \ has_bochner_integral M f x) \ has_bochner_integral M (\x. c *\<^sub>R f x) (c *\<^sub>R x)" - by (cases "c = 0") (auto simp add: has_bochner_integral_bounded_linear[OF bounded_linear_scaleR_right]) + by (cases "c = 0") (auto simp: has_bochner_integral_bounded_linear[OF bounded_linear_scaleR_right]) lemma has_bochner_integral_mult_left[intro]: fixes c :: "_::{real_normed_algebra,second_countable_topology}" shows "(c \ 0 \ has_bochner_integral M f x) \ has_bochner_integral M (\x. f x * c) (x * c)" - by (cases "c = 0") (auto simp add: has_bochner_integral_bounded_linear[OF bounded_linear_mult_left]) + by (cases "c = 0") (auto simp: has_bochner_integral_bounded_linear[OF bounded_linear_mult_left]) lemma has_bochner_integral_mult_right[intro]: fixes c :: "_::{real_normed_algebra,second_countable_topology}" shows "(c \ 0 \ has_bochner_integral M f x) \ has_bochner_integral M (\x. c * f x) (c * x)" - by (cases "c = 0") (auto simp add: has_bochner_integral_bounded_linear[OF bounded_linear_mult_right]) + by (cases "c = 0") (auto simp: has_bochner_integral_bounded_linear[OF bounded_linear_mult_right]) lemmas has_bochner_integral_divide = has_bochner_integral_bounded_linear[OF bounded_linear_divide] @@ -691,11 +680,11 @@ lemma has_bochner_integral_inner_left[intro]: "(c \ 0 \ has_bochner_integral M f x) \ has_bochner_integral M (\x. f x \ c) (x \ c)" - by (cases "c = 0") (auto simp add: has_bochner_integral_bounded_linear[OF bounded_linear_inner_left]) + by (cases "c = 0") (auto simp: has_bochner_integral_bounded_linear[OF bounded_linear_inner_left]) lemma has_bochner_integral_inner_right[intro]: "(c \ 0 \ has_bochner_integral M f x) \ has_bochner_integral M (\x. c \ f x) (c \ x)" - by (cases "c = 0") (auto simp add: has_bochner_integral_bounded_linear[OF bounded_linear_inner_right]) + by (cases "c = 0") (auto simp: has_bochner_integral_bounded_linear[OF bounded_linear_inner_right]) lemmas has_bochner_integral_minus = has_bochner_integral_bounded_linear[OF bounded_linear_minus[OF bounded_linear_ident]] @@ -857,16 +846,10 @@ qed (auto intro: g) lemma has_bochner_integral_eq_AE: - assumes f: "has_bochner_integral M f x" - and g: "has_bochner_integral M g y" - and ae: "AE x in M. f x = g x" + assumes "has_bochner_integral M f x" and "has_bochner_integral M g y" + and "AE x in M. f x = g x" shows "x = y" -proof - - from assms have "has_bochner_integral M g x" - by (auto intro: has_bochner_integralI_AE) - from this g show "x = y" - by (rule has_bochner_integral_eq) -qed + by (meson assms has_bochner_integral.simps has_bochner_integral_cong_AE has_bochner_integral_eq) lemma simple_bochner_integrable_restrict_space: fixes f :: "_ \ 'b::real_normed_vector" @@ -1082,12 +1065,7 @@ next assume not: "\ integrable M f" moreover have "\ integrable M (\x. T (f x))" - proof - assume "integrable M (\x. T (f x))" - from integrable_bounded_linear[OF T' this] not *[OF **] - show False - by auto - qed + by (metis (full_types) "*" "**" T' integrable_bounded_linear integrable_cong not) ultimately show ?thesis using T by (simp add: not_integrable_integral_eq linear_simps) qed @@ -1229,7 +1207,7 @@ show "(\i. ennreal (norm (u' x - u i x))) \ 0" by (simp add: tendsto_norm_zero_iff flip: ennreal_0) qed - qed (insert bnd w_nonneg, auto) + qed (use bnd w_nonneg in auto) then show ?thesis by simp qed @@ -1342,14 +1320,8 @@ shows "integrable M f \ g \ borel_measurable M \ (AE x in M. norm (g x) \ norm (f x)) \ integrable M g" unfolding integrable_iff_bounded -proof safe - assume "f \ borel_measurable M" "g \ borel_measurable M" - assume "AE x in M. norm (g x) \ norm (f x)" - then have "(\\<^sup>+ x. ennreal (norm (g x)) \M) \ (\\<^sup>+ x. ennreal (norm (f x)) \M)" - by (intro nn_integral_mono_AE) auto - also assume "(\\<^sup>+ x. ennreal (norm (f x)) \M) < \" - finally show "(\\<^sup>+ x. ennreal (norm (g x)) \M) < \" . -qed + by (smt (verit) AE_cong ennreal_le_iff nn_integral_mono_AE norm_ge_zero order_less_subst2 linorder_not_le order_less_le_trans) + lemma integrable_mult_indicator: fixes f :: "'a \ 'b::{banach, second_countable_topology}" @@ -1419,7 +1391,7 @@ proof cases assume *: "A \ space M \ sets M \ emeasure M (A \ space M) < \" have "integral\<^sup>L M (indicator A) = integral\<^sup>L M (indicator (A \ space M))" - by (intro integral_cong) (auto split: split_indicator) + by (metis (no_types, lifting) Int_iff indicator_simps integral_cong) also have "\ = measure M (A \ space M)" using * by (intro has_bochner_integral_integral_eq has_bochner_integral_real_indicator) auto finally show ?thesis . @@ -1434,6 +1406,26 @@ finally show ?thesis . qed +lemma integrable_discrete_difference_aux: + fixes f :: "'a \ 'b::{banach, second_countable_topology}" + assumes f: "integrable M f" and X: "countable X" + assumes null: "\x. x \ X \ emeasure M {x} = 0" + assumes sets: "\x. x \ X \ {x} \ sets M" + assumes eq: "\x. x \ space M \ x \ X \ f x = g x" + shows "integrable M g" + unfolding integrable_iff_bounded +proof + have fmeas: "f \ borel_measurable M" "(\\<^sup>+ x. ennreal (norm (f x)) \M) < \" + using f integrable_iff_bounded by auto + then show "g \ borel_measurable M" + using measurable_discrete_difference[where X=X] + by (smt (verit) UNIV_I X eq sets space_borel) + have "AE x in M. x \ X" + using AE_discrete_difference X null sets by blast + with fmeas show "(\\<^sup>+ x. ennreal (norm (g x)) \M) < \" + by (metis (mono_tags, lifting) AE_I2 AE_mp eq nn_integral_cong_AE) +qed + lemma integrable_discrete_difference: fixes f :: "'a \ 'b::{banach, second_countable_topology}" assumes X: "countable X" @@ -1441,22 +1433,7 @@ assumes sets: "\x. x \ X \ {x} \ sets M" assumes eq: "\x. x \ space M \ x \ X \ f x = g x" shows "integrable M f \ integrable M g" - unfolding integrable_iff_bounded -proof (rule conj_cong) - { assume "f \ borel_measurable M" then have "g \ borel_measurable M" - by (rule measurable_discrete_difference[where X=X]) (auto simp: assms) } - moreover - { assume "g \ borel_measurable M" then have "f \ borel_measurable M" - by (rule measurable_discrete_difference[where X=X]) (auto simp: assms) } - ultimately show "f \ borel_measurable M \ g \ borel_measurable M" .. -next - have "AE x in M. x \ X" - by (rule AE_discrete_difference) fact+ - then have "(\\<^sup>+ x. norm (f x) \M) = (\\<^sup>+ x. norm (g x) \M)" - by (intro nn_integral_cong_AE) (auto simp: eq) - then show "(\\<^sup>+ x. norm (f x) \M) < \ \ (\\<^sup>+ x. norm (g x) \M) < \" - by simp -qed + by (metis X eq integrable_discrete_difference_aux null sets) lemma integral_discrete_difference: fixes f :: "'a \ 'b::{banach, second_countable_topology}" @@ -1474,7 +1451,6 @@ proof (rule integral_cong_AE) show "f \ borel_measurable M" "g \ borel_measurable M" using f eq by (auto intro: borel_measurable_integrable) - have "AE x in M. x \ X" by (rule AE_discrete_difference) fact+ with AE_space show "AE x in M. f x = g x" @@ -1484,14 +1460,12 @@ lemma has_bochner_integral_discrete_difference: fixes f :: "'a \ 'b::{banach, second_countable_topology}" - assumes X: "countable X" - assumes null: "\x. x \ X \ emeasure M {x} = 0" - assumes sets: "\x. x \ X \ {x} \ sets M" - assumes eq: "\x. x \ space M \ x \ X \ f x = g x" + assumes "countable X" + assumes "\x. x \ X \ emeasure M {x} = 0" + assumes "\x. x \ X \ {x} \ sets M" + assumes "\x. x \ space M \ x \ X \ f x = g x" shows "has_bochner_integral M f x \ has_bochner_integral M g x" - using integrable_discrete_difference[of X M f g, OF assms] - using integral_discrete_difference[of X M f g, OF assms] - by (metis has_bochner_integral_iff) + by (metis assms has_bochner_integral_iff integrable_discrete_difference integral_discrete_difference) lemma fixes f :: "'a \ 'b::{banach, second_countable_topology}" and w :: "'a \ real" @@ -1547,7 +1521,7 @@ finally show "?d n \ (\\<^sup>+x. norm (s n x - f x) \M)" . qed show "(\n. \\<^sup>+x. norm (s n x - f x) \M) \ ennreal 0" - unfolding ennreal_0 + unfolding ennreal_0 apply (subst norm_minus_commute) proof (rule nn_integral_dominated_convergence_norm[where w=w]) show "\n. s n \ borel_measurable M" @@ -1687,7 +1661,7 @@ by (simp add: integrable_indicator_iff ennreal_indicator emeasure_eq_ennreal_measure) next case (mult f c) then show ?case - by (auto simp add: nn_integral_cmult ennreal_mult integral_nonneg_AE) + by (auto simp: nn_integral_cmult ennreal_mult integral_nonneg_AE) next case (add g f) then have "integrable M f" "integrable M g" @@ -1723,14 +1697,9 @@ qed lemma nn_integral_eq_integrable: - assumes f: "f \ M \\<^sub>M borel" "AE x in M. 0 \ f x" and "0 \ x" + assumes "f \ M \\<^sub>M borel" "AE x in M. 0 \ f x" and "0 \ x" shows "(\\<^sup>+x. f x \M) = ennreal x \ (integrable M f \ integral\<^sup>L M f = x)" -proof (safe intro!: nn_integral_eq_integral assms) - assume *: "(\\<^sup>+x. f x \M) = ennreal x" - with integrableI_nn_integral_finite[OF f this] nn_integral_eq_integral[of M f, OF _ f(2)] - show "integrable M f" "integral\<^sup>L M f = x" - by (simp_all add: * assms integral_nonneg_AE) -qed + by (metis assms ennreal_inj integrableI_nn_integral_finite integral_nonneg_AE nn_integral_eq_integral) lemma fixes f :: "_ \ _ \ 'a :: {banach, second_countable_topology}" @@ -1744,12 +1713,11 @@ proof - have 1: "integrable M (\x. \i. norm (f i x))" proof (rule integrableI_bounded) - have "(\\<^sup>+ x. ennreal (norm (\i. norm (f i x))) \M) = (\\<^sup>+ x. (\i. ennreal (norm (f i x))) \M)" - apply (intro nn_integral_cong_AE) - using summable - apply eventually_elim - apply (simp add: suminf_nonneg ennreal_suminf_neq_top) - done + have "\x. summable (\i. norm (f i x)) \ + ennreal (norm (\i. norm (f i x))) = (\i. ennreal (norm (f i x)))" + by (simp add: suminf_nonneg ennreal_suminf_neq_top) + then have "(\\<^sup>+ x. ennreal (norm (\i. norm (f i x))) \M) = (\\<^sup>+ x. (\i. ennreal (norm (f i x))) \M)" + using local.summable by (intro nn_integral_cong_AE) auto also have "\ = (\i. \\<^sup>+ x. norm (f i x) \M)" by (intro nn_integral_suminf) auto also have "\ = (\i. ennreal (\x. norm (f i x) \M))" @@ -1802,63 +1770,28 @@ using integral_norm_bound[of M f] by auto lemma integral_eq_nn_integral: - assumes [measurable]: "f \ borel_measurable M" - assumes nonneg: "AE x in M. 0 \ f x" + assumes "f \ borel_measurable M" + assumes "AE x in M. 0 \ f x" shows "integral\<^sup>L M f = enn2real (\\<^sup>+ x. ennreal (f x) \M)" -proof cases - assume *: "(\\<^sup>+ x. ennreal (f x) \M) = \" - also have "(\\<^sup>+ x. ennreal (f x) \M) = (\\<^sup>+ x. ennreal (norm (f x)) \M)" - using nonneg by (intro nn_integral_cong_AE) auto - finally have "\ integrable M f" - by (auto simp: integrable_iff_bounded) - then show ?thesis - by (simp add: * not_integrable_integral_eq) -next - assume "(\\<^sup>+ x. ennreal (f x) \M) \ \" - then have "integrable M f" - by (cases "\\<^sup>+ x. ennreal (f x) \M" rule: ennreal_cases) - (auto intro!: integrableI_nn_integral_finite assms) - from nn_integral_eq_integral[OF this] nonneg show ?thesis - by (simp add: integral_nonneg_AE) -qed + by (metis assms enn2real_ennreal enn2real_top infinity_ennreal_def integrableI_nonneg integral_nonneg_AE + less_top nn_integral_eq_integral not_integrable_integral_eq) lemma enn2real_nn_integral_eq_integral: - assumes eq: "AE x in M. f x = ennreal (g x)" and nn: "AE x in M. 0 \ g x" - and fin: "(\\<^sup>+x. f x \M) < top" - and [measurable]: "g \ M \\<^sub>M borel" + assumes "AE x in M. f x = ennreal (g x)" and nn: "AE x in M. 0 \ g x" + and "g \ M \\<^sub>M borel" shows "enn2real (\\<^sup>+x. f x \M) = (\x. g x \M)" -proof - - have "ennreal (enn2real (\\<^sup>+x. f x \M)) = (\\<^sup>+x. f x \M)" - using fin by (intro ennreal_enn2real) auto - also have "\ = (\\<^sup>+x. g x \M)" - using eq by (rule nn_integral_cong_AE) - also have "\ = (\x. g x \M)" - proof (rule nn_integral_eq_integral) - show "integrable M g" - proof (rule integrableI_bounded) - have "(\\<^sup>+ x. ennreal (norm (g x)) \M) = (\\<^sup>+ x. f x \M)" - using eq nn by (auto intro!: nn_integral_cong_AE elim!: eventually_elim2) - also note fin - finally show "(\\<^sup>+ x. ennreal (norm (g x)) \M) < \" - by simp - qed simp - qed fact - finally show ?thesis - using nn by (simp add: integral_nonneg_AE) -qed + by (metis assms integral_eq_nn_integral nn nn_integral_cong_AE) lemma has_bochner_integral_nn_integral: assumes "f \ borel_measurable M" "AE x in M. 0 \ f x" "0 \ x" assumes "(\\<^sup>+x. f x \M) = ennreal x" shows "has_bochner_integral M f x" - unfolding has_bochner_integral_iff - using assms by (auto simp: assms integral_eq_nn_integral intro: integrableI_nn_integral_finite) + by (metis assms has_bochner_integral_iff nn_integral_eq_integrable) lemma integrableI_simple_bochner_integrable: fixes f :: "'a \ 'b::{banach, second_countable_topology}" shows "simple_bochner_integrable M f \ integrable M f" - by (intro integrableI_sequence[where s="\_. f"] borel_measurable_simple_function) - (auto simp: zero_ennreal_def[symmetric] simple_bochner_integrable.simps) + using has_bochner_integral_simple_bochner_integrable integrable.intros by blast proposition integrable_induct[consumes 1, case_names base add lim, induct pred: integrable]: fixes f :: "'a \ 'b::{banach, second_countable_topology}" @@ -1896,7 +1829,7 @@ { fix i have "\z. {y. s' i z = y \ y \ s' i ` space M \ y \ 0 \ z \ space M} = (if z \ space M \ s' i z \ 0 then {s' i z} else {})" - by (auto simp add: s'_def split: split_indicator) + by (auto simp: s'_def split: split_indicator) then have "\z. s' i = (\z. \y\s' i`space M - {0}. indicator {x\space M. s' i x = y} z *\<^sub>R y)" using sf by (auto simp: fun_eq_iff simple_function_def s'_def) } note s'_eq = this @@ -1947,7 +1880,7 @@ by (simp add: nn_integral_0_iff_AE) with nonneg show "AE x in M. f x = 0" by auto -qed (auto simp add: integral_eq_zero_AE) +qed (auto simp: integral_eq_zero_AE) lemma integral_mono_AE: fixes f :: "'a \ real" @@ -1971,13 +1904,7 @@ fixes f :: "'a::euclidean_space \ 'b::{banach,second_countable_topology}" assumes "integrable M f" "integrable M g" "\x. x \ space M \ norm(f x) \ g x" shows "norm (\x. f x \M) \ (\x. g x \M)" -proof - - have "norm (\x. f x \M) \ (\x. norm (f x) \M)" - by (rule integral_norm_bound) - also have "... \ (\x. g x \M)" - using assms integrable_norm integral_mono by blast - finally show ?thesis . -qed + by (smt (verit, best) assms integrable_norm integral_mono integral_norm_bound) lemma integral_abs_bound_integral: fixes f :: "'a::euclidean_space \ real" @@ -1993,15 +1920,7 @@ fixes f::"_ \ real" assumes "integrable M f" "AE x in M. g x \ f x" "AE x in M. 0 \ f x" shows "(\x. g x \M) \ (\x. f x \M)" -proof (cases "integrable M g") - case True - show ?thesis by (rule integral_mono_AE, auto simp add: assms True) -next - case False - then have "(\x. g x \M) = 0" by (simp add: not_integrable_integral_eq) - also have "... \ (\x. f x \M)" by (simp add: integral_nonneg_AE[OF assms(3)]) - finally show ?thesis by simp -qed + by (metis assms integral_mono_AE integral_nonneg_AE not_integrable_integral_eq) lemma integral_mono': fixes f::"_ \ real" @@ -2022,21 +1941,11 @@ by (auto intro!: integrableI_bounded) qed -lemma integrableI_real_bounded: - assumes f: "f \ borel_measurable M" and ae: "AE x in M. 0 \ f x" and fin: "integral\<^sup>N M f < \" - shows "integrable M f" -proof (rule integrableI_bounded) - have "(\\<^sup>+ x. ennreal (norm (f x)) \M) = \\<^sup>+ x. ennreal (f x) \M" - using ae by (auto intro: nn_integral_cong_AE) - also note fin - finally show "(\\<^sup>+ x. ennreal (norm (f x)) \M) < \" . -qed fact - lemma nn_integral_nonneg_infinite: fixes f::"'a \ real" assumes "f \ borel_measurable M" "\ integrable M f" "AE x in M. f x \ 0" shows "(\\<^sup>+x. f x \M) = \" -using assms integrableI_real_bounded less_top by auto + using assms integrableI_nonneg less_top by auto lemma integral_real_bounded: assumes "0 \ r" "integral\<^sup>N M f \ ennreal r" @@ -2066,22 +1975,22 @@ fixes f:: "_ \ _ \ real" shows "\ finite I; I \ {}; \i. i \ I \ integrable M (f i) \ \ integrable M (\x. MIN i\I. f i x)" -by (induct rule: finite_ne_induct) simp+ + by (induct rule: finite_ne_induct) simp+ lemma integrable_MAX: fixes f:: "_ \ _ \ real" shows "\ finite I; I \ {}; \i. i \ I \ integrable M (f i) \ \ integrable M (\x. MAX i\I. f i x)" -by (induct rule: finite_ne_induct) simp+ + by (induct rule: finite_ne_induct) simp+ theorem integral_Markov_inequality: assumes [measurable]: "integrable M u" and "AE x in M. 0 \ u x" "0 < (c::real)" shows "(emeasure M) {x\space M. u x \ c} \ (1/c) * (\x. u x \M)" proof - have "(\\<^sup>+ x. ennreal(u x) * indicator (space M) x \M) \ (\\<^sup>+ x. u x \M)" - by (rule nn_integral_mono_AE, auto simp add: \c>0\ less_eq_real_def) - also have "... = (\x. u x \M)" - by (rule nn_integral_eq_integral, auto simp add: assms) + by (rule nn_integral_mono_AE, auto simp: \c>0\ less_eq_real_def) + also have "\ = (\x. u x \M)" + by (rule nn_integral_eq_integral, auto simp: assms) finally have *: "(\\<^sup>+ x. ennreal(u x) * indicator (space M) x \M) \ (\x. u x \M)" by simp @@ -2089,9 +1998,9 @@ using \c>0\ by (auto simp: ennreal_mult'[symmetric]) then have "emeasure M {x \ space M. u x \ c} = emeasure M ({x \ space M. ennreal(1/c) * u x \ 1})" by simp - also have "... \ ennreal(1/c) * (\\<^sup>+ x. ennreal(u x) * indicator (space M) x \M)" - by (rule nn_integral_Markov_inequality) (auto simp add: assms) - also have "... \ ennreal(1/c) * (\x. u x \M)" + also have "\ \ ennreal(1/c) * (\\<^sup>+ x. ennreal(u x) * indicator (space M) x \M)" + by (rule nn_integral_Markov_inequality) (auto simp: assms) + also have "\ \ ennreal(1/c) * (\x. u x \M)" by (rule mult_left_mono) (use * \c > 0\ in auto) finally show ?thesis using \0 by (simp add: ennreal_mult'[symmetric]) @@ -2109,8 +2018,7 @@ by (intro emeasure_eq_ennreal_measure [symmetric]) auto also note le finally show ?thesis - by (subst (asm) ennreal_le_iff) - (auto intro!: divide_nonneg_pos Bochner_Integration.integral_nonneg_AE assms) + by (simp add: assms divide_nonneg_pos integral_nonneg_AE) qed theorem%important (in finite_measure) second_moment_method: @@ -2149,9 +2057,10 @@ assumes "AE x in M. f x \ 0" "(\x. f x \M) > 0" and [measurable]: "f \ borel_measurable M" shows "\A e. A \ sets M \ e>0 \ emeasure M A > 0 \ (\x \ A. f x \ e)" -proof (rule not_AE_zero_E, auto simp add: assms) +proof (rule not_AE_zero_E, auto simp: assms) assume *: "AE x in M. f x = 0" - have "(\x. f x \M) = (\x. 0 \M)" by (rule integral_cong_AE, auto simp add: *) + have "(\x. f x \M) = (\x. 0 \M)" + by (rule integral_cong_AE, auto simp: *) then have "(\x. f x \M) = 0" by simp then show False using assms(2) by simp qed @@ -2163,28 +2072,31 @@ shows "((\n. (\x. u n x \M)) \ (\x. f x \M)) F" proof - have "((\n. norm((\x. u n x \M) - (\x. f x \M))) \ (0::ennreal)) F" - proof (rule tendsto_sandwich[of "\_. 0", where ?h = "\n. (\\<^sup>+x. norm(u n x - f x) \M)"], auto simp add: assms) + proof (rule tendsto_sandwich[of "\_. 0", where ?h = "\n. (\\<^sup>+x. norm(u n x - f x) \M)"], auto simp: assms) { fix n have "(\x. u n x \M) - (\x. f x \M) = (\x. u n x - f x \M)" - apply (rule Bochner_Integration.integral_diff[symmetric]) using assms by auto + by (simp add: assms) then have "norm((\x. u n x \M) - (\x. f x \M)) = norm (\x. u n x - f x \M)" by auto - also have "... \ (\x. norm(u n x - f x) \M)" + also have "\ \ (\x. norm(u n x - f x) \M)" by (rule integral_norm_bound) finally have "ennreal(norm((\x. u n x \M) - (\x. f x \M))) \ (\x. norm(u n x - f x) \M)" by simp - also have "... = (\\<^sup>+x. norm(u n x - f x) \M)" - apply (rule nn_integral_eq_integral[symmetric]) using assms by auto - finally have "norm((\x. u n x \M) - (\x. f x \M)) \ (\\<^sup>+x. norm(u n x - f x) \M)" by simp + also have "\ = (\\<^sup>+x. norm(u n x - f x) \M)" + by (simp add: assms nn_integral_eq_integral) + finally have "norm((\x. u n x \M) - (\x. f x \M)) \ (\\<^sup>+x. norm(u n x - f x) \M)" + by simp } then show "eventually (\n. norm((\x. u n x \M) - (\x. f x \M)) \ (\\<^sup>+x. norm(u n x - f x) \M)) F" by auto qed then have "((\n. norm((\x. u n x \M) - (\x. f x \M))) \ 0) F" by (simp flip: 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 + 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 text \The next lemma asserts that, if a sequence of functions converges in \L\<^sup>1\, then @@ -2219,7 +2131,7 @@ then have "(\x. norm(u (r n) x) \M) < (1/2)^n" unfolding r_def by auto moreover have "(\\<^sup>+x. norm(u (r n) x) \M) = (\x. norm(u (r n) x) \M)" - by (rule nn_integral_eq_integral, auto simp add: integrable_norm[OF assms(1)[of "r n"]]) + by (rule nn_integral_eq_integral, auto simp: integrable_norm[OF assms(1)[of "r n"]]) ultimately show ?thesis by (auto intro: ennreal_lessI) qed @@ -2231,31 +2143,32 @@ have A_bound: "emeasure M (A n) < (1/e) * ennreal((1/2)^n)" for n proof - have *: "indicator (A n) x \ (1/e) * ennreal(norm(u (r n) x))" for x - apply (cases "x \ A n") unfolding A_def using \0 < e\ by (auto simp: ennreal_mult[symmetric]) + using \0 < e\ by (cases "x \ A n") (auto simp: A_def ennreal_mult[symmetric]) have "emeasure M (A n) = (\\<^sup>+x. indicator (A n) x \M)" by auto - also have "... \ (\\<^sup>+x. (1/e) * ennreal(norm(u (r n) x)) \M)" - apply (rule nn_integral_mono) using * by auto - also have "... = (1/e) * (\\<^sup>+x. norm(u (r n) x) \M)" - apply (rule nn_integral_cmult) using \e > 0\ by auto - also have "... < (1/e) * ennreal((1/2)^n)" + also have "\ \ (\\<^sup>+x. (1/e) * ennreal(norm(u (r n) x)) \M)" + by (meson "*" nn_integral_mono) + also have "\ = (1/e) * (\\<^sup>+x. norm(u (r n) x) \M)" + using \e > 0\ by (force simp add: intro: nn_integral_cmult) + also have "\ < (1/e) * ennreal((1/2)^n)" using I[of n] \e > 0\ by (intro ennreal_mult_strict_left_mono) auto finally show ?thesis by simp qed have A_fin: "emeasure M (A n) < \" for n using \e > 0\ A_bound[of n] - by (auto simp add: ennreal_mult less_top[symmetric]) + by (auto simp: ennreal_mult less_top[symmetric]) have A_sum: "summable (\n. measure M (A n))" - proof (rule summable_comparison_test'[of "\n. (1/e) * (1/2)^n" 0]) - have "summable (\n. (1/(2::real))^n)" by (simp add: summable_geometric) + proof (rule summable_comparison_test') + have "summable (\n. (1/(2::real))^n)" + by (simp add: summable_geometric) then show "summable (\n. (1/e) * (1/2)^n)" using summable_mult by blast fix n::nat assume "n \ 0" have "norm(measure M (A n)) = measure M (A n)" by simp - also have "... = enn2real(emeasure M (A n))" unfolding measure_def by simp - also have "... < enn2real((1/e) * (1/2)^n)" + also have "\ = enn2real(emeasure M (A n))" unfolding measure_def by simp + also have "\ < enn2real((1/e) * (1/2)^n)" using A_bound[of n] \emeasure M (A n) < \\ \0 < e\ by (auto simp: emeasure_eq_ennreal_measure ennreal_mult[symmetric] ennreal_less_iff) - also have "... = (1/e) * (1/2)^n" + also have "\ = (1/e) * (1/2)^n" using \0 by auto finally show "norm(measure M (A n)) \ (1/e) * (1/2)^n" by simp qed @@ -2272,13 +2185,14 @@ then have "limsup (\n. ennreal (norm(u (r n) x))) \ e" by (simp add: Limsup_bounded) } - ultimately show "AE x in M. limsup (\n. ennreal (norm(u (r n) x))) \ 0 + ennreal e" by auto + ultimately show "AE x in M. limsup (\n. ennreal (norm(u (r n) x))) \ 0 + ennreal e" + by auto qed moreover { - fix x assume "limsup (\n. ennreal (norm(u (r n) x))) \ 0" - moreover then have "liminf (\n. ennreal (norm(u (r n) x))) \ 0" - by (rule order_trans[rotated]) (auto intro: Liminf_le_Limsup) + fix x assume x: "limsup (\n. ennreal (norm(u (r n) x))) \ 0" + moreover have "liminf (\n. ennreal (norm(u (r n) x))) \ 0" + by (metis x Liminf_le_Limsup le_zero_eq sequentially_bot) 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" @@ -2326,7 +2240,7 @@ unfolding lim using lim by (intro integral_dominated_convergence[where w="\x. 2 * norm (indicator \ x *\<^sub>R f x)"]) - (auto simp add: space_restrict_space integrable_restrict_space simp del: norm_scaleR + (auto simp: space_restrict_space integrable_restrict_space simp del: norm_scaleR split: split_indicator) qed qed @@ -2335,11 +2249,7 @@ lemma integral_empty: assumes "space M = {}" shows "integral\<^sup>L M f = 0" -proof - - have "(\ x. f x \M) = (\ x. 0 \M)" - by(rule integral_cong)(simp_all add: assms) - thus ?thesis by simp -qed + by (metis AE_I2 assms empty_iff integral_eq_zero_AE) subsection \Measure spaces with an associated density\ @@ -2379,7 +2289,7 @@ also have "\ = integral\<^sup>L (density M g) (indicator A)" using base by simp finally show ?case - using base g + using base g apply (simp add: int integral_nonneg_AE) apply (subst (asm) ennreal_inj) apply (auto intro!: integral_nonneg_AE) @@ -2409,7 +2319,7 @@ show "(\i. integral\<^sup>L M (\x. g x *\<^sub>R s i x)) \ integral\<^sup>L (density M g) f" unfolding lim(2)[symmetric] by (rule integral_dominated_convergence[where w="\x. 2 * norm (f x)"]) - (insert lim(3-5), auto) + (use lim in auto) qed qed qed (simp add: f g integrable_density) @@ -2438,8 +2348,7 @@ lemma integrable_distr: fixes f :: "'a \ 'b::{banach, second_countable_topology}" shows "T \ measurable M M' \ integrable (distr M M' T) f \ integrable M (\x. f (T x))" - by (subst integrable_distr_eq[symmetric, where g=T]) - (auto dest: borel_measurable_integrable) + by (metis integrable_distr_eq integrable_iff_bounded measurable_distr_eq1) lemma integral_distr: fixes f :: "'a \ 'b::{banach, second_countable_topology}" @@ -2488,7 +2397,7 @@ show "(\i. integral\<^sup>L M (\x. s i (g x))) \ integral\<^sup>L (distr M N g) f" unfolding lim(2)[symmetric] by (rule integral_dominated_convergence[where w="\x. 2 * norm (f x)"]) - (insert lim(3-5), auto) + (use lim in auto) qed qed qed (simp add: f g integrable_distr_eq) @@ -2532,8 +2441,8 @@ lemma integrable_count_space_nat_iff: fixes f :: "nat \ _::{banach,second_countable_topology}" shows "integrable (count_space UNIV) f \ summable (\x. norm (f x))" - by (auto simp add: integrable_iff_bounded nn_integral_count_space_nat ennreal_suminf_neq_top - intro: summable_suminf_not_top) + by (auto simp: integrable_iff_bounded nn_integral_count_space_nat ennreal_suminf_neq_top + intro: summable_suminf_not_top) lemma sums_integral_count_space_nat: fixes f :: "nat \ _::{banach,second_countable_topology}" @@ -2563,7 +2472,7 @@ lemma integral_count_space_nat: fixes f :: "nat \ _::{banach,second_countable_topology}" shows "integrable (count_space UNIV) f \ integral\<^sup>L (count_space UNIV) f = (\x. f x)" - using sums_integral_count_space_nat by (rule sums_unique) + using sums_integral_count_space_nat sums_unique by auto lemma integrable_bij_count_space: fixes f :: "'a \ 'b::{banach, second_countable_topology}" @@ -2575,11 +2484,8 @@ fixes f :: "'a \ 'b::{banach, second_countable_topology}" assumes g: "bij_betw g A B" shows "integral\<^sup>L (count_space A) (\x. f (g x)) = integral\<^sup>L (count_space B) f" - using g[THEN bij_betw_imp_funcset] - apply (subst distr_bij_count_space[OF g, symmetric]) - apply (intro integral_distr[symmetric]) - apply auto - done + using g[THEN bij_betw_imp_funcset] distr_bij_count_space [OF g] + by (metis borel_measurable_count_space integral_distr measurable_count_space_eq1 space_count_space) lemma has_bochner_integral_count_space_nat: fixes f :: "nat \ _::{banach,second_countable_topology}" @@ -2596,29 +2502,31 @@ proposition integrable_point_measure_finite: fixes g :: "'a \ 'b::{banach, second_countable_topology}" and f :: "'a \ real" - shows "finite A \ integrable (point_measure A f) g" - unfolding point_measure_def - apply (subst density_cong[where f'="\x. ennreal (max 0 (f x))"]) - apply (auto split: split_max simp: ennreal_neg) - apply (subst integrable_density) - apply (auto simp: AE_count_space integrable_count_space) - done + assumes "finite A" + shows "integrable (point_measure A f) g" +proof - + have "integrable (density (count_space A) (\x. ennreal (max 0 (f x)))) g" + using assms + by (simp add: integrable_count_space integrable_density ) + then show ?thesis + by (smt (verit) AE_I2 borel_measurable_count_space density_cong ennreal_neg point_measure_def) +qed subsection \Lebesgue integration on \<^const>\null_measure\\ lemma has_bochner_integral_null_measure_iff[iff]: "has_bochner_integral (null_measure M) f 0 \ f \ borel_measurable M" - by (auto simp add: has_bochner_integral.simps simple_bochner_integral_def[abs_def] + by (auto simp: has_bochner_integral.simps simple_bochner_integral_def[abs_def] intro!: exI[of _ "\n x. 0"] simple_bochner_integrable.intros) lemma integrable_null_measure_iff[iff]: "integrable (null_measure M) f \ f \ borel_measurable M" - by (auto simp add: integrable.simps) + by (auto simp: integrable.simps) lemma integral_null_measure[simp]: "integral\<^sup>L (null_measure M) f = 0" - by (cases "integrable (null_measure M) f") - (auto simp add: not_integrable_integral_eq has_bochner_integral_integral_eq) + using integral_norm_bound_ennreal not_integrable_integral_eq by fastforce subsection \Legacy lemmas for the real-valued Lebesgue integral\ + theorem real_lebesgue_integral_def: assumes f[measurable]: "integrable M f" shows "integral\<^sup>L M f = enn2real (\\<^sup>+x. f x \M) - enn2real (\\<^sup>+x. ennreal (- f x) \M)" @@ -2701,18 +2609,18 @@ using i by auto next show "(\\<^sup>+ x. ennreal (u x) \M) = \\<^sup>+ x. (SUP i. ennreal (f i x)) \M" - apply (rule nn_integral_cong_AE) - using lim mono nn u_nn - apply eventually_elim - apply (simp add: LIMSEQ_unique[OF _ LIMSEQ_SUP] incseq_def) - done + proof (rule nn_integral_cong_AE) + show "AE x in M. ennreal (u x) = (SUP i. ennreal (f i x))" + using lim mono nn u_nn + by eventually_elim (simp add: LIMSEQ_unique[OF _ LIMSEQ_SUP] incseq_def) + qed qed also have "\ = ennreal x" using mono i nn unfolding nn_integral_eq_integral[OF i pos] by (subst LIMSEQ_unique[OF LIMSEQ_SUP]) (auto simp: mono_def integral_nonneg_AE pos intro!: integral_mono_AE ilim) finally have "(\\<^sup>+ x. ennreal (u x) \M) = ennreal x" . moreover have "(\\<^sup>+ x. ennreal (- u x) \M) = 0" - using u u_nn by (subst nn_integral_0_iff_AE) (auto simp add: ennreal_neg) + using u u_nn by (subst nn_integral_0_iff_AE) (auto simp: ennreal_neg) ultimately show "integrable M u" "integral\<^sup>L M u = x" by (auto simp: real_integrable_def real_lebesgue_integral_def u) qed @@ -2777,7 +2685,7 @@ proof - { assume "emeasure M (space M) = \" "a \ 0" then have ?thesis - by (auto simp add: not_integrable_integral_eq ennreal_mult_less_top measure_def integrable_iff_bounded) } + by (auto simp: not_integrable_integral_eq ennreal_mult_less_top measure_def integrable_iff_bounded) } moreover { assume "a = 0" then have ?thesis by simp } moreover @@ -2795,17 +2703,14 @@ lemma (in finite_measure) integrable_const_bound: fixes f :: "'a \ 'b::{banach,second_countable_topology}" shows "AE x in M. norm (f x) \ B \ f \ borel_measurable M \ integrable M f" - apply (rule integrable_bound[OF integrable_const[of B], of f]) - apply assumption - apply (cases "0 \ B") - apply auto - done + using integrable_bound[OF integrable_const[of B], of f] + by (smt (verit, ccfv_SIG) eventually_elim2 real_norm_def) lemma (in finite_measure) integral_bounded_eq_bound_then_AE: assumes "AE x in M. f x \ (c::real)" "integrable M f" "(\x. f x \M) = c * measure M (space M)" shows "AE x in M. f x = c" - apply (rule integral_ineq_eq_0_then_AE) using assms by auto + using assms by (intro integral_ineq_eq_0_then_AE) auto lemma integral_indicator_finite_real: fixes f :: "'a \ real" @@ -3030,13 +2935,7 @@ fixes f :: "_ \ _::{banach, second_countable_topology}" assumes "integrable (M1 \\<^sub>M M2) f" shows "integrable (M2 \\<^sub>M M1) (\(x,y). f (y,x))" -proof - - interpret Q: pair_sigma_finite M2 M1 .. - have *: "(\(x,y). f (y,x)) = (\x. f (case x of (x,y)\(y,x)))" by (auto simp: fun_eq_iff) - show ?thesis unfolding * - by (rule integrable_distr[OF measurable_pair_swap']) - (simp add: distr_pair_swap[symmetric] assms) -qed + by (smt (verit) assms distr_pair_swap integrable_cong integrable_distr measurable_pair_swap' prod.case_distrib split_cong) lemma (in pair_sigma_finite) integrable_product_swap_iff: fixes f :: "_ \ _::{banach, second_countable_topology}" @@ -3051,11 +2950,7 @@ fixes f :: "_ \ _::{banach, second_countable_topology}" assumes f: "f \ borel_measurable (M1 \\<^sub>M M2)" shows "(\(x,y). f (y,x) \(M2 \\<^sub>M M1)) = integral\<^sup>L (M1 \\<^sub>M M2) f" -proof - - have *: "(\(x,y). f (y,x)) = (\x. f (case x of (x,y)\(y,x)))" by (auto simp: fun_eq_iff) - show ?thesis unfolding * - by (simp add: integral_distr[symmetric, OF measurable_pair_swap' f] distr_pair_swap[symmetric]) -qed + by (smt (verit) distr_pair_swap f integral_cong integral_distr measurable_pair_swap' prod.case_distrib split_cong) theorem (in pair_sigma_finite) Fubini_integrable: fixes f :: "_ \ _::{banach, second_countable_topology}" @@ -3180,14 +3075,13 @@ case (add f g) then have [measurable]: "f \ borel_measurable (M1 \\<^sub>M M2)" "g \ borel_measurable (M1 \\<^sub>M M2)" by auto - have "(\ x. \ y. f (x, y) + g (x, y) \M2 \M1) = + have "AE x in M1. LINT y|M2. f (x, y) + g (x, y) = + (LINT y|M2. f (x, y)) + (LINT y|M2. g (x, y))" + using AE_integrable_fst'[OF add(1)] AE_integrable_fst'[OF add(3)] + by eventually_elim simp + then have "(\ x. \ y. f (x, y) + g (x, y) \M2 \M1) = (\ x. (\ y. f (x, y) \M2) + (\ y. g (x, y) \M2) \M1)" - apply (rule integral_cong_AE) - apply simp_all - using AE_integrable_fst'[OF add(1)] AE_integrable_fst'[OF add(3)] - apply eventually_elim - apply simp - done + by (intro integral_cong_AE) auto also have "\ = (\ x. f x \(M1 \\<^sub>M M2)) + (\ x. g x \(M1 \\<^sub>M M2))" using integrable_fst'[OF add(1)] integrable_fst'[OF add(3)] add(2,4) by simp finally show ?case @@ -3278,10 +3172,7 @@ lemma (in product_sigma_finite) product_integral_singleton: fixes f :: "_ \ _::{banach, second_countable_topology}" shows "f \ borel_measurable (M i) \ (\x. f (x i) \Pi\<^sub>M {i} M) = integral\<^sup>L (M i) f" - apply (subst distr_singleton[symmetric]) - apply (subst integral_distr) - apply simp_all - done + by (metis (no_types) distr_singleton insert_iff integral_distr measurable_component_singleton) lemma (in product_sigma_finite) product_integral_fold: fixes f :: "_ \ _::{banach, second_countable_topology}" @@ -3302,12 +3193,12 @@ using measurable_comp[OF measurable_merge f_borel] by (simp add: comp_def) have f_int: "integrable (Pi\<^sub>M I M \\<^sub>M Pi\<^sub>M J M) ?f" by (rule integrable_distr[OF measurable_merge]) (simp add: distr_merge[OF IJ fin] f) - show ?thesis + have "LINT x|(Pi\<^sub>M I M \\<^sub>M Pi\<^sub>M J M). f (merge I J x) = + LINT x|Pi\<^sub>M I M. LINT y|Pi\<^sub>M J M. f (merge I J (x, y))" + by (simp add: P.integral_fst'[symmetric, OF f_int]) + then show ?thesis apply (subst distr_merge[symmetric, OF IJ fin]) - apply (subst integral_distr[OF measurable_merge f_borel]) - apply (subst P.integral_fst'[symmetric, OF f_int]) - apply simp - done + by (simp add: integral_distr[OF measurable_merge f_borel]) qed lemma (in product_sigma_finite) product_integral_insert: @@ -3413,15 +3304,11 @@ proof (intro LIMSEQ_unique) show "(\i. integral\<^sup>L N (s i)) \ integral\<^sup>L N f" apply (rule integral_dominated_convergence[where w="\x. 2 * norm (f x)"]) - using lim - apply auto - done + using lim by auto show "(\i. integral\<^sup>L N (s i)) \ integral\<^sup>L M f" unfolding lim apply (rule integral_dominated_convergence[where w="\x. 2 * norm (f x)"]) - using lim M N(2) - apply auto - done + using lim M N by auto qed qed qed (simp add: not_integrable_integral_eq integrable_subalgebra[OF assms]) diff -r ba2afdd29e1d -r 032a4344903e src/HOL/Analysis/Brouwer_Fixpoint.thy --- a/src/HOL/Analysis/Brouwer_Fixpoint.thy Tue Aug 01 11:27:55 2023 +0200 +++ b/src/HOL/Analysis/Brouwer_Fixpoint.thy Thu Aug 03 19:10:43 2023 +0200 @@ -391,22 +391,22 @@ qed fact locale kuhn_simplex = - fixes p n and base upd and s :: "(nat \ nat) set" + fixes p n and base upd and S :: "(nat \ nat) set" assumes base: "base \ {..< n} \ {..< p}" assumes base_out: "\i. n \ i \ base i = p" assumes upd: "bij_betw upd {..< n} {..< n}" - assumes s_pre: "s = (\i j. if j \ upd`{..< i} then Suc (base j) else base j) ` {.. n}" + assumes s_pre: "S = (\i j. if j \ upd`{..< i} then Suc (base j) else base j) ` {.. n}" begin definition "enum i j = (if j \ upd`{..< i} then Suc (base j) else base j)" -lemma s_eq: "s = enum ` {.. n}" +lemma s_eq: "S = enum ` {.. n}" unfolding s_pre enum_def[abs_def] .. lemma upd_space: "i < n \ upd i < n" using upd by (auto dest!: bij_betwE) -lemma s_space: "s \ {..< n} \ {.. p}" +lemma s_space: "S \ {..< n} \ {.. p}" proof - { fix i assume "i \ n" then have "enum i \ {..< n} \ {.. p}" proof (induct i) @@ -437,19 +437,19 @@ lemma enum_0: "enum 0 = base" by (simp add: enum_def[abs_def]) -lemma base_in_s: "base \ s" +lemma base_in_s: "base \ S" unfolding s_eq by (subst enum_0[symmetric]) auto -lemma enum_in: "i \ n \ enum i \ s" +lemma enum_in: "i \ n \ enum i \ S" unfolding s_eq by auto lemma one_step: - assumes a: "a \ s" "j < n" - assumes *: "\a'. a' \ s \ a' \ a \ a' j = p'" + assumes a: "a \ S" "j < n" + assumes *: "\a'. a' \ S \ a' \ a \ a' j = p'" shows "a j \ p'" proof assume "a j = p'" - with * a have "\a'. a' \ s \ a' j = p'" + with * a have "\a'. a' \ S \ a' j = p'" by auto then have "\i. i \ n \ enum i j = p'" unfolding s_eq by auto @@ -481,16 +481,16 @@ lemma enum_strict_mono: "i \ n \ j \ n \ enum i < enum j \ i < j" using enum_mono[of i j] enum_inj[of i j] by (auto simp: le_less) -lemma chain: "a \ s \ b \ s \ a \ b \ b \ a" +lemma chain: "a \ S \ b \ S \ a \ b \ b \ a" by (auto simp: s_eq enum_mono) -lemma less: "a \ s \ b \ s \ a i < b i \ a < b" +lemma less: "a \ S \ b \ S \ a i < b i \ a < b" using chain[of a b] by (auto simp: less_fun_def le_fun_def not_le[symmetric]) -lemma enum_0_bot: "a \ s \ a = enum 0 \ (\a'\s. a \ a')" +lemma enum_0_bot: "a \ S \ a = enum 0 \ (\a'\S. a \ a')" unfolding s_eq by (auto simp: enum_mono Ball_def) -lemma enum_n_top: "a \ s \ a = enum n \ (\a'\s. a' \ a)" +lemma enum_n_top: "a \ S \ a = enum n \ (\a'\S. a' \ a)" unfolding s_eq by (auto simp: enum_mono Ball_def) lemma enum_Suc: "i < n \ enum (Suc i) = (enum i)(upd i := Suc (enum i (upd i)))" @@ -499,51 +499,51 @@ lemma enum_eq_p: "i \ n \ n \ j \ enum i j = p" by (induct i) (auto simp: enum_Suc enum_0 base_out upd_space not_less[symmetric]) -lemma out_eq_p: "a \ s \ n \ j \ a j = p" +lemma out_eq_p: "a \ S \ n \ j \ a j = p" unfolding s_eq by (auto simp: enum_eq_p) -lemma s_le_p: "a \ s \ a j \ p" +lemma s_le_p: "a \ S \ a j \ p" using out_eq_p[of a j] s_space by (cases "j < n") auto -lemma le_Suc_base: "a \ s \ a j \ Suc (base j)" +lemma le_Suc_base: "a \ S \ a j \ Suc (base j)" unfolding s_eq by (auto simp: enum_def) -lemma base_le: "a \ s \ base j \ a j" +lemma base_le: "a \ S \ base j \ a j" unfolding s_eq by (auto simp: enum_def) lemma enum_le_p: "i \ n \ j < n \ enum i j \ p" using enum_in[of i] s_space by auto -lemma enum_less: "a \ s \ i < n \ enum i < a \ enum (Suc i) \ a" +lemma enum_less: "a \ S \ i < n \ enum i < a \ enum (Suc i) \ a" unfolding s_eq by (auto simp: enum_strict_mono enum_mono) lemma ksimplex_0: - "n = 0 \ s = {(\x. p)}" + "n = 0 \ S = {(\x. p)}" using s_eq enum_def base_out by auto lemma replace_0: - assumes "j < n" "a \ s" and p: "\x\s - {a}. x j = 0" and "x \ s" + assumes "j < n" "a \ S" and p: "\x\S - {a}. x j = 0" and "x \ S" shows "x \ a" proof cases assume "x \ a" have "a j \ 0" using assms by (intro one_step[where a=a]) auto - with less[OF \x\s\ \a\s\, of j] p[rule_format, of x] \x \ s\ \x \ a\ + with less[OF \x\S\ \a\S\, of j] p[rule_format, of x] \x \ S\ \x \ a\ show ?thesis by auto qed simp lemma replace_1: - assumes "j < n" "a \ s" and p: "\x\s - {a}. x j = p" and "x \ s" + assumes "j < n" "a \ S" and p: "\x\S - {a}. x j = p" and "x \ S" shows "a \ x" proof cases assume "x \ a" have "a j \ p" using assms by (intro one_step[where a=a]) auto - with enum_le_p[of _ j] \j < n\ \a\s\ + with enum_le_p[of _ j] \j < n\ \a\S\ have "a j < p" by (auto simp: less_le s_eq) - with less[OF \a\s\ \x\s\, of j] p[rule_format, of x] \x \ s\ \x \ a\ + with less[OF \a\S\ \x\S\, of j] p[rule_format, of x] \x \ S\ \x \ a\ show ?thesis by auto qed simp diff -r ba2afdd29e1d -r 032a4344903e src/HOL/Analysis/Change_Of_Vars.thy --- a/src/HOL/Analysis/Change_Of_Vars.thy Tue Aug 01 11:27:55 2023 +0200 +++ b/src/HOL/Analysis/Change_Of_Vars.thy Thu Aug 03 19:10:43 2023 +0200 @@ -48,10 +48,8 @@ by (metis \m \ n\ axis_index_axis eq_iff_diff_eq_0 negligible_hyperplane) moreover have "(cbox a ?c \ {x. ?mn \ x \ a $ m} \ (cbox a ?c \ {x. ?mn \ x \ b$m})) \ {x. ?mn \ x = b$m}" using \m \ n\ ab_ne - apply (auto simp: algebra_simps mem_box_cart inner_axis') - apply (drule_tac x=m in spec)+ - apply simp - done + apply (clarsimp simp: algebra_simps mem_box_cart inner_axis') + by (smt (verit, ccfv_SIG) interval_ne_empty_cart(1)) ultimately show "negligible (cbox a ?c \ {x. ?mn \ x \ a $ m} \ (cbox a ?c \ {x. ?mn \ x \ b$m}))" by (rule negligible_subset) show "?f ` cbox a b \ cbox a ?c \ {x. ?mn \ x \ a $ m} \ cbox a ?c \ {x. ?mn \ x \ b$m} = cbox a ?c" (is "?lhs = _") @@ -59,9 +57,8 @@ show "?lhs \ cbox a ?c" by (auto simp: mem_box_cart add_mono) (meson add_increasing2 an order_trans) show "cbox a ?c \ ?lhs" - apply (auto simp: algebra_simps image_iff inner_axis' lambda_add_Galois [OF \m \ n\]) - apply (auto simp: mem_box_cart split: if_split_asm) - done + apply (clarsimp simp: algebra_simps image_iff inner_axis' lambda_add_Galois [OF \m \ n\]) + by (smt (verit, del_insts) mem_box_cart(2) vec_lambda_beta) qed qed (fact fab) let ?d = "\ i. if i = m then a $ m - b $ m else 0" @@ -106,17 +103,15 @@ by (rule negligible_subset) qed have ac_ne: "cbox a ?c \ {}" - using ab_ne an - by (clarsimp simp: interval_eq_empty_cart) (meson add_less_same_cancel1 le_less_linear less_le_trans) + by (smt (verit, del_insts) ab_ne an interval_ne_empty_cart(1) vec_lambda_beta) have ax_ne: "cbox a (\ i. if i = m then a $ m + b $ n else b $ i) \ {}" using ab_ne an - by (clarsimp simp: interval_eq_empty_cart) (meson add_less_same_cancel1 le_less_linear less_le_trans) + by (smt (verit, ccfv_threshold) interval_ne_empty_cart(1) vec_lambda_beta) have eq3: "measure lebesgue (cbox a ?c) = measure lebesgue (cbox a (\ i. if i = m then a$m + b$n else b$i)) + measure lebesgue (cbox a b)" by (simp add: content_cbox_if_cart ab_ne ac_ne ax_ne algebra_simps prod.delta_remove if_distrib [of "\u. u - z" for z] prod.remove) show ?Q - using eq1 eq2 eq3 - by (simp add: algebra_simps) + using eq1 eq2 eq3 by (simp add: algebra_simps) qed @@ -214,10 +209,8 @@ by auto have nfS: "negligible (?f ` S)" by (rule negligible_subset [OF negligible_standard_hyperplane_cart]) (use \m k = 0\ in auto) - then have "(?f ` S) \ lmeasurable" - by (simp add: negligible_iff_measure) - with nfS show ?thesis - by (simp add: prm negligible_iff_measure0) + then show ?thesis + by (simp add: negligible_iff_measure prm) qed then show "(?f ` S) \ lmeasurable" ?MEQ by metis+ @@ -297,10 +290,7 @@ then have 1: "\det (matrix ?h)\ = 1" by (simp add: det_permute_columns permutes_swap_id sign_swap_id abs_mult) show "?h ` S \ lmeasurable \ ?Q ?h S" - proof - show "?h ` S \ lmeasurable" "?Q ?h S" - using measure_linear_sufficient [OF lin \S \ lmeasurable\] meq 1 by force+ - qed + using measure_linear_sufficient [OF lin \S \ lmeasurable\] meq 1 by force next fix m n :: "'n" and S :: "(real, 'n) vec set" assume "m \ n" and "S \ lmeasurable" @@ -355,7 +345,7 @@ by (rule ssubst) (rule measure_translation) also have "\ = measure lebesgue ((\v. \ i. if i = m then v $ m + v $ n else v $ i) ` cbox (?v +a) (?v + b))" by (metis (no_types, lifting) cbox_translation) - also have "\ = measure lebesgue ((+) (\ i. if i = n then - a $ n else 0) ` cbox a b)" + also have "\ = measure lebesgue ((+) ?v ` cbox a b)" apply (subst measure_shear_interval) using \m \ n\ ne apply auto apply (simp add: cbox_translation) @@ -617,9 +607,7 @@ have inj: "inj_on (\(x, y). ball x y) K" by (force simp: inj_on_def ball_eq_ball_iff dest: gt0) have disjnt: "disjoint ((\(x, y). ball x y) ` K)" - using pwC that - apply (clarsimp simp: pairwise_def case_prod_unfold ball_eq_ball_iff) - by (metis subsetD fst_conv snd_conv) + using pwC that pairwise_image pairwise_mono by fastforce have "?l \ (\i\K. ?\ (case i of (x, s) \ f ` (S \ ball x s)))" proof (rule measure_UNION_le [OF \finite K\], clarify) fix x r @@ -630,8 +618,7 @@ by (meson Int_lower1 S differentiable_on_subset f_diff fmeasurableD lmeasurable_ball order_refl sets.Int differentiable_image_in_sets_lebesgue) qed also have "\ \ (\(x,s) \ K. (B + e) * ?\ (ball x s))" - apply (rule sum_mono) - using Csub r \K \ C\ by auto + using Csub r \K \ C\ by (intro sum_mono) auto also have "\ = (B + e) * (\(x,s) \ K. ?\ (ball x s))" by (simp add: prod.case_distrib sum_distrib_left) also have "\ = (B + e) * sum ?\ ((\(x, y). ball x y) ` K)" @@ -641,8 +628,8 @@ by (subst measure_Union') (auto simp: disjnt measure_Union') also have "\ \ (B + e) * ?\ T" using \B \ 0\ \e > 0\ that apply simp - apply (rule measure_mono_fmeasurable [OF _ _ \T \ lmeasurable\]) - using Csub rT by force+ + using measure_mono_fmeasurable [OF _ _ \T \ lmeasurable\] Csub rT + by (smt (verit) SUP_least measure_nonneg measure_notin_sets mem_Collect_eq old.prod.case subset_iff) also have "\ \ (B + e) * (?\ S + d)" using \B \ 0\ \e > 0\ Tless by simp finally show ?thesis . @@ -732,10 +719,8 @@ using set_integrable_subset [OF aint_S] meas_t T_def by blast have Seq: "S = (\n. T n)" apply (auto simp: T_def) - apply (rule_tac x="nat(floor(abs(det(matrix(f' x))) / e))" in exI) - using that apply auto - using of_int_floor_le pos_le_divide_eq apply blast - by (metis add.commute pos_divide_less_eq real_of_int_floor_add_one_gt) + apply (rule_tac x = "nat \\det (matrix (f' x))\ / e\" in exI) + by (smt (verit, del_insts) divide_nonneg_nonneg floor_eq_iff of_nat_nat pos_divide_less_eq that zero_le_floor) have meas_ft: "f ` T n \ lmeasurable" for n proof (rule measurable_bounded_differentiable_image) show "T n \ lmeasurable" @@ -779,8 +764,8 @@ case (less m n) have False if "T n \ T m" "x \ T n" for x using \e > 0\ \m < n\ that - apply (auto simp: T_def mult.commute intro: less_le_trans dest!: subsetD) - by (metis add.commute less_le_trans nat_less_real_le not_le mult_le_cancel_iff2) + apply (auto simp: T_def mult.commute intro: less_le_trans dest!: subsetD) + by (smt (verit, best) mult_less_cancel_left_disj nat_less_real_le) then show ?case using less.prems by blast qed auto @@ -839,10 +824,10 @@ unfolding m_def proof (rule integral_subset_le) have "(\x. \det (matrix (f' x))\) absolutely_integrable_on (\k\n. T k)" - apply (rule set_integrable_subset [OF aint_S]) - apply (intro measurable meas_t fmeasurableD) - apply (force simp: Seq) - done + proof (rule set_integrable_subset [OF aint_S]) + show "\ (T ` {..n}) \ sets lebesgue" + by (intro measurable meas_t fmeasurableD) + qed (force simp: Seq) then show "(\x. \det (matrix (f' x))\) integrable_on (\k\n. T k)" using absolutely_integrable_on_def by blast qed (use Seq int in auto) @@ -905,9 +890,9 @@ let ?I = "\n::nat. cbox (vec (-n)) (vec n) \ S" let ?\ = "measure lebesgue" have "x \ cbox (vec (- real (nat \norm x\))) (vec (real (nat \norm x\)))" for x :: "real^'n::_" - apply (auto simp: mem_box_cart) - apply (metis abs_le_iff component_le_norm_cart minus_le_iff of_nat_ceiling order.trans) - by (meson abs_le_D1 norm_bound_component_le_cart real_nat_ceiling_ge) + apply (simp add: mem_box_cart) + by (smt (verit, best) Finite_Cartesian_Product.norm_nth_le nat_ceiling_le_eq + real_nat_ceiling_ge real_norm_def) then have Seq: "S = (\n. ?I n)" by auto have fIn: "f ` ?I n \ lmeasurable" @@ -918,14 +903,7 @@ moreover have "\x. x \ ?I n \ (f has_derivative f' x) (at x within ?I n)" by (meson Int_iff deriv has_derivative_subset subsetI) moreover have int_In: "(\x. \det (matrix (f' x))\) integrable_on ?I n" - proof - - have "(\x. \det (matrix (f' x))\) absolutely_integrable_on S" - using int absolutely_integrable_integrable_bound by force - then have "(\x. \det (matrix (f' x))\) absolutely_integrable_on ?I n" - by (metis (no_types) Int_lower1 In fmeasurableD inf_commute set_integrable_subset) - then show ?thesis - using absolutely_integrable_on_def by blast - qed + by (metis (mono_tags) Int_commute int integrable_altD(1) integrable_restrict_Int) ultimately have "f ` ?I n \ lmeasurable" "?\ (f ` ?I n) \ integral (?I n) (\x. \det (matrix (f' x))\)" using m_diff_image_weak by metis+ moreover have "integral (?I n) (\x. \det (matrix (f' x))\) \ integral S (\x. \det (matrix (f' x))\)" @@ -972,14 +950,16 @@ proof - have g: "g n x \ g (N + n) x" for N by (rule transitive_stepwise_le) (use inc_g in auto) - have "\na\N. g n x - f x \ dist (g na x) (f x)" for N - apply (rule_tac x="N+n" in exI) - using g [of N] by (auto simp: dist_norm) + have "\m\N. g n x - f x \ dist (g m x) (f x)" for N + proof + show "N \ N + n \ g n x - f x \ dist (g (N + n) x) (f x)" + using g [of N] by (auto simp: dist_norm) + qed with that show ?thesis using diff_gt_0_iff_gt by blast qed with lim show ?thesis - apply (auto simp: lim_sequentially) + unfolding lim_sequentially by (meson less_le_not_le not_le_imp_less) qed moreover @@ -993,10 +973,8 @@ fix k::real assume "k \ \" and k: "\k\ \ 2 ^ (2*n)" show "0 \ k/2^n * ?\ n k x" - using f \k \ \\ apply (auto simp: indicator_def field_split_simps Ints_def) - apply (drule spec [where x=x]) - using zero_le_power [of "2::real" n] mult_nonneg_nonneg [of "f x" "2^n"] - by linarith + using f \k \ \\ apply (clarsimp simp: indicator_def field_split_simps Ints_def) + by (smt (verit) int_less_real_le mult_nonneg_nonneg of_int_0 zero_le_power) qed show "?g n x \ ?g (Suc n) x" for n x proof - @@ -1049,19 +1027,12 @@ show "finite {k::real. k \ \ \ \k\ \ 2 ^ (2 * Suc n)}" by (rule finite_abs_int_segment) show "(*) 2 ` {k::real. k \ \ \ \k\ \ 2^(2*n)} \ (\x. 2*x + 1) ` {k \ \. \k\ \ 2^(2*n)} \ {k \ \. \k\ \ 2 ^ (2 * Suc n)}" - apply auto + apply (clarsimp simp: image_subset_iff) using one_le_power [of "2::real" "2*n"] by linarith have *: "\x \ (S \ T) - U; \x. x \ S \ x \ U; \x. x \ T \ x \ U\ \ P x" for S T U P by blast have "0 \ b" if "b \ \" "f x * (2 * 2^n) < b + 1" for b - proof - - have "0 \ f x * (2 * 2^n)" - by (simp add: f) - also have "\ < b+1" - by (simp add: that) - finally show "0 \ b" - using \b \ \\ by (auto simp: elim!: Ints_cases) - qed + by (smt (verit, ccfv_SIG) Ints_cases f int_le_real_less mult_nonneg_nonneg of_int_add one_le_power that) then show "0 \ b/2 ^ Suc n * indicator {y. b/2 ^ Suc n \ f y \ f y < (b + 1)/2 ^ Suc n} x" if "b \ {k \ \. \k\ \ 2 ^ (2 * Suc n)} - ((*) 2 ` {k \ \. \k\ \ 2 ^ (2*n)} \ (\x. 2*x + 1) ` {k \ \. \k\ \ 2 ^ (2*n)})" for b @@ -1125,10 +1096,8 @@ using N2 by (simp add: divide_simps mult.commute) linarith also have "\ \ \2^n\ * e" using that \e > 0\ by auto - finally have "dist (?m/2^n) (f x) < e" - by (simp add: dist_norm) - then show ?thesis - using eq by linarith + finally show ?thesis + using eq by (simp add: dist_real_def) qed then show "\no. \n\no. dist (\k | k \ \ \ \k\ \ 2 ^ (2*n). k * ?\ n k x/2^n) (f x) < e" by force @@ -1213,9 +1182,7 @@ ultimately have dim: "dim {x. f x = 0} = DIM('a)" by force then show ?thesis - using dim_eq_full - by (metis (mono_tags, lifting) eq_0_on_span eucl.span_Basis linear_axioms linear_eq_stdbasis - mem_Collect_eq module_hom_zero span_base span_raw_def) + by (metis (mono_tags, lifting) dim_eq_full UNIV_I eq_0_on_span mem_Collect_eq span_raw_def) qed lemma lemma_partial_derivatives: @@ -1368,8 +1335,7 @@ next case False have "\v \ (y - x)\ < norm v * e / 2 / real (CARD('m) ^ CARD('m)) * norm (x - y)" - apply (rule dless) - using False \y \ S\ d by (auto simp: norm_minus_commute) + by (metis Diff_iff False \y \ S\ d dless empty_iff insert_iff norm_minus_commute) also have "\ \ norm v * e * min d r / (2 * real CARD('m) ^ CARD('m))" using d r \e > 0\ by (simp add: field_simps norm_minus_commute mult_left_mono) finally show ?thesis . @@ -1500,8 +1466,7 @@ also have "\ = UNIV" proof - have "dim ?CA \ CARD('m)" - using dim_subset_UNIV[of ?CA] - by auto + using dim_subset_UNIV[of ?CA] by auto moreover have "False" if less: "dim ?CA < CARD('m)" proof - obtain d where "d \ 0" and d: "\y. y \ span ?CA \ orthogonal d y" @@ -1570,7 +1535,7 @@ proof (rule eventuallyI) fix n show "0 \ \d \ ((\ n - x) /\<^sub>R norm (\ n - x))\ - \" - using \le [of n] \Sx by (auto simp: abs_mult divide_simps) + using \le [of n] \Sx by (auto simp: abs_mult divide_simps) qed ultimately have "\ \ \d \ z\" using tendsto_lowerbound [where a=0] by fastforce @@ -1599,8 +1564,7 @@ also have "\ \ 1/(Suc j) * norm (\ p - x) + 1/(Suc i) * norm (\ p - x)" by (metis A Diff_iff \Sx dist_norm p add_mono) also have "\ \ 1/N * norm (\ p - x) + 1/N * norm (\ p - x)" - apply (intro add_mono mult_right_mono) - using ij \N > 0\ by (auto simp: field_simps) + using ij \N > 0\ by (intro add_mono mult_right_mono) (auto simp: field_simps) also have "\ = 2 / N * norm (\ p - x)" by simp finally have no_le: "norm ((A i - A j) *v (\ p - x)) \ 2 / N * norm (\ p - x)" . @@ -1807,8 +1771,12 @@ by metis qed also have "\ \ e * norm (y - x) / 4" - using \e > 0\ apply (simp add: norm_mult abs_mult) - by (metis component_le_norm_cart vector_minus_component) + proof - + have "\y $ n - x $ n\ \ norm (y - x)" + by (metis component_le_norm_cart vector_minus_component) + with \e > 0\ show ?thesis + by (simp add: norm_mult abs_mult) + qed finally show "norm ((matrix (f' x) - B) *v (y - x) - ((?A - B) *v (y - x))) \ e * norm (y - x) / 4" . show "0 < e * norm (y - x)" by (simp add: False \e > 0\) @@ -1869,12 +1837,7 @@ lemma sets_lebesgue_almost_borel: assumes "S \ sets lebesgue" obtains B N where "B \ sets borel" "negligible N" "B \ N = S" -proof - - obtain T N N' where "S = T \ N" "N \ N'" "N' \ null_sets lborel" "T \ sets borel" - using sets_completionE [OF assms] by auto - then show thesis - by (metis negligible_iff_null_sets negligible_subset null_sets_completionI that) -qed + by (metis assms negligible_iff_null_sets negligible_subset null_sets_completionI sets_completionE sets_lborel) lemma double_lebesgue_sets: assumes S: "S \ sets lebesgue" and T: "T \ sets lebesgue" and fim: "f ` S \ T" @@ -1906,8 +1869,7 @@ and 2 [rule_format]: "\U. negligible U \ U \ T \ {x \ S. f x \ U} \ sets lebesgue" and "U \ sets lebesgue" "U \ T" then obtain C N where C: "C \ sets borel \ negligible N \ C \ N = U" - using sets_lebesgue_almost_borel - by metis + using sets_lebesgue_almost_borel by metis then have "{x \ S. f x \ C} \ sets lebesgue" by (blast intro: 1) moreover have "{x \ S. f x \ N} \ sets lebesgue" @@ -1974,12 +1936,7 @@ and mS: "measure lebesgue S \ (2 * e) * (2 * m) ^ (CARD('n) - 1)" proof (rule Sard_lemma00 [of "norm a" "axis k (1::real)" "T-`P" m e]) have "norm a *\<^sub>R axis k 1 \ x = 0" if "T x \ P" for x - proof - - have "a \ T x = 0" - using P that by blast - then show ?thesis - by (metis (no_types, lifting) T a orthogonal_orthogonal_transformation orthogonal_def) - qed + by (smt (verit, del_insts) P T a mem_Collect_eq orthogonal_transformation_def subset_eq that) then show "T -` P \ {x. norm a *\<^sub>R axis k 1 \ x = 0}" by auto qed (use assms T in auto) @@ -1990,22 +1947,11 @@ have "{z. norm z \ m \ (\t\P. norm (z - t) \ e)} \ T ` {z. norm z \ m \ (\t\T -` P. norm (z - t) \ e)}" proof clarsimp fix x t - assume "norm x \ m" "t \ P" "norm (x - t) \ e" + assume \
: "norm x \ m" "t \ P" "norm (x - t) \ e" then have "norm (inv T x) \ m" using orthogonal_transformation_inv [OF T] by (simp add: orthogonal_transformation_norm) moreover have "\t\T -` P. norm (inv T x - t) \ e" - proof - have "T (inv T x - inv T t) = x - t" - using T linear_diff orthogonal_transformation_def - by (metis (no_types, opaque_lifting) Tinv) - then have "norm (inv T x - inv T t) = norm (x - t)" - by (metis T orthogonal_transformation_norm) - then show "norm (inv T x - inv T t) \ e" - using \norm (x - t) \ e\ by linarith - next - show "inv T t \ T -` P" - using \t \ P\ by force - qed + by (smt (verit, del_insts) T Tinv \
linear_diff orthogonal_transformation_def orthogonal_transformation_norm vimage_eq) ultimately show "x \ T ` {z. norm z \ m \ (\t\T -` P. norm (z - t) \ e)}" by force qed @@ -2124,8 +2070,8 @@ and subT: "{z. norm(z - f x) \ (2 * B) * norm(v - u) \ (\t \ range (f' x). norm(z - f x - t) \ d * norm(v - u))} \ T" and measT: "?\ T \ (2 * (d * norm(v - u))) * (2 * ((2 * B) * norm(v - u))) ^ (?n - 1)" (is "_ \ ?DVU") - apply (rule Sard_lemma1 [of "range (f' x)" "(2 * B) * norm(v - u)" "d * norm(v - u)" "f x"]) - using \B > 0\ \d > 0\ by simp_all + using Sard_lemma1 [of "range (f' x)" "(2 * B) * norm(v - u)" "d * norm(v - u)"] + using \B > 0\ \d > 0\ by auto show ?thesis proof (intro exI conjI) have "f ` (K \ S) \ {z. norm(z - f x) \ (2 * B) * norm(v - u) \ (\t \ range (f' x). norm(z - f x - t) \ d * norm(v - u))}" @@ -2150,18 +2096,13 @@ have "norm (f' x (y - x)) \ onorm (f' x) * norm (y - x)" using onorm [of "f' x" "y-x"] by (meson IntE lin_f' linear_linear x(1)) also have "\ \ B * norm (v - u)" - proof (rule mult_mono) - show "onorm (f' x) \ B" - using B x by blast - qed (use \B > 0\ yx_le in auto) + by (meson B IntE lin_f' linear_linear mult_mono' norm_ge_zero onorm_pos_le x(1) yx_le) finally show "norm (f' x (y - x)) \ B * norm (v - u)" . show "d * norm (y - x) \ B * norm (v - u)" using \B > 0\ by (auto intro: mult_mono [OF \d \ B\ yx_le]) qed show "\t. norm (f y - f x - f' x t) \ d * norm (v - u)" - apply (rule_tac x="y-x" in exI) - using \d > 0\ yx_le le_dyx mult_left_mono [where c=d] - by (meson order_trans mult_le_cancel_iff2) + by (smt (verit, best) \0 < d\ le_dyx mult_le_cancel_left_pos yx_le) qed with subT show "f ` (K \ S) \ T" by blast show "?\ T \ e / (2*c) ^ ?m * ?\ K" @@ -2355,8 +2296,12 @@ have "(\x. if x \ {t. h n (g t) = y} then ?D x else 0) integrable_on S" proof (rule measurable_bounded_by_integrable_imp_integrable) have "(\x. ?D x) \ borel_measurable (lebesgue_on ({t. h n (g t) = y} \ S))" - apply (intro borel_measurable_abs borel_measurable_det_Jacobian [OF h_lmeas, where f=g]) - by (meson der_g IntD2 has_derivative_subset inf_le2) + proof - + have "(\v. det (matrix (g' v))) \ borel_measurable (lebesgue_on (S \ {v. h n (g v) = y}))" + by (metis Int_lower1 S assms(4) borel_measurable_det_Jacobian measurable_restrict_mono) + then show ?thesis + by (simp add: Int_commute) + qed then have "(\x. if x \ {t. h n (g t) = y} \ S then ?D x else 0) \ borel_measurable lebesgue" by (rule borel_measurable_if_I [OF _ h_lmeas]) then show "(\x. if x \ {t. h n (g t) = y} then ?D x else 0) \ borel_measurable (lebesgue_on S)" @@ -2371,28 +2316,23 @@ then have int_det: "(\t. \det (matrix (g' t))\) integrable_on ({t. h n (g t) = y} \ S)" using integrable_restrict_Int by force have "(g ` ({t. h n (g t) = y} \ S)) \ lmeasurable" - apply (rule measurable_differentiable_image [OF h_lmeas]) - apply (blast intro: has_derivative_subset [OF der_g]) - apply (rule int_det) - done + by (blast intro: has_derivative_subset [OF der_g] measurable_differentiable_image [OF h_lmeas] int_det) moreover have "g ` ({t. h n (g t) = y} \ S) = {x. h n x = y} \ g ` S" by blast moreover have "measure lebesgue (g ` ({t. h n (g t) = y} \ S)) \ integral ({t. h n (g t) = y} \ S) (\t. \det (matrix (g' t))\)" - apply (rule measure_differentiable_image [OF h_lmeas _ int_det]) - apply (blast intro: has_derivative_subset [OF der_g]) - done + by (blast intro: has_derivative_subset [OF der_g] measure_differentiable_image [OF h_lmeas _ int_det]) ultimately show ?thesis using \y > 0\ integral_restrict_Int [of S "{t. h n (g t) = y}" "\t. \det (matrix (g' t))\ * y"] apply (simp add: integrable_on_indicator integral_indicator) apply (simp add: indicator_def of_bool_def if_distrib cong: if_cong) done qed - have hn_int: "h n integrable_on g ` S" - apply (subst hn_eq) - using yind by (force intro: integrable_sum [OF fin_R]) - then show ?thesis + show ?thesis proof + show "h n integrable_on g ` S" + apply (subst hn_eq) + using yind by (force intro: integrable_sum [OF fin_R]) have "?lhs = integral (g ` S) (\x. \y\range (h n). y * indicat_real {x. h n x = y} x)" by (metis hn_eq) also have "\ = (\y\range (h n). integral (g ` S) (\x. y * indicat_real {x. h n x = y} x))" @@ -2414,8 +2354,8 @@ next fix x assume "x \ S" - have "y * indicat_real {x. h n x = y} (g x) \ f (g x)" - by (metis \x \ S\ h_le_f indicator_simps(1) indicator_simps(2) mem_Collect_eq mult.commute mult.left_neutral mult_zero_left nonneg_fg) + then have "y * indicat_real {x. h n x = y} (g x) \ f (g x)" + by (metis (full_types) h_le_f indicator_simps mem_Collect_eq mult.right_neutral mult_zero_right nonneg_fg) with \y \ 0\ show "norm (?D x * y * indicat_real {x. h n x = y} (g x)) \ ?D x * f(g x)" by (simp add: abs_mult mult.assoc mult_left_mono) qed (use S det_int_fg in auto) @@ -2522,8 +2462,10 @@ by (rule Df_borel) finally have *: "(\x. \?D x\ * f (g x)) \ borel_measurable (lebesgue_on S')" by (simp add: borel_measurable_if_D) - have "?h \ borel_measurable (lebesgue_on S')" - by (intro * S' der_gS' borel_measurable_det_Jacobian measurable) (blast intro: der_gS') + have "(\v. det (matrix (g' v))) \ borel_measurable (lebesgue_on S')" + using S' borel_measurable_det_Jacobian der_gS' by blast + then have "?h \ borel_measurable (lebesgue_on S')" + using "*" borel_measurable_abs borel_measurable_inverse borel_measurable_scaleR by blast moreover have "?h x = f(g x)" if "x \ S'" for x using that by (auto simp: S'_def) ultimately have "(\x. f(g x)) \ borel_measurable (lebesgue_on S')" @@ -2586,8 +2528,8 @@ then have "f integrable_on g ` {x \ S. ?D x \ 0}" by (auto simp: image_iff elim!: integrable_eq) then show "f integrable_on g ` S" - apply (rule integrable_spike_set [OF _ empty_imp_negligible negligible_subset]) - using negg null by auto + using negg null + by (auto intro: integrable_spike_set [OF _ empty_imp_negligible negligible_subset]) have "integral (g ` S) f = integral (g ` {x \ S. ?D x \ 0}) f" using negg by (auto intro: negligible_subset integral_spike_set) also have "\ = integral (g ` {x \ S. ?D x \ 0}) (\x. if x \ g ` ?F then f x else 0)" @@ -2653,15 +2595,15 @@ have der_gP: "(g has_derivative g' x) (at x within ?P)" if "x \ ?P" for x using der_g has_derivative_subset that by force - have "f integrable_on g ` ?P \ integral (g ` ?P) f \ integral ?P ?D" - proof (rule integral_on_image_ubound_nonneg [OF _ der_gP]) - have "?D integrable_on {x \ S. 0 < ?D x}" - using Dgt - by (auto intro: set_lebesgue_integral_eq_integral [OF set_integrable_subset] intS) - then show "?D integrable_on ?P" - apply (rule integrable_spike_set) - by (auto simp: zero_less_mult_iff empty_imp_negligible) - qed auto + have "f integrable_on g ` ?P \ integral (g ` ?P) f \ integral ?P ?D" + proof (rule integral_on_image_ubound_nonneg [OF _ der_gP]) + show "?D integrable_on ?P" + proof (rule integrable_spike_set) + show "?D integrable_on {x \ S. 0 < ?D x}" + using Dgt + by (auto intro: set_lebesgue_integral_eq_integral [OF set_integrable_subset] intS) + qed (auto simp: zero_less_mult_iff empty_imp_negligible) + qed auto then have "f integrable_on g ` ?P" by metis moreover have "g ` ?P = {y \ g ` S. f y > 0}" @@ -2730,8 +2672,7 @@ if "x \ T" for x proof - have "matrix (h' x) ** matrix (g' (h x)) = mat 1" - using that id[OF that] der_g[of "h x"] gh[OF that] left_inverse_linear has_derivative_linear - by (subst matrix_compose[symmetric]) (force simp: matrix_id_mat_1 has_derivative_linear)+ + by (metis der_g der_h gh has_derivative_linear local.id matrix_compose matrix_id_mat_1 that) then have "\det (matrix (h' x))\ * \det (matrix (g' (h x)))\ = 1" by (metis abs_1 abs_mult det_I det_mul) then show ?thesis @@ -2740,10 +2681,7 @@ have "?D integrable_on (h ` T)" proof (intro set_lebesgue_integral_eq_integral absolutely_integrable_on_image_real) show "(\x. ?fgh x) absolutely_integrable_on T" - proof (subst absolutely_integrable_on_iff_nonneg) - show "(\x. ?fgh x) integrable_on T" - using ddf fT integrable_eq by force - qed (simp add: zero_le_mult_iff f0 gh) + by (smt (verit, del_insts) abs_absolutely_integrableI_1 ddf f0 fT integrable_eq) qed (use der_h in auto) with Seq show "(\x. ?D x) integrable_on S" by simp @@ -2755,9 +2693,8 @@ qed (use f0 gh der_h in auto) also have "\ = integral T f" by (force simp: ddf intro: integral_cong) - also have "\ \ b" - by (rule intf) - finally show "integral S (\x. ?D x) \ b" . + finally show "integral S (\x. ?D x) \ b" + using intf by linarith qed next assume R: ?rhs @@ -2850,15 +2787,16 @@ using "-" [of "integral {x \ S. f(g x) < 0} ?DN"] aN by (auto simp: set_lebesgue_integral_eq_integral has_integral_iff integrable_neg_iff) have faN: "f absolutely_integrable_on {y \ T. f y < 0}" - apply (rule absolutely_integrable_integrable_bound [where g = "\x. - f x"]) - using fN by (auto simp: integrable_neg_iff) + proof (rule absolutely_integrable_integrable_bound) + show "(\x. - f x) integrable_on {y \ T. f y < 0}" + using fN by (auto simp: integrable_neg_iff) + qed (use fN in auto) have fP: "f integrable_on {y \ T. f y > 0}" "integral {y \ T. f y > 0} f = integral {x \ S. f (g x) > 0} ?DP" using "+" [of "integral {x \ S. f(g x) > 0} ?DP"] aP by (auto simp: set_lebesgue_integral_eq_integral has_integral_iff integrable_neg_iff) have faP: "f absolutely_integrable_on {y \ T. f y > 0}" - apply (rule absolutely_integrable_integrable_bound [where g = f]) - using fP by auto + using fP(1) nonnegative_absolutely_integrable_1 by fastforce have fa: "f absolutely_integrable_on ({y \ T. f y < 0} \ {y \ T. f y > 0})" by (rule absolutely_integrable_Un [OF faN faP]) show ?rhs @@ -2904,8 +2842,7 @@ have aint: "f absolutely_integrable_on {y. y \ T \ 0 < (f y)}" "f absolutely_integrable_on {y. y \ T \ (f y) < 0}" and intT: "integral T f = b" - using set_integrable_subset [of _ T] TP TN RHS - by blast+ + using set_integrable_subset [of _ T] TP TN RHS by blast+ show ?lhs proof have fN: "f integrable_on {v \ T. f v < 0}" @@ -3162,10 +3099,7 @@ by (meson der_g differentiableI UnionI differentiable_on_def differentiable_on_subset rangeI subsetI) qed auto qed - next - show "(\n. integral (?U n) ?D) \ integral (\n. F n) ?D" - by (rule DU) - qed + qed (use DU in metis) next assume fs: "f absolutely_integrable_on (\x. g ` F x)" and b: "b = integral ((\x. g ` F x)) f" @@ -3194,8 +3128,7 @@ unfolding integrable_restrict_UNIV absolutely_integrable_on_def by (simp add: image_UN) { fix n::nat have "(norm \ f) absolutely_integrable_on (\m\n. g ` F m)" - apply (rule absolutely_integrable_norm) - using fgU by blast + using absolutely_integrable_norm fgU by blast then have "integral (?U n) (norm \ ?D) = integral (g ` ?U n) (norm \ f)" using iff [of n "?lift \ norm \ f" "integral (g ` ?U n) (?lift \ norm \ f)"] unfolding absolutely_integrable_on_1_iff integral_on_1_eq image_UN by (auto simp: o_def) @@ -3233,11 +3166,9 @@ unfolding absolutely_integrable_restrict_UNIV by simp show "integral (\n. F n) ?D = integral ((\x. g ` F x)) f" proof (rule LIMSEQ_unique) - show "(\n. integral (\m\n. g ` F m) f) \ integral (\x. g ` F x) f" - by (rule fgU) show "(\n. integral (\m\n. g ` F m) f) \ integral (\n. F n) ?D" unfolding D_int [symmetric] by (rule integral_countable_UN [OF Dai F_leb]) - qed + qed (use fgU in metis) qed qed diff -r ba2afdd29e1d -r 032a4344903e src/HOL/Analysis/Complex_Transcendental.thy --- a/src/HOL/Analysis/Complex_Transcendental.thy Tue Aug 01 11:27:55 2023 +0200 +++ b/src/HOL/Analysis/Complex_Transcendental.thy Thu Aug 03 19:10:43 2023 +0200 @@ -75,7 +75,7 @@ lemma continuous_within_exp: fixes z::"'a::{real_normed_field,banach}" shows "continuous (at z within s) exp" -by (simp add: continuous_at_imp_continuous_within) + by (simp add: continuous_at_imp_continuous_within) lemma holomorphic_on_exp [holomorphic_intros]: "exp holomorphic_on s" by (simp add: field_differentiable_within_exp holomorphic_on_def) @@ -153,7 +153,7 @@ theorem Euler: "exp(z) = of_real(exp(Re z)) * (of_real(cos(Im z)) + \ * of_real(sin(Im z)))" -by (cases z) (simp add: exp_add exp_Euler cos_of_real exp_of_real sin_of_real Complex_eq) + by (simp add: Complex_eq cis.code exp_eq_polar) lemma Re_sin: "Re(sin z) = sin(Re z) * (exp(Im z) + exp(-(Im z))) / 2" by (simp add: sin_exp_eq field_simps Re_divide Im_exp) @@ -568,8 +568,8 @@ lemma cos_eq: "cos x = cos y \ (\n \ \. x = y + 2*n*pi \ x = -y + 2*n*pi)" - using complex_cos_eq [of x y] - by (simp only: cos_of_real Re_complex_of_real of_real_add [symmetric] of_real_minus [symmetric] of_real_mult [symmetric] of_real_eq_iff) + using complex_cos_eq [of x y] unfolding cos_of_real + by (metis Re_complex_of_real of_real_add of_real_minus) lemma sinh_complex: fixes z :: complex @@ -635,7 +635,7 @@ lemma norm_cos_plus1_le: fixes z::complex shows "norm(1 + cos z) \ 2 * exp(norm z)" - by (smt (verit, best) exp_ge_add_one_self norm_cos_le norm_ge_zero norm_one norm_triangle_ineq) + by (metis mult_2 norm_cos_le norm_ge_zero norm_one norm_triangle_mono one_le_exp_iff) lemma sinh_conv_sin: "sinh z = -\ * sin (\*z)" by (simp add: sinh_field_def sin_i_times exp_minus) @@ -3854,8 +3854,7 @@ \ * (of_nat k * (of_real pi * 2)) + \ * (of_int z * (of_nat n * (of_real pi * 2)))) \ j mod n = k mod n" . note * = this show ?thesis - using assms - by (simp add: exp_eq field_split_simps *) + using assms by (simp add: exp_eq field_split_simps *) qed corollary bij_betw_roots_unity: diff -r ba2afdd29e1d -r 032a4344903e src/HOL/Analysis/Connected.thy --- a/src/HOL/Analysis/Connected.thy Tue Aug 01 11:27:55 2023 +0200 +++ b/src/HOL/Analysis/Connected.thy Thu Aug 03 19:10:43 2023 +0200 @@ -20,23 +20,12 @@ e1 \ e2 = {} \ e1 \ {} \ e2 \ {})" - unfolding connected_def openin_open - by safe blast+ + using connected_openin by blast lemma exists_diff: fixes P :: "'a set \ bool" shows "(\S. P (- S)) \ (\S. P S)" - (is "?lhs \ ?rhs") -proof - - have ?rhs if ?lhs - using that by blast - moreover have "P (- (- S))" if "P S" for S - proof - - have "S = - (- S)" by simp - with that show ?thesis by metis - qed - ultimately show ?thesis by metis -qed + by (metis boolean_algebra_class.boolean_algebra.double_compl) lemma connected_clopen: "connected S \ (\T. openin (top_of_set S) T \ @@ -80,10 +69,10 @@ by (auto simp: connected_component_def) lemma connected_component_refl: "x \ S \ connected_component S x x" - by (auto simp: connected_component_def) (use connected_sing in blast) - + using connected_component_def connected_sing by blast + lemma connected_component_refl_eq [simp]: "connected_component S x x \ x \ S" - by (auto simp: connected_component_refl) (auto simp: connected_component_def) + using connected_component_in connected_component_refl by blast lemma connected_component_sym: "connected_component S x y \ connected_component S y x" by (auto simp: connected_component_def) @@ -105,23 +94,7 @@ lemma connected_iff_eq_connected_component_set: "connected S \ (\x \ S. connected_component_set S x = S)" -proof (cases "S = {}") - case True - then show ?thesis by simp -next - case False - then obtain x where "x \ S" by auto - show ?thesis - proof - assume "connected S" - then show "\x \ S. connected_component_set S x = S" - by (force simp: connected_component_def) - next - assume "\x \ S. connected_component_set S x = S" - then show "connected S" - by (metis \x \ S\ connected_connected_component) - qed -qed + by (metis connected_component_def connected_component_in connected_connected_component mem_Collect_eq subsetI subset_antisym) lemma connected_component_subset: "connected_component_set S x \ S" using connected_component_in by blast @@ -165,34 +138,27 @@ unfolding closure_eq [symmetric] proof show "closure (connected_component_set S x) \ connected_component_set S x" - apply (rule connected_component_maximal) - apply (simp add: closure_def True) - apply (simp add: connected_imp_connected_closure) - apply (simp add: S closure_minimal connected_component_subset) - done - next - show "connected_component_set S x \ closure (connected_component_set S x)" - by (simp add: closure_subset) - qed + proof (rule connected_component_maximal) + show "x \ closure (connected_component_set S x)" + by (simp add: closure_def True) + show "connected (closure (connected_component_set S x))" + by (simp add: connected_imp_connected_closure) + show "closure (connected_component_set S x) \ S" + by (simp add: S closure_minimal connected_component_subset) + qed + qed (simp add: closure_subset) qed lemma connected_component_disjoint: "connected_component_set S a \ connected_component_set S b = {} \ a \ connected_component_set S b" - apply (auto simp: connected_component_eq) - using connected_component_eq connected_component_sym - apply blast - done + by (smt (verit) connected_component_eq connected_component_eq_empty connected_component_refl_eq + disjoint_iff_not_equal mem_Collect_eq) lemma connected_component_nonoverlap: "connected_component_set S a \ connected_component_set S b = {} \ a \ S \ b \ S \ connected_component_set S a \ connected_component_set S b" - apply (auto simp: connected_component_in) - using connected_component_refl_eq - apply blast - apply (metis connected_component_eq mem_Collect_eq) - apply (metis connected_component_eq mem_Collect_eq) - done + by (metis connected_component_disjoint connected_component_eq connected_component_eq_empty inf.idem) lemma connected_component_overlap: "connected_component_set S a \ connected_component_set S b \ {} \ @@ -205,13 +171,8 @@ lemma connected_component_eq_eq: "connected_component_set S x = connected_component_set S y \ x \ S \ y \ S \ x \ S \ y \ S \ connected_component S x y" - apply (cases "y \ S", simp) - apply (metis connected_component_eq connected_component_eq_empty connected_component_refl_eq mem_Collect_eq) - apply (cases "x \ S", simp) - apply (metis connected_component_eq_empty) - using connected_component_eq_empty - apply blast - done + by (metis connected_component_eq connected_component_eq_empty connected_component_refl mem_Collect_eq) + lemma connected_iff_connected_component_eq: "connected S \ (\x \ S. \y \ S. connected_component_set S x = connected_component_set S y)" @@ -219,19 +180,13 @@ lemma connected_component_idemp: "connected_component_set (connected_component_set S x) x = connected_component_set S x" - apply (rule subset_antisym) - apply (simp add: connected_component_subset) - apply (metis connected_component_eq_empty connected_component_maximal - connected_component_refl_eq connected_connected_component mem_Collect_eq set_eq_subset) - done + by (metis Int_absorb connected_component_disjoint connected_component_eq_empty connected_component_eq_self connected_connected_component) lemma connected_component_unique: "\x \ c; c \ S; connected c; \c'. \x \ c'; c' \ S; connected c'\ \ c' \ c\ \ connected_component_set S x = c" - apply (rule subset_antisym) - apply (meson connected_component_maximal connected_component_subset connected_connected_component contra_subsetD) - by (simp add: connected_component_maximal) + by (meson connected_component_maximal connected_component_subset connected_connected_component subsetD subset_antisym) lemma joinable_connected_component_eq: "\connected T; T \ S; @@ -241,25 +196,27 @@ by (metis (full_types) subsetD connected_component_eq connected_component_maximal disjoint_iff_not_equal) lemma Union_connected_component: "\(connected_component_set S ` S) = S" - apply (rule subset_antisym) - apply (simp add: SUP_least connected_component_subset) - using connected_component_refl_eq - by force - +proof + show "\(connected_component_set S ` S) \ S" + by (simp add: SUP_least connected_component_subset) +qed (use connected_component_refl_eq in force) lemma complement_connected_component_unions: "S - connected_component_set S x = \(connected_component_set S ` S - {connected_component_set S x})" - apply (subst Union_connected_component [symmetric], auto) - apply (metis connected_component_eq_eq connected_component_in) - by (metis connected_component_eq mem_Collect_eq) + (is "?lhs = ?rhs") +proof + show "?lhs \ ?rhs" + by (metis Diff_subset Diff_subset_conv Sup_insert Union_connected_component insert_Diff_single) + show "?rhs \ ?lhs" + by clarsimp (metis connected_component_eq_eq connected_component_in) +qed lemma connected_component_intermediate_subset: "\connected_component_set U a \ T; T \ U\ \ connected_component_set T a = connected_component_set U a" by (metis connected_component_idemp connected_component_mono subset_antisym) - lemma connected_component_homeomorphismI: assumes "homeomorphism A B f g" "connected_component A x y" shows "connected_component B (f x) (f y)" @@ -304,148 +261,121 @@ obtains x where "x \ U" "S = connected_component_set U x" using assms by (auto simp: components_def) -lemma Union_components [simp]: "\(components u) = u" - apply (rule subset_antisym) - using Union_connected_component components_def apply fastforce - apply (metis Union_connected_component components_def set_eq_subset) - done +lemma Union_components [simp]: "\(components U) = U" + by (simp add: Union_connected_component components_def) -lemma pairwise_disjoint_components: "pairwise (\X Y. X \ Y = {}) (components u)" - apply (simp add: pairwise_def) - apply (auto simp: components_iff) - apply (metis connected_component_eq_eq connected_component_in)+ - done +lemma pairwise_disjoint_components: "pairwise (\X Y. X \ Y = {}) (components U)" + unfolding pairwise_def + by (metis (full_types) components_iff connected_component_nonoverlap) -lemma in_components_nonempty: "c \ components s \ c \ {}" +lemma in_components_nonempty: "C \ components S \ C \ {}" by (metis components_iff connected_component_eq_empty) -lemma in_components_subset: "c \ components s \ c \ s" +lemma in_components_subset: "C \ components S \ C \ S" using Union_components by blast -lemma in_components_connected: "c \ components s \ connected c" +lemma in_components_connected: "C \ components S \ connected C" by (metis components_iff connected_connected_component) lemma in_components_maximal: - "c \ components s \ - c \ {} \ c \ s \ connected c \ (\d. d \ {} \ c \ d \ d \ s \ connected d \ d = c)" - apply (rule iffI) - apply (simp add: in_components_nonempty in_components_connected) - apply (metis (full_types) components_iff connected_component_eq_self connected_component_intermediate_subset connected_component_refl in_components_subset mem_Collect_eq rev_subsetD) - apply (metis bot.extremum_uniqueI components_iff connected_component_eq_empty connected_component_maximal connected_component_subset connected_connected_component subset_emptyI) - done + "C \ components S \ + C \ {} \ C \ S \ connected C \ (\d. d \ {} \ C \ d \ d \ S \ connected d \ d = C)" +(is "?lhs \ ?rhs") +proof + assume L: ?lhs + then have "C \ S" "connected C" + by (simp_all add: in_components_subset in_components_connected) + then show ?rhs + by (metis (full_types) L components_iff connected_component_maximal connected_component_refl empty_iff mem_Collect_eq subsetD subset_antisym) +next + show "?rhs \ ?lhs" + by (metis bot.extremum_uniqueI components_iff connected_component_eq_empty connected_component_maximal connected_component_subset connected_connected_component subset_emptyI) +qed + lemma joinable_components_eq: - "connected t \ t \ s \ c1 \ components s \ c2 \ components s \ c1 \ t \ {} \ c2 \ t \ {} \ c1 = c2" + "connected T \ T \ S \ c1 \ components S \ c2 \ components S \ c1 \ T \ {} \ c2 \ T \ {} \ c1 = c2" by (metis (full_types) components_iff joinable_connected_component_eq) -lemma closed_components: "\closed s; c \ components s\ \ closed c" +lemma closed_components: "\closed S; C \ components S\ \ closed C" by (metis closed_connected_component components_iff) lemma components_nonoverlap: - "\c \ components s; c' \ components s\ \ (c \ c' = {}) \ (c \ c')" - apply (auto simp: in_components_nonempty components_iff) - using connected_component_refl apply blast - apply (metis connected_component_eq_eq connected_component_in) - by (metis connected_component_eq mem_Collect_eq) + "\C \ components S; C' \ components S\ \ (C \ C' = {}) \ (C \ C')" + by (metis (full_types) components_iff connected_component_nonoverlap) -lemma components_eq: "\c \ components s; c' \ components s\ \ (c = c' \ c \ c' \ {})" +lemma components_eq: "\C \ components S; C' \ components S\ \ (C = C' \ C \ C' \ {})" by (metis components_nonoverlap) -lemma components_eq_empty [simp]: "components s = {} \ s = {}" +lemma components_eq_empty [simp]: "components S = {} \ S = {}" by (simp add: components_def) lemma components_empty [simp]: "components {} = {}" by simp -lemma connected_eq_connected_components_eq: "connected s \ (\c \ components s. \c' \ components s. c = c')" +lemma connected_eq_connected_components_eq: "connected S \ (\C \ components S. \C' \ components S. C = C')" by (metis (no_types, opaque_lifting) components_iff connected_component_eq_eq connected_iff_connected_component) -lemma components_eq_sing_iff: "components s = {s} \ connected s \ s \ {}" - apply (rule iffI) - using in_components_connected apply fastforce - apply safe - using Union_components apply fastforce - apply (metis components_iff connected_component_eq_self) - using in_components_maximal - apply auto - done +lemma components_eq_sing_iff: "components S = {S} \ connected S \ S \ {}" (is "?lhs \ ?rhs") +proof + show "?rhs \ ?lhs" + by (metis components_iff connected_component_eq_self equals0I insert_iff mk_disjoint_insert) +qed (use in_components_connected in fastforce) -lemma components_eq_sing_exists: "(\a. components s = {a}) \ connected s \ s \ {}" - apply (rule iffI) - using connected_eq_connected_components_eq apply fastforce - apply (metis components_eq_sing_iff) - done +lemma components_eq_sing_exists: "(\a. components S = {a}) \ connected S \ S \ {}" + by (metis Union_components ccpo_Sup_singleton components_eq_sing_iff) -lemma connected_eq_components_subset_sing: "connected s \ components s \ {s}" - by (metis Union_components components_empty components_eq_sing_iff connected_empty insert_subset order_refl subset_singletonD) +lemma connected_eq_components_subset_sing: "connected S \ components S \ {S}" + by (metis components_eq_empty components_eq_sing_iff connected_empty subset_singleton_iff) -lemma connected_eq_components_subset_sing_exists: "connected s \ (\a. components s \ {a})" - by (metis components_eq_sing_exists connected_eq_components_subset_sing empty_iff subset_iff subset_singletonD) +lemma connected_eq_components_subset_sing_exists: "connected S \ (\a. components S \ {a})" + by (metis components_eq_sing_exists connected_eq_components_subset_sing subset_singleton_iff) -lemma in_components_self: "s \ components s \ connected s \ s \ {}" +lemma in_components_self: "S \ components S \ connected S \ S \ {}" by (metis components_empty components_eq_sing_iff empty_iff in_components_connected insertI1) -lemma components_maximal: "\c \ components s; connected t; t \ s; c \ t \ {}\ \ t \ c" - apply (simp add: components_def ex_in_conv [symmetric], clarify) - by (meson connected_component_def connected_component_trans) +lemma components_maximal: "\C \ components S; connected T; T \ S; C \ T \ {}\ \ T \ C" + by (smt (verit, best) Int_Un_eq(4) Un_upper1 bot_eq_sup_iff connected_Un in_components_maximal inf.orderE sup.mono sup.orderI) -lemma exists_component_superset: "\t \ s; s \ {}; connected t\ \ \c. c \ components s \ t \ c" - apply (cases "t = {}", force) - apply (metis components_def ex_in_conv connected_component_maximal contra_subsetD image_eqI) - done +lemma exists_component_superset: "\T \ S; S \ {}; connected T\ \ \C. C \ components S \ T \ C" + by (meson componentsI connected_component_maximal equals0I subset_eq) -lemma components_intermediate_subset: "\s \ components u; s \ t; t \ u\ \ s \ components t" - apply (auto simp: components_iff) - apply (metis connected_component_eq_empty connected_component_intermediate_subset) - done +lemma components_intermediate_subset: "\S \ components U; S \ T; T \ U\ \ S \ components T" + by (smt (verit, best) dual_order.trans in_components_maximal) -lemma in_components_unions_complement: "c \ components s \ s - c = \(components s - {c})" +lemma in_components_unions_complement: "C \ components S \ S - C = \(components S - {C})" by (metis complement_connected_component_unions components_def components_iff) lemma connected_intermediate_closure: - assumes cs: "connected s" and st: "s \ t" and ts: "t \ closure s" - shows "connected t" -proof (rule connectedI) - fix A B - assume A: "open A" and B: "open B" and Alap: "A \ t \ {}" and Blap: "B \ t \ {}" - and disj: "A \ B \ t = {}" and cover: "t \ A \ B" - have disjs: "A \ B \ s = {}" - using disj st by auto - have "A \ closure s \ {}" - using Alap Int_absorb1 ts by blast - then have Alaps: "A \ s \ {}" - by (simp add: A open_Int_closure_eq_empty) - have "B \ closure s \ {}" - using Blap Int_absorb1 ts by blast - then have Blaps: "B \ s \ {}" - by (simp add: B open_Int_closure_eq_empty) - then show False - using cs [unfolded connected_def] A B disjs Alaps Blaps cover st - by blast -qed + assumes cs: "connected S" and st: "S \ T" and ts: "T \ closure S" + shows "connected T" + using assms unfolding connected_def + by (smt (verit) Int_assoc inf.absorb_iff2 inf_bot_left open_Int_closure_eq_empty) -lemma closedin_connected_component: "closedin (top_of_set s) (connected_component_set s x)" -proof (cases "connected_component_set s x = {}") +lemma closedin_connected_component: "closedin (top_of_set S) (connected_component_set S x)" +proof (cases "connected_component_set S x = {}") case True then show ?thesis by (metis closedin_empty) next case False - then obtain y where y: "connected_component s x y" - by blast - have *: "connected_component_set s x \ s \ closure (connected_component_set s x)" + then obtain y where y: "connected_component S x y" and "x \ S" + using connected_component_eq_empty by blast + have *: "connected_component_set S x \ S \ closure (connected_component_set S x)" by (auto simp: closure_def connected_component_in) - have "connected_component s x y \ s \ closure (connected_component_set s x) \ connected_component_set s x" - apply (rule connected_component_maximal, simp) - using closure_subset connected_component_in apply fastforce - using * connected_intermediate_closure apply blast+ - done + have **: "x \ closure (connected_component_set S x)" + by (simp add: \x \ S\ closure_def) + have "S \ closure (connected_component_set S x) \ connected_component_set S x" if "connected_component S x y" + proof (rule connected_component_maximal) + show "connected (S \ closure (connected_component_set S x))" + using "*" connected_intermediate_closure by blast + qed (use \x \ S\ ** in auto) with y * show ?thesis by (auto simp: closedin_closed) qed lemma closedin_component: - "C \ components s \ closedin (top_of_set s) C" + "C \ components S \ closedin (top_of_set S) C" using closedin_connected_component componentsE by blast @@ -454,37 +384,34 @@ lemma continuous_levelset_openin_cases: fixes f :: "_ \ 'b::t1_space" - shows "connected s \ continuous_on s f \ - openin (top_of_set s) {x \ s. f x = a} - \ (\x \ s. f x \ a) \ (\x \ s. f x = a)" + shows "connected S \ continuous_on S f \ + openin (top_of_set S) {x \ S. f x = a} + \ (\x \ S. f x \ a) \ (\x \ S. f x = a)" unfolding connected_clopen using continuous_closedin_preimage_constant by auto lemma continuous_levelset_openin: fixes f :: "_ \ 'b::t1_space" - shows "connected s \ continuous_on s f \ - openin (top_of_set s) {x \ s. f x = a} \ - (\x \ s. f x = a) \ (\x \ s. f x = a)" - using continuous_levelset_openin_cases[of s f ] + shows "connected S \ continuous_on S f \ + openin (top_of_set S) {x \ S. f x = a} \ + (\x \ S. f x = a) \ (\x \ S. f x = a)" + using continuous_levelset_openin_cases[of S f ] by meson lemma continuous_levelset_open: fixes f :: "_ \ 'b::t1_space" - assumes "connected s" - and "continuous_on s f" - and "open {x \ s. f x = a}" - and "\x \ s. f x = a" - shows "\x \ s. f x = a" - using continuous_levelset_openin[OF assms(1,2), of a, unfolded openin_open] - using assms (3,4) + assumes S: "connected S" "continuous_on S f" + and a: "open {x \ S. f x = a}" "\x \ S. f x = a" + shows "\x \ S. f x = a" + using a continuous_levelset_openin[OF S, of a, unfolded openin_open] by fast subsection\<^marker>\tag unimportant\ \Preservation of Connectedness\ lemma homeomorphic_connectedness: - assumes "s homeomorphic t" - shows "connected s \ connected t" + assumes "S homeomorphic T" + shows "connected S \ connected T" using assms unfolding homeomorphic_def homeomorphism_def by (metis connected_continuous_image) lemma connected_monotone_quotient_preimage: @@ -539,14 +466,7 @@ show ?thesis proof (rule connected_monotone_quotient_preimage [OF \connected C\ contf' eqC]) show "connected (S \ f -` C \ f -` {y})" if "y \ C" for y - proof - - have "S \ f -` C \ f -` {y} = S \ f -` {y}" - using that by blast - moreover have "connected (S \ f -` {y})" - using \C \ T\ connT that by blast - ultimately show ?thesis - by metis - qed + by (metis Int_assoc Int_empty_right Int_insert_right_if1 assms(6) connT in_mono that vimage_Int) have "\U. openin (top_of_set (S \ f -` C)) U \ openin (top_of_set C) (f ` U)" using open_map_restrict [OF _ ST \C \ T\] by metis @@ -572,14 +492,7 @@ show ?thesis proof (rule connected_monotone_quotient_preimage [OF \connected C\ contf' eqC]) show "connected (S \ f -` C \ f -` {y})" if "y \ C" for y - proof - - have "S \ f -` C \ f -` {y} = S \ f -` {y}" - using that by blast - moreover have "connected (S \ f -` {y})" - using \C \ T\ connT that by blast - ultimately show ?thesis - by metis - qed + by (metis Int_assoc Int_empty_right Int_insert_right_if1 \C \ T\ connT subsetD that vimage_Int) have "\U. closedin (top_of_set (S \ f -` C)) U \ closedin (top_of_set C) (f ` U)" using closed_map_restrict [OF _ ST \C \ T\] by metis @@ -615,15 +528,12 @@ then have clo: "closedin (top_of_set S) (S \ H1)" "closedin (top_of_set S) (S \ H2)" by (metis Un_upper1 closedin_closed_subset inf_commute)+ - have Seq: "S \ (H1 \ H2) = S" - by (simp add: H) - have "S \ ((S \ T) \ H1) \ S \ ((S \ T) \ H2) = S" - using Seq by auto + moreover have "S \ ((S \ T) \ H1) \ S \ ((S \ T) \ H2) = S" + using H by blast moreover have "H1 \ (S \ ((S \ T) \ H2)) = {}" using H by blast ultimately have "S \ H1 = {} \ S \ H2 = {}" - by (metis (no_types) H Int_assoc \S \ (H1 \ H2) = S\ \connected S\ - clo Seq connected_closedin inf_bot_right inf_le1) + by (smt (verit) Int_assoc \connected S\ connected_closedin_eq inf_commute inf_sup_absorb) then show "S \ H1 \ S \ H2" using H \connected S\ unfolding connected_closedin by blast next @@ -641,8 +551,7 @@ have "openin (top_of_set ((U - S) \ (S \ T))) H2" proof (rule openin_subtopology_Un) show "openin (top_of_set (S \ T)) H2" - using \H2 \ T\ apply (auto simp: openin_closedin_eq) - by (metis Diff_Diff_Int Diff_disjoint Diff_partition Diff_subset H Int_absorb1 Un_Diff) + by (metis Diff_cancel H Un_Diff Un_Diff_Int closedin_subset openin_closedin_eq topspace_euclidean_subtopology) then show "openin (top_of_set (U - S)) H2" by (meson H2T Un_upper2 opeT openin_subset_trans openin_trans) qed @@ -652,8 +561,7 @@ using H H2T cloT closedin_subset_trans by (blast intro: closedin_subtopology_Un closedin_trans) qed (simp add: H) - ultimately - have H2: "H2 = {} \ H2 = U" + ultimately have H2: "H2 = {} \ H2 = U" using Ueq \connected U\ unfolding connected_clopen by metis then have "H2 \ S" by (metis Diff_partition H Un_Diff_cancel Un_subset_iff \H2 \ T\ assms(3) inf.orderE opeT openin_imp_subset) @@ -674,7 +582,7 @@ fix H3 H4 assume clo3: "closedin (top_of_set (U - C)) H3" and clo4: "closedin (top_of_set (U - C)) H4" - and "H3 \ H4 = U - C" and "H3 \ H4 = {}" and "H3 \ {}" and "H4 \ {}" + and H34: "H3 \ H4 = U - C" "H3 \ H4 = {}" and "H3 \ {}" and "H4 \ {}" and * [rule_format]: "\H1 H2. \ closedin (top_of_set S) H1 \ \ closedin (top_of_set S) H2 \ @@ -687,16 +595,12 @@ have cCH3: "connected (C \ H3)" proof (rule connected_Un_clopen_in_complement [OF \connected C\ \connected U\ _ _ clo3]) show "openin (top_of_set (U - C)) H3" - apply (simp add: openin_closedin_eq \H3 \ U - C\) - apply (simp add: closedin_subtopology) - by (metis Diff_cancel Diff_triv Un_Diff clo4 \H3 \ H4 = {}\ \H3 \ H4 = U - C\ closedin_closed inf_commute sup_bot.left_neutral) + by (metis Diff_cancel Un_Diff Un_Diff_Int \H3 \ H4 = {}\ \H3 \ H4 = U - C\ ope4) qed (use clo3 \C \ U - S\ in auto) have cCH4: "connected (C \ H4)" proof (rule connected_Un_clopen_in_complement [OF \connected C\ \connected U\ _ _ clo4]) show "openin (top_of_set (U - C)) H4" - apply (simp add: openin_closedin_eq \H4 \ U - C\) - apply (simp add: closedin_subtopology) - by (metis Diff_cancel Int_commute Un_Diff Un_Diff_Int \H3 \ H4 = {}\ \H3 \ H4 = U - C\ clo3 closedin_closed) + by (metis Diff_cancel Diff_triv Int_Un_eq(2) Un_Diff H34 inf_commute ope3) qed (use clo4 \C \ U - S\ in auto) have "closedin (top_of_set S) (S \ H3)" "closedin (top_of_set S) (S \ H4)" using clo3 clo4 \S \ U\ \C \ U - S\ by (auto simp: closedin_closed) @@ -717,18 +621,17 @@ lemma continuous_disconnected_range_constant: assumes S: "connected S" and conf: "continuous_on S f" - and fim: "f ` S \ t" - and cct: "\y. y \ t \ connected_component_set t y = {y}" + and fim: "f \ S \ T" + and cct: "\y. y \ T \ connected_component_set T y = {y}" shows "f constant_on S" proof (cases "S = {}") case True then show ?thesis by (simp add: constant_on_def) next case False - { fix x assume "x \ S" - then have "f ` S \ {f x}" - by (metis connected_continuous_image conf connected_component_maximal fim image_subset_iff rev_image_eqI S cct) - } + then have "f ` S \ {f x}" if "x \ S" for x + by (metis PiE S cct connected_component_maximal connected_continuous_image [OF conf] fim image_eqI + image_subset_iff that) with False show ?thesis unfolding constant_on_def by blast qed @@ -740,19 +643,17 @@ \continuous_on S f; finite(f ` S)\ \ f constant_on S" shows "connected S" proof - - { fix t u - assume clt: "closedin (top_of_set S) t" - and clu: "closedin (top_of_set S) u" - and tue: "t \ u = {}" and tus: "t \ u = S" - have conif: "continuous_on S (\x. if x \ t then 0 else 1)" - apply (subst tus [symmetric]) - apply (rule continuous_on_cases_local) - using clt clu tue - apply (auto simp: tus) - done - have fi: "finite ((\x. if x \ t then 0 else 1) ` S)" + { fix T U + assume clt: "closedin (top_of_set S) T" + and clu: "closedin (top_of_set S) U" + and tue: "T \ U = {}" and tus: "T \ U = S" + have "continuous_on (T \ U) (\x. if x \ T then 0 else 1)" + using clt clu tue by (intro continuous_on_cases_local) (auto simp: tus) + then have conif: "continuous_on S (\x. if x \ T then 0 else 1)" + using tus by blast + have fi: "finite ((\x. if x \ T then 0 else 1) ` S)" by (rule finite_subset [of _ "{0,1}"]) auto - have "t = {} \ u = {}" + have "T = {} \ U = {}" using assms [OF conif fi] tus [symmetric] by (auto simp: Ball_def constant_on_def) (metis IntI empty_iff one_neq_zero tue) } diff -r ba2afdd29e1d -r 032a4344903e src/HOL/Analysis/Convex.thy --- a/src/HOL/Analysis/Convex.thy Tue Aug 01 11:27:55 2023 +0200 +++ b/src/HOL/Analysis/Convex.thy Thu Aug 03 19:10:43 2023 +0200 @@ -22,7 +22,7 @@ lemma convexI: assumes "\x y u v. x \ s \ y \ s \ 0 \ u \ 0 \ v \ u + v = 1 \ u *\<^sub>R x + v *\<^sub>R y \ s" shows "convex s" - using assms unfolding convex_def by fast + by (simp add: assms convex_def) lemma convexD: assumes "convex s" and "x \ s" and "y \ s" and "0 \ u" and "0 \ v" and "u + v = 1" @@ -31,25 +31,7 @@ lemma convex_alt: "convex s \ (\x\s. \y\s. \u. 0 \ u \ u \ 1 \ ((1 - u) *\<^sub>R x + u *\<^sub>R y) \ s)" (is "_ \ ?alt") -proof - show "convex s" if alt: ?alt - proof - - { - fix x y and u v :: real - assume mem: "x \ s" "y \ s" - assume "0 \ u" "0 \ v" - moreover - assume "u + v = 1" - then have "u = 1 - v" by auto - ultimately have "u *\<^sub>R x + v *\<^sub>R y \ s" - using alt [rule_format, OF mem] by auto - } - then show ?thesis - unfolding convex_def by auto - qed - show ?alt if "convex s" - using that by (auto simp: convex_def) -qed + by (smt (verit) convexD convexI) lemma convexD_alt: assumes "convex s" "a \ s" "b \ s" "0 \ u" "u \ 1" @@ -193,44 +175,31 @@ fixes C :: "'a::real_vector set" assumes "finite S" and "convex C" - and "(\ i \ S. a i) = 1" - assumes "\i. i \ S \ a i \ 0" - and "\i. i \ S \ y i \ C" + and a: "(\ i \ S. a i) = 1" "\i. i \ S \ a i \ 0" + and C: "\i. i \ S \ y i \ C" shows "(\ j \ S. a j *\<^sub>R y j) \ C" - using assms(1,3,4,5) -proof (induct arbitrary: a set: finite) + using \finite S\ a C +proof (induction arbitrary: a set: finite) case empty then show ?case by simp next - case (insert i S) note IH = this(3) - have "a i + sum a S = 1" - and "0 \ a i" - and "\j\S. 0 \ a j" - and "y i \ C" - and "\j\S. y j \ C" - using insert.hyps(1,2) insert.prems by simp_all + case (insert i S) then have "0 \ sum a S" by (simp add: sum_nonneg) have "a i *\<^sub>R y i + (\j\S. a j *\<^sub>R y j) \ C" proof (cases "sum a S = 0") - case True - with \a i + sum a S = 1\ have "a i = 1" - by simp - from sum_nonneg_0 [OF \finite S\ _ True] \\j\S. 0 \ a j\ have "\j\S. a j = 0" - by simp - show ?thesis using \a i = 1\ and \\j\S. a j = 0\ and \y i \ C\ - by simp + case True with insert show ?thesis + by (simp add: sum_nonneg_eq_0_iff) next case False with \0 \ sum a S\ have "0 < sum a S" by simp then have "(\j\S. (a j / sum a S) *\<^sub>R y j) \ C" - using \\j\S. 0 \ a j\ and \\j\S. y j \ C\ - by (simp add: IH sum_divide_distrib [symmetric]) - from \convex C\ and \y i \ C\ and this and \0 \ a i\ - and \0 \ sum a S\ and \a i + sum a S = 1\ + using insert + by (simp add: insert.IH flip: sum_divide_distrib) + with \convex C\ insert \0 \ sum a S\ have "a i *\<^sub>R y i + sum a S *\<^sub>R (\j\S. (a j / sum a S) *\<^sub>R y j) \ C" - by (rule convexD) + by (simp add: convex_def) then show ?thesis by (simp add: scaleR_sum_right False) qed @@ -239,18 +208,13 @@ qed lemma convex: - "convex S \ (\(k::nat) u x. (\i. 1\i \ i\k \ 0 \ u i \ x i \S) \ (sum u {1..k} = 1) - \ sum (\i. u i *\<^sub>R x i) {1..k} \ S)" -proof safe - fix k :: nat - fix u :: "nat \ real" - fix x - assume "convex S" - "\i. 1 \ i \ i \ k \ 0 \ u i \ x i \ S" - "sum u {1..k} = 1" - with convex_sum[of "{1 .. k}" S] show "(\j\{1 .. k}. u j *\<^sub>R x j) \ S" - by auto -next + "convex S \ + (\(k::nat) u x. (\i. 1\i \ i\k \ 0 \ u i \ x i \S) \ (sum u {1..k} = 1) + \ sum (\i. u i *\<^sub>R x i) {1..k} \ S)" + (is "?lhs = ?rhs") +proof + show "?lhs \ ?rhs" + by (metis (full_types) atLeastAtMost_iff convex_sum finite_atLeastAtMost) assume *: "\k u x. (\ i :: nat. 1 \ i \ i \ k \ 0 \ u i \ x i \ S) \ sum u {1..k} = 1 \ (\i = 1..k. u i *\<^sub>R (x i :: 'a)) \ S" { @@ -262,18 +226,14 @@ let ?x = "\i. if (i :: nat) = 1 then x else y" have "{1 :: nat .. 2} \ - {x. x = 1} = {2}" by auto - then have card: "card ({1 :: nat .. 2} \ - {x. x = 1}) = 1" - by simp - then have "sum ?u {1 .. 2} = 1" - using sum.If_cases[of "{(1 :: nat) .. 2}" "\ x. x = 1" "\ x. \" "\ x. 1 - \"] - by auto - with *[rule_format, of "2" ?u ?x] have S: "(\j \ {1..2}. ?u j *\<^sub>R ?x j) \ S" - using mu xy by auto + then have S: "(\j \ {1..2}. ?u j *\<^sub>R ?x j) \ S" + using sum.If_cases[of "{(1 :: nat) .. 2}" "\x. x = 1" "\x. \" "\x. 1 - \"] + using mu xy "*" by auto have grarr: "(\j \ {Suc (Suc 0)..2}. ?u j *\<^sub>R ?x j) = (1 - \) *\<^sub>R y" using sum.atLeast_Suc_atMost[of "Suc (Suc 0)" 2 "\ j. (1 - \) *\<^sub>R y"] by auto - from sum.atLeast_Suc_atMost[of "Suc 0" 2 "\ j. ?u j *\<^sub>R ?x j", simplified this] + with sum.atLeast_Suc_atMost have "(\j \ {1..2}. ?u j *\<^sub>R ?x j) = \ *\<^sub>R x + (1 - \) *\<^sub>R y" - by auto + by (smt (verit, best) Suc_1 Suc_eq_plus1 add_0 le_add1) then have "(1 - \) *\<^sub>R y + \ *\<^sub>R x \ S" using S by (auto simp: add.commute) } @@ -293,7 +253,7 @@ and "finite t" and "t \ S" "\x\t. 0 \ u x" "sum u t = 1" then show "(\x\t. u x *\<^sub>R x) \ S" - using convex_sum[of t S u "\ x. x"] by auto + by (simp add: convex_sum subsetD) next assume *: "\t. \ u. finite t \ t \ S \ (\x\t. 0 \ u x) \ sum u t = 1 \ (\x\t. u x *\<^sub>R x) \ S" @@ -330,7 +290,8 @@ assume *: "\x\T. 0 \ u x" "sum u T = 1" assume "T \ S" then have "S \ T = T" by auto - with sum[THEN spec[where x="\x. if x\T then u x else 0"]] * have "(\x\T. u x *\<^sub>R x) \ S" + with sum[THEN spec[where x="\x. if x\T then u x else 0"]] * + have "(\x\T. u x *\<^sub>R x) \ S" by (auto simp: assms sum.If_cases if_distrib if_distrib_arg) } moreover assume ?rhs ultimately show ?lhs @@ -357,30 +318,14 @@ f ((1 - t) *\<^sub>R x + t *\<^sub>R y) \ (1 - t) * f x + t * f y" shows "convex_on A f" unfolding convex_on_def -proof clarify - fix x y - fix u v :: real - assume A: "x \ A" "y \ A" "u \ 0" "v \ 0" "u + v = 1" - from A(5) have [simp]: "v = 1 - u" - by (simp add: algebra_simps) - from A(1-4) show "f (u *\<^sub>R x + v *\<^sub>R y) \ u * f x + v * f y" - using assms[of u y x] - by (cases "u = 0 \ u = 1") (auto simp: algebra_simps) -qed + by (smt (verit, del_insts) assms mult_cancel_right1 mult_eq_0_iff scaleR_collapse scaleR_eq_0_iff) lemma convex_on_linorderI [intro?]: fixes A :: "('a::{linorder,real_vector}) set" assumes "\t x y. t > 0 \ t < 1 \ x \ A \ y \ A \ x < y \ f ((1 - t) *\<^sub>R x + t *\<^sub>R y) \ (1 - t) * f x + t * f y" shows "convex_on A f" -proof - fix x y - fix t :: real - assume A: "x \ A" "y \ A" "t > 0" "t < 1" - with assms [of t x y] assms [of "1 - t" y x] - show "f ((1 - t) *\<^sub>R x + t *\<^sub>R y) \ (1 - t) * f x + t * f y" - by (cases x y rule: linorder_cases) (auto simp: algebra_simps) -qed + by (smt (verit, best) add.commute assms convex_onI distrib_left linorder_cases mult.commute mult_cancel_left2 scaleR_collapse) lemma convex_onD: assumes "convex_on A f" @@ -452,7 +397,7 @@ lemma convex_on_dist [intro]: fixes S :: "'a::real_normed_vector set" shows "convex_on S (\x. dist a x)" -proof (auto simp: convex_on_def dist_norm) +proof (clarsimp simp: convex_on_def dist_norm) fix x y assume "x \ S" "y \ S" fix u v :: real @@ -460,20 +405,18 @@ assume "0 \ v" assume "u + v = 1" have "a = u *\<^sub>R a + v *\<^sub>R a" - unfolding scaleR_left_distrib[symmetric] and \u + v = 1\ by simp - then have *: "a - (u *\<^sub>R x + v *\<^sub>R y) = (u *\<^sub>R (a - x)) + (v *\<^sub>R (a - y))" + by (metis \u + v = 1\ scaleR_left.add scaleR_one) + then have "a - (u *\<^sub>R x + v *\<^sub>R y) = (u *\<^sub>R (a - x)) + (v *\<^sub>R (a - y))" by (auto simp: algebra_simps) - show "norm (a - (u *\<^sub>R x + v *\<^sub>R y)) \ u * norm (a - x) + v * norm (a - y)" - unfolding * using norm_triangle_ineq[of "u *\<^sub>R (a - x)" "v *\<^sub>R (a - y)"] - using \0 \ u\ \0 \ v\ by auto + then show "norm (a - (u *\<^sub>R x + v *\<^sub>R y)) \ u * norm (a - x) + v * norm (a - y)" + by (smt (verit, best) \0 \ u\ \0 \ v\ norm_scaleR norm_triangle_ineq) qed subsection\<^marker>\tag unimportant\ \Arithmetic operations on sets preserve convexity\ lemma convex_linear_image: - assumes "linear f" - and "convex S" + assumes "linear f" and "convex S" shows "convex (f ` S)" proof - interpret f: linear f by fact @@ -482,8 +425,7 @@ qed lemma convex_linear_vimage: - assumes "linear f" - and "convex S" + assumes "linear f" and "convex S" shows "convex (f -` S)" proof - interpret f: linear f by fact @@ -494,32 +436,17 @@ lemma convex_scaling: assumes "convex S" shows "convex ((\x. c *\<^sub>R x) ` S)" -proof - - have "linear (\x. c *\<^sub>R x)" - by (simp add: linearI scaleR_add_right) - then show ?thesis - using \convex S\ by (rule convex_linear_image) -qed + by (simp add: assms convex_linear_image) lemma convex_scaled: assumes "convex S" shows "convex ((\x. x *\<^sub>R c) ` S)" -proof - - have "linear (\x. x *\<^sub>R c)" - by (simp add: linearI scaleR_add_left) - then show ?thesis - using \convex S\ by (rule convex_linear_image) -qed + by (simp add: assms convex_linear_image) lemma convex_negations: assumes "convex S" shows "convex ((\x. - x) ` S)" -proof - - have "linear (\x. - x)" - by (simp add: linearI) - then show ?thesis - using \convex S\ by (rule convex_linear_image) -qed + by (simp add: assms convex_linear_image module_hom_uminus) lemma convex_sums: assumes "convex S" @@ -572,124 +499,79 @@ fixes a :: "'a \ real" and y :: "'a \ 'b::real_vector" and f :: "'b \ real" - assumes "finite s" "s \ {}" + assumes "finite S" "S \ {}" and "convex_on C f" and "convex C" - and "(\ i \ s. a i) = 1" - and "\i. i \ s \ a i \ 0" - and "\i. i \ s \ y i \ C" - shows "f (\ i \ s. a i *\<^sub>R y i) \ (\ i \ s. a i * f (y i))" + and "(\ i \ S. a i) = 1" + and "\i. i \ S \ a i \ 0" + and "\i. i \ S \ y i \ C" + shows "f (\ i \ S. a i *\<^sub>R y i) \ (\ i \ S. a i * f (y i))" using assms -proof (induct s arbitrary: a rule: finite_ne_induct) +proof (induct S arbitrary: a rule: finite_ne_induct) case (singleton i) - then have ai: "a i = 1" - by auto then show ?case by auto next - case (insert i s) - then have "convex_on C f" - by simp - from this[unfolded convex_on_def, rule_format] - have conv: "\x y \. x \ C \ y \ C \ 0 \ \ \ \ \ 1 \ + case (insert i S) + then have yai: "y i \ C" "a i \ 0" + by auto + with insert have conv: "\x y \. x \ C \ y \ C \ 0 \ \ \ \ \ 1 \ f (\ *\<^sub>R x + (1 - \) *\<^sub>R y) \ \ * f x + (1 - \) * f y" - by simp + by (simp add: convex_on_def) show ?case proof (cases "a i = 1") case True - then have "(\ j \ s. a j) = 0" - using insert by auto - then have "\j. j \ s \ a j = 0" - using insert by (fastforce simp: sum_nonneg_eq_0_iff) - then show ?thesis - using insert by auto + with insert have "(\ j \ S. a j) = 0" + by auto + with insert show ?thesis + by (simp add: sum_nonneg_eq_0_iff) next case False - from insert have yai: "y i \ C" "a i \ 0" - by auto - have fis: "finite (insert i s)" - using insert by auto - then have ai1: "a i \ 1" - using sum_nonneg_leq_bound[of "insert i s" a] insert by simp - then have "a i < 1" - using False by auto + then have ai1: "a i < 1" + using sum_nonneg_leq_bound[of "insert i S" a] insert by force then have i0: "1 - a i > 0" by auto let ?a = "\j. a j / (1 - a i)" - have a_nonneg: "?a j \ 0" if "j \ s" for j + have a_nonneg: "?a j \ 0" if "j \ S" for j using i0 insert that by fastforce - have "(\ j \ insert i s. a j) = 1" + have "(\ j \ insert i S. a j) = 1" using insert by auto - then have "(\ j \ s. a j) = 1 - a i" + then have "(\ j \ S. a j) = 1 - a i" using sum.insert insert by fastforce - then have "(\ j \ s. a j) / (1 - a i) = 1" - using i0 by auto - then have a1: "(\ j \ s. ?a j) = 1" - unfolding sum_divide_distrib by simp - have "convex C" using insert by auto - then have asum: "(\ j \ s. ?a j *\<^sub>R y j) \ C" - using insert convex_sum [OF \finite s\ \convex C\ a1 a_nonneg] by auto - have asum_le: "f (\ j \ s. ?a j *\<^sub>R y j) \ (\ j \ s. ?a j * f (y j))" - using a_nonneg a1 insert by blast - have "f (\ j \ insert i s. a j *\<^sub>R y j) = f ((\ j \ s. a j *\<^sub>R y j) + a i *\<^sub>R y i)" - using sum.insert[of s i "\ j. a j *\<^sub>R y j", OF \finite s\ \i \ s\] insert - by (auto simp only: add.commute) - also have "\ = f (((1 - a i) * inverse (1 - a i)) *\<^sub>R (\ j \ s. a j *\<^sub>R y j) + a i *\<^sub>R y i)" + then have "(\ j \ S. a j) / (1 - a i) = 1" using i0 by auto - also have "\ = f ((1 - a i) *\<^sub>R (\ j \ s. (a j * inverse (1 - a i)) *\<^sub>R y j) + a i *\<^sub>R y i)" - using scaleR_right.sum[of "inverse (1 - a i)" "\ j. a j *\<^sub>R y j" s, symmetric] + then have a1: "(\ j \ S. ?a j) = 1" + unfolding sum_divide_distrib by simp + have asum: "(\ j \ S. ?a j *\<^sub>R y j) \ C" + using insert convex_sum [OF \finite S\ \convex C\ a1 a_nonneg] by auto + have asum_le: "f (\ j \ S. ?a j *\<^sub>R y j) \ (\ j \ S. ?a j * f (y j))" + using a_nonneg a1 insert by blast + have "f (\ j \ insert i S. a j *\<^sub>R y j) = f ((\ j \ S. a j *\<^sub>R y j) + a i *\<^sub>R y i)" + by (simp add: add.commute insert.hyps) + also have "\ = f (((1 - a i) * inverse (1 - a i)) *\<^sub>R (\ j \ S. a j *\<^sub>R y j) + a i *\<^sub>R y i)" + using i0 by auto + also have "\ = f ((1 - a i) *\<^sub>R (\ j \ S. (a j * inverse (1 - a i)) *\<^sub>R y j) + a i *\<^sub>R y i)" + using scaleR_right.sum[of "inverse (1 - a i)" "\ j. a j *\<^sub>R y j" S, symmetric] by (auto simp: algebra_simps) - also have "\ = f ((1 - a i) *\<^sub>R (\ j \ s. ?a j *\<^sub>R y j) + a i *\<^sub>R y i)" + also have "\ = f ((1 - a i) *\<^sub>R (\ j \ S. ?a j *\<^sub>R y j) + a i *\<^sub>R y i)" by (auto simp: divide_inverse) - also have "\ \ (1 - a i) *\<^sub>R f ((\ j \ s. ?a j *\<^sub>R y j)) + a i * f (y i)" - using conv[of "y i" "(\ j \ s. ?a j *\<^sub>R y j)" "a i", OF yai(1) asum yai(2) ai1] - by (auto simp: add.commute) - also have "\ \ (1 - a i) * (\ j \ s. ?a j * f (y j)) + a i * f (y i)" - using add_right_mono [OF mult_left_mono [of _ _ "1 - a i", - OF asum_le less_imp_le[OF i0]], of "a i * f (y i)"] - by simp - also have "\ = (\ j \ s. (1 - a i) * ?a j * f (y j)) + a i * f (y i)" - unfolding sum_distrib_left[of "1 - a i" "\ j. ?a j * f (y j)"] - using i0 by auto - also have "\ = (\ j \ s. a j * f (y j)) + a i * f (y i)" - using i0 by auto - also have "\ = (\ j \ insert i s. a j * f (y j))" + also have "\ \ (1 - a i) *\<^sub>R f ((\ j \ S. ?a j *\<^sub>R y j)) + a i * f (y i)" + using ai1 by (smt (verit) asum conv real_scaleR_def yai) + also have "\ \ (1 - a i) * (\ j \ S. ?a j * f (y j)) + a i * f (y i)" + using asum_le i0 by fastforce + also have "\ = (\ j \ S. a j * f (y j)) + a i * f (y i)" + using i0 by (auto simp: sum_distrib_left) + finally show ?thesis using insert by auto - finally show ?thesis - by simp qed qed lemma convex_on_alt: fixes C :: "'a::real_vector set" shows "convex_on C f \ - (\x \ C. \ y \ C. \ \ :: real. \ \ 0 \ \ \ 1 \ - f (\ *\<^sub>R x + (1 - \) *\<^sub>R y) \ \ * f x + (1 - \) * f y)" -proof safe - fix x y - fix \ :: real - assume *: "convex_on C f" "x \ C" "y \ C" "0 \ \" "\ \ 1" - from this[unfolded convex_on_def, rule_format] - have "0 \ u \ 0 \ v \ u + v = 1 \ f (u *\<^sub>R x + v *\<^sub>R y) \ u * f x + v * f y" for u v - by auto - from this [of "\" "1 - \", simplified] * - show "f (\ *\<^sub>R x + (1 - \) *\<^sub>R y) \ \ * f x + (1 - \) * f y" - by auto -next - assume *: "\x\C. \y\C. \\. 0 \ \ \ \ \ 1 \ - f (\ *\<^sub>R x + (1 - \) *\<^sub>R y) \ \ * f x + (1 - \) * f y" - { - fix x y - fix u v :: real - assume **: "x \ C" "y \ C" "u \ 0" "v \ 0" "u + v = 1" - then have[simp]: "1 - u = v" by auto - from *[rule_format, of x y u] - have "f (u *\<^sub>R x + v *\<^sub>R y) \ u * f x + v * f y" - using ** by auto - } - then show "convex_on C f" - unfolding convex_on_def by auto -qed + (\x \ C. \y \ C. \ \ :: real. \ \ 0 \ \ \ 1 \ + f (\ *\<^sub>R x + (1 - \) *\<^sub>R y) \ \ * f x + (1 - \) * f y)" + by (smt (verit) convex_on_def) lemma convex_on_diff: fixes f :: "real \ real" @@ -784,12 +666,10 @@ shows "f' x * (y - x) \ f y - f x" using assms proof - - have less_imp: "f y - f x \ f' x * (y - x)" "f' y * (x - y) \ f x - f y" + have "f y - f x \ f' x * (y - x)" "f' y * (x - y) \ f x - f y" if *: "x \ C" "y \ C" "y > x" for x y :: real proof - - from * have ge: "y - x > 0" "y - x \ 0" - by auto - from * have le: "x - y < 0" "x - y \ 0" + from * have ge: "y - x > 0" "y - x \ 0" and le: "x - y < 0" "x - y \ 0" by auto then obtain z1 where z1: "z1 > x" "z1 < y" "f y - f x = (y - x) * f' z1" using subsetD[OF atMostAtLeast_subset_convex[OF \convex C\ \x \ C\ \y \ C\ \x < y\], @@ -798,8 +678,6 @@ then have "z1 \ C" using atMostAtLeast_subset_convex \convex C\ \x \ C\ \y \ C\ \x < y\ by fastforce - from z1 have z1': "f x - f y = (x - y) * f' z1" - by (simp add: field_simps) obtain z2 where z2: "z2 > x" "z2 < z1" "f' z1 - f' x = (z1 - x) * f'' z2" using subsetD[OF atMostAtLeast_subset_convex[OF \convex C\ \x \ C\ \z1 \ C\ \x < z1\], THEN f'', THEN MVT2[OF \x < z1\, rule_format, unfolded atLeastAtMost_iff[symmetric]]] z1 @@ -808,75 +686,41 @@ using subsetD[OF atMostAtLeast_subset_convex[OF \convex C\ \z1 \ C\ \y \ C\ \z1 < y\], THEN f'', THEN MVT2[OF \z1 < y\, rule_format, unfolded atLeastAtMost_iff[symmetric]]] z1 by auto - have "f' y - (f x - f y) / (x - y) = f' y - f' z1" - using * z1' by auto - also have "\ = (y - z1) * f'' z3" - using z3 by auto - finally have cool': "f' y - (f x - f y) / (x - y) = (y - z1) * f'' z3" - by simp - have A': "y - z1 \ 0" - using z1 by auto + from z1 have "f x - f y = (x - y) * f' z1" + by (simp add: field_simps) + then have cool': "f' y - (f x - f y) / (x - y) = (y - z1) * f'' z3" + using le(1) z3(3) by auto have "z3 \ C" using z3 * atMostAtLeast_subset_convex \convex C\ \x \ C\ \z1 \ C\ \x < z1\ by fastforce then have B': "f'' z3 \ 0" using assms by auto - from A' B' have "(y - z1) * f'' z3 \ 0" - by auto - from cool' this have "f' y - (f x - f y) / (x - y) \ 0" - by auto - from mult_right_mono_neg[OF this le(2)] - have "f' y * (x - y) - (f x - f y) / (x - y) * (x - y) \ 0 * (x - y)" - by (simp add: algebra_simps) - then have "f' y * (x - y) - (f x - f y) \ 0" - using le by auto + with cool' have "f' y - (f x - f y) / (x - y) \ 0" + using z1 by auto then have res: "f' y * (x - y) \ f x - f y" - by auto - have "(f y - f x) / (y - x) - f' x = f' z1 - f' x" - using * z1 by auto - also have "\ = (z1 - x) * f'' z2" - using z2 by auto - finally have cool: "(f y - f x) / (y - x) - f' x = (z1 - x) * f'' z2" - by simp - have A: "z1 - x \ 0" - using z1 by auto + by (meson diff_ge_0_iff_ge le(1) neg_divide_le_eq) + have cool: "(f y - f x) / (y - x) - f' x = (z1 - x) * f'' z2" + using le(1) z1(3) z2(3) by auto have "z2 \ C" using z2 z1 * atMostAtLeast_subset_convex \convex C\ \z1 \ C\ \y \ C\ \z1 < y\ by fastforce - then have B: "f'' z2 \ 0" - using assms by auto - from A B have "(z1 - x) * f'' z2 \ 0" - by auto - with cool have "(f y - f x) / (y - x) - f' x \ 0" + with z1 assms have "(z1 - x) * f'' z2 \ 0" by auto - from mult_right_mono[OF this ge(2)] - have "(f y - f x) / (y - x) * (y - x) - f' x * (y - x) \ 0 * (y - x)" - by (simp add: algebra_simps) - then have "f y - f x - f' x * (y - x) \ 0" - using ge by auto then show "f y - f x \ f' x * (y - x)" "f' y * (x - y) \ f x - f y" - using res by auto + using that(3) z1(3) res cool by auto qed - show ?thesis - proof (cases "x = y") - case True - with x y show ?thesis by auto - next - case False - with less_imp x y show ?thesis - by (auto simp: neq_iff) - qed + then show ?thesis + using x y by fastforce qed lemma f''_ge0_imp_convex: fixes f :: "real \ real" - assumes conv: "convex C" - and f': "\x. x \ C \ DERIV f x :> (f' x)" - and f'': "\x. x \ C \ DERIV f' x :> (f'' x)" - and 0: "\x. x \ C \ f'' x \ 0" + assumes "convex C" + and "\x. x \ C \ DERIV f x :> (f' x)" + and "\x. x \ C \ DERIV f' x :> (f'' x)" + and "\x. x \ C \ f'' x \ 0" shows "convex_on C f" - using f''_imp_f'[OF conv f' f'' 0] assms pos_convex_function - by fastforce + by (metis assms f''_imp_f' pos_convex_function) lemma f''_le0_imp_concave: fixes f :: "real \ real" @@ -965,14 +809,23 @@ qed lemma convex_on_inverse: + fixes A :: "real set" assumes "A \ {0<..}" - shows "convex_on A (inverse :: real \ real)" -proof (rule convex_on_subset[OF _ assms], intro convex_on_realI[of _ _ "\x. -inverse (x^2)"]) - fix u v :: real - assume "u \ {0<..}" "v \ {0<..}" "u \ v" - with assms show "-inverse (u^2) \ -inverse (v^2)" - by (intro le_imp_neg_le le_imp_inverse_le power_mono) (simp_all) -qed (insert assms, auto intro!: derivative_eq_intros simp: field_split_simps power2_eq_square) + shows "convex_on A inverse" +proof - + have "convex_on {0::real<..} inverse" + proof (intro convex_on_realI) + fix u v :: real + assume "u \ {0<..}" "v \ {0<..}" "u \ v" + with assms show "-inverse (u^2) \ -inverse (v^2)" + by simp + next + show "\x. x \ {0<..} \ (inverse has_real_derivative - inverse (x\<^sup>2)) (at x)" + by (rule derivative_eq_intros | simp add: power2_eq_square)+ + qed auto + then show ?thesis + using assms convex_on_subset by blast +qed lemma convex_onD_Icc': assumes "convex_on {x..y} f" "c \ {x..y}" @@ -995,7 +848,7 @@ also from d have "\ = (f y - f x) / d * (c - x) + f x" by (simp add: field_simps) finally show ?thesis . -qed (insert assms(2), simp_all) +qed (use assms in auto) lemma convex_onD_Icc'': assumes "convex_on {x..y} f" "c \ {x..y}" @@ -1018,7 +871,7 @@ also from d have "\ = (f x - f y) / d * (y - c) + f y" by (simp add: field_simps) finally show ?thesis . -qed (insert assms(2), simp_all) +qed (use assms in auto) subsection \Some inequalities\ @@ -1161,59 +1014,34 @@ lemma cone_iff: assumes "S \ {}" - shows "cone S \ 0 \ S \ (\c. c > 0 \ ((*\<^sub>R) c) ` S = S)" -proof - + shows "cone S \ 0 \ S \ (\c. c > 0 \ ((*\<^sub>R) c) ` S = S)" (is "_ = ?rhs") +proof + assume "cone S" { - assume "cone S" - { - fix c :: real - assume "c > 0" - { - fix x - assume "x \ S" - then have "x \ ((*\<^sub>R) c) ` S" - unfolding image_def - using \cone S\ \c>0\ mem_cone[of S x "1/c"] - exI[of "(\t. t \ S \ x = c *\<^sub>R t)" "(1 / c) *\<^sub>R x"] - by auto - } - moreover - { - fix x - assume "x \ ((*\<^sub>R) c) ` S" - then have "x \ S" - using \0 < c\ \cone S\ mem_cone by fastforce - } - ultimately have "((*\<^sub>R) c) ` S = S" by blast - } - then have "0 \ S \ (\c. c > 0 \ ((*\<^sub>R) c) ` S = S)" - using \cone S\ cone_contains_0[of S] assms by auto + fix c :: real + assume "c > 0" + have "x \ ((*\<^sub>R) c) ` S" if "x \ S" for x + using \cone S\ \c>0\ mem_cone[of S x "1/c"] that + exI[of "(\t. t \ S \ x = c *\<^sub>R t)" "(1 / c) *\<^sub>R x"] + by auto + then have "((*\<^sub>R) c) ` S = S" + using \0 < c\ \cone S\ mem_cone by fastforce } - moreover - { - assume a: "0 \ S \ (\c. c > 0 \ ((*\<^sub>R) c) ` S = S)" - { - fix x - assume "x \ S" - fix c1 :: real - assume "c1 \ 0" - then have "c1 = 0 \ c1 > 0" by auto - then have "c1 *\<^sub>R x \ S" using a \x \ S\ by auto - } - then have "cone S" unfolding cone_def by auto - } - ultimately show ?thesis by blast + then show "0 \ S \ (\c. c > 0 \ ((*\<^sub>R) c) ` S = S)" + using \cone S\ cone_contains_0[of S] assms by auto +next + show "?rhs \ cone S" + by (metis Convex.cone_def imageI order_neq_le_trans scaleR_zero_left) qed lemma cone_hull_empty: "cone hull {} = {}" by (metis cone_empty cone_hull_eq) lemma cone_hull_empty_iff: "S = {} \ cone hull S = {}" - by (metis bot_least cone_hull_empty hull_subset xtrans(5)) + by (metis cone_hull_empty hull_subset subset_empty) lemma cone_hull_contains_0: "S \ {} \ 0 \ cone hull S" - using cone_cone_hull[of S] cone_contains_0[of "cone hull S"] cone_hull_empty_iff[of S] - by auto + by (metis cone_cone_hull cone_contains_0 cone_hull_empty_iff) lemma mem_cone_hull: assumes "x \ S" "c \ 0" @@ -1222,65 +1050,27 @@ proposition cone_hull_expl: "cone hull S = {c *\<^sub>R x | c x. c \ 0 \ x \ S}" (is "?lhs = ?rhs") -proof - - { - fix x - assume "x \ ?rhs" - then obtain cx :: real and xx where x: "x = cx *\<^sub>R xx" "cx \ 0" "xx \ S" - by auto - fix c :: real - assume c: "c \ 0" - then have "c *\<^sub>R x = (c * cx) *\<^sub>R xx" - using x by (simp add: algebra_simps) - moreover - have "c * cx \ 0" using c x by auto - ultimately - have "c *\<^sub>R x \ ?rhs" using x by auto - } - then have "cone ?rhs" - unfolding cone_def by auto - then have "?rhs \ Collect cone" - unfolding mem_Collect_eq by auto - { - fix x - assume "x \ S" - then have "1 *\<^sub>R x \ ?rhs" - using zero_le_one by blast - then have "x \ ?rhs" by auto - } - then have "S \ ?rhs" by auto - then have "?lhs \ ?rhs" - using \?rhs \ Collect cone\ hull_minimal[of S "?rhs" "cone"] by auto - moreover - { - fix x - assume "x \ ?rhs" - then obtain cx :: real and xx where x: "x = cx *\<^sub>R xx" "cx \ 0" "xx \ S" - by auto - then have "xx \ cone hull S" - using hull_subset[of S] by auto - then have "x \ ?lhs" - using x cone_cone_hull[of S] cone_def[of "cone hull S"] by auto - } - ultimately show ?thesis by auto -qed +proof + have "?rhs \ Collect cone" + using Convex.cone_def by fastforce + moreover have "S \ ?rhs" + by (smt (verit) mem_Collect_eq scaleR_one subsetI) + ultimately show "?lhs \ ?rhs" + using hull_minimal by blast +qed (use mem_cone_hull in auto) lemma convex_cone: - "convex s \ cone s \ (\x\s. \y\s. (x + y) \ s) \ (\x\s. \c\0. (c *\<^sub>R x) \ s)" + "convex S \ cone S \ (\x\S. \y\S. (x + y) \ S) \ (\x\S. \c\0. (c *\<^sub>R x) \ S)" (is "?lhs = ?rhs") proof - { fix x y - assume "x\s" "y\s" and ?lhs - then have "2 *\<^sub>R x \s" "2 *\<^sub>R y \ s" + assume "x\S" "y\S" and ?lhs + then have "2 *\<^sub>R x \S" "2 *\<^sub>R y \ S" "convex S" unfolding cone_def by auto - then have "x + y \ s" - using \?lhs\[unfolded convex_def, THEN conjunct1] - apply (erule_tac x="2*\<^sub>R x" in ballE) - apply (erule_tac x="2*\<^sub>R y" in ballE) - apply (erule_tac x="1/2" in allE, simp) - apply (erule_tac x="1/2" in allE, auto) - done + then have "x + y \ S" + using convexD [OF \convex S\, of "2*\<^sub>R x" "2*\<^sub>R y"] + by (smt (verit, ccfv_threshold) field_sum_of_halves scaleR_2 scaleR_half_double) } then show ?thesis unfolding convex_def cone_def by blast @@ -1315,13 +1105,12 @@ qed corollary%unimportant connected_UNIV[intro]: "connected (UNIV :: 'a::real_normed_vector set)" -by (simp add: convex_connected) + by (simp add: convex_connected) lemma convex_prod: assumes "\i. i \ Basis \ convex {x. P i x}" shows "convex {x. \i\Basis. P i (x\i)}" - using assms unfolding convex_def - by (auto simp: inner_add_left) + using assms by (auto simp: inner_add_left convex_def) lemma convex_positive_orthant: "convex {x::'a::euclidean_space. (\i\Basis. 0 \ x\i)}" by (rule convex_prod) (simp flip: atLeast_def) @@ -1329,9 +1118,7 @@ subsection \Convex hull\ lemma convex_convex_hull [iff]: "convex (convex hull s)" - unfolding hull_def - using convex_Inter[of "{t. convex t \ s \ t}"] - by auto + by (metis (mono_tags) convex_Inter hull_def mem_Collect_eq) lemma convex_hull_subset: "s \ convex hull t \ convex hull s \ convex hull t" @@ -1344,56 +1131,50 @@ lemma convex_hull_linear_image: assumes f: "linear f" - shows "f ` (convex hull s) = convex hull (f ` s)" + shows "f ` (convex hull S) = convex hull (f ` S)" proof - show "convex hull (f ` s) \ f ` (convex hull s)" + show "convex hull (f ` S) \ f ` (convex hull S)" by (intro hull_minimal image_mono hull_subset convex_linear_image assms convex_convex_hull) - show "f ` (convex hull s) \ convex hull (f ` s)" - proof (unfold image_subset_iff_subset_vimage, rule hull_minimal) - show "s \ f -` (convex hull (f ` s))" - by (fast intro: hull_inc) - show "convex (f -` (convex hull (f ` s)))" - by (intro convex_linear_vimage [OF f] convex_convex_hull) - qed + show "f ` (convex hull S) \ convex hull (f ` S)" + by (meson convex_convex_hull convex_linear_vimage f hull_minimal hull_subset image_subset_iff_subset_vimage) qed lemma in_convex_hull_linear_image: - assumes "linear f" - and "x \ convex hull s" - shows "f x \ convex hull (f ` s)" - using convex_hull_linear_image[OF assms(1)] assms(2) by auto + assumes "linear f" "x \ convex hull S" + shows "f x \ convex hull (f ` S)" + using assms convex_hull_linear_image image_eqI by blast lemma convex_hull_Times: - "convex hull (s \ t) = (convex hull s) \ (convex hull t)" + "convex hull (S \ T) = (convex hull S) \ (convex hull T)" proof - show "convex hull (s \ t) \ (convex hull s) \ (convex hull t)" + show "convex hull (S \ T) \ (convex hull S) \ (convex hull T)" by (intro hull_minimal Sigma_mono hull_subset convex_Times convex_convex_hull) - have "(x, y) \ convex hull (s \ t)" if x: "x \ convex hull s" and y: "y \ convex hull t" for x y + have "(x, y) \ convex hull (S \ T)" if x: "x \ convex hull S" and y: "y \ convex hull T" for x y proof (rule hull_induct [OF x], rule hull_induct [OF y]) - fix x y assume "x \ s" and "y \ t" - then show "(x, y) \ convex hull (s \ t)" + fix x y assume "x \ S" and "y \ T" + then show "(x, y) \ convex hull (S \ T)" by (simp add: hull_inc) next - fix x let ?S = "((\y. (0, y)) -` (\p. (- x, 0) + p) ` (convex hull s \ t))" + fix x let ?S = "((\y. (0, y)) -` (\p. (- x, 0) + p) ` (convex hull S \ T))" have "convex ?S" by (intro convex_linear_vimage convex_translation convex_convex_hull, simp add: linear_iff) - also have "?S = {y. (x, y) \ convex hull (s \ t)}" + also have "?S = {y. (x, y) \ convex hull (S \ T)}" by (auto simp: image_def Bex_def) - finally show "convex {y. (x, y) \ convex hull (s \ t)}" . + finally show "convex {y. (x, y) \ convex hull (S \ T)}" . next - show "convex {x. (x, y) \ convex hull s \ t}" + show "convex {x. (x, y) \ convex hull S \ T}" proof - - fix y let ?S = "((\x. (x, 0)) -` (\p. (0, - y) + p) ` (convex hull s \ t))" + fix y let ?S = "((\x. (x, 0)) -` (\p. (0, - y) + p) ` (convex hull S \ T))" have "convex ?S" by (intro convex_linear_vimage convex_translation convex_convex_hull, simp add: linear_iff) - also have "?S = {x. (x, y) \ convex hull (s \ t)}" + also have "?S = {x. (x, y) \ convex hull (S \ T)}" by (auto simp: image_def Bex_def) - finally show "convex {x. (x, y) \ convex hull (s \ t)}" . + finally show "convex {x. (x, y) \ convex hull (S \ T)}" . qed qed - then show "(convex hull s) \ (convex hull t) \ convex hull (s \ t)" + then show "(convex hull S) \ (convex hull T) \ convex hull (S \ T)" unfolding subset_eq split_paired_Ball_Sigma by blast qed @@ -1401,10 +1182,10 @@ subsubsection\<^marker>\tag unimportant\ \Stepping theorems for convex hulls of finite sets\ lemma convex_hull_empty[simp]: "convex hull {} = {}" - by (rule hull_unique) auto + by (simp add: hull_same) lemma convex_hull_singleton[simp]: "convex hull {a} = {a}" - by (rule hull_unique) auto + by (simp add: hull_same) lemma convex_hull_insert: fixes S :: "'a::real_vector set" @@ -1415,19 +1196,17 @@ proof (intro equalityI hull_minimal subsetI) fix x assume "x \ insert a S" - then have "\u\0. \v\0. u + v = 1 \ (\b. b \ convex hull S \ x = u *\<^sub>R a + v *\<^sub>R b)" + then show "x \ ?hull" unfolding insert_iff proof assume "x = a" then show ?thesis - by (rule_tac x=1 in exI) (use assms hull_subset in fastforce) + by (smt (verit, del_insts) add.right_neutral assms ex_in_conv hull_inc mem_Collect_eq scaleR_one scaleR_zero_left) next assume "x \ S" - with hull_subset[of S convex] show ?thesis + with hull_subset show ?thesis by force qed - then show "x \ ?hull" - by simp next fix x assume "x \ ?hull" @@ -1468,44 +1247,35 @@ next case False have "1 - (u * u1 + v * u2) = (u + v) - (u * u1 + v * u2)" - using as(3) obt1(3) obt2(3) by (auto simp: field_simps) - also have "\ = u * (v1 + u1 - u1) + v * (v2 + u2 - u2)" - using as(3) obt1(3) obt2(3) by (auto simp: field_simps) + by (simp add: as(3)) also have "\ = u * v1 + v * v2" - by simp - finally have **:"1 - (u * u1 + v * u2) = u * v1 + v * v2" by auto + by (smt (verit, ccfv_SIG) distrib_left mult_cancel_left1 obt1(3) obt2(3)) + finally have **:"1 - (u * u1 + v * u2) = u * v1 + v * v2" . let ?b = "((u * v1) / (u * v1 + v * v2)) *\<^sub>R b1 + ((v * v2) / (u * v1 + v * v2)) *\<^sub>R b2" have zeroes: "0 \ u * v1 + v * v2" "0 \ u * v1" "0 \ u * v1 + v * v2" "0 \ v * v2" - using as(1,2) obt1(1,2) obt2(1,2) by auto + using as obt1 obt2 by auto show ?thesis proof show "u *\<^sub>R x + v *\<^sub>R y = (u * u1) *\<^sub>R a + (v * u2) *\<^sub>R a + (?b - (u * u1) *\<^sub>R ?b - (v * u2) *\<^sub>R ?b)" unfolding xeq yeq * ** using False by (auto simp: scaleR_left_distrib scaleR_right_distrib) show "?b \ convex hull S" - using False zeroes obt1(4) obt2(4) - by (auto simp: convexD [OF convex_convex_hull] scaleR_left_distrib scaleR_right_distrib add_divide_distrib[symmetric] zero_le_divide_iff) + using False mem_convex_alt obt1(4) obt2(4) zeroes(2) zeroes(4) by fastforce qed qed then obtain b where b: "b \ convex hull S" "u *\<^sub>R x + v *\<^sub>R y = (u * u1) *\<^sub>R a + (v * u2) *\<^sub>R a + (b - (u * u1) *\<^sub>R b - (v * u2) *\<^sub>R b)" .. - - have u1: "u1 \ 1" - unfolding obt1(3)[symmetric] and not_le using obt1(2) by auto - have u2: "u2 \ 1" - unfolding obt2(3)[symmetric] and not_le using obt2(2) by auto + obtain u1: "u1 \ 1" and u2: "u2 \ 1" + using obt1 obt2 by auto have "u1 * u + u2 * v \ max u1 u2 * u + max u1 u2 * v" - proof (rule add_mono) - show "u1 * u \ max u1 u2 * u" "u2 * v \ max u1 u2 * v" - by (simp_all add: as mult_right_mono) - qed + by (smt (verit, ccfv_SIG) as mult_right_mono) also have "\ \ 1" unfolding distrib_left[symmetric] and as(3) using u1 u2 by auto finally have le1: "u1 * u + u2 * v \ 1" . show "u *\<^sub>R x + v *\<^sub>R y \ ?hull" proof (intro CollectI exI conjI) show "0 \ u * u1 + v * u2" - by (simp add: as(1) as(2) obt1(1) obt2(1)) + by (simp add: as obt1(1) obt2(1)) show "0 \ 1 - u * u1 - v * u2" by (simp add: le1 diff_diff_add mult.commute) qed (use b in \auto simp: algebra_simps\) @@ -1516,9 +1286,9 @@ "convex hull (insert a S) = (if S = {} then {a} else {(1 - u) *\<^sub>R a + u *\<^sub>R x |x u. 0 \ u \ u \ 1 \ x \ convex hull S})" - apply (auto simp: convex_hull_insert) - using diff_eq_eq apply fastforce - using diff_add_cancel diff_ge_0_iff_ge by blast + apply (simp add: convex_hull_insert) + using diff_add_cancel diff_ge_0_iff_ge + by (smt (verit, del_insts) Collect_cong) subsubsection\<^marker>\tag unimportant\ \Explicit expression for convex hull\ @@ -1608,87 +1378,56 @@ shows "convex hull p = {y. \S u. finite S \ S \ p \ (\x\S. 0 \ u x) \ sum u S = 1 \ sum (\v. u v *\<^sub>R v) S = y}" (is "?lhs = ?rhs") -proof - +proof (intro subset_antisym subsetI) + fix x + assume "x \ convex hull p" + then obtain k u y where + obt: "\i\{1::nat..k}. 0 \ u i \ y i \ p" "sum u {1..k} = 1" "(\i = 1..k. u i *\<^sub>R y i) = x" + unfolding convex_hull_indexed by auto + have fin: "finite {1..k}" by auto { - fix x - assume "x\?lhs" - then obtain k u y where - obt: "\i\{1::nat..k}. 0 \ u i \ y i \ p" "sum u {1..k} = 1" "(\i = 1..k. u i *\<^sub>R y i) = x" - unfolding convex_hull_indexed by auto - - have fin: "finite {1..k}" by auto - have fin': "\v. finite {i \ {1..k}. y i = v}" by auto - { - fix j - assume "j\{1..k}" - then have "y j \ p \ 0 \ sum u {i. Suc 0 \ i \ i \ k \ y i = y j}" - using obt(1)[THEN bspec[where x=j]] and obt(2) - by (metis (no_types, lifting) One_nat_def atLeastAtMost_iff mem_Collect_eq obt(1) sum_nonneg) - } - moreover - have "(\v\y ` {1..k}. sum u {i \ {1..k}. y i = v}) = 1" - unfolding sum.image_gen[OF fin, symmetric] using obt(2) by auto - moreover have "(\v\y ` {1..k}. sum u {i \ {1..k}. y i = v} *\<^sub>R v) = x" - using sum.image_gen[OF fin, of "\i. u i *\<^sub>R y i" y, symmetric] - unfolding scaleR_left.sum using obt(3) by auto - ultimately - have "\S u. finite S \ S \ p \ (\x\S. 0 \ u x) \ sum u S = 1 \ (\v\S. u v *\<^sub>R v) = x" - apply (rule_tac x="y ` {1..k}" in exI) - apply (rule_tac x="\v. sum u {i\{1..k}. y i = v}" in exI, auto) - done - then have "x\?rhs" by auto + fix j + assume "j\{1..k}" + then have "y j \ p \ 0 \ sum u {i. Suc 0 \ i \ i \ k \ y i = y j}" + by (metis (mono_tags, lifting) One_nat_def atLeastAtMost_iff mem_Collect_eq obt(1) sum_nonneg) } - moreover + moreover have "(\v\y ` {1..k}. sum u {i \ {1..k}. y i = v}) = 1" + unfolding sum.image_gen[OF fin, symmetric] using obt(2) by auto + moreover have "(\v\y ` {1..k}. sum u {i \ {1..k}. y i = v} *\<^sub>R v) = x" + using sum.image_gen[OF fin, of "\i. u i *\<^sub>R y i" y, symmetric] + unfolding scaleR_left.sum using obt(3) by auto + ultimately + have "\S u. finite S \ S \ p \ (\x\S. 0 \ u x) \ sum u S = 1 \ (\v\S. u v *\<^sub>R v) = x" + by (smt (verit, ccfv_SIG) imageE mem_Collect_eq obt(1) subsetI sum.cong sum.infinite sum_nonneg) + then show "x \ ?rhs" by auto +next + fix y + assume "y \ ?rhs" + then obtain S u where + S: "finite S" "S \ p" "\x\S. 0 \ u x" "sum u S = 1" "(\v\S. u v *\<^sub>R v) = y" + by auto + obtain f where f: "inj_on f {1..card S}" "f ` {1..card S} = S" + using ex_bij_betw_nat_finite_1[OF S(1)] unfolding bij_betw_def by auto + then have "0 \ u (f i)" "f i \ p" if "i \ {1..card S}" for i + using S \i \ {1..card S}\ by blast+ + moreover { fix y - assume "y\?rhs" - then obtain S u where - obt: "finite S" "S \ p" "\x\S. 0 \ u x" "sum u S = 1" "(\v\S. u v *\<^sub>R v) = y" - by auto - - obtain f where f: "inj_on f {1..card S}" "f ` {1..card S} = S" - using ex_bij_betw_nat_finite_1[OF obt(1)] unfolding bij_betw_def by auto - { - fix i :: nat - assume "i\{1..card S}" - then have "f i \ S" - using f(2) by blast - then have "0 \ u (f i)" "f i \ p" using obt(2,3) by auto - } - moreover have *: "finite {1..card S}" by auto - { - fix y - assume "y\S" - then obtain i where "i\{1..card S}" "f i = y" - using f using image_iff[of y f "{1..card S}"] - by auto - then have "{x. Suc 0 \ x \ x \ card S \ f x = y} = {i}" - using f(1) inj_onD by fastforce - then have "card {x. Suc 0 \ x \ x \ card S \ f x = y} = 1" by auto - then have "(\x\{x \ {1..card S}. f x = y}. u (f x)) = u y" - "(\x\{x \ {1..card S}. f x = y}. u (f x) *\<^sub>R f x) = u y *\<^sub>R y" - by (auto simp: sum_constant_scaleR) - } - then have "(\x = 1..card S. u (f x)) = 1" "(\i = 1..card S. u (f i) *\<^sub>R f i) = y" - unfolding sum.image_gen[OF *(1), of "\x. u (f x) *\<^sub>R f x" f] - and sum.image_gen[OF *(1), of "\x. u (f x)" f] - unfolding f - using sum.cong [of S S "\y. (\x\{x \ {1..card S}. f x = y}. u (f x) *\<^sub>R f x)" "\v. u v *\<^sub>R v"] - using sum.cong [of S S "\y. (\x\{x \ {1..card S}. f x = y}. u (f x))" u] - unfolding obt(4,5) - by auto - ultimately - have "\k u x. (\i\{1..k}. 0 \ u i \ x i \ p) \ sum u {1..k} = 1 \ - (\i::nat = 1..k. u i *\<^sub>R x i) = y" - apply (rule_tac x="card S" in exI) - apply (rule_tac x="u \ f" in exI) - apply (rule_tac x=f in exI, fastforce) - done - then have "y \ ?lhs" - unfolding convex_hull_indexed by auto + assume "y\S" + then obtain i where "i\{1..card S}" "f i = y" + by (metis f(2) image_iff) + then have "{x. Suc 0 \ x \ x \ card S \ f x = y} = {i}" + using f(1) inj_onD by fastforce + then have "(\x\{x \ {1..card S}. f x = y}. u (f x)) = u y" + "(\x\{x \ {1..card S}. f x = y}. u (f x) *\<^sub>R f x) = u y *\<^sub>R y" + by (simp_all add: sum_constant_scaleR \f i = y\) } - ultimately show ?thesis - unfolding set_eq_iff by blast + then have "(\x = 1..card S. u (f x)) = 1" "(\i = 1..card S. u (f i) *\<^sub>R f i) = y" + by (metis (mono_tags, lifting) S(4,5) f sum.reindex_cong)+ + ultimately + show "y \ convex hull p" + unfolding convex_hull_indexed + by (smt (verit, del_insts) mem_Collect_eq sum.cong) qed @@ -1832,10 +1571,7 @@ lemma aff_dim_convex_hull: fixes S :: "'n::euclidean_space set" shows "aff_dim (convex hull S) = aff_dim S" - using aff_dim_affine_hull[of S] convex_hull_subset_affine_hull[of S] - hull_subset[of S "convex"] aff_dim_subset[of S "convex hull S"] - aff_dim_subset[of "convex hull S" "affine hull S"] - by auto + by (smt (verit) aff_dim_affine_hull aff_dim_subset convex_hull_subset_affine_hull hull_subset) subsection \Caratheodory's theorem\ @@ -1856,40 +1592,31 @@ by (rule_tac ex_least_nat_le, auto) then obtain n where "?P n" and smallest: "\k ?P k" by blast - then obtain S u where obt: "finite S" "card S = n" "S\p" "\x\S. 0 \ u x" - "sum u S = 1" "(\v\S. u v *\<^sub>R v) = y" by auto + then obtain S u where S: "finite S" "card S = n" "S\p" + and u: "\x\S. 0 \ u x" "sum u S = 1" "(\v\S. u v *\<^sub>R v) = y" by auto have "card S \ aff_dim p + 1" proof (rule ccontr, simp only: not_le) assume "aff_dim p + 1 < card S" then have "affine_dependent S" - using affine_dependent_biggerset[OF obt(1)] independent_card_le_aff_dim not_less obt(3) - by blast + by (smt (verit) independent_card_le_aff_dim S(3)) then obtain w v where wv: "sum w S = 0" "v\S" "w v \ 0" "(\v\S. w v *\<^sub>R v) = 0" - using affine_dependent_explicit_finite[OF obt(1)] by auto + using affine_dependent_explicit_finite[OF S(1)] by auto define i where "i = (\v. (u v) / (- w v)) ` {v\S. w v < 0}" define t where "t = Min i" have "\x\S. w x < 0" - proof (rule ccontr, simp add: not_less) - assume as:"\x\S. 0 \ w x" - then have "sum w (S - {v}) \ 0" - by (meson Diff_iff sum_nonneg) - then have "sum w S > 0" - using as obt(1) sum_nonneg_eq_0_iff wv by blast - then show False using wv(1) by auto - qed + by (smt (verit, best) S(1) sum_pos2 wv) then have "i \ {}" unfolding i_def by auto then have "t \ 0" - using Min_ge_iff[of i 0] and obt(1) + using Min_ge_iff[of i 0] and S(1) u[unfolded le_less] unfolding t_def i_def - using obt(4)[unfolded le_less] by (auto simp: divide_le_0_iff) have t: "\v\S. u v + t * w v \ 0" proof fix v assume "v \ S" then have v: "0 \ u v" - using obt(4)[THEN bspec[where x=v]] by auto + using u(1) by blast show "0 \ u v + t * w v" proof (cases "w v < 0") case False @@ -1897,26 +1624,26 @@ next case True then have "t \ u v / (- w v)" - using \v\S\ obt unfolding t_def i_def by (auto intro: Min_le) + using \v\S\ S unfolding t_def i_def by (auto intro: Min_le) then show ?thesis unfolding real_0_le_add_iff using True neg_le_minus_divide_eq by auto qed qed obtain a where "a \ S" and "t = (\v. (u v) / (- w v)) a" and "w a < 0" - using Min_in[OF _ \i\{}\] and obt(1) unfolding i_def t_def by auto + using Min_in[OF _ \i\{}\] and S(1) unfolding i_def t_def by auto then have a: "a \ S" "u a + t * w a = 0" by auto have *: "\f. sum f (S - {a}) = sum f S - ((f a)::'b::ab_group_add)" - unfolding sum.remove[OF obt(1) \a\S\] by auto + unfolding sum.remove[OF S(1) \a\S\] by auto have "(\v\S. u v + t * w v) = 1" - unfolding sum.distrib wv(1) sum_distrib_left[symmetric] obt(5) by auto + by (metis add.right_neutral mult_zero_right sum.distrib sum_distrib_left u(2) wv(1)) moreover have "(\v\S. u v *\<^sub>R v + (t * w v) *\<^sub>R v) - (u a *\<^sub>R a + (t * w a) *\<^sub>R a) = y" - unfolding sum.distrib obt(6) scaleR_scaleR[symmetric] scaleR_right.sum [symmetric] wv(4) + unfolding sum.distrib u(3) scaleR_scaleR[symmetric] scaleR_right.sum [symmetric] wv(4) using a(2) [THEN eq_neg_iff_add_eq_0 [THEN iffD2]] by simp ultimately have "?P (n - 1)" apply (rule_tac x="(S - {a})" in exI) apply (rule_tac x="\v. u v + t * w v" in exI) - using obt(1-3) and t and a + using S t a apply (auto simp: * scaleR_left_distrib) done then show False @@ -1924,7 +1651,7 @@ qed then show "\S u. finite S \ S \ p \ card S \ aff_dim p + 1 \ (\x\S. 0 \ u x) \ sum u S = 1 \ (\v\S. u v *\<^sub>R v) = y" - using obt by auto + using S u by auto qed auto lemma caratheodory_aff_dim: @@ -1934,7 +1661,7 @@ proof have "\x S u. \finite S; S \ p; int (card S) \ aff_dim p + 1; \x\S. 0 \ u x; sum u S = 1\ \ (\v\S. u v *\<^sub>R v) \ convex hull S" - by (simp add: hull_subset convex_explicit [THEN iffD1, OF convex_convex_hull]) + by (metis (mono_tags, lifting) convex_convex_hull convex_explicit hull_subset) then show "?lhs \ ?rhs" by (subst convex_hull_caratheodory_aff_dim, auto) qed (use hull_mono in auto) @@ -1986,23 +1713,19 @@ lemma convex_hull_set_plus: "convex hull (S + T) = convex hull S + convex hull T" - unfolding set_plus_image - apply (subst convex_hull_linear_image [symmetric]) - apply (simp add: linear_iff scaleR_right_distrib) - apply (simp add: convex_hull_Times) - done + by (simp add: set_plus_image linear_iff scaleR_right_distrib convex_hull_Times + flip: convex_hull_linear_image) lemma translation_eq_singleton_plus: "(\x. a + x) ` T = {a} + T" unfolding set_plus_def by auto lemma convex_hull_translation: "convex hull ((\x. a + x) ` S) = (\x. a + x) ` (convex hull S)" - unfolding translation_eq_singleton_plus - by (simp only: convex_hull_set_plus convex_hull_singleton) + by (simp add: convex_hull_set_plus translation_eq_singleton_plus) lemma convex_hull_scaling: "convex hull ((\x. c *\<^sub>R x) ` S) = (\x. c *\<^sub>R x) ` (convex hull S)" - using linear_scaleR by (rule convex_hull_linear_image [symmetric]) + by (simp add: convex_hull_linear_image) lemma convex_hull_affinity: "convex hull ((\x. a + c *\<^sub>R x) ` S) = (\x. a + c *\<^sub>R x) ` (convex hull S)" @@ -2022,31 +1745,25 @@ fix u v :: real assume uv: "u \ 0" "v \ 0" "u + v = 1" then have *: "u *\<^sub>R x \ cone hull S" "v *\<^sub>R y \ cone hull S" - using cone_cone_hull[of S] xy cone_def[of "cone hull S"] by auto - from * obtain cx :: real and xx where x: "u *\<^sub>R x = cx *\<^sub>R xx" "cx \ 0" "xx \ S" - using cone_hull_expl[of S] by auto - from * obtain cy :: real and yy where y: "v *\<^sub>R y = cy *\<^sub>R yy" "cy \ 0" "yy \ S" + by (simp_all add: cone_cone_hull mem_cone uv xy) + then obtain cx :: real and xx + and cy :: real and yy where x: "u *\<^sub>R x = cx *\<^sub>R xx" "cx \ 0" "xx \ S" + and y: "v *\<^sub>R y = cy *\<^sub>R yy" "cy \ 0" "yy \ S" using cone_hull_expl[of S] by auto - { - assume "cx + cy \ 0" - then have "u *\<^sub>R x = 0" and "v *\<^sub>R y = 0" - using x y by auto - then have "u *\<^sub>R x + v *\<^sub>R y = 0" - by auto - then have "u *\<^sub>R x + v *\<^sub>R y \ cone hull S" - using cone_hull_contains_0[of S] \S \ {}\ by auto - } + + have "u *\<^sub>R x + v *\<^sub>R y \ cone hull S" if "cx + cy \ 0" + using "*"(1) nless_le that x(2) y by fastforce moreover - { - assume "cx + cy > 0" - then have "(cx / (cx + cy)) *\<^sub>R xx + (cy / (cx + cy)) *\<^sub>R yy \ S" - using assms mem_convex_alt[of S xx yy cx cy] x y by auto + have "u *\<^sub>R x + v *\<^sub>R y \ cone hull S" if "cx + cy > 0" + proof - + have "(cx / (cx + cy)) *\<^sub>R xx + (cy / (cx + cy)) *\<^sub>R yy \ S" + using assms mem_convex_alt[of S xx yy cx cy] x y that by auto then have "cx *\<^sub>R xx + cy *\<^sub>R yy \ cone hull S" using mem_cone_hull[of "(cx/(cx+cy)) *\<^sub>R xx + (cy/(cx+cy)) *\<^sub>R yy" S "cx+cy"] \cx+cy>0\ by (auto simp: scaleR_right_distrib) - then have "u *\<^sub>R x + v *\<^sub>R y \ cone hull S" + then show ?thesis using x y by auto - } + qed moreover have "cx + cy \ 0 \ cx + cy > 0" by auto ultimately show "u *\<^sub>R x + v *\<^sub>R y \ cone hull S" by blast qed @@ -2054,28 +1771,7 @@ lemma cone_convex_hull: assumes "cone S" shows "cone (convex hull S)" -proof (cases "S = {}") - case True - then show ?thesis by auto -next - case False - then have *: "0 \ S \ (\c. c > 0 \ (*\<^sub>R) c ` S = S)" - using cone_iff[of S] assms by auto - { - fix c :: real - assume "c > 0" - then have "(*\<^sub>R) c ` (convex hull S) = convex hull ((*\<^sub>R) c ` S)" - using convex_hull_scaling[of _ S] by auto - also have "\ = convex hull S" - using * \c > 0\ by auto - finally have "(*\<^sub>R) c ` (convex hull S) = convex hull S" - by auto - } - then have "0 \ convex hull S" "\c. c > 0 \ ((*\<^sub>R) c ` (convex hull S)) = (convex hull S)" - using * hull_subset[of S convex] by auto - then show ?thesis - using \S \ {}\ cone_iff[of "convex hull S"] by auto -qed + by (metis (no_types, lifting) affine_hull_convex_hull affine_hull_eq_empty assms cone_iff convex_hull_scaling hull_inc) subsection \Radon's theorem\ @@ -2084,29 +1780,17 @@ lemma Radon_ex_lemma: assumes "finite c" "affine_dependent c" shows "\u. sum u c = 0 \ (\v\c. u v \ 0) \ sum (\v. u v *\<^sub>R v) c = 0" -proof - - from assms(2)[unfolded affine_dependent_explicit] - obtain S u where - "finite S" "S \ c" "sum u S = 0" "\v\S. u v \ 0" "(\v\S. u v *\<^sub>R v) = 0" - by blast - then show ?thesis - apply (rule_tac x="\v. if v\S then u v else 0" in exI) - unfolding if_smult scaleR_zero_left - by (auto simp: Int_absorb1 sum.inter_restrict[OF \finite c\, symmetric]) -qed + using affine_dependent_explicit_finite assms by blast lemma Radon_s_lemma: assumes "finite S" and "sum f S = (0::real)" shows "sum f {x\S. 0 < f x} = - sum f {x\S. f x < 0}" proof - - have *: "\x. (if f x < 0 then f x else 0) + (if 0 < f x then f x else 0) = f x" + have "\x. (if f x < 0 then f x else 0) + (if 0 < f x then f x else 0) = f x" by auto - show ?thesis - unfolding add_eq_0_iff[symmetric] and sum.inter_filter[OF assms(1)] - and sum.distrib[symmetric] and * - using assms(2) - by assumption + then show ?thesis + using assms by (simp add: sum.inter_filter flip: sum.distrib add_eq_0_iff) qed lemma Radon_v_lemma: @@ -2115,19 +1799,15 @@ and "\x. g x = (0::real) \ f x = (0::'a::euclidean_space)" shows "(sum f {x\S. 0 < g x}) = - sum f {x\S. g x < 0}" proof - - have *: "\x. (if 0 < g x then f x else 0) + (if g x < 0 then f x else 0) = f x" - using assms(3) by auto - show ?thesis - unfolding eq_neg_iff_add_eq_0 and sum.inter_filter[OF assms(1)] - and sum.distrib[symmetric] and * - using assms(2) - apply assumption - done + have "\x. (if 0 < g x then f x else 0) + (if g x < 0 then f x else 0) = f x" + using assms by auto + then show ?thesis + using assms by (simp add: sum.inter_filter eq_neg_iff_add_eq_0 flip: sum.distrib add_eq_0_iff) qed lemma Radon_partition: assumes "finite C" "affine_dependent C" - shows "\m p. m \ p = {} \ m \ p = C \ (convex hull m) \ (convex hull p) \ {}" + shows "\M P. M \ P = {} \ M \ P = C \ (convex hull M) \ (convex hull P) \ {}" proof - obtain u v where uv: "sum u C = 0" "v\C" "u v \ 0" "(\v\C. u v *\<^sub>R v) = 0" using Radon_ex_lemma[OF assms] by auto @@ -2139,19 +1819,12 @@ case False then have "u v < 0" by auto then show ?thesis - proof (cases "\w\{x \ C. 0 < u x}. u w > 0") - case True - then show ?thesis - using sum_nonneg_eq_0_iff[of _ u, OF fin(1)] by auto - next - case False - then have "sum u C \ sum (\x. if x=v then u v else 0) C" - by (rule_tac sum_mono, auto) - then show ?thesis - unfolding sum.delta[OF assms(1)] using uv(2) and \u v < 0\ and uv(1) by auto - qed - qed (insert sum_nonneg_eq_0_iff[of _ u, OF fin(1)] uv(2-3), auto) - + by (smt (verit) assms(1) fin(1) mem_Collect_eq sum.neutral_const sum_mono_inv uv) + next + case True + with fin uv show "sum u {x \ C. 0 < u x} \ 0" + by (smt (verit) fin(1) mem_Collect_eq sum_nonneg_eq_0_iff uv) + qed then have *: "sum u {x\C. u x > 0} > 0" unfolding less_le by (metis (no_types, lifting) mem_Collect_eq sum_nonneg) moreover have "sum u ({x \ C. 0 < u x} \ {x \ C. u x < 0}) = sum u C" @@ -2188,97 +1861,85 @@ theorem Radon: assumes "affine_dependent c" - obtains m p where "m \ c" "p \ c" "m \ p = {}" "(convex hull m) \ (convex hull p) \ {}" -proof - - from assms[unfolded affine_dependent_explicit] - obtain S u where - "finite S" "S \ c" "sum u S = 0" "\v\S. u v \ 0" "(\v\S. u v *\<^sub>R v) = 0" - by blast - then have *: "finite S" "affine_dependent S" and S: "S \ c" - unfolding affine_dependent_explicit by auto - from Radon_partition[OF *] - obtain m p where "m \ p = {}" "m \ p = S" "convex hull m \ convex hull p \ {}" - by blast - with S show ?thesis - by (force intro: that[of p m]) -qed + obtains M P where "M \ c" "P \ c" "M \ P = {}" "(convex hull M) \ (convex hull P) \ {}" + by (smt (verit) Radon_partition affine_dependent_explicit affine_dependent_explicit_finite assms le_sup_iff) subsection \Helly's theorem\ lemma Helly_induct: - fixes f :: "'a::euclidean_space set set" - assumes "card f = n" + fixes \ :: "'a::euclidean_space set set" + assumes "card \ = n" and "n \ DIM('a) + 1" - and "\s\f. convex s" "\t\f. card t = DIM('a) + 1 \ \t \ {}" - shows "\f \ {}" + and "\S\\. convex S" "\T\\. card T = DIM('a) + 1 \ \T \ {}" + shows "\\ \ {}" using assms -proof (induction n arbitrary: f) +proof (induction n arbitrary: \) case 0 then show ?case by auto next case (Suc n) - have "finite f" - using \card f = Suc n\ by (auto intro: card_ge_0_finite) - show "\f \ {}" + have "finite \" + using \card \ = Suc n\ by (auto intro: card_ge_0_finite) + show "\\ \ {}" proof (cases "n = DIM('a)") case True then show ?thesis - by (simp add: Suc.prems(1) Suc.prems(4)) + by (simp add: Suc.prems) next case False - have "\(f - {s}) \ {}" if "s \ f" for s + have "\(\ - {S}) \ {}" if "S \ \" for S proof (rule Suc.IH[rule_format]) - show "card (f - {s}) = n" - by (simp add: Suc.prems(1) \finite f\ that) + show "card (\ - {S}) = n" + by (simp add: Suc.prems(1) \finite \\ that) show "DIM('a) + 1 \ n" using False Suc.prems(2) by linarith - show "\t. \t \ f - {s}; card t = DIM('a) + 1\ \ \t \ {}" + show "\t. \t \ \ - {S}; card t = DIM('a) + 1\ \ \t \ {}" by (simp add: Suc.prems(4) subset_Diff_insert) qed (use Suc in auto) - then have "\s\f. \x. x \ \(f - {s})" + then have "\S\\. \x. x \ \(\ - {S})" by blast - then obtain X where X: "\s. s\f \ X s \ \(f - {s})" + then obtain X where X: "\S. S\\ \ X S \ \(\ - {S})" by metis show ?thesis - proof (cases "inj_on X f") + proof (cases "inj_on X \") case False - then obtain s t where "s\t" and st: "s\f" "t\f" "X s = X t" + then obtain S T where "S\T" and st: "S\\" "T\\" "X S = X T" unfolding inj_on_def by auto - then have *: "\f = \(f - {s}) \ \(f - {t})" by auto + then have *: "\\ = \(\ - {S}) \ \(\ - {T})" by auto show ?thesis by (metis "*" X disjoint_iff_not_equal st) next case True - then obtain m p where mp: "m \ p = {}" "m \ p = X ` f" "convex hull m \ convex hull p \ {}" - using Radon_partition[of "X ` f"] and affine_dependent_biggerset[of "X ` f"] - unfolding card_image[OF True] and \card f = Suc n\ - using Suc(3) \finite f\ and False + then obtain M P where mp: "M \ P = {}" "M \ P = X ` \" "convex hull M \ convex hull P \ {}" + using Radon_partition[of "X ` \"] and affine_dependent_biggerset[of "X ` \"] + unfolding card_image[OF True] and \card \ = Suc n\ + using Suc(3) \finite \\ and False by auto - have "m \ X ` f" "p \ X ` f" + have "M \ X ` \" "P \ X ` \" using mp(2) by auto - then obtain g h where gh:"m = X ` g" "p = X ` h" "g \ f" "h \ f" + then obtain \ \ where gh:"M = X ` \" "P = X ` \" "\ \ \" "\ \ \" unfolding subset_image_iff by auto - then have "f \ (g \ h) = f" by auto - then have f: "f = g \ h" - using inj_on_Un_image_eq_iff[of X f "g \ h"] and True + then have "\ \ (\ \ \) = \" by auto + then have \: "\ = \ \ \" + using inj_on_Un_image_eq_iff[of X \ "\ \ \"] and True unfolding mp(2)[unfolded image_Un[symmetric] gh] by auto - have *: "g \ h = {}" - using gh(1) gh(2) local.mp(1) by blast - have "convex hull (X ` h) \ \g" "convex hull (X ` g) \ \h" - by (rule hull_minimal; use X * f in \auto simp: Suc.prems(3) convex_Inter\)+ + have *: "\ \ \ = {}" + using gh local.mp(1) by blast + have "convex hull (X ` \) \ \\" "convex hull (X ` \) \ \\" + by (rule hull_minimal; use X * \ in \auto simp: Suc.prems(3) convex_Inter\)+ then show ?thesis - unfolding f using mp(3)[unfolded gh] by blast + unfolding \ using mp(3)[unfolded gh] by blast qed qed qed theorem Helly: - fixes f :: "'a::euclidean_space set set" - assumes "card f \ DIM('a) + 1" "\s\f. convex s" - and "\t. \t\f; card t = DIM('a) + 1\ \ \t \ {}" - shows "\f \ {}" + fixes \ :: "'a::euclidean_space set set" + assumes "card \ \ DIM('a) + 1" "\s\\. convex s" + and "\t. \t\\; card t = DIM('a) + 1\ \ \t \ {}" + shows "\\ \ {}" using Helly_induct assms by blast subsection \Epigraphs of convex functions\ @@ -2371,25 +2032,18 @@ lemma convex_set_plus: assumes "convex S" and "convex T" shows "convex (S + T)" -proof - - have "convex (\x\ S. \y \ T. {x + y})" - using assms by (rule convex_sums) - moreover have "(\x\ S. \y \ T. {x + y}) = S + T" - unfolding set_plus_def by auto - finally show "convex (S + T)" . -qed + by (metis assms convex_hull_eq convex_hull_set_plus) lemma convex_set_sum: assumes "\i. i \ A \ convex (B i)" shows "convex (\i\A. B i)" -proof (cases "finite A") - case True then show ?thesis using assms - by induct (auto simp: convex_set_plus) -qed auto + using assms + by (induction A rule: infinite_finite_induct) (auto simp: convex_set_plus) lemma finite_set_sum: - assumes "finite A" and "\i\A. finite (B i)" shows "finite (\i\A. B i)" - using assms by (induct set: finite, simp, simp add: finite_set_plus) + assumes "\i\A. finite (B i)" shows "finite (\i\A. B i)" + using assms + by (induction A rule: infinite_finite_induct) (auto simp: finite_set_plus) lemma box_eq_set_sum_Basis: "{x. \i\Basis. x\i \ B i} = (\i\Basis. (\x. x *\<^sub>R i) ` (B i))" (is "?lhs = ?rhs") @@ -2401,7 +2055,7 @@ have "sum f Basis \ i \ B i" if "i \ Basis" and f: "\i\Basis. f i \ (\x. x *\<^sub>R i) ` B i" for i f proof - have "(\x\Basis - {i}. f x \ i) = 0" - proof (rule sum.neutral, intro strip) + proof (intro strip sum.neutral) show "f x \ i = 0" if "x \ Basis - {i}" for x using that f \i \ Basis\ inner_Basis that by fastforce qed @@ -2418,10 +2072,6 @@ lemma convex_hull_set_sum: "convex hull (\i\A. B i) = (\i\A. convex hull (B i))" -proof (cases "finite A") - assume "finite A" then show ?thesis - by (induct set: finite, simp, simp add: convex_hull_set_plus) -qed simp - + by (induction A rule: infinite_finite_induct) (auto simp: convex_hull_set_plus) end \ No newline at end of file diff -r ba2afdd29e1d -r 032a4344903e src/HOL/Analysis/Derivative.thy --- a/src/HOL/Analysis/Derivative.thy Tue Aug 01 11:27:55 2023 +0200 +++ b/src/HOL/Analysis/Derivative.thy Thu Aug 03 19:10:43 2023 +0200 @@ -163,27 +163,27 @@ lemma linear_imp_differentiable_on: fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" shows "linear f \ f differentiable_on S" -by (simp add: differentiable_on_def linear_imp_differentiable) + by (simp add: differentiable_on_def linear_imp_differentiable) lemma differentiable_on_minus [simp, derivative_intros]: - "f differentiable_on S \ (\z. -(f z)) differentiable_on S" -by (simp add: differentiable_on_def) + "f differentiable_on S \ (\z. -(f z)) differentiable_on S" + by (simp add: differentiable_on_def) lemma differentiable_on_add [simp, derivative_intros]: - "\f differentiable_on S; g differentiable_on S\ \ (\z. f z + g z) differentiable_on S" -by (simp add: differentiable_on_def) + "\f differentiable_on S; g differentiable_on S\ \ (\z. f z + g z) differentiable_on S" + by (simp add: differentiable_on_def) lemma differentiable_on_diff [simp, derivative_intros]: - "\f differentiable_on S; g differentiable_on S\ \ (\z. f z - g z) differentiable_on S" -by (simp add: differentiable_on_def) + "\f differentiable_on S; g differentiable_on S\ \ (\z. f z - g z) differentiable_on S" + by (simp add: differentiable_on_def) lemma differentiable_on_inverse [simp, derivative_intros]: fixes f :: "'a :: real_normed_vector \ 'b :: real_normed_field" shows "f differentiable_on S \ (\x. x \ S \ f x \ 0) \ (\x. inverse (f x)) differentiable_on S" -by (simp add: differentiable_on_def) + by (simp add: differentiable_on_def) lemma differentiable_on_scaleR [derivative_intros, simp]: - "\f differentiable_on S; g differentiable_on S\ \ (\x. f x *\<^sub>R g x) differentiable_on S" + "\f differentiable_on S; g differentiable_on S\ \ (\x. f x *\<^sub>R g x) differentiable_on S" unfolding differentiable_on_def by (blast intro: differentiable_scaleR) @@ -195,22 +195,22 @@ lemma differentiable_sqnorm_at [derivative_intros, simp]: fixes a :: "'a :: {real_normed_vector,real_inner}" shows "(\x. (norm x)\<^sup>2) differentiable (at a)" -by (force simp add: differentiable_def intro: has_derivative_sqnorm_at) + by (force simp add: differentiable_def intro: has_derivative_sqnorm_at) lemma differentiable_on_sqnorm [derivative_intros, simp]: fixes S :: "'a :: {real_normed_vector,real_inner} set" shows "(\x. (norm x)\<^sup>2) differentiable_on S" -by (simp add: differentiable_at_imp_differentiable_on) + by (simp add: differentiable_at_imp_differentiable_on) lemma differentiable_norm_at [derivative_intros, simp]: fixes a :: "'a :: {real_normed_vector,real_inner}" shows "a \ 0 \ norm differentiable (at a)" -using differentiableI has_derivative_norm by blast + using differentiableI has_derivative_norm by blast lemma differentiable_on_norm [derivative_intros, simp]: fixes S :: "'a :: {real_normed_vector,real_inner} set" shows "0 \ S \ norm differentiable_on S" -by (metis differentiable_at_imp_differentiable_on differentiable_norm_at) + by (metis differentiable_at_imp_differentiable_on differentiable_norm_at) subsection \Frechet derivative and Jacobian matrix\ @@ -480,12 +480,7 @@ assumes x: "x \ box a b" and f: "(f has_derivative f' ) (at x within box a b)" "(f has_derivative f'') (at x within box a b)" shows "f' = f''" -proof - - have "at x within box a b = at x" - by (metis x at_within_interior interior_open open_box) - with f show "f' = f''" - by (simp add: has_derivative_unique) -qed + by (metis at_within_open assms has_derivative_unique open_box) lemma frechet_derivative_at: "(f has_derivative f') (at x) \ f' = frechet_derivative f (at x)" @@ -613,7 +608,7 @@ shows "\x\{a..b}. f b - f a = f' x (b - a)" proof (cases "a = b") interpret bounded_linear "f' b" - using assms(2) assms(1) by auto + using assms by auto case True then show ?thesis by force @@ -944,10 +939,12 @@ and dn: "\z. z \ S \ norm (f' z) \ B" and "x \ S" "y \ S" shows "norm(f x - f y) \ B * norm(x - y)" - apply (rule differentiable_bound [OF cvs]) - apply (erule df [unfolded has_field_derivative_def]) - apply (rule onorm_le, simp_all add: norm_mult mult_right_mono assms) - done +proof (rule differentiable_bound [OF cvs]) + show "\x. x \ S \ (f has_derivative (*) (f' x)) (at x within S)" + by (simp add: df has_field_derivative_imp_has_derivative) + show "\x. x \ S \ onorm ((*) (f' x)) \ B" + by (metis (no_types, opaque_lifting) dn norm_mult onorm_le order.refl order_trans) +qed (use assms in auto) lemma differentiable_bound_segment: diff -r ba2afdd29e1d -r 032a4344903e src/HOL/Analysis/Elementary_Metric_Spaces.thy --- a/src/HOL/Analysis/Elementary_Metric_Spaces.thy Tue Aug 01 11:27:55 2023 +0200 +++ b/src/HOL/Analysis/Elementary_Metric_Spaces.thy Thu Aug 03 19:10:43 2023 +0200 @@ -1090,8 +1090,7 @@ proposition compact_eq_Bolzano_Weierstrass: fixes S :: "'a::metric_space set" shows "compact S \ (\T. infinite T \ T \ S \ (\x \ S. x islimpt T))" - using Bolzano_Weierstrass_imp_seq_compact Heine_Borel_imp_Bolzano_Weierstrass compact_eq_seq_compact_metric - by blast + by (meson Bolzano_Weierstrass_imp_seq_compact Heine_Borel_imp_Bolzano_Weierstrass seq_compact_imp_Heine_Borel) proposition Bolzano_Weierstrass_imp_bounded: "(\T. \infinite T; T \ S\ \ (\x \ S. x islimpt T)) \ bounded S" @@ -1511,10 +1510,8 @@ using bounded_imp_convergent_subsequence [OF \bounded (range f)\] by auto from f have fr: "\n. (f \ r) n \ S" by simp - have "l \ S" using \closed S\ fr l - by (rule closed_sequentially) - show "\l\S. \r. strict_mono r \ ((f \ r) \ l) sequentially" - using \l \ S\ r l by blast + show "\l\S. \r. strict_mono r \ (f \ r) \ l" + using assms(2) closed_sequentially fr l r by blast qed lemma compact_eq_bounded_closed: @@ -2829,7 +2826,7 @@ moreover have "(\i. (x \ r) (i + n)) \ l" using lr(3) by (rule LIMSEQ_ignore_initial_segment) ultimately show "l \ S n" - by (rule closed_sequentially) + by (metis closed_sequentially) qed then show ?thesis using that by blast diff -r ba2afdd29e1d -r 032a4344903e src/HOL/Analysis/Elementary_Normed_Spaces.thy --- a/src/HOL/Analysis/Elementary_Normed_Spaces.thy Tue Aug 01 11:27:55 2023 +0200 +++ b/src/HOL/Analysis/Elementary_Normed_Spaces.thy Thu Aug 03 19:10:43 2023 +0200 @@ -1776,9 +1776,9 @@ by blast have "?thesis1 \ ?thesis2 \ ?thesis3" apply (rule *) - using continuous_disconnected_range_constant apply metis - apply clarify - apply (frule discrete_subset_disconnected; blast) + using continuous_disconnected_range_constant + apply (metis image_subset_iff_funcset) + apply (smt (verit, best) discrete_subset_disconnected mem_Collect_eq subsetD subsetI) apply (blast dest: finite_implies_discrete) apply (blast intro!: finite_range_constant_imp_connected) done diff -r ba2afdd29e1d -r 032a4344903e src/HOL/Analysis/Elementary_Topology.thy --- a/src/HOL/Analysis/Elementary_Topology.thy Tue Aug 01 11:27:55 2023 +0200 +++ b/src/HOL/Analysis/Elementary_Topology.thy Thu Aug 03 19:10:43 2023 +0200 @@ -53,13 +53,14 @@ lemma topological_basis: "topological_basis B \ (\x. open x \ (\B'. B' \ B \ \B' = x))" - unfolding topological_basis_def - apply safe - apply fastforce - apply fastforce - apply (erule_tac x=x in allE, simp) - apply (rule_tac x="{x}" in exI, auto) - done + (is "?lhs = ?rhs") +proof + show "?lhs \ ?rhs" + by (meson local.open_Union subsetD topological_basis_def) + show "?rhs \ ?lhs" + unfolding topological_basis_def + by (metis cSup_singleton empty_subsetI insert_subset) +qed lemma topological_basis_iff: assumes "\B'. B' \ B \ open B'" @@ -89,7 +90,7 @@ assumes "\B'. B' \ B \ open B'" and "\O' x. open O' \ x \ O' \ \B'\B. x \ B' \ B' \ O'" shows "topological_basis B" - using assms by (subst topological_basis_iff) auto + by (simp add: assms topological_basis_iff) lemma topological_basisE: fixes O' @@ -97,17 +98,10 @@ and "open O'" and "x \ O'" obtains B' where "B' \ B" "x \ B'" "B' \ O'" -proof atomize_elim - from assms have "\B'. B'\B \ open B'" - by (simp add: topological_basis_def) - with topological_basis_iff assms - show "\B'. B' \ B \ x \ B' \ B' \ O'" - using assms by (simp add: Bex_def) -qed + by (metis assms topological_basis_def topological_basis_iff) lemma topological_basis_open: - assumes "topological_basis B" - and "X \ B" + assumes "topological_basis B" and "X \ B" shows "open X" using assms by (simp add: topological_basis_def) @@ -131,17 +125,9 @@ lemma basis_dense: fixes B :: "'a set set" and f :: "'a set \ 'a" - assumes "topological_basis B" - and choosefrom_basis: "\B'. B' \ {} \ f B' \ B'" + assumes "topological_basis B" and "\B'. B' \ {} \ f B' \ B'" shows "\X. open X \ X \ {} \ (\B' \ B. f B' \ X)" -proof (intro allI impI) - fix X :: "'a set" - assume "open X" and "X \ {}" - from topological_basisE[OF \topological_basis B\ \open X\ choosefrom_basis[OF \X \ {}\]] - obtain B' where "B' \ B" "f X \ B'" "B' \ X" . - then show "\B'\B. f B' \ X" - by (auto intro!: choosefrom_basis) -qed + by (metis assms equals0D in_mono topological_basisE) end @@ -149,27 +135,31 @@ assumes A: "topological_basis A" and B: "topological_basis B" shows "topological_basis ((\(a, b). a \ b) ` (A \ B))" - unfolding topological_basis_def -proof (safe, simp_all del: ex_simps add: subset_image_iff ex_simps(1)[symmetric]) - fix S :: "('a \ 'b) set" - assume "open S" - then show "\X\A \ B. (\(a,b)\X. a \ b) = S" - proof (safe intro!: exI[of _ "{x\A \ B. fst x \ snd x \ S}"]) - fix x y - assume "(x, y) \ S" - from open_prod_elim[OF \open S\ this] - obtain a b where a: "open a""x \ a" and b: "open b" "y \ b" and "a \ b \ S" - by (metis mem_Sigma_iff) - moreover - from A a obtain A0 where "A0 \ A" "x \ A0" "A0 \ a" - by (rule topological_basisE) - moreover - from B b obtain B0 where "B0 \ B" "y \ B0" "B0 \ b" - by (rule topological_basisE) - ultimately show "(x, y) \ (\(a, b)\{X \ A \ B. fst X \ snd X \ S}. a \ b)" - by (intro UN_I[of "(A0, B0)"]) auto - qed auto -qed (metis A B topological_basis_open open_Times) +proof - + have "\X\A \ B. (\(a,b)\X. a \ b) = S" if "open S" for S + proof - + have "(x, y) \ (\(a, b)\{X \ A \ B. fst X \ snd X \ S}. a \ b)" + if xy: "(x, y) \ S" for x y + proof - + obtain a b where a: "open a""x \ a" and b: "open b" "y \ b" and "a \ b \ S" + by (metis open_prod_elim[OF \open S\] xy mem_Sigma_iff) + moreover obtain A0 where "A0 \ A" "x \ A0" "A0 \ a" + using A a b topological_basisE by blast + moreover + from B b obtain B0 where "B0 \ B" "y \ B0" "B0 \ b" + by (rule topological_basisE) + ultimately show ?thesis + by (intro UN_I[of "(A0, B0)"]) auto + qed + then have "(\(a, b)\{x \ A \ B. fst x \ snd x \ S}. a \ b) = S" + by force + then show ?thesis + using subset_eq by force + qed + with A B show ?thesis + unfolding topological_basis_def + by (smt (verit) SigmaE imageE image_mono open_Times case_prod_conv) +qed subsection \Countable Basis\ @@ -189,8 +179,7 @@ lemma open_countable_basisE: assumes "p X" obtains B' where "B' \ B" "X = \B'" - using assms open_countable_basis_ex - by atomize_elim simp + using assms open_countable_basis_ex by auto lemma countable_dense_exists: "\D::'a set. countable D \ (\X. p X \ X \ {} \ (\d \ D. d \ X))" @@ -221,11 +210,10 @@ proof - obtain \ where \: "(\i::nat. x \ \ i \ open (\ i))" "(\S. open S \ x \ S \ (\i. \ i \ S))" using first_countable_basis[of x] by metis - show thesis - proof - show "countable (range \)" - by simp - qed (use \ in auto) + moreover have "countable (range \)" + by simp + ultimately show thesis + by (smt (verit, best) imageE rangeI that) qed lemma (in first_countable_topology) first_countable_basis_Int_stableE: @@ -262,8 +250,14 @@ fix S assume "open S" "x \ S" then obtain a where a: "a\\" "a \ S" using \ by blast - then show "\a\\. a \ S" using a \ - by (intro bexI[where x=a]) (auto simp: \_def intro: image_eqI[where x="{to_nat_on \ a}"]) + moreover have "a\\" + unfolding \_def + proof (rule image_eqI) + show "a = \ (from_nat_into \ ` {to_nat_on \ a})" + by (simp add: \ a) + qed auto + ultimately show "\a\\. a \ S" + by blast qed qed @@ -279,9 +273,9 @@ using range_from_nat_into_subset[OF \\ \ {}\] 1 by auto next fix S - assume "open S" "x\S" from 2[OF this] - show "\i. from_nat_into \ i \ S" - using subset_range_from_nat_into[OF \countable \\] by auto + assume "open S" "x\S" + then show "\i. from_nat_into \ i \ S" + by (metis "2" \countable \\ from_nat_into_surj) qed instance prod :: (first_countable_topology, first_countable_topology) first_countable_topology @@ -314,9 +308,8 @@ then obtain a' b' where a'b': "open a'" "open b'" "x \ a' \ b'" "a' \ b' \ S" by (rule open_prod_elim) moreover - from a'b' \(4)[of a'] B(4)[of b'] obtain a b where "a \ \" "a \ a'" "b \ B" "b \ b'" - by auto + by (meson B(4) \(4) a'b' mem_Times_iff) ultimately show "\a\(\(a, b). a \ b) ` (\ \ B). a \ S" by (auto intro!: bexI[of _ "a \ b"] bexI[of _ a] bexI[of _ b]) @@ -396,9 +389,7 @@ using univ_second_countable by blast define \ where "\ \ {S. S \ \ \ (\U. U \ \ \ S \ U)}" have "countable \" - apply (rule countable_subset [OF _ \countable \\]) - apply (force simp: \_def) - done + by (simp add: \_def \countable \\) have "\S. \U. S \ \ \ U \ \ \ S \ U" by (simp add: \_def) then obtain G where G: "\S. S \ \ \ G S \ \ \ S \ G S" @@ -408,12 +399,10 @@ moreover have "\\ \ \\" using \_def by blast ultimately have eq1: "\\ = \\" .. - have eq2: "\\ = \ (G ` \)" + moreover have "\\ = \ (G ` \)" using G eq1 by auto - show ?thesis - apply (rule_tac \' = "G ` \" in that) - using G \countable \\ - by (auto simp: eq1 eq2) + ultimately show ?thesis + by (metis G \countable \\ countable_image image_subset_iff that) qed lemma countable_disjoint_open_subsets: @@ -606,11 +595,8 @@ unfolding islimpt_UNIV_iff by (rule not_open_singleton) lemma closed_limpt: "closed S \ (\x. x islimpt S \ x \ S)" - unfolding closed_def - apply (subst open_subopen) - apply (simp add: islimpt_def subset_eq) - apply (metis ComplE ComplI) - done + unfolding closed_def open_subopen [of "-S"] + by (metis Compl_iff islimptE islimptI open_subopen subsetI) lemma islimpt_EMPTY[simp]: "\ x islimpt {}" by (auto simp: islimpt_def) @@ -625,25 +611,25 @@ lemma islimpt_insert: fixes x :: "'a::t1_space" - shows "x islimpt (insert a s) \ x islimpt s" + shows "x islimpt (insert a S) \ x islimpt S" proof - assume "x islimpt (insert a s)" - then show "x islimpt s" + assume "x islimpt (insert a S)" + then show "x islimpt S" by (metis closed_limpt closed_singleton empty_iff insert_iff insert_is_Un islimpt_Un islimpt_def) next - assume "x islimpt s" - then show "x islimpt (insert a s)" + assume "x islimpt S" + then show "x islimpt (insert a S)" by (rule islimpt_subset) auto qed lemma islimpt_finite: fixes x :: "'a::t1_space" - shows "finite s \ \ x islimpt s" + shows "finite S \ \ x islimpt S" by (induct set: finite) (simp_all add: islimpt_insert) lemma islimpt_Un_finite: fixes x :: "'a::t1_space" - shows "finite s \ x islimpt (s \ t) \ x islimpt t" + shows "finite S \ x islimpt (S \ T) \ x islimpt T" by (simp add: islimpt_Un islimpt_finite) lemma islimpt_eq_acc_point: @@ -659,10 +645,8 @@ next fix T assume *: "\U. l\U \ open U \ infinite (U \ S)" "l \ T" "open T" - then have "infinite (T \ S - {l})" - by auto then have "\x. x \ (T \ S - {l})" - unfolding ex_in_conv by (intro notI) simp + by (metis ex_in_conv finite.emptyI infinite_remove) then show "\y\S. y \ T \ y \ l" by auto qed @@ -684,11 +668,9 @@ have "infinite (A (Suc n) \ range f - f`{.. i})" using l A by auto then have "\x. x \ A (Suc n) \ range f - f`{.. i}" - unfolding ex_in_conv by (intro notI) simp - then have "\j. f j \ A (Suc n) \ j \ {.. i}" - by auto + by (metis all_not_in_conv finite.emptyI) then have "\a. i < a \ f a \ A (Suc n)" - by (auto simp: not_le) + by (force simp: linorder_not_le) then have "i < s n i" "f (s n i) \ A (Suc n)" unfolding s_def by (auto intro: someI2_ex) } @@ -728,17 +710,15 @@ lemma sequence_unique_limpt: fixes f :: "nat \ 'a::t2_space" - assumes "(f \ l) sequentially" + assumes f: "(f \ l) sequentially" and "l' islimpt (range f)" shows "l' = l" proof (rule ccontr) assume "l' \ l" obtain s t where "open s" "open t" "l' \ s" "l \ t" "s \ t = {}" using hausdorff [OF \l' \ l\] by auto - have "eventually (\n. f n \ t) sequentially" - using assms(1) \open t\ \l \ t\ by (rule topological_tendstoD) then obtain N where "\n\N. f n \ t" - unfolding eventually_sequentially by auto + by (meson f lim_explicit) have "UNIV = {.. {N..}" by auto @@ -752,10 +732,8 @@ using \l' \ s\ \open s\ by (rule islimptE) then obtain n where "N \ n" "f n \ s" "f n \ l'" by auto - with \\n\N. f n \ t\ have "f n \ s \ t" - by simp - with \s \ t = {}\ show False - by simp + with \\n\N. f n \ t\ \s \ t = {}\ show False + by blast qed (*could prove directly from islimpt_sequential_inj, but only for metric spaces*) @@ -826,18 +804,9 @@ qed lemma islimpt_image: - assumes "z islimpt g -` A \ B" "g z \ A" "z \ B" "open B" "continuous_on B g" + assumes "z islimpt g -` A \ B" "g z \ A" "z \ B" "continuous_on B g" shows "g z islimpt A" - unfolding islimpt_def -proof clarify - fix T assume T: "g z \ T" "open T" - have "z \ g -` T \ B" - using T assms by auto - moreover have "open (g -` T \ B)" - using T continuous_on_open_vimage assms by blast - ultimately show "\y\A. y \ T \ y \ g z" - using assms by (metis (mono_tags, lifting) IntD1 islimptE vimageE) -qed + by (smt (verit, best) IntD1 assms continuous_on_topological inf_le2 islimpt_def subset_eq vimageE) subsection \Interior of a Set\ @@ -905,12 +874,12 @@ fixes x :: "'a::perfect_space" assumes x: "x \ interior S" shows "x islimpt S" - using x islimpt_UNIV [of x] - unfolding interior_def islimpt_def - apply (clarsimp, rename_tac T T') - apply (drule_tac x="T \ T'" in spec) - apply (auto simp: open_Int) - done +proof - + obtain T where "x \ T" "T \ S" "open T" + using interior_subset x by blast + with x islimpt_UNIV [of x] show ?thesis + unfolding islimpt_def by (metis (full_types) Int_iff open_Int subsetD) +qed lemma open_imp_islimpt: fixes x::"'a:: perfect_space" @@ -950,12 +919,8 @@ assume "x \ interior S" with \x \ R\ \open R\ obtain y where "y \ R - S" unfolding interior_def by fast - from \open R\ \closed S\ have "open (R - S)" - by (rule open_Diff) - from \R \ S \ T\ have "R - S \ T" - by fast - from \y \ R - S\ \open (R - S)\ \R - S \ T\ \interior T = {}\ show False - unfolding interior_def by fast + then show False + by (metis Diff_subset_conv \R \ S \ T\ \open R\ cS empty_iff iT interiorI open_Diff) qed qed qed @@ -1039,7 +1004,7 @@ subsection \Closure of a Set\ definition\<^marker>\tag important\ closure :: "('a::topological_space) set \ 'a set" where -"closure S = S \ {x . x islimpt S}" + "closure S = S \ {x . x islimpt S}" lemma interior_closure: "interior S = - (closure (- S))" by (auto simp: interior_def closure_def islimpt_def) @@ -1102,18 +1067,8 @@ proof fix x assume *: "open S" "x \ S \ closure T" - have "x islimpt (S \ T)" if **: "x islimpt T" - proof (rule islimptI) - fix A - assume "x \ A" "open A" - with * have "x \ A \ S" "open (A \ S)" - by (simp_all add: open_Int) - with ** obtain y where "y \ T" "y \ A \ S" "y \ x" - by (rule islimptE) - then have "y \ S \ T" "y \ A \ y \ x" - by simp_all - then show "\y\(S \ T). y \ A \ y \ x" .. - qed + then have "x islimpt (S \ T)" if "x islimpt T" + by (metis IntD1 eventually_at_in_open' inf_commute islimpt_Int_eventually that) with * show "x \ closure (S \ T)" unfolding closure_def by blast qed @@ -1148,35 +1103,10 @@ lemma closure_open_Int_superset: assumes "open S" "S \ closure T" shows "closure(S \ T) = closure S" -proof - - have "closure S \ closure(S \ T)" - by (metis assms closed_closure closure_minimal inf.orderE open_Int_closure_subset) - then show ?thesis - by (simp add: closure_mono dual_order.antisym) -qed + by (metis Int_Un_distrib Un_Int_eq(4) assms closure_Un closure_closure open_Int_closure_subset sup.orderE) -lemma closure_Int: "closure (\I) \ \{closure S |S. S \ I}" -proof - - { - fix y - assume "y \ \I" - then have y: "\S \ I. y \ S" by auto - { - fix S - assume "S \ I" - then have "y \ closure S" - using closure_subset y by auto - } - then have "y \ \{closure S |S. S \ I}" - by auto - } - then have "\I \ \{closure S |S. S \ I}" - by auto - moreover have "closed (\{closure S |S. S \ I})" - unfolding closed_Inter closed_closure by auto - ultimately show ?thesis using closure_hull[of "\I"] - hull_minimal[of "\I" "\{closure S |S. S \ I}" "closed"] by auto -qed +lemma closure_Int: "closure (\I) \ \{closure S |S. S \ I}" + by (simp add: INF_greatest Inter_lower Setcompr_eq_image closure_mono) lemma islimpt_in_closure: "(x islimpt S) = (x\closure(S-{x}))" unfolding closure_def using islimpt_punctured by blast @@ -1202,7 +1132,7 @@ subsection \Frontier (also known as boundary)\ definition\<^marker>\tag important\ frontier :: "('a::topological_space) set \ 'a set" where -"frontier S = closure S - interior S" + "frontier S = closure S - interior S" lemma frontier_closed [iff]: "closed (frontier S)" by (simp add: frontier_def closed_Diff) @@ -1224,14 +1154,7 @@ lemma frontier_Int_closed: assumes "closed S" "closed T" shows "frontier(S \ T) = (frontier S \ T) \ (S \ frontier T)" -proof - - have "closure (S \ T) = T \ S" - using assms by (simp add: Int_commute closed_Int) - moreover have "T \ (closure S \ closure (- S)) = frontier S \ T" - by (simp add: Int_commute frontier_closures) - ultimately show ?thesis - by (simp add: Int_Un_distrib Int_assoc Int_left_commute assms frontier_closures) -qed + by (smt (verit, best) Diff_Int Int_Diff assms closed_Int closure_eq frontier_def inf_commute interior_Int) lemma frontier_subset_closed: "closed S \ frontier S \ S" by (metis frontier_def closure_closed Diff_subset) @@ -1240,16 +1163,7 @@ by (simp add: frontier_def) lemma frontier_subset_eq: "frontier S \ S \ closed S" -proof - - { - assume "frontier S \ S" - then have "closure S \ S" - using interior_subset unfolding frontier_def by auto - then have "closed S" - using closure_subset_eq by auto - } - then show ?thesis using frontier_subset_closed[of S] .. -qed + by (metis Diff_subset_conv closure_subset_eq frontier_def interior_subset subset_Un_eq) lemma frontier_complement [simp]: "frontier (- S) = frontier S" by (auto simp: frontier_def closure_complement interior_complement) @@ -1271,12 +1185,7 @@ by (simp add: Diff_mono frontier_interiors interior_mono interior_subset) lemma closure_Un_frontier: "closure S = S \ frontier S" -proof - - have "S \ interior S = S" - using interior_subset by auto - then show ?thesis - using closure_subset by (auto simp: frontier_def) -qed + by (simp add: closure_def frontier_closures sup_inf_distrib1) subsection\<^marker>\tag unimportant\ \Filters and the ``eventually true'' quantifier\ @@ -1284,22 +1193,8 @@ text \Identify Trivial limits, where we can't approach arbitrarily closely.\ lemma trivial_limit_within: "trivial_limit (at a within S) \ \ a islimpt S" -proof - assume "trivial_limit (at a within S)" - then show "\ a islimpt S" - unfolding trivial_limit_def - unfolding eventually_at_topological - unfolding islimpt_def - apply (clarsimp simp add: set_eq_iff) - apply (rename_tac T, rule_tac x=T in exI) - apply (clarsimp, drule_tac x=y in bspec, simp_all) - done -next - assume "\ a islimpt S" - then show "trivial_limit (at a within S)" unfolding trivial_limit_def eventually_at_topological islimpt_def - by metis -qed + by blast lemma trivial_limit_at_iff: "trivial_limit (at a) \ \ a islimpt UNIV" using trivial_limit_within [of a UNIV] by simp @@ -1312,13 +1207,13 @@ using islimpt_in_closure by (metis trivial_limit_within) lemma not_in_closure_trivial_limitI: - "x \ closure s \ trivial_limit (at x within s)" - using not_trivial_limit_within[of x s] + "x \ closure S \ trivial_limit (at x within S)" + using not_trivial_limit_within[of x S] by safe (metis Diff_empty Diff_insert0 closure_subset contra_subsetD) -lemma filterlim_at_within_closure_implies_filterlim: "filterlim f l (at x within s)" - if "x \ closure s \ filterlim f l (at x within s)" - by (metis bot.extremum filterlim_filtercomap filterlim_mono not_in_closure_trivial_limitI that) +lemma filterlim_at_within_closure_implies_filterlim: "filterlim f l (at x within S)" + if "x \ closure S \ filterlim f l (at x within S)" + by (metis bot.extremum filterlim_iff_le_filtercomap not_in_closure_trivial_limitI that) lemma at_within_eq_bot_iff: "at c within A = bot \ c \ closure (A - {c})" using not_trivial_limit_within[of c A] by blast @@ -1345,23 +1240,7 @@ lemma eventually_within_interior: assumes "x \ interior S" shows "eventually P (at x within S) \ eventually P (at x)" - (is "?lhs = ?rhs") -proof - from assms obtain T where T: "open T" "x \ T" "T \ S" .. - { - assume ?lhs - then obtain A where "open A" and "x \ A" and "\y\A. y \ x \ y \ S \ P y" - by (auto simp: eventually_at_topological) - with T have "open (A \ T)" and "x \ A \ T" and "\y \ A \ T. y \ x \ P y" - by auto - then show ?rhs - by (auto simp: eventually_at_topological) - next - assume ?rhs - then show ?lhs - by (auto elim: eventually_mono simp: eventually_at_filter) - } -qed + by (metis assms at_within_open_subset interior_subset open_interior) lemma at_within_interior: "x \ interior S \ at x within S = at x" unfolding filter_eq_iff by (intro allI eventually_within_interior) @@ -1443,25 +1322,7 @@ lemma closure_sequential: fixes l :: "'a::first_countable_topology" shows "l \ closure S \ (\x. (\n. x n \ S) \ (x \ l) sequentially)" - (is "?lhs = ?rhs") -proof - assume "?lhs" - moreover - { - assume "l \ S" - then have "?rhs" using tendsto_const[of l sequentially] by auto - } - moreover - { - assume "l islimpt S" - then have "?rhs" unfolding islimpt_sequential by auto - } - ultimately show "?rhs" - unfolding closure_def by auto -next - assume "?rhs" - then show "?lhs" unfolding closure_def islimpt_sequential by auto -qed + by (metis Diff_empty Diff_insert0 Un_iff closure_def islimpt_sequential mem_Collect_eq tendsto_const) lemma closed_sequential_limits: fixes S :: "'a::first_countable_topology set" @@ -1469,20 +1330,20 @@ by (metis closure_sequential closure_subset_eq subset_iff) lemma tendsto_If_within_closures: - assumes f: "x \ s \ (closure s \ closure t) \ - (f \ l x) (at x within s \ (closure s \ closure t))" - assumes g: "x \ t \ (closure s \ closure t) \ - (g \ l x) (at x within t \ (closure s \ closure t))" - assumes "x \ s \ t" - shows "((\x. if x \ s then f x else g x) \ l x) (at x within s \ t)" + assumes f: "x \ S \ (closure S \ closure T) \ + (f \ l x) (at x within S \ (closure S \ closure T))" + assumes g: "x \ T \ (closure S \ closure T) \ + (g \ l x) (at x within T \ (closure S \ closure T))" + assumes "x \ S \ T" + shows "((\x. if x \ S then f x else g x) \ l x) (at x within S \ T)" proof - - have *: "(s \ t) \ {x. x \ s} = s" "(s \ t) \ {x. x \ s} = t - s" + have *: "(S \ T) \ {x. x \ S} = S" "(S \ T) \ {x. x \ S} = T - S" by auto - have "(f \ l x) (at x within s)" + have "(f \ l x) (at x within S)" by (rule filterlim_at_within_closure_implies_filterlim) (use \x \ _\ in \auto simp: inf_commute closure_def intro: tendsto_within_subset[OF f]\) moreover - have "(g \ l x) (at x within t - s)" + have "(g \ l x) (at x within T - S)" by (rule filterlim_at_within_closure_implies_filterlim) (use \x \ _\ in \auto intro!: tendsto_within_subset[OF g] simp: closure_def intro: islimpt_subset\) @@ -1495,76 +1356,50 @@ lemma brouwer_compactness_lemma: fixes f :: "'a::topological_space \ 'b::real_normed_vector" - assumes "compact s" - and "continuous_on s f" - and "\ (\x\s. f x = 0)" - obtains d where "0 < d" and "\x\s. d \ norm (f x)" -proof (cases "s = {}") + assumes "compact S" + and "continuous_on S f" + and "\ (\x\S. f x = 0)" + obtains d where "0 < d" and "\x\S. d \ norm (f x)" +proof (cases "S = {}") case True show thesis by (rule that [of 1]) (auto simp: True) next case False - have "continuous_on s (norm \ f)" + have "continuous_on S (norm \ f)" by (rule continuous_intros continuous_on_norm assms(2))+ - with False obtain x where x: "x \ s" "\y\s. (norm \ f) x \ (norm \ f) y" + with False obtain x where x: "x \ S" "\y\S. (norm \ f) x \ (norm \ f) y" using continuous_attains_inf[OF assms(1), of "norm \ f"] unfolding o_def by auto - have "(norm \ f) x > 0" - using assms(3) and x(1) - by auto then show ?thesis - by (rule that) (insert x(2), auto simp: o_def) + by (metis assms(3) that comp_apply zero_less_norm_iff) qed subsubsection \Bolzano-Weierstrass property\ proposition Heine_Borel_imp_Bolzano_Weierstrass: - assumes "compact s" - and "infinite t" - and "t \ s" - shows "\x \ s. x islimpt t" + assumes "compact S" + and "infinite T" + and "T \ S" + shows "\x \ S. x islimpt T" proof (rule ccontr) - assume "\ (\x \ s. x islimpt t)" - then obtain f where f: "\x\s. x \ f x \ open (f x) \ (\y\t. y \ f x \ y = x)" - unfolding islimpt_def - using bchoice[of s "\ x T. x \ T \ open T \ (\y\t. y \ T \ y = x)"] - by auto - obtain g where g: "g \ {t. \x. x \ s \ t = f x}" "finite g" "s \ \g" - using assms(1)[unfolded compact_eq_Heine_Borel, THEN spec[where x="{t. \x. x\s \ t = f x}"]] + assume "\ (\x \ S. x islimpt T)" + then obtain f where f: "\x\S. x \ f x \ open (f x) \ (\y\T. y \ f x \ y = x)" + unfolding islimpt_def by metis + obtain g where g: "g \ {T. \x. x \ S \ T = f x}" "finite g" "S \ \g" + using assms(1)[unfolded compact_eq_Heine_Borel, THEN spec[where x="{T. \x. x\S \ T = f x}"]] using f by auto - from g(1,3) have g':"\x\g. \xa \ s. x = f xa" + then have g': "\x\g. \y \ S. x = f y" by auto - { - fix x y - assume "x \ t" "y \ t" "f x = f y" - then have "x \ f x" "y \ f x \ y = x" - using f[THEN bspec[where x=x]] and \t \ s\ by auto - then have "x = y" - using \f x = f y\ and f[THEN bspec[where x=y]] and \y \ t\ and \t \ s\ - by auto - } - then have "inj_on f t" - unfolding inj_on_def by simp - then have "infinite (f ` t)" + have "inj_on f T" + by (smt (verit, best) assms(3) f inj_onI subsetD) + then have "infinite (f ` T)" using assms(2) using finite_imageD by auto moreover - { - fix x - assume "x \ t" "f x \ g" - from g(3) assms(3) \x \ t\ obtain h where "h \ g" and "x \ h" - by auto - then obtain y where "y \ s" "h = f y" - using g'[THEN bspec[where x=h]] by auto - then have "y = x" - using f[THEN bspec[where x=y]] and \x\t\ and \x\h\[unfolded \h = f y\] - by auto - then have False - using \f x \ g\ \h \ g\ unfolding \h = f y\ - by auto - } - then have "f ` t \ g" by auto + have False if "x \ T" "f x \ g" for x + by (smt (verit) UnionE assms(3) f g' g(3) subsetD that) + then have "f ` T \ g" by auto ultimately show False using g(2) using finite_subset by auto qed @@ -1585,25 +1420,24 @@ qed lemma Bolzano_Weierstrass_imp_closed: - fixes s :: "'a::{first_countable_topology,t2_space} set" - assumes "\t. infinite t \ t \ s --> (\x \ s. x islimpt t)" - shows "closed s" + fixes S :: "'a::{first_countable_topology,t2_space} set" + assumes "\T. infinite T \ T \ S --> (\x \ S. x islimpt T)" + shows "closed S" proof - { fix x l - assume as: "\n::nat. x n \ s" "(x \ l) sequentially" - then have "l \ s" + assume as: "\n::nat. x n \ S" "(x \ l) sequentially" + have "l \ S" proof (cases "\n. x n \ l") case False - then show "l\s" using as(1) by auto + then show "l\S" using as(1) by auto next - case True note cas = this + case True with as(2) have "infinite (range x)" using sequence_infinite_lemma[of x l] by auto - then obtain l' where "l'\s" "l' islimpt (range x)" - using assms[THEN spec[where x="range x"]] as(1) by auto - then show "l\s" using sequence_unique_limpt[of x l l'] - using as cas by auto + then obtain l' where "l'\S" "l' islimpt (range x)" + using as(1) assms by auto + then show "l\S" using sequence_unique_limpt as True by auto qed } then show ?thesis @@ -1612,38 +1446,30 @@ lemma closure_insert: fixes x :: "'a::t1_space" - shows "closure (insert x s) = insert x (closure s)" - by (meson closed_closure closed_insert closure_minimal closure_subset dual_order.antisym insert_mono insert_subset) + shows "closure (insert x S) = insert x (closure S)" + by (metis closed_singleton closure_Un closure_closed insert_is_Un) lemma finite_not_islimpt_in_compact: assumes "compact A" "\z. z \ A \ \z islimpt B" shows "finite (A \ B)" -proof (rule ccontr) - assume "infinite (A \ B)" - have "\z\A. z islimpt A \ B" - by (rule Heine_Borel_imp_Bolzano_Weierstrass) (use assms \infinite _\ in auto) - hence "\z\A. z islimpt B" - using islimpt_subset by blast - thus False using assms(2) - by simp -qed + by (meson Heine_Borel_imp_Bolzano_Weierstrass assms inf_le1 inf_le2 islimpt_subset) text\In particular, some common special cases.\ lemma compact_Un [intro]: - assumes "compact s" - and "compact t" - shows " compact (s \ t)" + assumes "compact S" + and "compact T" + shows " compact (S \ T)" proof (rule compactI) fix f - assume *: "Ball f open" "s \ t \ \f" - from * \compact s\ obtain s' where "s' \ f \ finite s' \ s \ \s'" + assume *: "Ball f open" "S \ T \ \f" + from * \compact S\ obtain s' where "s' \ f \ finite s' \ S \ \s'" unfolding compact_eq_Heine_Borel by (auto elim!: allE[of _ f]) moreover - from * \compact t\ obtain t' where "t' \ f \ finite t' \ t \ \t'" + from * \compact T\ obtain t' where "t' \ f \ finite t' \ T \ \t'" unfolding compact_eq_Heine_Borel by (auto elim!: allE[of _ f]) - ultimately show "\f'\f. finite f' \ s \ t \ \f'" + ultimately show "\f'\f. finite f' \ S \ T \ \f'" by (auto intro!: exI[of _ "s' \ t'"]) qed @@ -1652,40 +1478,34 @@ lemma compact_UN [intro]: "finite A \ (\x. x \ A \ compact (B x)) \ compact (\x\A. B x)" - by (rule compact_Union) auto + by blast lemma closed_Int_compact [intro]: - assumes "closed s" - and "compact t" - shows "compact (s \ t)" - using compact_Int_closed [of t s] assms + assumes "closed S" and "compact T" + shows "compact (S \ T)" + using compact_Int_closed [of T S] assms by (simp add: Int_commute) lemma compact_Int [intro]: - fixes s t :: "'a :: t2_space set" - assumes "compact s" - and "compact t" - shows "compact (s \ t)" + fixes S T :: "'a :: t2_space set" + assumes "compact S" and "compact T" + shows "compact (S \ T)" using assms by (intro compact_Int_closed compact_imp_closed) lemma compact_sing [simp]: "compact {a}" unfolding compact_eq_Heine_Borel by auto lemma compact_insert [simp]: - assumes "compact s" - shows "compact (insert x s)" -proof - - have "compact ({x} \ s)" - using compact_sing assms by (rule compact_Un) - then show ?thesis by simp -qed + assumes "compact S" + shows "compact (insert x S)" + by (metis assms compact_Un compact_sing insert_is_Un) -lemma finite_imp_compact: "finite s \ compact s" +lemma finite_imp_compact: "finite S \ compact S" by (induct set: finite) simp_all lemma open_delete: - fixes s :: "'a::t1_space set" - shows "open s \ open (s - {x})" + fixes S :: "'a::t1_space set" + shows "open S \ open (S - {x})" by (simp add: open_Diff) @@ -1696,20 +1516,17 @@ proof safe assume x: "x \ closure X" fix S A - assume "open S" "x \ S" "X \ A = {}" "S \ A" + assume \
: "open S" "x \ S" "X \ A = {}" "S \ A" then have "x \ closure (-S)" - by (auto simp: closure_complement subset_eq[symmetric] intro: interiorI) + by (simp add: closed_open) with x have "x \ closure X - closure (-S)" by auto - also have "\ \ closure (X \ S)" - using \open S\ open_Int_closure_subset[of S X] by (simp add: closed_Compl ac_simps) - finally have "X \ S \ {}" by auto - then show False using \X \ A = {}\ \S \ A\ by auto + with \
show False + by (metis Compl_iff Diff_iff closure_mono disjoint_iff subsetD subsetI) next assume "\A S. S \ A \ open S \ x \ S \ X \ A \ {}" - from this[THEN spec, of "- X", THEN spec, of "- closure X"] - show "x \ closure X" - by (simp add: closure_subset open_Compl) + then show "x \ closure X" + by (metis ComplI Compl_disjoint closure_interior interior_subset open_interior) qed lemma compact_filter: @@ -1745,7 +1562,7 @@ proof safe fix P Q R S assume "eventually R F" "open S" "x \ S" - with open_Int_closure_eq_empty[of S "{x. R x}"] x[of "closure {x. R x}"] + with open_Int_closure_eq_empty[of S "{x. R x}"] x have "S \ {x. R x} \ {}" by (auto simp: Z_def) moreover assume "Ball S Q" "\x. Q x \ R x \ bot x" ultimately show False by (auto simp: set_eq_iff) @@ -1871,9 +1688,8 @@ subsubsection\Sequential compactness\ definition\<^marker>\tag important\ seq_compact :: "'a::topological_space set \ bool" where -"seq_compact S \ - (\f. (\n. f n \ S) - \ (\l\S. \r::nat\nat. strict_mono r \ ((f \ r) \ l) sequentially))" + "seq_compact S \ + (\f. (\n. f n \ S) \ (\l\S. \r::nat\nat. strict_mono r \ (f \ r) \ l))" lemma seq_compactI: assumes "\f. \n. f n \ S \ \l\S. \r::nat\nat. strict_mono r \ ((f \ r) \ l) sequentially" @@ -1886,38 +1702,30 @@ using assms unfolding seq_compact_def by fast lemma closed_sequentially: (* TODO: move upwards *) - assumes "closed s" and "\n. f n \ s" and "f \ l" - shows "l \ s" -proof (rule ccontr) - assume "l \ s" - with \closed s\ and \f \ l\ have "eventually (\n. f n \ - s) sequentially" - by (fast intro: topological_tendstoD) - with \\n. f n \ s\ show "False" - by simp -qed + assumes "closed S" and "\n. f n \ S" and "f \ l" + shows "l \ S" + by (metis Lim_in_closed_set assms eventually_sequentially trivial_limit_sequentially) lemma seq_compact_Int_closed: - assumes "seq_compact s" and "closed t" - shows "seq_compact (s \ t)" + assumes "seq_compact S" and "closed T" + shows "seq_compact (S \ T)" proof (rule seq_compactI) - fix f assume "\n::nat. f n \ s \ t" - hence "\n. f n \ s" and "\n. f n \ t" + fix f assume "\n::nat. f n \ S \ T" + hence "\n. f n \ S" and "\n. f n \ T" by simp_all - from \seq_compact s\ and \\n. f n \ s\ - obtain l r where "l \ s" and r: "strict_mono r" and l: "(f \ r) \ l" + from \seq_compact S\ and \\n. f n \ S\ + obtain l r where "l \ S" and r: "strict_mono r" and l: "(f \ r) \ l" by (rule seq_compactE) - from \\n. f n \ t\ have "\n. (f \ r) n \ t" + from \\n. f n \ T\ have "\n. (f \ r) n \ T" by simp - from \closed t\ and this and l have "l \ t" - by (rule closed_sequentially) - with \l \ s\ and r and l show "\l\s \ t. \r. strict_mono r \ (f \ r) \ l" - by fast + with \l \ S\ and r and l show "\l\S \ T. \r. strict_mono r \ (f \ r) \ l" + by (metis Int_iff \closed T\ closed_sequentially) qed lemma seq_compact_closed_subset: - assumes "closed s" and "s \ t" and "seq_compact t" - shows "seq_compact s" - using assms seq_compact_Int_closed [of t s] by (simp add: Int_absorb1) + assumes "closed S" and "S \ T" and "seq_compact T" + shows "seq_compact S" + using assms seq_compact_Int_closed [of T S] by (simp add: Int_absorb1) lemma seq_compact_imp_countably_compact: fixes U :: "'a :: first_countable_topology set" @@ -1951,9 +1759,9 @@ by auto from \x\U\ \U \ \A\ from_nat_into_surj[OF \countable A\] obtain n where "x \ from_nat_into A n" by auto - with r(2) A(1) from_nat_into[OF \A \ {}\, of n] + with r(2) A(1) from_nat_into[OF \A \ {}\] have "eventually (\i. X (r i) \ from_nat_into A n) sequentially" - unfolding tendsto_def by (auto simp: comp_def) + unfolding tendsto_def by fastforce then obtain N where "\i. N \ i \ X (r i) \ from_nat_into A n" by (auto simp: eventually_sequentially) moreover from X have "\i. n \ r i \ X (r i) \ from_nat_into A n" @@ -2037,95 +1845,81 @@ qed lemma countably_compact_imp_acc_point: - assumes "countably_compact s" - and "countable t" - and "infinite t" - and "t \ s" - shows "\x\s. \U. x\U \ open U \ infinite (U \ t)" + assumes "countably_compact S" + and "countable T" + and "infinite T" + and "T \ S" + shows "\x\S. \U. x\U \ open U \ infinite (U \ T)" proof (rule ccontr) - define C where "C = (\F. interior (F \ (- t))) ` {F. finite F \ F \ t }" - note \countably_compact s\ - moreover have "\t\C. open t" + define C where "C = (\F. interior (F \ (- T))) ` {F. finite F \ F \ T }" + note \countably_compact S\ + moreover have "\T\C. open T" by (auto simp: C_def) moreover - assume "\ (\x\s. \U. x\U \ open U \ infinite (U \ t))" - then have s: "\x. x \ s \ \U. x\U \ open U \ finite (U \ t)" by metis - have "s \ \C" - using \t \ s\ + assume "\ (\x\S. \U. x\U \ open U \ infinite (U \ T))" + then have S: "\x. x \ S \ \U. x\U \ open U \ finite (U \ T)" by metis + have "S \ \C" + using \T \ S\ unfolding C_def - apply (safe dest!: s) - apply (rule_tac a="U \ t" in UN_I) + apply (safe dest!: S) + apply (rule_tac a="U \ T" in UN_I) apply (auto intro!: interiorI simp add: finite_subset) done moreover - from \countable t\ have "countable C" + from \countable T\ have "countable C" unfolding C_def by (auto intro: countable_Collect_finite_subset) ultimately - obtain D where "D \ C" "finite D" "s \ \D" + obtain D where "D \ C" "finite D" "S \ \D" by (rule countably_compactE) - then obtain E where E: "E \ {F. finite F \ F \ t }" "finite E" - and s: "s \ (\F\E. interior (F \ (- t)))" + then obtain E where E: "E \ {F. finite F \ F \ T }" "finite E" + and S: "S \ (\F\E. interior (F \ (- T)))" by (metis (lifting) finite_subset_image C_def) - from s \t \ s\ have "t \ \E" + from S \T \ S\ have "T \ \E" using interior_subset by blast moreover have "finite (\E)" using E by auto - ultimately show False using \infinite t\ + ultimately show False using \infinite T\ by (auto simp: finite_subset) qed lemma countable_acc_point_imp_seq_compact: - fixes s :: "'a::first_countable_topology set" - assumes "\t. infinite t \ countable t \ t \ s \ - (\x\s. \U. x\U \ open U \ infinite (U \ t))" - shows "seq_compact s" -proof - - { - fix f :: "nat \ 'a" - assume f: "\n. f n \ s" - have "\l\s. \r. strict_mono r \ ((f \ r) \ l) sequentially" - proof (cases "finite (range f)") - case True - obtain l where "infinite {n. f n = f l}" - using pigeonhole_infinite[OF _ True] by auto - then obtain r :: "nat \ nat" where "strict_mono r" and fr: "\n. f (r n) = f l" - using infinite_enumerate by blast - then have "strict_mono r \ (f \ r) \ f l" - by (simp add: fr o_def) - with f show "\l\s. \r. strict_mono r \ (f \ r) \ l" - by auto - next - case False - with f assms have "\x\s. \U. x\U \ open U \ infinite (U \ range f)" - by auto - then obtain l where "l \ s" "\U. l\U \ open U \ infinite (U \ range f)" .. - from this(2) have "\r. strict_mono r \ ((f \ r) \ l) sequentially" - using acc_point_range_imp_convergent_subsequence[of l f] by auto - with \l \ s\ show "\l\s. \r. strict_mono r \ ((f \ r) \ l) sequentially" .. - qed - } - then show ?thesis - unfolding seq_compact_def by auto + fixes S :: "'a::first_countable_topology set" + assumes "\T. \infinite T; countable T; T \ S\ \ \x\S. \U. x\U \ open U \ infinite (U \ T)" + shows "seq_compact S" + unfolding seq_compact_def +proof (intro strip) + fix f :: "nat \ 'a" + assume f: "\n. f n \ S" + show "\l\S. \r. strict_mono r \ ((f \ r) \ l) sequentially" + proof (cases "finite (range f)") + case True + obtain l where "infinite {n. f n = f l}" + using pigeonhole_infinite[OF _ True] by auto + then obtain r :: "nat \ nat" where "strict_mono r" and fr: "\n. f (r n) = f l" + using infinite_enumerate by blast + then have "strict_mono r \ (f \ r) \ f l" + by (simp add: fr o_def) + with f show "\l\S. \r. strict_mono r \ (f \ r) \ l" + by auto + next + case False + with f assms obtain l where "l \ S" "\U. l\U \ open U \ infinite (U \ range f)" + by (metis image_subset_iff uncountable_def) + with \l \ S\ show "\l\S. \r. strict_mono r \ ((f \ r) \ l) sequentially" + by (meson acc_point_range_imp_convergent_subsequence) + qed qed lemma seq_compact_eq_countably_compact: fixes U :: "'a :: first_countable_topology set" shows "seq_compact U \ countably_compact U" - using - countable_acc_point_imp_seq_compact - countably_compact_imp_acc_point - seq_compact_imp_countably_compact - by metis + by (metis countable_acc_point_imp_seq_compact countably_compact_imp_acc_point seq_compact_imp_countably_compact) lemma seq_compact_eq_acc_point: - fixes s :: "'a :: first_countable_topology set" - shows "seq_compact s \ - (\t. infinite t \ countable t \ t \ s --> (\x\s. \U. x\U \ open U \ infinite (U \ t)))" - using - countable_acc_point_imp_seq_compact[of s] - countably_compact_imp_acc_point[of s] - seq_compact_imp_countably_compact[of s] - by metis + fixes S :: "'a :: first_countable_topology set" + shows "seq_compact S \ + (\T. infinite T \ countable T \ T \ S --> (\x\S. \U. x\U \ open U \ infinite (U \ T)))" + by (metis countable_acc_point_imp_seq_compact countably_compact_imp_acc_point seq_compact_imp_countably_compact) lemma seq_compact_eq_compact: fixes U :: "'a :: second_countable_topology set" @@ -2133,74 +1927,72 @@ using seq_compact_eq_countably_compact countably_compact_eq_compact by blast proposition Bolzano_Weierstrass_imp_seq_compact: - fixes s :: "'a::{t1_space, first_countable_topology} set" - shows "\t. infinite t \ t \ s \ (\x \ s. x islimpt t) \ seq_compact s" + fixes S :: "'a::{t1_space, first_countable_topology} set" + shows "(\T. \infinite T; T \ S\ \\x \ S. x islimpt T) \ seq_compact S" by (rule countable_acc_point_imp_seq_compact) (metis islimpt_eq_acc_point) subsection\<^marker>\tag unimportant\ \Cartesian products\ -lemma seq_compact_Times: "seq_compact s \ seq_compact t \ seq_compact (s \ t)" +lemma seq_compact_Times: + assumes "seq_compact S" "seq_compact T" + shows "seq_compact (S \ T)" unfolding seq_compact_def - apply clarify - apply (drule_tac x="fst \ f" in spec) - apply (drule mp, simp add: mem_Times_iff) - apply (clarify, rename_tac l1 r1) - apply (drule_tac x="snd \ f \ r1" in spec) - apply (drule mp, simp add: mem_Times_iff) - apply (clarify, rename_tac l2 r2) - apply (rule_tac x="(l1, l2)" in rev_bexI, simp) - apply (rule_tac x="r1 \ r2" in exI) - apply (rule conjI, simp add: strict_mono_def) - apply (drule_tac f=r2 in LIMSEQ_subseq_LIMSEQ, assumption) - apply (drule (1) tendsto_Pair) back - apply (simp add: o_def) - done +proof clarify + fix h :: "nat \ 'a \ 'b" + assume "\n. h n \ S \ T" + then have *: "\n. (fst \ h) n \ S" "\n. (snd \ h) n \ T" + by (simp_all add: mem_Times_iff) + then obtain lS and rS :: "nat\nat" + where "lS\S" "strict_mono rS" and lS: "(fst \ h \ rS) \ lS" + using assms seq_compact_def by metis + then obtain lT and rT :: "nat\nat" + where "lT\T" "strict_mono rT" and lT: "(snd \ h \ rS \ rT) \ lT" + using assms seq_compact_def * + by (metis (mono_tags, lifting) comp_apply) + have "strict_mono (rS \ rT)" + by (simp add: \strict_mono rS\ \strict_mono rT\ strict_mono_o) + moreover have "(h \ (rS \ rT)) \ (lS,lT)" + using tendsto_Pair [OF LIMSEQ_subseq_LIMSEQ [OF lS \strict_mono rT\] lT] + by (simp add: o_def) + ultimately show "\l\S \ T. \r. strict_mono r \ (h \ r) \ l" + using \lS \ S\ \lT \ T\ by blast +qed lemma compact_Times: - assumes "compact s" "compact t" - shows "compact (s \ t)" + assumes "compact S" "compact T" + shows "compact (S \ T)" proof (rule compactI) - fix C - assume C: "\t\C. open t" "s \ t \ \C" - have "\x\s. \a. open a \ x \ a \ (\d\C. finite d \ a \ t \ \d)" + fix \ + assume C: "\T\\. open T" "S \ T \ \\" + have "\x\S. \A. open A \ x \ A \ (\D\\. finite D \ A \ T \ \D)" proof fix x - assume "x \ s" - have "\y\t. \a b c. c \ C \ open a \ open b \ x \ a \ y \ b \ a \ b \ c" (is "\y\t. ?P y") - proof - fix y - assume "y \ t" - with \x \ s\ C obtain c where "c \ C" "(x, y) \ c" "open c" by auto - then show "?P y" by (auto elim!: open_prod_elim) - qed - then obtain a b c where b: "\y. y \ t \ open (b y)" - and c: "\y. y \ t \ c y \ C \ open (a y) \ open (b y) \ x \ a y \ y \ b y \ a y \ b y \ c y" + assume "x \ S" + have "\y\T. \A B C. C \ \ \ open A \ open B \ x \ A \ y \ B \ A \ B \ C" + by (smt (verit, ccfv_threshold) C UnionE \x \ S\ mem_Sigma_iff open_prod_def subsetD) + then obtain a b c where b: "\y. y \ T \ open (b y)" + and c: "\y. y \ T \ c y \ \ \ open (a y) \ open (b y) \ x \ a y \ y \ b y \ a y \ b y \ c y" by metis - then have "\y\t. open (b y)" "t \ (\y\t. b y)" by auto - with compactE_image[OF \compact t\] obtain D where D: "D \ t" "finite D" "t \ (\y\D. b y)" + then have "\y\T. open (b y)" "T \ (\y\T. b y)" by auto + with compactE_image[OF \compact T\] obtain D where D: "D \ T" "finite D" "T \ (\y\D. b y)" by metis - moreover from D c have "(\y\D. a y) \ t \ (\y\D. c y)" + moreover from D c have "(\y\D. a y) \ T \ (\y\D. c y)" by (fastforce simp: subset_eq) - ultimately show "\a. open a \ x \ a \ (\d\C. finite d \ a \ t \ \d)" + ultimately show "\a. open a \ x \ a \ (\d\\. finite d \ a \ T \ \d)" using c by (intro exI[of _ "c`D"] exI[of _ "\(a`D)"] conjI) (auto intro!: open_INT) qed - then obtain a d where a: "\x. x\s \ open (a x)" "s \ (\x\s. a x)" - and d: "\x. x \ s \ d x \ C \ finite (d x) \ a x \ t \ \(d x)" + then obtain a d where a: "\x. x\S \ open (a x)" "S \ (\x\S. a x)" + and d: "\x. x \ S \ d x \ \ \ finite (d x) \ a x \ T \ \(d x)" unfolding subset_eq UN_iff by metis moreover - from compactE_image[OF \compact s\ a] - obtain e where e: "e \ s" "finite e" and s: "s \ (\x\e. a x)" + from compactE_image[OF \compact S\ a] + obtain e where e: "e \ S" "finite e" and S: "S \ (\x\e. a x)" by auto moreover - { - from s have "s \ t \ (\x\e. a x \ t)" - by auto - also have "\ \ (\x\e. \(d x))" - using d \e \ s\ by (intro UN_mono) auto - finally have "s \ t \ (\x\e. \(d x))" . - } - ultimately show "\C'\C. finite C' \ s \ t \ \C'" + have "S \ T \ (\x\e. \(d x))" + by (smt (verit, del_insts) S SigmaE UN_iff d e(1) mem_Sigma_iff subset_eq) + ultimately show "\C'\\. finite C' \ S \ T \ \C'" by (intro exI[of _ "(\x\e. d x)"]) (auto simp: subset_eq) qed @@ -2248,9 +2040,9 @@ have "open {.. C) psi" by (auto intro!: continuous_intros simp: psi_def split_beta') - from this[unfolded continuous_on_open_invariant, rule_format, OF \open {..] - obtain W where W: "open W" "W \ U \ C = W0 \ U \ C" - unfolding W0_eq by blast + then obtain W where W: "open W" "W \ U \ C = W0 \ U \ C" + unfolding W0_eq + by (metis \open {.. continuous_on_open_invariant inf_right_idem) have "{x0} \ C \ W \ U \ C" unfolding W by (auto simp: W0_def psi_def \0 < e\) @@ -2265,20 +2057,9 @@ fix t assume t: "t \ C" have "dist (fx (x, t)) (fx (x0, t)) = psi (x, t)" by (auto simp: psi_def) - also - { - have "(x, t) \ X0 \ C" - using t x - by auto - also note \\ \ W\ - finally have "(x, t) \ W" . - with t x have "(x, t) \ W \ U \ C" - by blast - also note \W \ U \ C = W0 \ U \ C\ - finally have "psi (x, t) < e" - by (auto simp: W0_def) - } - finally show "dist (fx (x, t)) (fx (x0, t)) \ e" by simp + also have "psi (x, t) < e" + using W(2) W0_def X0(3) t x by fastforce + finally show "dist (fx (x,t)) (fx (x0,t)) \ e" by simp qed from X0(1,2) this show ?thesis .. qed @@ -2296,15 +2077,15 @@ lemmas continuous_on = continuous_on_def \ \legacy theorem name\ lemma continuous_within_subset: - "continuous (at x within s) f \ t \ s \ continuous (at x within t) f" + "continuous (at x within S) f \ t \ S \ continuous (at x within t) f" unfolding continuous_within by(metis tendsto_within_subset) lemma continuous_on_interior: - "continuous_on s f \ x \ interior s \ continuous (at x) f" + "continuous_on S f \ x \ interior S \ continuous (at x) f" by (metis continuous_on_eq_continuous_at continuous_on_subset interiorE) lemma continuous_on_eq: - "\continuous_on s f; \x. x \ s \ f x = g x\ \ continuous_on s g" + "\continuous_on S f; \x. x \ S \ f x = g x\ \ continuous_on S g" unfolding continuous_on_def tendsto_def eventually_at_topological by simp @@ -2312,39 +2093,38 @@ lemma continuous_within_sequentiallyI: fixes f :: "'a::{first_countable_topology, t2_space} \ 'b::topological_space" - assumes "\u::nat \ 'a. u \ a \ (\n. u n \ s) \ (\n. f (u n)) \ f a" - shows "continuous (at a within s) f" + assumes "\u::nat \ 'a. u \ a \ (\n. u n \ S) \ (\n. f (u n)) \ f a" + shows "continuous (at a within S) f" using assms unfolding continuous_within tendsto_def[where l = "f a"] by (auto intro!: sequentially_imp_eventually_within) lemma continuous_within_tendsto_compose: fixes f::"'a::t2_space \ 'b::topological_space" - assumes "continuous (at a within s) f" - "eventually (\n. x n \ s) F" - "(x \ a) F " + assumes f: "continuous (at a within S) f" + and "eventually (\n. x n \ S) F" "(x \ a) F " shows "((\n. f (x n)) \ f a) F" proof - - have *: "filterlim x (inf (nhds a) (principal s)) F" - using assms(2) assms(3) unfolding at_within_def filterlim_inf by (auto simp: filterlim_principal eventually_mono) + have *: "filterlim x (inf (nhds a) (principal S)) F" + by (simp add: assms filterlim_inf filterlim_principal) show ?thesis - by (auto simp: assms(1) continuous_within[symmetric] tendsto_at_within_iff_tendsto_nhds[symmetric] intro!: filterlim_compose[OF _ *]) + using "*" f continuous_within filterlim_compose tendsto_at_within_iff_tendsto_nhds by blast qed lemma continuous_within_tendsto_compose': fixes f::"'a::t2_space \ 'b::topological_space" - assumes "continuous (at a within s) f" - "\n. x n \ s" + assumes "continuous (at a within S) f" + "\n. x n \ S" "(x \ a) F " shows "((\n. f (x n)) \ f a) F" - by (auto intro!: continuous_within_tendsto_compose[OF assms(1)] simp add: assms) + using always_eventually assms continuous_within_tendsto_compose by blast lemma continuous_within_sequentially: fixes f :: "'a::{first_countable_topology, t2_space} \ 'b::topological_space" - shows "continuous (at a within s) f \ - (\x. (\n::nat. x n \ s) \ (x \ a) sequentially + shows "continuous (at a within S) f \ + (\x. (\n::nat. x n \ S) \ (x \ a) sequentially \ ((f \ x) \ f a) sequentially)" - using continuous_within_tendsto_compose'[of a s f _ sequentially] - continuous_within_sequentiallyI[of a s f] + using continuous_within_tendsto_compose'[of a S f _ sequentially] + continuous_within_sequentiallyI[of a S f] by (auto simp: o_def) lemma continuous_at_sequentiallyI: @@ -2361,42 +2141,26 @@ lemma continuous_on_sequentiallyI: fixes f :: "'a::{first_countable_topology, t2_space} \ 'b::topological_space" - assumes "\u a. (\n. u n \ s) \ a \ s \ u \ a \ (\n. f (u n)) \ f a" - shows "continuous_on s f" + assumes "\u a. (\n. u n \ S) \ a \ S \ u \ a \ (\n. f (u n)) \ f a" + shows "continuous_on S f" using assms unfolding continuous_on_eq_continuous_within - using continuous_within_sequentiallyI[of _ s f] by auto + using continuous_within_sequentiallyI[of _ S f] by auto lemma continuous_on_sequentially: fixes f :: "'a::{first_countable_topology, t2_space} \ 'b::topological_space" - shows "continuous_on s f \ - (\x. \a \ s. (\n. x(n) \ s) \ (x \ a) sequentially - --> ((f \ x) \ f a) sequentially)" - (is "?lhs = ?rhs") -proof - assume ?rhs - then show ?lhs - using continuous_within_sequentially[of _ s f] - unfolding continuous_on_eq_continuous_within - by auto -next - assume ?lhs - then show ?rhs - unfolding continuous_on_eq_continuous_within - using continuous_within_sequentially[of _ s f] - by auto -qed + shows "continuous_on S f \ + (\x. \a \ S. (\n. x(n) \ S) \ (x \ a) sequentially + \ ((f \ x) \ f a) sequentially)" + by (meson continuous_on_eq_continuous_within continuous_within_sequentially) text \Continuity in terms of open preimages.\ lemma continuous_at_open: - "continuous (at x) f \ (\t. open t \ f x \ t --> (\s. open s \ x \ s \ (\x' \ s. (f x') \ t)))" - unfolding continuous_within_topological [of x UNIV f] - unfolding imp_conjL - by (intro all_cong imp_cong ex_cong conj_cong refl) auto + "continuous (at x) f \ (\t. open t \ f x \ t \ (\S. open S \ x \ S \ (\x' \ S. (f x') \ t)))" + by (metis UNIV_I continuous_within_topological) lemma continuous_imp_tendsto: - assumes "continuous (at x0) f" - and "x \ x0" + assumes "continuous (at x0) f" and "x \ x0" shows "(f \ x) \ (f x0)" proof (rule topological_tendstoI) fix S @@ -2411,20 +2175,20 @@ subsection \Homeomorphisms\ -definition\<^marker>\tag important\ "homeomorphism s t f g \ - (\x\s. (g(f x) = x)) \ (f ` s = t) \ continuous_on s f \ - (\y\t. (f(g y) = y)) \ (g ` t = s) \ continuous_on t g" +definition\<^marker>\tag important\ "homeomorphism S T f g \ + (\x\S. (g(f x) = x)) \ (f ` S = T) \ continuous_on S f \ + (\y\T. (f(g y) = y)) \ (g ` T = S) \ continuous_on T g" lemma homeomorphismI [intro?]: assumes "continuous_on S f" "continuous_on T g" - "f ` S \ T" "g ` T \ S" "\x. x \ S \ g(f x) = x" "\y. y \ T \ f(g y) = y" - shows "homeomorphism S T f g" + "f ` S \ T" "g ` T \ S" "\x. x \ S \ g(f x) = x" "\y. y \ T \ f(g y) = y" + shows "homeomorphism S T f g" using assms by (force simp: homeomorphism_def) lemma homeomorphism_translation: fixes a :: "'a :: real_normed_vector" shows "homeomorphism ((+) a ` S) S ((+) (- a)) ((+) a)" -unfolding homeomorphism_def by (auto simp: algebra_simps continuous_intros) + unfolding homeomorphism_def by (auto simp: algebra_simps continuous_intros) lemma homeomorphism_ident: "homeomorphism T T (\a. a) (\a. a)" by (rule homeomorphismI) auto @@ -2469,43 +2233,25 @@ "S homeomorphic {} \ S = {}" "{} homeomorphic S \ S = {}" by (auto simp: homeomorphic_def homeomorphism_def) -lemma homeomorphic_refl: "s homeomorphic s" - unfolding homeomorphic_def homeomorphism_def - using continuous_on_id - apply (rule_tac x = "(\x. x)" in exI) - apply (rule_tac x = "(\x. x)" in exI) - apply blast - done +lemma homeomorphic_refl: "S homeomorphic S" + using homeomorphic_def homeomorphism_ident by fastforce -lemma homeomorphic_sym: "s homeomorphic t \ t homeomorphic s" +lemma homeomorphic_sym: "S homeomorphic T \ T homeomorphic S" unfolding homeomorphic_def homeomorphism_def by blast lemma homeomorphic_trans [trans]: - assumes "S homeomorphic T" - and "T homeomorphic U" - shows "S homeomorphic U" - using assms - unfolding homeomorphic_def -by (metis homeomorphism_compose) + assumes "S homeomorphic T" and "T homeomorphic U" + shows "S homeomorphic U" + using assms unfolding homeomorphic_def + by (metis homeomorphism_compose) lemma homeomorphic_minimal: - "s homeomorphic t \ - (\f g. (\x\s. f(x) \ t \ (g(f(x)) = x)) \ - (\y\t. g(y) \ s \ (f(g(y)) = y)) \ - continuous_on s f \ continuous_on t g)" - (is "?lhs = ?rhs") -proof - assume ?lhs - then show ?rhs - by (fastforce simp: homeomorphic_def homeomorphism_def) -next - assume ?rhs - then show ?lhs - apply clarify - unfolding homeomorphic_def homeomorphism_def - by (metis equalityI image_subset_iff subsetI) - qed + "S homeomorphic T \ + (\f g. (\x\S. f(x) \ T \ (g(f(x)) = x)) \ + (\y\T. g(y) \ S \ (f(g(y)) = y)) \ + continuous_on S f \ continuous_on T g)" + by (smt (verit, ccfv_threshold) homeomorphic_def homeomorphismI homeomorphism_def image_eqI image_subset_iff) lemma homeomorphicI [intro?]: "\f ` S = T; g ` T = S; @@ -2517,8 +2263,7 @@ lemma homeomorphism_of_subsets: "\homeomorphism S T f g; S' \ S; T'' \ T; f ` S' = T'\ \ homeomorphism S' T' f g" -apply (auto simp: homeomorphism_def elim!: continuous_on_subset) -by (metis subsetD imageI) + by (smt (verit, del_insts) continuous_on_subset homeomorphismI homeomorphism_def imageE subset_eq) lemma homeomorphism_apply1: "\homeomorphism S T f g; x \ S\ \ g(f x) = x" by (simp add: homeomorphism_def) @@ -2555,77 +2300,37 @@ proof assume "S homeomorphic T" with assms show ?rhs - apply (auto simp: homeomorphic_def homeomorphism_def) - apply (metis finite_imageI) - by (metis card_image_le finite_imageI le_antisym) + by (metis (full_types) card_image_le finite_imageI homeomorphic_def homeomorphism_def le_antisym) next assume R: ?rhs with finite_same_card_bij obtain h where "bij_betw h S T" by auto with R show ?lhs - apply (auto simp: homeomorphic_def homeomorphism_def continuous_on_finite) - apply (rule_tac x=h in exI) - apply (rule_tac x="inv_into S h" in exI) - apply (auto simp: bij_betw_inv_into_left bij_betw_inv_into_right bij_betw_imp_surj_on inv_into_into bij_betwE) - apply (metis bij_betw_def bij_betw_inv_into) - done + apply (simp only: homeomorphic_def homeomorphism_def continuous_on_finite) + by (smt (verit, ccfv_SIG) bij_betw_imp_surj_on bij_betw_inv_into bij_betw_inv_into_left bij_betw_inv_into_right) qed text \Relatively weak hypotheses if a set is compact.\ lemma homeomorphism_compact: fixes f :: "'a::topological_space \ 'b::t2_space" - assumes "compact s" "continuous_on s f" "f ` s = t" "inj_on f s" - shows "\g. homeomorphism s t f g" + assumes "compact S" "continuous_on S f" "f ` S = T" "inj_on f S" + shows "\g. homeomorphism S T f g" proof - - define g where "g x = (SOME y. y\s \ f y = x)" for x - have g: "\x\s. g (f x) = x" - using assms(3) assms(4)[unfolded inj_on_def] unfolding g_def by auto - { - fix y - assume "y \ t" - then obtain x where x:"f x = y" "x\s" - using assms(3) by auto - then have "g (f x) = x" using g by auto - then have "f (g y) = y" unfolding x(1)[symmetric] by auto - } - then have g':"\x\t. f (g x) = x" by auto - moreover - { - fix x - have "x\s \ x \ g ` t" - using g[THEN bspec[where x=x]] - unfolding image_iff - using assms(3) - by (auto intro!: bexI[where x="f x"]) - moreover - { - assume "x\g ` t" - then obtain y where y:"y\t" "g y = x" by auto - then obtain x' where x':"x'\s" "f x' = y" - using assms(3) by auto - then have "x \ s" - unfolding g_def - using someI2[of "\b. b\s \ f b = y" x' "\x. x\s"] - unfolding y(2)[symmetric] and g_def - by auto - } - ultimately have "x\s \ x \ g ` t" .. - } - then have "g ` t = s" by auto - ultimately show ?thesis - unfolding homeomorphism_def homeomorphic_def - using assms continuous_on_inv by fastforce + obtain g where g: "\x\S. g (f x) = x" "\x\T. f (g x) = x" "g ` T = S" + using assms the_inv_into_f_f by fastforce + with assms show ?thesis + unfolding homeomorphism_def homeomorphic_def by (metis continuous_on_inv) qed lemma homeomorphic_compact: fixes f :: "'a::topological_space \ 'b::t2_space" - shows "compact s \ continuous_on s f \ (f ` s = t) \ inj_on f s \ s homeomorphic t" + shows "compact S \ continuous_on S f \ (f ` S = T) \ inj_on f S \ S homeomorphic T" unfolding homeomorphic_def by (metis homeomorphism_compact) text\Preservation of topological properties.\ -lemma homeomorphic_compactness: "s homeomorphic t \ (compact s \ compact t)" +lemma homeomorphic_compactness: "S homeomorphic T \ (compact S \ compact T)" unfolding homeomorphic_def homeomorphism_def by (metis compact_continuous_image) @@ -2639,8 +2344,8 @@ proof (rule islimptI) fix T assume "open T" "a \ T" - from open_right[OF this \a < b\] - obtain c where c: "a < c" "{a.. T" by auto + then obtain c where c: "a < c" "{a.. T" + by (meson assms open_right) with assms dense[of a "min c b"] show "\y\{a<.. T \ y \ a" by (metis atLeastLessThan_iff greaterThanLessThan_iff min_less_iff_conj diff -r ba2afdd29e1d -r 032a4344903e src/HOL/Analysis/Henstock_Kurzweil_Integration.thy --- a/src/HOL/Analysis/Henstock_Kurzweil_Integration.thy Tue Aug 01 11:27:55 2023 +0200 +++ b/src/HOL/Analysis/Henstock_Kurzweil_Integration.thy Thu Aug 03 19:10:43 2023 +0200 @@ -133,16 +133,13 @@ proof - from rational_boxes[OF assms, of c] obtain a b where ab: "c \ box a b" "box a b \ ball c r" by auto - from ab have "0 < content (box a b)" + then have "0 < content (box a b)" by (subst measure_lborel_box_eq) (auto intro!: prod_pos simp: algebra_simps box_def) have "emeasure lborel (box a b) \ emeasure lborel (ball c r)" using ab by (intro emeasure_mono) auto - also have "emeasure lborel (box a b) = ennreal (content (box a b))" - using emeasure_lborel_box_finite[of a b] by (intro emeasure_eq_ennreal_measure) auto - also have "emeasure lborel (ball c r) = ennreal (content (ball c r))" - using emeasure_lborel_ball_finite[of c r] by (intro emeasure_eq_ennreal_measure) auto - finally show ?thesis - using \content (box a b) > 0\ by simp + then show ?thesis + using \content (box a b) > 0\ + by (smt (verit, best) Sigma_Algebra.measure_def emeasure_lborel_ball_finite enn2real_mono infinity_ennreal_def) qed lemma content_cball_pos: @@ -151,12 +148,10 @@ proof - from rational_boxes[OF assms, of c] obtain a b where ab: "c \ box a b" "box a b \ ball c r" by auto - from ab have "0 < content (box a b)" + then have "0 < content (box a b)" by (subst measure_lborel_box_eq) (auto intro!: prod_pos simp: algebra_simps box_def) - have "emeasure lborel (box a b) \ emeasure lborel (ball c r)" + have "emeasure lborel (box a b) \ emeasure lborel (cball c r)" using ab by (intro emeasure_mono) auto - also have "\ \ emeasure lborel (cball c r)" - by (intro emeasure_mono) auto also have "emeasure lborel (box a b) = ennreal (content (box a b))" using emeasure_lborel_box_finite[of a b] by (intro emeasure_eq_ennreal_measure) auto also have "emeasure lborel (cball c r) = ennreal (content (cball c r))" @@ -209,7 +204,7 @@ assumes "content (cbox a b) = 0" and "p tagged_division_of (cbox a b)" shows "(\(x,K)\p. content K *\<^sub>R f x) = (0::'a::real_normed_vector)" -proof (rule sum.neutral, rule) +proof (intro sum.neutral strip) fix y assume y: "y \ p" obtain x K where xk: "y = (x, K)" @@ -257,7 +252,7 @@ qed lemma content_real_eq_0: "content {a..b::real} = 0 \ a \ b" - by (metis atLeastatMost_empty_iff2 content_empty content_real diff_self eq_iff le_cases le_iff_diff_le_0) + by simp lemma property_empty_interval: "\a b. content (cbox a b) = 0 \ P (cbox a b) \ P {}" using content_empty unfolding empty_as_interval by auto @@ -332,7 +327,7 @@ definition "integral i f = (SOME y. (f has_integral y) i \ \ f integrable_on i \ y=0)" lemma integrable_integral[intro]: "f integrable_on i \ (f has_integral (integral i f)) i" - unfolding integrable_on_def integral_def by (metis (mono_tags, lifting) someI_ex) + unfolding integrable_on_def integral_def by (metis (mono_tags, lifting)) lemma not_integrable_integral: "\ f integrable_on i \ integral i f = 0" unfolding integrable_on_def integral_def by blast @@ -351,7 +346,7 @@ lemma has_integral_unique_cbox: fixes f :: "'n::euclidean_space \ 'a::real_normed_vector" shows "(f has_integral k1) (cbox a b) \ (f has_integral k2) (cbox a b) \ k1 = k2" - by (auto simp: has_integral_cbox intro: tendsto_unique[OF division_filter_not_empty]) + by (meson division_filter_not_empty has_integral_cbox tendsto_unique) lemma has_integral_unique: fixes f :: "'n::euclidean_space \ 'a::real_normed_vector" @@ -400,9 +395,7 @@ lemma eq_integralD: "integral k f = y \ (f has_integral y) k \ \ f integrable_on k \ y=0" unfolding integral_def integrable_on_def - apply (erule subst) - apply (rule someI_ex) - by blast + by (metis (mono_tags, lifting)) lemma has_integral_const [intro]: fixes a b :: "'a::euclidean_space" @@ -423,12 +416,12 @@ lemma integral_const [simp]: fixes a b :: "'a::euclidean_space" shows "integral (cbox a b) (\x. c) = content (cbox a b) *\<^sub>R c" - by (rule integral_unique) (rule has_integral_const) + by blast lemma integral_const_real [simp]: fixes a b :: real shows "integral {a..b} (\x. c) = content {a..b} *\<^sub>R c" - by (metis box_real(2) integral_const) + by blast lemma has_integral_is_0_cbox: fixes f :: "'n::euclidean_space \ 'a::real_normed_vector" @@ -685,7 +678,7 @@ lemma integral_0 [simp]: "integral S (\x::'n::euclidean_space. 0::'m::real_normed_vector) = 0" - by (rule integral_unique has_integral_0)+ + by auto lemma integral_add: "f integrable_on S \ g integrable_on S \ integral S (\x. f x + g x) = integral S f + integral S g" @@ -704,17 +697,10 @@ lemma integral_mult: fixes K::real shows "f integrable_on X \ K * integral X f = integral X (\x. K * f x)" - unfolding real_scaleR_def[symmetric] integral_cmul .. + by simp lemma integral_neg [simp]: "integral S (\x. - f x) = - integral S f" -proof (cases "f integrable_on S") - case True then show ?thesis - by (simp add: has_integral_neg integrable_integral integral_unique) -next - case False then have "\ (\x. - f x) integrable_on S" - using has_integral_neg [of "(\x. - f x)" _ S ] by auto - with False show ?thesis by (simp add: not_integrable_integral) -qed + by (metis eq_integralD equation_minus_iff has_integral_iff has_integral_neg_iff neg_equal_0_iff_equal) lemma integral_diff: "f integrable_on S \ g integrable_on S \ integral S (\x. f x - g x) = integral S f - integral S g" @@ -787,15 +773,8 @@ assumes "finite T" and "\a. a \ T \ ((f a) has_integral (i a)) S" shows "((\x. sum (\a. f a x) T) has_integral (sum i T)) S" - using assms(1) subset_refl[of T] -proof (induct rule: finite_subset_induct) - case empty - then show ?case by auto -next - case (insert x F) - with assms show ?case - by (simp add: has_integral_add) -qed + using \finite T\ subset_refl[of T] + by (induct rule: finite_subset_induct) (use assms in \auto simp: has_integral_add\) lemma integral_sum: "\finite I; \a. a \ I \ f a integrable_on S\ \ @@ -808,11 +787,11 @@ lemma has_integral_eq: assumes "\x. x \ s \ f x = g x" - and "(f has_integral k) s" + and f: "(f has_integral k) s" shows "(g has_integral k) s" - using has_integral_diff[OF assms(2), of "\x. f x - g x" 0] + using has_integral_diff[OF f, of "\x. f x - g x" 0] using has_integral_is_0[of s "\x. f x - g x"] - using assms(1) + using assms by auto lemma integrable_eq: "\f integrable_on s; \x. x \ s \ f x = g x\ \ g integrable_on s" @@ -822,8 +801,7 @@ lemma has_integral_cong: assumes "\x. x \ s \ f x = g x" shows "(f has_integral i) s = (g has_integral i) s" - using has_integral_eq[of s f g] has_integral_eq[of s g f] assms - by auto + by (metis assms has_integral_eq) lemma integrable_cong: assumes "\x. x \ A \ f x = g x" @@ -855,25 +833,19 @@ fixes f :: "_ \ 'b :: {comm_ring,real_algebra_1,real_normed_vector}" assumes "f integrable_on s" shows "(\x. f x * of_real c) integrable_on s" -using integrable_on_cmult_left [OF assms] by (simp add: mult.commute) + using integrable_on_cmult_left [OF assms] by (simp add: mult.commute) lemma integrable_on_cmult_right_iff [simp]: fixes f :: "_ \ 'b :: {comm_ring,real_algebra_1,real_normed_vector}" assumes "c \ 0" shows "(\x. f x * of_real c) integrable_on s \ f integrable_on s" -using integrable_on_cmult_left_iff [OF assms] by (simp add: mult.commute) - -lemma integrable_on_cdivide: - fixes f :: "_ \ 'b :: real_normed_field" - assumes "f integrable_on s" - shows "(\x. f x / of_real c) integrable_on s" -by (simp add: integrable_on_cmult_right divide_inverse assms flip: of_real_inverse) + using integrable_on_cmult_left_iff [OF assms] by (simp add: mult.commute) lemma integrable_on_cdivide_iff [simp]: fixes f :: "_ \ 'b :: real_normed_field" assumes "c \ 0" shows "(\x. f x / of_real c) integrable_on s \ f integrable_on s" -by (simp add: divide_inverse assms flip: of_real_inverse) + by (simp add: divide_inverse assms flip: of_real_inverse) lemma has_integral_null [intro]: "content(cbox a b) = 0 \ (f has_integral 0) (cbox a b)" unfolding has_integral_cbox @@ -902,7 +874,7 @@ unfolding integrable_on_def by auto lemma integral_empty[simp]: "integral {} f = 0" - by (rule integral_unique) (rule has_integral_empty) + by blast lemma has_integral_refl[intro]: fixes a :: "'a::euclidean_space" @@ -919,7 +891,7 @@ unfolding integrable_on_def by auto lemma integral_refl [simp]: "integral (cbox a a) f = 0" - by (rule integral_unique) auto + by auto lemma integral_singleton [simp]: "integral {a} f = 0" by auto @@ -1000,7 +972,6 @@ by (drule integrable_linear[OF _ bounded_linear_inner_left[of y]]) (simp add: o_def) - subsection \Cauchy-type criterion for integrability\ proposition integrable_Cauchy: @@ -1634,10 +1605,10 @@ corollary integrable_bound: fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" assumes "0 \ B" - and "f integrable_on (cbox a b)" - and "\x. x\cbox a b \ norm (f x) \ B" - shows "norm (integral (cbox a b) f) \ B * content (cbox a b)" -by (metis integrable_integral has_integral_bound assms) + and "f integrable_on (cbox a b)" + and "\x. x\cbox a b \ norm (f x) \ B" + shows "norm (integral (cbox a b) f) \ B * content (cbox a b)" + by (metis integrable_integral has_integral_bound assms) subsection \Similar theorems about relationship among components\ @@ -1752,21 +1723,14 @@ and "(f has_integral i) S" and "\x. x \ S \ 0 \ (f x)\k" shows "0 \ i\k" - using has_integral_component_le[OF assms(1) has_integral_0 assms(2)] - using assms(3-) - by auto + by (metis (no_types, lifting) assms euclidean_all_zero_iff has_integral_0 has_integral_component_le) lemma integral_component_nonneg: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "k \ Basis" and "\x. x \ S \ 0 \ (f x)\k" shows "0 \ (integral S f)\k" -proof (cases "f integrable_on S") - case True show ?thesis - using True assms has_integral_component_nonneg by blast -next - case False then show ?thesis by (simp add: not_integrable_integral) -qed + by (smt (verit, ccfv_threshold) assms eq_integralD euclidean_all_zero_iff has_integral_component_nonneg) lemma has_integral_component_neg: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" @@ -1774,8 +1738,7 @@ and "(f has_integral i) S" and "\x. x \ S \ (f x)\k \ 0" shows "i\k \ 0" - using has_integral_component_le[OF assms(1,2) has_integral_0] assms(2-) - by auto + by (metis (no_types, lifting) assms has_integral_0 has_integral_component_le inner_zero_left) lemma has_integral_component_lbound: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" @@ -1808,8 +1771,7 @@ and "\x\{a..b}. B \ f(x)\k" and "k \ Basis" shows "B * content {a..b} \ (integral {a..b} f)\k" - using assms - by (metis box_real(2) integral_component_lbound) + using assms by (metis box_real(2) integral_component_lbound) lemma integral_component_ubound: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" @@ -1825,8 +1787,7 @@ and "\x\{a..b}. f x\k \ B" and "k \ Basis" shows "(integral {a..b} f)\k \ B * content {a..b}" - using assms - by (metis box_real(2) integral_component_ubound) + using assms by (metis box_real(2) integral_component_ubound) subsection \Uniform limit of integrable functions is integrable\ @@ -1863,7 +1824,7 @@ then obtain M where "M \ 0" and M: "1 / (real M) < e/4 / content (cbox a b)" by (metis inverse_eq_divide real_arch_inverse) show "\M. \m\M. \n\M. dist (h m) (h n) < e" - proof (rule exI [where x=M], clarify) + proof (intro exI strip) fix m n assume m: "M \ m" and n: "M \ n" have "e/4>0" using \e>0\ by auto @@ -2611,8 +2572,7 @@ fixes f :: "real \ 'a::banach" assumes "continuous_on (closed_segment a b) f" shows "f integrable_on (closed_segment a b)" - using assms - by (auto intro!: integrable_continuous_interval simp: closed_segment_eq_real_ivl) + by (metis assms closed_segment_eq_real_ivl integrable_continuous_interval) subsection \Specialization of additivity to one dimension\ @@ -2773,7 +2733,7 @@ lemma ident_integrable_on: fixes a::real shows "(\x. x) integrable_on {a..b}" -by (metis atLeastatMost_empty_iff integrable_on_def has_integral_empty ident_has_integral) + using continuous_on_id integrable_continuous_real by blast lemma integral_sin [simp]: fixes a::real @@ -2862,7 +2822,7 @@ assumes f'[derivative_intros]: "\x. x \ S \ (f has_derivative f' x) (at x within S)" assumes line_in: "\t. t \ {0..1} \ x + t *\<^sub>R y \ S" - shows "f (x + y) - f x = integral {0..1} (\t. f' (x + t *\<^sub>R y) y)" (is ?th1) + shows "f (x + y) - f x = integral {0..1} (\t. f' (x + t *\<^sub>R y) y)" proof - from assms have subset: "(\xa. x + xa *\<^sub>R y) ` {0..1} \ S" by auto note [derivative_intros] = @@ -2879,7 +2839,7 @@ linear_cmul[OF has_derivative_linear[OF f'], symmetric] intro!: derivative_eq_intros) from fundamental_theorem_of_calculus[rule_format, OF _ this] - show ?th1 + show ?thesis by (auto intro!: integral_unique[symmetric]) qed @@ -2969,26 +2929,20 @@ using atLeastatMost_empty'[simp del] by (simp add: i_def g_def Dg_def) also - have one: "(- 1) ^ p' * (- 1) ^ p' = (1::real)" - and "{.. {i. p = Suc i} = {p - 1}" - for p' - using \p > 0\ - by (auto simp: power_mult_distrib[symmetric]) + have one: "(- 1) ^ p' * (- 1) ^ p' = (1::real)" "{.. {i. p = Suc i} = {p - 1}" for p' + using \p > 0\ by (auto simp: power_mult_distrib) then have "?sum b = f b" using Suc_pred'[OF \p > 0\] by (simp add: diff_eq_eq Dg_def power_0_left le_Suc_eq if_distrib if_distribR sum.If_cases f0) also - have "{..x. p - x - 1) ` {.. (\x. p - x - 1) ` {..iR Df i a)" - by (rule sum.reindex_cong) (auto simp add: inj_on_def Dg_def one) + proof (rule sum.reindex_cong) + have "\i. i < p \ \jp>0\ diff_Suc_less diff_diff_cancel less_or_eq_imp_le) + then show "{..x. p - x - 1) ` {..i. f integrable_on i" using order_refl by (rule operative_integrableI) show ?thesis - proof (cases "cbox c d = {}") - case True - then show ?thesis - using division [symmetric] f by (auto simp: comm_monoid_set_F_and) - next - case False - then show ?thesis - by (metis cd comm_monoid_set_F_and division division_of_finite f partial_division_extend_1) - qed + by (metis cd division division_of_finite empty f partial_division_extend_1 remove) qed lemma integrable_subinterval_real: @@ -3151,8 +3097,7 @@ by (auto simp: split: if_split_asm) then have "f integrable_on cbox a b" using ac cb by (auto split: if_split_asm) - with * - show ?thesis + with * show ?thesis using ac cb by (auto simp add: integrable_on_def integral_unique split: if_split_asm) qed @@ -3168,10 +3113,8 @@ moreover have "(f has_integral integral {c..b} f) {c..b}" using ab \a \ c\ integrable_subinterval_real by fastforce - ultimately have "(f has_integral integral {a..c} f + integral {c..b} f) {a..b}" - using \a \ c\ \c \ b\ has_integral_combine by blast - then show ?thesis - by (simp add: has_integral_integrable_integral) + ultimately show ?thesis + by (smt (verit, best) assms has_integral_combine integral_unique) qed lemma integrable_combine: @@ -3181,9 +3124,7 @@ and "f integrable_on {a..c}" and "f integrable_on {c..b}" shows "f integrable_on {a..b}" - using assms - unfolding integrable_on_def - by (auto intro!: has_integral_combine) + using assms has_integral_combine by blast lemma integral_minus_sets: fixes f::"real \ 'a::banach" @@ -3304,7 +3245,7 @@ assumes "continuous_on {a..b} f" and "x \ {a..b}" shows "((\u. integral {a..u} f) has_vector_derivative f(x)) (at x within {a..b})" -using assms integral_has_vector_derivative_continuous_at [OF integrable_continuous_real] + using assms integral_has_vector_derivative_continuous_at [OF integrable_continuous_real] by (fastforce simp: continuous_on_eq_continuous_within) lemma integral_has_real_derivative: @@ -3396,29 +3337,29 @@ show "d fine (\(x, k). (g x, g ` k)) ` p" using finep unfolding fine_def d'_def by auto next - fix x k - assume xk: "(x, k) \ p" - show "g x \ g ` k" + fix x K + assume xk: "(x, K) \ p" + show "g x \ g ` K" using p(2)[OF xk] by auto - show "\u v. g ` k = cbox u v" + show "\u v. g ` K = cbox u v" using p(4)[OF xk] using assms(5-6) by auto fix x' K' u - assume xk': "(x', K') \ p" and u: "u \ interior (g ` k)" "u \ interior (g ` K')" - have "interior k \ interior K' \ {}" + assume xk': "(x', K') \ p" and u: "u \ interior (g ` K)" "u \ interior (g ` K')" + have "interior K \ interior K' \ {}" proof - assume "interior k \ interior K' = {}" - moreover have "u \ g ` (interior k \ interior K')" + assume "interior K \ interior K' = {}" + moreover have "u \ g ` (interior K \ interior K')" using interior_image_subset[OF \inj g\ contg] u unfolding image_Int[OF inj(1)] by blast ultimately show False by blast qed - then have same: "(x, k) = (x', K')" + then have same: "(x, K) = (x', K')" using ptag xk' xk by blast then show "g x = g x'" by auto - show "g u \ g ` K'"if "u \ k" for u + show "g u \ g ` K'"if "u \ K" for u using that same by auto - show "g u \ g ` k"if "u \ K'" for u + show "g u \ g ` K"if "u \ K'" for u using that same by auto next fix x @@ -3528,8 +3469,7 @@ moreover from True have *: "\i. (m *\<^sub>R b + c) \ i - (m *\<^sub>R a + c) \ i = m *\<^sub>R (b-a) \ i" by (simp add: inner_simps field_simps) ultimately show ?thesis - by (simp add: image_affinity_cbox True content_cbox' - prod.distrib inner_diff_left) + by (simp add: image_affinity_cbox True content_cbox' prod.distrib inner_diff_left) next case False with \cbox a b \ {}\ have "cbox (m *\<^sub>R b + c) (m *\<^sub>R a + c) \ {}" @@ -3578,7 +3518,8 @@ assumes "c \ 0" shows "((\x. c *\<^sub>R f x) has_integral (c *\<^sub>R I)) A \ (f has_integral I) A" using assms has_integral_cmul[of f I A c] - has_integral_cmul[of "\x. c *\<^sub>R f x" "c *\<^sub>R I" A "inverse c"] by (auto simp: field_simps) + has_integral_cmul[of "\x. c *\<^sub>R f x" "c *\<^sub>R I" A "inverse c"] + by (auto simp: field_simps) lemma has_integral_cmul_iff': assumes "c \ 0" @@ -3607,7 +3548,8 @@ by (subst image_smult_cbox) simp_all also have "(\x. - ((1 / m) *\<^sub>R c) + x) ` \ = cbox ((a - c) /\<^sub>R m) ((b - c) /\<^sub>R m)" by (subst cbox_translation [symmetric]) (simp add: field_simps vector_add_divide_simps) - finally show ?thesis using \m > 0\ by (simp add: field_simps) + finally show ?thesis using \m > 0\ + by (simp add: field_simps) qed lemma has_integral_affinity_iff: @@ -3619,7 +3561,7 @@ proof assume ?lhs from has_integral_affinity'[OF this, of "1 / m" "-c /\<^sub>R m"] and \m > 0\ - show ?rhs by (simp add: vector_add_divide_simps) (simp add: field_simps) + show ?rhs by (simp add: vector_add_divide_simps) (simp add: field_simps) next assume ?rhs from has_integral_affinity'[OF this, of m c] and \m > 0\ @@ -3708,12 +3650,7 @@ shows "uminus ` cbox a b = cbox (-b) (-a)" proof - have "x \ uminus ` cbox a b" if "x \ cbox (- b) (- a)" for x - proof - - have "-x \ cbox a b" - using that by (auto simp: mem_box) - then show ?thesis - by force - qed + by (smt (verit) add.inverse_inverse image_iff inner_minus_left mem_box(2) that) then show ?thesis by (auto simp: mem_box) qed @@ -3727,9 +3664,7 @@ lemma has_integral_reflect_lemma_real[intro]: assumes "(f has_integral i) {a..b::real}" shows "((\x. f(-x)) has_integral i) {-b .. -a}" - using assms - unfolding box_real[symmetric] - by (rule has_integral_reflect_lemma) + by (metis has_integral_reflect_lemma interval_cbox assms) lemma has_integral_reflect[simp]: "((\x. f (-x)) has_integral i) (cbox (-b) (-a)) \ (f has_integral i) (cbox a b)" @@ -3832,9 +3767,8 @@ proof (rule add_mono) have "norm ((c - a) *\<^sub>R f' a) \ norm (l *\<^sub>R f' a)" by (auto intro: mult_right_mono [OF lel]) - also have "... \ e * (b-a) / 8" - by (rule l) - finally show "norm ((c - a) *\<^sub>R f' a) \ e * (b-a) / 8" . + with l show "norm ((c - a) *\<^sub>R f' a) \ e * (b-a) / 8" + by linarith next have "norm (f c - f a) < e * (b-a) / 8" proof (cases "a = c") @@ -3959,9 +3893,8 @@ by (auto intro: sum_norm_le) also have "... \ (\n\p - ?A. e * (case n of (x, k) \ Sup k - Inf k)/2)" using non by (fastforce intro: sum_mono) - finally have I: "norm (\(x, k)\p - ?A. - content k *\<^sub>R f' x - (f (Sup k) - f (Inf k))) - \ (\n\p - ?A. e * (case n of (x, k) \ Sup k - Inf k))/2" + finally have I: "norm (\(x, k)\p - ?A. content k *\<^sub>R f' x - (f (Sup k) - f (Inf k))) + \ (\n\p - ?A. e * (case n of (x, k) \ Sup k - Inf k))/2" by (simp add: sum_divide_distrib) have II: "norm (\(x, k)\p \ ?A. content k *\<^sub>R f' x - (f (Sup k) - f (Inf k))) - (\n\p \ ?A. e * (case n of (x, k) \ Sup k - Inf k)) @@ -3971,7 +3904,7 @@ proof - obtain u v where uv: "k = cbox u v" by (meson Int_iff xkp p(4)) - with p(2) that uv have "cbox u v \ {}" + with p that have "cbox u v \ {}" by blast then show "0 \ e * ((Sup k) - (Inf k))" unfolding uv using e by (auto simp add: field_simps) @@ -3987,10 +3920,8 @@ proof - have xk: "(x,K) \ p" and k0: "content K = 0" using that by auto - then obtain u v where uv: "K = cbox u v" - using p(4) by blast - then have "u = v" - using xk k0 p(2) by force + then obtain u v where uv: "K = cbox u v" "u = v" + using xk k0 p by fastforce then show "content K *\<^sub>R (f' x) - (f ((Sup K)) - f ((Inf K))) = 0" using xk unfolding uv by auto qed @@ -4021,7 +3952,7 @@ also have "... \ e * (b - a) / 4 + e * (b - a) / 4" proof (rule norm_triangle_le [OF add_mono]) have pa: "\v. k = cbox a v \ a \ v" if "(a, k) \ p" for k - using p(2) p(3) p(4) that by fastforce + using p that by fastforce show "norm (\(x,K) \ ?B a. content K *\<^sub>R f' x - (f (Sup K) - f (Inf K))) \ e * (b - a) / 4" proof (intro norm_le; clarsimp) fix K K' @@ -4059,7 +3990,7 @@ qed (use ab e in auto) next have pb: "\v. k = cbox v b \ b \ v" if "(b, k) \ p" for k - using p(2) p(3) p(4) that by fastforce + using p that by fastforce show "norm (\(x,K) \ ?B b. content K *\<^sub>R f' x - (f (Sup K) - f (Inf K))) \ e * (b - a) / 4" proof (intro norm_le; clarsimp) fix K K' @@ -4071,7 +4002,7 @@ unfolding v v' by (auto simp: mem_box) then have "interior (box (max v v') b) \ interior K \ interior K'" using interior_Int interior_mono by blast - moreover have " ((b + ?v)/2) \ box ?v b" + moreover have "((b + ?v)/2) \ box ?v b" using ne0 unfolding v v' content_eq_0 not_le by (auto simp: mem_box) ultimately have "((b + ?v)/2) \ interior K \ interior K'" unfolding interior_open[OF open_box] by auto @@ -4181,18 +4112,11 @@ hence e3: "0 < e/3 / norm (f c)" using \e>0\ by simp moreover have "norm (f c) * norm (c - t) < e/3" if "t < c" and "c - e/3 / norm (f c) < t" for t - proof - - have "norm (c - t) < e/3 / norm (f c)" - using that by auto - then show "norm (f c) * norm (c - t) < e/3" - by (metis e3 mult.commute norm_not_less_zero pos_less_divide_eq zero_less_divide_iff) - qed + unfolding real_norm_def + by (smt (verit) False divide_right_mono nonzero_mult_div_cancel_left norm_eq_zero norm_ge_zero that) ultimately show ?thesis using that by auto - next - case True then show ?thesis - using \e > 0\ that by auto - qed + qed (use \e > 0\ in auto) let ?SUM = "\p. (\(x,K) \ p. content K *\<^sub>R f x)" have e3: "e/3 > 0" @@ -4288,7 +4212,7 @@ ultimately have cwt: "c - w < t" by (auto simp add: field_simps) have eq: "integral {a..c} f - integral {a..t} f = -(((c - t) *\<^sub>R f c + ?SUM p) - - integral {a..c} f) + (?SUM p - integral {a..t} f) + (c - t) *\<^sub>R f c" + integral {a..c} f) + (?SUM p - integral {a..t} f) + (c - t) *\<^sub>R f c" by auto have "norm (integral {a..c} f - integral {a..t} f) < e/3 + e/3 + e/3" unfolding eq @@ -4373,15 +4297,8 @@ show "0 < min d1 d2" using \0 < d1\ \0 < d2\ by simp show "dist (integral {a..y} f) (integral {a..x} f) < e" - if "y \ {a..b}" "dist y x < min d1 d2" for y - proof (cases "y < x") - case True - with that d1 show ?thesis by (auto simp: dist_commute dist_norm) - next - case False - with that d2 show ?thesis - by (auto simp: dist_commute dist_norm) - qed + if "dist y x < min d1 d2" for y + by (smt (verit) d1 d2 dist_norm dist_real_def norm_minus_commute that) qed qed qed @@ -4562,18 +4479,7 @@ assumes "connected S" "open S" "finite K" "continuous_on S f" and "\x\S-K. (f has_derivative (\h. 0)) (at x within S)" shows "f constant_on S" -proof (cases "S = {}") - case True - then show ?thesis - by (simp add: constant_on_def) -next - case False - then obtain c where "c \ S" - by (metis equals0I) - then show ?thesis - unfolding constant_on_def - by (metis has_derivative_zero_unique_strong_connected assms ) -qed + by (smt (verit, best) assms constant_on_def has_derivative_zero_unique_strong_connected) lemma DERIV_zero_connected_constant_on: fixes f :: "'a::{real_normed_field,euclidean_space} \ 'a" @@ -4594,13 +4500,8 @@ fixes f :: "'a::{real_normed_field,euclidean_space} \ 'a" assumes "\z. z \ S \ (f has_field_derivative 0) (at z)" and S: "connected S" "open S" shows "f constant_on S" -proof - - have *: "continuous_on S f" - using assms(1) by (intro DERIV_continuous_on[of _ _ "\_. 0"]) - (use assms in \auto simp: at_within_open[of _ S]\) - show ?thesis - using DERIV_zero_connected_constant_on[OF S finite.emptyI *] assms(1) by blast -qed + using DERIV_zero_connected_constant_on [where K="Basis"] + by (metis DERIV_isCont Diff_iff assms continuous_at_imp_continuous_on eucl.finite_Basis) subsection \Integrating characteristic function of an interval\ @@ -4653,19 +4554,11 @@ using integrable_spike_interior[where f=f] by (meson g_def has_integral_integrable intf) moreover have "integral (cbox c d) g = i" - proof (rule has_integral_unique[OF has_integral_spike_interior intf]) - show "\x. x \ box c d \ f x = g x" - by (auto simp: g_def) - show "(g has_integral integral (cbox c d) g) (cbox c d)" - by (rule integrable_integral[OF intg]) - qed + by (meson g_def has_integral_iff has_integral_spike_interior intf) ultimately have "F (\A. if g integrable_on A then Some (integral A g) else None) p = Some i" by (metis (full_types, lifting) division_of_finite inp iterate pdiv remove right_neutral) - then - have "(g has_integral i) (cbox a b)" - by (metis integrable_on_def integral_unique operat option.inject option.simps(3)) with False show ?thesis - by blast + by (metis integrable_integral not_None_eq operat option.inject) qed qed @@ -4799,17 +4692,13 @@ and "g integrable_on S" and "\x. x \ S \ f x \ g x" shows "integral S f \ integral S g" - by (rule has_integral_le[OF assms(1,2)[unfolded has_integral_integral] assms(3)]) + by (meson assms has_integral_le integrable_integral) lemma has_integral_nonneg: fixes f :: "'n::euclidean_space \ real" - assumes "(f has_integral i) S" - and "\x. x \ S \ 0 \ f x" + assumes "(f has_integral i) S" and "\x. x \ S \ 0 \ f x" shows "0 \ i" - using has_integral_component_nonneg[of 1 f i S] - unfolding o_def - using assms - by auto + using assms has_integral_0 has_integral_le by blast lemma integral_nonneg: fixes f :: "'n::euclidean_space \ real" @@ -4862,16 +4751,11 @@ lemma has_integral_on_superset: fixes f :: "'n::euclidean_space \ 'a::banach" - assumes f: "(f has_integral i) S" + assumes "(f has_integral i) S" and "\x. x \ S \ f x = 0" and "S \ T" shows "(f has_integral i) T" -proof - - have "(\x. if x \ S then f x else 0) = (\x. if x \ T then f x else 0)" - using assms by fastforce - with f show ?thesis - by (simp only: has_integral_restrict_UNIV [symmetric, of f]) -qed + by (smt (verit, ccfv_SIG) assms has_integral_cong has_integral_restrict) lemma integrable_on_superset: fixes f :: "'n::euclidean_space \ 'a::banach" @@ -4906,14 +4790,14 @@ lemma has_integral_subset_component_le: fixes f :: "'n::euclidean_space \ 'm::euclidean_space" assumes k: "k \ Basis" - and as: "S \ T" "(f has_integral i) S" "(f has_integral j) T" "\x. x\T \ 0 \ f(x)\k" + and "S \ T" "(f has_integral i) S" "(f has_integral j) T" "\x. x\T \ 0 \ f(x)\k" shows "i\k \ j\k" proof - have \
: "((\x. if x \ S then f x else 0) has_integral i) UNIV" "((\x. if x \ T then f x else 0) has_integral j) UNIV" by (simp_all add: assms) show ?thesis - using as by (force intro!: has_integral_component_le[OF k \
]) + using assms by (force intro!: has_integral_component_le[OF k \
]) qed subsection\Integrals on set differences\ @@ -4961,13 +4845,7 @@ assume R: ?r show ?l unfolding negligible_def - proof safe - fix a b - have "negligible (s \ cbox a b)" - by (simp add: R) - then show "(indicator s has_integral 0) (cbox a b)" - by (meson Diff_iff Int_iff has_integral_negligible indicator_simps(2)) - qed + by (metis Diff_iff Int_iff R has_integral_negligible indicator_simps(2)) qed (simp add: negligible_Int) lemma negligible_translation: @@ -4994,8 +4872,8 @@ lemma negligible_translation_rev: assumes "negligible ((+) c ` S)" - shows "negligible S" -by (metis negligible_translation [OF assms, of "-c"] translation_galois) + shows "negligible S" + by (metis negligible_translation [OF assms, of "-c"] translation_galois) lemma negligible_atLeastAtMostI: "b \ a \ negligible {a..(b::real)}" using negligible_insert by fastforce @@ -5077,12 +4955,7 @@ lemma has_integral_Icc_iff_Ioo: fixes f :: "real \ 'a :: banach" shows "(f has_integral I) {a..b} \ (f has_integral I) {a<.. {a..b} - {a<.. 0}" - by (rule negligible_subset [of "{a,b}"]) auto - show "negligible {x \ {a<.. 0}" - by (rule negligible_subset [of "{}"]) auto -qed + by (metis box_real(1) cbox_interval has_integral_open_interval) lemma integrable_on_Icc_iff_Ioo: fixes f :: "real \ 'a :: banach" @@ -5099,8 +4972,7 @@ and "(f has_integral j) t" and "\x\t. 0 \ f x" shows "i \ j" - using has_integral_subset_component_le[OF _ assms(1), of 1 f i j] - using assms + using assms has_integral_subset_component_le[OF _ assms(1), of 1 f i j] by auto lemma integral_subset_component_le: @@ -5166,7 +5038,6 @@ show "cbox a b \ cbox ?a ?b" by (force simp: mem_box) qed - fix e :: real assume "e > 0" with lhs show "\B>0. \a b. ball 0 B \ cbox a b \ @@ -5659,12 +5530,7 @@ shows "(f has_integral (sum (\i. integral i f) \)) K" proof - have "f integrable_on L" if "L \ \" for L - proof - - have "L \ S" - using \K \ S\ \ that by blast - then show "f integrable_on L" - using that by (metis (no_types) f \ division_ofD(4) integrable_on_subcbox) - qed + by (smt (verit, best) assms cbox_division_memE f integrable_on_subcbox subset_trans that) then show ?thesis by (meson \ has_integral_combine_division has_integral_integrable_integral) qed @@ -5691,12 +5557,7 @@ shows "f integrable_on i" proof - have "f integrable_on i" if "i \ \" for i -proof - - have "i \ S" - using assms that by auto - then show "f integrable_on i" - using that by (metis (no_types) \ f division_ofD(4) integrable_on_subcbox) -qed + by (smt (verit, best) assms cbox_division_memE f integrable_on_subcbox order_trans that) then show ?thesis using \ integrable_combine_division by blast qed @@ -5711,12 +5572,7 @@ shows "(f has_integral (\(x,k)\p. i k)) S" proof - have *: "(f has_integral (\k\snd`p. integral k f)) S" - proof - - have "snd ` p division_of S" - by (simp add: assms(1) division_of_tagged_division) - with assms show ?thesis - by (metis (mono_tags, lifting) has_integral_combine_division has_integral_integrable_integral imageE prod.collapse) - qed + by (smt (verit, del_insts) assms division_of_tagged_division has_integral_combine_division has_integral_iff imageE prod.collapse) also have "(\k\snd`p. integral k f) = (\(x, k)\p. integral k f)" by (intro sum.over_tagged_division_lemma[OF assms(1), symmetric] integral_null) (simp add: content_eq_0_interior) @@ -5850,11 +5706,9 @@ have "L \ snd ` p" using \(x,L) \ p\ image_iff by fastforce then have "L \ q" "K \ q" "L \ K" - using that(1,3) q(1) unfolding r_def by auto - with q'(5) have "interior L = {}" - using interior_mono[OF \L \ K\] by blast - then show "content L *\<^sub>R f x = 0" - unfolding uv content_eq_0_interior[symmetric] by auto + using that q(1) unfolding r_def by auto + with q'(5) show "content L *\<^sub>R f x = 0" + by (metis \L \ K\ content_eq_0_interior inf.orderE interior_Int scaleR_eq_0_iff uv) qed show "finite (\(qq ` r))" by (meson finite_UN qq \finite r\ tagged_division_of_finite) @@ -5868,11 +5722,9 @@ using \(x, M) \ qq L\ \L \ r\ kl(2) by blast have empty: "interior (K \ L) = {}" by (metis DiffD1 interior_Int q'(5) r_def KL r) - have "interior M = {}" - by (metis (no_types, lifting) Int_assoc empty inf.absorb_iff2 interior_Int kl(1) subset_empty x r) - then show "content M *\<^sub>R f x = 0" - unfolding uv content_eq_0_interior[symmetric] - by auto + with that kl show "content M *\<^sub>R f x = 0" + by (metis content_eq_0_interior dual_order.refl inf.orderE inf_mono interior_mono + scaleR_eq_0_iff subset_empty uv x) qed ultimately have "norm (?SUM p + sum ?SUM (qq ` r) - integral (cbox a b) f) < e" apply (subst (asm) sum.Union_comp) @@ -5886,12 +5738,9 @@ have norm_le: "norm (cp - ip) \ e + k" if "norm ((cp + cr) - i) < e" "norm (cr - ir) < k" "ip + ir = i" for ir ip i cr cp::'a - proof - - from that show ?thesis - using norm_triangle_le[of "cp + cr - i" "- (cr - ir)"] - unfolding that(3)[symmetric] norm_minus_cancel - by (auto simp add: algebra_simps) - qed + using norm_triangle_le[of "cp + cr - i" "- (cr - ir)"] that + unfolding that(3)[symmetric] norm_minus_cancel + by (auto simp add: algebra_simps) have "?lhs = norm (?SUM p - (\(x, k)\p. integral k f))" unfolding split_def sum_subtractf .. @@ -5911,11 +5760,9 @@ proof - obtain u v where uv: "l = cbox u v" using inp p'(4) by blast - have "content (cbox u v) = 0" - unfolding content_eq_0_interior using that p(1) uv - by (auto dest: tagged_partial_division_ofD) then show ?thesis - using uv by blast + using uv that p + by (metis content_eq_0_interior dual_order.refl inf.orderE integral_null ne tagged_partial_division_ofD(5)) qed then have "(\(x, K)\p. integral K f) = (\K\snd ` p. integral K f)" apply (subst sum.reindex_nontrivial [OF \finite p\]) @@ -6526,14 +6373,10 @@ lemma has_integral_norm_bound_integral_component: fixes f :: "'n::euclidean_space \ 'a::banach" fixes g :: "'n \ 'b::euclidean_space" - assumes f: "(f has_integral i) S" - and g: "(g has_integral j) S" + assumes "(f has_integral i) S" and "(g has_integral j) S" and "\x. x \ S \ norm (f x) \ (g x)\k" shows "norm i \ j\k" - using integral_norm_bound_integral_component[of f S g k] - unfolding integral_unique[OF f] integral_unique[OF g] - using assms - by auto + by (metis assms has_integral_integrable integral_norm_bound_integral_component integral_unique) lemma uniformly_convergent_improper_integral: @@ -6631,7 +6474,7 @@ show ?thesis unfolding continuous_on_def - proof (safe intro!: tendstoI) + proof (intro strip tendstoI) fix e'::real and x assume "e' > 0" define e where "e = e' / (content (cbox a b) + 1)" @@ -6712,7 +6555,7 @@ case (elim x) from elim have "0 < norm (x - x0)" by simp have "closed_segment x0 x \ U" - by (rule \convex U\[unfolded convex_contains_segment, rule_format, OF \x0 \ U\ \x \ U\]) + by (simp add: assms closed_segment_subset elim(4)) from elim have [intro]: "x \ U" by auto have "?F x - ?F x0 - ?dF (x - x0) = integral (cbox a b) (\y. f x y - f x0 y - fx x0 y (x - x0))" @@ -6724,16 +6567,14 @@ also { fix t assume t: "t \ (cbox a b)" + then have deriv: + "((\x. f x t) has_derivative (fx y t)) (at y within X0 \ U)" + if "y \ X0 \ U" for y + using fx has_derivative_subset that by fastforce have seg: "\t. t \ {0..1} \ x0 + t *\<^sub>R (x - x0) \ X0 \ U" using \closed_segment x0 x \ U\ \closed_segment x0 x \ X0\ by (force simp: closed_segment_def algebra_simps) - from t have deriv: - "((\x. f x t) has_derivative (fx y t)) (at y within X0 \ U)" - if "y \ X0 \ U" for y - unfolding has_vector_derivative_def[symmetric] - using that \x \ X0\ - by (intro has_derivative_subset[OF fx]) auto have "\x. x \ X0 \ U \ onorm (blinfun_apply (fx x t) - (fx x0 t)) \ e" using fx_bound t by (auto simp add: norm_blinfun_def fun_diff_def blinfun.bilinear_simps[symmetric]) @@ -6848,7 +6689,7 @@ have fi[simp]: "f n integrable_on (cbox a b)" for n by (auto intro!: integrable_continuous assms) then obtain I where I: "\n. (f n has_integral I n) (cbox a b)" - by atomize_elim (auto simp: integrable_on_def intro!: choice) + unfolding integrable_on_def by metis moreover have gi[simp]: "g integrable_on (cbox a b)" @@ -6885,8 +6726,7 @@ using elim by (intro integral_norm_bound_integral) (auto intro!: integrable_diff) also have "\ < e" - using \0 < e\ - by (simp add: e'_def) + using \0 < e\ by (simp add: e'_def) finally show ?case . qed qed @@ -6935,7 +6775,7 @@ "\x. x\{a<.. (g has_vector_derivative g' x) (at x)" assumes "((\x. prod (f x) (g' x)) has_integral (prod (f b) (g b) - prod (f a) (g a) - y)) {a..b}" shows "((\x. prod (f' x) (g x)) has_integral y) {a..b}" - by (rule integration_by_parts_interior_strong[of _ "{}" _ _ f g f' g']) (insert assms, simp_all) + by (rule integration_by_parts_interior_strong[of _ "{}" _ _ f g f' g']) (use assms in simp_all) lemma integration_by_parts: fixes prod :: "_ \ _ \ 'b :: banach" @@ -6945,7 +6785,7 @@ "\x. x\{a..b} \ (g has_vector_derivative g' x) (at x)" assumes "((\x. prod (f x) (g' x)) has_integral (prod (f b) (g b) - prod (f a) (g a) - y)) {a..b}" shows "((\x. prod (f' x) (g x)) has_integral y) {a..b}" - by (rule integration_by_parts_interior[of _ _ _ f g f' g']) (insert assms, simp_all) + by (rule integration_by_parts_interior[of _ _ _ f g f' g']) (use assms in simp_all) lemma integrable_by_parts_interior_strong: fixes prod :: "_ \ _ \ 'b :: banach" @@ -6973,7 +6813,7 @@ "\x. x\{a<.. (g has_vector_derivative g' x) (at x)" assumes "(\x. prod (f x) (g' x)) integrable_on {a..b}" shows "(\x. prod (f' x) (g x)) integrable_on {a..b}" - by (rule integrable_by_parts_interior_strong[of _ "{}" _ _ f g f' g']) (insert assms, simp_all) + by (rule integrable_by_parts_interior_strong[of _ "{}" _ _ f g f' g']) (use assms in simp_all) lemma integrable_by_parts: fixes prod :: "_ \ _ \ 'b :: banach" @@ -6983,7 +6823,7 @@ "\x. x\{a..b} \ (g has_vector_derivative g' x) (at x)" assumes "(\x. prod (f x) (g' x)) integrable_on {a..b}" shows "(\x. prod (f' x) (g x)) integrable_on {a..b}" - by (rule integrable_by_parts_interior_strong[of _ "{}" _ _ f g f' g']) (insert assms, simp_all) + by (rule integrable_by_parts_interior_strong[of _ "{}" _ _ f g f' g']) (use assms in simp_all) subsection \Integration by substitution\ @@ -7433,7 +7273,7 @@ (\k. integral {c..} (f k)) \ integral {c..} (\x. exp (-a*x))" proof (intro monotone_convergence_increasing allI ballI) fix k ::nat - have "(\x. exp (-a*x)) integrable_on {c..of_real k}" (is ?P) + have "(\x. exp (-a*x)) integrable_on {c..of_real k}" unfolding f_def by (auto intro!: continuous_intros integrable_continuous_real) hence "(f k) integrable_on {c..of_real k}" by (rule integrable_eq) (simp add: f_def) @@ -7444,7 +7284,8 @@ have "sequentially \ principal {nat \x\..}" unfolding at_top_def by (simp add: Inf_lower) also have "{nat \x\..} \ {k. x \ real k}" by auto also have "inf (principal \) (principal {k. \x \ real k}) = - principal ({k. x \ real k} \ {k. \x \ real k})" by simp + principal ({k. x \ real k} \ {k. \x \ real k})" + by simp also have "{k. x \ real k} \ {k. \x \ real k} = {}" by blast finally have "inf sequentially (principal {k. \x \ real k}) = bot" by (simp add: inf.coboundedI1 bot_unique) @@ -7468,7 +7309,7 @@ have "(\k. exp (-a*c)/a - exp (-a * of_nat k)/a) \ exp (-a*c)/a - 0/a" by (intro tendsto_intros filterlim_compose[OF exp_at_bot] filterlim_tendsto_neg_mult_at_bot[OF tendsto_const] filterlim_real_sequentially)+ - (insert a, simp_all) + (use a in simp_all) moreover from eventually_gt_at_top[of "nat \c\"] have "eventually (\k. of_nat k > c) sequentially" by eventually_elim linarith @@ -7493,30 +7334,24 @@ define f where "f = (\k x. if x \ {inverse (of_nat (Suc k))..c} then x powr a else 0)" define F where "F = (\k. if inverse (of_nat (Suc k)) \ c then c powr (a+1)/(a+1) - inverse (real (Suc k)) powr (a+1)/(a+1) else 0)" - { - fix k :: nat - have "(f k has_integral F k) {0..c}" - proof (cases "inverse (of_nat (Suc k)) \ c") - case True - { - fix x assume x: "x \ inverse (1 + real k)" - have "0 < inverse (1 + real k)" by simp - also note x - finally have "x > 0" . - } note x = this - hence "((\x. x powr a) has_integral c powr (a + 1) / (a + 1) - + have has_integral_f: "(f k has_integral F k) {0..c}" for k::nat + proof (cases "inverse (of_nat (Suc k)) \ c") + case True + have x: "x > 0" if "x \ inverse (1 + real k)" for x + by (smt (verit) that inverse_Suc of_nat_Suc) + hence "((\x. x powr a) has_integral c powr (a + 1) / (a + 1) - inverse (real (Suc k)) powr (a + 1) / (a + 1)) {inverse (real (Suc k))..c}" - using True a by (intro fundamental_theorem_of_calculus) - (auto intro!: derivative_eq_intros continuous_on_powr' continuous_on_const - simp: has_real_derivative_iff_has_vector_derivative [symmetric]) - with True show ?thesis unfolding f_def F_def by (subst has_integral_restrict) simp_all - next - case False - thus ?thesis unfolding f_def F_def by (subst has_integral_restrict) auto - qed - } note has_integral_f = this - have integral_f: "integral {0..c} (f k) = F k" for k - using has_integral_f[of k] by (rule integral_unique) + using True a by (intro fundamental_theorem_of_calculus) + (auto intro!: derivative_eq_intros continuous_on_powr' continuous_on_const + simp: has_real_derivative_iff_has_vector_derivative [symmetric]) + with True show ?thesis unfolding f_def F_def by (subst has_integral_restrict) simp_all + next + case False + thus ?thesis unfolding f_def F_def + by (subst has_integral_restrict) auto + qed + then have integral_f: "integral {0..c} (f k) = F k" for k + by blast have A: "(\x. x powr a) integrable_on {0..c} \ (\k. integral {0..c} (f k)) \ integral {0..c} (\x. x powr a)" @@ -7527,9 +7362,8 @@ fix k :: nat and x :: real { assume x: "inverse (real (Suc k)) \ x" - have "inverse (real (Suc (Suc k))) \ inverse (real (Suc k))" by (simp add: field_simps) - also note x - finally have "inverse (real (Suc (Suc k))) \ x" . + then have "inverse (real (Suc (Suc k))) \ x" + using dual_order.trans by fastforce } thus "f k x \ f (Suc k) x" by (auto simp: f_def simp del: of_nat_Suc) next @@ -7607,7 +7441,7 @@ case False have "(f n has_integral 0) {a}" by (rule has_integral_refl) hence "(f n has_integral 0) {a..}" - by (rule has_integral_on_superset) (insert False, simp_all add: f_def) + using False f_def by force with False show ?thesis by (simp add: integral_unique) qed @@ -7639,7 +7473,7 @@ with assms have "a powr (e + 1) \ n powr (e + 1)" by (intro powr_mono2') simp_all with assms show ?thesis by (auto simp: divide_simps F_def integral_f) - qed (insert assms, simp add: integral_f F_def field_split_simps) + qed (use assms in \simp add: integral_f F_def field_split_simps\) thus "bounded (range(\k. integral {a..} (f k)))" unfolding bounded_iff by (intro exI[of _ "-F a"]) auto qed @@ -7654,9 +7488,9 @@ filterlim_ident filterlim_real_sequentially | simp)+) hence "(\n. F n - F a) \ -F a" by simp ultimately have "(\n. integral {a..} (f n)) \ -F a" by (blast intro: Lim_transform_eventually) - from conjunct2[OF *] and this - have "integral {a..} (\x. x powr e) = -F a" by (rule LIMSEQ_unique) - with conjunct1[OF *] show ?thesis + then have "integral {a..} (\x. x powr e) = -F a" + using "*" LIMSEQ_unique by blast + with * show ?thesis by (simp add: has_integral_integral F_def) qed diff -r ba2afdd29e1d -r 032a4344903e src/HOL/Analysis/Homeomorphism.thy --- a/src/HOL/Analysis/Homeomorphism.thy Tue Aug 01 11:27:55 2023 +0200 +++ b/src/HOL/Analysis/Homeomorphism.thy Thu Aug 03 19:10:43 2023 +0200 @@ -210,7 +210,7 @@ ultimately have *: "continuous (at x within affine hull S) (\x. surf (proj x))" by (simp add: continuous_within Lim_transform_within_set continuous_on_eq_continuous_within) show ?thesis - by (intro continuous_within_subset [where s = "affine hull S", OF _ Int_lower2] continuous_intros *) + by (intro continuous_within_subset [where S = "affine hull S", OF _ Int_lower2] continuous_intros *) qed have cont_nosp2: "continuous_on ?CBALL (\x. norm x *\<^sub>R ((surf o proj) x))" by (simp add: continuous_on_eq_continuous_within cont_nosp) diff -r ba2afdd29e1d -r 032a4344903e src/HOL/Analysis/Homotopy.thy --- a/src/HOL/Analysis/Homotopy.thy Tue Aug 01 11:27:55 2023 +0200 +++ b/src/HOL/Analysis/Homotopy.thy Thu Aug 03 19:10:43 2023 +0200 @@ -712,7 +712,7 @@ proposition homotopic_loops_subset: "\homotopic_loops S p q; S \ t\ \ homotopic_loops t p q" - by (fastforce simp add: homotopic_loops) + by (fastforce simp: homotopic_loops) proposition homotopic_loops_eq: "\path p; path_image p \ S; pathfinish p = pathstart p; \t. t \ {0..1} \ p(t) = q(t)\ @@ -1040,17 +1040,15 @@ have "homotopic_paths S (subpath u w g +++ subpath w v g) ((subpath u v g +++ subpath v w g) +++ subpath w v g)" by (simp add: hom homotopic_paths_join homotopic_paths_sym wvg) - also have "homotopic_paths S ((subpath u v g +++ subpath v w g) +++ subpath w v g) - (subpath u v g +++ subpath v w g +++ subpath w v g)" + also have "homotopic_paths S \ (subpath u v g +++ subpath v w g +++ subpath w v g)" using wvg vug \path g\ by (metis homotopic_paths_assoc homotopic_paths_sym path_image_subpath_commute path_subpath pathfinish_subpath pathstart_subpath u v w) - also have "homotopic_paths S (subpath u v g +++ subpath v w g +++ subpath w v g) - (subpath u v g +++ linepath (pathfinish (subpath u v g)) (pathfinish (subpath u v g)))" + also have "homotopic_paths S \ (subpath u v g +++ linepath (pathfinish (subpath u v g)) (pathfinish (subpath u v g)))" using wvg vug \path g\ by (metis homotopic_paths_join homotopic_paths_linv homotopic_paths_refl path_image_subpath_commute path_subpath pathfinish_subpath pathstart_join pathstart_subpath reversepath_subpath u v) - also have "homotopic_paths S (subpath u v g +++ linepath (pathfinish (subpath u v g)) (pathfinish (subpath u v g))) (subpath u v g)" + also have "homotopic_paths S \ (subpath u v g)" using vug \path g\ by (metis homotopic_paths_rid path_image_subpath_commute path_subpath u v) finally have "homotopic_paths S (subpath u w g +++ subpath w v g) (subpath u v g)" . then show ?thesis @@ -1077,11 +1075,11 @@ qed lemma homotopic_loops_imp_path_component_value: - "\homotopic_loops S p q; 0 \ t; t \ 1\ \ path_component S (p t) (q t)" -apply (clarsimp simp: homotopic_loops_def homotopic_with_def path_defs) -apply (rule_tac x="h \ (\u. (u, t))" in exI) -apply (fastforce elim!: continuous_on_subset intro!: continuous_intros) -done + "\homotopic_loops S p q; 0 \ t; t \ 1\ \ path_component S (p t) (q t)" + apply (clarsimp simp: homotopic_loops_def homotopic_with_def path_defs) + apply (rule_tac x="h \ (\u. (u, t))" in exI) + apply (fastforce elim!: continuous_on_subset intro!: continuous_intros) + done lemma homotopic_points_eq_path_component: "homotopic_loops S (linepath a a) (linepath b b) \ path_component S a b" @@ -1109,12 +1107,12 @@ lemma simply_connected_imp_path_connected: fixes S :: "_::real_normed_vector set" shows "simply_connected S \ path_connected S" -by (simp add: simply_connected_def path_connected_eq_homotopic_points) + by (simp add: simply_connected_def path_connected_eq_homotopic_points) lemma simply_connected_imp_connected: fixes S :: "_::real_normed_vector set" shows "simply_connected S \ connected S" -by (simp add: path_connected_imp_connected simply_connected_imp_path_connected) + by (simp add: path_connected_imp_connected simply_connected_imp_path_connected) lemma simply_connected_eq_contractible_loop_any: fixes S :: "_::real_normed_vector set" @@ -1123,13 +1121,10 @@ \ homotopic_loops S p (linepath a a))" (is "?lhs = ?rhs") proof - assume ?lhs then show ?rhs - unfolding simply_connected_def by force -next assume ?rhs then show ?lhs unfolding simply_connected_def by (metis pathfinish_in_path_image subsetD homotopic_loops_trans homotopic_loops_sym) -qed +qed (force simp: simply_connected_def) lemma simply_connected_eq_contractible_loop_some: fixes S :: "_::real_normed_vector set" @@ -1154,15 +1149,7 @@ S = {} \ (\a \ S. \p. path p \ path_image p \ S \ pathfinish p = pathstart p \ homotopic_loops S p (linepath a a))" - (is "?lhs = ?rhs") -proof (cases "S = {}") - case True then show ?thesis by force -next - case False - then obtain a where "a \ S" by blast - then show ?thesis - by (meson False homotopic_loops_sym homotopic_loops_trans simply_connected_def simply_connected_eq_contractible_loop_any) -qed + by (meson ex_in_conv homotopic_loops_sym homotopic_loops_trans simply_connected_def simply_connected_eq_contractible_loop_any) lemma simply_connected_eq_contractible_path: fixes S :: "_::real_normed_vector set" @@ -1207,14 +1194,12 @@ using that by (smt (verit, best) homotopic_paths_join homotopic_paths_linv homotopic_paths_rid homotopic_paths_sym homotopic_paths_trans pathstart_linepath) - also have "homotopic_paths S (p +++ reversepath q +++ q) - ((p +++ reversepath q) +++ q)" + also have "homotopic_paths S \ ((p +++ reversepath q) +++ q)" by (simp add: that homotopic_paths_assoc) - also have "homotopic_paths S ((p +++ reversepath q) +++ q) - (linepath (pathstart q) (pathstart q) +++ q)" + also have "homotopic_paths S \ (linepath (pathstart q) (pathstart q) +++ q)" using * [of "p +++ reversepath q"] that by (simp add: homotopic_paths_assoc homotopic_paths_join path_image_join) - also have "homotopic_paths S (linepath (pathstart q) (pathstart q) +++ q) q" + also have "homotopic_paths S \ q" using that homotopic_paths_lid by blast finally show ?thesis . qed @@ -1238,7 +1223,7 @@ have "path (fst \ p)" by (simp add: continuous_on_fst Path_Connected.path_continuous_image [OF \path p\]) moreover have "path_image (fst \ p) \ S" - using that by (force simp add: path_image_def) + using that by (force simp: path_image_def) ultimately have p1: "homotopic_loops S (fst \ p) (linepath a a)" using S that by (simp add: simply_connected_eq_contractible_loop_any pathfinish_def pathstart_def) @@ -1291,12 +1276,12 @@ corollary contractible_imp_connected: fixes S :: "_::real_normed_vector set" shows "contractible S \ connected S" -by (simp add: contractible_imp_simply_connected simply_connected_imp_connected) + by (simp add: contractible_imp_simply_connected simply_connected_imp_connected) lemma contractible_imp_path_connected: fixes S :: "_::real_normed_vector set" shows "contractible S \ path_connected S" -by (simp add: contractible_imp_simply_connected simply_connected_imp_path_connected) + by (simp add: contractible_imp_simply_connected simply_connected_imp_path_connected) lemma nullhomotopic_through_contractible: fixes S :: "_::topological_space set" @@ -1351,7 +1336,7 @@ with \path_connected U\ show ?thesis by blast qed then have "homotopic_with_canon (\h. True) S U (\x. c2) (\x. c1)" - by (auto simp add: path_component homotopic_constant_maps) + by (auto simp: path_component homotopic_constant_maps) then show ?thesis using c1 c2 homotopic_with_symD homotopic_with_trans by blast qed @@ -1419,7 +1404,7 @@ using \a \ S\ unfolding homotopic_with_def apply (rule_tac x="\y. (1 - (fst y)) *\<^sub>R snd y + (fst y) *\<^sub>R a" in exI) - apply (force simp add: P intro: continuous_intros) + apply (force simp: P intro: continuous_intros) done then show ?thesis using that by blast @@ -1516,18 +1501,17 @@ where "locally P S \ \w x. openin (top_of_set S) w \ x \ w - \ (\u v. openin (top_of_set S) u \ P v \ x \ u \ u \ v \ v \ w)" + \ (\U V. openin (top_of_set S) U \ P V \ x \ U \ U \ V \ V \ w)" lemma locallyI: assumes "\w x. \openin (top_of_set S) w; x \ w\ - \ \u v. openin (top_of_set S) u \ P v \ x \ u \ u \ v \ v \ w" + \ \U V. openin (top_of_set S) U \ P V \ x \ U \ U \ V \ V \ w" shows "locally P S" using assms by (force simp: locally_def) lemma locallyE: assumes "locally P S" "openin (top_of_set S) w" "x \ w" - obtains u v where "openin (top_of_set S) u" - "P v" "x \ u" "u \ v" "v \ w" + obtains U V where "openin (top_of_set S) U" "P V" "x \ U" "U \ V" "V \ w" using assms unfolding locally_def by meson lemma locally_mono: @@ -1636,28 +1620,26 @@ using \f ` S = T\ f \V \ S\ by auto have contvf: "continuous_on V f" using \V \ S\ continuous_on_subset f(3) by blast - have "f ` V \ W" - using uv using Int_lower2 gw image_subsetI mem_Collect_eq subset_iff by auto - then have contvg: "continuous_on (f ` V) g" - using \W \ T\ continuous_on_subset [OF g(3)] by blast - have "V \ g ` f ` V" - by (metis \V \ S\ hom homeomorphism_def homeomorphism_of_subsets order_refl) - then have homv: "homeomorphism V (f ` V) f g" - using \V \ S\ f by (auto simp: homeomorphism_def contvf contvg) have "openin (top_of_set (g ` T)) U" using \g ` T = S\ by (simp add: osu) then have "openin (top_of_set T) (T \ g -` U)" using \continuous_on T g\ continuous_on_open [THEN iffD1] by blast moreover have "\V. Q V \ y \ (T \ g -` U) \ (T \ g -` U) \ V \ V \ W" proof (intro exI conjI) + show "f ` V \ W" + using uv using Int_lower2 gw image_subsetI mem_Collect_eq subset_iff by auto + then have contvg: "continuous_on (f ` V) g" + using \W \ T\ continuous_on_subset [OF g(3)] by blast + have "V \ g ` f ` V" + by (metis \V \ S\ hom homeomorphism_def homeomorphism_of_subsets order_refl) + then have homv: "homeomorphism V (f ` V) f g" + using \V \ S\ f by (auto simp: homeomorphism_def contvf contvg) show "Q (f ` V)" using Q homv \P V\ by blast show "y \ T \ g -` U" using T(2) \y \ W\ \g y \ U\ by blast show "T \ g -` U \ f ` V" using g(1) image_iff uv(3) by fastforce - show "f ` V \ W" - using \f ` V \ W\ by blast qed ultimately show "\U. openin (top_of_set T) U \ (\v. Q v \ y \ U \ U \ v \ v \ W)" by meson @@ -1692,8 +1674,7 @@ fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes f: "linear f" "inj f" and iff: "\S. P (f ` S) \ Q S" shows "locally P (f ` S) \ locally Q S" - using homeomorphism_locally [of "f`S" _ _ f] linear_homeomorphism_image [OF f] - by (metis (no_types, lifting) homeomorphism_image2 iff) + by (smt (verit) f homeomorphism_image2 homeomorphism_locally iff linear_homeomorphism_image) lemma locally_open_map_image: fixes f :: "'a::real_normed_vector \ 'b::real_normed_vector" @@ -1782,8 +1763,8 @@ shows "R a b" proof - have "\a b c. \a \ S; b \ S; c \ S; R a b\ \ R a c" - apply (rule connected_induction_simple [OF \connected S\], simp_all) - by (meson local.sym local.trans opI openin_imp_subset subsetCE) + by (smt (verit, ccfv_threshold) connected_induction_simple [OF \connected S\] + assms openin_imp_subset subset_eq) then show ?thesis by (metis etc opI) qed @@ -1833,13 +1814,13 @@ compact v \ x \ u \ u \ v \ v \ S \ T" if "open T" "x \ S" "x \ T" for x T proof - - obtain u v where uv: "x \ u" "u \ v" "v \ S" "compact v" "openin (top_of_set S) u" + obtain U V where uv: "x \ U" "U \ V" "V \ S" "compact V" "openin (top_of_set S) U" using r [OF \x \ S\] by auto obtain e where "e>0" and e: "cball x e \ T" using open_contains_cball \open T\ \x \ T\ by blast show ?thesis - apply (rule_tac x="(S \ ball x e) \ u" in exI) - apply (rule_tac x="cball x e \ v" in exI) + apply (rule_tac x="(S \ ball x e) \ U" in exI) + apply (rule_tac x="cball x e \ V" in exI) using that \e > 0\ e uv apply auto done @@ -1860,16 +1841,8 @@ shows "locally compact S \ (\x \ S. \U. x \ U \ openin (top_of_set S) U \ compact(closure U) \ closure U \ S)" - (is "?lhs = ?rhs") -proof - assume ?lhs - then show ?rhs - by (meson bounded_subset closure_minimal compact_closure compact_imp_bounded - compact_imp_closed dual_order.trans locally_compactE) -next - assume ?rhs then show ?lhs - by (meson closure_subset locally_compact) -qed + by (smt (verit, ccfv_threshold) bounded_subset closure_closed closure_mono closure_subset + compact_closure compact_imp_closed order.trans locally_compact) lemma locally_compact_Int_cball: fixes S :: "'a :: heine_borel set" @@ -2008,8 +1981,9 @@ then show ?rhs unfolding locally_def apply (elim all_forward imp_forward asm_rl exE) - apply (rule_tac x = "u \ ball x 1" in exI) - apply (rule_tac x = "v \ cball x 1" in exI) + apply (rename_tac U V) + apply (rule_tac x = "U \ ball x 1" in exI) + apply (rule_tac x = "V \ cball x 1" in exI) apply (force intro: openin_trans) done next @@ -2019,7 +1993,7 @@ lemma locally_compact_openin_Un: fixes S :: "'a::euclidean_space set" - assumes LCS: "locally compact S" and LCT:"locally compact T" + assumes LCS: "locally compact S" and LCT: "locally compact T" and opS: "openin (top_of_set (S \ T)) S" and opT: "openin (top_of_set (S \ T)) T" shows "locally compact (S \ T)" @@ -2453,7 +2427,7 @@ "\v x. \openin (top_of_set S) v; x \ v\ \ \u. openin (top_of_set S) u \ path_connected u \ x \ u \ u \ v" shows "locally path_connected S" - by (force simp add: locally_def dest: assms) + by (force simp: locally_def dest: assms) lemma locally_path_connected_2: assumes "locally path_connected S" @@ -2525,7 +2499,7 @@ proof assume ?lhs then show ?rhs - by (fastforce simp add: locally_connected) + by (fastforce simp: locally_connected) next assume ?rhs have *: "\T. openin (top_of_set S) T \ x \ T \ T \ c" @@ -3577,7 +3551,7 @@ show thesis using homotopic_with_compose_continuous_map_right [OF homotopic_with_compose_continuous_map_left [OF b g] f] - by (force simp add: that) + by (force simp: that) qed lemma nullhomotopic_into_contractible_space: @@ -3901,7 +3875,7 @@ by (rule homotopic_with_compose_continuous_left [where Y=T]) (use f in \auto simp: hom homotopic_with_symD\) ultimately show ?thesis - using that homotopic_with_trans by (fastforce simp add: o_def) + using that homotopic_with_trans by (fastforce simp: o_def) qed lemma homotopy_eqv_cohomotopic_triviality_null: @@ -3942,7 +3916,7 @@ by (rule homotopic_with_compose_continuous_right [where X=T]) (use f in \auto simp: hom homotopic_with_symD\) ultimately show ?thesis - using homotopic_with_trans by (fastforce simp add: o_def) + using homotopic_with_trans by (fastforce simp: o_def) qed lemma homotopy_eqv_homotopic_triviality_null: @@ -5002,7 +4976,7 @@ using \ by (auto simp: pairwise_def bij_betw_def inj_on_def) qed (use affine_hull_open assms that in auto) then show ?thesis - using \ \P \ U\ bij_betwE by (fastforce simp add: intro!: that) + using \ \P \ U\ bij_betwE by (fastforce simp: intro!: that) next case False with DIM_positive have "DIM('a) = 1" @@ -5089,7 +5063,7 @@ using \ by (auto simp: pairwise_def bij_betw_def inj_on_def) qed then show ?thesis - using \ \P \ U\ bij_betwE by (fastforce simp add: intro!: that) + using \ \P \ U\ bij_betwE by (fastforce simp: intro!: that) next case False with aff_dim_geq [of S] consider "aff_dim S = -1" | "aff_dim S = 0" | "aff_dim S = 1" by linarith @@ -5173,9 +5147,9 @@ using \T \ affine hull S\ h by auto qed show "\x. x \ T \ g' (f' x) = x" - using h j hom homeomorphism_apply1 by (fastforce simp add: f'_def g'_def) + using h j hom homeomorphism_apply1 by (fastforce simp: f'_def g'_def) show "\y. y \ T \ f' (g' y) = y" - using h j hom homeomorphism_apply2 by (fastforce simp add: f'_def g'_def) + using h j hom homeomorphism_apply2 by (fastforce simp: f'_def g'_def) qed next have \
: "\x y. \x \ affine hull S; h x = h y; y \ S\ \ x \ S" @@ -5241,7 +5215,7 @@ and d: "\x'. \dist x' a \ r; x' \ a; dist x' x < d\ \ dist (g x') (g x) < e" using contg False x \e>0\ - unfolding continuous_on_iff by (fastforce simp add: dist_commute intro: that) + unfolding continuous_on_iff by (fastforce simp: dist_commute intro: that) show ?thesis using \d > 0\ \x \ a\ by (rule_tac x="min d (norm(x - a))" in exI) @@ -5287,7 +5261,7 @@ next assume ?rhs then show ?lhs - using equal continuous_on_const by (force simp add: homotopic_with) + using equal continuous_on_const by (force simp: homotopic_with) qed next case greater diff -r ba2afdd29e1d -r 032a4344903e src/HOL/Analysis/Polytope.thy --- a/src/HOL/Analysis/Polytope.thy Tue Aug 01 11:27:55 2023 +0200 +++ b/src/HOL/Analysis/Polytope.thy Thu Aug 03 19:10:43 2023 +0200 @@ -1355,7 +1355,7 @@ case False have "x = (((1 - v) * ub) *\<^sub>R b + (v * uc) *\<^sub>R c) /\<^sub>R ux" by (metis \ux \ 0\ uxx mult.commute right_inverse scaleR_one scaleR_scaleR) - also have "... = (1 - v * uc / ux) *\<^sub>R b + (v * uc / ux) *\<^sub>R c" + also have "\ = (1 - v * uc / ux) *\<^sub>R b + (v * uc / ux) *\<^sub>R c" using \ux \ 0\ equx apply (auto simp: field_split_simps) by (metis add.commute add_diff_eq add_divide_distrib diff_add_cancel scaleR_add_left) finally have "x = (1 - v * uc / ux) *\<^sub>R b + (v * uc / ux) *\<^sub>R c" . @@ -1501,7 +1501,7 @@ proof have "?lhs \ convex hull {x. x extreme_point_of S}" using Krein_Milman_Minkowski assms by blast - also have "... \ ?rhs" + also have "\ \ ?rhs" proof (rule hull_mono) show "{x. x extreme_point_of S} \ frontier S" using closure_subset @@ -1511,7 +1511,7 @@ next have "?rhs \ convex hull S" by (metis Diff_subset \compact S\ closure_closed compact_eq_bounded_closed frontier_def hull_mono) - also have "... \ ?lhs" + also have "\ \ ?lhs" by (simp add: \convex S\ hull_same) finally show "?rhs \ ?lhs" . qed @@ -1626,12 +1626,11 @@ done lemma polyhedron_UNIV [iff]: "polyhedron UNIV" - unfolding polyhedron_def - by (rule_tac x="{}" in exI) auto + using polyhedron_def by auto lemma polyhedron_Inter [intro,simp]: - "\finite F; \S. S \ F \ polyhedron S\ \ polyhedron(\F)" -by (induction F rule: finite_induct) auto + "\finite F; \S. S \ F \ polyhedron S\ \ polyhedron(\F)" + by (induction F rule: finite_induct) auto lemma polyhedron_empty [iff]: "polyhedron ({} :: 'a :: euclidean_space set)" @@ -1640,10 +1639,7 @@ have "\a. a \ 0 \ (\b. {x. i \ x \ -1} = {x. a \ x \ b})" by (rule_tac x="i" in exI) (force simp: i_def SOME_Basis nonzero_Basis) moreover have "\a b. a \ 0 \ {x. -i \ x \ - 1} = {x. a \ x \ b}" - apply (rule_tac x="-i" in exI) - apply (rule_tac x="-1" in exI) - apply (simp add: i_def SOME_Basis nonzero_Basis) - done + by (metis Basis_zero SOME_Basis i_def neg_0_equal_iff_equal) ultimately show ?thesis unfolding polyhedron_def by (rule_tac x="{{x. i \ x \ -1}, {x. -i \ x \ -1}}" in exI) force @@ -1664,7 +1660,7 @@ lemma polyhedron_halfspace_ge: fixes a :: "'a :: euclidean_space" shows "polyhedron {x. a \ x \ b}" -using polyhedron_halfspace_le [of "-a" "-b"] by simp + using polyhedron_halfspace_le [of "-a" "-b"] by simp lemma polyhedron_hyperplane: fixes a :: "'a :: euclidean_space" @@ -1679,7 +1675,7 @@ lemma affine_imp_polyhedron: fixes S :: "'a :: euclidean_space set" shows "affine S \ polyhedron S" -by (metis affine_hull_eq polyhedron_Inter polyhedron_hyperplane affine_hull_finite_intersection_hyperplanes [of S]) + by (metis affine_hull_finite_intersection_hyperplanes hull_same polyhedron_Inter polyhedron_hyperplane) lemma polyhedron_imp_closed: fixes S :: "'a :: euclidean_space set" @@ -1694,7 +1690,7 @@ lemma polyhedron_affine_hull: fixes S :: "'a :: euclidean_space set" shows "polyhedron(affine hull S)" -by (simp add: affine_imp_polyhedron) + by (simp add: affine_imp_polyhedron) subsection\Canonical polyhedron representation making facial structure explicit\ @@ -1704,14 +1700,7 @@ shows "polyhedron S \ (\F. finite F \ S = (affine hull S) \ \F \ (\h \ F. \a b. a \ 0 \ h = {x. a \ x \ b}))" - (is "?lhs = ?rhs") -proof - assume ?lhs then show ?rhs - using hull_subset polyhedron_def by fastforce -next - assume ?rhs then show ?lhs - by (metis polyhedron_Int polyhedron_Inter polyhedron_affine_hull polyhedron_halfspace_le) -qed + by (metis hull_subset inf.absorb_iff2 polyhedron_Int polyhedron_affine_hull polyhedron_def) proposition rel_interior_polyhedron_explicit: assumes "finite F" @@ -1744,13 +1733,13 @@ define \ where "\ = min (1/2) (e / 2 / norm(z - x))" have "norm (\ *\<^sub>R x - \ *\<^sub>R z) = norm (\ *\<^sub>R (x - z))" by (simp add: \_def algebra_simps norm_mult) - also have "... = \ * norm (x - z)" + also have "\ = \ * norm (x - z)" using \e > 0\ by (simp add: \_def) - also have "... < e" + also have "\ < e" using \z \ x\ \e > 0\ by (simp add: \_def min_def field_split_simps norm_minus_commute) finally have lte: "norm (\ *\<^sub>R x - \ *\<^sub>R z) < e" . have \_aff: "\ *\<^sub>R z + (1 - \) *\<^sub>R x \ affine hull S" - by (metis \x \ S\ add.commute affine_affine_hull diff_add_cancel hull_inc mem_affine zaff) + by (simp add: \x \ S\ hull_inc mem_affine zaff) have "\ *\<^sub>R z + (1 - \) *\<^sub>R x \ S" using ins [OF _ \_aff] by (simp add: algebra_simps lte) then obtain l where l: "0 < l" "l < 1" and ls: "(l *\<^sub>R z + (1 - l) *\<^sub>R x) \ S" @@ -1809,7 +1798,7 @@ have "a h \ 0" and "h = {x. a h \ x \ b h}" "h \ \F = \F" using \h \ F\ ab by auto then have "(affine hull S) \ {x. a h \ x \ b h} \ {}" - by (metis (no_types) affine_hull_eq_empty inf.absorb_iff2 inf_assoc inf_bot_right inf_commute seq that(2)) + by (metis affine_hull_eq_empty inf.absorb_iff1 inf_assoc inf_bot_left seq that(2)) moreover have "\ (affine hull S \ {x. a h \ x \ b h})" using \h = {x. a h \ x \ b h}\ that(2) by blast ultimately show ?thesis @@ -1822,14 +1811,14 @@ affine hull S \ {x. a h \ x \ b h} = affine hull S \ h \ (\w \ affine hull S. (w + a h) \ affine hull S)" by metis - have seq2: "S = affine hull S \ (\h\{h \ F. \ affine hull S \ h}. {x. a h \ x \ b h})" - by (subst seq) (auto simp: ab INT_extend_simps) + let ?F = "(\h. {x. a h \ x \ b h}) ` {h \ F. \ affine hull S \ h}" show ?thesis - apply (rule_tac x="(\h. {x. a h \ x \ b h}) ` {h. h \ F \ \(affine hull S \ h)}" in exI) - apply (intro conjI seq2) - using \finite F\ apply force - using ab apply blast - done + proof (intro exI conjI) + show "finite ?F" + using \finite F\ by force + show "S = affine hull S \ \ ?F" + by (subst seq) (auto simp: ab INT_extend_simps) + qed (use ab in blast) qed next assume ?rhs then show ?lhs @@ -1887,12 +1876,7 @@ (\F. finite F \ S = (affine hull S) \ \F \ (\h \ F. \a b. a \ 0 \ h = {x. a \ x \ b}) \ (\F'. F' \ F \ S \ (affine hull S) \ \F'))" - (is "?lhs = ?rhs") -proof - assume ?lhs - then show ?rhs - by (force simp: polyhedron_Int_affine_parallel_minimal elim!: ex_forward) -qed (auto simp: polyhedron_Int_affine elim!: ex_forward) + by (metis polyhedron_Int_affine polyhedron_Int_affine_parallel_minimal) proposition facet_of_polyhedron_explicit: assumes "finite F" @@ -1928,7 +1912,7 @@ have "x \ h" using that xint by auto then have able: "a h \ x \ b h" using faceq that by blast - also have "... < a h \ z" using \z \ h\ faceq [OF that] xint by auto + also have "\ < a h \ z" using \z \ h\ faceq [OF that] xint by auto finally have xltz: "a h \ x < a h \ z" . define l where "l = (b h - a h \ x) / (a h \ z - a h \ x)" define w where "w = (1 - l) *\<^sub>R x + l *\<^sub>R z" @@ -1942,14 +1926,15 @@ moreover have "l * (a i \ z) \ l * b i" proof (rule mult_left_mono) show "a i \ z \ b i" - by (metis Diff_insert_absorb Inter_iff Set.set_insert \h \ F\ faceq insertE mem_Collect_eq that zint) + by (metis DiffI Inter_iff empty_iff faceq insertE mem_Collect_eq that zint) qed (use \0 < l\ in auto) ultimately show ?thesis by (simp add: w_def algebra_simps) qed have weq: "a h \ w = b h" using xltz unfolding w_def l_def by (simp add: algebra_simps) (simp add: field_simps) - have faceS: "S \ {x. a h \ x = b h} face_of S" + let ?F = "{x. a h \ x = b h}" + have faceS: "S \ ?F face_of S" proof (rule face_of_Int_supporting_hyperplane_le) show "\x. x \ S \ a h \ x \ b h" using faceq seq that by fastforce @@ -1960,15 +1945,18 @@ using \a h \ w = b h\ awlt faceq less_eq_real_def by blast ultimately have "w \ S" using seq by blast - with weq have ne: "S \ {x. a h \ x = b h} \ {}" by blast - moreover have "affine hull (S \ {x. a h \ x = b h}) = (affine hull S) \ {x. a h \ x = b h}" + with weq have ne: "S \ ?F \ {}" by blast + moreover have "affine hull (S \ ?F) = (affine hull S) \ ?F" proof - show "affine hull (S \ {x. a h \ x = b h}) \ affine hull S \ {x. a h \ x = b h}" - apply (intro Int_greatest hull_mono Int_lower1) - apply (metis affine_hull_eq affine_hyperplane hull_mono inf_le2) - done + show "affine hull (S \ ?F) \ affine hull S \ ?F" + proof - + have "affine hull (S \ ?F) \ affine hull S" + by (simp add: hull_mono) + then show ?thesis + by (simp add: affine_hyperplane subset_hull) + qed next - show "affine hull S \ {x. a h \ x = b h} \ affine hull (S \ {x. a h \ x = b h})" + show "affine hull S \ ?F \ affine hull (S \ ?F)" proof fix y assume yaff: "y \ affine hull S \ {y. a h \ y = b h}" @@ -2014,7 +2002,7 @@ case False with \0 < inff\ have "inff * (a j \ y - a j \ w) \ 0" by (simp add: mult_le_0_iff) - also have "... < b j - a j \ w" + also have "\ < b j - a j \ w" by (simp add: awlt that) finally show ?thesis by simp qed @@ -2050,7 +2038,7 @@ by (metis yeq affine_affine_hull [simplified affine_alt, rule_format, OF waff caff]) qed qed - ultimately have "aff_dim (affine hull (S \ {x. a h \ x = b h})) = aff_dim S - 1" + ultimately have "aff_dim (affine hull (S \ ?F)) = aff_dim S - 1" using \b h < a h \ z\ zaff by (force simp: aff_dim_affine_Int_hyperplane) then show ?thesis by (simp add: ne faceS facet_of_def) @@ -2755,7 +2743,7 @@ by (meson \0 < a\ divide_less_eq floor_eq_iff) have "?n * a \ a + x" using \a>0\ by (simp add: distrib_right floor_divide_lower) - also have "... < y" + also have "\ < y" by (rule 1) finally have "?n * a < y" . with x show ?thesis @@ -2767,7 +2755,7 @@ by (meson \0 < a\ divide_less_eq floor_eq_iff) have "?n * a \ a + y" using \a>0\ by (simp add: distrib_right floor_divide_lower) - also have "... < x" + also have "\ < x" by (rule 2) finally have "?n * a < x" . then show ?thesis @@ -2891,7 +2879,7 @@ using B \X \ \'\ eq that by blast+ have "norm (x - y) \ (\b\Basis. \(x-y) \ b\)" by (rule norm_le_l1) - also have "... \ of_nat (DIM('a)) * (e / 2 / DIM('a))" + also have "\ \ of_nat (DIM('a)) * (e / 2 / DIM('a))" proof (rule sum_bounded_above) fix i::'a assume "i \ Basis" @@ -2930,12 +2918,12 @@ using I' [OF \n \ C\ refl] n by auto qed qed - also have "... = e / 2" + also have "\ = e / 2" by simp finally show ?thesis . qed qed (use \0 < e\ in force) - also have "... < e" + also have "\ < e" by (simp add: \0 < e\) finally show ?thesis . qed @@ -3240,7 +3228,7 @@ using \K \ \\ C\ by blast have "K \ rel_frontier C" by (simp add: \K \ rel_frontier C\) - also have "... \ C" + also have "\ \ C" by (simp add: \closed C\ rel_frontier_def subset_iff) finally have "K \ C" . have "L \ C face_of C" @@ -3292,7 +3280,7 @@ by (auto simp: \\ affine_dependent I\ aff_independent_finite finite_imp_compact) moreover have "F face_of convex hull insert ?z I" by (metis S \F face_of S\ \K = convex hull I\ convex_hull_eq_empty convex_hull_insert_segments hull_hull) - ultimately obtain J where "J \ insert ?z I" "F = convex hull J" + ultimately obtain J where J: "J \ insert ?z I" "F = convex hull J" using face_of_convex_hull_subset [of "insert ?z I" F] by auto show ?thesis proof (cases "?z \ J") @@ -3320,7 +3308,7 @@ case False then have "F \ \" using face_of_convex_hull_affine_independent [OF \\ affine_dependent I\] - by (metis Int_absorb2 Int_insert_right_if0 \F = convex hull J\ \J \ insert ?z I\ \K = convex hull I\ face\ inf_le2 \K \ \\) + by (metis J \K = convex hull I\ face\ subset_insert \K \ \\) then show "F \ \ \ ?\" by blast qed @@ -3366,7 +3354,7 @@ using ahK_C_disjoint \C \ \\ \K \ \\ \K \ rel_frontier C\ affine_hull_convex_hull z by blast have "X \ K face_of K" by (simp add: XY(1) \K \ \\ faceI\ inf_commute) - also have "... face_of convex hull insert ?z K" + also have "\ face_of convex hull insert ?z K" by (metis I Keq \?z \ affine hull I\ aff_independent_finite convex_convex_hull face_of_convex_hull_insert face_of_refl hull_insert) finally have "X \ K face_of convex hull insert ?z K" . then show ?thesis @@ -3413,7 +3401,7 @@ using \L \ rel_frontier D\ by auto have "convex hull insert (SOME z. z \ rel_interior C) (K \ L) face_of convex hull insert (SOME z. z \ rel_interior C) K" - by (metis face_of_polytope_insert2 "*" IntI \C \ \\ aff_independent_finite ahK_C_disjoint empty_iff faceI\ polytope_def z \K \ \\ \L \ \\\K \ rel_frontier C\) + by (metis IntI \C \ \\ \K \ \\ \K \ rel_frontier C\ \L \ \\ ahK_C_disjoint empty_iff faceI\ face_of_polytope_insert2 simpl\ simplex_imp_polytope z) then show ?thesis using True X Y \K \ rel_frontier C\ \L \ rel_frontier C\ \convex C\ \convex K\ \convex L\ convex_hull_insert_Int_eq z by force next @@ -3442,7 +3430,7 @@ qed finally have CD: "C \ (rel_interior D) = {}" . have zKC: "(convex hull insert ?z K) \ C" - by (metis DiffE \C \ \\ \K \ rel_frontier C\ closed\ closure_closed convex\ hull_minimal insert_subset rel_frontier_def rel_interior_subset subset_iff z) + by (metis \K \ C\ \convex C\ in_mono insert_subsetI rel_interior_subset subset_hull z) have "disjnt (convex hull insert (SOME z. z \ rel_interior C) K) (rel_interior D)" using zKC CD by (force simp: disjnt_def) then have eq: "convex hull (insert ?z K) \ convex hull (insert ?w L) = @@ -3454,7 +3442,7 @@ by (simp add: \C \ \\ convex\) have "convex hull (insert ?z K) \ L = L \ convex hull (insert ?z K)" by blast - also have "... = convex hull K \ L" + also have "\ = convex hull K \ L" proof (subst Int_convex_hull_insert_rel_exterior [OF \convex C\ \K \ C\ z]) have "(C \ D) \ rel_interior C = {}" proof (rule face_of_disjoint_rel_interior) @@ -3482,26 +3470,26 @@ finally have chKL: "convex hull (insert ?z K) \ L = convex hull K \ L" . have "convex hull insert ?z K \ convex hull L face_of K" by (simp add: \K \ \\ \L \ \\ ch_id chKL faceI\) - also have "... face_of convex hull insert ?z K" + also have "\ face_of convex hull insert ?z K" proof - obtain I where I: "\ affine_dependent I" "K = convex hull I" using * [OF \K \ \\] by auto then have "\a. a \ rel_interior C \ a \ affine hull I" using ahK_C_disjoint \C \ \\ \K \ \\ \K \ rel_frontier C\ affine_hull_convex_hull by blast then show ?thesis - by (metis I affine_independent_insert face_of_convex_hull_affine_independent hull_insert subset_insertI z) + by (metis I \convex K\ aff_independent_finite face_of_convex_hull_insert_eq face_of_refl hull_insert z) qed finally have 1: "convex hull insert ?z K \ convex hull L face_of convex hull insert ?z K" . have "convex hull insert ?z K \ convex hull L face_of L" by (metis \K \ \\ \L \ \\ chKL ch_id faceI\ inf_commute) - also have "... face_of convex hull insert ?w L" + also have "\ face_of convex hull insert ?w L" proof - obtain I where I: "\ affine_dependent I" "L = convex hull I" using * [OF \L \ \\] by auto then have "\a. a \ rel_interior D \ a \ affine hull I" using \D \ \\ \L \ \\ \L \ rel_frontier D\ affine_hull_convex_hull ahK_C_disjoint by blast then show ?thesis - by (metis I aff_independent_finite convex_convex_hull face_of_convex_hull_insert face_of_refl hull_insert w) + by (metis I \convex L\ aff_independent_finite face_of_convex_hull_insert face_of_refl hull_insert w) qed finally have 2: "convex hull insert ?z K \ convex hull L face_of convex hull insert ?w L" . show ?thesis @@ -3767,18 +3755,13 @@ if "T \ {U - C |C. C \ \ \ aff_dim C < aff_dim U}" for T using that dense_complement_convex_closed \closed U\ \convex U\ by auto qed - also have "... \ closure ?lhs" + also have "\ \ closure ?lhs" proof - obtain C where "C \ \" "aff_dim C < aff_dim U" by (metis False Sup_upper aff_dim_subset eq eq_iff not_le) - have "\X. X \ \ \ aff_dim X = aff_dim U \ x \ X" + then have "\X. X \ \ \ aff_dim X = aff_dim U \ x \ X" if "\V. (\C. V = U - C \ C \ \ \ aff_dim C < aff_dim U) \ x \ V" for x - proof - - have "x \ U \ x \ \\" - using \C \ \\ \aff_dim C < aff_dim U\ eq that by blast - then show ?thesis - by (metis Diff_iff Sup_upper Union_iff aff_dim_subset dual_order.order_iff_strict eq that) - qed + by (metis Diff_iff Sup_upper UnionE aff_dim_subset eq order_less_le that) then show ?thesis by (auto intro!: closure_mono) qed diff -r ba2afdd29e1d -r 032a4344903e src/HOL/Analysis/Smooth_Paths.thy --- a/src/HOL/Analysis/Smooth_Paths.thy Tue Aug 01 11:27:55 2023 +0200 +++ b/src/HOL/Analysis/Smooth_Paths.thy Thu Aug 03 19:10:43 2023 +0200 @@ -16,11 +16,8 @@ proof - have "path_image \ homeomorphic {0..1::real}" by (simp add: assms homeomorphic_arc_image_interval) - then - show ?thesis - apply (rule path_connected_complement_homeomorphic_convex_compact) - apply (auto simp: assms) - done + then show ?thesis + by (intro path_connected_complement_homeomorphic_convex_compact) (auto simp: assms) qed lemma connected_arc_complement: @@ -39,15 +36,11 @@ using assms connected_arc_image connected_convex_1_gen inside_convex by blast next case False + then have "connected (- path_image \)" + by (metis DIM_ge_Suc0 One_nat_def Suc_1 antisym assms connected_arc_complement not_less_eq_eq) + then show ?thesis - proof (rule inside_bounded_complement_connected_empty) - show "connected (- path_image \)" - apply (rule connected_arc_complement [OF assms]) - using False - by (metis DIM_ge_Suc0 One_nat_def Suc_1 not_less_eq_eq order_class.order.antisym) - show "bounded (path_image \)" - by (simp add: assms bounded_arc_image) - qed + by (simp add: assms bounded_arc_image inside_bounded_complement_connected_empty) qed lemma inside_simple_curve_imp_closed: @@ -59,18 +52,18 @@ subsection \Piecewise differentiability of paths\ lemma continuous_on_joinpaths_D1: - "continuous_on {0..1} (g1 +++ g2) \ continuous_on {0..1} g1" - apply (rule continuous_on_eq [of _ "(g1 +++ g2) \ ((*)(inverse 2))"]) - apply (rule continuous_intros | simp)+ - apply (auto elim!: continuous_on_subset simp: joinpaths_def) - done + assumes "continuous_on {0..1} (g1 +++ g2)" + shows "continuous_on {0..1} g1" +proof (rule continuous_on_eq) + have "continuous_on {0..1/2} (g1 +++ g2)" + using assms continuous_on_subset split_01 by auto + then show "continuous_on {0..1} (g1 +++ g2 \ (*) (inverse 2))" + by (intro continuous_intros) force +qed (auto simp: joinpaths_def) lemma continuous_on_joinpaths_D2: "\continuous_on {0..1} (g1 +++ g2); pathfinish g1 = pathstart g2\ \ continuous_on {0..1} g2" - apply (rule continuous_on_eq [of _ "(g1 +++ g2) \ (\x. inverse 2*x + 1/2)"]) - apply (rule continuous_intros | simp)+ - apply (auto elim!: continuous_on_subset simp add: joinpaths_def pathfinish_def pathstart_def Ball_def) - done + using path_def path_join by blast lemma piecewise_differentiable_D1: assumes "(g1 +++ g2) piecewise_differentiable_on {0..1}" @@ -145,8 +138,8 @@ proof (rule differentiable_transform_within) show "g1 +++ g2 \ (*) (inverse 2) differentiable at x" using that g12D - apply (simp only: joinpaths_def) - by (rule differentiable_chain_at derivative_intros | force)+ + unfolding joinpaths_def + by (intro differentiable_chain_at derivative_intros | force)+ show "\x'. \dist x' x < dist (x/2) (1/2)\ \ (g1 +++ g2 \ (*) (inverse 2)) x' = g1 x'" using that by (auto simp: dist_real_def joinpaths_def) @@ -173,10 +166,9 @@ qed have "continuous_on ({0..1} - insert 1 ((*) 2 ` S)) ((\x. 1/2 * vector_derivative (g1 \ (*)2) (at x)) \ (*)(1/2))" - apply (rule continuous_intros)+ using coDhalf - apply (simp add: scaleR_conv_of_real image_set_diff image_image) - done + apply (intro continuous_intros) + by (simp add: scaleR_conv_of_real image_set_diff image_image) then have con_g1: "continuous_on ({0..1} - insert 1 ((*) 2 ` S)) (\x. vector_derivative g1 (at x))" by (rule continuous_on_eq) (simp add: scaleR_conv_of_real) have "continuous_on {0..1} g1" @@ -201,7 +193,7 @@ proof (rule differentiable_transform_within) show "g1 +++ g2 \ (\x. (x + 1) / 2) differentiable at x" using g12D that - apply (simp only: joinpaths_def) + unfolding joinpaths_def apply (drule_tac x= "(x+1) / 2" in bspec, force simp: field_split_simps) apply (rule differentiable_chain_at derivative_intros | force)+ done @@ -238,10 +230,7 @@ have "continuous_on {0..1} g2" using continuous_on_joinpaths_D2 assms piecewise_C1_differentiable_on_def by blast with \finite S\ show ?thesis - apply (clarsimp simp add: piecewise_C1_differentiable_on_def C1_differentiable_on_eq) - apply (rule_tac x="insert 0 ((\x. 2 * x - 1) ` S)" in exI) - apply (simp add: g2D con_g2) - done + by (meson C1_differentiable_on_eq con_g2 finite_imageI finite_insert g2D piecewise_C1_differentiable_on_def) qed @@ -305,12 +294,8 @@ next show "vector_derivative g (at t) * deriv f (g t) = vector_derivative (f \ g) (at t)" when "t \ {0..1} - S" for t - proof (rule vector_derivative_chain_at_general[symmetric]) - show "g differentiable at t" by (meson C1_differentiable_on_eq g_diff that) - next - have "g t\path_image g" using that DiffD1 image_eqI path_image_def by metis - then show "f field_differentiable at (g t)" using der by auto - qed + by (metis C1_differentiable_on_eq DiffD1 der g_diff imageI path_image_def that + vector_derivative_chain_at_general) qed ultimately have "f \ g C1_differentiable_on {0..1} - S" using C1_differentiable_on_eq by blast @@ -379,10 +364,8 @@ by (auto intro!: continuous_intros finite_vimageI [where h = "(\x. 2*x - 1)"] inj_onI simp: image_affinity_atLeastAtMost_diff continuous_on_joinpaths) ultimately show ?thesis - apply (simp only: valid_path_def continuous_on_joinpaths joinpaths_def) - apply (rule piecewise_C1_differentiable_cases) - apply (auto simp: o_def) - done + unfolding valid_path_def continuous_on_joinpaths joinpaths_def + by (intro piecewise_C1_differentiable_cases) (auto simp: o_def) qed lemma valid_path_join_D1: @@ -406,31 +389,24 @@ assumes "valid_path g" "pathfinish g = pathstart g" "a \ {0..1}" shows "valid_path(shiftpath a g)" using assms - apply (auto simp: valid_path_def shiftpath_alt_def) - apply (rule piecewise_C1_differentiable_cases) - apply (auto simp: algebra_simps) - apply (rule piecewise_C1_differentiable_affine [of g 1 a, simplified o_def scaleR_one]) - apply (auto simp: pathfinish_def pathstart_def elim: piecewise_C1_differentiable_on_subset) - apply (rule piecewise_C1_differentiable_affine [of g 1 "a-1", simplified o_def scaleR_one algebra_simps]) - apply (auto simp: pathfinish_def pathstart_def elim: piecewise_C1_differentiable_on_subset) + unfolding valid_path_def shiftpath_alt_def + apply (intro piecewise_C1_differentiable_cases) + apply (simp_all add: add.commute) + apply (rule piecewise_C1_differentiable_affine [of g 1 a, simplified o_def scaleR_one]) + apply (force simp: pathfinish_def pathstart_def elim: piecewise_C1_differentiable_on_subset) + apply (rule piecewise_C1_differentiable_affine [of g 1 "a-1", simplified o_def scaleR_one algebra_simps]) + apply (auto simp: pathfinish_def pathstart_def elim: piecewise_C1_differentiable_on_subset) done lemma vector_derivative_linepath_within: "x \ {0..1} \ vector_derivative (linepath a b) (at x within {0..1}) = b - a" - apply (rule vector_derivative_within_cbox [of 0 "1::real", simplified]) - apply (auto simp: has_vector_derivative_linepath_within) - done + by (simp add: has_vector_derivative_linepath_within vector_derivative_at_within_ivl) lemma vector_derivative_linepath_at [simp]: "vector_derivative (linepath a b) (at x) = b - a" by (simp add: has_vector_derivative_linepath_within vector_derivative_at) lemma valid_path_linepath [iff]: "valid_path (linepath a b)" - apply (simp add: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq continuous_on_linepath) - apply (rule_tac x="{}" in exI) - apply (simp add: differentiable_on_def differentiable_def) - using has_vector_derivative_def has_vector_derivative_linepath_within - apply (fastforce simp add: continuous_on_eq_continuous_within) - done + using C1_differentiable_on_eq piecewise_C1_differentiable_on_def valid_path_def by fastforce lemma valid_path_subpath: fixes g :: "real \ 'a :: real_normed_vector" @@ -443,15 +419,19 @@ by (force intro: C1_differentiable_on_const C1_differentiable_imp_piecewise) next case False - have "(g \ (\x. ((v-u) * x + u))) piecewise_C1_differentiable_on {0..1}" - apply (rule piecewise_C1_differentiable_compose) - apply (simp add: C1_differentiable_imp_piecewise) - apply (simp add: image_affinity_atLeastAtMost) - using assms False - apply (auto simp: algebra_simps valid_path_def piecewise_C1_differentiable_on_subset) - apply (subst Int_commute) - apply (auto simp: inj_on_def algebra_simps crossproduct_eq finite_vimage_IntI) - done + let ?f = "\x. ((v-u) * x + u)" + have "(g \ ?f) piecewise_C1_differentiable_on {0..1}" + proof (rule piecewise_C1_differentiable_compose) + show "?f piecewise_C1_differentiable_on {0..1}" + by (simp add: C1_differentiable_imp_piecewise) + have "g piecewise_C1_differentiable_on (if u \ v then {u..v} else {v..u})" + using assms piecewise_C1_differentiable_on_subset valid_path_def by force + then show "g piecewise_C1_differentiable_on ?f ` {0..1}" + by (simp add: image_affinity_atLeastAtMost split: if_split_asm) + show "\x. finite ({0..1} \ ?f -` {x})" + using False + by (simp add: Int_commute [of "{0..1}"] inj_on_def crossproduct_eq finite_vimage_IntI) + qed then show ?thesis by (auto simp: o_def valid_path_def subpath_def) qed diff -r ba2afdd29e1d -r 032a4344903e src/HOL/Analysis/Tagged_Division.thy --- a/src/HOL/Analysis/Tagged_Division.thy Tue Aug 01 11:27:55 2023 +0200 +++ b/src/HOL/Analysis/Tagged_Division.thy Thu Aug 03 19:10:43 2023 +0200 @@ -14,99 +14,56 @@ and "\i. i \ S \ finite (T i)" shows "(\i\S. sum (x i) (T i)) = (\(i, j)\Sigma S T. x i j)" using assms -proof induct - case empty - then show ?case - by simp -next - case (insert a S) - show ?case - by (simp add: insert.hyps(1) insert.prems sum.Sigma) -qed + by induction (auto simp: sum.Sigma) lemmas scaleR_simps = scaleR_zero_left scaleR_minus_left scaleR_left_diff_distrib scaleR_zero_right scaleR_minus_right scaleR_right_diff_distrib scaleR_eq_0_iff scaleR_cancel_left scaleR_cancel_right scaleR_add_right scaleR_add_left real_vector_class.scaleR_one -subsection\<^marker>\tag unimportant\ \Sundries\ - - -text\A transitive relation is well-founded if all initial segments are finite.\ -lemma wf_finite_segments: - assumes "irrefl r" and "trans r" and "\x. finite {y. (y, x) \ r}" - shows "wf (r)" - apply (simp add: trans_wf_iff wf_iff_acyclic_if_finite converse_def assms) - using acyclic_def assms irrefl_def trans_Restr by fastforce - -text\For creating values between \<^term>\u\ and \<^term>\v\.\ -lemma scaling_mono: - fixes u::"'a::linordered_field" - assumes "u \ v" "0 \ r" "r \ s" - shows "u + r * (v - u) / s \ v" -proof - - have "r/s \ 1" using assms - using divide_le_eq_1 by fastforce - then have "(r/s) * (v - u) \ 1 * (v - u)" - by (meson diff_ge_0_iff_ge mult_right_mono \u \ v\) - then show ?thesis - by (simp add: field_simps) -qed - subsection \Some useful lemmas about intervals\ lemma interior_subset_union_intervals: - assumes "i = cbox a b" - and "j = cbox c d" - and "interior j \ {}" + fixes a b c d + defines "i \ cbox a b" + defines "j \ cbox c d" + assumes "interior j \ {}" and "i \ j \ S" and "interior i \ interior j = {}" shows "interior i \ interior S" -proof - - have "box a b \ cbox c d = {}" - using Int_interval_mixed_eq_empty[of c d a b] assms - unfolding interior_cbox by auto - moreover - have "box a b \ cbox c d \ S" - apply (rule order_trans,rule box_subset_cbox) - using assms by auto - ultimately - show ?thesis - unfolding assms interior_cbox - by auto (metis IntI UnE empty_iff interior_maximal open_box subsetCE subsetI) -qed + by (smt (verit, del_insts) IntI Int_interval_mixed_eq_empty UnE assms empty_iff interior_cbox interior_maximal interior_subset open_interior subset_eq) lemma interior_Union_subset_cbox: - assumes "finite f" - assumes f: "\s. s \ f \ \a b. s = cbox a b" "\s. s \ f \ interior s \ t" - and t: "closed t" - shows "interior (\f) \ t" + assumes "finite \" + assumes \: "\S. S \ \ \ \a b. S = cbox a b" "\S. S \ \ \ interior S \ T" + and T: "closed T" + shows "interior (\\) \ T" proof - - have [simp]: "s \ f \ closed s" for s - using f by auto - define E where "E = {s\f. interior s = {}}" - then have "finite E" "E \ {s\f. interior s = {}}" - using \finite f\ by auto - then have "interior (\f) = interior (\(f - E))" + have clo[simp]: "S \ \ \ closed S" for S + using \ by auto + define E where "E = {s\\. interior s = {}}" + then have "finite E" "E \ {s\\. interior s = {}}" + using \finite \\ by auto + then have "interior (\\) = interior (\(\ - E))" proof (induction E rule: finite_subset_induct') case (insert s f') - have "interior (\(f - insert s f') \ s) = interior (\(f - insert s f'))" - using insert.hyps \finite f\ by (intro interior_closed_Un_empty_interior) auto - also have "\(f - insert s f') \ s = \(f - f')" + have "interior (\(\ - insert s f') \ s) = interior (\(\ - insert s f'))" + using insert.hyps \finite \\ by (intro interior_closed_Un_empty_interior) auto + also have "\(\ - insert s f') \ s = \(\ - f')" using insert.hyps by auto finally show ?case by (simp add: insert.IH) qed simp - also have "\ \ \(f - E)" + also have "\ \ \(\ - E)" by (rule interior_subset) - also have "\ \ t" + also have "\ \ T" proof (rule Union_least) - fix s assume "s \ f - E" - with f[of s] obtain a b where s: "s \ f" "s = cbox a b" "box a b \ {}" + fix s assume "s \ \ - E" + with \[of s] obtain a b where s: "s \ \" "s = cbox a b" "box a b \ {}" by (fastforce simp: E_def) - have "closure (interior s) \ closure t" - by (intro closure_mono f \s \ f\) - with s \closed t\ show "s \ t" + have "closure (interior s) \ closure T" + by (intro closure_mono \ \s \ \\) + with s \closed T\ show "s \ T" by simp qed finally show ?thesis . @@ -195,14 +152,12 @@ definition\<^marker>\tag important\ "gauge \ \ (\x. x \ \ x \ open (\ x))" lemma gaugeI: - assumes "\x. x \ \ x" - and "\x. open (\ x)" + assumes "\x. x \ \ x" and "\x. open (\ x)" shows "gauge \" using assms unfolding gauge_def by auto lemma gaugeD[dest]: - assumes "gauge \" - shows "x \ \ x" + assumes "gauge \" shows "x \ \ x" and "open (\ x)" using assms unfolding gauge_def by auto @@ -213,7 +168,7 @@ unfolding gauge_def by auto lemma gauge_trivial[intro!]: "gauge (\x. ball x 1)" - by (rule gauge_ball) auto + by auto lemma gauge_Int[intro]: "gauge \1 \ gauge \2 \ gauge (\x. \1 x \ \2 x)" unfolding gauge_def by auto @@ -221,8 +176,7 @@ lemma gauge_reflect: fixes \ :: "'a::euclidean_space \ 'a set" shows "gauge \ \ gauge (\x. uminus ` \ (- x))" - using equation_minus_iff - by (auto simp: gauge_def surj_def intro!: open_surjective_linear_image linear_uminus) + by (metis (mono_tags, lifting) gauge_def imageI open_negations minus_minus) lemma gauge_Inter: assumes "finite S" @@ -233,7 +187,7 @@ by auto show ?thesis unfolding gauge_def unfolding * - using assms unfolding Ball_def Inter_iff mem_Collect_eq gauge_def by auto + by (simp add: assms gaugeD open_INT) qed lemma gauge_existence_lemma: @@ -290,18 +244,6 @@ "s division_of cbox a (a::'a::euclidean_space) \ s = {cbox a a}" (is "?l = ?r") proof - assume ?r - moreover - { fix K - assume "s = {{a}}" "K\s" - then have "\x y. K = cbox x y" - apply (rule_tac x=a in exI)+ - apply force - done - } - ultimately show ?l - unfolding division_of_def cbox_idem by auto -next assume ?l have "x = {a}" if "x \ s" for x by (metis \s division_of cbox a a\ cbox_idem division_ofD(2) division_ofD(3) subset_singletonD that) @@ -309,10 +251,10 @@ using \s division_of cbox a a\ by auto ultimately show ?r unfolding cbox_idem by auto -qed +qed (use cbox_idem in blast) lemma elementary_empty: obtains p where "p division_of {}" - unfolding division_of_trivial by auto + by simp lemma elementary_interval: obtains p where "p division_of (cbox a b)" by (metis division_of_trivial division_of_self) @@ -334,31 +276,24 @@ and "q \ p" shows "q division_of (\q)" proof (rule division_ofI) - note * = division_ofD[OF assms(1)] show "finite q" - using "*"(1) assms(2) infinite_super by auto - { - fix k - assume "k \ q" - then have kp: "k \ p" - using assms(2) by auto - show "k \ \q" - using \k \ q\ by auto - show "\a b. k = cbox a b" - using *(4)[OF kp] by auto - show "k \ {}" - using *(3)[OF kp] by auto - } + using assms finite_subset by blast +next + fix k + assume "k \ q" + show "k \ \q" + using \k \ q\ by auto + show "\a b. k = cbox a b" "k \ {}" + using assms \k \ q\ by blast+ +next fix k1 k2 assume "k1 \ q" "k2 \ q" "k1 \ k2" - then have **: "k1 \ p" "k2 \ p" "k1 \ k2" - using assms(2) by auto - show "interior k1 \ interior k2 = {}" - using *(5)[OF **] by auto + then show "interior k1 \ interior k2 = {}" + using assms by blast qed auto lemma division_of_union_self[intro]: "p division_of s \ p division_of (\p)" - unfolding division_of_def by auto + by blast lemma division_inter: fixes s1 s2 :: "'a::euclidean_space set" @@ -379,25 +314,20 @@ ultimately show "finite ?A" by auto have *: "\s. \{x\s. x \ {}} = \s" by auto - show "\?A = s1 \ s2" + show UA: "\?A = s1 \ s2" unfolding * using division_ofD(6)[OF assms(1)] and division_ofD(6)[OF assms(2)] by auto { fix k - assume "k \ ?A" + assume kA: "k \ ?A" then obtain k1 k2 where k: "k = k1 \ k2" "k1 \ p1" "k2 \ p2" "k \ {}" by auto then show "k \ {}" by auto show "k \ s1 \ s2" - using division_ofD(2)[OF assms(1) k(2)] and division_ofD(2)[OF assms(2) k(3)] - unfolding k by auto - obtain a1 b1 where k1: "k1 = cbox a1 b1" - using division_ofD(4)[OF assms(1) k(2)] by blast - obtain a2 b2 where k2: "k2 = cbox a2 b2" - using division_ofD(4)[OF assms(2) k(3)] by blast + using UA kA by blast show "\a b. k = cbox a b" - unfolding k k1 k2 unfolding Int_interval by auto + using k by (metis (no_types, lifting) Int_interval assms division_ofD(4)) } fix k1 k2 assume "k1 \ ?A" @@ -407,17 +337,9 @@ then obtain x2 y2 where k2: "k2 = x2 \ y2" "x2 \ p1" "y2 \ p2" "k2 \ {}" by auto assume "k1 \ k2" - then have th: "x1 \ x2 \ y1 \ y2" - unfolding k1 k2 by auto - have *: "interior x1 \ interior x2 = {} \ interior y1 \ interior y2 = {} \ - interior (x1 \ y1) \ interior x1 \ interior (x1 \ y1) \ interior y1 \ - interior (x2 \ y2) \ interior x2 \ interior (x2 \ y2) \ interior y2 \ - interior (x1 \ y1) \ interior (x2 \ y2) = {}" by auto - show "interior k1 \ interior k2 = {}" + then show "interior k1 \ interior k2 = {}" unfolding k1 k2 - apply (rule *) - using assms division_ofD(5) k1 k2(2) k2(3) th apply auto - done + using assms division_ofD(5) k1 k2 by auto qed qed @@ -439,33 +361,20 @@ lemma elementary_Int: fixes s t :: "'a::euclidean_space set" - assumes "p1 division_of s" - and "p2 division_of t" + assumes "p1 division_of s" and "p2 division_of t" shows "\p. p division_of (s \ t)" using assms division_inter by blast lemma elementary_Inter: - assumes "finite f" - and "f \ {}" - and "\s\f. \p. p division_of (s::('a::euclidean_space) set)" - shows "\p. p division_of (\f)" + assumes "finite \" + and "\ \ {}" + and "\s\\. \p. p division_of (s::('a::euclidean_space) set)" + shows "\p. p division_of (\\)" using assms -proof (induct f rule: finite_induct) - case (insert x f) - show ?case - proof (cases "f = {}") - case True - then show ?thesis - unfolding True using insert by auto - next - case False - obtain p where "p division_of \f" - using insert(3)[OF False insert(5)[unfolded ball_simps,THEN conjunct2]] .. - moreover obtain px where "px division_of x" - using insert(5)[rule_format,OF insertI1] .. - ultimately show ?thesis - by (simp add: elementary_Int Inter_insert) - qed +proof (induct \ rule: finite_induct) + case (insert x \) + then show ?case + by (metis cInf_singleton complete_lattice_class.Inf_insert elementary_Int insert_iff) qed auto lemma division_disjoint_union: @@ -476,40 +385,11 @@ proof (rule division_ofI) note d1 = division_ofD[OF assms(1)] note d2 = division_ofD[OF assms(2)] - show "finite (p1 \ p2)" - using d1(1) d2(1) by auto - show "\(p1 \ p2) = s1 \ s2" - using d1(6) d2(6) by auto - { - fix k1 k2 - assume as: "k1 \ p1 \ p2" "k2 \ p1 \ p2" "k1 \ k2" - moreover - let ?g="interior k1 \ interior k2 = {}" - { - assume as: "k1\p1" "k2\p2" - have ?g - using interior_mono[OF d1(2)[OF as(1)]] interior_mono[OF d2(2)[OF as(2)]] - using assms(3) by blast - } - moreover - { - assume as: "k1\p2" "k2\p1" - have ?g - using interior_mono[OF d1(2)[OF as(2)]] interior_mono[OF d2(2)[OF as(1)]] - using assms(3) by blast - } - ultimately show ?g - using d1(5)[OF _ _ as(3)] and d2(5)[OF _ _ as(3)] by auto - } - fix k - assume k: "k \ p1 \ p2" - show "k \ s1 \ s2" - using k d1(2) d2(2) by auto - show "k \ {}" - using k d1(3) d2(3) by auto - show "\a b. k = cbox a b" - using k d1(4) d2(4) by auto -qed + fix k1 k2 + assume "k1 \ p1 \ p2" "k2 \ p1 \ p2" "k1 \ k2" + with assms show "interior k1 \ interior k2 = {}" + by (smt (verit, best) IntE Un_iff disjoint_iff_not_equal division_ofD(2) division_ofD(5) inf.orderE interior_Int) +qed (use division_ofD assms in auto) lemma partial_division_extend_1: fixes a b c d :: "'a::euclidean_space" @@ -533,71 +413,59 @@ show "finite p" unfolding p_def by (auto intro!: finite_PiE) { - fix k - assume "k \ p" - then obtain f where f: "f \ Basis \\<^sub>E {(a, c), (c, d), (d, b)}" and k: "k = ?B f" + fix K + assume "K \ p" + then obtain f where f: "f \ Basis \\<^sub>E {(a, c), (c, d), (d, b)}" and k: "K = ?B f" by (auto simp: p_def) - then show "\a b. k = cbox a b" - by auto - have "k \ cbox a b \ k \ {}" - proof (simp add: k box_eq_empty subset_box not_less, safe) - fix i :: 'a - assume i: "i \ Basis" - with f have "f i = (a, c) \ f i = (c, d) \ f i = (d, b)" - by (auto simp: PiE_iff) - with i ord[of i] - show "a \ i \ fst (f i) \ i" "snd (f i) \ i \ b \ i" "fst (f i) \ i \ snd (f i) \ i" - by auto - qed - then show "k \ {}" "k \ cbox a b" + then show "\a b. K = cbox a b" by auto { fix l assume "l \ p" then obtain g where g: "g \ Basis \\<^sub>E {(a, c), (c, d), (d, b)}" and l: "l = ?B g" by (auto simp: p_def) - assume "l \ k" + assume "l \ K" have "\i\Basis. f i \ g i" - proof (rule ccontr) - assume "\ ?thesis" - with f g have "f = g" - by (auto simp: PiE_iff extensional_def fun_eq_iff) - with \l \ k\ show False - by (simp add: l k) - qed + using \l \ K\ l k f g by auto then obtain i where *: "i \ Basis" "f i \ g i" .. then have "f i = (a, c) \ f i = (c, d) \ f i = (d, b)" - "g i = (a, c) \ g i = (c, d) \ g i = (d, b)" + "g i = (a, c) \ g i = (c, d) \ g i = (d, b)" using f g by (auto simp: PiE_iff) - with * ord[of i] show "interior l \ interior k = {}" + with * ord[of i] show "interior l \ interior K = {}" by (auto simp add: l k disjoint_interval intro!: bexI[of _ i]) } - note \k \ cbox a b\ + have "a \ i \ fst (f i) \ i" "snd (f i) \ i \ b \ i" "fst (f i) \ i \ snd (f i) \ i" + if "i \ Basis" for i + proof - + have "f i = (a, c) \ f i = (c, d) \ f i = (d, b)" + using that f by (auto simp: PiE_iff) + with that ord[of i] + show "a \ i \ fst (f i) \ i" "snd (f i) \ i \ b \ i" "fst (f i) \ i \ snd (f i) \ i" + by auto + qed + then show "K \ {}" "K \ cbox a b" + by (auto simp add: k box_eq_empty subset_box not_less) } moreover - { - fix x assume x: "x \ cbox a b" - have "\i\Basis. \l. x \ i \ {fst l \ i .. snd l \ i} \ l \ {(a, c), (c, d), (d, b)}" - proof - fix i :: 'a - assume "i \ Basis" - with x ord[of i] + have "\k\p. x \ k" if x: "x \ cbox a b" for x + proof - + have "\l. x \ i \ {fst l \ i .. snd l \ i} \ l \ {(a, c), (c, d), (d, b)}" if "i \ Basis" for i + proof - have "(a \ i \ x \ i \ x \ i \ c \ i) \ (c \ i \ x \ i \ x \ i \ d \ i) \ (d \ i \ x \ i \ x \ i \ b \ i)" + using that x ord[of i] by (auto simp: cbox_def) then show "\l. x \ i \ {fst l \ i .. snd l \ i} \ l \ {(a, c), (c, d), (d, b)}" by auto qed then obtain f where f: "\i\Basis. x \ i \ {fst (f i) \ i..snd (f i) \ i} \ f i \ {(a, c), (c, d), (d, b)}" - unfolding bchoice_iff .. - moreover from f have "restrict f Basis \ Basis \\<^sub>E {(a, c), (c, d), (d, b)}" - by auto - moreover from f have "x \ ?B (restrict f Basis)" + by metis + moreover from f have "x \ ?B (restrict f Basis)" "restrict f Basis \ Basis \\<^sub>E {(a,c), (c,d), (d,b)}" by (auto simp: mem_box) - ultimately have "\k\p. x \ k" + ultimately show ?thesis unfolding p_def by blast - } + qed ultimately show "\p = cbox a b" by auto qed @@ -608,14 +476,12 @@ obtains q where "p \ q" "q division_of cbox a (b::'a::euclidean_space)" proof (cases "p = {}") case True - obtain q where "q division_of (cbox a b)" - by (rule elementary_interval) then show ?thesis - using True that by blast + using elementary_interval that by auto next case False note p = division_ofD[OF assms(1)] - have div_cbox: "\k\p. \q. q division_of cbox a b \ k \ q" + have "\k\p. \q. q division_of cbox a b \ k \ q" proof fix k assume kp: "k \ p" @@ -629,64 +495,55 @@ then show "\q. q division_of cbox a b \ k \ q" unfolding k by auto qed - obtain q where q: "\x. x \ p \ q x division_of cbox a b" "\x. x \ p \ x \ q x" - using bchoice[OF div_cbox] by blast + then obtain q where q: "\x. x \ p \ q x division_of cbox a b" "\x. x \ p \ x \ q x" + by metis have "q x division_of \(q x)" if x: "x \ p" for x - apply (rule division_ofI) - using division_ofD[OF q(1)[OF x]] by auto + using q(1) x by blast then have di: "\x. x \ p \ \d. d division_of \(q x - {x})" by (meson Diff_subset division_of_subset) - have "\d. d division_of \((\i. \(q i - {i})) ` p)" - apply (rule elementary_Inter [OF finite_imageI[OF p(1)]]) - apply (auto dest: di simp: False elementary_Inter [OF finite_imageI[OF p(1)]]) - done - then obtain d where d: "d division_of \((\i. \(q i - {i})) ` p)" .. + have "\s\(\i. \ (q i - {i})) ` p. \d. d division_of s" + using di by blast + then obtain d where d: "d division_of \((\i. \(q i - {i})) ` p)" + by (meson False elementary_Inter finite_imageI image_is_empty p(1)) have "d \ p division_of cbox a b" proof - have te: "\S f T. S \ {} \ \i\S. f i \ i = T \ T = \(f ` S) \ \S" by auto have cbox_eq: "cbox a b = \((\i. \(q i - {i})) ` p) \ \p" proof (rule te[OF False], clarify) fix i - assume i: "i \ p" - show "\(q i - {i}) \ i = cbox a b" - using division_ofD(6)[OF q(1)[OF i]] using q(2)[OF i] by auto + assume "i \ p" + then show "\(q i - {i}) \ i = cbox a b" + by (metis Un_commute complete_lattice_class.Sup_insert division_ofD(6) insert_Diff q) qed - { fix K - assume K: "K \ p" - note qk=division_ofD[OF q(1)[OF K]] - have *: "\u T S. T \ S = {} \ u \ S \ u \ T = {}" + have [simp]: "interior (\i\p. \(q i - {i})) \ interior K = {}" if K: "K \ p" for K + proof - + note qk = division_ofD[OF q(1)[OF K]] + have *: "\U T S. T \ S = {} \ U \ S \ U \ T = {}" by auto - have "interior (\i\p. \(q i - {i})) \ interior K = {}" + show ?thesis proof (rule *[OF Int_interior_Union_intervals]) show "\T. T\q K - {K} \ interior K \ interior T = {}" - using qk(5) using q(2)[OF K] by auto + using K q(2) qk(5) by auto show "interior (\i\p. \(q i - {i})) \ interior (\(q K - {K}))" - apply (rule interior_mono)+ - using K by auto - qed (use qk in auto)} note [simp] = this + by (meson INF_lower K interior_mono) + qed (use qk in auto) + qed show "d \ p division_of (cbox a b)" - unfolding cbox_eq - apply (rule division_disjoint_union[OF d assms(1)]) - apply (rule Int_interior_Union_intervals) - apply (rule p open_interior ballI)+ - apply simp_all - done + by (simp add: Int_interior_Union_intervals assms(1) cbox_eq d division_disjoint_union p(1) p(4)) qed then show ?thesis by (meson Un_upper2 that) qed lemma elementary_bounded[dest]: - fixes S :: "'a::euclidean_space set" shows "p division_of S \ bounded S" unfolding division_of_def by (metis bounded_Union bounded_cbox) lemma elementary_subset_cbox: - "p division_of S \ \a b. S \ cbox a (b::'a::euclidean_space)" + "p division_of S \ \a b. S \ cbox a b" by (meson bounded_subset_cbox_symmetric elementary_bounded) proposition division_union_intervals_exists: - fixes a b :: "'a::euclidean_space" assumes "cbox a b \ {}" obtains p where "(insert (cbox a b) p) division_of (cbox a b \ cbox c d)" proof (cases "cbox c d = {}") @@ -707,26 +564,19 @@ obtain p where pd: "p division_of cbox c d" and "cbox u v \ p" by (rule partial_division_extend_1[OF uv_sub False[unfolded uv]]) note p = this division_ofD[OF pd] - have "interior (cbox a b \ \(p - {cbox u v})) = interior(cbox u v \ \(p - {cbox u v}))" - apply (rule arg_cong[of _ _ interior]) - using p(8) uv by auto + have "interior (cbox a b \ \(p - {cbox u v})) = interior(cbox u v) \ interior (\(p - {cbox u v}))" + by (metis Diff_subset Int_absorb1 Int_assoc Sup_subset_mono interior_Int p(8) uv) also have "\ = {}" - unfolding interior_Int - apply (rule Int_interior_Union_intervals) using p(6) p(7)[OF p(2)] \finite p\ - apply auto - done - finally have [simp]: "interior (cbox a b) \ interior (\(p - {cbox u v})) = {}" by simp + by (intro Int_interior_Union_intervals) auto + finally have disj: "interior (cbox a b) \ interior (\(p - {cbox u v})) = {}" + by simp have cbe: "cbox a b \ cbox c d = cbox a b \ \(p - {cbox u v})" using p(8) unfolding uv[symmetric] by auto have "insert (cbox a b) (p - {cbox u v}) division_of cbox a b \ \(p - {cbox u v})" - proof - - have "{cbox a b} division_of cbox a b" - by (simp add: assms division_of_self) - then show "insert (cbox a b) (p - {cbox u v}) division_of cbox a b \ \(p - {cbox u v})" - by (metis (no_types) Diff_subset \interior (cbox a b) \ interior (\(p - {cbox u v})) = {}\ division_disjoint_union division_of_subset insert_is_Un p(1) p(8)) - qed - with that[of "p - {cbox u v}"] show ?thesis by (simp add: cbe) + by (metis Diff_subset assms disj division_disjoint_union division_of_self division_of_subset insert_is_Un p(8) pd) + with that[of "p - {cbox u v}"] show ?thesis + by (simp add: cbe) qed qed @@ -735,7 +585,7 @@ and "\p. p \ f \ p division_of (\p)" and "\k1 k2. k1 \ \f \ k2 \ \f \ k1 \ k2 \ interior k1 \ interior k2 = {}" shows "\f division_of \(\f)" - using assms by (auto intro!: division_ofI) + using assms by (auto intro!: division_ofI) lemma elementary_union_interval: fixes a b :: "'a::euclidean_space" @@ -756,21 +606,13 @@ proof (cases "interior (cbox a b) = {}") case True show ?thesis - apply (rule that [of "insert (cbox a b) p", OF division_ofI]) - using pdiv(1-4) True False apply auto - apply (metis IntI equals0D pdiv(5)) - by (metis disjoint_iff_not_equal pdiv(5)) + by (metis True assms division_disjoint_union elementary_interval inf_bot_left that) next - case False + case nonempty: False have "\K\p. \q. (insert (cbox a b) q) division_of (cbox a b \ K)" - proof - fix K - assume kp: "K \ p" - from pdiv(4)[OF kp] obtain c d where "K = cbox c d" by blast - then show "\q. (insert (cbox a b) q) division_of (cbox a b \ K)" - by (meson \cbox a b \ {}\ division_union_intervals_exists) - qed - from bchoice[OF this] obtain q where "\x\p. insert (cbox a b) (q x) division_of (cbox a b) \ x" .. + by (metis \cbox a b \ {}\ division_union_intervals_exists pdiv(4)) + then obtain q where "\x\p. insert (cbox a b) (q x) division_of (cbox a b) \ x" + by metis note q = division_ofD[OF this[rule_format]] let ?D = "\{insert (cbox a b) (q K) | K. K \ p}" show thesis @@ -814,11 +656,11 @@ by (metis \interior (cbox a b) \ {}\ K q(2) x interior_subset_union_intervals) moreover obtain c d where c_d: "K' = cbox c d" - using q(4)[OF x'(2,1)] by blast + using q(4) x'(1) x'(2) by presburger have "interior K' \ interior (cbox a b) = {}" using as'(2) q(5) x' by blast then have "interior K' \ interior x'" - by (metis \interior (cbox a b) \ {}\ c_d interior_subset_union_intervals q(2) x'(1) x'(2)) + by (metis nonempty c_d interior_subset_union_intervals q(2) x') moreover have "interior x \ interior x' = {}" by (meson False assms division_ofD(5) x'(2) x(2)) ultimately show ?thesis @@ -831,39 +673,34 @@ lemma elementary_unions_intervals: - assumes fin: "finite f" - and "\s. s \ f \ \a b. s = cbox a (b::'a::euclidean_space)" - obtains p where "p division_of (\f)" + assumes fin: "finite \" + and "\s. s \ \ \ \a b. s = cbox a (b::'a::euclidean_space)" + obtains p where "p division_of (\\)" proof - - have "\p. p division_of (\f)" - proof (induct_tac f rule:finite_subset_induct) + have "\p. p division_of (\\)" + proof (induct_tac \ rule:finite_subset_induct) show "\p. p division_of \{}" using elementary_empty by auto next fix x F - assume as: "finite F" "x \ F" "\p. p division_of \F" "x\f" - from this(3) obtain p where p: "p division_of \F" .. - from assms(2)[OF as(4)] obtain a b where x: "x = cbox a b" by blast - have *: "\F = \p" - using division_ofD[OF p] by auto - show "\p. p division_of \(insert x F)" - using elementary_union_interval[OF p[unfolded *], of a b] - unfolding Union_insert x * by metis - qed (insert assms, auto) + assume as: "finite F" "x \ F" "\p. p division_of \F" "x\\" + then obtain a b where x: "x = cbox a b" + by (meson assms(2)) + then show "\p. p division_of \(insert x F)" + using elementary_union_interval by (metis Union_insert as(3) division_ofD(6)) + qed (use assms in auto) then show ?thesis using that by auto qed lemma elementary_union: - fixes S T :: "'a::euclidean_space set" assumes "ps division_of S" "pt division_of T" obtains p where "p division_of (S \ T)" proof - - have *: "S \ T = \ps \ \pt" + have "S \ T = \ps \ \pt" using assms unfolding division_of_def by auto + with elementary_unions_intervals[of "ps \ pt"] assms show ?thesis - apply (rule elementary_unions_intervals[of "ps \ pt"]) - using assms apply auto - by (simp add: * that) + by (metis Un_iff Union_Un_distrib division_of_def finite_Un that) qed lemma partial_division_extend: @@ -881,10 +718,7 @@ by (metis ab dual_order.trans partial_division_extend_interval divp(6)) note r1 = this division_ofD[OF this(2)] obtain p' where "p' division_of \(r1 - p)" - apply (rule elementary_unions_intervals[of "r1 - p"]) - using r1(3,6) - apply auto - done + by (metis Diff_subset division_of_subset r1(2) r1(8)) then obtain r2 where r2: "r2 division_of (\(r1 - p)) \ (\q)" by (metis assms(2) divq(6) elementary_Int) { @@ -894,11 +728,7 @@ unfolding r1 using ab by (meson division_contains r1(2) subsetCE) moreover have "R \ p" - proof - assume "R \ p" - then have "x \ S" using divp(2) r by auto - then show False using x by auto - qed + using divp(6) r(2) x(2) by blast ultimately have "x\\(r1 - p)" by auto } then have Teq: "T = \p \ (\(r1 - p) \ \q)" @@ -923,14 +753,11 @@ then have div: "p \ r2 division_of \p \ \(r1 - p) \ \q" by (simp add: assms(1) division_disjoint_union divp(6) r2) show ?thesis - apply (rule that[of "p \ r2"]) - apply (auto simp: div Teq) - done + by (metis Teq div sup_ge1 that) qed lemma division_split: - fixes a :: "'a::euclidean_space" assumes "p division_of (cbox a b)" and k: "k\Basis" shows "{l \ {x. x\k \ c} | l. l \ p \ l \ {x. x\k \ c} \ {}} division_of(cbox a b \ {x. x\k \ c})" @@ -954,11 +781,10 @@ using l p(2) uv by force show "K \ {}" by (simp add: l) - show "\a b. K = cbox a b" - apply (simp add: l uv p(2-3)[OF l(2)]) - apply (subst interval_split[OF k]) - apply (auto intro: order.trans) - done + have "\a b. cbox u v \ {x. x \ k \ c} = cbox a b" + unfolding interval_split[OF k] by (auto intro: order.trans) + then show "\a b. K = cbox a b" + by (simp add: l(1) uv) fix K' assume "K' \ ?p1" then obtain l' where l': "K' = l' \ {x. x \ k \ c}" "l' \ p" "l' \ {x. x \ k \ c} \ {}" @@ -978,11 +804,10 @@ using l p(2) uv by force show "K \ {}" by (simp add: l) - show "\a b. K = cbox a b" - apply (simp add: l uv p(2-3)[OF l(2)]) - apply (subst interval_split[OF k]) - apply (auto intro: order.trans) - done + have "\a b. cbox u v \ {x. c \ x \ k} = cbox a b" + unfolding interval_split[OF k] by (auto intro: order.trans) + then show "\a b. K = cbox a b" + by (simp add: l(1) uv) fix K' assume "K' \ ?p2" then obtain l' where l': "K' = l' \ {x. c \ x \ k}" "l' \ p" "l' \ {x. c \ x \ k} \ {}" @@ -1037,9 +862,9 @@ and "(\{K. \x. (x,K) \ s} = i)" shows "s tagged_division_of i" unfolding tagged_division_of - using assms by fastforce + using assms by fastforce -lemma tagged_division_ofD[dest]: (*FIXME USE A LOCALE*) +lemma tagged_division_ofD[dest]: assumes "s tagged_division_of i" shows "finite s" and "\x K. (x,K) \ s \ x \ K" @@ -1057,18 +882,14 @@ note assm = tagged_division_ofD[OF assms] show "\(snd ` s) = i" "finite (snd ` s)" using assm by auto - fix k - assume k: "k \ snd ` s" - then obtain xk where xk: "(xk, k) \ s" - by auto - then show "k \ i" "k \ {}" "\a b. k = cbox a b" + fix K + assume k: "K \ snd ` s" + then show "K \ i" "K \ {}" "\a b. K = cbox a b" using assm by fastforce+ - fix k' - assume k': "k' \ snd ` s" "k \ k'" - from this(1) obtain xk' where xk': "(xk', k') \ s" - by auto - then show "interior k \ interior k' = {}" - using assm(5) k'(2) xk by blast + fix K' + assume k': "K' \ snd ` s" "K \ K'" + then show "interior K \ interior K' = {}" + by (metis (no_types, lifting) assms imageE k prod.collapse tagged_division_ofD(5)) qed lemma partial_division_of_tagged_division: @@ -1078,23 +899,18 @@ note assm = tagged_partial_division_ofD[OF assms] show "finite (snd ` s)" "\(snd ` s) = \(snd ` s)" using assm by auto - fix k - assume k: "k \ snd ` s" - then obtain xk where xk: "(xk, k) \ s" - by auto - then show "k \ {}" "\a b. k = cbox a b" "k \ \(snd ` s)" - using assm by auto - fix k' - assume k': "k' \ snd ` s" "k \ k'" - from this(1) obtain xk' where xk': "(xk', k') \ s" - by auto - then show "interior k \ interior k' = {}" - using assm(5) k'(2) xk by auto + fix K + assume K: "K \ snd ` s" + then show "K \ {}" "\a b. K = cbox a b" "K \ \(snd ` s)" + using assm by fastforce+ + fix K' + assume "K' \ snd ` s" "K \ K'" + then show "interior K \ interior K' = {}" + using assm(5) K by force qed lemma tagged_partial_division_subset: - assumes "s tagged_partial_division_of i" - and "t \ s" + assumes "s tagged_partial_division_of i" and "t \ s" shows "t tagged_partial_division_of i" using assms finite_subset[OF assms(2)] unfolding tagged_partial_division_of_def @@ -1116,8 +932,7 @@ by (rule tagged_division_ofI) auto lemma tagged_division_of_self_real: "x \ {a .. b::real} \ {(x,{a .. b})} tagged_division_of {a .. b}" - unfolding box_real[symmetric] - by (rule tagged_division_of_self) + by (metis box_real(2) tagged_division_of_self) lemma tagged_division_Un: assumes "p1 tagged_division_of s1" @@ -1131,20 +946,16 @@ using p1(1) p2(1) by auto show "\{k. \x. (x, k) \ p1 \ p2} = s1 \ s2" using p1(6) p2(6) by blast - fix x k - assume xk: "(x, k) \ p1 \ p2" - show "x \ k" "\a b. k = cbox a b" - using xk p1(2,4) p2(2,4) by auto - show "k \ s1 \ s2" - using xk p1(3) p2(3) by blast - fix x' k' - assume xk': "(x', k') \ p1 \ p2" "(x, k) \ (x', k')" + fix x K + assume xK: "(x, K) \ p1 \ p2" + show "x \ K" "\a b. K = cbox a b" "K \ s1 \ s2" + using xK p1 p2 by auto + fix x' K' + assume xk': "(x', K') \ p1 \ p2" "(x, K) \ (x', K')" have *: "\a b. a \ s1 \ b \ s2 \ interior a \ interior b = {}" using assms(3) interior_mono by blast - show "interior k \ interior k' = {}" - apply (cases "(x, k) \ p1") - apply (meson "*" UnE assms(1) assms(2) p1(5) tagged_division_ofD(3) xk'(1) xk'(2)) - by (metis "*" UnE assms(1) assms(2) inf_sup_aci(1) p2(5) tagged_division_ofD(3) xk xk'(1) xk'(2)) + with assms show "interior K \ interior K' = {}" + by (metis Un_iff inf_commute p1(3) p2(3) tagged_division_ofD(5) xK xk') qed lemma tagged_division_Union: @@ -1186,21 +997,16 @@ lemma tagged_partial_division_of_Union_self: assumes "p tagged_partial_division_of s" shows "p tagged_division_of (\(snd ` p))" - apply (rule tagged_division_ofI) using tagged_partial_division_ofD[OF assms] - apply auto - done + by (intro tagged_division_ofI) auto lemma tagged_division_of_union_self: assumes "p tagged_division_of s" shows "p tagged_division_of (\(snd ` p))" - apply (rule tagged_division_ofI) using tagged_division_ofD[OF assms] - apply auto - done + by (intro tagged_division_ofI) auto lemma tagged_division_Un_interval: - fixes a :: "'a::euclidean_space" assumes "p1 tagged_division_of (cbox a b \ {x. x\k \ (c::real)})" and "p2 tagged_division_of (cbox a b \ {x. x\k \ c})" and k: "k \ Basis" @@ -1208,13 +1014,10 @@ proof - have *: "cbox a b = (cbox a b \ {x. x\k \ c}) \ (cbox a b \ {x. x\k \ c})" by auto - show ?thesis - apply (subst *) - apply (rule tagged_division_Un[OF assms(1-2)]) - unfolding interval_split[OF k] interior_cbox - using k - apply (auto simp add: box_def elim!: ballE[where x=k]) - done + have "interior (cbox a b \ {x. x \ k \ c}) \ interior (cbox a b \ {x. c \ x \ k}) = {}" + using k by (force simp: interval_split[OF k] box_def) + with assms show ?thesis + by (metis "*" tagged_division_Un) qed lemma tagged_division_Un_interval_real: @@ -1223,35 +1026,23 @@ and "p2 tagged_division_of ({a .. b} \ {x. x\k \ c})" and k: "k \ Basis" shows "(p1 \ p2) tagged_division_of {a .. b}" - using assms - unfolding box_real[symmetric] - by (rule tagged_division_Un_interval) + using assms by (metis box_real(2) tagged_division_Un_interval) lemma tagged_division_split_left_inj: assumes d: "d tagged_division_of i" - and tags: "(x1, K1) \ d" "(x2, K2) \ d" - and "K1 \ K2" - and eq: "K1 \ {x. x \ k \ c} = K2 \ {x. x \ k \ c}" - shows "interior (K1 \ {x. x\k \ c}) = {}" -proof - - have "interior (K1 \ K2) = {} \ (x2, K2) = (x1, K1)" - using tags d by (metis (no_types) interior_Int tagged_division_ofD(5)) - then show ?thesis - using eq \K1 \ K2\ by (metis (no_types) inf_assoc inf_bot_left inf_left_idem interior_Int old.prod.inject) -qed + and tags: "(x1, K1) \ d" "(x2, K2) \ d" + and "K1 \ K2" + and eq: "K1 \ {x. x \ k \ c} = K2 \ {x. x \ k \ c}" + shows "interior (K1 \ {x. x\k \ c}) = {}" + by (smt (verit) Int_Un_eq(1) Un_Int_distrib interior_Int prod.inject sup_bot.right_neutral tagged_division_ofD(5) assms) lemma tagged_division_split_right_inj: assumes d: "d tagged_division_of i" - and tags: "(x1, K1) \ d" "(x2, K2) \ d" - and "K1 \ K2" - and eq: "K1 \ {x. x\k \ c} = K2 \ {x. x\k \ c}" - shows "interior (K1 \ {x. x\k \ c}) = {}" -proof - - have "interior (K1 \ K2) = {} \ (x2, K2) = (x1, K1)" - using tags d by (metis (no_types) interior_Int tagged_division_ofD(5)) - then show ?thesis - using eq \K1 \ K2\ by (metis (no_types) inf_assoc inf_bot_left inf_left_idem interior_Int old.prod.inject) -qed + and tags: "(x1, K1) \ d" "(x2, K2) \ d" + and "K1 \ K2" + and eq: "K1 \ {x. x\k \ c} = K2 \ {x. x\k \ c}" + shows "interior (K1 \ {x. x\k \ c}) = {}" + by (smt (verit) Int_Un_eq(1) Un_Int_distrib interior_Int prod.inject sup_bot.right_neutral tagged_division_ofD(5) assms) lemma (in comm_monoid_set) over_tagged_division_lemma: assumes "p tagged_division_of i" @@ -1272,14 +1063,10 @@ using assm(4)[of "fst x" "snd x"] \x\p\ by auto have "(fst x, snd y) \ p" "(fst x, snd y) \ y" using \x \ p\ \x \ y\ \snd x = snd y\ [symmetric] by auto - with \x\p\ \y\p\ have "interior (snd x) \ interior (snd y) = {}" - by (intro assm(5)[of "fst x" _ "fst y"]) auto - then have "box a b = {}" - unfolding \snd x = snd y\[symmetric] ab by auto - then have "d (cbox a b) = \<^bold>1" - using assm(2)[of "fst x" "snd x"] \x\p\ ab[symmetric] by (intro assms(2)) auto + with \x\p\ \y\p\ have "box a b = {}" + by (metis \snd x = snd y\ ab assm(5) inf.idem interior_cbox prod.collapse) then show "d (snd x) = \<^bold>1" - unfolding ab by auto + by (simp add: ab assms(2)) qed qed @@ -1311,11 +1098,8 @@ by standard (auto simp: lift_option_def ac_simps split: bind_split) qed -lemma comm_monoid_and: "comm_monoid HOL.conj True" - by standard auto - lemma comm_monoid_set_and: "comm_monoid_set HOL.conj True" - by (rule comm_monoid_set.intro) (fact comm_monoid_and) + by (simp add: comm_monoid_set.intro conj.comm_monoid_axioms) paragraph \Misc\ @@ -1323,9 +1107,7 @@ lemma interval_real_split: "{a .. b::real} \ {x. x \ c} = {a .. min b c}" "{a .. b} \ {x. c \ x} = {max a c .. b}" - apply (metis Int_atLeastAtMostL1 atMost_def) - apply (metis Int_atLeastAtMostL2 atLeast_def) - done + by auto lemma bgauge_existence_lemma: "(\x\s. \d::real. 0 < d \ q d x) \ (\x. \d>0. x\s \ q d x)" by (meson zero_less_one) @@ -1375,16 +1157,11 @@ proof - obtain u v where l: "l = cbox u v" using \l \ d\ assms(1) by blast - have *: "\i\Basis. u \ i \ (\i\Basis. (if i = k then min (v \ k) c else v \ i) *\<^sub>R i) \ i" - using that(6) unfolding l interval_split[OF k] box_ne_empty that . - have **: "\i\Basis. u\i \ v\i" + have "\i\Basis. u\i \ v\i" using l using that(6) unfolding box_ne_empty[symmetric] by auto - show ?thesis - apply (rule bexI[OF _ \l \ d\]) - using that(1-3,5) \x \ Basis\ - unfolding l interval_bounds[OF **] interval_bounds[OF *] interval_split[OF k] that - apply (auto split: if_split_asm) - done + then show ?thesis + using that \x \ Basis\ unfolding l interval_split[OF k] + by (force split: if_split_asm) qed moreover have "\x y. \y < (if x = k then c else b \ x)\ \ y < b \ x" using \c < b\k\ by (auto split: if_split_asm) @@ -1403,16 +1180,11 @@ proof - obtain u v where l: "l = cbox u v" using \l \ d\ assm(4) by blast - have *: "\i\Basis. (\i\Basis. (if i = k then max (u \ k) c else u \ i) *\<^sub>R i) \ i \ v \ i" - using that(6) unfolding l interval_split[OF k] box_ne_empty that . - have **: "\i\Basis. u\i \ v\i" + have "\i\Basis. u\i \ v\i" using l using that(6) unfolding box_ne_empty[symmetric] by auto - show "\i\d. interval_lowerbound i \ x = y \ interval_upperbound i \ x = y" - apply (rule bexI[OF _ \l \ d\]) - using that(1-3,5) \x \ Basis\ - unfolding l interval_bounds[OF **] interval_bounds[OF *] interval_split[OF k] that - apply (auto split: if_split_asm) - done + then show "\i\d. interval_lowerbound i \ x = y \ interval_upperbound i \ x = y" + using that \x \ Basis\ unfolding l interval_split[OF k] + by (force split: if_split_asm) qed ultimately show ?t2 unfolding division_points_def interval_split[OF k, of a b] @@ -1520,8 +1292,6 @@ shows "(\l. l \ {x. \x\k - c\ \ e}) ` {l\p. l \ {x. \x\k - c\ \ e} \ {}} division_of (cbox a b \ {x. \x\k - c\ \ e})" proof - - have *: "\x c. \x - c\ \ e \ x \ c - e \ x \ c + e" - by auto have **: "\p q p' q'. p division_of q \ p = p' \ q = q' \ p' division_of q'" by auto note division_split(1)[OF assms, where c="c+e",unfolded interval_split[OF k]] @@ -1533,7 +1303,8 @@ apply (rule equalityI) apply blast apply clarsimp - apply (rule_tac x="xa \ {x. c + e \ x \ k}" in exI) + apply (rename_tac S) + apply (rule_tac x="S \ {x. c + e \ x \ k}" in exI) apply auto done by (simp add: interval_split k interval_doublesplit) @@ -1547,16 +1318,8 @@ and Basis_imp: "\a b c k. k \ Basis \ g (cbox a b) = g (cbox a b \ {x. x\k \ c}) \<^bold>* g (cbox a b \ {x. x\k \ c})" begin -lemma empty [simp]: - "g {} = \<^bold>1" -proof - - have *: "cbox One (-One) = ({}::'b set)" - by (auto simp: box_eq_empty inner_sum_left inner_Basis sum.If_cases ex_in_conv) - moreover have "box One (-One) = ({}::'b set)" - using box_subset_cbox[of One "-One"] by (auto simp: *) - ultimately show ?thesis - using box_empty_imp [of One "-One"] by simp -qed +lemma empty [simp]: "g {} = \<^bold>1" + by (metis box_empty_imp box_subset_cbox empty_as_interval subset_empty) lemma division: "F g d = g (cbox a b)" if "d division_of (cbox a b)" @@ -1566,21 +1329,22 @@ using that proof (induction C arbitrary: a b d rule: less_induct) case (less a b d) show ?case - proof cases - assume "box a b = {}" + proof (cases "box a b = {}") + case True { fix k assume "k\d" then obtain a' b' where k: "k = cbox a' b'" using division_ofD(4)[OF less.prems] by blast - with \k\d\ division_ofD(2)[OF less.prems] have "cbox a' b' \ cbox a b" - by auto + then have "cbox a' b' \ cbox a b" + using \k \ d\ less.prems by blast then have "box a' b' \ box a b" unfolding subset_box by auto then have "g k = \<^bold>1" - using box_empty_imp [of a' b'] k by (simp add: \box a b = {}\) } - then show "box a b = {} \ F g d = g (cbox a b)" + using box_empty_imp [of a' b'] k by (simp add: True) + } + with True show "F g d = g (cbox a b)" by (auto intro!: neutral simp: box_empty_imp) next - assume "box a b \ {}" + case False then have ab: "\i\Basis. a\i < b\i" and ab': "\i\Basis. a\i \ b\i" by (auto simp: box_ne_empty) show "F g d = g (cbox a b)" @@ -1600,11 +1364,10 @@ note *[OF this(1)] *[OF this(2)] note this[unfolded interval_bounds[OF uv(1)]] moreover have "a\j \ u\j" "v\j \ b\j" - using division_ofD(2,2,3)[OF \d division_of cbox a b\ as] - apply (metis j subset_box(1) uv(1)) - by (metis \cbox u v \ cbox a b\ j subset_box(1) uv(1)) + by (meson as division_ofD(2) j less.prems subset_box(1) uv(1))+ ultimately have "u\j = a\j \ v\j = a\j \ u\j = b\j \ v\j = b\j \ u\j = a\j \ v\j = b\j" - unfolding not_less de_Morgan_disj using ab[rule_format,of j] uv(2) j by force } + using uv(2) by force + } then have d': "\i\d. \u v. i = cbox u v \ (\j\Basis. u\j = a\j \ v\j = a\j \ u\j = b\j \ v\j = b\j \ u\j = a\j \ v\j = b\j)" unfolding forall_in_division[OF less.prems] by blast @@ -1624,9 +1387,9 @@ proof safe fix j :: 'b assume j: "j \ Basis" - note i(2)[unfolded uv mem_box,rule_format,of j] + note i(2)[unfolded uv mem_box] then show "u \ j = a \ j" and "v \ j = b \ j" - using uv(2)[rule_format,of j] j by (auto simp: inner_simps) + by (smt (verit) False box_midpoint j mem_box(1) uv(2))+ qed then have "i = cbox a b" using uv by auto then show ?thesis using i by auto @@ -1650,17 +1413,11 @@ unfolding euclidean_eq_iff[where 'a='b] by auto then have "u\j = v\j" using uv(2)[rule_format,OF j] by auto - then have "box u v = {}" - using j unfolding box_eq_empty by (auto intro!: bexI[of _ j]) then show "g x = \<^bold>1" - unfolding uv(1) by (rule box_empty_imp) + by (smt (verit) box_empty_imp box_eq_empty(1) j uv(1)) qed then show "F g d = g (cbox a b)" - using division_ofD[OF less.prems] - apply (subst deq) - apply (subst insert) - apply auto - done + by (metis deq division_of_finite insert_Diff_single insert_remove less.prems right_neutral) next case False then have "\x. x \ division_points (cbox a b) d" @@ -1748,12 +1505,7 @@ proposition tagged_division: assumes "d tagged_division_of (cbox a b)" shows "F (\(_, l). g l) d = g (cbox a b)" -proof - - have "F (\(_, k). g k) d = F g (snd ` d)" - using assms box_empty_imp by (rule over_tagged_division_lemma) - then show ?thesis - unfolding assms [THEN division_of_tagged_division, THEN division] . - qed + by (metis assms box_empty_imp division division_of_tagged_division over_tagged_division_lemma) end @@ -1778,14 +1530,14 @@ from that have [simp]: "k = 1" by simp from neutral [of 0 1] neutral [of a a for a] coalesce_less - have [simp]: "g {} = \<^bold>1" "\a. g {a} = \<^bold>1" - "\a b c. a < c \ c < b \ g {a..c} \<^bold>* g {c..b} = g {a..b}" - by auto + have [simp]: "g {} = \<^bold>1" "\a. g {a} = \<^bold>1" + "\a b c. a < c \ c < b \ g {a..c} \<^bold>* g {c..b} = g {a..b}" + by auto have "g {a..b} = g {a..min b c} \<^bold>* g {max a c..b}" - by (auto simp: min_def max_def le_less) + by (auto simp: min_def max_def le_less) then show "g (cbox a b) = g (cbox a b \ {x. x \ k \ c}) \<^bold>* g (cbox a b \ {x. c \ x \ k})" by (simp add: atMost_def [symmetric] atLeast_def [symmetric]) -qed + qed qed show "box = (greaterThanLessThan :: real \ _)" and "cbox = (atLeastAtMost :: real \ _)" @@ -1795,17 +1547,7 @@ lemma coalesce_less_eq: "g {a..c} \<^bold>* g {c..b} = g {a..b}" if "a \ c" "c \ b" - proof (cases "c = a \ c = b") - case False - with that have "a < c" "c < b" - by auto - then show ?thesis - by (rule coalesce_less) - next - case True - with that box_empty_imp [of a a] box_empty_imp [of b b] show ?thesis - by safe simp_all - qed + by (metis coalesce_less commute left_neutral less_eq_real_def neutral that) end @@ -1819,22 +1561,22 @@ show "g {a..b} = z" if "b \ a" for a b using that box_empty_imp by simp show "f (g {a..c}) (g {c..b}) = g {a..b}" if "a < c" "c < b" for a b c - using that - using Basis_imp [of 1 a b c] + using that Basis_imp [of 1 a b c] by (simp_all add: atMost_def [symmetric] atLeast_def [symmetric] max_def min_def) -qed + qed qed subsection \Special case of additivity we need for the FTC\ (*fix add explanation here *) + lemma additive_tagged_division_1: fixes f :: "real \ 'a::real_normed_vector" assumes "a \ b" and "p tagged_division_of {a..b}" - shows "sum (\(x,k). f(Sup k) - f(Inf k)) p = f b - f a" + shows "sum (\(x,K). f(Sup K) - f(Inf K)) p = f b - f a" proof - - let ?f = "(\k::(real) set. if k = {} then 0 else f(interval_upperbound k) - f(interval_lowerbound k))" + let ?f = "(\K::real set. if K = {} then 0 else f(interval_upperbound K) - f(interval_lowerbound K))" interpret operative_real plus 0 ?f rewrites "comm_monoid_set.F (+) 0 = sum" by standard[1] (auto simp add: sum_def) @@ -1843,12 +1585,7 @@ have **: "cbox a b \ {}" using assms(1) by auto then have "f b - f a = (\(x, l)\p. if l = {} then 0 else f (interval_upperbound l) - f (interval_lowerbound l))" - proof - - have "(if cbox a b = {} then 0 else f (interval_upperbound (cbox a b)) - f (interval_lowerbound (cbox a b))) = f b - f a" - using assms by auto - then show ?thesis - using p_td assms by (simp add: tagged_division) - qed + using assms(2) tagged_division by force then show ?thesis using assms by (auto intro!: sum.cong) qed @@ -1897,12 +1634,16 @@ obtain pfn where pfn: "\x. x \ I \ pfn x tagged_division_of x" "\x. x \ I \ d fine pfn x" - using bchoice[OF assms(2)] by auto + using assms by metis show thesis - apply (rule_tac p="\(pfn ` I)" in that) - using assms(1) assms(3) assms(4) pfn(1) tagged_division_Union apply force - by (metis (mono_tags, lifting) fine_Union imageE pfn(2)) + proof + show "\ (pfn ` I) tagged_division_of i" + using assms pfn(1) tagged_division_Union by force + show "d fine \ (pfn ` I)" + by (metis (mono_tags, lifting) fine_Union imageE pfn(2)) + qed qed + (* FIXME structure here, do not have one lemma in each subsection *) subsection\<^marker>\tag unimportant\ \The set we're concerned with must be closed\ @@ -1910,9 +1651,11 @@ lemma division_of_closed: fixes i :: "'n::euclidean_space set" shows "s division_of i \ closed i" - unfolding division_of_def by fastforce + by blast (* FIXME structure here, do not have one lemma in each subsection *) + subsection \General bisection principle for intervals; might be useful elsewhere\ + (* FIXME move? *) lemma interval_bisection_step: fixes type :: "'a::euclidean_space" @@ -1952,11 +1695,10 @@ assume "x \ ?A" then obtain c d where x: "x = cbox c d" - "\i. i \ Basis \ - c \ i = a \ i \ d \ i = ?ab i \ c \ i = ?ab i \ d \ i = b \ i" + "\i. i \ Basis \ c \ i = a \ i \ d \ i = ?ab i \ c \ i = ?ab i \ d \ i = b \ i" by blast have "c = (\i\Basis. (if c \ i = a \ i then a \ i else ?ab i) *\<^sub>R i)" - "d = (\i\Basis. (if c \ i = a \ i then ?ab i else b \ i) *\<^sub>R i)" + "d = (\i\Basis. (if c \ i = a \ i then ?ab i else b \ i) *\<^sub>R i)" using x(2) ab by (fastforce simp add: euclidean_eq_iff[where 'a='a])+ then show "x \ ?B" unfolding x by (rule_tac x="{i. i\Basis \ c\i = a\i}" in image_eqI) auto @@ -1968,7 +1710,7 @@ assume "S \ ?A" then obtain c d where s: "S = cbox c d" - "\i. i \ Basis \ c \ i = a \ i \ d \ i = ?ab i \ c \ i = ?ab i \ d \ i = b \ i" + "\i. i \ Basis \ c \ i = a \ i \ d \ i = ?ab i \ c \ i = ?ab i \ d \ i = b \ i" by blast show "P S" unfolding s using ab s(2) by (fastforce intro!: that) @@ -1986,8 +1728,7 @@ then obtain i where "c\i \ e\i \ d\i \ f\i" and i': "i \ Basis" unfolding euclidean_eq_iff[where 'a='a] by auto then have i: "c\i \ e\i" "d\i \ f\i" - using s(2) t(2) apply fastforce - using t(2)[OF i'] \c \ i \ e \ i \ d \ i \ f \ i\ i' s(2) t(2) by fastforce + using s(2) t(2) by fastforce+ have *: "\s t. (\a. a \ s \ a \ t \ False) \ s \ t = {}" by auto show "interior S \ interior T = {}" @@ -1995,10 +1736,10 @@ proof (rule *) fix x assume "x \ box c d" "x \ box e f" - then have x: "c\i < d\i" "e\i < f\i" "c\i < f\i" "e\i < d\i" - unfolding mem_box using i' by force+ - show False using s(2)[OF i'] t(2)[OF i'] and i x - by auto + then have "c\i < f\i" "e\i < d\i" + unfolding mem_box using i' by force+ + with i i' show False + using s(2) t(2) by fastforce qed qed also have "\?A = cbox a b" @@ -2024,7 +1765,7 @@ by (force simp add: mem_box) qed finally show thesis - by (metis (no_types, lifting) assms(3) that) + by (metis (no_types, lifting) assms(3) that) qed lemma interval_bisection: @@ -2066,9 +1807,10 @@ snd (f x) \ i \ snd x \ i \ 2 * (snd (f x) \ i - fst (f x) \ i) \ snd x \ i - fst x \ i)" by metis define AB A B where ab_def: "AB n = (f ^^ n) (a,b)" "A n = fst(AB n)" "B n = snd(AB n)" for n - have [simp]: "A 0 = a" "B 0 = b" and ABRAW: "\n. \ P (cbox (A(Suc n)) (B(Suc n))) \ - (\i\Basis. A(n)\i \ A(Suc n)\i \ A(Suc n)\i \ B(Suc n)\i \ B(Suc n)\i \ B(n)\i \ - 2 * (B(Suc n)\i - A(Suc n)\i) \ B(n)\i - A(n)\i)" (is "\n. ?P n") + have [simp]: "A 0 = a" "B 0 = b" + and ABRAW: "\n. \ P (cbox (A(Suc n)) (B(Suc n))) \ + (\i\Basis. A(n)\i \ A(Suc n)\i \ A(Suc n)\i \ B(Suc n)\i \ B(Suc n)\i \ B(n)\i \ + 2 * (B(Suc n)\i - A(Suc n)\i) \ B(n)\i - A(n)\i)" (is "\n. ?P n") proof - show "A 0 = a" "B 0 = b" unfolding ab_def by auto @@ -2080,13 +1822,9 @@ unfolding S using \\ P (cbox a b)\ f by auto next case (Suc n) - show ?case + then show ?case unfolding S - apply (rule f[rule_format]) - using Suc - unfolding S - apply auto - done + by (force intro!: f[rule_format]) qed qed then have AB: "A(n)\i \ A(Suc n)\i" "A(Suc n)\i \ B(Suc n)\i" @@ -2109,10 +1847,9 @@ also have "\ \ sum (\i. B n\i - A n\i) Basis" proof (rule sum_mono) fix i :: 'a - assume i: "i \ Basis" - show "\(x - y) \ i\ \ B n \ i - A n \ i" - using xy[unfolded mem_box,THEN bspec, OF i] - by (auto simp: inner_diff_left) + assume "i \ Basis" + with xy show "\(x - y) \ i\ \ B n \ i - A n \ i" + by (smt (verit, best) inner_diff_left mem_box(2)) qed also have "\ \ sum (\i. b\i - a\i) Basis / 2^n" unfolding sum_divide_distrib @@ -2136,23 +1873,21 @@ finally show "dist x y < e" . qed qed - { - fix n m :: nat - assume "m \ n" then have "cbox (A n) (B n) \ cbox (A m) (B m)" - proof (induction rule: inc_induct) - case (step i) - show ?case - using AB by (intro order_trans[OF step.IH] subset_box_imp) auto - qed simp - } note ABsubset = this + have ABsubset: "cbox (A n) (B n) \ cbox (A m) (B m)" if "m \ n" for m n + using that + proof (induction rule: inc_induct) + case (step i) show ?case + by (smt (verit, ccfv_SIG) ABRAW in_mono mem_box(2) step.IH subsetI) + qed simp have "\n. cbox (A n) (B n) \ {}" by (meson AB dual_order.trans interval_not_empty) then obtain x0 where x0: "\n. x0 \ cbox (A n) (B n)" using decreasing_closed_nest [OF closed_cbox] ABsubset interv by blast show thesis - proof (rule that[rule_format, of x0]) - show "x0\cbox a b" + proof (intro that strip) + show "x0 \ cbox a b" using \A 0 = a\ \B 0 = b\ x0 by blast + next fix e :: real assume "e > 0" from interv[OF this] obtain n @@ -2170,17 +1905,14 @@ moreover have "cbox (A n) (B n) \ cbox a b" using ABsubset \A 0 = a\ \B 0 = b\ by blast ultimately show "\c d. x0 \ cbox c d \ cbox c d \ ball x0 e \ cbox c d \ cbox a b \ \ P (cbox c d)" - apply (rule_tac x="A n" in exI) - apply (rule_tac x="B n" in exI) - apply (auto simp: x0) - done + by (meson x0) qed qed subsection \Cousin's lemma\ -lemma fine_division_exists: (*FIXME rename(?) *) +lemma fine_division_exists: fixes a b :: "'a::euclidean_space" assumes "gauge g" obtains p where "p tagged_division_of (cbox a b)" "g fine p" @@ -2199,23 +1931,21 @@ cbox c d \ ball x e \ cbox c d \ (cbox a b) \ \ (\p. p tagged_division_of cbox c d \ g fine p)" - apply (rule interval_bisection[of "\s. \p. p tagged_division_of s \ g fine p", OF _ _ False]) - apply (simp add: fine_def) - apply (metis tagged_division_Un fine_Un) - apply auto - done + proof (rule interval_bisection[of "\s. \p. p tagged_division_of s \ _ p", OF _ _ False]) + show "\p. p tagged_division_of {} \ g fine p" + by (simp add: fine_def) + qed (meson tagged_division_Un fine_Un)+ obtain e where e: "e > 0" "ball x e \ g x" - using gaugeD[OF assms, of x] unfolding open_contains_ball by auto - from x(2)[OF e(1)] - obtain c d where c_d: "x \ cbox c d" + by (meson assms gauge_def openE) + then obtain c d where c_d: "x \ cbox c d" "cbox c d \ ball x e" "cbox c d \ cbox a b" "\ (\p. p tagged_division_of cbox c d \ g fine p)" - by blast + by (meson x(2)) have "g fine {(x, cbox c d)}" unfolding fine_def using e using c_d(2) by auto then show ?thesis - using tagged_division_of_self[OF c_d(1)] using c_d by auto + using tagged_division_of_self[OF c_d(1)] using c_d by simp qed lemma fine_division_exists_real: @@ -2232,7 +1962,7 @@ and "gauge d" obtains q where "q tagged_division_of (cbox a b)" and "d fine q" - and "\(x,k) \ p. k \ d(x) \ (x,k) \ q" + and "\(x,K) \ p. K \ d(x) \ (x,K) \ q" proof - have p: "finite p" "p tagged_partial_division_of (cbox a b)" using ptag tagged_division_of_def by blast+ @@ -2250,48 +1980,48 @@ obtain q1 where q1: "q1 tagged_division_of \{k. \x. (x, k) \ p}" and "d fine q1" and q1I: "\x k. \(x, k)\p; k \ d x\ \ (x, k) \ q1" - using case_prodD tagged_partial_division_subset[OF insert(4) subset_insertI] - by (metis (mono_tags, lifting) insert.hyps(3) insert.prems(2)) + using insert + by (metis (mono_tags, lifting) case_prodD subset_insertI tagged_partial_division_subset) have *: "\{l. \y. (y,l) \ insert xk p} = k \ \{l. \y. (y,l) \ p}" unfolding xk by auto note p = tagged_partial_division_ofD[OF insert(4)] obtain u v where uv: "k = cbox u v" using p(4) xk by blast - have "finite {k. \x. (x, k) \ p}" - apply (rule finite_subset[of _ "snd ` p"]) - using image_iff apply fastforce - using insert.hyps(1) by blast + have "{K. \x. (x, K) \ p} \ snd ` p" + by force + then have "finite {K. \x. (x, K) \ p}" + using finite_surj insert.hyps(1) by blast then have int: "interior (cbox u v) \ interior (\{k. \x. (x, k) \ p}) = {}" proof (rule Int_interior_Union_intervals) show "open (interior (cbox u v))" by auto - show "\T. T \ {k. \x. (x, k) \ p} \ \a b. T = cbox a b" + show "\T. T \ {K. \x. (x, K) \ p} \ \a b. T = cbox a b" using p(4) by auto - show "\T. T \ {k. \x. (x, k) \ p} \ interior (cbox u v) \ interior T = {}" - by clarify (metis insert.hyps(2) insert_iff interior_cbox p(5) uv xk) + show "\T. T \ {K. \x. (x, K) \ p} \ interior (cbox u v) \ interior T = {}" + using insert.hyps(2) p(5) uv xk by blast qed show ?case proof (cases "cbox u v \ d x") case True have "{(x, cbox u v)} tagged_division_of cbox u v" by (simp add: p(2) uv xk tagged_division_of_self) - then have "{(x, cbox u v)} \ q1 tagged_division_of \{k. \x. (x, k) \ insert xk p}" - unfolding * uv by (metis (no_types, lifting) int q1 tagged_division_Un) - with True show ?thesis - apply (rule_tac x="{(x,cbox u v)} \ q1" in exI) - using \d fine q1\ fine_def q1I uv xk apply fastforce - done + then have "{(x, cbox u v)} \ q1 tagged_division_of \{K. \x. (x, K) \ insert xk p}" + by (smt (verit, best) "*" int q1 tagged_division_Un uv) + moreover have "d fine ({(x,cbox u v)} \ q1)" + using True \d fine q1\ fine_def by fastforce + ultimately show ?thesis + by (metis (no_types, lifting) case_prodI2 insert_iff insert_is_Un q1I uv xk) next case False obtain q2 where q2: "q2 tagged_division_of cbox u v" "d fine q2" using fine_division_exists[OF assms(2)] by blast show ?thesis - apply (rule_tac x="q2 \ q1" in exI) - apply (intro conjI) - unfolding * uv - apply (rule tagged_division_Un q2 q1 int fine_Un)+ - apply (auto intro: q1 q2 fine_Un \d fine q1\ simp add: False q1I uv xk) - done + proof (intro exI conjI) + show "q2 \ q1 tagged_division_of \ {k. \x. (x, k) \ insert xk p}" + by (smt (verit, best) "*" int q1 q2(1) tagged_division_Un uv) + show "d fine q2 \ q1" + by (simp add: \d fine q1\ fine_Un q2(2)) + qed (use False uv xk q1I in auto) qed qed with p obtain q where q: "q tagged_division_of \{k. \x. (x, k) \ p}" "d fine q" "\(x, k)\p. k \ d x \ (x, k) \ q" @@ -2335,8 +2065,7 @@ show "\?D0 \ cbox a b" apply (simp add: UN_subset_iff) apply (intro conjI allI ballI subset_box_imp) - apply (simp add: field_simps) - apply (auto intro: mult_right_mono aibi) + apply (simp add: field_simps aibi mult_right_mono) apply (force simp: aibi scaling_mono nat_less_real_le dest: PiE_mem intro: mult_right_mono) done next @@ -2352,22 +2081,19 @@ fix v w m and n::nat assume "m \ n" \ \WLOG we can assume \<^term>\m \ n\, when the first disjunct becomes impossible\ have "?K0(n,w) \ ?K0(m,v) \ interior(?K0(m,v)) \ interior(?K0(n,w)) = {}" + apply (rule ccontr) apply (simp add: subset_box disjoint_interval) - apply (rule ccontr) apply (clarsimp simp add: aibi mult_le_cancel_right divide_le_cancel not_less not_le) apply (drule_tac x=i in bspec, assumption) using \m\n\ realff [of _ _ "1+_"] realff [of "1+_"_ "1+_"] apply (auto simp: divide_simps add.commute not_le nat_le_iff_add realff) - apply (simp_all add: power_add) - apply (metis (no_types, opaque_lifting) mult_Suc mult_less_cancel2 not_less_eq mult.assoc) - apply (metis (no_types, opaque_lifting) mult_Suc mult_less_cancel2 not_less_eq mult.assoc) + apply (metis (no_types, opaque_lifting) power_add mult_Suc mult_less_cancel2 not_less_eq mult.assoc)+ done then show "?K0(m,v) \ ?K0(n,w) \ ?K0(n,w) \ ?K0(m,v) \ interior(?K0(m,v)) \ interior(?K0(n,w)) = {}" by meson qed auto show "\A B. \A \ ?D0; B \ ?D0\ \ A \ B \ B \ A \ interior A \ interior B = {}" - apply (erule imageE SigmaE)+ - using * by simp + using * by fastforce next show "\K \ ?D0. x \ K \ K \ g x" if "x \ S" for x proof (simp only: bex_simps split_paired_Bex_Sigma) @@ -2389,9 +2115,8 @@ by (meson sum_bounded_above that) also have "... = \ / 2" by (simp add: field_split_simps) - also have "... < \" - by (simp add: \0 < \\) - finally show ?thesis . + finally show ?thesis + using \0 < \\ by linarith qed then show ?thesis by (rule_tac e = "\ / 2 / DIM('a)" in that) (simp_all add: \0 < \\ dist_norm subsetD [OF \]) @@ -2404,13 +2129,7 @@ then have "norm (b - a) < e * 2^n" by (auto simp: field_split_simps) then have bai: "b \ i - a \ i < e * 2 ^ n" if "i \ Basis" for i - proof - - have "b \ i - a \ i \ norm (b - a)" - by (metis abs_of_nonneg dual_order.trans inner_diff_left linear norm_ge_zero Basis_le_norm that) - also have "... < e * 2 ^ n" - using \norm (b - a) < e * 2 ^ n\ by blast - finally show ?thesis . - qed + by (smt (verit, del_insts) Basis_le_norm diff_add_cancel inner_simps(1) that) have D: "(a + n \ x \ x \ a + m) \ (a + n \ y \ y \ a + m) \ abs(x - y) \ m - n" for a m n x and y::real by auto @@ -2438,9 +2157,9 @@ next have "a \ i + real ?k * (b \ i - a \ i) / 2 ^ n \ a \ i + (2 ^ n * (x \ i - a \ i) / (b \ i - a \ i)) * (b \ i - a \ i) / 2 ^ n" + using aibi [OF \i \ Basis\] xab 2 apply (intro add_left_mono mult_right_mono divide_right_mono of_nat_floor) - using aibi [OF \i \ Basis\] xab 2 - apply (simp_all add: \i \ Basis\ mem_box field_split_simps) + apply (auto simp: \i \ Basis\ mem_box field_split_simps) done also have "... = x \ i" using abi_less by (simp add: field_split_simps) @@ -2449,8 +2168,8 @@ have "x \ i \ a \ i + (2 ^ n * (x \ i - a \ i) / (b \ i - a \ i)) * (b \ i - a \ i) / 2 ^ n" using abi_less by (simp add: field_split_simps) also have "... \ a \ i + (real ?k + 1) * (b \ i - a \ i) / 2 ^ n" + using aibi [OF \i \ Basis\] xab apply (intro add_left_mono mult_right_mono divide_right_mono of_nat_floor) - using aibi [OF \i \ Basis\] xab apply (auto simp: \i \ Basis\ mem_box divide_simps) done finally show "x \ i \ a \ i + (real ?k + 1) * (b \ i - a \ i) / 2 ^ n" . @@ -2540,7 +2259,7 @@ qed let ?\ = "{K \ \. \K'. K' \ \ \ K \ K' \ \(K \ K')}" show ?thesis - proof (rule that) + proof show "countable ?\" by (blast intro: countable_subset [OF _ count]) show "\?\ \ cbox a b" @@ -2597,6 +2316,6 @@ lemma eventually_division_filter_tagged_division: "eventually (\p. p tagged_division_of s) (division_filter s)" - unfolding eventually_division_filter by (intro exI[of _ "\x. ball x 1"]) auto + using eventually_division_filter by auto end diff -r ba2afdd29e1d -r 032a4344903e src/HOL/Complex_Analysis/Conformal_Mappings.thy --- a/src/HOL/Complex_Analysis/Conformal_Mappings.thy Tue Aug 01 11:27:55 2023 +0200 +++ b/src/HOL/Complex_Analysis/Conformal_Mappings.thy Thu Aug 03 19:10:43 2023 +0200 @@ -1751,7 +1751,7 @@ have "(*) ((cmod z)\<^sup>2) integrable_on {0..1}" by (metis ident_integrable_on integrable_0 integrable_eq integrable_on_cmult_iff lambda_zero) then show "(\t. (norm z)\<^sup>2 * t / (r - norm z) * C) integrable_on {0..1}" - using integrable_on_cmult_right[where 'b=real, simplified] integrable_on_cdivide [where 'b=real, simplified] + using integrable_on_cmult_right[where 'b=real, simplified] integrable_on_divide [where 'b=real, simplified] by blast qed (simp add: norm_mult power2_eq_square 4) then have int_le: "norm (f z - deriv f 0 * z) \ (norm z)\<^sup>2 * norm(deriv f 0) / ((r - norm z))"