# HG changeset patch # User paulson # Date 1626462792 -3600 # Node ID 4cca14dc577c16c51252c01e85ea7f6799230b56 # Parent 1a0a536b8aafd19bf01d5bd5dda384c5b560e739# Parent df976eefcba097c6a2119a7bbe31c139fb076ff9 merged diff -r 1a0a536b8aaf -r 4cca14dc577c src/HOL/Analysis/Path_Connected.thy --- a/src/HOL/Analysis/Path_Connected.thy Fri Jul 16 15:42:52 2021 +0200 +++ b/src/HOL/Analysis/Path_Connected.thy Fri Jul 16 20:13:12 2021 +0100 @@ -137,6 +137,24 @@ subsection\<^marker>\tag unimportant\\Basic lemmas about paths\ +lemma path_of_real: "path complex_of_real" + unfolding path_def by (intro continuous_intros) + +lemma path_const: "path (\t. a)" for a::"'a::real_normed_vector" + unfolding path_def by (intro continuous_intros) + +lemma path_minus: "path g \ path (\t. - g t)" for g::"real\'a::real_normed_vector" + unfolding path_def by (intro continuous_intros) + +lemma path_add: "\path f; path g\ \ path (\t. f t + g t)" for f::"real\'a::real_normed_vector" + unfolding path_def by (intro continuous_intros) + +lemma path_diff: "\path f; path g\ \ path (\t. f t - g t)" for f::"real\'a::real_normed_vector" + unfolding path_def by (intro continuous_intros) + +lemma path_mult: "\path f; path g\ \ path (\t. f t * g t)" for f::"real\'a::real_normed_field" + unfolding path_def by (intro continuous_intros) + lemma pathin_iff_path_real [simp]: "pathin euclideanreal g \ path g" by (simp add: pathin_def path_def) diff -r 1a0a536b8aaf -r 4cca14dc577c src/HOL/Analysis/Starlike.thy --- a/src/HOL/Analysis/Starlike.thy Fri Jul 16 15:42:52 2021 +0200 +++ b/src/HOL/Analysis/Starlike.thy Fri Jul 16 20:13:12 2021 +0100 @@ -1493,16 +1493,7 @@ and "affine T" and "rel_interior S \ T \ {}" shows "closure (S \ T) = closure S \ T" -proof - - have "affine hull T = T" - using assms by auto - then have "rel_interior T = T" - using rel_interior_affine_hull[of T] by metis - moreover have "closure T = T" - using assms affine_closed[of T] by auto - ultimately show ?thesis - using convex_closure_inter_two[of S T] assms affine_imp_convex by auto -qed + by (metis affine_imp_convex assms convex_closure_inter_two rel_interior_affine rel_interior_eq_closure) lemma connected_component_1_gen: fixes S :: "'a :: euclidean_space set" @@ -1523,16 +1514,7 @@ and "affine T" and "rel_interior S \ T \ {}" shows "rel_interior (S \ T) = rel_interior S \ T" -proof - - have "affine hull T = T" - using assms by auto - then have "rel_interior T = T" - using rel_interior_affine_hull[of T] by metis - moreover have "closure T = T" - using assms affine_closed[of T] by auto - ultimately show ?thesis - using convex_rel_interior_inter_two[of S T] assms affine_imp_convex by auto -qed + by (simp add: affine_imp_convex assms convex_rel_interior_inter_two rel_interior_affine) lemma convex_affine_rel_frontier_Int: fixes S T :: "'n::euclidean_space set" @@ -1547,49 +1529,8 @@ lemma rel_interior_convex_Int_affine: fixes S :: "'a::euclidean_space set" assumes "convex S" "affine T" "interior S \ T \ {}" - shows "rel_interior(S \ T) = interior S \ T" -proof - - obtain a where aS: "a \ interior S" and aT:"a \ T" - using assms by force - have "rel_interior S = interior S" - by (metis (no_types) aS affine_hull_nonempty_interior equals0D rel_interior_interior) - then show ?thesis - by (metis (no_types) affine_imp_convex assms convex_rel_interior_inter_two hull_same rel_interior_affine_hull) -qed - -lemma closure_convex_Int_affine: - fixes S :: "'a::euclidean_space set" - assumes "convex S" "affine T" "rel_interior S \ T \ {}" - shows "closure(S \ T) = closure S \ T" -proof - have "closure (S \ T) \ closure T" - by (simp add: closure_mono) - also have "... \ T" - by (simp add: affine_closed assms) - finally show "closure(S \ T) \ closure S \ T" - by (simp add: closure_mono) -next - obtain a where "a \ rel_interior S" "a \ T" - using assms by auto - then have ssT: "subspace ((\x. (-a)+x) ` T)" and "a \ S" - using affine_diffs_subspace rel_interior_subset assms by blast+ - show "closure S \ T \ closure (S \ T)" - proof - fix x assume "x \ closure S \ T" - show "x \ closure (S \ T)" - proof (cases "x = a") - case True - then show ?thesis - using \a \ S\ \a \ T\ closure_subset by fastforce - next - case False - then have "x \ closure(open_segment a x)" - by auto - then show ?thesis - using \x \ closure S \ T\ assms convex_affine_closure_Int by blast - qed - qed -qed + shows "rel_interior(S \ T) = interior S \ T" + by (metis Int_emptyI assms convex_affine_rel_interior_Int empty_iff interior_rel_interior_gen) lemma subset_rel_interior_convex: fixes S T :: "'n::euclidean_space set" diff -r 1a0a536b8aaf -r 4cca14dc577c src/HOL/Complex_Analysis/Conformal_Mappings.thy --- a/src/HOL/Complex_Analysis/Conformal_Mappings.thy Fri Jul 16 15:42:52 2021 +0200 +++ b/src/HOL/Complex_Analysis/Conformal_Mappings.thy Fri Jul 16 20:13:12 2021 +0100 @@ -82,6 +82,38 @@ thus ?thesis by simp qed +corollary analytic_continuation': + assumes "f holomorphic_on S" "open S" "connected S" + and "U \ S" "\ \ S" "\ islimpt U" + and "f constant_on U" + shows "f constant_on S" +proof - + obtain c where c: "\x. x \ U \ f x - c = 0" + by (metis \f constant_on U\ constant_on_def diff_self) + have "(\z. f z - c) holomorphic_on S" + using assms by (intro holomorphic_intros) + with c analytic_continuation assms have "\x. x \ S \ f x - c = 0" + by blast + then show ?thesis + unfolding constant_on_def by force +qed + +lemma holomorphic_compact_finite_zeros: + assumes S: "f holomorphic_on S" "open S" "connected S" + and "compact K" "K \ S" + and "\ f constant_on S" + shows "finite {z\K. f z = 0}" +proof (rule ccontr) + assume "infinite {z\K. f z = 0}" + then obtain z where "z \ K" and z: "z islimpt {z\K. f z = 0}" + using \compact K\ by (auto simp: compact_eq_Bolzano_Weierstrass) + moreover have "{z\K. f z = 0} \ S" + using \K \ S\ by blast + ultimately show False + using assms analytic_continuation [OF S] unfolding constant_on_def + by blast +qed + subsection\Open mapping theorem\ lemma holomorphic_contract_to_zero: diff -r 1a0a536b8aaf -r 4cca14dc577c src/HOL/Conditionally_Complete_Lattices.thy --- a/src/HOL/Conditionally_Complete_Lattices.thy Fri Jul 16 15:42:52 2021 +0200 +++ b/src/HOL/Conditionally_Complete_Lattices.thy Fri Jul 16 20:13:12 2021 +0100 @@ -427,6 +427,25 @@ "A \ {} \ bdd_above A \ bdd_below A \ Inf A \ Sup A" by (auto intro!: cSup_upper2[of "SOME a. a \ A"] intro: someI cInf_lower) + context + fixes f :: "'a \ 'b::conditionally_complete_lattice" + assumes "mono f" + begin + + lemma mono_cInf: "\bdd_below A; A\{}\ \ f (Inf A) \ (INF x\A. f x)" + by (simp add: \mono f\ conditionally_complete_lattice_class.cINF_greatest cInf_lower monoD) + + lemma mono_cSup: "\bdd_above A; A\{}\ \ (SUP x\A. f x) \ f (Sup A)" + by (simp add: \mono f\ conditionally_complete_lattice_class.cSUP_least cSup_upper monoD) + + lemma mono_cINF: "\bdd_below (A`I); I\{}\ \ f (INF i\I. A i) \ (INF x\I. f (A x))" + by (simp add: \mono f\ conditionally_complete_lattice_class.cINF_greatest cINF_lower monoD) + + lemma mono_cSUP: "\bdd_above (A`I); I\{}\ \ (SUP x\I. f (A x)) \ f (SUP i\I. A i)" + by (simp add: \mono f\ conditionally_complete_lattice_class.cSUP_least cSUP_upper monoD) + + end + end instance complete_lattice \ conditionally_complete_lattice diff -r 1a0a536b8aaf -r 4cca14dc577c src/HOL/Deriv.thy --- a/src/HOL/Deriv.thy Fri Jul 16 15:42:52 2021 +0200 +++ b/src/HOL/Deriv.thy Fri Jul 16 20:13:12 2021 +0100 @@ -959,6 +959,12 @@ shows "(f has_vector_derivative x) F \ ((\x. f x * a) has_vector_derivative (x * a)) F" by (rule bounded_linear.has_vector_derivative[OF bounded_linear_mult_left]) +lemma has_vector_derivative_divide[derivative_intros]: + fixes a :: "'a::real_normed_field" + shows "(f has_vector_derivative x) F \ ((\x. f x / a) has_vector_derivative (x / a)) F" + using has_vector_derivative_mult_left [of f x F "inverse a"] + by (simp add: field_class.field_divide_inverse) + subsection \Derivatives\ diff -r 1a0a536b8aaf -r 4cca14dc577c src/HOL/Groups.thy --- a/src/HOL/Groups.thy Fri Jul 16 15:42:52 2021 +0200 +++ b/src/HOL/Groups.thy Fri Jul 16 20:13:12 2021 +0100 @@ -639,6 +639,11 @@ end +lemma mono_add: + fixes a :: "'a::ordered_ab_semigroup_add" + shows "mono ((+) a)" + by (simp add: add_left_mono monoI) + text \Strict monotonicity in both arguments\ class strict_ordered_ab_semigroup_add = ordered_ab_semigroup_add + assumes add_strict_mono: "a < b \ c < d \ a + c < b + d" diff -r 1a0a536b8aaf -r 4cca14dc577c src/HOL/Real_Vector_Spaces.thy --- a/src/HOL/Real_Vector_Spaces.thy Fri Jul 16 15:42:52 2021 +0200 +++ b/src/HOL/Real_Vector_Spaces.thy Fri Jul 16 20:13:12 2021 +0100 @@ -764,6 +764,9 @@ finally show ?thesis by simp qed +lemma bdd_below_norm_image: "bdd_below (norm ` A)" + by (meson bdd_belowI2 norm_ge_zero) + end class real_normed_algebra = real_algebra + real_normed_vector + diff -r 1a0a536b8aaf -r 4cca14dc577c src/HOL/Rings.thy --- a/src/HOL/Rings.thy Fri Jul 16 15:42:52 2021 +0200 +++ b/src/HOL/Rings.thy Fri Jul 16 20:13:12 2021 +0100 @@ -1884,6 +1884,11 @@ end +lemma mono_mult: + fixes a :: "'a::ordered_semiring" + shows "a \ 0 \ mono ((*) a)" + by (simp add: mono_def mult_left_mono) + class ordered_semiring_0 = semiring_0 + ordered_semiring begin