# HG changeset patch # User paulson # Date 1634298171 -3600 # Node ID 67d87d224e0004daeac3d050d00376f270c2742a # Parent c434f4e749471bcecd58c273442b39dda954d1dd A few new lemmas plus some refinements diff -r c434f4e74947 -r 67d87d224e00 src/HOL/Analysis/Complex_Analysis_Basics.thy --- a/src/HOL/Analysis/Complex_Analysis_Basics.thy Thu Oct 14 15:24:28 2021 +0200 +++ b/src/HOL/Analysis/Complex_Analysis_Basics.thy Fri Oct 15 12:42:51 2021 +0100 @@ -144,6 +144,25 @@ "open s \ f holomorphic_on s \ (\x \ s. \f'. DERIV f x :> f')" by (auto simp: holomorphic_on_def field_differentiable_def has_field_derivative_def at_within_open [of _ s]) +lemma holomorphic_on_UN_open: + assumes "\n. n \ I \ f holomorphic_on A n" "\n. n \ I \ open (A n)" + shows "f holomorphic_on (\n\I. A n)" +proof - + have "f field_differentiable at z within (\n\I. A n)" if "z \ (\n\I. A n)" for z + proof - + from that obtain n where "n \ I" "z \ A n" + by blast + hence "f holomorphic_on A n" "open (A n)" + by (simp add: assms)+ + with \z \ A n\ have "f field_differentiable at z" + by (auto simp: holomorphic_on_open field_differentiable_def) + thus ?thesis + by (meson field_differentiable_at_within) + qed + thus ?thesis + by (auto simp: holomorphic_on_def) +qed + lemma holomorphic_on_imp_continuous_on: "f holomorphic_on s \ continuous_on s f" by (metis field_differentiable_imp_continuous_at continuous_on_eq_continuous_within holomorphic_on_def) @@ -181,6 +200,18 @@ lemma holomorphic_on_id [simp, holomorphic_intros]: "id holomorphic_on s" unfolding id_def by (rule holomorphic_on_ident) +lemma constant_on_imp_holomorphic_on: + assumes "f constant_on A" + shows "f holomorphic_on A" +proof - + from assms obtain c where c: "\x\A. f x = c" + unfolding constant_on_def by blast + have "f holomorphic_on A \ (\_. c) holomorphic_on A" + by (intro holomorphic_cong) (use c in auto) + thus ?thesis + by simp +qed + lemma holomorphic_on_compose: "f holomorphic_on s \ g holomorphic_on (f ` s) \ (g o f) holomorphic_on s" using field_differentiable_compose_within[of f _ s g] diff -r c434f4e74947 -r 67d87d224e00 src/HOL/Analysis/Complex_Transcendental.thy --- a/src/HOL/Analysis/Complex_Transcendental.thy Thu Oct 14 15:24:28 2021 +0200 +++ b/src/HOL/Analysis/Complex_Transcendental.thy Fri Oct 15 12:42:51 2021 +0100 @@ -2258,8 +2258,8 @@ by (induct n) (auto simp: ac_simps powr_add) lemma powr_complexnumeral [simp]: - fixes x::complex shows "x \ 0 \ x powr (numeral n) = x ^ (numeral n)" - by (metis of_nat_numeral powr_complexpow) + fixes x::complex shows "x powr (numeral n) = x ^ (numeral n)" + by (metis of_nat_numeral power_zero_numeral powr_nat) lemma cnj_powr: assumes "Im a = 0 \ Re a \ 0" diff -r c434f4e74947 -r 67d87d224e00 src/HOL/Analysis/Elementary_Topology.thy --- a/src/HOL/Analysis/Elementary_Topology.thy Thu Oct 14 15:24:28 2021 +0200 +++ b/src/HOL/Analysis/Elementary_Topology.thy Fri Oct 15 12:42:51 2021 +0100 @@ -1597,11 +1597,20 @@ lemma closure_insert: fixes x :: "'a::t1_space" shows "closure (insert x s) = insert x (closure s)" - apply (rule closure_unique) - apply (rule insert_mono [OF closure_subset]) - apply (rule closed_insert [OF closed_closure]) - apply (simp add: closure_minimal) - done + by (meson closed_closure closed_insert closure_minimal closure_subset dual_order.antisym insert_mono insert_subset) + +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 text\In particular, some common special cases.\ diff -r c434f4e74947 -r 67d87d224e00 src/HOL/Analysis/Polytope.thy --- a/src/HOL/Analysis/Polytope.thy Thu Oct 14 15:24:28 2021 +0200 +++ b/src/HOL/Analysis/Polytope.thy Fri Oct 15 12:42:51 2021 +0100 @@ -434,11 +434,10 @@ then obtain c where c: "\F'. \finite F'; F' \ T; card F' \ DIM('a) + 2\ \ c F' \ T \ c F' \ (\F') \ (\F')" by metis - define d where "d = rec_nat {c{}} (\n r. insert (c r) r)" - have [simp]: "d 0 = {c {}}" - by (simp add: d_def) - have dSuc [simp]: "\n. d (Suc n) = insert (c (d n)) (d n)" - by (simp add: d_def) + define d where "d \ \n. ((\r. insert (c r) r)^^n) {c{}}" + note d_def [simp] + have dSuc: "\n. d (Suc n) = insert (c (d n)) (d n)" + by simp have dn_notempty: "d n \ {}" for n by (induction n) auto have dn_le_Suc: "d n \ T \ finite(d n) \ card(d n) \ Suc n" if "n \ DIM('a) + 2" for n diff -r c434f4e74947 -r 67d87d224e00 src/HOL/Limits.thy --- a/src/HOL/Limits.thy Thu Oct 14 15:24:28 2021 +0200 +++ b/src/HOL/Limits.thy Fri Oct 15 12:42:51 2021 +0100 @@ -2849,9 +2849,8 @@ and local: "\x. a \ x \ x \ b \ \d>0. \a b. a \ x \ x \ b \ b - a < d \ P a b" shows "P a b" proof - - define bisect where "bisect = - rec_nat (a, b) (\n (x, y). if P x ((x+y) / 2) then ((x+y)/2, y) else (x, (x+y)/2))" - define l u where "l n = fst (bisect n)" and "u n = snd (bisect n)" for n + define bisect where "bisect \ \(x,y). if P x ((x+y) / 2) then ((x+y)/2, y) else (x, (x+y)/2)" + define l u where "l n \ fst ((bisect^^n)(a,b))" and "u n \ snd ((bisect^^n)(a,b))" for n have l[simp]: "l 0 = a" "\n. l (Suc n) = (if P (l n) ((l n + u n) / 2) then (l n + u n) / 2 else l n)" and u[simp]: "u 0 = b" "\n. u (Suc n) = (if P (l n) ((l n + u n) / 2) then u n else (l n + u n) / 2)" by (simp_all add: l_def u_def bisect_def split: prod.split)