# HG changeset patch # User paulson # Date 1677777438 0 # Node ID 2c86ea8961b598e071f611a1886c7d13d32de33d # Parent df1150ec6cd770e94ae05a9e74e4bb8597631e0c Some new lemmas. Some tidying up diff -r df1150ec6cd7 -r 2c86ea8961b5 src/HOL/Analysis/Convex.thy --- a/src/HOL/Analysis/Convex.thy Thu Mar 02 11:34:54 2023 +0000 +++ b/src/HOL/Analysis/Convex.thy Thu Mar 02 17:17:18 2023 +0000 @@ -2369,9 +2369,6 @@ using hull_inc u by fastforce qed -lemma inner_sum_Basis[simp]: "i \ Basis \ (\Basis) \ i = 1" - by (simp add: inner_sum_left sum.If_cases inner_Basis) - lemma convex_set_plus: assumes "convex S" and "convex T" shows "convex (S + T)" proof - diff -r df1150ec6cd7 -r 2c86ea8961b5 src/HOL/Analysis/Elementary_Metric_Spaces.thy --- a/src/HOL/Analysis/Elementary_Metric_Spaces.thy Thu Mar 02 11:34:54 2023 +0000 +++ b/src/HOL/Analysis/Elementary_Metric_Spaces.thy Thu Mar 02 17:17:18 2023 +0000 @@ -895,6 +895,17 @@ by (rule bounded_subset) qed +lemma continuous_on_compact_bound: + assumes "compact A" "continuous_on A f" + obtains B where "B \ 0" "\x. x \ A \ norm (f x) \ B" +proof - + have "compact (f ` A)" by (metis assms compact_continuous_image) + then obtain B where "\x\A. norm (f x) \ B" + by (auto dest!: compact_imp_bounded simp: bounded_iff) + hence "max B 0 \ 0" and "\x\A. norm (f x) \ max B 0" by auto + thus ?thesis using that by blast +qed + lemma closure_Int_ball_not_empty: assumes "S \ closure T" "x \ S" "r > 0" shows "T \ ball x r \ {}" diff -r df1150ec6cd7 -r 2c86ea8961b5 src/HOL/Analysis/Euclidean_Space.thy --- a/src/HOL/Analysis/Euclidean_Space.thy Thu Mar 02 11:34:54 2023 +0000 +++ b/src/HOL/Analysis/Euclidean_Space.thy Thu Mar 02 17:17:18 2023 +0000 @@ -53,6 +53,9 @@ lemma (in euclidean_space) sgn_Basis: "u \ Basis \ sgn u = u" unfolding sgn_div_norm by (simp add: scaleR_one) +lemma inner_sum_Basis[simp]: "i \ Basis \ inner (\Basis) i = 1" + by (simp add: inner_sum_left sum.If_cases inner_Basis) + lemma (in euclidean_space) Basis_zero [simp]: "0 \ Basis" proof assume "0 \ Basis" thus "False" diff -r df1150ec6cd7 -r 2c86ea8961b5 src/HOL/Analysis/Topology_Euclidean_Space.thy --- a/src/HOL/Analysis/Topology_Euclidean_Space.thy Thu Mar 02 11:34:54 2023 +0000 +++ b/src/HOL/Analysis/Topology_Euclidean_Space.thy Thu Mar 02 17:17:18 2023 +0000 @@ -840,6 +840,76 @@ by (metis assms continuous_on_compose continuous_on_swap swap_cbox_Pair) qed +lemma open_contains_cbox: + fixes x :: "'a :: euclidean_space" + assumes "open A" "x \ A" + obtains a b where "cbox a b \ A" "x \ box a b" "\i\Basis. a \ i < b \ i" +proof - + from assms obtain R where R: "R > 0" "ball x R \ A" + by (auto simp: open_contains_ball) + define r :: real where "r = R / (2 * sqrt DIM('a))" + from \R > 0\ have [simp]: "r > 0" by (auto simp: r_def) + define d :: 'a where "d = r *\<^sub>R Topology_Euclidean_Space.One" + have "cbox (x - d) (x + d) \ A" + proof safe + fix y assume y: "y \ cbox (x - d) (x + d)" + have "dist x y = sqrt (\i\Basis. (dist (x \ i) (y \ i))\<^sup>2)" + by (subst euclidean_dist_l2) (auto simp: L2_set_def) + also from y have "sqrt (\i\Basis. (dist (x \ i) (y \ i))\<^sup>2) \ sqrt (\i\(Basis::'a set). r\<^sup>2)" + by (intro real_sqrt_le_mono sum_mono power_mono) + (auto simp: dist_norm d_def cbox_def algebra_simps) + also have "\ = sqrt (DIM('a) * r\<^sup>2)" by simp + also have "DIM('a) * r\<^sup>2 = (R / 2) ^ 2" + by (simp add: r_def power_divide) + also have "sqrt \ = R / 2" + using \R > 0\ by simp + also from \R > 0\ have "\ < R" by simp + finally have "y \ ball x R" by simp + with R show "y \ A" by blast + qed + thus ?thesis + using that[of "x - d" "x + d"] by (auto simp: algebra_simps d_def box_def) +qed + +lemma open_contains_box: + fixes x :: "'a :: euclidean_space" + assumes "open A" "x \ A" + obtains a b where "box a b \ A" "x \ box a b" "\i\Basis. a \ i < b \ i" + by (meson assms box_subset_cbox dual_order.trans open_contains_cbox) + +lemma inner_image_box: + assumes "(i :: 'a :: euclidean_space) \ Basis" + assumes "\i\Basis. a \ i < b \ i" + shows "(\x. x \ i) ` box a b = {a \ i<.. i}" +proof safe + fix x assume x: "x \ {a \ i<.. i}" + let ?y = "(\j\Basis. (if i = j then x else (a + b) \ j / 2) *\<^sub>R j)" + from x assms have "?y \ i \ (\x. x \ i) ` box a b" + by (intro imageI) (auto simp: box_def algebra_simps) + also have "?y \ i = (\j\Basis. (if i = j then x else (a + b) \ j / 2) * (j \ i))" + by (simp add: inner_sum_left) + also have "\ = (\j\Basis. if i = j then x else 0)" + by (intro sum.cong) (auto simp: inner_not_same_Basis assms) + also have "\ = x" using assms by simp + finally show "x \ (\x. x \ i) ` box a b" . +qed (insert assms, auto simp: box_def) + +lemma inner_image_cbox: + assumes "(i :: 'a :: euclidean_space) \ Basis" + assumes "\i\Basis. a \ i \ b \ i" + shows "(\x. x \ i) ` cbox a b = {a \ i..b \ i}" +proof safe + fix x assume x: "x \ {a \ i..b \ i}" + let ?y = "(\j\Basis. (if i = j then x else a \ j) *\<^sub>R j)" + from x assms have "?y \ i \ (\x. x \ i) ` cbox a b" + by (intro imageI) (auto simp: cbox_def) + also have "?y \ i = (\j\Basis. (if i = j then x else a \ j) * (j \ i))" + by (simp add: inner_sum_left) + also have "\ = (\j\Basis. if i = j then x else 0)" + by (intro sum.cong) (auto simp: inner_not_same_Basis assms) + also have "\ = x" using assms by simp + finally show "x \ (\x. x \ i) ` cbox a b" . +qed (insert assms, auto simp: cbox_def) subsection \General Intervals\ diff -r df1150ec6cd7 -r 2c86ea8961b5 src/HOL/Archimedean_Field.thy --- a/src/HOL/Archimedean_Field.thy Thu Mar 02 11:34:54 2023 +0000 +++ b/src/HOL/Archimedean_Field.thy Thu Mar 02 17:17:18 2023 +0000 @@ -417,33 +417,7 @@ lemma floor_divide_of_nat_eq: "\of_nat m / of_nat n\ = of_nat (m div n)" for m n :: nat -proof (cases "n = 0") - case True - then show ?thesis by simp -next - case False - then have *: "\of_nat (m mod n) / of_nat n :: 'a\ = 0" - by (auto intro: floor_unique) - have "(of_nat m :: 'a) = of_nat (m div n * n + m mod n)" - by simp - also have "\ = (of_nat (m div n) + of_nat (m mod n) / of_nat n) * of_nat n" - using False by (simp only: of_nat_add) (simp add: field_simps) - finally have "(of_nat m / of_nat n :: 'a) = \ / of_nat n" - by simp - then have "(of_nat m / of_nat n :: 'a) = of_nat (m div n) + of_nat (m mod n) / of_nat n" - using False by (simp only:) simp - then have "\of_nat m / of_nat n :: 'a\ = \of_nat (m div n) + of_nat (m mod n) / of_nat n :: 'a\" - by simp - then have "\of_nat m / of_nat n :: 'a\ = \of_nat (m mod n) / of_nat n + of_nat (m div n) :: 'a\" - by (simp add: ac_simps) - moreover have "(of_nat (m div n) :: 'a) = of_int (of_nat (m div n))" - by simp - ultimately have "\of_nat m / of_nat n :: 'a\ = - \of_nat (m mod n) / of_nat n :: 'a\ + of_nat (m div n)" - by (simp only: floor_add_int) - with * show ?thesis - by simp -qed + by (metis floor_divide_of_int_eq of_int_of_nat_eq unique_euclidean_semiring_with_nat_class.of_nat_div) lemma floor_divide_lower: fixes q :: "'a::floor_ceiling" diff -r df1150ec6cd7 -r 2c86ea8961b5 src/HOL/Library/Landau_Symbols.thy --- a/src/HOL/Library/Landau_Symbols.thy Thu Mar 02 11:34:54 2023 +0000 +++ b/src/HOL/Library/Landau_Symbols.thy Thu Mar 02 17:17:18 2023 +0000 @@ -2052,6 +2052,17 @@ shows "(\x. f1 x / f2 x) \[F] (\x. g1 x / g2 x)" using asymp_equiv_mult[OF assms(1) asymp_equiv_inverse[OF assms(2)]] by (simp add: field_simps) +lemma asymp_equivD_strong: + assumes "f \[F] g" "eventually (\x. f x \ 0 \ g x \ 0) F" + shows "((\x. f x / g x) \ 1) F" +proof - + from assms(1) have "((\x. if f x = 0 \ g x = 0 then 1 else f x / g x) \ 1) F" + by (rule asymp_equivD) + also have "?this \ ?thesis" + by (intro filterlim_cong eventually_mono[OF assms(2)]) auto + finally show ?thesis . +qed + lemma asymp_equiv_compose [asymp_equiv_intros]: assumes "f \[G] g" "filterlim h G F" shows "f \ h \[F] g \ h" diff -r df1150ec6cd7 -r 2c86ea8961b5 src/HOL/Real.thy --- a/src/HOL/Real.thy Thu Mar 02 11:34:54 2023 +0000 +++ b/src/HOL/Real.thy Thu Mar 02 17:17:18 2023 +0000 @@ -1067,11 +1067,7 @@ by (metis floor_divide_of_nat_eq of_int_floor_le of_int_of_nat_eq) lemma real_of_nat_div3: "real n / real x - real (n div x) \ 1" for n x :: nat -proof (cases "x = 0") - case False - then show ?thesis - by (metis of_int_of_nat_eq real_of_int_div3 zdiv_int) -qed auto + by (metis of_int_of_nat_eq real_of_int_div3 of_nat_div) lemma real_of_nat_div4: "real (n div x) \ real n / real x" for n x :: nat using real_of_nat_div2 [of n x] by simp diff -r df1150ec6cd7 -r 2c86ea8961b5 src/HOL/Transcendental.thy --- a/src/HOL/Transcendental.thy Thu Mar 02 11:34:54 2023 +0000 +++ b/src/HOL/Transcendental.thy Thu Mar 02 17:17:18 2023 +0000 @@ -2492,6 +2492,9 @@ for a x :: "'a::{ln,real_normed_field}" by (simp add: divide_inverse powr_minus) +lemma powr_sum: "x \ 0 \ finite A \ x powr sum f A = (\y\A. x powr f y)" + by (simp add: powr_def exp_sum sum_distrib_right) + lemma divide_powr_uminus: "a / b powr c = a * b powr (- c)" for a b c :: real by (simp add: powr_minus_divide)