# HG changeset patch # User paulson # Date 1672533955 0 # Node ID 30182f9e1818c220110389775eb4f5de35789da7 # Parent 8d8af7e92c5e676c88f0e9c8c3885ffedee0ca4f Big simplifications of old proofs diff -r 8d8af7e92c5e -r 30182f9e1818 src/HOL/Analysis/Affine.thy --- a/src/HOL/Analysis/Affine.thy Sat Dec 31 11:09:19 2022 +0000 +++ b/src/HOL/Analysis/Affine.thy Sun Jan 01 00:45:55 2023 +0000 @@ -51,7 +51,7 @@ using dim_unique[of B B "card B"] assms span_superset[of B] by auto have "dim B \ card (Basis :: 'a set)" using dim_subset_UNIV[of B] by simp - from ex_card[OF this] obtain d :: "'a set" where d: "d \ Basis" and t: "card d = dim B" + from obtain_subset_with_card_n[OF this] obtain d :: "'a set" where d: "d \ Basis" and t: "card d = dim B" by auto let ?t = "{x::'a::euclidean_space. \i\Basis. i \ d \ x\i = 0}" have "\f. linear f \ f ` B = d \ f ` span B = ?t \ inj_on f (span B)" diff -r 8d8af7e92c5e -r 30182f9e1818 src/HOL/Analysis/Determinants.thy --- a/src/HOL/Analysis/Determinants.thy Sat Dec 31 11:09:19 2022 +0000 +++ b/src/HOL/Analysis/Determinants.thy Sun Jan 01 00:45:55 2023 +0000 @@ -1025,7 +1025,7 @@ then show ?thesis using \A *v axis k 1 = a\ that by auto next - from ex_card[OF 2] obtain h i::'n where "h \ i" + from obtain_subset_with_card_n[OF 2] obtain h i::'n where "h \ i" by (auto simp add: eval_nat_numeral card_Suc_eq) then obtain j where "j \ k" by (metis (full_types)) diff -r 8d8af7e92c5e -r 30182f9e1818 src/HOL/Analysis/Linear_Algebra.thy --- a/src/HOL/Analysis/Linear_Algebra.thy Sat Dec 31 11:09:19 2022 +0000 +++ b/src/HOL/Analysis/Linear_Algebra.thy Sun Jan 01 00:45:55 2023 +0000 @@ -68,38 +68,17 @@ fixes a :: "'a::ab_group_add" assumes "(\x. a + x) ` A = (\x. a + x) ` B" shows "A = B" -proof - - have "(\x. -a + x) ` ((\x. a + x) ` A) = (\x. - a + x) ` ((\x. a + x) ` B)" - using assms by auto - then show ?thesis - using translation_assoc[of "-a" a A] translation_assoc[of "-a" a B] by auto -qed + using assms translation_assoc by fastforce lemma translation_galois: fixes a :: "'a::ab_group_add" shows "T = ((\x. a + x) ` S) \ S = ((\x. (- a) + x) ` T)" - using translation_assoc[of "-a" a S] - apply auto - using translation_assoc[of a "-a" T] - apply auto - done + by (metis add.right_inverse group_cancel.rule0 translation_invert translation_assoc) lemma translation_inverse_subset: assumes "((\x. - a + x) ` V) \ (S :: 'n::ab_group_add set)" shows "V \ ((\x. a + x) ` S)" -proof - - { - fix x - assume "x \ V" - then have "x-a \ S" using assms by auto - then have "x \ {a + v |v. v \ S}" - apply auto - apply (rule exI[of _ "x-a"], simp) - done - then have "x \ ((\x. a+x) ` S)" by auto - } - then show ?thesis by auto -qed + by (metis assms subset_image_iff translation_galois) subsection\<^marker>\tag unimportant\ \More interesting properties of the norm\ @@ -120,41 +99,25 @@ qed lemma vector_eq: "x = y \ x \ x = x \ y \ y \ y = x \ x" - (is "?lhs \ ?rhs") -proof - assume ?lhs - then show ?rhs by simp -next - assume ?rhs - then have "x \ x - x \ y = 0 \ x \ y - y \ y = 0" - by simp - then have "x \ (x - y) = 0 \ y \ (x - y) = 0" - by (simp add: inner_diff inner_commute) - then have "(x - y) \ (x - y) = 0" - by (simp add: field_simps inner_diff inner_commute) - then show "x = y" by simp -qed + by (metis (no_types, opaque_lifting) inner_commute inner_diff_right inner_eq_zero_iff right_minus_eq) lemma norm_triangle_half_r: - "norm (y - x1) < e / 2 \ norm (y - x2) < e / 2 \ norm (x1 - x2) < e" + "norm (y - x1) < e/2 \ norm (y - x2) < e/2 \ norm (x1 - x2) < e" using dist_triangle_half_r unfolding dist_norm[symmetric] by auto lemma norm_triangle_half_l: - assumes "norm (x - y) < e / 2" - and "norm (x' - y) < e / 2" + assumes "norm (x - y) < e/2" and "norm (x' - y) < e/2" shows "norm (x - x') < e" - using dist_triangle_half_l[OF assms[unfolded dist_norm[symmetric]]] - unfolding dist_norm[symmetric] . + by (metis assms dist_norm dist_triangle_half_l) lemma abs_triangle_half_r: fixes y :: "'a::linordered_field" - shows "abs (y - x1) < e / 2 \ abs (y - x2) < e / 2 \ abs (x1 - x2) < e" + shows "abs (y - x1) < e/2 \ abs (y - x2) < e/2 \ abs (x1 - x2) < e" by linarith lemma abs_triangle_half_l: fixes y :: "'a::linordered_field" - assumes "abs (x - y) < e / 2" - and "abs (x' - y) < e / 2" + assumes "abs (x - y) < e/2" and "abs (x' - y) < e/2" shows "abs (x - x') < e" using assms by linarith @@ -163,41 +126,15 @@ and "finite S \ sum f (insert x S) = (if x \ S then sum f S else f x + sum f S)" by (auto simp add: insert_absorb) -lemma vector_eq_ldot: "(\x. x \ y = x \ z) \ y = z" -proof - assume "\x. x \ y = x \ z" - then have "\x. x \ (y - z) = 0" - by (simp add: inner_diff) - then have "(y - z) \ (y - z) = 0" .. - then show "y = z" by simp -qed simp - -lemma vector_eq_rdot: "(\z. x \ z = y \ z) \ x = y" -proof - assume "\z. x \ z = y \ z" - then have "\z. (x - y) \ z = 0" - by (simp add: inner_diff) - then have "(x - y) \ (x - y) = 0" .. - then show "x = y" by simp -qed simp +lemma vector_eq_ldot: "(\x. x \ y = x \ z) \ y = z" and vector_eq_rdot: "(\z. x \ z = y \ z) \ x = y" + by (metis inner_commute vector_eq)+ subsection \Substandard Basis\ lemma ex_card: assumes "n \ card A" shows "\S\A. card S = n" -proof (cases "finite A") - case True - from ex_bij_betw_nat_finite[OF this] obtain f where f: "bij_betw f {0..n \ card A\ have "{..< n} \ {..< card A}" "inj_on f {..< n}" - by (auto simp: bij_betw_def intro: subset_inj_on) - ultimately have "f ` {..< n} \ A" "card (f ` {..< n}) = n" - by (auto simp: bij_betw_def card_image) - then show ?thesis by blast -next - case False - with \n \ card A\ show ?thesis by force -qed + by (meson assms obtain_subset_with_card_n) lemma subspace_substandard: "subspace {x::'a::euclidean_space. (\i\Basis. P i \ x\i = 0)}" by (auto simp: subspace_def inner_add_left) @@ -271,7 +208,7 @@ lemma norm_add_Pythagorean: assumes "orthogonal a b" - shows "norm(a + b) ^ 2 = norm a ^ 2 + norm b ^ 2" + shows "(norm (a + b))\<^sup>2 = (norm a)\<^sup>2 + (norm b)\<^sup>2" proof - from assms have "(a - (0 - b)) \ (a - (0 - b)) = a \ a - (0 - b \ b)" by (simp add: algebra_simps orthogonal_def inner_commute) @@ -300,12 +237,7 @@ lemma\<^marker>\tag unimportant\ orthogonal_transformation: "orthogonal_transformation f \ linear f \ (\v. norm (f v) = norm v)" - unfolding orthogonal_transformation_def - apply auto - apply (erule_tac x=v in allE)+ - apply (simp add: norm_eq_sqrt_inner) - apply (simp add: dot_norm linear_add[symmetric]) - done + by (smt (verit, ccfv_threshold) dot_norm linear_add norm_eq_sqrt_inner orthogonal_transformation_def) lemma\<^marker>\tag unimportant\ orthogonal_transformation_id [simp]: "orthogonal_transformation (\x. x)" by (simp add: linear_iff orthogonal_transformation_def) @@ -431,15 +363,8 @@ next fix h assume "\x y. inner (f x) y = inner x (h y)" - then have "\x y. inner x (g y) = inner x (h y)" - using assms by simp - then have "\x y. inner x (g y - h y) = 0" - by (simp add: inner_diff_right) - then have "\y. inner (g y - h y) (g y - h y) = 0" - by simp - then have "\y. h y = g y" - by simp - then show "h = g" by (simp add: ext) + then show "h = g" + by (metis assms ext vector_eq_ldot) qed text \TODO: The following lemmas about adjoints should hold for any @@ -524,32 +449,18 @@ proof - from Basis_le_norm[OF that, of x] show "norm (?g i) \ norm (f i) * norm x" - unfolding norm_scaleR by (metis mult.commute mult_left_mono norm_ge_zero) + unfolding norm_scaleR by (metis mult.commute mult_left_mono norm_ge_zero) qed from sum_norm_le[of _ ?g, OF th] show "norm (f x) \ ?B * norm x" - unfolding th0 sum_distrib_right by metis + by (simp add: sum_distrib_right th0) qed qed lemma linear_conv_bounded_linear: fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" shows "linear f \ bounded_linear f" -proof - assume "linear f" - then interpret f: linear f . - show "bounded_linear f" - proof - have "\B. \x. norm (f x) \ B * norm x" - using \linear f\ by (rule linear_bounded) - then show "\K. \x. norm (f x) \ norm x * K" - by (simp add: mult.commute) - qed -next - assume "bounded_linear f" - then interpret f: bounded_linear f . - show "linear f" .. -qed + by (metis mult.commute bounded_linear_axioms.intro bounded_linear_def linear_bounded) lemmas linear_linear = linear_conv_bounded_linear[symmetric] @@ -562,17 +473,11 @@ fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" assumes lf: "linear f" obtains B where "B > 0" "\x. norm (f x) \ B * norm x" -proof - - have "\B > 0. \x. norm (f x) \ norm x * B" - using lf unfolding linear_conv_bounded_linear - by (rule bounded_linear.pos_bounded) - with that show ?thesis - by (auto simp: mult.commute) -qed + by (metis bounded_linear.pos_bounded lf linear_linear mult.commute) lemma linear_invertible_bounded_below_pos: fixes f :: "'a::real_normed_vector \ 'b::euclidean_space" - assumes "linear f" "linear g" "g \ f = id" + assumes "linear f" "linear g" and gf: "g \ f = id" obtains B where "B > 0" "\x. B * norm x \ norm(f x)" proof - obtain B where "B > 0" and B: "\x. norm (g x) \ B * norm x" @@ -582,13 +487,8 @@ show "0 < 1/B" by (simp add: \B > 0\) show "1/B * norm x \ norm (f x)" for x - proof - - have "1/B * norm x = 1/B * norm (g (f x))" - using assms by (simp add: pointfree_idE) - also have "\ \ norm (f x)" - using B [of "f x"] by (simp add: \B > 0\ mult.commute pos_divide_le_eq) - finally show ?thesis . - qed + by (smt (verit, ccfv_SIG) B \0 < B\ gf comp_apply divide_inverse id_apply inverse_eq_divide + less_divide_eq mult.commute) qed qed @@ -663,15 +563,10 @@ fixes h :: "'a::euclidean_space \ 'b::euclidean_space \ 'c::real_normed_vector" assumes bh: "bilinear h" shows "\B > 0. \x y. norm (h x y) \ B * norm x * norm y" -proof - - have "\B > 0. \x y. norm (h x y) \ norm x * norm y * B" - using bh [unfolded bilinear_conv_bounded_bilinear] - by (rule bounded_bilinear.pos_bounded) - then show ?thesis - by (simp only: ac_simps) -qed + by (metis mult.assoc bh bilinear_conv_bounded_bilinear bounded_bilinear.pos_bounded mult.commute) -lemma bounded_linear_imp_has_derivative: "bounded_linear f \ (f has_derivative f) net" +lemma bounded_linear_imp_has_derivative: + "bounded_linear f \ (f has_derivative f) net" by (auto simp add: has_derivative_def linear_diff linear_linear linear_def dest: bounded_linear.linear) @@ -723,8 +618,7 @@ assumes "pairwise orthogonal S" and "\y. y \ S \ orthogonal x y" shows "pairwise orthogonal (insert x S)" - using assms unfolding pairwise_def - by (auto simp add: orthogonal_commute) + using assms by (auto simp: pairwise_def orthogonal_commute) lemma basis_orthogonal: fixes B :: "'a::real_inner set" @@ -735,9 +629,7 @@ proof (induct rule: finite_induct) case empty then show ?case - apply (rule exI[where x="{}"]) - apply (auto simp add: pairwise_def) - done + using pairwise_empty by blast next case (insert a B) note fB = \finite B\ and aB = \a \ B\ @@ -748,21 +640,15 @@ let ?C = "insert ?a C" from C(1) have fC: "finite ?C" by simp - from fB aB C(1,2) have cC: "card ?C \ card (insert a B)" - by (simp add: card_insert_if) + have cC: "card ?C \ card (insert a B)" + using C aB card_insert_if local.insert(1) by fastforce { fix x k have th0: "\(a::'a) b c. a - (b - c) = c + (a - b)" by (simp add: field_simps) have "x - k *\<^sub>R (a - (\x\C. (x \ a / (x \ x)) *\<^sub>R x)) \ span C \ x - k *\<^sub>R a \ span C" - apply (simp only: scaleR_right_diff_distrib th0) - apply (rule span_add_eq) - apply (rule span_scale) - apply (rule span_sum) - apply (rule span_scale) - apply (rule span_base) - apply assumption - done + unfolding scaleR_right_diff_distrib th0 + by (intro span_add_eq span_scale span_sum span_base) } then have SC: "span ?C = span (insert a B)" unfolding set_eq_iff span_breakdown_eq C(3)[symmetric] by auto @@ -773,15 +659,14 @@ by blast have fth: "finite (C - {y})" using C by simp - have "orthogonal ?a y" + have "y \ 0 \ \x\C - {y}. x \ a * (x \ y) / (x \ x) = 0" + using \pairwise orthogonal C\ + by (metis Cy DiffE div_0 insertCI mult_zero_right orthogonal_def pairwise_insert) + then have "orthogonal ?a y" unfolding orthogonal_def unfolding inner_diff inner_sum_left right_minus_eq unfolding sum.remove [OF \finite C\ \y \ C\] - apply (clarsimp simp add: inner_commute[of y a]) - apply (rule sum.neutral) - apply clarsimp - apply (rule C(4)[unfolded pairwise_def orthogonal_def, rule_format]) - using \y \ C\ by auto + by (auto simp add: sum.neutral inner_commute[of y a]) } with \pairwise orthogonal C\ have CPO: "pairwise orthogonal ?C" by (rule pairwise_orthogonal_insert) @@ -792,8 +677,7 @@ lemma orthogonal_basis_exists: fixes V :: "('a::euclidean_space) set" - shows "\B. independent B \ B \ span V \ V \ span B \ - (card B = dim V) \ pairwise orthogonal B" + shows "\B. independent B \ B \ span V \ V \ span B \ (card B = dim V) \ pairwise orthogonal B" proof - from basis_exists[of V] obtain B where B: "B \ V" "independent B" "V \ span B" "card B = dim V" @@ -807,18 +691,15 @@ by (metis span_superset span_mono subset_trans) from span_mono[OF B(3)] C have SVC: "span V \ span C" by (simp add: span_span) - from card_le_dim_spanning[OF CSV SVC C(1)] C(2,3) fB - have iC: "independent C" - by (simp) from C fB have "card C \ dim V" by simp moreover have "dim V \ card C" using span_card_ge_dim[OF CSV SVC C(1)] by simp - ultimately have CdV: "card C = dim V" + ultimately have "card C = dim V" using C(1) by simp - from C B CSV CdV iC show ?thesis - by auto + with C B CSV show ?thesis + by (metis SVC card_eq_dim dim_span) qed text \Low-dimensional subset is in a hyperplane (weak orthogonal complement).\ @@ -831,22 +712,15 @@ from sU obtain a where a: "a \ span S" by blast from orthogonal_basis_exists obtain B where - B: "independent B" "B \ span S" "S \ span B" - "card B = dim S" "pairwise orthogonal B" + B: "independent B" "B \ span S" "S \ span B" "card B = dim S" "pairwise orthogonal B" by blast from B have fB: "finite B" "card B = dim S" using independent_bound by auto - from span_mono[OF B(2)] span_mono[OF B(3)] have sSB: "span S = span B" - by (simp add: span_span) + by (simp add: B span_eq) let ?a = "a - sum (\b. (a \ b / (b \ b)) *\<^sub>R b) B" have "sum (\b. (a \ b / (b \ b)) *\<^sub>R b) B \ span S" - unfolding sSB - apply (rule span_sum) - apply (rule span_scale) - apply (rule span_base) - apply assumption - done + by (simp add: sSB span_base span_mul span_sum) with a have a0:"?a \ 0" by auto have "?a \ x = 0" if "x\span B" for x @@ -861,23 +735,19 @@ by blast have fth: "finite (B - {x})" using fB by simp - have "?a \ x = 0" + have "(\b\B - {x}. a \ b * (b \ x) / (b \ b)) = 0" if "x \ 0" + by (smt (verit) B' B(5) DiffD2 divide_eq_0_iff inner_real_def inner_zero_right insertCI orthogonal_def pairwise_insert sum.neutral) + then have "?a \ x = 0" apply (subst B') using fB fth unfolding sum_clauses(2)[OF fth] - apply simp unfolding inner_simps - apply (clarsimp simp add: inner_add inner_sum_left) - apply (rule sum.neutral, rule ballI) - apply (simp only: inner_commute) - apply (auto simp add: x field_simps - intro: B(5)[unfolded pairwise_def orthogonal_def, rule_format]) - done + by (auto simp add: inner_add_left inner_diff_left inner_sum_left) } then show "?a \ x = 0" if "x \ B" for x using that by blast qed - with a0 show ?thesis - unfolding sSB by (auto intro: exI[where x="?a"]) + with a0 sSB show ?thesis + by blast qed lemma span_not_univ_subset_hyperplane: @@ -890,19 +760,7 @@ fixes S :: "'a::euclidean_space set" assumes d: "dim S < DIM('a)" shows "\a::'a. a \ 0 \ span S \ {x. a \ x = 0}" -proof - - { - assume "span S = UNIV" - then have "dim (span S) = dim (UNIV :: ('a) set)" - by simp - then have "dim S = DIM('a)" - by (metis Euclidean_Space.dim_UNIV dim_span) - with d have False by arith - } - then have th: "span S \ UNIV" - by blast - from span_not_univ_subset_hyperplane[OF th] show ?thesis . -qed + using d dim_eq_full nless_le span_not_univ_subset_hyperplane by blast lemma linear_eq_stdbasis: fixes f :: "'a::euclidean_space \ _" @@ -910,8 +768,7 @@ and lg: "linear g" and fg: "\b. b \ Basis \ f b = g b" shows "f = g" - using linear_eq_on_span[OF lf lg, of Basis] fg - by auto + using linear_eq_on_span[OF lf lg, of Basis] fg by auto text \Similar results for bilinear functions.\ @@ -932,15 +789,9 @@ span_add Ball_def intro: bilinear_ladd[OF bf]) have sfg: "\x. x \ B \ subspace {a. f x a = g x a}" - apply (auto simp add: subspace_def) - using bf bg unfolding bilinear_def linear_iff - apply (auto simp add: span_zero bilinear_rzero[OF bf] bilinear_rzero[OF bg] - span_add Ball_def - intro: bilinear_ladd[OF bf]) - done + by (auto simp: subspace_def bf bg bilinear_rzero bilinear_radd bilinear_rmul) have "\y\ span C. f x y = g x y" if "x \ span B" for x - apply (rule span_induct [OF that sp]) - using fg sfg span_induct by blast + using span_induct [OF that sp] fg sfg span_induct by blast then show ?thesis using SB TC assms by auto qed @@ -972,8 +823,7 @@ fixes x :: "'a::euclidean_space" shows "finite {\x \ i\ |i. i \ Basis}" and "{\x \ i\ |i. i \ Basis} \ {}" - unfolding infnorm_set_image - by auto + unfolding infnorm_set_image by auto lemma infnorm_pos_le: fixes x :: "'a::euclidean_space" @@ -1010,16 +860,7 @@ by (metis infnorm_neg minus_diff_eq) lemma absdiff_infnorm: "\infnorm x - infnorm y\ \ infnorm (x - y)" -proof - - have *: "\(nx::real) n ny. nx \ n + ny \ ny \ n + nx \ \nx - ny\ \ n" - by arith - show ?thesis - proof (rule *) - from infnorm_triangle[of "x - y" " y"] infnorm_triangle[of "x - y" "-x"] - show "infnorm x \ infnorm (x - y) + infnorm y" "infnorm y \ infnorm (x - y) + infnorm x" - by (simp_all add: field_simps infnorm_neg) - qed -qed + by (smt (verit, del_insts) diff_add_cancel infnorm_sub infnorm_triangle) lemma real_abs_infnorm: "\infnorm x\ = infnorm x" using infnorm_pos_le[of x] by arith @@ -1065,9 +906,9 @@ by (simp add: zero_le_mult_iff infnorm_pos_le) have "x \ x \ (\b\Basis. x \ b * (x \ b))" by (metis euclidean_inner order_refl) - also have "... \ DIM('a) * \infnorm x\\<^sup>2" + also have "\ \ DIM('a) * \infnorm x\\<^sup>2" by (rule sum_bounded_above) (metis Basis_le_infnorm abs_le_square_iff power2_eq_square real_abs_infnorm) - also have "... \ (sqrt DIM('a) * infnorm x)\<^sup>2" + also have "\ \ (sqrt DIM('a) * infnorm x)\<^sup>2" by (simp add: power_mult_distrib) finally show "x \ x \ (sqrt DIM('a) * infnorm x)\<^sup>2" . qed @@ -1091,11 +932,11 @@ then show ?thesis by auto next - case False + case False from inner_eq_zero_iff[of "norm y *\<^sub>R x - norm x *\<^sub>R y"] have "?rhs \ (norm y * (norm y * norm x * norm x - norm x * (x \ y)) - - norm x * (norm y * (y \ x) - norm x * norm y * norm y) = 0)" + norm x * (norm y * (y \ x) - norm x * norm y * norm y) = 0)" using False unfolding inner_simps by (auto simp add: power2_norm_eq_inner[symmetric] power2_eq_square inner_commute field_simps) also have "\ \ (2 * norm x * norm y * (norm x * norm y - x \ y) = 0)" @@ -1108,20 +949,9 @@ lemma norm_cauchy_schwarz_abs_eq: "\x \ y\ = norm x * norm y \ norm x *\<^sub>R y = norm y *\<^sub>R x \ norm x *\<^sub>R y = - norm y *\<^sub>R x" - (is "?lhs \ ?rhs") -proof - - have th: "\(x::real) a. a \ 0 \ \x\ = a \ x = a \ x = - a" - by arith - have "?rhs \ norm x *\<^sub>R y = norm y *\<^sub>R x \ norm (- x) *\<^sub>R y = norm y *\<^sub>R (- x)" - by simp - also have "\ \ (x \ y = norm x * norm y \ (- x) \ y = norm x * norm y)" - unfolding norm_cauchy_schwarz_eq[symmetric] - unfolding norm_minus_cancel norm_scaleR .. - also have "\ \ ?lhs" - unfolding th[OF mult_nonneg_nonneg, OF norm_ge_zero[of x] norm_ge_zero[of y]] inner_simps - by auto - finally show ?thesis .. -qed + using norm_cauchy_schwarz_eq [symmetric, of x y] + using norm_cauchy_schwarz_eq [symmetric, of "-x" y] Cauchy_Schwarz_ineq2 [of x y] + by auto lemma norm_triangle_eq: fixes x y :: "'a::real_inner" @@ -1137,9 +967,7 @@ have "norm (x + y) = norm x + norm y \ (norm (x + y))\<^sup>2 = (norm x + norm y)\<^sup>2" by simp also have "\ \ norm x *\<^sub>R y = norm y *\<^sub>R x" - unfolding norm_cauchy_schwarz_eq[symmetric] - unfolding power2_norm_eq_inner inner_simps - by (simp add: power2_norm_eq_inner[symmetric] power2_eq_square inner_commute field_simps) + by (smt (verit, best) dot_norm inner_real_def inner_simps norm_cauchy_schwarz_eq power2_eq_square) finally show ?thesis . qed @@ -1147,11 +975,7 @@ fixes x y z :: "'a::real_inner" shows "dist x z = dist x y + dist y z \ norm (x - y) *\<^sub>R (y - z) = norm (y - z) *\<^sub>R (x - y)" -proof - - have *: "x - y + (y - z) = x - z" by auto - show ?thesis unfolding dist_norm norm_triangle_eq[of "x - y" "y - z", unfolded *] - by (auto simp:norm_minus_commute) -qed + by (metis (no_types, lifting) add_diff_eq diff_add_cancel dist_norm norm_triangle_eq) subsection \Collinearity\ @@ -1163,7 +987,7 @@ proof assume ?lhs then show ?rhs - unfolding collinear_def by (metis Groups.add_ac(2) diff_add_cancel) + unfolding collinear_def by (metis add.commute diff_add_cancel) next assume ?rhs then obtain u v where *: "\x. x \ S \ \c. x = u + c *\<^sub>R v" @@ -1192,8 +1016,7 @@ using \v \ 0\ by blast qed then show ?thesis - apply (clarsimp simp: collinear_def) - by (metis scaleR_zero_right vector_fraction_eq_iff) + by (metis collinear_def) qed lemma collinear_subset: "\collinear T; S \ T\ \ collinear S" @@ -1206,9 +1029,7 @@ by (simp add: collinear_def) lemma collinear_2 [iff]: "collinear {x, y}" - apply (simp add: collinear_def) - apply (rule exI[where x="x - y"]) - by (metis minus_diff_eq scaleR_left.minus scaleR_one) + by (simp add: collinear_def) (metis minus_diff_eq scaleR_left.minus scaleR_one) lemma collinear_lemma: "collinear {0, x, y} \ x = 0 \ y = 0 \ (\c. y = c *\<^sub>R x)" (is "?lhs \ ?rhs") @@ -1237,15 +1058,8 @@ then obtain c where c: "y = c *\<^sub>R x" using False by blast show ?lhs - unfolding collinear_def c - apply (rule exI[where x=x]) - apply auto - apply (rule exI[where x="- 1"], simp) - apply (rule exI[where x= "-c"], simp) - apply (rule exI[where x=1], simp) - apply (rule exI[where x="1 - c"], simp add: scaleR_left_diff_distrib) - apply (rule exI[where x="c - 1"], simp add: scaleR_left_diff_distrib) - done + apply (simp add: collinear_def c) + by (metis (mono_tags, lifting) scaleR_left.minus scaleR_left_diff_distrib scaleR_one) qed qed @@ -1291,11 +1105,8 @@ fixes x y :: "'a::real_inner" assumes "norm (x + y) = norm x + norm y" shows "collinear{0,x,y}" -proof (cases "x = 0 \ y = 0") - case False - with assms show ?thesis - by (meson norm_cauchy_schwarz_abs_eq norm_cauchy_schwarz_equal norm_triangle_eq) -qed (use collinear_lemma in blast) + using assms norm_cauchy_schwarz_abs_eq norm_cauchy_schwarz_equal norm_triangle_eq + by blast subsection\Properties of special hyperplanes\ @@ -1315,7 +1126,7 @@ proof - have "x = (\b\Basis. (x \ b) *\<^sub>R b)" by (simp add: euclidean_representation) - also have "... = (\b \ Basis - {k}. (x \ b) *\<^sub>R b)" + also have "\ = (\b \ Basis - {k}. (x \ b) *\<^sub>R b)" by (auto simp: sum.remove [of _ k] inner_commute assms that) finally have "x = (\b\Basis - {k}. (x \ b) *\<^sub>R b)" . then show ?thesis @@ -1331,11 +1142,7 @@ lemma dim_special_hyperplane: fixes k :: "'n::euclidean_space" shows "k \ Basis \ dim {x. k \ x = 0} = DIM('n) - 1" -apply (simp add: special_hyperplane_span) -apply (rule dim_unique [OF subset_refl]) -apply (auto simp: independent_substdbasis) -apply (metis member_remove remove_def span_base) -done + by (metis Diff_subset card_Diff_singleton indep_card_eq_dim_span independent_substdbasis special_hyperplane_span) proposition dim_hyperplane: fixes a :: "'a::euclidean_space" @@ -1358,20 +1165,17 @@ using \independent B\ independent_bound by blast have "UNIV \ span (insert a B)" proof fix y::'a - obtain r z where z: "y = r *\<^sub>R a + z" "a \ z = 0" - apply (rule_tac r="(a \ y) / (a \ a)" and z = "y - ((a \ y) / (a \ a)) *\<^sub>R a" in that) - using assms - by (auto simp: algebra_simps) - show "y \ span (insert a B)" - by (metis (mono_tags, lifting) z Bsub span_eq_iff - add_diff_cancel_left' mem_Collect_eq span0 span_breakdown_eq span_subspace subspB) + obtain r z where "y = r *\<^sub>R a + z" "a \ z = 0" + by (metis add.commute diff_add_cancel vector_sub_project_orthogonal) + then show "y \ span (insert a B)" + by (metis (mono_tags, lifting) Bsub add_diff_cancel_left' + mem_Collect_eq span0 span_breakdown_eq span_eq subspB) qed - then have dima: "DIM('a) = dim(insert a B)" + then have "DIM('a) = dim(insert a B)" by (metis independent_Basis span_Basis dim_eq_card top.extremum_uniqueI) then show ?thesis - by (metis (mono_tags, lifting) Bsub Diff_insert_absorb \a \ span B\ ind card0 - card_Diff_singleton dim_span indep_card_eq_dim_span insertI1 subsetCE - subspB) + by (metis One_nat_def \a \ span B\ \finite B\ card0 card_insert_disjoint + diff_Suc_Suc diff_zero dim_eq_card_independent ind span_base) qed lemma lowdim_eq_hyperplane: @@ -1379,14 +1183,10 @@ assumes "dim S = DIM('a) - 1" obtains a where "a \ 0" and "span S = {x. a \ x = 0}" proof - - have dimS: "dim S < DIM('a)" - by (simp add: assms) - then obtain b where b: "b \ 0" "span S \ {a. b \ a = 0}" - using lowdim_subset_hyperplane [of S] by fastforce - show ?thesis - apply (rule that[OF b(1)]) - apply (rule subspace_dim_equal) - by (auto simp: assms b dim_hyperplane subspace_hyperplane) + obtain b where b: "b \ 0" "span S \ {a. b \ a = 0}" + by (metis DIM_positive assms diff_less zero_less_one lowdim_subset_hyperplane) + then show ?thesis + by (metis assms dim_hyperplane dim_span dim_subset subspace_dim_equal subspace_hyperplane subspace_span that) qed lemma dim_eq_hyperplane: @@ -1409,10 +1209,9 @@ using a by (force simp: span_explicit) then have "a \ a = a \ (\v\T. U v *\<^sub>R v)" by simp - also have "... = 0" + also have "\ = 0" apply (simp add: inner_sum_right) - apply (rule comm_monoid_add_class.sum.neutral) - by (metis "0" DiffE \T \ S - {a}\ mult_not_zero singletonI subsetCE \a \ S\) + by (smt (verit) "0" DiffE \T \ S - {a}\ in_mono insertCI mult_not_zero sum.neutral that(1)) finally show ?thesis using \0 \ S\ \a \ S\ by auto qed @@ -1424,14 +1223,8 @@ fixes S :: "'a::euclidean_space set" assumes "pairwise orthogonal S" shows "finite S" -proof - - have "independent (S - {0})" - apply (rule pairwise_orthogonal_independent) - apply (metis Diff_iff assms pairwise_def) - by blast - then show ?thesis - by (meson independent_imp_finite infinite_remove) -qed + by (metis Set.set_insert assms finite_insert independent_bound pairwise_insert + pairwise_orthogonal_independent) lemma subspace_orthogonal_to_vector: "subspace {y. orthogonal x y}" by (simp add: subspace_def orthogonal_clauses) @@ -1457,7 +1250,7 @@ proof - have "a \ x = (\y\S. if y = x then y \ a else 0)" by (simp add: \finite S\ inner_commute that) - also have "... = (\b\S. b \ a * (b \ x) / (b \ b))" + also have "\ = (\b\S. b \ a * (b \ x) / (b \ b))" apply (rule sum.cong [OF refl], simp) by (meson S orthogonal_def pairwise_def that) finally show ?thesis @@ -1487,20 +1280,15 @@ (auto simp: Gram_Schmidt_step a'_def insert.prems orthogonal_commute pairwise_orthogonal_insert span_clauses) have orthS: "\x. x \ S \ a' \ x = 0" - apply (simp add: a'_def) - using Gram_Schmidt_step [OF \pairwise orthogonal S\] - apply (force simp: orthogonal_def inner_commute span_superset [THEN subsetD]) - done + using Gram_Schmidt_step a'_def insert.prems orthogonal_commute orthogonal_def span_base by blast have "span (S \ insert a' U) = span (insert a' (S \ T))" using spanU by simp - also have "... = span (insert a (S \ T))" - apply (rule eq_span_insert_eq) - apply (simp add: a'_def span_neg span_sum span_base span_mul) - done - also have "... = span (S \ insert a T)" + also have "\ = span (insert a (S \ T))" + by (simp add: a'_def span_neg span_sum span_base span_mul eq_span_insert_eq) + also have "\ = span (S \ insert a T)" by simp finally show ?case - by (rule_tac x="insert a' U" in exI) (use orthU in auto) + using orthU by blast qed @@ -1524,14 +1312,12 @@ obtains U where "U \ (insert 0 S) = {}" "pairwise orthogonal (S \ U)" "span (S \ U) = span (S \ T)" proof - - obtain U where "pairwise orthogonal (S \ U)" "span (S \ U) = span (S \ T)" + obtain U where U: "pairwise orthogonal (S \ U)" "span (S \ U) = span (S \ T)" using orthogonal_extension assms by blast - then show ?thesis - apply (rule_tac U = "U - (insert 0 S)" in that) - apply blast - apply (force simp: pairwise_def) - apply (metis Un_Diff_cancel Un_insert_left span_redundant span_zero) - done + moreover have "pairwise orthogonal (S \ (U - insert 0 S))" + by (smt (verit, best) Un_Diff_Int Un_iff U pairwise_def) + ultimately show ?thesis + by (metis Diff_disjoint Un_Diff_cancel Un_insert_left inf_commute span_insert_0 that) qed subsection\Decomposing a vector into parts in orthogonal subspaces\ @@ -1542,27 +1328,14 @@ fixes S :: "'a :: euclidean_space set" assumes "subspace S" obtains B where "B \ S" "pairwise orthogonal B" "span B = S" -proof - - obtain B where "B \ S" "independent B" "S \ span B" "card B = dim S" - using basis_exists by blast - with orthogonal_extension [of "{}" B] - show ?thesis - by (metis Un_empty_left assms pairwise_empty span_superset span_subspace that) -qed + by (metis assms basis_orthogonal basis_subspace_exists span_eq) lemma orthogonal_basis_subspace: fixes S :: "'a :: euclidean_space set" assumes "subspace S" obtains B where "0 \ B" "B \ S" "pairwise orthogonal B" "independent B" "card B = dim S" "span B = S" -proof - - obtain B where "B \ S" "pairwise orthogonal B" "span B = S" - using assms orthogonal_spanningset_subspace by blast - then show ?thesis - apply (rule_tac B = "B - {0}" in that) - apply (auto simp: indep_card_eq_dim_span pairwise_subset pairwise_orthogonal_independent elim: pairwise_subset) - done -qed + by (metis assms dependent_zero orthogonal_basis_exists span_eq span_eq_iff) proposition orthonormal_basis_subspace: fixes S :: "'a :: euclidean_space set" @@ -1609,7 +1382,7 @@ obtain B where "B \ span S" and orthB: "pairwise orthogonal B" and "\x. x \ B \ norm x = 1" and "independent B" "card B = dim S" "span B = span S" - by (rule orthonormal_basis_subspace [of "span S", OF subspace_span]) (auto) + by (metis dim_span orthonormal_basis_subspace subspace_span) with assms obtain u where spanBT: "span B \ span T" and "u \ span B" "u \ span T" by auto obtain C where orthBC: "pairwise orthogonal (B \ C)" and spanBC: "span (B \ C) = span (B \ {u})" @@ -1630,9 +1403,8 @@ then obtain x where "x \ C" "x \ 0" "x \ B" by blast then have "x \ span T" - by (metis (no_types, lifting) Un_insert_right Un_upper2 \u \ span T\ spanBT spanBC - \u \ span T\ insert_subset span_superset span_mono - span_span subsetCE subset_trans sup_bot.comm_neutral) + by (smt (verit, ccfv_SIG) Set.set_insert \u \ span T\ empty_subsetI insert_subset + le_sup_iff spanBC spanBT span_mono span_span span_superset subset_trans) moreover have "orthogonal x y" if "y \ span B" for y using that proof (rule span_induct) @@ -1652,8 +1424,7 @@ obtains x where "x \ 0" "\y. y \ span S \ orthogonal x y" proof - have "span S \ UNIV" - by (metis (mono_tags) UNIV_I assms inner_eq_zero_iff less_le lowdim_subset_hyperplane - mem_Collect_eq top.extremum_strict top.not_eq_extremum) + by (metis assms dim_eq_full order_less_imp_not_less top.not_eq_extremum) with orthogonal_to_subspace_exists_gen [of S UNIV] that show ?thesis by (auto) qed @@ -1683,11 +1454,8 @@ have orth: "orthogonal (x - ?a) w" if "w \ span S" for w by (simp add: Gram_Schmidt_step \pairwise orthogonal T\ \span T = span S\ orthogonal_commute that) - show ?thesis - apply (rule_tac y = "?a" and z = "x - ?a" in that) - apply (meson \T \ span S\ span_scale span_sum subsetCE) - apply (fact orth, simp) - done + with that[of ?a "x-?a"] \T \ span S\ show ?thesis + by (simp add: span_mul span_sum subsetD) qed lemma orthogonal_subspace_decomp_unique: @@ -1702,15 +1470,15 @@ moreover have "\a b. \a \ span S; b \ span T\ \ orthogonal a b" by (meson orth orthogonal_commute orthogonal_to_span) ultimately have "0 = x' - x" - by (metis (full_types) add_diff_cancel_left' ST diff_right_commute orthogonal_clauses(10) orthogonal_clauses(5) orthogonal_self) + using assms + by (metis add.commute add_diff_cancel_right' diff_right_commute orthogonal_self span_diff) with assms show ?thesis by auto qed lemma vector_in_orthogonal_spanningset: fixes a :: "'a::euclidean_space" obtains S where "a \ S" "pairwise orthogonal S" "span S = UNIV" - by (metis UNIV_I Un_iff empty_iff insert_subset orthogonal_extension pairwise_def - pairwise_orthogonal_insert span_UNIV subsetI subset_antisym) + by (metis UnI1 Un_UNIV_right insertI1 orthogonal_extension pairwise_singleton span_UNIV) lemma vector_in_orthogonal_basis: fixes a :: "'a::euclidean_space" @@ -1755,17 +1523,15 @@ using \pairwise orthogonal S\ by (auto simp: pairwise_def orthogonal_def) show "\x. x \ (\x. x /\<^sub>R norm x) ` S \ norm x = 1" using \0 \ S\ by (auto simp: field_split_simps) - then show "independent ?S" + then show ind: "independent ?S" by (metis \pairwise orthogonal ((\x. x /\<^sub>R norm x) ` S)\ norm_zero pairwise_orthogonal_independent zero_neq_one) have "inj_on (\x. x /\<^sub>R norm x) S" unfolding inj_on_def by (metis (full_types) S(1) \0 \ S\ inverse_nonzero_iff_nonzero norm_eq_zero orthogonal_scaleR orthogonal_self pairwise_def) then show "card ?S = DIM('a)" by (simp add: card_image S) - show "span ?S = UNIV" - by (metis (no_types) \0 \ S\ \finite S\ \span S = UNIV\ - field_class.field_inverse_zero inverse_inverse_eq less_irrefl span_image_scale - zero_less_norm_iff) + then show "span ?S = UNIV" + by (metis ind dim_eq_card dim_eq_full) qed qed @@ -1786,12 +1552,10 @@ by (auto simp add: span_Un image_def) also have "dim \ = dim {x + y |x y. x \ span A \ y \ span B}" by (auto intro!: arg_cong [where f=dim]) - also have "... = dim {x + y |x y. x \ span A \ y \ span B} + dim(span A \ span B)" - by (auto simp: dest: 0) - also have "... = dim (span A) + dim (span B)" - by (rule dim_sums_Int) (auto) - also have "... = dim A + dim B" - by (simp) + also have "\ = dim {x + y |x y. x \ span A \ y \ span B} + dim(span A \ span B)" + by (auto dest: 0) + also have "\ = dim A + dim B" + using dim_sums_Int by fastforce finally show ?thesis . qed @@ -1810,17 +1574,16 @@ proof - obtain y z where "x = y + z" "y \ span A" and orth: "\w. w \ span A \ orthogonal z w" using orthogonal_subspace_decomp_exists [of A x] that by auto + moreover have "y \ span B" using \y \ span A\ assms(3) span_mono by blast - then have "z \ {a \ B. \x. x \ A \ orthogonal x a}" - apply simp - using \x = y + z\ assms(1) assms(2) orth orthogonal_commute span_add_eq - span_eq_iff that by blast + ultimately have "z \ B \ (\x. x \ A \ orthogonal x z)" + using assms by (metis orthogonal_commute span_add_eq span_eq_iff that) then have z: "z \ span {y \ B. \x\A. orthogonal x y}" - by (meson span_superset subset_iff) + by (simp add: span_base) then show ?thesis - apply (auto simp: span_Un image_def \x = y + z\ \y \ span A\) - using \y \ span A\ add.commute by blast + by (smt (verit, best) \x = y + z\ \y \ span A\ le_sup_iff span_add_eq span_subspace_induct + span_superset subset_iff subspace_span) qed show "span B \ span ({y \ B. \x\A. orthogonal x y} \ A)" by (rule span_minimal) (auto intro: * span_minimal) @@ -1845,12 +1608,8 @@ qed lemma linear_continuous_at: - assumes "bounded_linear f" - shows "continuous (at a) f" - unfolding continuous_at using assms - apply (rule bounded_linear.tendsto) - apply (rule tendsto_ident_at) - done + "bounded_linear f \continuous (at a) f" + by (simp add: bounded_linear.isUCont isUCont_isCont) lemma linear_continuous_within: "bounded_linear f \ continuous (at x within s) f" @@ -1869,21 +1628,7 @@ using linear_bounded_pos [OF \linear h\] by blast show ?thesis unfolding tendsto_iff - proof (intro allI impI) - show "\\<^sub>F x in F. dist (h (f x)) (h l) < e" if "e > 0" for e - proof - - have "\\<^sub>F x in F. dist (f x) l < e/B" - by (simp add: \0 < B\ assms(1) tendstoD that) - then show ?thesis - unfolding dist_norm - proof (rule eventually_mono) - show "norm (h (f x) - h l) < e" if "norm (f x - l) < e / B" for x - using that B - apply (simp add: field_split_simps) - by (metis \linear h\ le_less_trans linear_diff) - qed - qed - qed + by (simp add: assms bounded_linear.tendsto linear_linear tendstoD) qed lemma linear_continuous_compose: diff -r 8d8af7e92c5e -r 30182f9e1818 src/HOL/Analysis/Path_Connected.thy --- a/src/HOL/Analysis/Path_Connected.thy Sat Dec 31 11:09:19 2022 +0000 +++ b/src/HOL/Analysis/Path_Connected.thy Sun Jan 01 00:45:55 2023 +0000 @@ -2497,7 +2497,7 @@ by (simp add: inner_commute) qed obtain S :: "'a set" where "S \ Basis" and "card S = Suc (Suc 0)" - using ex_card[OF assms] + using obtain_subset_with_card_n[OF assms] by auto then obtain b0 b1 :: 'a where "b0 \ Basis" and "b1 \ Basis" and "b0 \ b1" unfolding card_Suc_eq by auto diff -r 8d8af7e92c5e -r 30182f9e1818 src/HOL/Analysis/Topology_Euclidean_Space.thy --- a/src/HOL/Analysis/Topology_Euclidean_Space.thy Sat Dec 31 11:09:19 2022 +0000 +++ b/src/HOL/Analysis/Topology_Euclidean_Space.thy Sun Jan 01 00:45:55 2023 +0000 @@ -1307,18 +1307,11 @@ assume i: "i \ Basis" have "dist (x - (e / 2) *\<^sub>R i) x < e" and "dist (x + (e / 2) *\<^sub>R i) x < e" - unfolding dist_norm - apply auto - unfolding norm_minus_cancel - using norm_Basis[OF i] \e>0\ - apply auto - done + using norm_Basis[OF i] \e>0\ by (auto simp: dist_norm) then have "a \ i \ (x - (e / 2) *\<^sub>R i) \ i" and "(x + (e / 2) *\<^sub>R i) \ i \ b \ i" using e[THEN spec[where x="x - (e/2) *\<^sub>R i"]] and e[THEN spec[where x="x + (e/2) *\<^sub>R i"]] - unfolding mem_box - using i - by blast+ + unfolding mem_box using i by blast+ then have "a \ i < x \ i" and "x \ i < b \ i" using \e>0\ i by (auto simp: inner_diff_left inner_Basis inner_add_left) @@ -1349,8 +1342,7 @@ lemma bounded_box [simp]: fixes a :: "'a::euclidean_space" shows "bounded (box a b)" - using bounded_cbox[of a b] box_subset_cbox[of a b] bounded_subset[of "cbox a b" "box a b"] - by simp + by (metis bounded_cbox bounded_interior interior_cbox) lemma not_interval_UNIV [simp]: fixes a :: "'a::euclidean_space" @@ -1385,12 +1377,7 @@ have "a \ i = e * (a \ i) + (1 - e) * (a \ i)" unfolding left_diff_distrib by simp also have "\ < e * (x \ i) + (1 - e) * (y \ i)" - proof (rule add_less_le_mono) - show "e * (a \ i) < e * (x \ i)" - using \0 < e\ i mem_box(1) x by auto - show "(1 - e) * (a \ i) \ (1 - e) * (y \ i)" - by (meson diff_ge_0_iff_ge \e \ 1\ i mem_box(2) mult_left_mono y) - qed + by (smt (verit, best) e i mem_box mult_le_cancel_left_pos mult_left_mono x y) finally have "a \ i < (e *\<^sub>R x + (1 - e) *\<^sub>R y) \ i" unfolding inner_simps by auto moreover @@ -1398,12 +1385,7 @@ have "b \ i = e * (b\i) + (1 - e) * (b\i)" unfolding left_diff_distrib by simp also have "\ > e * (x \ i) + (1 - e) * (y \ i)" - proof (rule add_less_le_mono) - show "e * (x \ i) < e * (b \ i)" - using \0 < e\ i mem_box(1) x by auto - show "(1 - e) * (y \ i) \ (1 - e) * (b \ i)" - by (meson diff_ge_0_iff_ge \e \ 1\ i mem_box(2) mult_left_mono y) - qed + by (smt (verit, best) e i mem_box mult_le_cancel_left_pos mult_left_mono x y) finally have "(e *\<^sub>R x + (1 - e) *\<^sub>R y) \ i < b \ i" unfolding inner_simps by auto } @@ -1427,7 +1409,7 @@ let ?c = "(1 / 2) *\<^sub>R (a + b)" { fix x - assume as:"x \ cbox a b" + assume as: "x \ cbox a b" define f where [abs_def]: "f n = x + (inverse (real n + 1)) *\<^sub>R (?c - x)" for n { fix n @@ -1445,22 +1427,16 @@ } moreover { - assume "\ (f \ x) sequentially" - { - fix e :: real - assume "e > 0" - then obtain N :: nat where N: "inverse (real (N + 1)) < e" - using reals_Archimedean by auto - have "inverse (real n + 1) < e" if "N \ n" for n - by (auto intro!: that le_less_trans [OF _ N]) - then have "\N::nat. \n\N. inverse (real n + 1) < e" by auto - } - then have "((\n. inverse (real n + 1)) \ 0) sequentially" + have "\N::nat. \n\N. inverse (real n + 1) < \" if "\ > 0" for \ + using reals_Archimedean [of \] that + by (metis inverse_inverse_eq inverse_less_imp_less nat_le_real_less order_less_trans + reals_Archimedean2) + then have "(\n. inverse (real n + 1)) \ 0" unfolding lim_sequentially by(auto simp: dist_norm) - then have "(f \ x) sequentially" + then have "f \ x" unfolding f_def - using tendsto_add[OF tendsto_const, of "\n::nat. (inverse (real n + 1)) *\<^sub>R ((1 / 2) *\<^sub>R (a + b) - x)" 0 sequentially x] - using tendsto_scaleR [OF _ tendsto_const, of "\n::nat. inverse (real n + 1)" 0 sequentially "((1 / 2) *\<^sub>R (a + b) - x)"] + using tendsto_add[OF tendsto_const, of "\n. (inverse (real n + 1)) *\<^sub>R ((1 / 2) *\<^sub>R (a + b) - x)" 0 sequentially x] + using tendsto_scaleR [OF _ tendsto_const, of "\n. inverse (real n + 1)" 0 sequentially "((1 / 2) *\<^sub>R (a + b) - x)"] by auto } ultimately have "x \ closure (box a b)" @@ -1491,12 +1467,7 @@ fixes S :: "('a::euclidean_space) set" assumes "bounded S" obtains a where "S \ cbox (-a) a" -proof - - obtain a where "S \ box (-a) a" - using bounded_subset_box_symmetric[OF assms] by auto - then show ?thesis - by (meson box_subset_cbox dual_order.trans that) -qed + by (meson assms bounded_subset_box_symmetric box_subset_cbox order.trans) lemma frontier_cbox: fixes a b :: "'a::euclidean_space" @@ -1506,16 +1477,7 @@ lemma frontier_box: fixes a b :: "'a::euclidean_space" shows "frontier (box a b) = (if box a b = {} then {} else cbox a b - box a b)" -proof (cases "box a b = {}") - case True - then show ?thesis - using frontier_empty by auto -next - case False - then show ?thesis - unfolding frontier_def and closure_box[OF False] and interior_open[OF open_box] - by auto -qed + by (simp add: frontier_def interior_open open_box) lemma Int_interval_mixed_eq_empty: fixes a :: "'a::euclidean_space" @@ -1549,28 +1511,21 @@ with l have "eventually (\n. \i\Basis. dist (f (r n) \ i) (l \ i) < e / (real_of_nat DIM('a))) sequentially" by simp moreover - { - fix n + { fix n assume n: "\i\Basis. dist (f (r n) \ i) (l \ i) < e / (real_of_nat DIM('a))" have "dist (f (r n)) l \ (\i\Basis. dist (f (r n) \ i) (l \ i))" - apply (subst euclidean_dist_l2) - using zero_le_dist - apply (rule L2_set_le_sum) - done + using L2_set_le_sum [OF zero_le_dist] by (subst euclidean_dist_l2) also have "\ < (\i\(Basis::'a set). e / (real_of_nat DIM('a)))" - apply (rule sum_strict_mono) - using n - apply auto - done + by (meson eucl.finite_Basis n nonempty_Basis sum_strict_mono) finally have "dist (f (r n)) l < e" by auto } - ultimately have "eventually (\n. dist (f (r n)) l < e) sequentially" + ultimately have "\\<^sub>F n in sequentially. dist (f (r n)) l < e" by (rule eventually_mono) } - then have *: "((f \ r) \ l) sequentially" + then have *: "(f \ r) \ l" unfolding o_def tendsto_iff by simp - with r show "\l r. strict_mono r \ ((f \ r) \ l) sequentially" + with r show "\l r. strict_mono r \ (f \ r) \ l" by auto qed @@ -1592,10 +1547,8 @@ fix A::"'a set" assume "open A" show "\B'\B. \B' = A" - apply (rule exI[of _ "{b\B. b \ A}"]) - apply (subst (3) open_UNION_box[OF \open A\]) - apply (auto simp: a b B_def) - done + using open_UNION_box[OF \open A\] + by (smt (verit, ccfv_threshold) B_def a b image_iff mem_Collect_eq subsetI) qed ultimately have "topological_basis B" @@ -1658,10 +1611,11 @@ by (simp add: False b cSUP_least) finally have bi: "x \ i \ b \ i" . show "x \ i \ (\x. x \ i) ` S" - apply (rule_tac x="\j\Basis. (if j = i then x \ i else a \ j) *\<^sub>R j" in image_eqI) + apply (rule_tac x="\j\Basis. (((\)a)(i := x \ j))j *\<^sub>R j" in image_eqI) apply (simp add: i) apply (rule mem_is_intervalI [OF \is_interval S\ \a \ S\ \b \ S\]) - using i ai bi apply force + using i ai bi + apply force done qed have "S = cbox a b" @@ -1714,10 +1668,7 @@ assume ?lhs then show ?rhs unfolding tendsto_def - apply clarify - apply (drule_tac x="{s. s \ i \ S}" in spec) - apply (auto simp: open_preimage_inner) - done + by (smt (verit) eventually_elim2 mem_Collect_eq open_preimage_inner) next assume R: ?rhs then have "\e. e > 0 \ \i\Basis. \\<^sub>F x in F. dist (f x \ i) (l \ i) < e" @@ -1735,19 +1686,15 @@ have "L2_set (\i. dist (f x \ i) (l \ i)) Basis \ sum (\i. dist (f x \ i) (l \ i)) Basis" by (simp add: L2_set_le_sum) also have "... < DIM('b) * (e / real DIM('b))" - apply (rule sum_bounded_above_strict) - using that by auto + by (meson DIM_positive sum_bounded_above_strict that) also have "... = e" by (simp add: field_simps) finally show "L2_set (\i. dist (f x \ i) (l \ i)) Basis < e" . qed have "\\<^sub>F x in F. \i\Basis. dist (f x \ i) (l \ i) < e / DIM('b)" - apply (rule R') - using \0 < e\ by simp + by (simp add: R' \0 < e\) then show "\\<^sub>F x in F. dist (f x) l < e" - apply (rule eventually_mono) - apply (subst euclidean_dist_l2) - using * by blast + by eventually_elim (metis (full_types) "*" euclidean_dist_l2) qed qed @@ -1759,22 +1706,21 @@ corollary continuous_on_componentwise: fixes S :: "'a :: t2_space set" shows "continuous_on S f \ (\i \ Basis. continuous_on S (\x. (f x \ i)))" - apply (simp add: continuous_on_eq_continuous_within) - using continuous_componentwise by blast + by (metis continuous_componentwise continuous_on_eq_continuous_within) lemma linear_componentwise_iff: - "(linear f') \ (\i\Basis. linear (\x. f' x \ i))" - apply (auto simp: linear_iff inner_left_distrib) - apply (metis inner_left_distrib euclidean_eq_iff) - by (metis euclidean_eqI inner_scaleR_left) + "linear f' \ (\i\Basis. linear (\x. f' x \ i))" (is "?lhs \ ?rhs") +proof + show "?lhs \ ?rhs" + by (simp add: Real_Vector_Spaces.linear_iff inner_left_distrib) + show "?rhs \ ?lhs" + by (simp add: linear_iff) (metis euclidean_eqI inner_left_distrib inner_scaleR_left) +qed lemma bounded_linear_componentwise_iff: "(bounded_linear f') \ (\i\Basis. bounded_linear (\x. f' x \ i))" (is "?lhs = ?rhs") proof - assume ?lhs then show ?rhs - by (simp add: bounded_linear_inner_left_comp) -next assume ?rhs then have "(\i\Basis. \K. \x. \f' x \ i\ \ norm x * K)" "linear f'" by (auto simp: bounded_linear_def bounded_linear_axioms_def linear_componentwise_iff [symmetric] ball_conj_distrib) @@ -1792,7 +1738,7 @@ qed then show ?lhs by (force simp: bounded_linear_def bounded_linear_axioms_def \linear f'\) -qed +qed (simp add: bounded_linear_inner_left_comp) subsection\<^marker>\tag unimportant\ \Continuous Extension\ @@ -1853,8 +1799,7 @@ obtain d where d: "0 < d" "\x'. x' \ cbox a b \ dist x' (clamp a b x) < d \ dist (f x') (f (clamp a b x)) < e" by force - show "\d>0. \x'. dist x' x < d \ - dist (f (clamp a b x')) (f (clamp a b x)) < e" + show "\d>0. \x'. dist x' x < d \ dist (f (clamp a b x')) (f (clamp a b x)) < e" using le by (auto intro!: d clamp_in_interval dist_clamps_le_dist_args[THEN le_less_trans]) qed @@ -1876,8 +1821,7 @@ from bounded obtain c where f_bound: "\x\f ` cbox a b. dist undefined x \ c" by (auto simp: bounded_any_center[where a=undefined]) then show ?thesis - by (auto intro!: exI[where x=c] clamp_in_interval[OF le[rule_format]] - simp: bounded_any_center[where a=undefined]) + by (metis bounded bounded_subset clamp_in_interval image_mono image_subsetI le range_composition) qed (auto simp: clamp_empty_interval image_def) @@ -1888,9 +1832,7 @@ fixes x a b :: "'a::euclidean_space" assumes x: "x \ cbox a b" shows "ext_cont f a b x = f x" - using assms - unfolding ext_cont_def - by (auto simp: clamp_def mem_box intro!: euclidean_eqI[where 'a='a] arg_cong[where f=f]) + using assms by (simp add: ext_cont_def) lemma continuous_on_ext_cont[continuous_intros]: "continuous_on (cbox a b) f \ continuous_on S (ext_cont f a b)" @@ -1909,18 +1851,16 @@ and Un: "\S. open S \ \U. U \ \ \ S = \U" using univ_second_countable by blast have *: "infinite (range (\n. ball (0::'a) (inverse(Suc n))))" - apply (rule Infinite_Set.range_inj_infinite) - apply (simp add: inj_on_def ball_eq_ball_iff) - done + by (simp add: inj_on_def ball_eq_ball_iff Infinite_Set.range_inj_infinite) have "infinite \" proof assume "finite \" then have "finite (Union ` (Pow \))" by simp - then have "finite (range (\n. ball (0::'a) (inverse(Suc n))))" - apply (rule rev_finite_subset) + moreover have "range (\n. ball 0 (inverse (real (Suc n)))) \ \ ` Pow \" by (metis (no_types, lifting) PowI image_eqI image_subset_iff Un [OF open_ball]) - with * show False by simp + ultimately show False + by (metis finite_subset *) qed obtain f :: "nat \ 'a set" where "\ = range f" "inj f" by (blast intro: countable_as_injective_image [OF \countable \\ \infinite \\]) @@ -1931,9 +1871,7 @@ using \inj f\ \\ = range f\ apply force done show ?thesis - apply (rule that [OF \inj f\ _ *]) - apply (auto simp: \\ = range f\ opn) - done + using "*" \\ = range f\ \inj f\ opn that by force qed proposition separable: @@ -1969,14 +1907,8 @@ using \0 < e\ \ \x \ S\ by auto next case False - then obtain C where "C \ \" by blast - show ?thesis - proof - show "dist (f C) x < e" - by (metis Int_iff Union_iff \ \C \ \\ dist_commute f mem_ball subsetCE) - show "C \ \" - using \\ \ \\ \C \ \\ by blast - qed + then show ?thesis + by (metis IntI Union_iff \ \0 < e\ \x \ S\ dist_commute dist_self f inf_le2 mem_ball subset_eq) qed qed qed @@ -2001,8 +1933,7 @@ then show "norm (x - y) \ 2*r" by simp qed (simp add: that) have "2*r = dist (a + r *\<^sub>R (SOME i. i \ Basis)) (a - r *\<^sub>R (SOME i. i \ Basis))" - apply (simp add: dist_norm) - by (metis abs_of_nonneg mult.right_neutral norm_numeral norm_scaleR norm_some_Basis real_norm_def scaleR_2 that) + using \0 \ r\ that by (simp add: dist_norm flip: scaleR_2) also have "... \ diameter (cball a r)" apply (rule diameter_bounded_bound) using that by (auto simp: dist_norm) @@ -2023,8 +1954,8 @@ lemma diameter_closed_interval [simp]: "diameter {a..b} = (if b < a then 0 else b-a)" proof - - have "{a .. b} = cball ((a+b)/2) ((b-a)/2)" - by (auto simp: dist_norm abs_if field_split_simps split: if_split_asm) + have "{a..b} = cball ((a+b)/2) ((b-a)/2)" + using atLeastAtMost_eq_cball by blast then show ?thesis by simp qed @@ -2032,7 +1963,7 @@ lemma diameter_open_interval [simp]: "diameter {a<.. open A" proof assume "open(f ` A)" - then have "open(f -` (f ` A))" - using assms by (force simp: linear_continuous_at linear_conv_bounded_linear continuous_open_vimage) then show "open A" - by (simp add: assms bij_is_inj inj_vimage_image_eq) + by (metis assms bij_is_inj continuous_open_vimage inj_vimage_image_eq linear_continuous_at linear_linear) next assume "open A" then show "open(f ` A)" @@ -2114,20 +2043,9 @@ corollary interior_bijective_linear_image: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "linear f" "bij f" - shows "interior (f ` S) = f ` interior S" (is "?lhs = ?rhs") -proof safe - fix x - assume x: "x \ ?lhs" - then obtain T where "open T" and "x \ T" and "T \ f ` S" - by (metis interiorE) - then show "x \ ?rhs" - by (metis (no_types, opaque_lifting) assms subsetD interior_maximal open_bijective_linear_image_eq subset_image_iff) -next - fix x - assume x: "x \ interior S" - then show "f x \ interior (f ` S)" - by (meson assms imageI image_mono interiorI interior_subset open_bijective_linear_image_eq open_interior) -qed + shows "interior (f ` S) = f ` interior S" + by (smt (verit) assms bij_is_inj inj_image_subset_iff interior_maximal interior_subset + open_bijective_linear_image_eq open_interior subset_antisym subset_imageE) lemma interior_injective_linear_image: fixes f :: "'a::euclidean_space \ 'a::euclidean_space" @@ -2234,10 +2152,9 @@ shows "closed(f ` s)" proof - obtain e where "e > 0" and e: "\x\s. e * norm x \ norm (f x)" - using injective_imp_isometric[OF assms(4,1,2,3)] by auto - show ?thesis - using complete_isometric_image[OF \e>0\ assms(1,2) e] and assms(4) - unfolding complete_eq_closed[symmetric] by auto + using assms injective_imp_isometric by blast + with assms show ?thesis + by (meson complete_eq_closed complete_isometric_image) qed @@ -2274,9 +2191,8 @@ show "closedin (top_of_set (range f)) (f ` S)" using continuous_closedin_preimage [OF confg cgf] by simp show "closed (range f)" - apply (rule closed_injective_image_subspace) - using f apply (auto simp: linear_linear linear_injective_0) - done + using closed_injective_image_subspace f linear_conv_bounded_linear + linear_injective_0 subspace_UNIV by blast qed qed @@ -2289,25 +2205,28 @@ lemma closure_injective_linear_image: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" shows "\linear f; inj f\ \ f ` (closure S) = closure (f ` S)" - apply (rule subset_antisym) - apply (simp add: closure_linear_image_subset) - by (simp add: closure_minimal closed_injective_linear_image closure_subset image_mono) + by (simp add: closed_injective_linear_image closure_linear_image_subset + closure_minimal closure_subset image_mono subset_antisym) lemma closure_bounded_linear_image: - fixes f :: "'a::euclidean_space \ 'b::euclidean_space" - shows "\linear f; bounded S\ \ f ` (closure S) = closure (f ` S)" - apply (rule subset_antisym, simp add: closure_linear_image_subset) - apply (rule closure_minimal, simp add: closure_subset image_mono) - by (meson bounded_closure closed_closure compact_continuous_image compact_eq_bounded_closed linear_continuous_on linear_conv_bounded_linear) + fixes f :: "'a::euclidean_space \ 'b::euclidean_space" + assumes "linear f" "bounded S" + shows "f ` (closure S) = closure (f ` S)" (is "?lhs = ?rhs") +proof + show "?lhs \ ?rhs" + using assms closure_linear_image_subset by blast + show "?rhs \ ?lhs" + using assms by (meson closure_minimal closure_subset compact_closure compact_eq_bounded_closed + compact_continuous_image image_mono linear_continuous_on linear_linear) +qed lemma closure_scaleR: fixes S :: "'a::real_normed_vector set" - shows "((*\<^sub>R) c) ` (closure S) = closure (((*\<^sub>R) c) ` S)" + shows "((*\<^sub>R) c) ` (closure S) = closure (((*\<^sub>R) c) ` S)" (is "?lhs = ?rhs") proof - show "((*\<^sub>R) c) ` (closure S) \ closure (((*\<^sub>R) c) ` S)" - using bounded_linear_scaleR_right - by (rule closure_bounded_linear_image_subset) - show "closure (((*\<^sub>R) c) ` S) \ ((*\<^sub>R) c) ` (closure S)" + show "?lhs \ ?rhs" + using bounded_linear_scaleR_right by (rule closure_bounded_linear_image_subset) + show "?rhs \ ?lhs" by (intro closure_minimal image_mono closure_subset closed_scaling closed_closure) qed @@ -2326,73 +2245,55 @@ qed lemma closed_subspace: - fixes s :: "'a::euclidean_space set" - assumes "subspace s" - shows "closed s" + fixes S :: "'a::euclidean_space set" + assumes "subspace S" + shows "closed S" proof - - have "dim s \ card (Basis :: 'a set)" + have "dim S \ card (Basis :: 'a set)" using dim_subset_UNIV by auto - with ex_card[OF this] obtain d :: "'a set" where t: "card d = dim s" and d: "d \ Basis" - by auto + with obtain_subset_with_card_n + obtain d :: "'a set" where cd: "card d = dim S" and d: "d \ Basis" + by metis let ?t = "{x::'a. \i\Basis. i \ d \ x\i = 0}" - have "\f. linear f \ f ` {x::'a. \i\Basis. i \ d \ x \ i = 0} = s \ + have "\f. linear f \ f ` {x::'a. \i\Basis. i \ d \ x \ i = 0} = S \ inj_on f {x::'a. \i\Basis. i \ d \ x \ i = 0}" - using dim_substandard[of d] t d assms + using dim_substandard[of d] cd d assms by (intro subspace_isomorphism[OF subspace_substandard[of "\i. i \ d"]]) (auto simp: inner_Basis) then obtain f where f: "linear f" - "f ` {x. \i\Basis. i \ d \ x \ i = 0} = s" + "f ` {x. \i\Basis. i \ d \ x \ i = 0} = S" "inj_on f {x. \i\Basis. i \ d \ x \ i = 0}" by blast interpret f: bounded_linear f using f by (simp add: linear_conv_bounded_linear) have "x \ ?t \ f x = 0 \ x = 0" for x using f.zero d f(3)[THEN inj_onD, of x 0] by auto - moreover have "closed ?t" by (rule closed_substandard) - moreover have "subspace ?t" by (rule subspace_substandard) - ultimately show ?thesis - using closed_injective_image_subspace[of ?t f] - unfolding f(2) using f(1) unfolding linear_conv_bounded_linear by auto + then show ?thesis + using closed_injective_image_subspace[of ?t f] closed_substandard subspace_substandard + using f(2) f.bounded_linear_axioms by force qed -lemma complete_subspace: "subspace s \ complete s" - for s :: "'a::euclidean_space set" +lemma complete_subspace: "subspace S \ complete S" + for S :: "'a::euclidean_space set" using complete_eq_closed closed_subspace by auto -lemma closed_span [iff]: "closed (span s)" - for s :: "'a::euclidean_space set" +lemma closed_span [iff]: "closed (span S)" + for S :: "'a::euclidean_space set" by (simp add: closed_subspace) -lemma dim_closure [simp]: "dim (closure s) = dim s" (is "?dc = ?d") - for s :: "'a::euclidean_space set" -proof - - have "?dc \ ?d" - using closure_minimal[OF span_superset, of s] - using closed_subspace[OF subspace_span, of s] - using dim_subset[of "closure s" "span s"] - by simp - then show ?thesis - using dim_subset[OF closure_subset, of s] - by simp -qed +lemma dim_closure [simp]: "dim (closure S) = dim S" (is "?dc = ?d") + for S :: "'a::euclidean_space set" + by (metis closed_span closure_minimal closure_subset dim_eq_span span_eq_dim span_superset subset_le_dim) subsection \Set Distance\ lemma setdist_compact_closed: fixes A :: "'a::heine_borel set" - assumes A: "compact A" and B: "closed B" + assumes "compact A" "closed B" and "A \ {}" "B \ {}" shows "\x \ A. \y \ B. dist x y = setdist A B" -proof - - obtain x where "x \ A" "setdist A B = infdist x B" - by (metis A assms(3) setdist_attains_inf setdist_sym) - moreover - obtain y where"y \ B" "infdist x B = dist x y" - using B \B \ {}\ infdist_attains_inf by blast - ultimately show ?thesis - using \x \ A\ \y \ B\ by auto -qed + by (metis assms infdist_attains_inf setdist_attains_inf setdist_sym) lemma setdist_closed_compact: fixes S :: "'a::heine_borel set" @@ -2406,14 +2307,10 @@ assumes S: "compact S" and T: "closed T" shows "setdist S T = 0 \ S = {} \ T = {} \ S \ T \ {}" proof (cases "S = {} \ T = {}") - case True - then show ?thesis - by force -next case False then show ?thesis - by (metis S T disjoint_iff_not_equal in_closed_iff_infdist_zero setdist_attains_inf setdist_eq_0I setdist_sym) -qed + by (metis S T disjoint_iff in_closed_iff_infdist_zero setdist_attains_inf setdist_eq_0I setdist_sym) +qed auto corollary setdist_gt_0_compact_closed: assumes S: "compact S" and T: "closed T"