diff --git a/src/HOL/Analysis/Complex_Analysis_Basics.thy b/src/HOL/Analysis/Complex_Analysis_Basics.thy --- a/src/HOL/Analysis/Complex_Analysis_Basics.thy +++ b/src/HOL/Analysis/Complex_Analysis_Basics.thy @@ -1,864 +1,895 @@ (* Author: John Harrison, Marco Maggesi, Graziano Gentili, Gianni Ciolli, Valentina Bruno Ported from "hol_light/Multivariate/canal.ml" by L C Paulson (2014) *) section \Complex Analysis Basics\ text \Definitions of analytic and holomorphic functions, limit theorems, complex differentiation\ theory Complex_Analysis_Basics imports Derivative "HOL-Library.Nonpos_Ints" begin subsection\<^marker>\tag unimportant\\General lemmas\ lemma nonneg_Reals_cmod_eq_Re: "z \ \\<^sub>\\<^sub>0 \ norm z = Re z" by (simp add: complex_nonneg_Reals_iff cmod_eq_Re) lemma fact_cancel: fixes c :: "'a::real_field" shows "of_nat (Suc n) * c / (fact (Suc n)) = c / (fact n)" using of_nat_neq_0 by force lemma vector_derivative_cnj_within: assumes "at x within A \ bot" and "f differentiable at x within A" shows "vector_derivative (\z. cnj (f z)) (at x within A) = cnj (vector_derivative f (at x within A))" (is "_ = cnj ?D") proof - let ?D = "vector_derivative f (at x within A)" from assms have "(f has_vector_derivative ?D) (at x within A)" by (subst (asm) vector_derivative_works) hence "((\x. cnj (f x)) has_vector_derivative cnj ?D) (at x within A)" by (rule has_vector_derivative_cnj) thus ?thesis using assms by (auto dest: vector_derivative_within) qed lemma vector_derivative_cnj: assumes "f differentiable at x" shows "vector_derivative (\z. cnj (f z)) (at x) = cnj (vector_derivative f (at x))" using assms by (intro vector_derivative_cnj_within) auto lemma shows open_halfspace_Re_lt: "open {z. Re(z) < b}" and open_halfspace_Re_gt: "open {z. Re(z) > b}" and closed_halfspace_Re_ge: "closed {z. Re(z) \ b}" and closed_halfspace_Re_le: "closed {z. Re(z) \ b}" and closed_halfspace_Re_eq: "closed {z. Re(z) = b}" and open_halfspace_Im_lt: "open {z. Im(z) < b}" and open_halfspace_Im_gt: "open {z. Im(z) > b}" and closed_halfspace_Im_ge: "closed {z. Im(z) \ b}" and closed_halfspace_Im_le: "closed {z. Im(z) \ b}" and closed_halfspace_Im_eq: "closed {z. Im(z) = b}" by (intro open_Collect_less closed_Collect_le closed_Collect_eq continuous_on_Re continuous_on_Im continuous_on_id continuous_on_const)+ lemma closed_complex_Reals: "closed (\ :: complex set)" proof - have "(\ :: complex set) = {z. Im z = 0}" by (auto simp: complex_is_Real_iff) then show ?thesis by (metis closed_halfspace_Im_eq) qed lemma closed_Real_halfspace_Re_le: "closed (\ \ {w. Re w \ x})" by (simp add: closed_Int closed_complex_Reals closed_halfspace_Re_le) lemma closed_nonpos_Reals_complex [simp]: "closed (\\<^sub>\\<^sub>0 :: complex set)" proof - have "\\<^sub>\\<^sub>0 = \ \ {z. Re(z) \ 0}" using complex_nonpos_Reals_iff complex_is_Real_iff by auto then show ?thesis by (metis closed_Real_halfspace_Re_le) qed lemma closed_Real_halfspace_Re_ge: "closed (\ \ {w. x \ Re(w)})" using closed_halfspace_Re_ge by (simp add: closed_Int closed_complex_Reals) lemma closed_nonneg_Reals_complex [simp]: "closed (\\<^sub>\\<^sub>0 :: complex set)" proof - have "\\<^sub>\\<^sub>0 = \ \ {z. Re(z) \ 0}" using complex_nonneg_Reals_iff complex_is_Real_iff by auto then show ?thesis by (metis closed_Real_halfspace_Re_ge) qed lemma closed_real_abs_le: "closed {w \ \. \Re w\ \ r}" proof - have "{w \ \. \Re w\ \ r} = (\ \ {w. Re w \ r}) \ (\ \ {w. Re w \ -r})" by auto then show "closed {w \ \. \Re w\ \ r}" by (simp add: closed_Int closed_Real_halfspace_Re_ge closed_Real_halfspace_Re_le) qed lemma real_lim: fixes l::complex assumes "(f \ l) F" and "\ trivial_limit F" and "eventually P F" and "\a. P a \ f a \ \" shows "l \ \" proof (rule Lim_in_closed_set[OF closed_complex_Reals _ assms(2,1)]) show "eventually (\x. f x \ \) F" using assms(3, 4) by (auto intro: eventually_mono) qed lemma real_lim_sequentially: fixes l::complex shows "(f \ l) sequentially \ (\N. \n\N. f n \ \) \ l \ \" by (rule real_lim [where F=sequentially]) (auto simp: eventually_sequentially) lemma real_series: fixes l::complex shows "f sums l \ (\n. f n \ \) \ l \ \" unfolding sums_def by (metis real_lim_sequentially sum_in_Reals) lemma Lim_null_comparison_Re: assumes "eventually (\x. norm(f x) \ Re(g x)) F" "(g \ 0) F" shows "(f \ 0) F" by (rule Lim_null_comparison[OF assms(1)] tendsto_eq_intros assms(2))+ simp subsection\Holomorphic functions\ definition\<^marker>\tag important\ holomorphic_on :: "[complex \ complex, complex set] \ bool" (infixl "(holomorphic'_on)" 50) where "f holomorphic_on s \ \x\s. f field_differentiable (at x within s)" named_theorems\<^marker>\tag important\ holomorphic_intros "structural introduction rules for holomorphic_on" lemma holomorphic_onI [intro?]: "(\x. x \ s \ f field_differentiable (at x within s)) \ f holomorphic_on s" by (simp add: holomorphic_on_def) lemma holomorphic_onD [dest?]: "\f holomorphic_on s; x \ s\ \ f field_differentiable (at x within s)" by (simp add: holomorphic_on_def) lemma holomorphic_on_imp_differentiable_on: "f holomorphic_on s \ f differentiable_on s" unfolding holomorphic_on_def differentiable_on_def by (simp add: field_differentiable_imp_differentiable) lemma holomorphic_on_imp_differentiable_at: "\f holomorphic_on s; open s; x \ s\ \ f field_differentiable (at x)" using at_within_open holomorphic_on_def by fastforce lemma holomorphic_on_empty [holomorphic_intros]: "f holomorphic_on {}" by (simp add: holomorphic_on_def) lemma holomorphic_on_open: "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) lemma holomorphic_closedin_preimage_constant: assumes "f holomorphic_on D" shows "closedin (top_of_set D) {z\D. f z = a}" by (simp add: assms continuous_closedin_preimage_constant holomorphic_on_imp_continuous_on) lemma holomorphic_closed_preimage_constant: assumes "f holomorphic_on UNIV" shows "closed {z. f z = a}" using holomorphic_closedin_preimage_constant [OF assms] by simp lemma holomorphic_on_subset [elim]: "f holomorphic_on s \ t \ s \ f holomorphic_on t" unfolding holomorphic_on_def by (metis field_differentiable_within_subset subsetD) lemma holomorphic_transform: "\f holomorphic_on s; \x. x \ s \ f x = g x\ \ g holomorphic_on s" by (metis field_differentiable_transform_within linordered_field_no_ub holomorphic_on_def) lemma holomorphic_cong: "s = t ==> (\x. x \ s \ f x = g x) \ f holomorphic_on s \ g holomorphic_on t" by (metis holomorphic_transform) lemma holomorphic_on_linear [simp, holomorphic_intros]: "((*) c) holomorphic_on s" unfolding holomorphic_on_def by (metis field_differentiable_linear) lemma holomorphic_on_const [simp, holomorphic_intros]: "(\z. c) holomorphic_on s" unfolding holomorphic_on_def by (metis field_differentiable_const) lemma holomorphic_on_ident [simp, holomorphic_intros]: "(\x. x) holomorphic_on s" unfolding holomorphic_on_def by (metis field_differentiable_ident) 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] by (auto simp: holomorphic_on_def) lemma holomorphic_on_compose_gen: "f holomorphic_on s \ g holomorphic_on t \ f ` s \ t \ (g o f) holomorphic_on s" by (metis holomorphic_on_compose holomorphic_on_subset) lemma holomorphic_on_balls_imp_entire: assumes "\bdd_above A" "\r. r \ A \ f holomorphic_on ball c r" shows "f holomorphic_on B" proof (rule holomorphic_on_subset) show "f holomorphic_on UNIV" unfolding holomorphic_on_def proof fix z :: complex from \\bdd_above A\ obtain r where r: "r \ A" "r > norm (z - c)" by (meson bdd_aboveI not_le) with assms(2) have "f holomorphic_on ball c r" by blast moreover from r have "z \ ball c r" by (auto simp: dist_norm norm_minus_commute) ultimately show "f field_differentiable at z" by (auto simp: holomorphic_on_def at_within_open[of _ "ball c r"]) qed qed auto lemma holomorphic_on_balls_imp_entire': assumes "\r. r > 0 \ f holomorphic_on ball c r" shows "f holomorphic_on B" proof (rule holomorphic_on_balls_imp_entire) { fix M :: real have "\x. x > max M 0" by (intro gt_ex) hence "\x>0. x > M" by auto } thus "\bdd_above {(0::real)<..}" unfolding bdd_above_def by (auto simp: not_le) qed (insert assms, auto) lemma holomorphic_on_minus [holomorphic_intros]: "f holomorphic_on s \ (\z. -(f z)) holomorphic_on s" by (metis field_differentiable_minus holomorphic_on_def) lemma holomorphic_on_add [holomorphic_intros]: "\f holomorphic_on s; g holomorphic_on s\ \ (\z. f z + g z) holomorphic_on s" unfolding holomorphic_on_def by (metis field_differentiable_add) lemma holomorphic_on_diff [holomorphic_intros]: "\f holomorphic_on s; g holomorphic_on s\ \ (\z. f z - g z) holomorphic_on s" unfolding holomorphic_on_def by (metis field_differentiable_diff) lemma holomorphic_on_mult [holomorphic_intros]: "\f holomorphic_on s; g holomorphic_on s\ \ (\z. f z * g z) holomorphic_on s" unfolding holomorphic_on_def by (metis field_differentiable_mult) lemma holomorphic_on_inverse [holomorphic_intros]: "\f holomorphic_on s; \z. z \ s \ f z \ 0\ \ (\z. inverse (f z)) holomorphic_on s" unfolding holomorphic_on_def by (metis field_differentiable_inverse) lemma holomorphic_on_divide [holomorphic_intros]: "\f holomorphic_on s; g holomorphic_on s; \z. z \ s \ g z \ 0\ \ (\z. f z / g z) holomorphic_on s" unfolding holomorphic_on_def by (metis field_differentiable_divide) lemma holomorphic_on_power [holomorphic_intros]: "f holomorphic_on s \ (\z. (f z)^n) holomorphic_on s" unfolding holomorphic_on_def by (metis field_differentiable_power) lemma holomorphic_on_sum [holomorphic_intros]: "(\i. i \ I \ (f i) holomorphic_on s) \ (\x. sum (\i. f i x) I) holomorphic_on s" unfolding holomorphic_on_def by (metis field_differentiable_sum) lemma holomorphic_on_prod [holomorphic_intros]: "(\i. i \ I \ (f i) holomorphic_on s) \ (\x. prod (\i. f i x) I) holomorphic_on s" by (induction I rule: infinite_finite_induct) (auto intro: holomorphic_intros) lemma holomorphic_pochhammer [holomorphic_intros]: "f holomorphic_on A \ (\s. pochhammer (f s) n) holomorphic_on A" by (induction n) (auto intro!: holomorphic_intros simp: pochhammer_Suc) lemma holomorphic_on_scaleR [holomorphic_intros]: "f holomorphic_on A \ (\x. c *\<^sub>R f x) holomorphic_on A" by (auto simp: scaleR_conv_of_real intro!: holomorphic_intros) lemma holomorphic_on_Un [holomorphic_intros]: assumes "f holomorphic_on A" "f holomorphic_on B" "open A" "open B" shows "f holomorphic_on (A \ B)" using assms by (auto simp: holomorphic_on_def at_within_open[of _ A] at_within_open[of _ B] at_within_open[of _ "A \ B"] open_Un) lemma holomorphic_on_If_Un [holomorphic_intros]: assumes "f holomorphic_on A" "g holomorphic_on B" "open A" "open B" assumes "\z. z \ A \ z \ B \ f z = g z" shows "(\z. if z \ A then f z else g z) holomorphic_on (A \ B)" (is "?h holomorphic_on _") proof (intro holomorphic_on_Un) note \f holomorphic_on A\ also have "f holomorphic_on A \ ?h holomorphic_on A" by (intro holomorphic_cong) auto finally show \ . next note \g holomorphic_on B\ also have "g holomorphic_on B \ ?h holomorphic_on B" using assms by (intro holomorphic_cong) auto finally show \ . qed (insert assms, auto) lemma holomorphic_derivI: "\f holomorphic_on S; open S; x \ S\ \ (f has_field_derivative deriv f x) (at x within T)" by (metis DERIV_deriv_iff_field_differentiable at_within_open holomorphic_on_def has_field_derivative_at_within) lemma complex_derivative_transform_within_open: "\f holomorphic_on s; g holomorphic_on s; open s; z \ s; \w. w \ s \ f w = g w\ \ deriv f z = deriv g z" unfolding holomorphic_on_def by (rule DERIV_imp_deriv) (metis DERIV_deriv_iff_field_differentiable has_field_derivative_transform_within_open at_within_open) lemma holomorphic_nonconstant: assumes holf: "f holomorphic_on S" and "open S" "\ \ S" "deriv f \ \ 0" shows "\ f constant_on S" by (rule nonzero_deriv_nonconstant [of f "deriv f \" \ S]) (use assms in \auto simp: holomorphic_derivI\) subsection\Analyticity on a set\ definition\<^marker>\tag important\ analytic_on (infixl "(analytic'_on)" 50) where "f analytic_on S \ \x \ S. \e. 0 < e \ f holomorphic_on (ball x e)" named_theorems\<^marker>\tag important\ analytic_intros "introduction rules for proving analyticity" lemma analytic_imp_holomorphic: "f analytic_on S \ f holomorphic_on S" by (simp add: at_within_open [OF _ open_ball] analytic_on_def holomorphic_on_def) (metis centre_in_ball field_differentiable_at_within) lemma analytic_on_open: "open S \ f analytic_on S \ f holomorphic_on S" apply (auto simp: analytic_imp_holomorphic) apply (auto simp: analytic_on_def holomorphic_on_def) by (metis holomorphic_on_def holomorphic_on_subset open_contains_ball) lemma analytic_on_imp_differentiable_at: "f analytic_on S \ x \ S \ f field_differentiable (at x)" apply (auto simp: analytic_on_def holomorphic_on_def) by (metis open_ball centre_in_ball field_differentiable_within_open) lemma analytic_on_subset: "f analytic_on S \ T \ S \ f analytic_on T" by (auto simp: analytic_on_def) lemma analytic_on_Un: "f analytic_on (S \ T) \ f analytic_on S \ f analytic_on T" by (auto simp: analytic_on_def) lemma analytic_on_Union: "f analytic_on (\\) \ (\T \ \. f analytic_on T)" by (auto simp: analytic_on_def) lemma analytic_on_UN: "f analytic_on (\i\I. S i) \ (\i\I. f analytic_on (S i))" by (auto simp: analytic_on_def) lemma analytic_on_holomorphic: "f analytic_on S \ (\T. open T \ S \ T \ f holomorphic_on T)" (is "?lhs = ?rhs") proof - have "?lhs \ (\T. open T \ S \ T \ f analytic_on T)" proof safe assume "f analytic_on S" then show "\T. open T \ S \ T \ f analytic_on T" apply (simp add: analytic_on_def) apply (rule exI [where x="\{U. open U \ f analytic_on U}"], auto) apply (metis open_ball analytic_on_open centre_in_ball) by (metis analytic_on_def) next fix T assume "open T" "S \ T" "f analytic_on T" then show "f analytic_on S" by (metis analytic_on_subset) qed also have "... \ ?rhs" by (auto simp: analytic_on_open) finally show ?thesis . qed lemma analytic_on_linear [analytic_intros,simp]: "((*) c) analytic_on S" by (auto simp add: analytic_on_holomorphic) lemma analytic_on_const [analytic_intros,simp]: "(\z. c) analytic_on S" by (metis analytic_on_def holomorphic_on_const zero_less_one) lemma analytic_on_ident [analytic_intros,simp]: "(\x. x) analytic_on S" by (simp add: analytic_on_def gt_ex) lemma analytic_on_id [analytic_intros]: "id analytic_on S" unfolding id_def by (rule analytic_on_ident) lemma analytic_on_compose: assumes f: "f analytic_on S" and g: "g analytic_on (f ` S)" shows "(g o f) analytic_on S" unfolding analytic_on_def proof (intro ballI) fix x assume x: "x \ S" then obtain e where e: "0 < e" and fh: "f holomorphic_on ball x e" using f by (metis analytic_on_def) obtain e' where e': "0 < e'" and gh: "g holomorphic_on ball (f x) e'" using g by (metis analytic_on_def g image_eqI x) have "isCont f x" by (metis analytic_on_imp_differentiable_at field_differentiable_imp_continuous_at f x) with e' obtain d where d: "0 < d" and fd: "f ` ball x d \ ball (f x) e'" by (auto simp: continuous_at_ball) have "g \ f holomorphic_on ball x (min d e)" apply (rule holomorphic_on_compose) apply (metis fh holomorphic_on_subset min.bounded_iff order_refl subset_ball) by (metis fd gh holomorphic_on_subset image_mono min.cobounded1 subset_ball) then show "\e>0. g \ f holomorphic_on ball x e" by (metis d e min_less_iff_conj) qed lemma analytic_on_compose_gen: "f analytic_on S \ g analytic_on T \ (\z. z \ S \ f z \ T) \ g o f analytic_on S" by (metis analytic_on_compose analytic_on_subset image_subset_iff) lemma analytic_on_neg [analytic_intros]: "f analytic_on S \ (\z. -(f z)) analytic_on S" by (metis analytic_on_holomorphic holomorphic_on_minus) lemma analytic_on_add [analytic_intros]: assumes f: "f analytic_on S" and g: "g analytic_on S" shows "(\z. f z + g z) analytic_on S" unfolding analytic_on_def proof (intro ballI) fix z assume z: "z \ S" then obtain e where e: "0 < e" and fh: "f holomorphic_on ball z e" using f by (metis analytic_on_def) obtain e' where e': "0 < e'" and gh: "g holomorphic_on ball z e'" using g by (metis analytic_on_def g z) have "(\z. f z + g z) holomorphic_on ball z (min e e')" apply (rule holomorphic_on_add) apply (metis fh holomorphic_on_subset min.bounded_iff order_refl subset_ball) by (metis gh holomorphic_on_subset min.bounded_iff order_refl subset_ball) then show "\e>0. (\z. f z + g z) holomorphic_on ball z e" by (metis e e' min_less_iff_conj) qed lemma analytic_on_diff [analytic_intros]: assumes f: "f analytic_on S" and g: "g analytic_on S" shows "(\z. f z - g z) analytic_on S" unfolding analytic_on_def proof (intro ballI) fix z assume z: "z \ S" then obtain e where e: "0 < e" and fh: "f holomorphic_on ball z e" using f by (metis analytic_on_def) obtain e' where e': "0 < e'" and gh: "g holomorphic_on ball z e'" using g by (metis analytic_on_def g z) have "(\z. f z - g z) holomorphic_on ball z (min e e')" apply (rule holomorphic_on_diff) apply (metis fh holomorphic_on_subset min.bounded_iff order_refl subset_ball) by (metis gh holomorphic_on_subset min.bounded_iff order_refl subset_ball) then show "\e>0. (\z. f z - g z) holomorphic_on ball z e" by (metis e e' min_less_iff_conj) qed lemma analytic_on_mult [analytic_intros]: assumes f: "f analytic_on S" and g: "g analytic_on S" shows "(\z. f z * g z) analytic_on S" unfolding analytic_on_def proof (intro ballI) fix z assume z: "z \ S" then obtain e where e: "0 < e" and fh: "f holomorphic_on ball z e" using f by (metis analytic_on_def) obtain e' where e': "0 < e'" and gh: "g holomorphic_on ball z e'" using g by (metis analytic_on_def g z) have "(\z. f z * g z) holomorphic_on ball z (min e e')" apply (rule holomorphic_on_mult) apply (metis fh holomorphic_on_subset min.bounded_iff order_refl subset_ball) by (metis gh holomorphic_on_subset min.bounded_iff order_refl subset_ball) then show "\e>0. (\z. f z * g z) holomorphic_on ball z e" by (metis e e' min_less_iff_conj) qed lemma analytic_on_inverse [analytic_intros]: assumes f: "f analytic_on S" and nz: "(\z. z \ S \ f z \ 0)" shows "(\z. inverse (f z)) analytic_on S" unfolding analytic_on_def proof (intro ballI) fix z assume z: "z \ S" then obtain e where e: "0 < e" and fh: "f holomorphic_on ball z e" using f by (metis analytic_on_def) have "continuous_on (ball z e) f" by (metis fh holomorphic_on_imp_continuous_on) then obtain e' where e': "0 < e'" and nz': "\y. dist z y < e' \ f y \ 0" by (metis open_ball centre_in_ball continuous_on_open_avoid e z nz) have "(\z. inverse (f z)) holomorphic_on ball z (min e e')" apply (rule holomorphic_on_inverse) apply (metis fh holomorphic_on_subset min.cobounded2 min.commute subset_ball) by (metis nz' mem_ball min_less_iff_conj) then show "\e>0. (\z. inverse (f z)) holomorphic_on ball z e" by (metis e e' min_less_iff_conj) qed lemma analytic_on_divide [analytic_intros]: assumes f: "f analytic_on S" and g: "g analytic_on S" and nz: "(\z. z \ S \ g z \ 0)" shows "(\z. f z / g z) analytic_on S" unfolding divide_inverse by (metis analytic_on_inverse analytic_on_mult f g nz) lemma analytic_on_power [analytic_intros]: "f analytic_on S \ (\z. (f z) ^ n) analytic_on S" by (induct n) (auto simp: analytic_on_mult) lemma analytic_on_sum [analytic_intros]: "(\i. i \ I \ (f i) analytic_on S) \ (\x. sum (\i. f i x) I) analytic_on S" by (induct I rule: infinite_finite_induct) (auto simp: analytic_on_add) lemma deriv_left_inverse: assumes "f holomorphic_on S" and "g holomorphic_on T" and "open S" and "open T" and "f ` S \ T" and [simp]: "\z. z \ S \ g (f z) = z" and "w \ S" shows "deriv f w * deriv g (f w) = 1" proof - have "deriv f w * deriv g (f w) = deriv g (f w) * deriv f w" by (simp add: algebra_simps) also have "... = deriv (g o f) w" using assms by (metis analytic_on_imp_differentiable_at analytic_on_open deriv_chain image_subset_iff) also have "... = deriv id w" proof (rule complex_derivative_transform_within_open [where s=S]) show "g \ f holomorphic_on S" by (rule assms holomorphic_on_compose_gen holomorphic_intros)+ qed (use assms in auto) also have "... = 1" by simp finally show ?thesis . qed subsection\<^marker>\tag unimportant\\Analyticity at a point\ lemma analytic_at_ball: "f analytic_on {z} \ (\e. 0 f holomorphic_on ball z e)" by (metis analytic_on_def singleton_iff) lemma analytic_at: "f analytic_on {z} \ (\s. open s \ z \ s \ f holomorphic_on s)" by (metis analytic_on_holomorphic empty_subsetI insert_subset) lemma analytic_on_analytic_at: "f analytic_on s \ (\z \ s. f analytic_on {z})" by (metis analytic_at_ball analytic_on_def) lemma analytic_at_two: "f analytic_on {z} \ g analytic_on {z} \ (\s. open s \ z \ s \ f holomorphic_on s \ g holomorphic_on s)" (is "?lhs = ?rhs") proof assume ?lhs then obtain s t where st: "open s" "z \ s" "f holomorphic_on s" "open t" "z \ t" "g holomorphic_on t" by (auto simp: analytic_at) show ?rhs apply (rule_tac x="s \ t" in exI) using st apply (auto simp: holomorphic_on_subset) done next assume ?rhs then show ?lhs by (force simp add: analytic_at) qed subsection\<^marker>\tag unimportant\\Combining theorems for derivative with ``analytic at'' hypotheses\ lemma assumes "f analytic_on {z}" "g analytic_on {z}" shows complex_derivative_add_at: "deriv (\w. f w + g w) z = deriv f z + deriv g z" and complex_derivative_diff_at: "deriv (\w. f w - g w) z = deriv f z - deriv g z" and complex_derivative_mult_at: "deriv (\w. f w * g w) z = f z * deriv g z + deriv f z * g z" proof - obtain s where s: "open s" "z \ s" "f holomorphic_on s" "g holomorphic_on s" using assms by (metis analytic_at_two) show "deriv (\w. f w + g w) z = deriv f z + deriv g z" apply (rule DERIV_imp_deriv [OF DERIV_add]) using s apply (auto simp: holomorphic_on_open field_differentiable_def DERIV_deriv_iff_field_differentiable) done show "deriv (\w. f w - g w) z = deriv f z - deriv g z" apply (rule DERIV_imp_deriv [OF DERIV_diff]) using s apply (auto simp: holomorphic_on_open field_differentiable_def DERIV_deriv_iff_field_differentiable) done show "deriv (\w. f w * g w) z = f z * deriv g z + deriv f z * g z" apply (rule DERIV_imp_deriv [OF DERIV_mult']) using s apply (auto simp: holomorphic_on_open field_differentiable_def DERIV_deriv_iff_field_differentiable) done qed lemma deriv_cmult_at: "f analytic_on {z} \ deriv (\w. c * f w) z = c * deriv f z" by (auto simp: complex_derivative_mult_at) lemma deriv_cmult_right_at: "f analytic_on {z} \ deriv (\w. f w * c) z = deriv f z * c" by (auto simp: complex_derivative_mult_at) subsection\<^marker>\tag unimportant\\Complex differentiation of sequences and series\ (* TODO: Could probably be simplified using Uniform_Limit *) lemma has_complex_derivative_sequence: fixes S :: "complex set" assumes cvs: "convex S" and df: "\n x. x \ S \ (f n has_field_derivative f' n x) (at x within S)" and conv: "\e. 0 < e \ \N. \n x. n \ N \ x \ S \ norm (f' n x - g' x) \ e" and "\x l. x \ S \ ((\n. f n x) \ l) sequentially" shows "\g. \x \ S. ((\n. f n x) \ g x) sequentially \ (g has_field_derivative (g' x)) (at x within S)" proof - from assms obtain x l where x: "x \ S" and tf: "((\n. f n x) \ l) sequentially" by blast { fix e::real assume e: "e > 0" then obtain N where N: "\n\N. \x. x \ S \ cmod (f' n x - g' x) \ e" by (metis conv) have "\N. \n\N. \x\S. \h. cmod (f' n x * h - g' x * h) \ e * cmod h" proof (rule exI [of _ N], clarify) fix n y h assume "N \ n" "y \ S" then have "cmod (f' n y - g' y) \ e" by (metis N) then have "cmod h * cmod (f' n y - g' y) \ cmod h * e" by (auto simp: antisym_conv2 mult_le_cancel_left norm_triangle_ineq2) then show "cmod (f' n y * h - g' y * h) \ e * cmod h" by (simp add: norm_mult [symmetric] field_simps) qed } note ** = this show ?thesis unfolding has_field_derivative_def proof (rule has_derivative_sequence [OF cvs _ _ x]) show "(\n. f n x) \ l" by (rule tf) next show "\e. e > 0 \ \\<^sub>F n in sequentially. \x\S. \h. cmod (f' n x * h - g' x * h) \ e * cmod h" unfolding eventually_sequentially by (blast intro: **) qed (metis has_field_derivative_def df) qed lemma has_complex_derivative_series: fixes S :: "complex set" assumes cvs: "convex S" and df: "\n x. x \ S \ (f n has_field_derivative f' n x) (at x within S)" and conv: "\e. 0 < e \ \N. \n x. n \ N \ x \ S \ cmod ((\i e" and "\x l. x \ S \ ((\n. f n x) sums l)" shows "\g. \x \ S. ((\n. f n x) sums g x) \ ((g has_field_derivative g' x) (at x within S))" proof - from assms obtain x l where x: "x \ S" and sf: "((\n. f n x) sums l)" by blast { fix e::real assume e: "e > 0" then obtain N where N: "\n x. n \ N \ x \ S \ cmod ((\i e" by (metis conv) have "\N. \n\N. \x\S. \h. cmod ((\i e * cmod h" proof (rule exI [of _ N], clarify) fix n y h assume "N \ n" "y \ S" then have "cmod ((\i e" by (metis N) then have "cmod h * cmod ((\i cmod h * e" by (auto simp: antisym_conv2 mult_le_cancel_left norm_triangle_ineq2) then show "cmod ((\i e * cmod h" by (simp add: norm_mult [symmetric] field_simps sum_distrib_left) qed } note ** = this show ?thesis unfolding has_field_derivative_def proof (rule has_derivative_series [OF cvs _ _ x]) fix n x assume "x \ S" then show "((f n) has_derivative (\z. z * f' n x)) (at x within S)" by (metis df has_field_derivative_def mult_commute_abs) next show " ((\n. f n x) sums l)" by (rule sf) next show "\e. e>0 \ \\<^sub>F n in sequentially. \x\S. \h. cmod ((\i e * cmod h" unfolding eventually_sequentially by (blast intro: **) qed qed subsection\<^marker>\tag unimportant\ \Taylor on Complex Numbers\ lemma sum_Suc_reindex: fixes f :: "nat \ 'a::ab_group_add" shows "sum f {0..n} = f 0 - f (Suc n) + sum (\i. f (Suc i)) {0..n}" by (induct n) auto lemma field_Taylor: assumes S: "convex S" and f: "\i x. x \ S \ i \ n \ (f i has_field_derivative f (Suc i) x) (at x within S)" and B: "\x. x \ S \ norm (f (Suc n) x) \ B" and w: "w \ S" and z: "z \ S" shows "norm(f 0 z - (\i\n. f i w * (z-w) ^ i / (fact i))) \ B * norm(z - w)^(Suc n) / fact n" proof - have wzs: "closed_segment w z \ S" using assms by (metis convex_contains_segment) { fix u assume "u \ closed_segment w z" then have "u \ S" by (metis wzs subsetD) have "(\i\n. f i u * (- of_nat i * (z-u)^(i - 1)) / (fact i) + f (Suc i) u * (z-u)^i / (fact i)) = f (Suc n) u * (z-u) ^ n / (fact n)" proof (induction n) case 0 show ?case by simp next case (Suc n) have "(\i\Suc n. f i u * (- of_nat i * (z-u) ^ (i - 1)) / (fact i) + f (Suc i) u * (z-u) ^ i / (fact i)) = f (Suc n) u * (z-u) ^ n / (fact n) + f (Suc (Suc n)) u * ((z-u) * (z-u) ^ n) / (fact (Suc n)) - f (Suc n) u * ((1 + of_nat n) * (z-u) ^ n) / (fact (Suc n))" using Suc by simp also have "... = f (Suc (Suc n)) u * (z-u) ^ Suc n / (fact (Suc n))" proof - have "(fact(Suc n)) * (f(Suc n) u *(z-u) ^ n / (fact n) + f(Suc(Suc n)) u *((z-u) *(z-u) ^ n) / (fact(Suc n)) - f(Suc n) u *((1 + of_nat n) *(z-u) ^ n) / (fact(Suc n))) = ((fact(Suc n)) *(f(Suc n) u *(z-u) ^ n)) / (fact n) + ((fact(Suc n)) *(f(Suc(Suc n)) u *((z-u) *(z-u) ^ n)) / (fact(Suc n))) - ((fact(Suc n)) *(f(Suc n) u *(of_nat(Suc n) *(z-u) ^ n))) / (fact(Suc n))" by (simp add: algebra_simps del: fact_Suc) also have "... = ((fact (Suc n)) * (f (Suc n) u * (z-u) ^ n)) / (fact n) + (f (Suc (Suc n)) u * ((z-u) * (z-u) ^ n)) - (f (Suc n) u * ((1 + of_nat n) * (z-u) ^ n))" by (simp del: fact_Suc) also have "... = (of_nat (Suc n) * (f (Suc n) u * (z-u) ^ n)) + (f (Suc (Suc n)) u * ((z-u) * (z-u) ^ n)) - (f (Suc n) u * ((1 + of_nat n) * (z-u) ^ n))" by (simp only: fact_Suc of_nat_mult ac_simps) simp also have "... = f (Suc (Suc n)) u * ((z-u) * (z-u) ^ n)" by (simp add: algebra_simps) finally show ?thesis by (simp add: mult_left_cancel [where c = "(fact (Suc n))", THEN iffD1] del: fact_Suc) qed finally show ?case . qed then have "((\v. (\i\n. f i v * (z - v)^i / (fact i))) has_field_derivative f (Suc n) u * (z-u) ^ n / (fact n)) (at u within S)" apply (intro derivative_eq_intros) apply (blast intro: assms \u \ S\) apply (rule refl)+ apply (auto simp: field_simps) done } note sum_deriv = this { fix u assume u: "u \ closed_segment w z" then have us: "u \ S" by (metis wzs subsetD) have "norm (f (Suc n) u) * norm (z - u) ^ n \ norm (f (Suc n) u) * norm (u - z) ^ n" by (metis norm_minus_commute order_refl) also have "... \ norm (f (Suc n) u) * norm (z - w) ^ n" by (metis mult_left_mono norm_ge_zero power_mono segment_bound [OF u]) also have "... \ B * norm (z - w) ^ n" by (metis norm_ge_zero zero_le_power mult_right_mono B [OF us]) finally have "norm (f (Suc n) u) * norm (z - u) ^ n \ B * norm (z - w) ^ n" . } note cmod_bound = this have "(\i\n. f i z * (z - z) ^ i / (fact i)) = (\i\n. (f i z / (fact i)) * 0 ^ i)" by simp also have "\ = f 0 z / (fact 0)" by (subst sum_zero_power) simp finally have "norm (f 0 z - (\i\n. f i w * (z - w) ^ i / (fact i))) \ norm ((\i\n. f i w * (z - w) ^ i / (fact i)) - (\i\n. f i z * (z - z) ^ i / (fact i)))" by (simp add: norm_minus_commute) also have "... \ B * norm (z - w) ^ n / (fact n) * norm (w - z)" apply (rule field_differentiable_bound [where f' = "\w. f (Suc n) w * (z - w)^n / (fact n)" and S = "closed_segment w z", OF convex_closed_segment]) apply (auto simp: DERIV_subset [OF sum_deriv wzs] norm_divide norm_mult norm_power divide_le_cancel cmod_bound) done also have "... \ B * norm (z - w) ^ Suc n / (fact n)" by (simp add: algebra_simps norm_minus_commute) finally show ?thesis . qed lemma complex_Taylor: assumes S: "convex S" and f: "\i x. x \ S \ i \ n \ (f i has_field_derivative f (Suc i) x) (at x within S)" and B: "\x. x \ S \ cmod (f (Suc n) x) \ B" and w: "w \ S" and z: "z \ S" shows "cmod(f 0 z - (\i\n. f i w * (z-w) ^ i / (fact i))) \ B * cmod(z - w)^(Suc n) / fact n" using assms by (rule field_Taylor) text\Something more like the traditional MVT for real components\ lemma complex_mvt_line: assumes "\u. u \ closed_segment w z \ (f has_field_derivative f'(u)) (at u)" shows "\u. u \ closed_segment w z \ Re(f z) - Re(f w) = Re(f'(u) * (z - w))" proof - have twz: "\t. (1 - t) *\<^sub>R w + t *\<^sub>R z = w + t *\<^sub>R (z - w)" by (simp add: real_vector.scale_left_diff_distrib real_vector.scale_right_diff_distrib) note assms[unfolded has_field_derivative_def, derivative_intros] show ?thesis apply (cut_tac mvt_simple [of 0 1 "Re o f o (\t. (1 - t) *\<^sub>R w + t *\<^sub>R z)" "\u. Re o (\h. f'((1 - u) *\<^sub>R w + u *\<^sub>R z) * h) o (\t. t *\<^sub>R (z - w))"]) apply auto apply (rule_tac x="(1 - x) *\<^sub>R w + x *\<^sub>R z" in exI) apply (auto simp: closed_segment_def twz) [] apply (intro derivative_eq_intros has_derivative_at_withinI, simp_all) apply (simp add: fun_eq_iff real_vector.scale_right_diff_distrib) apply (force simp: twz closed_segment_def) done qed lemma complex_Taylor_mvt: assumes "\i x. \x \ closed_segment w z; i \ n\ \ ((f i) has_field_derivative f (Suc i) x) (at x)" shows "\u. u \ closed_segment w z \ Re (f 0 z) = Re ((\i = 0..n. f i w * (z - w) ^ i / (fact i)) + (f (Suc n) u * (z-u)^n / (fact n)) * (z - w))" proof - { fix u assume u: "u \ closed_segment w z" have "(\i = 0..n. (f (Suc i) u * (z-u) ^ i - of_nat i * (f i u * (z-u) ^ (i - Suc 0))) / (fact i)) = f (Suc 0) u - (f (Suc (Suc n)) u * ((z-u) ^ Suc n) - (of_nat (Suc n)) * (z-u) ^ n * f (Suc n) u) / (fact (Suc n)) + (\i = 0..n. (f (Suc (Suc i)) u * ((z-u) ^ Suc i) - of_nat (Suc i) * (f (Suc i) u * (z-u) ^ i)) / (fact (Suc i)))" by (subst sum_Suc_reindex) simp also have "... = f (Suc 0) u - (f (Suc (Suc n)) u * ((z-u) ^ Suc n) - (of_nat (Suc n)) * (z-u) ^ n * f (Suc n) u) / (fact (Suc n)) + (\i = 0..n. f (Suc (Suc i)) u * ((z-u) ^ Suc i) / (fact (Suc i)) - f (Suc i) u * (z-u) ^ i / (fact i))" by (simp only: diff_divide_distrib fact_cancel ac_simps) also have "... = f (Suc 0) u - (f (Suc (Suc n)) u * (z-u) ^ Suc n - of_nat (Suc n) * (z-u) ^ n * f (Suc n) u) / (fact (Suc n)) + f (Suc (Suc n)) u * (z-u) ^ Suc n / (fact (Suc n)) - f (Suc 0) u" by (subst sum_Suc_diff) auto also have "... = f (Suc n) u * (z-u) ^ n / (fact n)" by (simp only: algebra_simps diff_divide_distrib fact_cancel) finally have "(\i = 0..n. (f (Suc i) u * (z - u) ^ i - of_nat i * (f i u * (z-u) ^ (i - Suc 0))) / (fact i)) = f (Suc n) u * (z - u) ^ n / (fact n)" . then have "((\u. \i = 0..n. f i u * (z - u) ^ i / (fact i)) has_field_derivative f (Suc n) u * (z - u) ^ n / (fact n)) (at u)" apply (intro derivative_eq_intros)+ apply (force intro: u assms) apply (rule refl)+ apply (auto simp: ac_simps) done } then show ?thesis apply (cut_tac complex_mvt_line [of w z "\u. \i = 0..n. f i u * (z-u) ^ i / (fact i)" "\u. (f (Suc n) u * (z-u)^n / (fact n))"]) apply (auto simp add: intro: open_closed_segment) done qed end diff --git a/src/HOL/Analysis/Complex_Transcendental.thy b/src/HOL/Analysis/Complex_Transcendental.thy --- a/src/HOL/Analysis/Complex_Transcendental.thy +++ b/src/HOL/Analysis/Complex_Transcendental.thy @@ -1,4100 +1,4100 @@ section \Complex Transcendental Functions\ text\By John Harrison et al. Ported from HOL Light by L C Paulson (2015)\ theory Complex_Transcendental imports Complex_Analysis_Basics Summation_Tests "HOL-Library.Periodic_Fun" begin subsection\Möbius transformations\ (* TODO: Figure out what to do with Möbius transformations *) definition\<^marker>\tag important\ "moebius a b c d \ (\z. (a*z+b) / (c*z+d :: 'a :: field))" theorem moebius_inverse: assumes "a * d \ b * c" "c * z + d \ 0" shows "moebius d (-b) (-c) a (moebius a b c d z) = z" proof - from assms have "(-c) * moebius a b c d z + a \ 0" unfolding moebius_def by (simp add: field_simps) with assms show ?thesis unfolding moebius_def by (simp add: moebius_def divide_simps) (simp add: algebra_simps)? qed lemma moebius_inverse': assumes "a * d \ b * c" "c * z - a \ 0" shows "moebius a b c d (moebius d (-b) (-c) a z) = z" using assms moebius_inverse[of d a "-b" "-c" z] by (auto simp: algebra_simps) lemma cmod_add_real_less: assumes "Im z \ 0" "r\0" shows "cmod (z + r) < cmod z + \r\" proof (cases z) case (Complex x y) then have "0 < y * y" using assms mult_neg_neg by force with assms have "r * x / \r\ < sqrt (x*x + y*y)" by (simp add: real_less_rsqrt power2_eq_square) then show ?thesis using assms Complex apply (simp add: cmod_def) apply (rule power2_less_imp_less, auto) apply (simp add: power2_eq_square field_simps) done qed lemma cmod_diff_real_less: "Im z \ 0 \ x\0 \ cmod (z - x) < cmod z + \x\" using cmod_add_real_less [of z "-x"] by simp lemma cmod_square_less_1_plus: assumes "Im z = 0 \ \Re z\ < 1" shows "(cmod z)\<^sup>2 < 1 + cmod (1 - z\<^sup>2)" proof (cases "Im z = 0 \ Re z = 0") case True with assms abs_square_less_1 show ?thesis by (force simp add: Re_power2 Im_power2 cmod_def) next case False with cmod_diff_real_less [of "1 - z\<^sup>2" "1"] show ?thesis by (simp add: norm_power Im_power2) qed subsection\<^marker>\tag unimportant\\The Exponential Function\ lemma norm_exp_i_times [simp]: "norm (exp(\ * of_real y)) = 1" by simp lemma norm_exp_imaginary: "norm(exp z) = 1 \ Re z = 0" by simp lemma field_differentiable_within_exp: "exp field_differentiable (at z within s)" using DERIV_exp field_differentiable_at_within field_differentiable_def by blast 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) lemma holomorphic_on_exp [holomorphic_intros]: "exp holomorphic_on s" by (simp add: field_differentiable_within_exp holomorphic_on_def) lemma holomorphic_on_exp' [holomorphic_intros]: "f holomorphic_on s \ (\x. exp (f x)) holomorphic_on s" using holomorphic_on_compose[OF _ holomorphic_on_exp] by (simp add: o_def) subsection\Euler and de Moivre formulas\ text\The sine series times \<^term>\i\\ lemma sin_i_eq: "(\n. (\ * sin_coeff n) * z^n) sums (\ * sin z)" proof - have "(\n. \ * sin_coeff n *\<^sub>R z^n) sums (\ * sin z)" using sin_converges sums_mult by blast then show ?thesis by (simp add: scaleR_conv_of_real field_simps) qed theorem exp_Euler: "exp(\ * z) = cos(z) + \ * sin(z)" proof - have "(\n. (cos_coeff n + \ * sin_coeff n) * z^n) = (\n. (\ * z) ^ n /\<^sub>R (fact n))" proof fix n show "(cos_coeff n + \ * sin_coeff n) * z^n = (\ * z) ^ n /\<^sub>R (fact n)" by (auto simp: cos_coeff_def sin_coeff_def scaleR_conv_of_real field_simps elim!: evenE oddE) qed also have "... sums (exp (\ * z))" by (rule exp_converges) finally have "(\n. (cos_coeff n + \ * sin_coeff n) * z^n) sums (exp (\ * z))" . moreover have "(\n. (cos_coeff n + \ * sin_coeff n) * z^n) sums (cos z + \ * sin z)" using sums_add [OF cos_converges [of z] sin_i_eq [of z]] by (simp add: field_simps scaleR_conv_of_real) ultimately show ?thesis using sums_unique2 by blast qed corollary\<^marker>\tag unimportant\ exp_minus_Euler: "exp(-(\ * z)) = cos(z) - \ * sin(z)" using exp_Euler [of "-z"] by simp lemma sin_exp_eq: "sin z = (exp(\ * z) - exp(-(\ * z))) / (2*\)" by (simp add: exp_Euler exp_minus_Euler) lemma sin_exp_eq': "sin z = \ * (exp(-(\ * z)) - exp(\ * z)) / 2" by (simp add: exp_Euler exp_minus_Euler) lemma cos_exp_eq: "cos z = (exp(\ * z) + exp(-(\ * z))) / 2" by (simp add: exp_Euler exp_minus_Euler) 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) 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) lemma Im_sin: "Im(sin z) = cos(Re z) * (exp(Im z) - exp(-(Im z))) / 2" by (simp add: sin_exp_eq field_simps Im_divide Re_exp) lemma Re_cos: "Re(cos z) = cos(Re z) * (exp(Im z) + exp(-(Im z))) / 2" by (simp add: cos_exp_eq field_simps Re_divide Re_exp) lemma Im_cos: "Im(cos z) = sin(Re z) * (exp(-(Im z)) - exp(Im z)) / 2" by (simp add: cos_exp_eq field_simps Im_divide Im_exp) lemma Re_sin_pos: "0 < Re z \ Re z < pi \ Re (sin z) > 0" by (auto simp: Re_sin Im_sin add_pos_pos sin_gt_zero) lemma Im_sin_nonneg: "Re z = 0 \ 0 \ Im z \ 0 \ Im (sin z)" by (simp add: Re_sin Im_sin algebra_simps) lemma Im_sin_nonneg2: "Re z = pi \ Im z \ 0 \ 0 \ Im (sin z)" by (simp add: Re_sin Im_sin algebra_simps) subsection\<^marker>\tag unimportant\\Relationships between real and complex trigonometric and hyperbolic functions\ lemma real_sin_eq [simp]: "Re(sin(of_real x)) = sin x" by (simp add: sin_of_real) lemma real_cos_eq [simp]: "Re(cos(of_real x)) = cos x" by (simp add: cos_of_real) lemma DeMoivre: "(cos z + \ * sin z) ^ n = cos(n * z) + \ * sin(n * z)" by (metis exp_Euler [symmetric] exp_of_nat_mult mult.left_commute) lemma exp_cnj: "cnj (exp z) = exp (cnj z)" proof - have "(\n. cnj (z ^ n /\<^sub>R (fact n))) = (\n. (cnj z)^n /\<^sub>R (fact n))" by auto also have "... sums (exp (cnj z))" by (rule exp_converges) finally have "(\n. cnj (z ^ n /\<^sub>R (fact n))) sums (exp (cnj z))" . moreover have "(\n. cnj (z ^ n /\<^sub>R (fact n))) sums (cnj (exp z))" by (metis exp_converges sums_cnj) ultimately show ?thesis using sums_unique2 by blast qed lemma cnj_sin: "cnj(sin z) = sin(cnj z)" by (simp add: sin_exp_eq exp_cnj field_simps) lemma cnj_cos: "cnj(cos z) = cos(cnj z)" by (simp add: cos_exp_eq exp_cnj field_simps) lemma field_differentiable_at_sin: "sin field_differentiable at z" using DERIV_sin field_differentiable_def by blast lemma field_differentiable_within_sin: "sin field_differentiable (at z within S)" by (simp add: field_differentiable_at_sin field_differentiable_at_within) lemma field_differentiable_at_cos: "cos field_differentiable at z" using DERIV_cos field_differentiable_def by blast lemma field_differentiable_within_cos: "cos field_differentiable (at z within S)" by (simp add: field_differentiable_at_cos field_differentiable_at_within) lemma holomorphic_on_sin: "sin holomorphic_on S" by (simp add: field_differentiable_within_sin holomorphic_on_def) lemma holomorphic_on_cos: "cos holomorphic_on S" by (simp add: field_differentiable_within_cos holomorphic_on_def) lemma holomorphic_on_sin' [holomorphic_intros]: assumes "f holomorphic_on A" shows "(\x. sin (f x)) holomorphic_on A" using holomorphic_on_compose[OF assms holomorphic_on_sin] by (simp add: o_def) lemma holomorphic_on_cos' [holomorphic_intros]: assumes "f holomorphic_on A" shows "(\x. cos (f x)) holomorphic_on A" using holomorphic_on_compose[OF assms holomorphic_on_cos] by (simp add: o_def) subsection\<^marker>\tag unimportant\\More on the Polar Representation of Complex Numbers\ lemma exp_Complex: "exp(Complex r t) = of_real(exp r) * Complex (cos t) (sin t)" by (simp add: Complex_eq exp_add exp_Euler exp_of_real sin_of_real cos_of_real) lemma exp_eq_1: "exp z = 1 \ Re(z) = 0 \ (\n::int. Im(z) = of_int (2 * n) * pi)" (is "?lhs = ?rhs") proof assume "exp z = 1" then have "Re z = 0" by (metis exp_eq_one_iff norm_exp_eq_Re norm_one) with \?lhs\ show ?rhs by (metis Re_exp complex_Re_of_int cos_one_2pi_int exp_zero mult.commute mult_numeral_1 numeral_One of_int_mult of_int_numeral) next assume ?rhs then show ?lhs using Im_exp Re_exp complex_eq_iff by (simp add: cos_one_2pi_int cos_one_sin_zero mult.commute) qed lemma exp_eq: "exp w = exp z \ (\n::int. w = z + (of_int (2 * n) * pi) * \)" (is "?lhs = ?rhs") proof - have "exp w = exp z \ exp (w-z) = 1" by (simp add: exp_diff) also have "... \ (Re w = Re z \ (\n::int. Im w - Im z = of_int (2 * n) * pi))" by (simp add: exp_eq_1) also have "... \ ?rhs" by (auto simp: algebra_simps intro!: complex_eqI) finally show ?thesis . qed lemma exp_complex_eqI: "\Im w - Im z\ < 2*pi \ exp w = exp z \ w = z" by (auto simp: exp_eq abs_mult) lemma exp_integer_2pi: assumes "n \ \" shows "exp((2 * n * pi) * \) = 1" proof - have "exp((2 * n * pi) * \) = exp 0" using assms unfolding Ints_def exp_eq by auto also have "... = 1" by simp finally show ?thesis . qed lemma exp_plus_2pin [simp]: "exp (z + \ * (of_int n * (of_real pi * 2))) = exp z" by (simp add: exp_eq) lemma exp_integer_2pi_plus1: assumes "n \ \" shows "exp(((2 * n + 1) * pi) * \) = - 1" proof - from assms obtain n' where [simp]: "n = of_int n'" by (auto simp: Ints_def) have "exp(((2 * n + 1) * pi) * \) = exp (pi * \)" using assms by (subst exp_eq) (auto intro!: exI[of _ n'] simp: algebra_simps) also have "... = - 1" by simp finally show ?thesis . qed lemma inj_on_exp_pi: fixes z::complex shows "inj_on exp (ball z pi)" proof (clarsimp simp: inj_on_def exp_eq) fix y n assume "dist z (y + 2 * of_int n * of_real pi * \) < pi" "dist z y < pi" then have "dist y (y + 2 * of_int n * of_real pi * \) < pi+pi" using dist_commute_lessI dist_triangle_less_add by blast then have "norm (2 * of_int n * of_real pi * \) < 2*pi" by (simp add: dist_norm) then show "n = 0" by (auto simp: norm_mult) qed lemma cmod_add_squared: fixes r1 r2::real assumes "r1 \ 0" "r2 \ 0" shows "(cmod (r1 * exp (\ * \1) + r2 * exp (\ * \2)))\<^sup>2 = r1\<^sup>2 + r2\<^sup>2 + 2 * r1 * r2 * cos (\1 - \2)" (is "(cmod (?z1 + ?z2))\<^sup>2 = ?rhs") proof - have "(cmod (?z1 + ?z2))\<^sup>2 = (?z1 + ?z2) * cnj (?z1 + ?z2)" by (rule complex_norm_square) also have "\ = (?z1 * cnj ?z1 + ?z2 * cnj ?z2) + (?z1 * cnj ?z2 + cnj ?z1 * ?z2)" by (simp add: algebra_simps) also have "\ = (norm ?z1)\<^sup>2 + (norm ?z2)\<^sup>2 + 2 * Re (?z1 * cnj ?z2)" unfolding complex_norm_square [symmetric] cnj_add_mult_eq_Re by simp also have "\ = ?rhs" by (simp add: norm_mult) (simp add: exp_Euler complex_is_Real_iff [THEN iffD1] cos_diff algebra_simps) finally show ?thesis using of_real_eq_iff by blast qed lemma cmod_diff_squared: fixes r1 r2::real assumes "r1 \ 0" "r2 \ 0" shows "(cmod (r1 * exp (\ * \1) - r2 * exp (\ * \2)))\<^sup>2 = r1\<^sup>2 + r2\<^sup>2 - 2*r1*r2*cos (\1 - \2)" (is "(cmod (?z1 - ?z2))\<^sup>2 = ?rhs") proof - have "exp (\ * (\2 + pi)) = - exp (\ * \2)" by (simp add: exp_Euler cos_plus_pi sin_plus_pi) then have "(cmod (?z1 - ?z2))\<^sup>2 = cmod (?z1 + r2 * exp (\ * (\2 + pi))) ^2" by simp also have "\ = r1\<^sup>2 + r2\<^sup>2 + 2*r1*r2*cos (\1 - (\2 + pi))" using assms cmod_add_squared by blast also have "\ = ?rhs" by (simp add: add.commute diff_add_eq_diff_diff_swap) finally show ?thesis . qed lemma polar_convergence: fixes R::real assumes "\j. r j > 0" "R > 0" shows "((\j. r j * exp (\ * \ j)) \ (R * exp (\ * \))) \ (r \ R) \ (\k. (\j. \ j - of_int (k j) * (2 * pi)) \ \)" (is "(?z \ ?Z) = ?rhs") proof assume L: "?z \ ?Z" have rR: "r \ R" using tendsto_norm [OF L] assms by (auto simp: norm_mult abs_of_pos) moreover obtain k where "(\j. \ j - of_int (k j) * (2 * pi)) \ \" proof - have "cos (\ j - \) = ((r j)\<^sup>2 + R\<^sup>2 - (norm(?z j - ?Z))\<^sup>2) / (2 * R * r j)" for j apply (subst cmod_diff_squared) using assms by (auto simp: field_split_simps less_le) moreover have "(\j. ((r j)\<^sup>2 + R\<^sup>2 - (norm(?z j - ?Z))\<^sup>2) / (2 * R * r j)) \ ((R\<^sup>2 + R\<^sup>2 - (norm(?Z - ?Z))\<^sup>2) / (2 * R * R))" by (intro L rR tendsto_intros) (use \R > 0\ in force) moreover have "((R\<^sup>2 + R\<^sup>2 - (norm(?Z - ?Z))\<^sup>2) / (2 * R * R)) = 1" using \R > 0\ by (simp add: power2_eq_square field_split_simps) ultimately have "(\j. cos (\ j - \)) \ 1" by auto then show ?thesis using that cos_diff_limit_1 by blast qed ultimately show ?rhs by metis next assume R: ?rhs show "?z \ ?Z" proof (rule tendsto_mult) show "(\x. complex_of_real (r x)) \ of_real R" using R by (auto simp: tendsto_of_real_iff) obtain k where "(\j. \ j - of_int (k j) * (2 * pi)) \ \" using R by metis then have "(\j. complex_of_real (\ j - of_int (k j) * (2 * pi))) \ of_real \" using tendsto_of_real_iff by force then have "(\j. exp (\ * of_real (\ j - of_int (k j) * (2 * pi)))) \ exp (\ * \)" using tendsto_mult [OF tendsto_const] isCont_exp isCont_tendsto_compose by blast moreover have "exp (\ * of_real (\ j - of_int (k j) * (2 * pi))) = exp (\ * \ j)" for j unfolding exp_eq by (rule_tac x="- k j" in exI) (auto simp: algebra_simps) ultimately show "(\j. exp (\ * \ j)) \ exp (\ * \)" by auto qed qed lemma sin_cos_eq_iff: "sin y = sin x \ cos y = cos x \ (\n::int. y = x + 2 * pi * n)" proof - { assume "sin y = sin x" "cos y = cos x" then have "cos (y-x) = 1" using cos_add [of y "-x"] by simp then have "\n::int. y-x = 2 * pi * n" using cos_one_2pi_int by auto } then show ?thesis apply (auto simp: sin_add cos_add) apply (metis add.commute diff_add_cancel) done qed lemma exp_i_ne_1: assumes "0 < x" "x < 2*pi" shows "exp(\ * of_real x) \ 1" proof assume "exp (\ * of_real x) = 1" then have "exp (\ * of_real x) = exp 0" by simp then obtain n where "\ * of_real x = (of_int (2 * n) * pi) * \" by (simp only: Ints_def exp_eq) auto then have "of_real x = (of_int (2 * n) * pi)" by (metis complex_i_not_zero mult.commute mult_cancel_left of_real_eq_iff real_scaleR_def scaleR_conv_of_real) then have "x = (of_int (2 * n) * pi)" by simp then show False using assms by (cases n) (auto simp: zero_less_mult_iff mult_less_0_iff) qed lemma sin_eq_0: fixes z::complex shows "sin z = 0 \ (\n::int. z = of_real(n * pi))" by (simp add: sin_exp_eq exp_eq) lemma cos_eq_0: fixes z::complex shows "cos z = 0 \ (\n::int. z = of_real(n * pi) + of_real pi/2)" using sin_eq_0 [of "z - of_real pi/2"] by (simp add: sin_diff algebra_simps) lemma cos_eq_1: fixes z::complex shows "cos z = 1 \ (\n::int. z = of_real(2 * n * pi))" proof - have "cos z = cos (2*(z/2))" by simp also have "... = 1 - 2 * sin (z/2) ^ 2" by (simp only: cos_double_sin) finally have [simp]: "cos z = 1 \ sin (z/2) = 0" by simp show ?thesis by (auto simp: sin_eq_0) qed lemma csin_eq_1: fixes z::complex shows "sin z = 1 \ (\n::int. z = of_real(2 * n * pi) + of_real pi/2)" using cos_eq_1 [of "z - of_real pi/2"] by (simp add: cos_diff algebra_simps) lemma csin_eq_minus1: fixes z::complex shows "sin z = -1 \ (\n::int. z = of_real(2 * n * pi) + 3/2*pi)" (is "_ = ?rhs") proof - have "sin z = -1 \ sin (-z) = 1" by (simp add: equation_minus_iff) also have "... \ (\n::int. -z = of_real(2 * n * pi) + of_real pi/2)" by (simp only: csin_eq_1) also have "... \ (\n::int. z = - of_real(2 * n * pi) - of_real pi/2)" by (rule iff_exI) (metis add.inverse_inverse add_uminus_conv_diff minus_add_distrib) also have "... = ?rhs" apply safe apply (rule_tac [2] x="-(x+1)" in exI) apply (rule_tac x="-(x+1)" in exI) apply (simp_all add: algebra_simps) done finally show ?thesis . qed lemma ccos_eq_minus1: fixes z::complex shows "cos z = -1 \ (\n::int. z = of_real(2 * n * pi) + pi)" using csin_eq_1 [of "z - of_real pi/2"] by (simp add: sin_diff algebra_simps equation_minus_iff) lemma sin_eq_1: "sin x = 1 \ (\n::int. x = (2 * n + 1 / 2) * pi)" (is "_ = ?rhs") proof - have "sin x = 1 \ sin (complex_of_real x) = 1" by (metis of_real_1 one_complex.simps(1) real_sin_eq sin_of_real) also have "... \ (\n::int. complex_of_real x = of_real(2 * n * pi) + of_real pi/2)" by (simp only: csin_eq_1) also have "... \ (\n::int. x = of_real(2 * n * pi) + of_real pi/2)" by (rule iff_exI) (auto simp: algebra_simps intro: injD [OF inj_of_real [where 'a = complex]]) also have "... = ?rhs" by (auto simp: algebra_simps) finally show ?thesis . qed lemma sin_eq_minus1: "sin x = -1 \ (\n::int. x = (2*n + 3/2) * pi)" (is "_ = ?rhs") proof - have "sin x = -1 \ sin (complex_of_real x) = -1" by (metis Re_complex_of_real of_real_def scaleR_minus1_left sin_of_real) also have "... \ (\n::int. complex_of_real x = of_real(2 * n * pi) + 3/2*pi)" by (simp only: csin_eq_minus1) also have "... \ (\n::int. x = of_real(2 * n * pi) + 3/2*pi)" by (rule iff_exI) (auto simp: algebra_simps intro: injD [OF inj_of_real [where 'a = complex]]) also have "... = ?rhs" by (auto simp: algebra_simps) finally show ?thesis . qed lemma cos_eq_minus1: "cos x = -1 \ (\n::int. x = (2*n + 1) * pi)" (is "_ = ?rhs") proof - have "cos x = -1 \ cos (complex_of_real x) = -1" by (metis Re_complex_of_real of_real_def scaleR_minus1_left cos_of_real) also have "... \ (\n::int. complex_of_real x = of_real(2 * n * pi) + pi)" by (simp only: ccos_eq_minus1) also have "... \ (\n::int. x = of_real(2 * n * pi) + pi)" by (rule iff_exI) (auto simp: algebra_simps intro: injD [OF inj_of_real [where 'a = complex]]) also have "... = ?rhs" by (auto simp: algebra_simps) finally show ?thesis . qed lemma dist_exp_i_1: "norm(exp(\ * of_real t) - 1) = 2 * \sin(t / 2)\" proof - have "sqrt (2 - cos t * 2) = 2 * \sin (t / 2)\" using cos_double_sin [of "t/2"] by (simp add: real_sqrt_mult) then show ?thesis by (simp add: exp_Euler cmod_def power2_diff sin_of_real cos_of_real algebra_simps) qed lemma sin_cx_2pi [simp]: "\z = of_int m; even m\ \ sin (z * complex_of_real pi) = 0" by (simp add: sin_eq_0) lemma cos_cx_2pi [simp]: "\z = of_int m; even m\ \ cos (z * complex_of_real pi) = 1" using cos_eq_1 by auto lemma complex_sin_eq: fixes w :: complex shows "sin w = sin z \ (\n \ \. w = z + of_real(2*n*pi) \ w = -z + of_real((2*n + 1)*pi))" (is "?lhs = ?rhs") proof assume ?lhs then have "sin w - sin z = 0" by (auto simp: algebra_simps) then have "sin ((w - z) / 2)*cos ((w + z) / 2) = 0" by (auto simp: sin_diff_sin) then consider "sin ((w - z) / 2) = 0" | "cos ((w + z) / 2) = 0" using mult_eq_0_iff by blast then show ?rhs proof cases case 1 then show ?thesis by (simp add: sin_eq_0 algebra_simps) (metis Ints_of_int of_real_of_int_eq) next case 2 then show ?thesis by (simp add: cos_eq_0 algebra_simps) (metis Ints_of_int of_real_of_int_eq) qed next assume ?rhs then consider n::int where "w = z + of_real (2 * of_int n * pi)" | n::int where " w = -z + of_real ((2 * of_int n + 1) * pi)" using Ints_cases by blast then show ?lhs proof cases case 1 then show ?thesis using Periodic_Fun.sin.plus_of_int [of z n] by (auto simp: algebra_simps) next case 2 then show ?thesis using Periodic_Fun.sin.plus_of_int [of "-z" "n"] apply (simp add: algebra_simps) by (metis add.commute add.inverse_inverse add_diff_cancel_left diff_add_cancel sin_plus_pi) qed qed lemma complex_cos_eq: fixes w :: complex shows "cos w = cos z \ (\n \ \. w = z + of_real(2*n*pi) \ w = -z + of_real(2*n*pi))" (is "?lhs = ?rhs") proof assume ?lhs then have "cos w - cos z = 0" by (auto simp: algebra_simps) then have "sin ((w + z) / 2) * sin ((z - w) / 2) = 0" by (auto simp: cos_diff_cos) then consider "sin ((w + z) / 2) = 0" | "sin ((z - w) / 2) = 0" using mult_eq_0_iff by blast then show ?rhs proof cases case 1 then obtain n where "w + z = of_int n * (complex_of_real pi * 2)" by (auto simp: sin_eq_0 algebra_simps) then have "w = -z + of_real(2 * of_int n * pi)" by (auto simp: algebra_simps) then show ?thesis using Ints_of_int by blast next case 2 then obtain n where "z = w + of_int n * (complex_of_real pi * 2)" by (auto simp: sin_eq_0 algebra_simps) then have "w = z + complex_of_real (2 * of_int(-n) * pi)" by (auto simp: algebra_simps) then show ?thesis using Ints_of_int by blast qed next assume ?rhs then obtain n::int where w: "w = z + of_real (2* of_int n*pi) \ w = -z + of_real(2*n*pi)" using Ints_cases by (metis of_int_mult of_int_numeral) then show ?lhs using Periodic_Fun.cos.plus_of_int [of z n] apply (simp add: algebra_simps) by (metis cos.plus_of_int cos_minus minus_add_cancel mult.commute) qed lemma sin_eq: "sin x = sin y \ (\n \ \. x = y + 2*n*pi \ x = -y + (2*n + 1)*pi)" using complex_sin_eq [of x y] by (simp only: sin_of_real Re_complex_of_real of_real_add [symmetric] of_real_minus [symmetric] of_real_mult [symmetric] of_real_eq_iff) 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) lemma sinh_complex: fixes z :: complex shows "(exp z - inverse (exp z)) / 2 = -\ * sin(\ * z)" by (simp add: sin_exp_eq field_split_simps exp_minus) lemma sin_i_times: fixes z :: complex shows "sin(\ * z) = \ * ((exp z - inverse (exp z)) / 2)" using sinh_complex by auto lemma sinh_real: fixes x :: real shows "of_real((exp x - inverse (exp x)) / 2) = -\ * sin(\ * of_real x)" by (simp add: exp_of_real sin_i_times) lemma cosh_complex: fixes z :: complex shows "(exp z + inverse (exp z)) / 2 = cos(\ * z)" by (simp add: cos_exp_eq field_split_simps exp_minus exp_of_real) lemma cosh_real: fixes x :: real shows "of_real((exp x + inverse (exp x)) / 2) = cos(\ * of_real x)" by (simp add: cos_exp_eq field_split_simps exp_minus exp_of_real) lemmas cos_i_times = cosh_complex [symmetric] lemma norm_cos_squared: "norm(cos z) ^ 2 = cos(Re z) ^ 2 + (exp(Im z) - inverse(exp(Im z))) ^ 2 / 4" proof (cases z) case (Complex x1 x2) then show ?thesis apply (simp only: cos_add cmod_power2 cos_of_real sin_of_real Complex_eq) apply (simp add: cos_exp_eq sin_exp_eq exp_minus exp_of_real Re_divide Im_divide power_divide) apply (simp only: left_diff_distrib [symmetric] power_mult_distrib sin_squared_eq) apply (simp add: power2_eq_square field_split_simps) done qed lemma norm_sin_squared: "norm(sin z) ^ 2 = (exp(2 * Im z) + inverse(exp(2 * Im z)) - 2 * cos(2 * Re z)) / 4" proof (cases z) case (Complex x1 x2) then show ?thesis apply (simp only: sin_add cmod_power2 cos_of_real sin_of_real cos_double_cos exp_double Complex_eq) apply (simp add: cos_exp_eq sin_exp_eq exp_minus exp_of_real Re_divide Im_divide power_divide) apply (simp only: left_diff_distrib [symmetric] power_mult_distrib cos_squared_eq) apply (simp add: power2_eq_square field_split_simps) done qed lemma exp_uminus_Im: "exp (- Im z) \ exp (cmod z)" using abs_Im_le_cmod linear order_trans by fastforce lemma norm_cos_le: fixes z::complex shows "norm(cos z) \ exp(norm z)" proof - have "Im z \ cmod z" using abs_Im_le_cmod abs_le_D1 by auto then have "exp (- Im z) + exp (Im z) \ exp (cmod z) * 2" by (metis exp_uminus_Im add_mono exp_le_cancel_iff mult_2_right) then show ?thesis by (force simp add: cos_exp_eq norm_divide intro: order_trans [OF norm_triangle_ineq]) qed lemma norm_cos_plus1_le: fixes z::complex shows "norm(1 + cos z) \ 2 * exp(norm z)" proof - have mono: "\u w z::real. (1 \ w | 1 \ z) \ (w \ u & z \ u) \ 2 + w + z \ 4 * u" by arith have *: "Im z \ cmod z" using abs_Im_le_cmod abs_le_D1 by auto have triangle3: "\x y z. norm(x + y + z) \ norm(x) + norm(y) + norm(z)" by (simp add: norm_add_rule_thm) have "norm(1 + cos z) = cmod (1 + (exp (\ * z) + exp (- (\ * z))) / 2)" by (simp add: cos_exp_eq) also have "... = cmod ((2 + exp (\ * z) + exp (- (\ * z))) / 2)" by (simp add: field_simps) also have "... = cmod (2 + exp (\ * z) + exp (- (\ * z))) / 2" by (simp add: norm_divide) finally show ?thesis by (metis exp_eq_one_iff exp_le_cancel_iff mult_2 norm_cos_le norm_ge_zero norm_one norm_triangle_mono) qed lemma sinh_conv_sin: "sinh z = -\ * sin (\*z)" by (simp add: sinh_field_def sin_i_times exp_minus) lemma cosh_conv_cos: "cosh z = cos (\*z)" by (simp add: cosh_field_def cos_i_times exp_minus) lemma tanh_conv_tan: "tanh z = -\ * tan (\*z)" by (simp add: tanh_def sinh_conv_sin cosh_conv_cos tan_def) lemma sin_conv_sinh: "sin z = -\ * sinh (\*z)" by (simp add: sinh_conv_sin) lemma cos_conv_cosh: "cos z = cosh (\*z)" by (simp add: cosh_conv_cos) lemma tan_conv_tanh: "tan z = -\ * tanh (\*z)" by (simp add: tan_def sin_conv_sinh cos_conv_cosh tanh_def) lemma sinh_complex_eq_iff: "sinh (z :: complex) = sinh w \ (\n\\. z = w - 2 * \ * of_real n * of_real pi \ z = -(2 * complex_of_real n + 1) * \ * complex_of_real pi - w)" (is "_ = ?rhs") proof - have "sinh z = sinh w \ sin (\ * z) = sin (\ * w)" by (simp add: sinh_conv_sin) also have "\ \ ?rhs" by (subst complex_sin_eq) (force simp: field_simps complex_eq_iff) finally show ?thesis . qed subsection\<^marker>\tag unimportant\\Taylor series for complex exponential, sine and cosine\ declare power_Suc [simp del] lemma Taylor_exp_field: fixes z::"'a::{banach,real_normed_field}" shows "norm (exp z - (\i\n. z ^ i / fact i)) \ exp (norm z) * (norm z ^ Suc n) / fact n" proof (rule field_Taylor[of _ n "\k. exp" "exp (norm z)" 0 z, simplified]) show "convex (closed_segment 0 z)" by (rule convex_closed_segment [of 0 z]) next fix k x assume "x \ closed_segment 0 z" "k \ n" show "(exp has_field_derivative exp x) (at x within closed_segment 0 z)" using DERIV_exp DERIV_subset by blast next fix x assume x: "x \ closed_segment 0 z" have "norm (exp x) \ exp (norm x)" by (rule norm_exp) also have "norm x \ norm z" using x by (auto simp: closed_segment_def intro!: mult_left_le_one_le) finally show "norm (exp x) \ exp (norm z)" by simp qed auto lemma Taylor_exp: "norm(exp z - (\k\n. z ^ k / (fact k))) \ exp\Re z\ * (norm z) ^ (Suc n) / (fact n)" proof (rule complex_Taylor [of _ n "\k. exp" "exp\Re z\" 0 z, simplified]) show "convex (closed_segment 0 z)" by (rule convex_closed_segment [of 0 z]) next fix k x assume "x \ closed_segment 0 z" "k \ n" show "(exp has_field_derivative exp x) (at x within closed_segment 0 z)" using DERIV_exp DERIV_subset by blast next fix x assume "x \ closed_segment 0 z" then obtain u where u: "x = complex_of_real u * z" "0 \ u" "u \ 1" by (auto simp: closed_segment_def scaleR_conv_of_real) then have "u * Re z \ \Re z\" by (metis abs_ge_self abs_ge_zero mult.commute mult.right_neutral mult_mono) then show "Re x \ \Re z\" by (simp add: u) qed auto lemma assumes "0 \ u" "u \ 1" shows cmod_sin_le_exp: "cmod (sin (u *\<^sub>R z)) \ exp \Im z\" and cmod_cos_le_exp: "cmod (cos (u *\<^sub>R z)) \ exp \Im z\" proof - have mono: "\u w z::real. w \ u \ z \ u \ (w + z)/2 \ u" by simp have *: "(cmod (exp (\ * (u * z))) + cmod (exp (- (\ * (u * z)))) ) / 2 \ exp \Im z\" proof (rule mono) show "cmod (exp (\ * (u * z))) \ exp \Im z\" using assms by (auto simp: abs_if mult_left_le_one_le not_less intro: order_trans [of _ 0]) show "cmod (exp (- (\ * (u * z)))) \ exp \Im z\" using assms by (auto simp: abs_if mult_left_le_one_le mult_nonneg_nonpos intro: order_trans [of _ 0]) qed have "cmod (sin (u *\<^sub>R z)) = cmod (exp (\ * (u * z)) - exp (- (\ * (u * z)))) / 2" by (auto simp: scaleR_conv_of_real norm_mult norm_power sin_exp_eq norm_divide) also have "... \ (cmod (exp (\ * (u * z))) + cmod (exp (- (\ * (u * z)))) ) / 2" by (intro divide_right_mono norm_triangle_ineq4) simp also have "... \ exp \Im z\" by (rule *) finally show "cmod (sin (u *\<^sub>R z)) \ exp \Im z\" . have "cmod (cos (u *\<^sub>R z)) = cmod (exp (\ * (u * z)) + exp (- (\ * (u * z)))) / 2" by (auto simp: scaleR_conv_of_real norm_mult norm_power cos_exp_eq norm_divide) also have "... \ (cmod (exp (\ * (u * z))) + cmod (exp (- (\ * (u * z)))) ) / 2" by (intro divide_right_mono norm_triangle_ineq) simp also have "... \ exp \Im z\" by (rule *) finally show "cmod (cos (u *\<^sub>R z)) \ exp \Im z\" . qed lemma Taylor_sin: "norm(sin z - (\k\n. complex_of_real (sin_coeff k) * z ^ k)) \ exp\Im z\ * (norm z) ^ (Suc n) / (fact n)" proof - have mono: "\u w z::real. w \ u \ z \ u \ w + z \ u*2" by arith have *: "cmod (sin z - (\i\n. (-1) ^ (i div 2) * (if even i then sin 0 else cos 0) * z ^ i / (fact i))) \ exp \Im z\ * cmod z ^ Suc n / (fact n)" proof (rule complex_Taylor [of "closed_segment 0 z" n "\k x. (-1)^(k div 2) * (if even k then sin x else cos x)" "exp\Im z\" 0 z, simplified]) fix k x show "((\x. (- 1) ^ (k div 2) * (if even k then sin x else cos x)) has_field_derivative (- 1) ^ (Suc k div 2) * (if odd k then sin x else cos x)) (at x within closed_segment 0 z)" apply (auto simp: power_Suc) apply (intro derivative_eq_intros | simp)+ done next fix x assume "x \ closed_segment 0 z" then show "cmod ((- 1) ^ (Suc n div 2) * (if odd n then sin x else cos x)) \ exp \Im z\" by (auto simp: closed_segment_def norm_mult norm_power cmod_sin_le_exp cmod_cos_le_exp) qed have **: "\k. complex_of_real (sin_coeff k) * z ^ k = (-1)^(k div 2) * (if even k then sin 0 else cos 0) * z^k / of_nat (fact k)" by (auto simp: sin_coeff_def elim!: oddE) show ?thesis by (simp add: ** order_trans [OF _ *]) qed lemma Taylor_cos: "norm(cos z - (\k\n. complex_of_real (cos_coeff k) * z ^ k)) \ exp\Im z\ * (norm z) ^ Suc n / (fact n)" proof - have mono: "\u w z::real. w \ u \ z \ u \ w + z \ u*2" by arith have *: "cmod (cos z - (\i\n. (-1) ^ (Suc i div 2) * (if even i then cos 0 else sin 0) * z ^ i / (fact i))) \ exp \Im z\ * cmod z ^ Suc n / (fact n)" proof (rule complex_Taylor [of "closed_segment 0 z" n "\k x. (-1)^(Suc k div 2) * (if even k then cos x else sin x)" "exp\Im z\" 0 z, simplified]) fix k x assume "x \ closed_segment 0 z" "k \ n" show "((\x. (- 1) ^ (Suc k div 2) * (if even k then cos x else sin x)) has_field_derivative (- 1) ^ Suc (k div 2) * (if odd k then cos x else sin x)) (at x within closed_segment 0 z)" apply (auto simp: power_Suc) apply (intro derivative_eq_intros | simp)+ done next fix x assume "x \ closed_segment 0 z" then show "cmod ((- 1) ^ Suc (n div 2) * (if odd n then cos x else sin x)) \ exp \Im z\" by (auto simp: closed_segment_def norm_mult norm_power cmod_sin_le_exp cmod_cos_le_exp) qed have **: "\k. complex_of_real (cos_coeff k) * z ^ k = (-1)^(Suc k div 2) * (if even k then cos 0 else sin 0) * z^k / of_nat (fact k)" by (auto simp: cos_coeff_def elim!: evenE) show ?thesis by (simp add: ** order_trans [OF _ *]) qed declare power_Suc [simp] text\32-bit Approximation to e\ lemma e_approx_32: "\exp(1) - 5837465777 / 2147483648\ \ (inverse(2 ^ 32)::real)" using Taylor_exp [of 1 14] exp_le apply (simp add: sum_distrib_right in_Reals_norm Re_exp atMost_nat_numeral fact_numeral) apply (simp only: pos_le_divide_eq [symmetric]) done lemma e_less_272: "exp 1 < (272/100::real)" using e_approx_32 by (simp add: abs_if split: if_split_asm) lemma ln_272_gt_1: "ln (272/100) > (1::real)" by (metis e_less_272 exp_less_cancel_iff exp_ln_iff less_trans ln_exp) text\Apparently redundant. But many arguments involve integers.\ lemma ln3_gt_1: "ln 3 > (1::real)" by (simp add: less_trans [OF ln_272_gt_1]) subsection\The argument of a complex number (HOL Light version)\ definition\<^marker>\tag important\ is_Arg :: "[complex,real] \ bool" where "is_Arg z r \ z = of_real(norm z) * exp(\ * of_real r)" definition\<^marker>\tag important\ Arg2pi :: "complex \ real" where "Arg2pi z \ if z = 0 then 0 else THE t. 0 \ t \ t < 2*pi \ is_Arg z t" lemma is_Arg_2pi_iff: "is_Arg z (r + of_int k * (2 * pi)) \ is_Arg z r" by (simp add: algebra_simps is_Arg_def) lemma is_Arg_eqI: assumes r: "is_Arg z r" and s: "is_Arg z s" and rs: "abs (r-s) < 2*pi" and "z \ 0" shows "r = s" proof - have zr: "z = (cmod z) * exp (\ * r)" and zs: "z = (cmod z) * exp (\ * s)" using r s by (auto simp: is_Arg_def) with \z \ 0\ have eq: "exp (\ * r) = exp (\ * s)" by (metis mult_eq_0_iff mult_left_cancel) have "\ * r = \ * s" by (rule exp_complex_eqI) (use rs in \auto simp: eq exp_complex_eqI\) then show ?thesis by simp qed text\This function returns the angle of a complex number from its representation in polar coordinates. Due to periodicity, its range is arbitrary. \<^term>\Arg2pi\ follows HOL Light in adopting the interval \[0,2\)\. But we have the same periodicity issue with logarithms, and it is usual to adopt the same interval for the complex logarithm and argument functions. Further on down, we shall define both functions for the interval \(-\,\]\. The present version is provided for compatibility.\ lemma Arg2pi_0 [simp]: "Arg2pi(0) = 0" by (simp add: Arg2pi_def) lemma Arg2pi_unique_lemma: assumes z: "is_Arg z t" and z': "is_Arg z t'" and t: "0 \ t" "t < 2*pi" and t': "0 \ t'" "t' < 2*pi" and nz: "z \ 0" shows "t' = t" proof - have [dest]: "\x y z::real. x\0 \ x+y < z \ y * of_real t') = of_real (cmod z) * exp (\ * of_real t)" by (metis z z' is_Arg_def) then have "exp (\ * of_real t') = exp (\ * of_real t)" by (metis nz mult_left_cancel mult_zero_left z is_Arg_def) then have "sin t' = sin t \ cos t' = cos t" by (metis cis.simps cis_conv_exp) then obtain n::int where n: "t' = t + 2 * n * pi" by (auto simp: sin_cos_eq_iff) then have "n=0" by (cases n) (use t t' in \auto simp: mult_less_0_iff algebra_simps\) then show "t' = t" by (simp add: n) qed lemma Arg2pi: "0 \ Arg2pi z \ Arg2pi z < 2*pi \ is_Arg z (Arg2pi z)" proof (cases "z=0") case True then show ?thesis by (simp add: Arg2pi_def is_Arg_def) next case False obtain t where t: "0 \ t" "t < 2*pi" and ReIm: "Re z / cmod z = cos t" "Im z / cmod z = sin t" using sincos_total_2pi [OF complex_unit_circle [OF False]] by blast have z: "is_Arg z t" unfolding is_Arg_def using t False ReIm by (intro complex_eqI) (auto simp: exp_Euler sin_of_real cos_of_real field_split_simps) show ?thesis apply (simp add: Arg2pi_def False) apply (rule theI [where a=t]) using t z False apply (auto intro: Arg2pi_unique_lemma) done qed corollary\<^marker>\tag unimportant\ shows Arg2pi_ge_0: "0 \ Arg2pi z" and Arg2pi_lt_2pi: "Arg2pi z < 2*pi" and Arg2pi_eq: "z = of_real(norm z) * exp(\ * of_real(Arg2pi z))" using Arg2pi is_Arg_def by auto lemma complex_norm_eq_1_exp: "norm z = 1 \ exp(\ * of_real (Arg2pi z)) = z" by (metis Arg2pi_eq cis_conv_exp mult.left_neutral norm_cis of_real_1) lemma Arg2pi_unique: "\of_real r * exp(\ * of_real a) = z; 0 < r; 0 \ a; a < 2*pi\ \ Arg2pi z = a" by (rule Arg2pi_unique_lemma [unfolded is_Arg_def, OF _ Arg2pi_eq]) (use Arg2pi [of z] in \auto simp: norm_mult\) lemma cos_Arg2pi: "cmod z * cos (Arg2pi z) = Re z" and sin_Arg2pi: "cmod z * sin (Arg2pi z) = Im z" using Arg2pi_eq [of z] cis_conv_exp Re_rcis Im_rcis unfolding rcis_def by metis+ lemma Arg2pi_minus: assumes "z \ 0" shows "Arg2pi (-z) = (if Arg2pi z < pi then Arg2pi z + pi else Arg2pi z - pi)" apply (rule Arg2pi_unique [of "norm z", OF complex_eqI]) using cos_Arg2pi sin_Arg2pi Arg2pi_ge_0 Arg2pi_lt_2pi [of z] assms by (auto simp: Re_exp Im_exp) lemma Arg2pi_times_of_real [simp]: assumes "0 < r" shows "Arg2pi (of_real r * z) = Arg2pi z" proof (cases "z=0") case False show ?thesis by (rule Arg2pi_unique [of "r * norm z"]) (use Arg2pi False assms is_Arg_def in auto) qed auto lemma Arg2pi_times_of_real2 [simp]: "0 < r \ Arg2pi (z * of_real r) = Arg2pi z" by (metis Arg2pi_times_of_real mult.commute) lemma Arg2pi_divide_of_real [simp]: "0 < r \ Arg2pi (z / of_real r) = Arg2pi z" by (metis Arg2pi_times_of_real2 less_numeral_extra(3) nonzero_eq_divide_eq of_real_eq_0_iff) lemma Arg2pi_le_pi: "Arg2pi z \ pi \ 0 \ Im z" proof (cases "z=0") case False have "0 \ Im z \ 0 \ Im (of_real (cmod z) * exp (\ * complex_of_real (Arg2pi z)))" by (metis Arg2pi_eq) also have "... = (0 \ Im (exp (\ * complex_of_real (Arg2pi z))))" using False by (simp add: zero_le_mult_iff) also have "... \ Arg2pi z \ pi" by (simp add: Im_exp) (metis Arg2pi_ge_0 Arg2pi_lt_2pi sin_lt_zero sin_ge_zero not_le) finally show ?thesis by blast qed auto lemma Arg2pi_lt_pi: "0 < Arg2pi z \ Arg2pi z < pi \ 0 < Im z" proof (cases "z=0") case False have "0 < Im z \ 0 < Im (of_real (cmod z) * exp (\ * complex_of_real (Arg2pi z)))" by (metis Arg2pi_eq) also have "... = (0 < Im (exp (\ * complex_of_real (Arg2pi z))))" using False by (simp add: zero_less_mult_iff) also have "... \ 0 < Arg2pi z \ Arg2pi z < pi" (is "_ = ?rhs") proof - have "0 < sin (Arg2pi z) \ ?rhs" by (meson Arg2pi_ge_0 Arg2pi_lt_2pi less_le_trans not_le sin_le_zero sin_x_le_x) then show ?thesis by (auto simp: Im_exp sin_gt_zero) qed finally show ?thesis by blast qed auto lemma Arg2pi_eq_0: "Arg2pi z = 0 \ z \ \ \ 0 \ Re z" proof (cases "z=0") case False have "z \ \ \ 0 \ Re z \ z \ \ \ 0 \ Re (of_real (cmod z) * exp (\ * complex_of_real (Arg2pi z)))" by (metis Arg2pi_eq) also have "... \ z \ \ \ 0 \ Re (exp (\ * complex_of_real (Arg2pi z)))" using False by (simp add: zero_le_mult_iff) also have "... \ Arg2pi z = 0" proof - have [simp]: "Arg2pi z = 0 \ z \ \" using Arg2pi_eq [of z] by (auto simp: Reals_def) moreover have "\z \ \; 0 \ cos (Arg2pi z)\ \ Arg2pi z = 0" by (metis Arg2pi_lt_pi Arg2pi_ge_0 Arg2pi_le_pi cos_pi complex_is_Real_iff leD less_linear less_minus_one_simps(2) minus_minus neg_less_eq_nonneg order_refl) ultimately show ?thesis by (auto simp: Re_exp) qed finally show ?thesis by blast qed auto corollary\<^marker>\tag unimportant\ Arg2pi_gt_0: assumes "z \ \\<^sub>\\<^sub>0" shows "Arg2pi z > 0" using Arg2pi_eq_0 Arg2pi_ge_0 assms dual_order.strict_iff_order unfolding nonneg_Reals_def by fastforce lemma Arg2pi_eq_pi: "Arg2pi z = pi \ z \ \ \ Re z < 0" using Arg2pi_le_pi [of z] Arg2pi_lt_pi [of z] Arg2pi_eq_0 [of z] Arg2pi_ge_0 [of z] by (fastforce simp: complex_is_Real_iff) lemma Arg2pi_eq_0_pi: "Arg2pi z = 0 \ Arg2pi z = pi \ z \ \" using Arg2pi_eq_0 Arg2pi_eq_pi not_le by auto lemma Arg2pi_of_real: "Arg2pi (of_real r) = (if r<0 then pi else 0)" using Arg2pi_eq_0_pi Arg2pi_eq_pi by fastforce lemma Arg2pi_real: "z \ \ \ Arg2pi z = (if 0 \ Re z then 0 else pi)" using Arg2pi_eq_0 Arg2pi_eq_0_pi by auto lemma Arg2pi_inverse: "Arg2pi(inverse z) = (if z \ \ then Arg2pi z else 2*pi - Arg2pi z)" proof (cases "z=0") case False show ?thesis apply (rule Arg2pi_unique [of "inverse (norm z)"]) using Arg2pi_eq False Arg2pi_ge_0 [of z] Arg2pi_lt_2pi [of z] Arg2pi_eq_0 [of z] by (auto simp: Arg2pi_real in_Reals_norm exp_diff field_simps) qed auto lemma Arg2pi_eq_iff: assumes "w \ 0" "z \ 0" shows "Arg2pi w = Arg2pi z \ (\x. 0 < x & w = of_real x * z)" using assms Arg2pi_eq [of z] Arg2pi_eq [of w] apply auto apply (rule_tac x="norm w / norm z" in exI) apply (simp add: field_split_simps) by (metis mult.commute mult.left_commute) lemma Arg2pi_inverse_eq_0: "Arg2pi(inverse z) = 0 \ Arg2pi z = 0" by (metis Arg2pi_eq_0 Arg2pi_inverse inverse_inverse_eq) lemma Arg2pi_divide: assumes "w \ 0" "z \ 0" "Arg2pi w \ Arg2pi z" shows "Arg2pi(z / w) = Arg2pi z - Arg2pi w" apply (rule Arg2pi_unique [of "norm(z / w)"]) using assms Arg2pi_eq Arg2pi_ge_0 [of w] Arg2pi_lt_2pi [of z] apply (auto simp: exp_diff norm_divide field_simps) done lemma Arg2pi_le_div_sum: assumes "w \ 0" "z \ 0" "Arg2pi w \ Arg2pi z" shows "Arg2pi z = Arg2pi w + Arg2pi(z / w)" by (simp add: Arg2pi_divide assms) lemma Arg2pi_le_div_sum_eq: assumes "w \ 0" "z \ 0" shows "Arg2pi w \ Arg2pi z \ Arg2pi z = Arg2pi w + Arg2pi(z / w)" using assms by (auto simp: Arg2pi_ge_0 intro: Arg2pi_le_div_sum) lemma Arg2pi_diff: assumes "w \ 0" "z \ 0" shows "Arg2pi w - Arg2pi z = (if Arg2pi z \ Arg2pi w then Arg2pi(w / z) else Arg2pi(w/z) - 2*pi)" using assms Arg2pi_divide Arg2pi_inverse [of "w/z"] Arg2pi_eq_0_pi by (force simp add: Arg2pi_ge_0 Arg2pi_divide not_le split: if_split_asm) lemma Arg2pi_add: assumes "w \ 0" "z \ 0" shows "Arg2pi w + Arg2pi z = (if Arg2pi w + Arg2pi z < 2*pi then Arg2pi(w * z) else Arg2pi(w * z) + 2*pi)" using assms Arg2pi_diff [of "w*z" z] Arg2pi_le_div_sum_eq [of z "w*z"] apply (auto simp: Arg2pi_ge_0 Arg2pi_divide not_le) apply (metis Arg2pi_lt_2pi add.commute) apply (metis (no_types) Arg2pi add.commute diff_0 diff_add_cancel diff_less_eq diff_minus_eq_add not_less) done lemma Arg2pi_times: assumes "w \ 0" "z \ 0" shows "Arg2pi (w * z) = (if Arg2pi w + Arg2pi z < 2*pi then Arg2pi w + Arg2pi z else (Arg2pi w + Arg2pi z) - 2*pi)" using Arg2pi_add [OF assms] by auto lemma Arg2pi_cnj_eq_inverse: "z\0 \ Arg2pi (cnj z) = Arg2pi (inverse z)" apply (simp add: Arg2pi_eq_iff field_split_simps complex_norm_square [symmetric]) by (metis of_real_power zero_less_norm_iff zero_less_power) lemma Arg2pi_cnj: "Arg2pi(cnj z) = (if z \ \ then Arg2pi z else 2*pi - Arg2pi z)" proof (cases "z=0") case False then show ?thesis by (simp add: Arg2pi_cnj_eq_inverse Arg2pi_inverse) qed auto lemma Arg2pi_exp: "0 \ Im z \ Im z < 2*pi \ Arg2pi(exp z) = Im z" by (rule Arg2pi_unique [of "exp(Re z)"]) (auto simp: exp_eq_polar) lemma complex_split_polar: obtains r a::real where "z = complex_of_real r * (cos a + \ * sin a)" "0 \ r" "0 \ a" "a < 2*pi" using Arg2pi cis.ctr cis_conv_exp unfolding Complex_eq is_Arg_def by fastforce lemma Re_Im_le_cmod: "Im w * sin \ + Re w * cos \ \ cmod w" proof (cases w rule: complex_split_polar) case (1 r a) with sin_cos_le1 [of a \] show ?thesis apply (simp add: norm_mult cmod_unit_one) by (metis (no_types, opaque_lifting) abs_le_D1 distrib_left mult.commute mult.left_commute mult_left_le) qed subsection\<^marker>\tag unimportant\\Analytic properties of tangent function\ lemma cnj_tan: "cnj(tan z) = tan(cnj z)" by (simp add: cnj_cos cnj_sin tan_def) lemma field_differentiable_at_tan: "cos z \ 0 \ tan field_differentiable at z" unfolding field_differentiable_def using DERIV_tan by blast lemma field_differentiable_within_tan: "cos z \ 0 \ tan field_differentiable (at z within s)" using field_differentiable_at_tan field_differentiable_at_within by blast lemma continuous_within_tan: "cos z \ 0 \ continuous (at z within s) tan" using continuous_at_imp_continuous_within isCont_tan by blast lemma continuous_on_tan [continuous_intros]: "(\z. z \ s \ cos z \ 0) \ continuous_on s tan" by (simp add: continuous_at_imp_continuous_on) lemma holomorphic_on_tan: "(\z. z \ s \ cos z \ 0) \ tan holomorphic_on s" by (simp add: field_differentiable_within_tan holomorphic_on_def) subsection\The principal branch of the Complex logarithm\ instantiation complex :: ln begin definition\<^marker>\tag important\ ln_complex :: "complex \ complex" where "ln_complex \ \z. THE w. exp w = z & -pi < Im(w) & Im(w) \ pi" text\NOTE: within this scope, the constant Ln is not yet available!\ lemma assumes "z \ 0" shows exp_Ln [simp]: "exp(ln z) = z" and mpi_less_Im_Ln: "-pi < Im(ln z)" and Im_Ln_le_pi: "Im(ln z) \ pi" proof - obtain \ where z: "z / (cmod z) = Complex (cos \) (sin \)" using complex_unimodular_polar [of "z / (norm z)"] assms by (auto simp: norm_divide field_split_simps) obtain \ where \: "- pi < \" "\ \ pi" "sin \ = sin \" "cos \ = cos \" using sincos_principal_value [of "\"] assms by (auto simp: norm_divide field_split_simps) have "exp(ln z) = z & -pi < Im(ln z) & Im(ln z) \ pi" unfolding ln_complex_def apply (rule theI [where a = "Complex (ln(norm z)) \"]) using z assms \ apply (auto simp: field_simps exp_complex_eqI exp_eq_polar cis.code) done then show "exp(ln z) = z" "-pi < Im(ln z)" "Im(ln z) \ pi" by auto qed lemma Ln_exp [simp]: assumes "-pi < Im(z)" "Im(z) \ pi" shows "ln(exp z) = z" proof (rule exp_complex_eqI) show "\Im (ln (exp z)) - Im z\ < 2 * pi" using assms mpi_less_Im_Ln [of "exp z"] Im_Ln_le_pi [of "exp z"] by auto qed auto subsection\<^marker>\tag unimportant\\Relation to Real Logarithm\ lemma Ln_of_real: assumes "0 < z" shows "ln(of_real z::complex) = of_real(ln z)" proof - have "ln(of_real (exp (ln z))::complex) = ln (exp (of_real (ln z)))" by (simp add: exp_of_real) also have "... = of_real(ln z)" using assms by (subst Ln_exp) auto finally show ?thesis using assms by simp qed corollary\<^marker>\tag unimportant\ Ln_in_Reals [simp]: "z \ \ \ Re z > 0 \ ln z \ \" by (auto simp: Ln_of_real elim: Reals_cases) corollary\<^marker>\tag unimportant\ Im_Ln_of_real [simp]: "r > 0 \ Im (ln (of_real r)) = 0" by (simp add: Ln_of_real) lemma cmod_Ln_Reals [simp]: "z \ \ \ 0 < Re z \ cmod (ln z) = norm (ln (Re z))" using Ln_of_real by force lemma Ln_Reals_eq: "\x \ \; Re x > 0\ \ ln x = of_real (ln (Re x))" using Ln_of_real by force lemma Ln_1 [simp]: "ln 1 = (0::complex)" proof - have "ln (exp 0) = (0::complex)" by (simp add: del: exp_zero) then show ?thesis by simp qed lemma Ln_eq_zero_iff [simp]: "x \ \\<^sub>\\<^sub>0 \ ln x = 0 \ x = 1" for x::complex by auto (metis exp_Ln exp_zero nonpos_Reals_zero_I) instance by intro_classes (rule ln_complex_def Ln_1) end abbreviation Ln :: "complex \ complex" where "Ln \ ln" lemma Ln_eq_iff: "w \ 0 \ z \ 0 \ (Ln w = Ln z \ w = z)" by (metis exp_Ln) lemma Ln_unique: "exp(z) = w \ -pi < Im(z) \ Im(z) \ pi \ Ln w = z" using Ln_exp by blast lemma Re_Ln [simp]: "z \ 0 \ Re(Ln z) = ln(norm z)" by (metis exp_Ln ln_exp norm_exp_eq_Re) corollary\<^marker>\tag unimportant\ ln_cmod_le: assumes z: "z \ 0" shows "ln (cmod z) \ cmod (Ln z)" using norm_exp [of "Ln z", simplified exp_Ln [OF z]] by (metis Re_Ln complex_Re_le_cmod z) proposition\<^marker>\tag unimportant\ exists_complex_root: fixes z :: complex assumes "n \ 0" obtains w where "z = w ^ n" proof (cases "z=0") case False then show ?thesis by (rule_tac w = "exp(Ln z / n)" in that) (simp add: assms exp_of_nat_mult [symmetric]) qed (use assms in auto) corollary\<^marker>\tag unimportant\ exists_complex_root_nonzero: fixes z::complex assumes "z \ 0" "n \ 0" obtains w where "w \ 0" "z = w ^ n" by (metis exists_complex_root [of n z] assms power_0_left) subsection\<^marker>\tag unimportant\\Derivative of Ln away from the branch cut\ lemma assumes "z \ \\<^sub>\\<^sub>0" shows has_field_derivative_Ln: "(Ln has_field_derivative inverse(z)) (at z)" and Im_Ln_less_pi: "Im (Ln z) < pi" proof - have znz [simp]: "z \ 0" using assms by auto then have "Im (Ln z) \ pi" by (metis (no_types) Im_exp Ln_in_Reals assms complex_nonpos_Reals_iff complex_is_Real_iff exp_Ln mult_zero_right not_less pi_neq_zero sin_pi znz) then show *: "Im (Ln z) < pi" using assms Im_Ln_le_pi by (simp add: le_neq_trans) let ?U = "{w. -pi < Im(w) \ Im(w) < pi}" have 1: "open ?U" by (simp add: open_Collect_conj open_halfspace_Im_gt open_halfspace_Im_lt) have 2: "\x. x \ ?U \ (exp has_derivative blinfun_apply(Blinfun ((*) (exp x)))) (at x)" by (simp add: bounded_linear_Blinfun_apply bounded_linear_mult_right has_field_derivative_imp_has_derivative) have 3: "continuous_on ?U (\x. Blinfun ((*) (exp x)))" unfolding blinfun_mult_right.abs_eq [symmetric] by (intro continuous_intros) have 4: "Ln z \ ?U" by (auto simp: mpi_less_Im_Ln *) have 5: "Blinfun ((*) (inverse z)) o\<^sub>L Blinfun ((*) (exp (Ln z))) = id_blinfun" by (rule blinfun_eqI) (simp add: bounded_linear_mult_right bounded_linear_Blinfun_apply) obtain U' V g g' where "open U'" and sub: "U' \ ?U" and "Ln z \ U'" "open V" "z \ V" and hom: "homeomorphism U' V exp g" and g: "\y. y \ V \ (g has_derivative (g' y)) (at y)" and g': "\y. y \ V \ g' y = inv ((*) (exp (g y)))" and bij: "\y. y \ V \ bij ((*) (exp (g y)))" using inverse_function_theorem [OF 1 2 3 4 5] by (simp add: bounded_linear_Blinfun_apply bounded_linear_mult_right) blast show "(Ln has_field_derivative inverse(z)) (at z)" unfolding has_field_derivative_def proof (rule has_derivative_transform_within_open) show g_eq_Ln: "g y = Ln y" if "y \ V" for y proof - obtain x where "y = exp x" "x \ U'" using hom homeomorphism_image1 that \y \ V\ by blast then show ?thesis using sub hom homeomorphism_apply1 by fastforce qed have "0 \ V" by (meson exp_not_eq_zero hom homeomorphism_def) then have "\y. y \ V \ g' y = inv ((*) y)" by (metis exp_Ln g' g_eq_Ln) then have g': "g' z = (\x. x/z)" by (metis (no_types, opaque_lifting) bij \z \ V\ bij_inv_eq_iff exp_Ln g_eq_Ln nonzero_mult_div_cancel_left znz) show "(g has_derivative (*) (inverse z)) (at z)" using g [OF \z \ V\] g' by (simp add: \z \ V\ field_class.field_divide_inverse has_derivative_imp_has_field_derivative has_field_derivative_imp_has_derivative) qed (auto simp: \z \ V\ \open V\) qed declare has_field_derivative_Ln [derivative_intros] declare has_field_derivative_Ln [THEN DERIV_chain2, derivative_intros] lemma field_differentiable_at_Ln: "z \ \\<^sub>\\<^sub>0 \ Ln field_differentiable at z" using field_differentiable_def has_field_derivative_Ln by blast lemma field_differentiable_within_Ln: "z \ \\<^sub>\\<^sub>0 \ Ln field_differentiable (at z within S)" using field_differentiable_at_Ln field_differentiable_within_subset by blast lemma continuous_at_Ln: "z \ \\<^sub>\\<^sub>0 \ continuous (at z) Ln" by (simp add: field_differentiable_imp_continuous_at field_differentiable_within_Ln) lemma isCont_Ln' [simp,continuous_intros]: "\isCont f z; f z \ \\<^sub>\\<^sub>0\ \ isCont (\x. Ln (f x)) z" by (blast intro: isCont_o2 [OF _ continuous_at_Ln]) lemma continuous_within_Ln [continuous_intros]: "z \ \\<^sub>\\<^sub>0 \ continuous (at z within S) Ln" using continuous_at_Ln continuous_at_imp_continuous_within by blast lemma continuous_on_Ln [continuous_intros]: "(\z. z \ S \ z \ \\<^sub>\\<^sub>0) \ continuous_on S Ln" by (simp add: continuous_at_imp_continuous_on continuous_within_Ln) lemma continuous_on_Ln' [continuous_intros]: "continuous_on S f \ (\z. z \ S \ f z \ \\<^sub>\\<^sub>0) \ continuous_on S (\x. Ln (f x))" by (rule continuous_on_compose2[OF continuous_on_Ln, of "UNIV - nonpos_Reals" S f]) auto lemma holomorphic_on_Ln [holomorphic_intros]: "S \ \\<^sub>\\<^sub>0 = {} \ Ln holomorphic_on S" by (simp add: disjoint_iff field_differentiable_within_Ln holomorphic_on_def) lemma holomorphic_on_Ln' [holomorphic_intros]: "(\z. z \ A \ f z \ \\<^sub>\\<^sub>0) \ f holomorphic_on A \ (\z. Ln (f z)) holomorphic_on A" using holomorphic_on_compose_gen[OF _ holomorphic_on_Ln, of f A "- \\<^sub>\\<^sub>0"] by (auto simp: o_def) lemma tendsto_Ln [tendsto_intros]: fixes L F assumes "(f \ L) F" "L \ \\<^sub>\\<^sub>0" shows "((\x. Ln (f x)) \ Ln L) F" proof - have "nhds L \ filtermap f F" using assms(1) by (simp add: filterlim_def) moreover have "\\<^sub>F y in nhds L. y \ - \\<^sub>\\<^sub>0" using eventually_nhds_in_open[of "- \\<^sub>\\<^sub>0" L] assms by (auto simp: open_Compl) ultimately have "\\<^sub>F y in filtermap f F. y \ - \\<^sub>\\<^sub>0" by (rule filter_leD) moreover have "continuous_on (-\\<^sub>\\<^sub>0) Ln" by (rule continuous_on_Ln) auto ultimately show ?thesis using continuous_on_tendsto_compose[of "- \\<^sub>\\<^sub>0" Ln f L F] assms by (simp add: eventually_filtermap) qed lemma divide_ln_mono: fixes x y::real assumes "3 \ x" "x \ y" shows "x / ln x \ y / ln y" proof (rule exE [OF complex_mvt_line [of x y "\z. z / Ln z" "\z. 1/(Ln z) - 1/(Ln z)^2"]]; clarsimp simp add: closed_segment_Reals closed_segment_eq_real_ivl assms) show "\u. \x \ u; u \ y\ \ ((\z. z / Ln z) has_field_derivative 1 / Ln u - 1 / (Ln u)\<^sup>2) (at u)" using \3 \ x\ by (force intro!: derivative_eq_intros simp: field_simps power_eq_if) show "x / ln x \ y / ln y" if "Re (y / Ln y) - Re (x / Ln x) = (Re (1 / Ln u) - Re (1 / (Ln u)\<^sup>2)) * (y - x)" and x: "x \ u" "u \ y" for u proof - have eq: "y / ln y = (1 / ln u - 1 / (ln u)\<^sup>2) * (y - x) + x / ln x" using that \3 \ x\ by (auto simp: Ln_Reals_eq in_Reals_norm group_add_class.diff_eq_eq) show ?thesis using exp_le \3 \ x\ x by (simp add: eq) (simp add: power_eq_if divide_simps ln_ge_iff) qed qed theorem Ln_series: fixes z :: complex assumes "norm z < 1" shows "(\n. (-1)^Suc n / of_nat n * z^n) sums ln (1 + z)" (is "(\n. ?f n * z^n) sums _") proof - let ?F = "\z. \n. ?f n * z^n" and ?F' = "\z. \n. diffs ?f n * z^n" have r: "conv_radius ?f = 1" by (intro conv_radius_ratio_limit_nonzero[of _ 1]) (simp_all add: norm_divide LIMSEQ_Suc_n_over_n del: of_nat_Suc) have "\c. \z\ball 0 1. ln (1 + z) - ?F z = c" proof (rule has_field_derivative_zero_constant) fix z :: complex assume z': "z \ ball 0 1" hence z: "norm z < 1" by simp define t :: complex where "t = of_real (1 + norm z) / 2" from z have t: "norm z < norm t" "norm t < 1" unfolding t_def by (simp_all add: field_simps norm_divide del: of_real_add) have "Re (-z) \ norm (-z)" by (rule complex_Re_le_cmod) also from z have "... < 1" by simp finally have "((\z. ln (1 + z)) has_field_derivative inverse (1+z)) (at z)" by (auto intro!: derivative_eq_intros simp: complex_nonpos_Reals_iff) moreover have "(?F has_field_derivative ?F' z) (at z)" using t r by (intro termdiffs_strong[of _ t] summable_in_conv_radius) simp_all ultimately have "((\z. ln (1 + z) - ?F z) has_field_derivative (inverse (1 + z) - ?F' z)) (at z within ball 0 1)" by (intro derivative_intros) (simp_all add: at_within_open[OF z']) also have "(\n. of_nat n * ?f n * z ^ (n - Suc 0)) sums ?F' z" using t r by (intro diffs_equiv termdiff_converges[OF t(1)] summable_in_conv_radius) simp_all from sums_split_initial_segment[OF this, of 1] have "(\i. (-z) ^ i) sums ?F' z" by (simp add: power_minus[of z] del: of_nat_Suc) hence "?F' z = inverse (1 + z)" using z by (simp add: sums_iff suminf_geometric divide_inverse) also have "inverse (1 + z) - inverse (1 + z) = 0" by simp finally show "((\z. ln (1 + z) - ?F z) has_field_derivative 0) (at z within ball 0 1)" . qed simp_all then obtain c where c: "\z. z \ ball 0 1 \ ln (1 + z) - ?F z = c" by blast from c[of 0] have "c = 0" by (simp only: powser_zero) simp with c[of z] assms have "ln (1 + z) = ?F z" by simp moreover have "summable (\n. ?f n * z^n)" using assms r by (intro summable_in_conv_radius) simp_all ultimately show ?thesis by (simp add: sums_iff) qed lemma Ln_series': "cmod z < 1 \ (\n. - ((-z)^n) / of_nat n) sums ln (1 + z)" by (drule Ln_series) (simp add: power_minus') lemma ln_series': assumes "abs (x::real) < 1" shows "(\n. - ((-x)^n) / of_nat n) sums ln (1 + x)" proof - from assms have "(\n. - ((-of_real x)^n) / of_nat n) sums ln (1 + complex_of_real x)" by (intro Ln_series') simp_all also have "(\n. - ((-of_real x)^n) / of_nat n) = (\n. complex_of_real (- ((-x)^n) / of_nat n))" by (rule ext) simp also from assms have "ln (1 + complex_of_real x) = of_real (ln (1 + x))" by (subst Ln_of_real [symmetric]) simp_all finally show ?thesis by (subst (asm) sums_of_real_iff) qed lemma Ln_approx_linear: fixes z :: complex assumes "norm z < 1" shows "norm (ln (1 + z) - z) \ norm z^2 / (1 - norm z)" proof - let ?f = "\n. (-1)^Suc n / of_nat n" from assms have "(\n. ?f n * z^n) sums ln (1 + z)" using Ln_series by simp moreover have "(\n. (if n = 1 then 1 else 0) * z^n) sums z" using powser_sums_if[of 1] by simp ultimately have "(\n. (?f n - (if n = 1 then 1 else 0)) * z^n) sums (ln (1 + z) - z)" by (subst left_diff_distrib, intro sums_diff) simp_all from sums_split_initial_segment[OF this, of "Suc 1"] have "(\i. (-(z^2)) * inverse (2 + of_nat i) * (- z)^i) sums (Ln (1 + z) - z)" by (simp add: power2_eq_square mult_ac power_minus[of z] divide_inverse) hence "(Ln (1 + z) - z) = (\i. (-(z^2)) * inverse (of_nat (i+2)) * (-z)^i)" by (simp add: sums_iff) also have A: "summable (\n. norm z^2 * (inverse (real_of_nat (Suc (Suc n))) * cmod z ^ n))" by (rule summable_mult, rule summable_comparison_test_ev[OF _ summable_geometric[of "norm z"]]) (auto simp: assms field_simps intro!: always_eventually) hence "norm (\i. (-(z^2)) * inverse (of_nat (i+2)) * (-z)^i) \ (\i. norm (-(z^2) * inverse (of_nat (i+2)) * (-z)^i))" by (intro summable_norm) (auto simp: norm_power norm_inverse norm_mult mult_ac simp del: of_nat_add of_nat_Suc) also have "norm ((-z)^2 * (-z)^i) * inverse (of_nat (i+2)) \ norm ((-z)^2 * (-z)^i) * 1" for i by (intro mult_left_mono) (simp_all add: field_split_simps) hence "(\i. norm (-(z^2) * inverse (of_nat (i+2)) * (-z)^i)) \ (\i. norm (-(z^2) * (-z)^i))" using A assms unfolding norm_power norm_inverse norm_divide norm_mult apply (intro suminf_le summable_mult summable_geometric) apply (auto simp: norm_power field_simps simp del: of_nat_add of_nat_Suc) done also have "... = norm z^2 * (\i. norm z^i)" using assms by (subst suminf_mult [symmetric]) (auto intro!: summable_geometric simp: norm_mult norm_power) also have "(\i. norm z^i) = inverse (1 - norm z)" using assms by (subst suminf_geometric) (simp_all add: divide_inverse) also have "norm z^2 * ... = norm z^2 / (1 - norm z)" by (simp add: divide_inverse) finally show ?thesis . qed subsection\<^marker>\tag unimportant\\Quadrant-type results for Ln\ lemma cos_lt_zero_pi: "pi/2 < x \ x < 3*pi/2 \ cos x < 0" using cos_minus_pi cos_gt_zero_pi [of "x-pi"] by simp lemma Re_Ln_pos_lt: assumes "z \ 0" shows "\Im(Ln z)\ < pi/2 \ 0 < Re(z)" proof - { fix w assume "w = Ln z" then have w: "Im w \ pi" "- pi < Im w" using Im_Ln_le_pi [of z] mpi_less_Im_Ln [of z] assms by auto have "\Im w\ < pi/2 \ 0 < Re(exp w)" proof assume "\Im w\ < pi/2" then show "0 < Re(exp w)" by (auto simp: Re_exp cos_gt_zero_pi split: if_split_asm) next assume R: "0 < Re(exp w)" then have "\Im w\ \ pi/2" by (metis cos_minus cos_pi_half mult_eq_0_iff Re_exp abs_if order_less_irrefl) then show "\Im w\ < pi/2" using cos_lt_zero_pi [of "-(Im w)"] cos_lt_zero_pi [of "(Im w)"] w R by (force simp: Re_exp zero_less_mult_iff abs_if not_less_iff_gr_or_eq) qed } then show ?thesis using assms by auto qed lemma Re_Ln_pos_le: assumes "z \ 0" shows "\Im(Ln z)\ \ pi/2 \ 0 \ Re(z)" proof - { fix w assume "w = Ln z" then have w: "Im w \ pi" "- pi < Im w" using Im_Ln_le_pi [of z] mpi_less_Im_Ln [of z] assms by auto then have "\Im w\ \ pi/2 \ 0 \ Re(exp w)" using cos_lt_zero_pi [of "- (Im w)"] cos_lt_zero_pi [of "(Im w)"] not_le by (auto simp: Re_exp zero_le_mult_iff abs_if intro: cos_ge_zero) } then show ?thesis using assms by auto qed lemma Im_Ln_pos_lt: assumes "z \ 0" shows "0 < Im(Ln z) \ Im(Ln z) < pi \ 0 < Im(z)" proof - { fix w assume "w = Ln z" then have w: "Im w \ pi" "- pi < Im w" using Im_Ln_le_pi [of z] mpi_less_Im_Ln [of z] assms by auto then have "0 < Im w \ Im w < pi \ 0 < Im(exp w)" using sin_gt_zero [of "- (Im w)"] sin_gt_zero [of "(Im w)"] less_linear by (fastforce simp add: Im_exp zero_less_mult_iff) } then show ?thesis using assms by auto qed lemma Im_Ln_pos_le: assumes "z \ 0" shows "0 \ Im(Ln z) \ Im(Ln z) \ pi \ 0 \ Im(z)" proof - { fix w assume "w = Ln z" then have w: "Im w \ pi" "- pi < Im w" using Im_Ln_le_pi [of z] mpi_less_Im_Ln [of z] assms by auto then have "0 \ Im w \ Im w \ pi \ 0 \ Im(exp w)" using sin_ge_zero [of "- (Im w)"] sin_ge_zero [of "abs(Im w)"] sin_zero_pi_iff [of "Im w"] by (force simp: Im_exp zero_le_mult_iff sin_ge_zero) } then show ?thesis using assms by auto qed lemma Re_Ln_pos_lt_imp: "0 < Re(z) \ \Im(Ln z)\ < pi/2" by (metis Re_Ln_pos_lt less_irrefl zero_complex.simps(1)) lemma Im_Ln_pos_lt_imp: "0 < Im(z) \ 0 < Im(Ln z) \ Im(Ln z) < pi" by (metis Im_Ln_pos_lt not_le order_refl zero_complex.simps(2)) text\A reference to the set of positive real numbers\ lemma Im_Ln_eq_0: "z \ 0 \ (Im(Ln z) = 0 \ 0 < Re(z) \ Im(z) = 0)" by (metis Im_complex_of_real Im_exp Ln_in_Reals Re_Ln_pos_lt Re_Ln_pos_lt_imp Re_complex_of_real complex_is_Real_iff exp_Ln exp_of_real pi_gt_zero) lemma Im_Ln_eq_pi: "z \ 0 \ (Im(Ln z) = pi \ Re(z) < 0 \ Im(z) = 0)" by (metis Im_Ln_eq_0 Im_Ln_pos_le Im_Ln_pos_lt add.left_neutral complex_eq less_eq_real_def mult_zero_right not_less_iff_gr_or_eq pi_ge_zero pi_neq_zero rcis_zero_arg rcis_zero_mod) subsection\<^marker>\tag unimportant\\More Properties of Ln\ lemma cnj_Ln: assumes "z \ \\<^sub>\\<^sub>0" shows "cnj(Ln z) = Ln(cnj z)" proof (cases "z=0") case False show ?thesis proof (rule exp_complex_eqI) have "\Im (cnj (Ln z)) - Im (Ln (cnj z))\ \ \Im (cnj (Ln z))\ + \Im (Ln (cnj z))\" by (rule abs_triangle_ineq4) also have "... < pi + pi" proof - have "\Im (cnj (Ln z))\ < pi" by (simp add: False Im_Ln_less_pi abs_if assms minus_less_iff mpi_less_Im_Ln) moreover have "\Im (Ln (cnj z))\ \ pi" by (meson abs_le_iff complex_cnj_zero_iff less_eq_real_def minus_less_iff False Im_Ln_le_pi mpi_less_Im_Ln) ultimately show ?thesis by simp qed finally show "\Im (cnj (Ln z)) - Im (Ln (cnj z))\ < 2 * pi" by simp show "exp (cnj (Ln z)) = exp (Ln (cnj z))" by (metis False complex_cnj_zero_iff exp_Ln exp_cnj) qed qed (use assms in auto) lemma Ln_inverse: assumes "z \ \\<^sub>\\<^sub>0" shows "Ln(inverse z) = -(Ln z)" proof (cases "z=0") case False show ?thesis proof (rule exp_complex_eqI) have "\Im (Ln (inverse z)) - Im (- Ln z)\ \ \Im (Ln (inverse z))\ + \Im (- Ln z)\" by (rule abs_triangle_ineq4) also have "... < pi + pi" proof - have "\Im (Ln (inverse z))\ < pi" by (simp add: False Im_Ln_less_pi abs_if assms minus_less_iff mpi_less_Im_Ln) moreover have "\Im (- Ln z)\ \ pi" using False Im_Ln_le_pi mpi_less_Im_Ln by fastforce ultimately show ?thesis by simp qed finally show "\Im (Ln (inverse z)) - Im (- Ln z)\ < 2 * pi" by simp show "exp (Ln (inverse z)) = exp (- Ln z)" by (simp add: False exp_minus) qed qed (use assms in auto) lemma Ln_minus1 [simp]: "Ln(-1) = \ * pi" proof (rule exp_complex_eqI) show "\Im (Ln (- 1)) - Im (\ * pi)\ < 2 * pi" using Im_Ln_le_pi [of "-1"] mpi_less_Im_Ln [of "-1"] by auto qed auto lemma Ln_ii [simp]: "Ln \ = \ * of_real pi/2" using Ln_exp [of "\ * (of_real pi/2)"] unfolding exp_Euler by simp lemma Ln_minus_ii [simp]: "Ln(-\) = - (\ * pi/2)" proof - have "Ln(-\) = Ln(inverse \)" by simp also have "... = - (Ln \)" using Ln_inverse by blast also have "... = - (\ * pi/2)" by simp finally show ?thesis . qed lemma Ln_times: assumes "w \ 0" "z \ 0" shows "Ln(w * z) = (if Im(Ln w + Ln z) \ -pi then (Ln(w) + Ln(z)) + \ * of_real(2*pi) else if Im(Ln w + Ln z) > pi then (Ln(w) + Ln(z)) - \ * of_real(2*pi) else Ln(w) + Ln(z))" using pi_ge_zero Im_Ln_le_pi [of w] Im_Ln_le_pi [of z] using assms mpi_less_Im_Ln [of w] mpi_less_Im_Ln [of z] by (auto simp: exp_add exp_diff sin_double cos_double exp_Euler intro!: Ln_unique) corollary\<^marker>\tag unimportant\ Ln_times_simple: "\w \ 0; z \ 0; -pi < Im(Ln w) + Im(Ln z); Im(Ln w) + Im(Ln z) \ pi\ \ Ln(w * z) = Ln(w) + Ln(z)" by (simp add: Ln_times) corollary\<^marker>\tag unimportant\ Ln_times_of_real: "\r > 0; z \ 0\ \ Ln(of_real r * z) = ln r + Ln(z)" using mpi_less_Im_Ln Im_Ln_le_pi by (force simp: Ln_times) corollary\<^marker>\tag unimportant\ Ln_times_of_nat: "\r > 0; z \ 0\ \ Ln(of_nat r * z :: complex) = ln (of_nat r) + Ln(z)" using Ln_times_of_real[of "of_nat r" z] by simp corollary\<^marker>\tag unimportant\ Ln_times_Reals: "\r \ Reals; Re r > 0; z \ 0\ \ Ln(r * z) = ln (Re r) + Ln(z)" using Ln_Reals_eq Ln_times_of_real by fastforce corollary\<^marker>\tag unimportant\ Ln_divide_of_real: "\r > 0; z \ 0\ \ Ln(z / of_real r) = Ln(z) - ln r" using Ln_times_of_real [of "inverse r" z] by (simp add: ln_inverse Ln_of_real mult.commute divide_inverse of_real_inverse [symmetric] del: of_real_inverse) corollary\<^marker>\tag unimportant\ Ln_prod: fixes f :: "'a \ complex" assumes "finite A" "\x. x \ A \ f x \ 0" shows "\n. Ln (prod f A) = (\x \ A. Ln (f x) + (of_int (n x) * (2 * pi)) * \)" using assms proof (induction A) case (insert x A) then obtain n where n: "Ln (prod f A) = (\x\A. Ln (f x) + of_real (of_int (n x) * (2 * pi)) * \)" by auto define D where "D \ Im (Ln (f x)) + Im (Ln (prod f A))" define q::int where "q \ (if D \ -pi then 1 else if D > pi then -1 else 0)" have "prod f A \ 0" "f x \ 0" by (auto simp: insert.hyps insert.prems) with insert.hyps pi_ge_zero show ?case by (rule_tac x="n(x:=q)" in exI) (force simp: Ln_times q_def D_def n intro!: sum.cong) qed auto lemma Ln_minus: assumes "z \ 0" shows "Ln(-z) = (if Im(z) \ 0 \ \(Re(z) < 0 \ Im(z) = 0) then Ln(z) + \ * pi else Ln(z) - \ * pi)" (is "_ = ?rhs") using Im_Ln_le_pi [of z] mpi_less_Im_Ln [of z] assms Im_Ln_eq_pi [of z] Im_Ln_pos_lt [of z] by (fastforce simp: exp_add exp_diff exp_Euler intro!: Ln_unique) lemma Ln_inverse_if: assumes "z \ 0" shows "Ln (inverse z) = (if z \ \\<^sub>\\<^sub>0 then -(Ln z) + \ * 2 * complex_of_real pi else -(Ln z))" proof (cases "z \ \\<^sub>\\<^sub>0") case False then show ?thesis by (simp add: Ln_inverse) next case True then have z: "Im z = 0" "Re z < 0" "- z \ \\<^sub>\\<^sub>0" using assms complex_eq_iff complex_nonpos_Reals_iff by auto have "Ln(inverse z) = Ln(- (inverse (-z)))" by simp also have "... = Ln (inverse (-z)) + \ * complex_of_real pi" using assms z by (simp add: Ln_minus divide_less_0_iff) also have "... = - Ln (- z) + \ * complex_of_real pi" using z Ln_inverse by presburger also have "... = - (Ln z) + \ * 2 * complex_of_real pi" using Ln_minus assms z by auto finally show ?thesis by (simp add: True) qed lemma Ln_times_ii: assumes "z \ 0" shows "Ln(\ * z) = (if 0 \ Re(z) | Im(z) < 0 then Ln(z) + \ * of_real pi/2 else Ln(z) - \ * of_real(3 * pi/2))" using Im_Ln_le_pi [of z] mpi_less_Im_Ln [of z] assms Im_Ln_eq_pi [of z] Im_Ln_pos_lt [of z] Re_Ln_pos_le [of z] by (simp add: Ln_times) auto lemma Ln_of_nat [simp]: "0 < n \ Ln (of_nat n) = of_real (ln (of_nat n))" by (subst of_real_of_nat_eq[symmetric], subst Ln_of_real[symmetric]) simp_all lemma Ln_of_nat_over_of_nat: assumes "m > 0" "n > 0" shows "Ln (of_nat m / of_nat n) = of_real (ln (of_nat m) - ln (of_nat n))" proof - have "of_nat m / of_nat n = (of_real (of_nat m / of_nat n) :: complex)" by simp also from assms have "Ln ... = of_real (ln (of_nat m / of_nat n))" by (simp add: Ln_of_real[symmetric]) also from assms have "... = of_real (ln (of_nat m) - ln (of_nat n))" by (simp add: ln_div) finally show ?thesis . qed subsection\The Argument of a Complex Number\ text\Unlike in HOL Light, it's defined for the same interval as the complex logarithm: \(-\,\]\.\ lemma Arg_eq_Im_Ln: assumes "z \ 0" shows "Arg z = Im (Ln z)" proof (rule cis_Arg_unique) show "sgn z = cis (Im (Ln z))" by (metis assms exp_Ln exp_eq_polar nonzero_mult_div_cancel_left norm_eq_zero norm_exp_eq_Re of_real_eq_0_iff sgn_eq) show "- pi < Im (Ln z)" by (simp add: assms mpi_less_Im_Ln) show "Im (Ln z) \ pi" by (simp add: Im_Ln_le_pi assms) qed text \The 1990s definition of argument coincides with the more recent one\ lemma\<^marker>\tag important\ Arg_def: shows "Arg z = (if z = 0 then 0 else Im (Ln z))" by (simp add: Arg_eq_Im_Ln Arg_zero) lemma Arg_of_real [simp]: "Arg (of_real r) = (if r<0 then pi else 0)" by (simp add: Im_Ln_eq_pi Arg_def) lemma mpi_less_Arg: "-pi < Arg z" and Arg_le_pi: "Arg z \ pi" by (auto simp: Arg_def mpi_less_Im_Ln Im_Ln_le_pi) lemma assumes "z \ 0" shows Arg_eq: "z = of_real(norm z) * exp(\ * Arg z)" using assms exp_Ln exp_eq_polar by (auto simp: Arg_def) lemma is_Arg_Arg: "z \ 0 \ is_Arg z (Arg z)" by (simp add: Arg_eq is_Arg_def) lemma Argument_exists: assumes "z \ 0" and R: "R = {r-pi<..r+pi}" obtains s where "is_Arg z s" "s\R" proof - let ?rp = "r - Arg z + pi" define k where "k \ \?rp / (2 * pi)\" have "(Arg z + of_int k * (2 * pi)) \ R" using floor_divide_lower [of "2*pi" ?rp] floor_divide_upper [of "2*pi" ?rp] by (auto simp: k_def algebra_simps R) then show ?thesis using Arg_eq \z \ 0\ is_Arg_2pi_iff is_Arg_def that by blast qed lemma Argument_exists_unique: assumes "z \ 0" and R: "R = {r-pi<..r+pi}" obtains s where "is_Arg z s" "s\R" "\t. \is_Arg z t; t\R\ \ s=t" proof - obtain s where s: "is_Arg z s" "s\R" using Argument_exists [OF assms] . moreover have "\t. \is_Arg z t; t\R\ \ s=t" using assms s by (auto simp: is_Arg_eqI) ultimately show thesis using that by blast qed lemma Argument_Ex1: assumes "z \ 0" and R: "R = {r-pi<..r+pi}" shows "\!s. is_Arg z s \ s \ R" using Argument_exists_unique [OF assms] by metis lemma Arg_divide: assumes "w \ 0" "z \ 0" shows "is_Arg (z / w) (Arg z - Arg w)" using Arg_eq [of z] Arg_eq [of w] Arg_eq [of "norm(z / w)"] assms by (auto simp: is_Arg_def norm_divide field_simps exp_diff Arg_of_real) lemma Arg_unique_lemma: assumes z: "is_Arg z t" and z': "is_Arg z t'" and t: "- pi < t" "t \ pi" and t': "- pi < t'" "t' \ pi" and nz: "z \ 0" shows "t' = t" using Arg2pi_unique_lemma [of "- (inverse z)"] proof - have "pi - t' = pi - t" proof (rule Arg2pi_unique_lemma [of "- (inverse z)"]) have "- (inverse z) = - (inverse (of_real(norm z) * exp(\ * t)))" by (metis is_Arg_def z) also have "... = (cmod (- inverse z)) * exp (\ * (pi - t))" by (auto simp: field_simps exp_diff norm_divide) finally show "is_Arg (- inverse z) (pi - t)" unfolding is_Arg_def . have "- (inverse z) = - (inverse (of_real(norm z) * exp(\ * t')))" by (metis is_Arg_def z') also have "... = (cmod (- inverse z)) * exp (\ * (pi - t'))" by (auto simp: field_simps exp_diff norm_divide) finally show "is_Arg (- inverse z) (pi - t')" unfolding is_Arg_def . qed (use assms in auto) then show ?thesis by simp qed lemma complex_norm_eq_1_exp_eq: "norm z = 1 \ exp(\ * (Arg z)) = z" by (metis Arg_eq exp_not_eq_zero exp_zero mult.left_neutral norm_zero of_real_1 norm_exp_i_times) lemma Arg_unique: "\of_real r * exp(\ * a) = z; 0 < r; -pi < a; a \ pi\ \ Arg z = a" by (rule Arg_unique_lemma [unfolded is_Arg_def, OF _ Arg_eq]) (use mpi_less_Arg Arg_le_pi in \auto simp: norm_mult\) lemma Arg_minus: assumes "z \ 0" shows "Arg (-z) = (if Arg z \ 0 then Arg z + pi else Arg z - pi)" proof - have [simp]: "cmod z * cos (Arg z) = Re z" using assms Arg_eq [of z] by (metis Re_exp exp_Ln norm_exp_eq_Re Arg_def) have [simp]: "cmod z * sin (Arg z) = Im z" using assms Arg_eq [of z] by (metis Im_exp exp_Ln norm_exp_eq_Re Arg_def) show ?thesis apply (rule Arg_unique [of "norm z", OF complex_eqI]) using mpi_less_Arg [of z] Arg_le_pi [of z] assms by (auto simp: Re_exp Im_exp) qed lemma Arg_times_of_real [simp]: assumes "0 < r" shows "Arg (of_real r * z) = Arg z" proof (cases "z=0") case True then show ?thesis by (simp add: Arg_def) next case False with Arg_eq assms show ?thesis by (auto simp: mpi_less_Arg Arg_le_pi intro!: Arg_unique [of "r * norm z"]) qed lemma Arg_times_of_real2 [simp]: "0 < r \ Arg (z * of_real r) = Arg z" by (metis Arg_times_of_real mult.commute) lemma Arg_divide_of_real [simp]: "0 < r \ Arg (z / of_real r) = Arg z" by (metis Arg_times_of_real2 less_numeral_extra(3) nonzero_eq_divide_eq of_real_eq_0_iff) lemma Arg_less_0: "0 \ Arg z \ 0 \ Im z" using Im_Ln_le_pi Im_Ln_pos_le by (simp add: Arg_def) lemma Arg_eq_pi: "Arg z = pi \ Re z < 0 \ Im z = 0" by (auto simp: Arg_def Im_Ln_eq_pi) lemma Arg_lt_pi: "0 < Arg z \ Arg z < pi \ 0 < Im z" using Arg_less_0 [of z] Im_Ln_pos_lt by (auto simp: order.order_iff_strict Arg_def) lemma Arg_eq_0: "Arg z = 0 \ z \ \ \ 0 \ Re z" using complex_is_Real_iff by (simp add: Arg_def Im_Ln_eq_0) (metis less_eq_real_def of_real_Re of_real_def scale_zero_left) corollary\<^marker>\tag unimportant\ Arg_ne_0: assumes "z \ \\<^sub>\\<^sub>0" shows "Arg z \ 0" using assms by (auto simp: nonneg_Reals_def Arg_eq_0) lemma Arg_eq_pi_iff: "Arg z = pi \ z \ \ \ Re z < 0" proof (cases "z=0") case False then show ?thesis using Arg_eq_0 [of "-z"] Arg_eq_pi complex_is_Real_iff by blast qed (simp add: Arg_def) lemma Arg_eq_0_pi: "Arg z = 0 \ Arg z = pi \ z \ \" using Arg_eq_pi_iff Arg_eq_0 by force lemma Arg_real: "z \ \ \ Arg z = (if 0 \ Re z then 0 else pi)" using Arg_eq_0 Arg_eq_0_pi by auto lemma Arg_inverse: "Arg(inverse z) = (if z \ \ then Arg z else - Arg z)" proof (cases "z \ \") case True then show ?thesis by simp (metis Arg2pi_inverse Arg2pi_real Arg_real Reals_inverse) next case False then have z: "Arg z < pi" "z \ 0" using Arg_eq_0_pi Arg_le_pi by (auto simp: less_eq_real_def) show ?thesis apply (rule Arg_unique [of "inverse (norm z)"]) using False z mpi_less_Arg [of z] Arg_eq [of z] by (auto simp: exp_minus field_simps) qed lemma Arg_eq_iff: assumes "w \ 0" "z \ 0" shows "Arg w = Arg z \ (\x. 0 < x \ w = of_real x * z)" (is "?lhs = ?rhs") proof assume ?lhs then have "w = complex_of_real (cmod w / cmod z) * z" by (metis Arg_eq assms divide_divide_eq_right eq_divide_eq exp_not_eq_zero of_real_divide) then show ?rhs using assms divide_pos_pos zero_less_norm_iff by blast qed auto lemma Arg_inverse_eq_0: "Arg(inverse z) = 0 \ Arg z = 0" by (metis Arg_eq_0 Arg_inverse inverse_inverse_eq) lemma Arg_cnj_eq_inverse: "z\0 \ Arg (cnj z) = Arg (inverse z)" using Arg2pi_cnj_eq_inverse Arg2pi_eq_iff Arg_eq_iff by auto lemma Arg_cnj: "Arg(cnj z) = (if z \ \ then Arg z else - Arg z)" by (metis Arg_cnj_eq_inverse Arg_inverse Reals_0 complex_cnj_zero) lemma Arg_exp: "-pi < Im z \ Im z \ pi \ Arg(exp z) = Im z" by (rule Arg_unique [of "exp(Re z)"]) (auto simp: exp_eq_polar) lemma Ln_Arg: "z\0 \ Ln(z) = ln(norm z) + \ * Arg(z)" by (metis Arg_def Re_Ln complex_eq) lemma continuous_at_Arg: assumes "z \ \\<^sub>\\<^sub>0" shows "continuous (at z) Arg" proof - have [simp]: "(\z. Im (Ln z)) \z\ Arg z" using Arg_def assms continuous_at by fastforce show ?thesis unfolding continuous_at proof (rule Lim_transform_within_open) show "\w. \w \ - \\<^sub>\\<^sub>0; w \ z\ \ Im (Ln w) = Arg w" by (metis Arg_def Compl_iff nonpos_Reals_zero_I) qed (use assms in auto) qed lemma continuous_within_Arg: "z \ \\<^sub>\\<^sub>0 \ continuous (at z within S) Arg" using continuous_at_Arg continuous_at_imp_continuous_within by blast subsection\The Unwinding Number and the Ln product Formula\ text\Note that in this special case the unwinding number is -1, 0 or 1. But it's always an integer.\ lemma is_Arg_exp_Im: "is_Arg (exp z) (Im z)" using exp_eq_polar is_Arg_def norm_exp_eq_Re by auto lemma is_Arg_exp_diff_2pi: assumes "is_Arg (exp z) \" shows "\k. Im z - of_int k * (2 * pi) = \" proof (intro exI is_Arg_eqI) let ?k = "\(Im z - \) / (2 * pi)\" show "is_Arg (exp z) (Im z - real_of_int ?k * (2 * pi))" by (metis diff_add_cancel is_Arg_2pi_iff is_Arg_exp_Im) show "\Im z - real_of_int ?k * (2 * pi) - \\ < 2 * pi" using floor_divide_upper [of "2*pi" "Im z - \"] floor_divide_lower [of "2*pi" "Im z - \"] by (auto simp: algebra_simps abs_if) qed (auto simp: is_Arg_exp_Im assms) lemma Arg_exp_diff_2pi: "\k. Im z - of_int k * (2 * pi) = Arg (exp z)" using is_Arg_exp_diff_2pi [OF is_Arg_Arg] by auto lemma unwinding_in_Ints: "(z - Ln(exp z)) / (of_real(2*pi) * \) \ \" using Arg_exp_diff_2pi [of z] by (force simp: Ints_def image_def field_simps Arg_def intro!: complex_eqI) definition\<^marker>\tag important\ unwinding :: "complex \ int" where "unwinding z \ THE k. of_int k = (z - Ln(exp z)) / (of_real(2*pi) * \)" lemma unwinding: "of_int (unwinding z) = (z - Ln(exp z)) / (of_real(2*pi) * \)" using unwinding_in_Ints [of z] unfolding unwinding_def Ints_def by force lemma unwinding_2pi: "(2*pi) * \ * unwinding(z) = z - Ln(exp z)" by (simp add: unwinding) lemma Ln_times_unwinding: "w \ 0 \ z \ 0 \ Ln(w * z) = Ln(w) + Ln(z) - (2*pi) * \ * unwinding(Ln w + Ln z)" using unwinding_2pi by (simp add: exp_add) lemma arg_conv_arctan: assumes "Re z > 0" shows "Arg z = arctan (Im z / Re z)" proof (rule cis_Arg_unique) show "sgn z = cis (arctan (Im z / Re z))" proof (rule complex_eqI) have "Re (cis (arctan (Im z / Re z))) = 1 / sqrt (1 + (Im z)\<^sup>2 / (Re z)\<^sup>2)" by (simp add: cos_arctan power_divide) also have "1 + Im z ^ 2 / Re z ^ 2 = norm z ^ 2 / Re z ^ 2" using assms by (simp add: cmod_def field_simps) also have "1 / sqrt \ = Re z / norm z" using assms by (simp add: real_sqrt_divide) finally show "Re (sgn z) = Re (cis (arctan (Im z / Re z)))" by simp next have "Im (cis (arctan (Im z / Re z))) = Im z / (Re z * sqrt (1 + (Im z)\<^sup>2 / (Re z)\<^sup>2))" by (simp add: sin_arctan field_simps) also have "1 + Im z ^ 2 / Re z ^ 2 = norm z ^ 2 / Re z ^ 2" using assms by (simp add: cmod_def field_simps) also have "Im z / (Re z * sqrt \) = Im z / norm z" using assms by (simp add: real_sqrt_divide) finally show "Im (sgn z) = Im (cis (arctan (Im z / Re z)))" by simp qed next show "arctan (Im z / Re z) > -pi" by (rule le_less_trans[OF _ arctan_lbound]) auto next have "arctan (Im z / Re z) < pi / 2" by (rule arctan_ubound) also have "\ \ pi" by simp finally show "arctan (Im z / Re z) \ pi" by simp qed subsection\<^marker>\tag unimportant\\Relation between Ln and Arg2pi, and hence continuity of Arg2pi\ lemma Arg2pi_Ln: assumes "0 < Arg2pi z" shows "Arg2pi z = Im(Ln(-z)) + pi" proof (cases "z = 0") case True with assms show ?thesis by simp next case False then have "z / of_real(norm z) = exp(\ * of_real(Arg2pi z))" using Arg2pi [of z] by (metis is_Arg_def abs_norm_cancel nonzero_mult_div_cancel_left norm_of_real zero_less_norm_iff) then have "- z / of_real(norm z) = exp (\ * (of_real (Arg2pi z) - pi))" using cis_conv_exp cis_pi by (auto simp: exp_diff algebra_simps) then have "ln (- z / of_real(norm z)) = ln (exp (\ * (of_real (Arg2pi z) - pi)))" by simp also have "... = \ * (of_real(Arg2pi z) - pi)" using Arg2pi [of z] assms pi_not_less_zero by auto finally have "Arg2pi z = Im (Ln (- z / of_real (cmod z))) + pi" by simp also have "... = Im (Ln (-z) - ln (cmod z)) + pi" by (metis diff_0_right minus_diff_eq zero_less_norm_iff Ln_divide_of_real False) also have "... = Im (Ln (-z)) + pi" by simp finally show ?thesis . qed lemma continuous_at_Arg2pi: assumes "z \ \\<^sub>\\<^sub>0" shows "continuous (at z) Arg2pi" proof - have *: "isCont (\z. Im (Ln (- z)) + pi) z" by (rule Complex.isCont_Im isCont_Ln' continuous_intros | simp add: assms complex_is_Real_iff)+ have [simp]: "Im x \ 0 \ Im (Ln (- x)) + pi = Arg2pi x" for x using Arg2pi_Ln by (simp add: Arg2pi_gt_0 complex_nonneg_Reals_iff) consider "Re z < 0" | "Im z \ 0" using assms using complex_nonneg_Reals_iff not_le by blast then have [simp]: "(\z. Im (Ln (- z)) + pi) \z\ Arg2pi z" using "*" by (simp add: Arg2pi_Ln Arg2pi_gt_0 assms continuous_within) show ?thesis unfolding continuous_at proof (rule Lim_transform_within_open) show "\x. \x \ - \\<^sub>\\<^sub>0; x \ z\ \ Im (Ln (- x)) + pi = Arg2pi x" by (auto simp add: Arg2pi_Ln [OF Arg2pi_gt_0] complex_nonneg_Reals_iff) qed (use assms in auto) qed text\Relation between Arg2pi and arctangent in upper halfplane\ lemma Arg2pi_arctan_upperhalf: assumes "0 < Im z" shows "Arg2pi z = pi/2 - arctan(Re z / Im z)" proof (cases "z = 0") case False show ?thesis proof (rule Arg2pi_unique [of "norm z"]) show "(cmod z) * exp (\ * (pi / 2 - arctan (Re z / Im z))) = z" apply (rule complex_eqI) using assms norm_complex_def [of z, symmetric] unfolding exp_Euler cos_diff sin_diff sin_of_real cos_of_real by (simp_all add: field_simps real_sqrt_divide sin_arctan cos_arctan) qed (use False arctan [of "Re z / Im z"] in auto) qed (use assms in auto) lemma Arg2pi_eq_Im_Ln: assumes "0 \ Im z" "0 < Re z" shows "Arg2pi z = Im (Ln z)" proof (cases "Im z = 0") case True then show ?thesis using Arg2pi_eq_0 Ln_in_Reals assms(2) complex_is_Real_iff by auto next case False then have *: "Arg2pi z > 0" using Arg2pi_gt_0 complex_is_Real_iff by blast then have "z \ 0" by auto with * assms False show ?thesis by (subst Arg2pi_Ln) (auto simp: Ln_minus) qed lemma continuous_within_upperhalf_Arg2pi: assumes "z \ 0" shows "continuous (at z within {z. 0 \ Im z}) Arg2pi" proof (cases "z \ \\<^sub>\\<^sub>0") case False then show ?thesis using continuous_at_Arg2pi continuous_at_imp_continuous_within by auto next case True then have z: "z \ \" "0 < Re z" using assms by (auto simp: complex_nonneg_Reals_iff complex_is_Real_iff complex_neq_0) then have [simp]: "Arg2pi z = 0" "Im (Ln z) = 0" by (auto simp: Arg2pi_eq_0 Im_Ln_eq_0 assms complex_is_Real_iff) show ?thesis proof (clarsimp simp add: continuous_within Lim_within dist_norm) fix e::real assume "0 < e" moreover have "continuous (at z) (\x. Im (Ln x))" using z by (simp add: continuous_at_Ln complex_nonpos_Reals_iff) ultimately obtain d where d: "d>0" "\x. x \ z \ cmod (x - z) < d \ \Im (Ln x)\ < e" by (auto simp: continuous_within Lim_within dist_norm) { fix x assume "cmod (x - z) < Re z / 2" then have "\Re x - Re z\ < Re z / 2" by (metis le_less_trans abs_Re_le_cmod minus_complex.simps(1)) then have "0 < Re x" using z by linarith } then show "\d>0. \x. 0 \ Im x \ x \ z \ cmod (x - z) < d \ \Arg2pi x\ < e" apply (rule_tac x="min d (Re z / 2)" in exI) using z d by (auto simp: Arg2pi_eq_Im_Ln) qed qed lemma continuous_on_upperhalf_Arg2pi: "continuous_on ({z. 0 \ Im z} - {0}) Arg2pi" unfolding continuous_on_eq_continuous_within by (metis DiffE Diff_subset continuous_within_subset continuous_within_upperhalf_Arg2pi insertCI) lemma open_Arg2pi2pi_less_Int: assumes "0 \ s" "t \ 2*pi" shows "open ({y. s < Arg2pi y} \ {y. Arg2pi y < t})" proof - have 1: "continuous_on (UNIV - \\<^sub>\\<^sub>0) Arg2pi" using continuous_at_Arg2pi continuous_at_imp_continuous_within by (auto simp: continuous_on_eq_continuous_within) have 2: "open (UNIV - \\<^sub>\\<^sub>0 :: complex set)" by (simp add: open_Diff) have "open ({z. s < z} \ {z. z < t})" using open_lessThan [of t] open_greaterThan [of s] by (metis greaterThan_def lessThan_def open_Int) moreover have "{y. s < Arg2pi y} \ {y. Arg2pi y < t} \ - \\<^sub>\\<^sub>0" using assms by (auto simp: Arg2pi_real complex_nonneg_Reals_iff complex_is_Real_iff) ultimately show ?thesis using continuous_imp_open_vimage [OF 1 2, of "{z. Re z > s} \ {z. Re z < t}"] by auto qed lemma open_Arg2pi2pi_gt: "open {z. t < Arg2pi z}" proof (cases "t < 0") case True then have "{z. t < Arg2pi z} = UNIV" using Arg2pi_ge_0 less_le_trans by auto then show ?thesis by simp next case False then show ?thesis using open_Arg2pi2pi_less_Int [of t "2*pi"] Arg2pi_lt_2pi by auto qed lemma closed_Arg2pi2pi_le: "closed {z. Arg2pi z \ t}" using open_Arg2pi2pi_gt [of t] by (simp add: closed_def Set.Collect_neg_eq [symmetric] not_le) subsection\<^marker>\tag unimportant\\Complex Powers\ lemma powr_to_1 [simp]: "z powr 1 = (z::complex)" by (simp add: powr_def) lemma powr_nat: fixes n::nat and z::complex shows "z powr n = (if z = 0 then 0 else z^n)" by (simp add: exp_of_nat_mult powr_def) lemma norm_powr_real: "w \ \ \ 0 < Re w \ norm(w powr z) = exp(Re z * ln(Re w))" using Ln_Reals_eq norm_exp_eq_Re by (auto simp: Im_Ln_eq_0 powr_def norm_complex_def) lemma powr_complexpow [simp]: fixes x::complex shows "x \ 0 \ x powr (of_nat n) = x^n" 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" shows "cnj (a powr b) = cnj a powr cnj b" proof (cases "a = 0") case False with assms have "a \ \\<^sub>\\<^sub>0" by (auto simp: complex_eq_iff complex_nonpos_Reals_iff) with False show ?thesis by (simp add: powr_def exp_cnj cnj_Ln) qed simp lemma powr_real_real: assumes "w \ \" "z \ \" "0 < Re w" shows "w powr z = exp(Re z * ln(Re w))" proof - have "w \ 0" using assms by auto with assms show ?thesis by (simp add: powr_def Ln_Reals_eq of_real_exp) qed lemma powr_of_real: fixes x::real and y::real shows "0 \ x \ of_real x powr (of_real y::complex) = of_real (x powr y)" by (simp_all add: powr_def exp_eq_polar) lemma powr_of_int: fixes z::complex and n::int assumes "z\(0::complex)" shows "z powr of_int n = (if n\0 then z^nat n else inverse (z^nat (-n)))" by (metis assms not_le of_int_of_nat powr_complexpow powr_minus) lemma powr_Reals_eq: "\x \ \; y \ \; Re x \ 0\ \ x powr y = of_real (Re x powr Re y)" by (metis of_real_Re powr_of_real) lemma norm_powr_real_mono: "\w \ \; 1 < Re w\ \ cmod(w powr z1) \ cmod(w powr z2) \ Re z1 \ Re z2" by (auto simp: powr_def algebra_simps Reals_def Ln_of_real) lemma powr_times_real: "\x \ \; y \ \; 0 \ Re x; 0 \ Re y\ \ (x * y) powr z = x powr z * y powr z" by (auto simp: Reals_def powr_def Ln_times exp_add algebra_simps less_eq_real_def Ln_of_real) lemma Re_powr_le: "r \ \\<^sub>\\<^sub>0 \ Re (r powr z) \ Re r powr Re z" by (auto simp: powr_def nonneg_Reals_def order_trans [OF complex_Re_le_cmod]) lemma fixes w::complex shows Reals_powr [simp]: "\w \ \\<^sub>\\<^sub>0; z \ \\ \ w powr z \ \" and nonneg_Reals_powr [simp]: "\w \ \\<^sub>\\<^sub>0; z \ \\ \ w powr z \ \\<^sub>\\<^sub>0" by (auto simp: nonneg_Reals_def Reals_def powr_of_real) lemma powr_neg_real_complex: "(- of_real x) powr a = (-1) powr (of_real (sgn x) * a) * of_real x powr (a :: complex)" proof (cases "x = 0") assume x: "x \ 0" hence "(-x) powr a = exp (a * ln (-of_real x))" by (simp add: powr_def) also from x have "ln (-of_real x) = Ln (of_real x) + of_real (sgn x) * pi * \" by (simp add: Ln_minus Ln_of_real) also from x have "exp (a * ...) = cis pi powr (of_real (sgn x) * a) * of_real x powr a" by (simp add: powr_def exp_add algebra_simps Ln_of_real cis_conv_exp) also note cis_pi finally show ?thesis by simp qed simp_all lemma has_field_derivative_powr: fixes z :: complex assumes "z \ \\<^sub>\\<^sub>0" shows "((\z. z powr s) has_field_derivative (s * z powr (s - 1))) (at z)" proof (cases "z=0") case False then have \
: "exp (s * Ln z) * inverse z = exp ((s - 1) * Ln z)" by (simp add: divide_complex_def exp_diff left_diff_distrib') show ?thesis unfolding powr_def proof (rule has_field_derivative_transform_within) show "((\z. exp (s * Ln z)) has_field_derivative s * (if z = 0 then 0 else exp ((s - 1) * Ln z))) (at z)" by (intro derivative_eq_intros | simp add: assms False \
)+ qed (use False in auto) qed (use assms in auto) declare has_field_derivative_powr[THEN DERIV_chain2, derivative_intros] lemma has_field_derivative_powr_of_int: fixes z :: complex assumes gderiv:"(g has_field_derivative gd) (at z within S)" and "g z\0" shows "((\z. g z powr of_int n) has_field_derivative (n * g z powr (of_int n - 1) * gd)) (at z within S)" proof - define dd where "dd = of_int n * g z powr (of_int (n - 1)) * gd" obtain e where "e>0" and e_dist:"\y\S. dist z y < e \ g y \ 0" using DERIV_continuous[OF gderiv,THEN continuous_within_avoid] \g z\0\ by auto have ?thesis when "n\0" proof - define dd' where "dd' = of_int n * g z ^ (nat n - 1) * gd" have "dd=dd'" proof (cases "n=0") case False then have "n-1 \0" using \n\0\ by auto then have "g z powr (of_int (n - 1)) = g z ^ (nat n - 1)" using powr_of_int[OF \g z\0\,of "n-1"] by (simp add: nat_diff_distrib') then show ?thesis unfolding dd_def dd'_def by simp qed (simp add:dd_def dd'_def) then have "((\z. g z powr of_int n) has_field_derivative dd) (at z within S) \ ((\z. g z powr of_int n) has_field_derivative dd') (at z within S)" by simp also have "... \ ((\z. g z ^ nat n) has_field_derivative dd') (at z within S)" proof (rule has_field_derivative_cong_eventually) show "\\<^sub>F x in at z within S. g x powr of_int n = g x ^ nat n" unfolding eventually_at apply (rule exI[where x=e]) using powr_of_int that \e>0\ e_dist by (simp add: dist_commute) qed (use powr_of_int \g z\0\ that in simp) also have "..." unfolding dd'_def using gderiv that by (auto intro!: derivative_eq_intros) finally have "((\z. g z powr of_int n) has_field_derivative dd) (at z within S)" . then show ?thesis unfolding dd_def by simp qed moreover have ?thesis when "n<0" proof - define dd' where "dd' = of_int n / g z ^ (nat (1 - n)) * gd" have "dd=dd'" proof - have "g z powr of_int (n - 1) = inverse (g z ^ nat (1-n))" using powr_of_int[OF \g z\0\,of "n-1"] that by auto then show ?thesis unfolding dd_def dd'_def by (simp add: divide_inverse) qed then have "((\z. g z powr of_int n) has_field_derivative dd) (at z within S) \ ((\z. g z powr of_int n) has_field_derivative dd') (at z within S)" by simp also have "... \ ((\z. inverse (g z ^ nat (-n))) has_field_derivative dd') (at z within S)" proof (rule has_field_derivative_cong_eventually) show "\\<^sub>F x in at z within S. g x powr of_int n = inverse (g x ^ nat (- n))" unfolding eventually_at apply (rule exI[where x=e]) using powr_of_int that \e>0\ e_dist by (simp add: dist_commute) qed (use powr_of_int \g z\0\ that in simp) also have "..." proof - have "nat (- n) + nat (1 - n) - Suc 0 = nat (- n) + nat (- n)" by auto then show ?thesis unfolding dd'_def using gderiv that \g z\0\ by (auto intro!: derivative_eq_intros simp add:field_split_simps power_add[symmetric]) qed finally have "((\z. g z powr of_int n) has_field_derivative dd) (at z within S)" . then show ?thesis unfolding dd_def by simp qed ultimately show ?thesis by force qed lemma field_differentiable_powr_of_int: fixes z :: complex assumes gderiv: "g field_differentiable (at z within S)" and "g z \ 0" shows "(\z. g z powr of_int n) field_differentiable (at z within S)" using has_field_derivative_powr_of_int assms(2) field_differentiable_def gderiv by blast lemma holomorphic_on_powr_of_int [holomorphic_intros]: assumes holf: "f holomorphic_on S" and 0: "\z. z\S \ f z \ 0" shows "(\z. (f z) powr of_int n) holomorphic_on S" proof (cases "n\0") case True then have "?thesis \ (\z. (f z) ^ nat n) holomorphic_on S" by (metis (no_types, lifting) 0 holomorphic_cong powr_of_int) moreover have "(\z. (f z) ^ nat n) holomorphic_on S" using holf by (auto intro: holomorphic_intros) ultimately show ?thesis by auto next case False then have "?thesis \ (\z. inverse (f z) ^ nat (-n)) holomorphic_on S" by (metis (no_types, lifting) "0" holomorphic_cong power_inverse powr_of_int) moreover have "(\z. inverse (f z) ^ nat (-n)) holomorphic_on S" using assms by (auto intro!:holomorphic_intros) ultimately show ?thesis by auto qed lemma has_field_derivative_powr_right [derivative_intros]: "w \ 0 \ ((\z. w powr z) has_field_derivative Ln w * w powr z) (at z)" unfolding powr_def by (intro derivative_eq_intros | simp)+ lemma field_differentiable_powr_right [derivative_intros]: fixes w::complex shows "w \ 0 \ (\z. w powr z) field_differentiable (at z)" using field_differentiable_def has_field_derivative_powr_right by blast lemma holomorphic_on_powr_right [holomorphic_intros]: assumes "f holomorphic_on s" shows "(\z. w powr (f z)) holomorphic_on s" proof (cases "w = 0") case False with assms show ?thesis unfolding holomorphic_on_def field_differentiable_def by (metis (full_types) DERIV_chain' has_field_derivative_powr_right) qed simp lemma holomorphic_on_divide_gen [holomorphic_intros]: assumes f: "f holomorphic_on s" and g: "g holomorphic_on s" and 0: "\z z'. \z \ s; z' \ s\ \ g z = 0 \ g z' = 0" shows "(\z. f z / g z) holomorphic_on s" proof (cases "\z\s. g z = 0") case True with 0 have "g z = 0" if "z \ s" for z using that by blast then show ?thesis using g holomorphic_transform by auto next case False with 0 have "g z \ 0" if "z \ s" for z using that by blast with holomorphic_on_divide show ?thesis using f g by blast qed lemma norm_powr_real_powr: "w \ \ \ 0 \ Re w \ cmod (w powr z) = Re w powr Re z" by (metis dual_order.order_iff_strict norm_powr_real norm_zero of_real_0 of_real_Re powr_def) lemma tendsto_powr_complex: fixes f g :: "_ \ complex" assumes a: "a \ \\<^sub>\\<^sub>0" assumes f: "(f \ a) F" and g: "(g \ b) F" shows "((\z. f z powr g z) \ a powr b) F" proof - from a have [simp]: "a \ 0" by auto from f g a have "((\z. exp (g z * ln (f z))) \ a powr b) F" (is ?P) by (auto intro!: tendsto_intros simp: powr_def) also { have "eventually (\z. z \ 0) (nhds a)" by (intro t1_space_nhds) simp_all with f have "eventually (\z. f z \ 0) F" using filterlim_iff by blast } hence "?P \ ((\z. f z powr g z) \ a powr b) F" by (intro tendsto_cong refl) (simp_all add: powr_def mult_ac) finally show ?thesis . qed lemma tendsto_powr_complex_0: fixes f g :: "'a \ complex" assumes f: "(f \ 0) F" and g: "(g \ b) F" and b: "Re b > 0" shows "((\z. f z powr g z) \ 0) F" proof (rule tendsto_norm_zero_cancel) define h where "h = (\z. if f z = 0 then 0 else exp (Re (g z) * ln (cmod (f z)) + abs (Im (g z)) * pi))" { fix z :: 'a assume z: "f z \ 0" define c where "c = abs (Im (g z)) * pi" from mpi_less_Im_Ln[OF z] Im_Ln_le_pi[OF z] have "abs (Im (Ln (f z))) \ pi" by simp from mult_left_mono[OF this, of "abs (Im (g z))"] have "abs (Im (g z) * Im (ln (f z))) \ c" by (simp add: abs_mult c_def) hence "-Im (g z) * Im (ln (f z)) \ c" by simp hence "norm (f z powr g z) \ h z" by (simp add: powr_def field_simps h_def c_def) } hence le: "norm (f z powr g z) \ h z" for z by (cases "f x = 0") (simp_all add: h_def) have g': "(g \ b) (inf F (principal {z. f z \ 0}))" by (rule tendsto_mono[OF _ g]) simp_all have "((\x. norm (f x)) \ 0) (inf F (principal {z. f z \ 0}))" by (subst tendsto_norm_zero_iff, rule tendsto_mono[OF _ f]) simp_all moreover { have "filterlim (\x. norm (f x)) (principal {0<..}) (principal {z. f z \ 0})" by (auto simp: filterlim_def) hence "filterlim (\x. norm (f x)) (principal {0<..}) (inf F (principal {z. f z \ 0}))" by (rule filterlim_mono) simp_all } ultimately have norm: "filterlim (\x. norm (f x)) (at_right 0) (inf F (principal {z. f z \ 0}))" by (simp add: filterlim_inf at_within_def) have A: "LIM x inf F (principal {z. f z \ 0}). Re (g x) * -ln (cmod (f x)) :> at_top" by (rule filterlim_tendsto_pos_mult_at_top tendsto_intros g' b filterlim_compose[OF filterlim_uminus_at_top_at_bot] filterlim_compose[OF ln_at_0] norm)+ have B: "LIM x inf F (principal {z. f z \ 0}). -\Im (g x)\ * pi + -(Re (g x) * ln (cmod (f x))) :> at_top" by (rule filterlim_tendsto_add_at_top tendsto_intros g')+ (insert A, simp_all) have C: "(h \ 0) F" unfolding h_def by (intro filterlim_If tendsto_const filterlim_compose[OF exp_at_bot]) (insert B, auto simp: filterlim_uminus_at_bot algebra_simps) show "((\x. norm (f x powr g x)) \ 0) F" by (rule Lim_null_comparison[OF always_eventually C]) (insert le, auto) qed lemma tendsto_powr_complex' [tendsto_intros]: fixes f g :: "_ \ complex" assumes "a \ \\<^sub>\\<^sub>0 \ (a = 0 \ Re b > 0)" and "(f \ a) F" "(g \ b) F" shows "((\z. f z powr g z) \ a powr b) F" using assms tendsto_powr_complex tendsto_powr_complex_0 by fastforce lemma tendsto_neg_powr_complex_of_real: assumes "filterlim f at_top F" and "Re s < 0" shows "((\x. complex_of_real (f x) powr s) \ 0) F" proof - have "((\x. norm (complex_of_real (f x) powr s)) \ 0) F" proof (rule Lim_transform_eventually) from assms(1) have "eventually (\x. f x \ 0) F" by (auto simp: filterlim_at_top) thus "eventually (\x. f x powr Re s = norm (of_real (f x) powr s)) F" by eventually_elim (simp add: norm_powr_real_powr) from assms show "((\x. f x powr Re s) \ 0) F" by (intro tendsto_neg_powr) qed thus ?thesis by (simp add: tendsto_norm_zero_iff) qed lemma tendsto_neg_powr_complex_of_nat: assumes "filterlim f at_top F" and "Re s < 0" shows "((\x. of_nat (f x) powr s) \ 0) F" proof - have "((\x. of_real (real (f x)) powr s) \ 0) F" using assms(2) by (intro filterlim_compose[OF _ tendsto_neg_powr_complex_of_real] filterlim_compose[OF _ assms(1)] filterlim_real_sequentially filterlim_ident) auto thus ?thesis by simp qed lemma continuous_powr_complex: assumes "f (netlimit F) \ \\<^sub>\\<^sub>0" "continuous F f" "continuous F g" shows "continuous F (\z. f z powr g z :: complex)" using assms unfolding continuous_def by (intro tendsto_powr_complex) simp_all lemma isCont_powr_complex [continuous_intros]: assumes "f z \ \\<^sub>\\<^sub>0" "isCont f z" "isCont g z" shows "isCont (\z. f z powr g z :: complex) z" using assms unfolding isCont_def by (intro tendsto_powr_complex) simp_all lemma continuous_on_powr_complex [continuous_intros]: assumes "A \ {z. Re (f z) \ 0 \ Im (f z) \ 0}" assumes "\z. z \ A \ f z = 0 \ Re (g z) > 0" assumes "continuous_on A f" "continuous_on A g" shows "continuous_on A (\z. f z powr g z)" unfolding continuous_on_def proof fix z assume z: "z \ A" show "((\z. f z powr g z) \ f z powr g z) (at z within A)" proof (cases "f z = 0") case False from assms(1,2) z have "Re (f z) \ 0 \ Im (f z) \ 0" "f z = 0 \ Re (g z) > 0" by auto with assms(3,4) z show ?thesis by (intro tendsto_powr_complex') (auto elim!: nonpos_Reals_cases simp: complex_eq_iff continuous_on_def) next case True with assms z show ?thesis by (auto intro!: tendsto_powr_complex_0 simp: continuous_on_def) qed qed subsection\<^marker>\tag unimportant\\Some Limits involving Logarithms\ lemma lim_Ln_over_power: fixes s::complex assumes "0 < Re s" shows "(\n. Ln (of_nat n) / of_nat n powr s) \ 0" proof (simp add: lim_sequentially dist_norm, clarify) fix e::real assume e: "0 < e" have "\xo>0. \x\xo. 0 < e * 2 + (e * Re s * 2 - 2) * x + e * (Re s)\<^sup>2 * x\<^sup>2" proof (rule_tac x="2/(e * (Re s)\<^sup>2)" in exI, safe) show "0 < 2 / (e * (Re s)\<^sup>2)" using e assms by (simp add: field_simps) next fix x::real assume x: "2 / (e * (Re s)\<^sup>2) \ x" have "2 / (e * (Re s)\<^sup>2) > 0" using e assms by simp with x have "x > 0" by linarith then have "x * 2 \ e * (x\<^sup>2 * (Re s)\<^sup>2)" using e assms x by (auto simp: power2_eq_square field_simps) also have "... < e * (2 + (x * (Re s * 2) + x\<^sup>2 * (Re s)\<^sup>2))" using e assms \x > 0\ by (auto simp: power2_eq_square field_simps add_pos_pos) finally show "0 < e * 2 + (e * Re s * 2 - 2) * x + e * (Re s)\<^sup>2 * x\<^sup>2" by (auto simp: algebra_simps) qed then have "\xo>0. \x\xo. x / e < 1 + (Re s * x) + (1/2) * (Re s * x)^2" using e by (simp add: field_simps) then have "\xo>0. \x\xo. x / e < exp (Re s * x)" using assms by (force intro: less_le_trans [OF _ exp_lower_Taylor_quadratic]) then obtain xo where "xo > 0" and xo: "\x. x \ xo \ x < e * exp (Re s * x)" using e by (auto simp: field_simps) have "norm (Ln (of_nat n) / of_nat n powr s) < e" if "n \ nat \exp xo\" for n proof - have "ln (real n) \ xo" using that exp_gt_zero ln_ge_iff [of n] nat_ceiling_le_eq by fastforce then show ?thesis using e xo [of "ln n"] by (auto simp: norm_divide norm_powr_real field_split_simps) qed then show "\no. \n\no. norm (Ln (of_nat n) / of_nat n powr s) < e" by blast qed lemma lim_Ln_over_n: "((\n. Ln(of_nat n) / of_nat n) \ 0) sequentially" using lim_Ln_over_power [of 1] by simp lemma lim_ln_over_power: fixes s :: real assumes "0 < s" shows "((\n. ln n / (n powr s)) \ 0) sequentially" proof - have "(\n. ln (Suc n) / (Suc n) powr s) \ 0" using lim_Ln_over_power [of "of_real s", THEN filterlim_sequentially_Suc [THEN iffD2]] assms by (simp add: lim_sequentially dist_norm Ln_Reals_eq norm_powr_real_powr norm_divide) then show ?thesis using filterlim_sequentially_Suc[of "\n::nat. ln n / n powr s"] by auto qed lemma lim_ln_over_n [tendsto_intros]: "((\n. ln(real_of_nat n) / of_nat n) \ 0) sequentially" using lim_ln_over_power [of 1] by auto lemma lim_log_over_n [tendsto_intros]: "(\n. log k n/n) \ 0" proof - have *: "log k n/n = (1/ln k) * (ln n / n)" for n unfolding log_def by auto have "(\n. (1/ln k) * (ln n / n)) \ (1/ln k) * 0" by (intro tendsto_intros) then show ?thesis unfolding * by auto qed lemma lim_1_over_complex_power: assumes "0 < Re s" shows "(\n. 1 / of_nat n powr s) \ 0" proof (rule Lim_null_comparison) have "\n>0. 3 \ n \ 1 \ ln (real_of_nat n)" using ln_272_gt_1 by (force intro: order_trans [of _ "ln (272/100)"]) then show "\\<^sub>F x in sequentially. cmod (1 / of_nat x powr s) \ cmod (Ln (of_nat x) / of_nat x powr s)" by (auto simp: norm_divide field_split_simps eventually_sequentially) show "(\n. cmod (Ln (of_nat n) / of_nat n powr s)) \ 0" using lim_Ln_over_power [OF assms] by (metis tendsto_norm_zero_iff) qed lemma lim_1_over_real_power: fixes s :: real assumes "0 < s" shows "((\n. 1 / (of_nat n powr s)) \ 0) sequentially" using lim_1_over_complex_power [of "of_real s", THEN filterlim_sequentially_Suc [THEN iffD2]] assms apply (subst filterlim_sequentially_Suc [symmetric]) by (simp add: lim_sequentially dist_norm Ln_Reals_eq norm_powr_real_powr norm_divide) lemma lim_1_over_Ln: "((\n. 1 / Ln(of_nat n)) \ 0) sequentially" proof (clarsimp simp add: lim_sequentially dist_norm norm_divide field_split_simps) fix r::real assume "0 < r" have ir: "inverse (exp (inverse r)) > 0" by simp obtain n where n: "1 < of_nat n * inverse (exp (inverse r))" using ex_less_of_nat_mult [of _ 1, OF ir] by auto then have "exp (inverse r) < of_nat n" by (simp add: field_split_simps) then have "ln (exp (inverse r)) < ln (of_nat n)" by (metis exp_gt_zero less_trans ln_exp ln_less_cancel_iff) with \0 < r\ have "1 < r * ln (real_of_nat n)" by (simp add: field_simps) moreover have "n > 0" using n using neq0_conv by fastforce ultimately show "\no. \k. Ln (of_nat k) \ 0 \ no \ k \ 1 < r * cmod (Ln (of_nat k))" using n \0 < r\ by (rule_tac x=n in exI) (force simp: field_split_simps intro: less_le_trans) qed lemma lim_1_over_ln: "((\n. 1 / ln(real_of_nat n)) \ 0) sequentially" using lim_1_over_Ln [THEN filterlim_sequentially_Suc [THEN iffD2]] apply (subst filterlim_sequentially_Suc [symmetric]) by (simp add: lim_sequentially dist_norm Ln_Reals_eq norm_powr_real_powr norm_divide) lemma lim_ln1_over_ln: "(\n. ln(Suc n) / ln n) \ 1" proof (rule Lim_transform_eventually) have "(\n. ln(1 + 1/n) / ln n) \ 0" proof (rule Lim_transform_bound) show "(inverse o real) \ 0" by (metis comp_def lim_inverse_n lim_explicit) show "\\<^sub>F n in sequentially. norm (ln (1 + 1 / n) / ln n) \ norm ((inverse \ real) n)" proof fix n::nat assume n: "3 \ n" then have "ln 3 \ ln n" and ln0: "0 \ ln n" by auto with ln3_gt_1 have "1/ ln n \ 1" by (simp add: field_split_simps) moreover have "ln (1 + 1 / real n) \ 1/n" by (simp add: ln_add_one_self_le_self) ultimately have "ln (1 + 1 / real n) * (1 / ln n) \ (1/n) * 1" by (intro mult_mono) (use n in auto) then show "norm (ln (1 + 1 / n) / ln n) \ norm ((inverse \ real) n)" by (simp add: field_simps ln0) qed qed then show "(\n. 1 + ln(1 + 1/n) / ln n) \ 1" by (metis (full_types) add.right_neutral tendsto_add_const_iff) show "\\<^sub>F k in sequentially. 1 + ln (1 + 1 / k) / ln k = ln(Suc k) / ln k" by (simp add: field_split_simps ln_div eventually_sequentiallyI [of 2]) qed lemma lim_ln_over_ln1: "(\n. ln n / ln(Suc n)) \ 1" proof - have "(\n. inverse (ln(Suc n) / ln n)) \ inverse 1" by (rule tendsto_inverse [OF lim_ln1_over_ln]) auto then show ?thesis by simp qed subsection\<^marker>\tag unimportant\\Relation between Square Root and exp/ln, hence its derivative\ lemma csqrt_exp_Ln: assumes "z \ 0" shows "csqrt z = exp(Ln(z) / 2)" proof - have "(exp (Ln z / 2))\<^sup>2 = (exp (Ln z))" by (metis exp_double nonzero_mult_div_cancel_left times_divide_eq_right zero_neq_numeral) also have "... = z" using assms exp_Ln by blast finally have "csqrt z = csqrt ((exp (Ln z / 2))\<^sup>2)" by simp also have "... = exp (Ln z / 2)" apply (rule csqrt_square) using cos_gt_zero_pi [of "(Im (Ln z) / 2)"] Im_Ln_le_pi mpi_less_Im_Ln assms by (fastforce simp: Re_exp Im_exp ) finally show ?thesis using assms csqrt_square by simp qed lemma csqrt_inverse: assumes "z \ \\<^sub>\\<^sub>0" shows "csqrt (inverse z) = inverse (csqrt z)" proof (cases "z=0") case False then show ?thesis using assms csqrt_exp_Ln Ln_inverse exp_minus by (simp add: csqrt_exp_Ln Ln_inverse exp_minus) qed auto lemma cnj_csqrt: assumes "z \ \\<^sub>\\<^sub>0" shows "cnj(csqrt z) = csqrt(cnj z)" proof (cases "z=0") case False then show ?thesis by (simp add: assms cnj_Ln csqrt_exp_Ln exp_cnj) qed auto lemma has_field_derivative_csqrt: assumes "z \ \\<^sub>\\<^sub>0" shows "(csqrt has_field_derivative inverse(2 * csqrt z)) (at z)" proof - have z: "z \ 0" using assms by auto then have *: "inverse z = inverse (2*z) * 2" by (simp add: field_split_simps) have [simp]: "exp (Ln z / 2) * inverse z = inverse (csqrt z)" by (simp add: z field_simps csqrt_exp_Ln [symmetric]) (metis power2_csqrt power2_eq_square) have "Im z = 0 \ 0 < Re z" using assms complex_nonpos_Reals_iff not_less by blast with z have "((\z. exp (Ln z / 2)) has_field_derivative inverse (2 * csqrt z)) (at z)" by (force intro: derivative_eq_intros * simp add: assms) then show ?thesis proof (rule has_field_derivative_transform_within) show "\x. dist x z < cmod z \ exp (Ln x / 2) = csqrt x" by (metis csqrt_exp_Ln dist_0_norm less_irrefl) qed (use z in auto) qed lemma field_differentiable_at_csqrt: "z \ \\<^sub>\\<^sub>0 \ csqrt field_differentiable at z" using field_differentiable_def has_field_derivative_csqrt by blast lemma field_differentiable_within_csqrt: "z \ \\<^sub>\\<^sub>0 \ csqrt field_differentiable (at z within s)" using field_differentiable_at_csqrt field_differentiable_within_subset by blast lemma continuous_at_csqrt: "z \ \\<^sub>\\<^sub>0 \ continuous (at z) csqrt" by (simp add: field_differentiable_within_csqrt field_differentiable_imp_continuous_at) corollary\<^marker>\tag unimportant\ isCont_csqrt' [simp]: "\isCont f z; f z \ \\<^sub>\\<^sub>0\ \ isCont (\x. csqrt (f x)) z" by (blast intro: isCont_o2 [OF _ continuous_at_csqrt]) lemma continuous_within_csqrt: "z \ \\<^sub>\\<^sub>0 \ continuous (at z within s) csqrt" by (simp add: field_differentiable_imp_continuous_at field_differentiable_within_csqrt) lemma continuous_on_csqrt [continuous_intros]: "(\z. z \ s \ z \ \\<^sub>\\<^sub>0) \ continuous_on s csqrt" by (simp add: continuous_at_imp_continuous_on continuous_within_csqrt) lemma holomorphic_on_csqrt: "(\z. z \ s \ z \ \\<^sub>\\<^sub>0) \ csqrt holomorphic_on s" by (simp add: field_differentiable_within_csqrt holomorphic_on_def) lemma continuous_within_closed_nontrivial: "closed s \ a \ s ==> continuous (at a within s) f" using open_Compl by (force simp add: continuous_def eventually_at_topological filterlim_iff open_Collect_neg) lemma continuous_within_csqrt_posreal: "continuous (at z within (\ \ {w. 0 \ Re(w)})) csqrt" proof (cases "z \ \\<^sub>\\<^sub>0") case True have [simp]: "Im z = 0" and 0: "Re z < 0 \ z = 0" using True cnj.code complex_cnj_zero_iff by (auto simp: Complex_eq complex_nonpos_Reals_iff) fastforce show ?thesis using 0 proof assume "Re z < 0" then show ?thesis by (auto simp: continuous_within_closed_nontrivial [OF closed_Real_halfspace_Re_ge]) next assume "z = 0" moreover have "\e. 0 < e \ \x'\\ \ {w. 0 \ Re w}. cmod x' < e^2 \ cmod (csqrt x') < e" by (auto simp: Reals_def real_less_lsqrt) ultimately show ?thesis using zero_less_power by (fastforce simp: continuous_within_eps_delta) qed qed (blast intro: continuous_within_csqrt) subsection\Complex arctangent\ text\The branch cut gives standard bounds in the real case.\ definition\<^marker>\tag important\ Arctan :: "complex \ complex" where "Arctan \ \z. (\/2) * Ln((1 - \*z) / (1 + \*z))" lemma Arctan_def_moebius: "Arctan z = \/2 * Ln (moebius (-\) 1 \ 1 z)" by (simp add: Arctan_def moebius_def add_ac) lemma Ln_conv_Arctan: assumes "z \ -1" shows "Ln z = -2*\ * Arctan (moebius 1 (- 1) (- \) (- \) z)" proof - have "Arctan (moebius 1 (- 1) (- \) (- \) z) = \/2 * Ln (moebius (- \) 1 \ 1 (moebius 1 (- 1) (- \) (- \) z))" by (simp add: Arctan_def_moebius) also from assms have "\ * z \ \ * (-1)" by (subst mult_left_cancel) simp hence "\ * z - -\ \ 0" by (simp add: eq_neg_iff_add_eq_0) from moebius_inverse'[OF _ this, of 1 1] have "moebius (- \) 1 \ 1 (moebius 1 (- 1) (- \) (- \) z) = z" by simp finally show ?thesis by (simp add: field_simps) qed lemma Arctan_0 [simp]: "Arctan 0 = 0" by (simp add: Arctan_def) lemma Im_complex_div_lemma: "Im((1 - \*z) / (1 + \*z)) = 0 \ Re z = 0" by (auto simp: Im_complex_div_eq_0 algebra_simps) lemma Re_complex_div_lemma: "0 < Re((1 - \*z) / (1 + \*z)) \ norm z < 1" by (simp add: Re_complex_div_gt_0 algebra_simps cmod_def power2_eq_square) lemma tan_Arctan: assumes "z\<^sup>2 \ -1" shows [simp]:"tan(Arctan z) = z" proof - have "1 + \*z \ 0" by (metis assms complex_i_mult_minus i_squared minus_unique power2_eq_square power2_minus) moreover have "1 - \*z \ 0" by (metis assms complex_i_mult_minus i_squared power2_eq_square power2_minus right_minus_eq) ultimately show ?thesis by (simp add: Arctan_def tan_def sin_exp_eq cos_exp_eq exp_minus csqrt_exp_Ln [symmetric] divide_simps power2_eq_square [symmetric]) qed lemma Arctan_tan [simp]: assumes "\Re z\ < pi/2" shows "Arctan(tan z) = z" proof - have "Ln ((1 - \ * tan z) / (1 + \ * tan z)) = 2 * z / \" proof (rule Ln_unique) have ge_pi2: "\n::int. \of_int (2*n + 1) * pi/2\ \ pi/2" by (case_tac n rule: int_cases) (auto simp: abs_mult) have "exp (\*z)*exp (\*z) = -1 \ exp (2*\*z) = -1" by (metis distrib_right exp_add mult_2) also have "... \ exp (2*\*z) = exp (\*pi)" using cis_conv_exp cis_pi by auto also have "... \ exp (2*\*z - \*pi) = 1" by (metis (no_types) diff_add_cancel diff_minus_eq_add exp_add exp_minus_inverse mult.commute) also have "... \ Re(\*2*z - \*pi) = 0 \ (\n::int. Im(\*2*z - \*pi) = of_int (2 * n) * pi)" by (simp add: exp_eq_1) also have "... \ Im z = 0 \ (\n::int. 2 * Re z = of_int (2*n + 1) * pi)" by (simp add: algebra_simps) also have "... \ False" using assms ge_pi2 apply (auto simp: algebra_simps) by (metis abs_mult_pos not_less of_nat_less_0_iff of_nat_numeral) finally have "exp (\*z)*exp (\*z) + 1 \ 0" by (auto simp: add.commute minus_unique) then show "exp (2 * z / \) = (1 - \ * tan z) / (1 + \ * tan z)" apply (simp add: tan_def sin_exp_eq cos_exp_eq exp_minus divide_simps) by (simp add: algebra_simps flip: power2_eq_square exp_double) qed (use assms in auto) then show ?thesis by (auto simp: Arctan_def) qed lemma assumes "Re z = 0 \ \Im z\ < 1" shows Re_Arctan_bounds: "\Re(Arctan z)\ < pi/2" and has_field_derivative_Arctan: "(Arctan has_field_derivative inverse(1 + z\<^sup>2)) (at z)" proof - have nz0: "1 + \*z \ 0" using assms by (metis abs_one add_diff_cancel_left' complex_i_mult_minus diff_0 i_squared imaginary_unit.simps less_asym neg_equal_iff_equal) have "z \ -\" using assms by auto then have zz: "1 + z * z \ 0" by (metis abs_one assms i_squared imaginary_unit.simps less_irrefl minus_unique square_eq_iff) have nz1: "1 - \*z \ 0" using assms by (force simp add: i_times_eq_iff) have nz2: "inverse (1 + \*z) \ 0" using assms by (metis Im_complex_div_lemma Re_complex_div_lemma cmod_eq_Im divide_complex_def less_irrefl mult_zero_right zero_complex.simps(1) zero_complex.simps(2)) have nzi: "((1 - \*z) * inverse (1 + \*z)) \ 0" using nz1 nz2 by auto have "Im ((1 - \*z) / (1 + \*z)) = 0 \ 0 < Re ((1 - \*z) / (1 + \*z))" apply (simp add: divide_complex_def) apply (simp add: divide_simps split: if_split_asm) using assms apply (auto simp: algebra_simps abs_square_less_1 [unfolded power2_eq_square]) done then have *: "((1 - \*z) / (1 + \*z)) \ \\<^sub>\\<^sub>0" by (auto simp add: complex_nonpos_Reals_iff) show "\Re(Arctan z)\ < pi/2" unfolding Arctan_def divide_complex_def using mpi_less_Im_Ln [OF nzi] by (auto simp: abs_if intro!: Im_Ln_less_pi * [unfolded divide_complex_def]) show "(Arctan has_field_derivative inverse(1 + z\<^sup>2)) (at z)" unfolding Arctan_def scaleR_conv_of_real apply (intro derivative_eq_intros | simp add: nz0 *)+ using nz1 zz apply (simp add: field_split_simps power2_eq_square) apply algebra done qed lemma field_differentiable_at_Arctan: "(Re z = 0 \ \Im z\ < 1) \ Arctan field_differentiable at z" using has_field_derivative_Arctan by (auto simp: field_differentiable_def) lemma field_differentiable_within_Arctan: "(Re z = 0 \ \Im z\ < 1) \ Arctan field_differentiable (at z within s)" using field_differentiable_at_Arctan field_differentiable_at_within by blast declare has_field_derivative_Arctan [derivative_intros] declare has_field_derivative_Arctan [THEN DERIV_chain2, derivative_intros] lemma continuous_at_Arctan: "(Re z = 0 \ \Im z\ < 1) \ continuous (at z) Arctan" by (simp add: field_differentiable_imp_continuous_at field_differentiable_within_Arctan) lemma continuous_within_Arctan: "(Re z = 0 \ \Im z\ < 1) \ continuous (at z within s) Arctan" using continuous_at_Arctan continuous_at_imp_continuous_within by blast lemma continuous_on_Arctan [continuous_intros]: "(\z. z \ s \ Re z = 0 \ \Im z\ < 1) \ continuous_on s Arctan" by (auto simp: continuous_at_imp_continuous_on continuous_within_Arctan) lemma holomorphic_on_Arctan: "(\z. z \ s \ Re z = 0 \ \Im z\ < 1) \ Arctan holomorphic_on s" by (simp add: field_differentiable_within_Arctan holomorphic_on_def) theorem Arctan_series: assumes z: "norm (z :: complex) < 1" defines "g \ \n. if odd n then -\*\^n / n else 0" defines "h \ \z n. (-1)^n / of_nat (2*n+1) * (z::complex)^(2*n+1)" shows "(\n. g n * z^n) sums Arctan z" and "h z sums Arctan z" proof - define G where [abs_def]: "G z = (\n. g n * z^n)" for z have summable: "summable (\n. g n * u^n)" if "norm u < 1" for u proof (cases "u = 0") assume u: "u \ 0" have "(\n. ereal (norm (h u n) / norm (h u (Suc n)))) = (\n. ereal (inverse (norm u)^2) * ereal ((2 + inverse (real (Suc n))) / (2 - inverse (real (Suc n)))))" proof fix n have "ereal (norm (h u n) / norm (h u (Suc n))) = ereal (inverse (norm u)^2) * ereal (((2*Suc n+1) / (Suc n)) / ((2*Suc n-1) / (Suc n)))" by (simp add: h_def norm_mult norm_power norm_divide field_split_simps power2_eq_square eval_nat_numeral del: of_nat_add of_nat_Suc) also have "of_nat (2*Suc n+1) / of_nat (Suc n) = (2::real) + inverse (real (Suc n))" by (auto simp: field_split_simps simp del: of_nat_Suc) simp_all? also have "of_nat (2*Suc n-1) / of_nat (Suc n) = (2::real) - inverse (real (Suc n))" by (auto simp: field_split_simps simp del: of_nat_Suc) simp_all? finally show "ereal (norm (h u n) / norm (h u (Suc n))) = ereal (inverse (norm u)^2) * ereal ((2 + inverse (real (Suc n))) / (2 - inverse (real (Suc n))))" . qed also have "\ \ ereal (inverse (norm u)^2) * ereal ((2 + 0) / (2 - 0))" by (intro tendsto_intros LIMSEQ_inverse_real_of_nat) simp_all finally have "liminf (\n. ereal (cmod (h u n) / cmod (h u (Suc n)))) = inverse (norm u)^2" by (intro lim_imp_Liminf) simp_all moreover from power_strict_mono[OF that, of 2] u have "inverse (norm u)^2 > 1" by (simp add: field_split_simps) ultimately have A: "liminf (\n. ereal (cmod (h u n) / cmod (h u (Suc n)))) > 1" by simp from u have "summable (h u)" by (intro summable_norm_cancel[OF ratio_test_convergence[OF _ A]]) (auto simp: h_def norm_divide norm_mult norm_power simp del: of_nat_Suc intro!: mult_pos_pos divide_pos_pos always_eventually) thus "summable (\n. g n * u^n)" by (subst summable_mono_reindex[of "\n. 2*n+1", symmetric]) (auto simp: power_mult strict_mono_def g_def h_def elim!: oddE) qed (simp add: h_def) have "\c. \u\ball 0 1. Arctan u - G u = c" proof (rule has_field_derivative_zero_constant) fix u :: complex assume "u \ ball 0 1" hence u: "norm u < 1" by (simp) define K where "K = (norm u + 1) / 2" from u and abs_Im_le_cmod[of u] have Im_u: "\Im u\ < 1" by linarith from u have K: "0 \ K" "norm u < K" "K < 1" by (simp_all add: K_def) hence "(G has_field_derivative (\n. diffs g n * u ^ n)) (at u)" unfolding G_def by (intro termdiffs_strong[of _ "of_real K"] summable) simp_all also have "(\n. diffs g n * u^n) = (\n. if even n then (\*u)^n else 0)" by (intro ext) (simp_all del: of_nat_Suc add: g_def diffs_def power_mult_distrib) also have "suminf \ = (\n. (-(u^2))^n)" by (subst suminf_mono_reindex[of "\n. 2*n", symmetric]) (auto elim!: evenE simp: strict_mono_def power_mult power_mult_distrib) also from u have "norm u^2 < 1^2" by (intro power_strict_mono) simp_all hence "(\n. (-(u^2))^n) = inverse (1 + u^2)" by (subst suminf_geometric) (simp_all add: norm_power inverse_eq_divide) finally have "(G has_field_derivative inverse (1 + u\<^sup>2)) (at u)" . from DERIV_diff[OF has_field_derivative_Arctan this] Im_u u show "((\u. Arctan u - G u) has_field_derivative 0) (at u within ball 0 1)" by (simp_all add: at_within_open[OF _ open_ball]) qed simp_all then obtain c where c: "\u. norm u < 1 \ Arctan u - G u = c" by auto from this[of 0] have "c = 0" by (simp add: G_def g_def) with c z have "Arctan z = G z" by simp with summable[OF z] show "(\n. g n * z^n) sums Arctan z" unfolding G_def by (simp add: sums_iff) thus "h z sums Arctan z" by (subst (asm) sums_mono_reindex[of "\n. 2*n+1", symmetric]) (auto elim!: oddE simp: strict_mono_def power_mult g_def h_def) qed text \A quickly-converging series for the logarithm, based on the arctangent.\ theorem ln_series_quadratic: assumes x: "x > (0::real)" shows "(\n. (2*((x - 1) / (x + 1)) ^ (2*n+1) / of_nat (2*n+1))) sums ln x" proof - define y :: complex where "y = of_real ((x-1)/(x+1))" from x have x': "complex_of_real x \ of_real (-1)" by (subst of_real_eq_iff) auto from x have "\x - 1\ < \x + 1\" by linarith hence "norm (complex_of_real (x - 1) / complex_of_real (x + 1)) < 1" by (simp add: norm_divide del: of_real_add of_real_diff) hence "norm (\ * y) < 1" unfolding y_def by (subst norm_mult) simp hence "(\n. (-2*\) * ((-1)^n / of_nat (2*n+1) * (\*y)^(2*n+1))) sums ((-2*\) * Arctan (\*y))" by (intro Arctan_series sums_mult) simp_all also have "(\n. (-2*\) * ((-1)^n / of_nat (2*n+1) * (\*y)^(2*n+1))) = (\n. (-2*\) * ((-1)^n * (\*y*(-y\<^sup>2)^n)/of_nat (2*n+1)))" by (intro ext) (simp_all add: power_mult power_mult_distrib) also have "\ = (\n. 2*y* ((-1) * (-y\<^sup>2))^n/of_nat (2*n+1))" by (intro ext, subst power_mult_distrib) (simp add: algebra_simps power_mult) also have "\ = (\n. 2*y^(2*n+1) / of_nat (2*n+1))" by (subst power_add, subst power_mult) (simp add: mult_ac) also have "\ = (\n. of_real (2*((x-1)/(x+1))^(2*n+1) / of_nat (2*n+1)))" by (intro ext) (simp add: y_def) also have "\ * y = (of_real x - 1) / (-\ * (of_real x + 1))" by (subst divide_divide_eq_left [symmetric]) (simp add: y_def) also have "\ = moebius 1 (-1) (-\) (-\) (of_real x)" by (simp add: moebius_def algebra_simps) also from x' have "-2*\*Arctan \ = Ln (of_real x)" by (intro Ln_conv_Arctan [symmetric]) simp_all also from x have "\ = ln x" by (rule Ln_of_real) finally show ?thesis by (subst (asm) sums_of_real_iff) qed subsection\<^marker>\tag unimportant\ \Real arctangent\ lemma Im_Arctan_of_real [simp]: "Im (Arctan (of_real x)) = 0" proof - have ne: "1 + x\<^sup>2 \ 0" by (metis power_one sum_power2_eq_zero_iff zero_neq_one) have ne1: "1 + \ * complex_of_real x \ 0" using Complex_eq complex_eq_cancel_iff2 by fastforce have "Re (Ln ((1 - \ * x) * inverse (1 + \ * x))) = 0" apply (rule norm_exp_imaginary) using ne apply (simp add: ne1 cmod_def) apply (auto simp: field_split_simps) apply algebra done then show ?thesis unfolding Arctan_def divide_complex_def by (simp add: complex_eq_iff) qed lemma arctan_eq_Re_Arctan: "arctan x = Re (Arctan (of_real x))" proof (rule arctan_unique) have "(1 - \ * x) / (1 + \ * x) \ \\<^sub>\\<^sub>0" by (auto simp: Im_complex_div_lemma complex_nonpos_Reals_iff) then show "- (pi / 2) < Re (Arctan (complex_of_real x))" by (simp add: Arctan_def Im_Ln_less_pi) next have *: " (1 - \*x) / (1 + \*x) \ 0" by (simp add: field_split_simps) ( simp add: complex_eq_iff) show "Re (Arctan (complex_of_real x)) < pi / 2" using mpi_less_Im_Ln [OF *] by (simp add: Arctan_def) next have "tan (Re (Arctan (of_real x))) = Re (tan (Arctan (of_real x)))" by (auto simp: tan_def Complex.Re_divide Re_sin Re_cos Im_sin Im_cos field_simps power2_eq_square) also have "... = x" proof - have "(complex_of_real x)\<^sup>2 \ - 1" by (metis diff_0_right minus_diff_eq mult_zero_left not_le of_real_1 of_real_eq_iff of_real_minus of_real_power power2_eq_square real_minus_mult_self_le zero_less_one) then show ?thesis by simp qed finally show "tan (Re (Arctan (complex_of_real x))) = x" . qed lemma Arctan_of_real: "Arctan (of_real x) = of_real (arctan x)" unfolding arctan_eq_Re_Arctan divide_complex_def by (simp add: complex_eq_iff) lemma Arctan_in_Reals [simp]: "z \ \ \ Arctan z \ \" by (metis Reals_cases Reals_of_real Arctan_of_real) declare arctan_one [simp] lemma arctan_less_pi4_pos: "x < 1 \ arctan x < pi/4" by (metis arctan_less_iff arctan_one) lemma arctan_less_pi4_neg: "-1 < x \ -(pi/4) < arctan x" by (metis arctan_less_iff arctan_minus arctan_one) lemma arctan_less_pi4: "\x\ < 1 \ \arctan x\ < pi/4" by (metis abs_less_iff arctan_less_pi4_pos arctan_minus) lemma arctan_le_pi4: "\x\ \ 1 \ \arctan x\ \ pi/4" by (metis abs_le_iff arctan_le_iff arctan_minus arctan_one) lemma abs_arctan: "\arctan x\ = arctan \x\" by (simp add: abs_if arctan_minus) lemma arctan_add_raw: assumes "\arctan x + arctan y\ < pi/2" shows "arctan x + arctan y = arctan((x + y) / (1 - x * y))" proof (rule arctan_unique [symmetric]) show 12: "- (pi / 2) < arctan x + arctan y" "arctan x + arctan y < pi / 2" using assms by linarith+ show "tan (arctan x + arctan y) = (x + y) / (1 - x * y)" using cos_gt_zero_pi [OF 12] by (simp add: arctan tan_add) qed lemma arctan_inverse: assumes "0 < x" shows "arctan(inverse x) = pi/2 - arctan x" proof - have "arctan(inverse x) = arctan(inverse(tan(arctan x)))" by (simp add: arctan) also have "... = arctan (tan (pi / 2 - arctan x))" by (simp add: tan_cot) also have "... = pi/2 - arctan x" proof - have "0 < pi - arctan x" using arctan_ubound [of x] pi_gt_zero by linarith with assms show ?thesis by (simp add: Transcendental.arctan_tan) qed finally show ?thesis . qed lemma arctan_add_small: assumes "\x * y\ < 1" shows "(arctan x + arctan y = arctan((x + y) / (1 - x * y)))" proof (cases "x = 0 \ y = 0") case False with assms have "\x\ < inverse \y\" by (simp add: field_split_simps abs_mult) with False have "\arctan x\ < pi / 2 - \arctan y\" using assms by (auto simp add: abs_arctan arctan_inverse [symmetric] arctan_less_iff) then show ?thesis by (intro arctan_add_raw) linarith qed auto lemma abs_arctan_le: fixes x::real shows "\arctan x\ \ \x\" proof - have 1: "\x. x \ \ \ cmod (inverse (1 + x\<^sup>2)) \ 1" by (simp add: norm_divide divide_simps in_Reals_norm complex_is_Real_iff power2_eq_square) have "cmod (Arctan w - Arctan z) \ 1 * cmod (w-z)" if "w \ \" "z \ \" for w z apply (rule field_differentiable_bound [OF convex_Reals, of Arctan _ 1]) apply (rule has_field_derivative_at_within [OF has_field_derivative_Arctan]) using 1 that by (auto simp: Reals_def) then have "cmod (Arctan (of_real x) - Arctan 0) \ 1 * cmod (of_real x - 0)" using Reals_0 Reals_of_real by blast then show ?thesis by (simp add: Arctan_of_real) qed lemma arctan_le_self: "0 \ x \ arctan x \ x" by (metis abs_arctan_le abs_of_nonneg zero_le_arctan_iff) lemma abs_tan_ge: "\x\ < pi/2 \ \x\ \ \tan x\" by (metis abs_arctan_le abs_less_iff arctan_tan minus_less_iff) lemma arctan_bounds: assumes "0 \ x" "x < 1" shows arctan_lower_bound: "(\k<2 * n. (- 1) ^ k * (1 / real (k * 2 + 1) * x ^ (k * 2 + 1))) \ arctan x" (is "(\k<_. (- 1)^ k * ?a k) \ _") and arctan_upper_bound: "arctan x \ (\k<2 * n + 1. (- 1) ^ k * (1 / real (k * 2 + 1) * x ^ (k * 2 + 1)))" proof - have tendsto_zero: "?a \ 0" proof (rule tendsto_eq_rhs) show "(\k. 1 / real (k * 2 + 1) * x ^ (k * 2 + 1)) \ 0 * 0" using assms by (intro tendsto_mult real_tendsto_divide_at_top) (auto simp: filterlim_real_sequentially filterlim_sequentially_iff_filterlim_real intro!: real_tendsto_divide_at_top tendsto_power_zero filterlim_real_sequentially tendsto_eq_intros filterlim_at_top_mult_tendsto_pos filterlim_tendsto_add_at_top) qed simp have nonneg: "0 \ ?a n" for n by (force intro!: divide_nonneg_nonneg mult_nonneg_nonneg zero_le_power assms) have le: "?a (Suc n) \ ?a n" for n by (rule mult_mono[OF _ power_decreasing]) (auto simp: field_split_simps assms less_imp_le) from summable_Leibniz'(4)[of ?a, OF tendsto_zero nonneg le, of n] summable_Leibniz'(2)[of ?a, OF tendsto_zero nonneg le, of n] assms show "(\k<2*n. (- 1)^ k * ?a k) \ arctan x" "arctan x \ (\k<2 * n + 1. (- 1)^ k * ?a k)" by (auto simp: arctan_series) qed subsection\<^marker>\tag unimportant\ \Bounds on pi using real arctangent\ lemma pi_machin: "pi = 16 * arctan (1 / 5) - 4 * arctan (1 / 239)" using machin by simp lemma pi_approx: "3.141592653588 \ pi" "pi \ 3.1415926535899" unfolding pi_machin using arctan_bounds[of "1/5" 4] arctan_bounds[of "1/239" 4] by (simp_all add: eval_nat_numeral) lemma pi_gt3: "pi > 3" using pi_approx by simp subsection\Inverse Sine\ definition\<^marker>\tag important\ Arcsin :: "complex \ complex" where "Arcsin \ \z. -\ * Ln(\ * z + csqrt(1 - z\<^sup>2))" lemma Arcsin_body_lemma: "\ * z + csqrt(1 - z\<^sup>2) \ 0" using power2_csqrt [of "1 - z\<^sup>2"] by (metis add.inverse_inverse complex_i_mult_minus diff_0 diff_add_cancel diff_minus_eq_add mult.assoc mult.commute numeral_One power2_eq_square zero_neq_numeral) lemma Arcsin_range_lemma: "\Re z\ < 1 \ 0 < Re(\ * z + csqrt(1 - z\<^sup>2))" using Complex.cmod_power2 [of z, symmetric] by (simp add: real_less_rsqrt algebra_simps Re_power2 cmod_square_less_1_plus) lemma Re_Arcsin: "Re(Arcsin z) = Im (Ln (\ * z + csqrt(1 - z\<^sup>2)))" by (simp add: Arcsin_def) lemma Im_Arcsin: "Im(Arcsin z) = - ln (cmod (\ * z + csqrt (1 - z\<^sup>2)))" by (simp add: Arcsin_def Arcsin_body_lemma) lemma one_minus_z2_notin_nonpos_Reals: assumes "Im z = 0 \ \Re z\ < 1" shows "1 - z\<^sup>2 \ \\<^sub>\\<^sub>0" proof (cases "Im z = 0") case True with assms show ?thesis by (simp add: complex_nonpos_Reals_iff flip: abs_square_less_1) next case False have "\ (Im z)\<^sup>2 \ - 1" using False power2_less_eq_zero_iff by fastforce with False show ?thesis by (auto simp add: complex_nonpos_Reals_iff Re_power2 Im_power2) qed lemma isCont_Arcsin_lemma: assumes le0: "Re (\ * z + csqrt (1 - z\<^sup>2)) \ 0" and "(Im z = 0 \ \Re z\ < 1)" shows False proof (cases "Im z = 0") case True then show ?thesis using assms by (fastforce simp: cmod_def abs_square_less_1 [symmetric]) next case False have leim: "(cmod (1 - z\<^sup>2) + (1 - Re (z\<^sup>2))) / 2 \ (Im z)\<^sup>2" using le0 sqrt_le_D by fastforce have neq: "(cmod z)\<^sup>2 \ 1 + cmod (1 - z\<^sup>2)" proof (clarsimp simp add: cmod_def) assume "(Re z)\<^sup>2 + (Im z)\<^sup>2 = 1 + sqrt ((1 - Re (z\<^sup>2))\<^sup>2 + (Im (z\<^sup>2))\<^sup>2)" then have "((Re z)\<^sup>2 + (Im z)\<^sup>2 - 1)\<^sup>2 = ((1 - Re (z\<^sup>2))\<^sup>2 + (Im (z\<^sup>2))\<^sup>2)" by simp then show False using False by (simp add: power2_eq_square algebra_simps) qed moreover have 2: "(Im z)\<^sup>2 = (1 + ((Im z)\<^sup>2 + cmod (1 - z\<^sup>2)) - (Re z)\<^sup>2) / 2" using leim cmod_power2 [of z] norm_triangle_ineq2 [of "z^2" 1] by (simp add: norm_power Re_power2 norm_minus_commute [of 1]) ultimately show False by (simp add: Re_power2 Im_power2 cmod_power2) qed lemma isCont_Arcsin: assumes "(Im z = 0 \ \Re z\ < 1)" shows "isCont Arcsin z" proof - have 1: "\ * z + csqrt (1 - z\<^sup>2) \ \\<^sub>\\<^sub>0" by (metis isCont_Arcsin_lemma assms complex_nonpos_Reals_iff) have 2: "1 - z\<^sup>2 \ \\<^sub>\\<^sub>0" by (simp add: one_minus_z2_notin_nonpos_Reals assms) show ?thesis using assms unfolding Arcsin_def by (intro isCont_Ln' isCont_csqrt' continuous_intros 1 2) qed lemma isCont_Arcsin' [simp]: shows "isCont f z \ (Im (f z) = 0 \ \Re (f z)\ < 1) \ isCont (\x. Arcsin (f x)) z" by (blast intro: isCont_o2 [OF _ isCont_Arcsin]) lemma sin_Arcsin [simp]: "sin(Arcsin z) = z" proof - have "\*z*2 + csqrt (1 - z\<^sup>2)*2 = 0 \ (\*z)*2 + csqrt (1 - z\<^sup>2)*2 = 0" by (simp add: algebra_simps) \ \Cancelling a factor of 2\ moreover have "... \ (\*z) + csqrt (1 - z\<^sup>2) = 0" by (metis Arcsin_body_lemma distrib_right no_zero_divisors zero_neq_numeral) ultimately show ?thesis apply (simp add: sin_exp_eq Arcsin_def Arcsin_body_lemma exp_minus divide_simps) apply (simp add: algebra_simps) apply (simp add: power2_eq_square [symmetric] algebra_simps) done qed lemma Re_eq_pihalf_lemma: "\Re z\ = pi/2 \ Im z = 0 \ Re ((exp (\*z) + inverse (exp (\*z))) / 2) = 0 \ 0 \ Im ((exp (\*z) + inverse (exp (\*z))) / 2)" apply (simp add: cos_i_times [symmetric] Re_cos Im_cos abs_if del: eq_divide_eq_numeral1) by (metis cos_minus cos_pi_half) lemma Re_less_pihalf_lemma: assumes "\Re z\ < pi / 2" shows "0 < Re ((exp (\*z) + inverse (exp (\*z))) / 2)" proof - have "0 < cos (Re z)" using assms using cos_gt_zero_pi by auto then show ?thesis by (simp add: cos_i_times [symmetric] Re_cos Im_cos add_pos_pos) qed lemma Arcsin_sin: assumes "\Re z\ < pi/2 \ (\Re z\ = pi/2 \ Im z = 0)" shows "Arcsin(sin z) = z" proof - have "Arcsin(sin z) = - (\ * Ln (csqrt (1 - (\ * (exp (\*z) - inverse (exp (\*z))))\<^sup>2 / 4) - (inverse (exp (\*z)) - exp (\*z)) / 2))" by (simp add: sin_exp_eq Arcsin_def exp_minus power_divide) also have "... = - (\ * Ln (csqrt (((exp (\*z) + inverse (exp (\*z)))/2)\<^sup>2) - (inverse (exp (\*z)) - exp (\*z)) / 2))" by (simp add: field_simps power2_eq_square) also have "... = - (\ * Ln (((exp (\*z) + inverse (exp (\*z)))/2) - (inverse (exp (\*z)) - exp (\*z)) / 2))" apply (subst csqrt_square) using assms Re_eq_pihalf_lemma Re_less_pihalf_lemma by auto also have "... = - (\ * Ln (exp (\*z)))" by (simp add: field_simps power2_eq_square) also have "... = z" using assms by (auto simp: abs_if simp del: eq_divide_eq_numeral1 split: if_split_asm) finally show ?thesis . qed lemma Arcsin_unique: "\sin z = w; \Re z\ < pi/2 \ (\Re z\ = pi/2 \ Im z = 0)\ \ Arcsin w = z" by (metis Arcsin_sin) lemma Arcsin_0 [simp]: "Arcsin 0 = 0" by (metis Arcsin_sin norm_zero pi_half_gt_zero real_norm_def sin_zero zero_complex.simps(1)) lemma Arcsin_1 [simp]: "Arcsin 1 = pi/2" by (metis Arcsin_sin Im_complex_of_real Re_complex_of_real numeral_One of_real_numeral pi_half_ge_zero real_sqrt_abs real_sqrt_pow2 real_sqrt_power sin_of_real sin_pi_half) lemma Arcsin_minus_1 [simp]: "Arcsin(-1) = - (pi/2)" by (metis Arcsin_1 Arcsin_sin Im_complex_of_real Re_complex_of_real abs_of_nonneg of_real_minus pi_half_ge_zero power2_minus real_sqrt_abs sin_Arcsin sin_minus) lemma has_field_derivative_Arcsin: assumes "Im z = 0 \ \Re z\ < 1" shows "(Arcsin has_field_derivative inverse(cos(Arcsin z))) (at z)" proof - have "(sin (Arcsin z))\<^sup>2 \ 1" using assms one_minus_z2_notin_nonpos_Reals by force then have "cos (Arcsin z) \ 0" by (metis diff_0_right power_zero_numeral sin_squared_eq) then show ?thesis by (rule has_field_derivative_inverse_basic [OF DERIV_sin _ _ open_ball [of z 1]]) (auto intro: isCont_Arcsin assms) qed declare has_field_derivative_Arcsin [derivative_intros] declare has_field_derivative_Arcsin [THEN DERIV_chain2, derivative_intros] lemma field_differentiable_at_Arcsin: "(Im z = 0 \ \Re z\ < 1) \ Arcsin field_differentiable at z" using field_differentiable_def has_field_derivative_Arcsin by blast lemma field_differentiable_within_Arcsin: "(Im z = 0 \ \Re z\ < 1) \ Arcsin field_differentiable (at z within s)" using field_differentiable_at_Arcsin field_differentiable_within_subset by blast lemma continuous_within_Arcsin: "(Im z = 0 \ \Re z\ < 1) \ continuous (at z within s) Arcsin" using continuous_at_imp_continuous_within isCont_Arcsin by blast lemma continuous_on_Arcsin [continuous_intros]: "(\z. z \ s \ Im z = 0 \ \Re z\ < 1) \ continuous_on s Arcsin" by (simp add: continuous_at_imp_continuous_on) lemma holomorphic_on_Arcsin: "(\z. z \ s \ Im z = 0 \ \Re z\ < 1) \ Arcsin holomorphic_on s" by (simp add: field_differentiable_within_Arcsin holomorphic_on_def) subsection\Inverse Cosine\ definition\<^marker>\tag important\ Arccos :: "complex \ complex" where "Arccos \ \z. -\ * Ln(z + \ * csqrt(1 - z\<^sup>2))" lemma Arccos_range_lemma: "\Re z\ < 1 \ 0 < Im(z + \ * csqrt(1 - z\<^sup>2))" using Arcsin_range_lemma [of "-z"] by simp lemma Arccos_body_lemma: "z + \ * csqrt(1 - z\<^sup>2) \ 0" using Arcsin_body_lemma [of z] by (metis Arcsin_body_lemma complex_i_mult_minus diff_minus_eq_add power2_minus right_minus_eq) lemma Re_Arccos: "Re(Arccos z) = Im (Ln (z + \ * csqrt(1 - z\<^sup>2)))" by (simp add: Arccos_def) lemma Im_Arccos: "Im(Arccos z) = - ln (cmod (z + \ * csqrt (1 - z\<^sup>2)))" by (simp add: Arccos_def Arccos_body_lemma) text\A very tricky argument to find!\ lemma isCont_Arccos_lemma: assumes eq0: "Im (z + \ * csqrt (1 - z\<^sup>2)) = 0" and "(Im z = 0 \ \Re z\ < 1)" shows False proof (cases "Im z = 0") case True then show ?thesis using assms by (fastforce simp add: cmod_def abs_square_less_1 [symmetric]) next case False have Imz: "Im z = - sqrt ((1 + ((Im z)\<^sup>2 + cmod (1 - z\<^sup>2)) - (Re z)\<^sup>2) / 2)" using eq0 abs_Re_le_cmod [of "1-z\<^sup>2"] by (simp add: Re_power2 algebra_simps) have "(cmod z)\<^sup>2 - 1 \ cmod (1 - z\<^sup>2)" proof (clarsimp simp add: cmod_def) assume "(Re z)\<^sup>2 + (Im z)\<^sup>2 - 1 = sqrt ((1 - Re (z\<^sup>2))\<^sup>2 + (Im (z\<^sup>2))\<^sup>2)" then have "((Re z)\<^sup>2 + (Im z)\<^sup>2 - 1)\<^sup>2 = ((1 - Re (z\<^sup>2))\<^sup>2 + (Im (z\<^sup>2))\<^sup>2)" by simp then show False using False by (simp add: power2_eq_square algebra_simps) qed moreover have "(Im z)\<^sup>2 = (1 + ((Im z)\<^sup>2 + cmod (1 - z\<^sup>2)) - (Re z)\<^sup>2) / 2" using abs_Re_le_cmod [of "1-z\<^sup>2"] by (subst Imz) (simp add: Re_power2) ultimately show False by (simp add: cmod_power2) qed lemma isCont_Arccos: assumes "(Im z = 0 \ \Re z\ < 1)" shows "isCont Arccos z" proof - have "z + \ * csqrt (1 - z\<^sup>2) \ \\<^sub>\\<^sub>0" by (metis complex_nonpos_Reals_iff isCont_Arccos_lemma assms) with assms show ?thesis unfolding Arccos_def by (simp_all add: one_minus_z2_notin_nonpos_Reals assms) qed lemma isCont_Arccos' [simp]: "isCont f z \ (Im (f z) = 0 \ \Re (f z)\ < 1) \ isCont (\x. Arccos (f x)) z" by (blast intro: isCont_o2 [OF _ isCont_Arccos]) lemma cos_Arccos [simp]: "cos(Arccos z) = z" proof - have "z*2 + \ * (2 * csqrt (1 - z\<^sup>2)) = 0 \ z*2 + \ * csqrt (1 - z\<^sup>2)*2 = 0" by (simp add: algebra_simps) \ \Cancelling a factor of 2\ moreover have "... \ z + \ * csqrt (1 - z\<^sup>2) = 0" by (metis distrib_right mult_eq_0_iff zero_neq_numeral) ultimately show ?thesis by (simp add: cos_exp_eq Arccos_def Arccos_body_lemma exp_minus field_simps flip: power2_eq_square) qed lemma Arccos_cos: assumes "0 < Re z \ Re z < pi \ Re z = 0 \ 0 \ Im z \ Re z = pi \ Im z \ 0" shows "Arccos(cos z) = z" proof - have *: "((\ - (exp (\ * z))\<^sup>2 * \) / (2 * exp (\ * z))) = sin z" by (simp add: sin_exp_eq exp_minus field_simps power2_eq_square) have "1 - (exp (\ * z) + inverse (exp (\ * z)))\<^sup>2 / 4 = ((\ - (exp (\ * z))\<^sup>2 * \) / (2 * exp (\ * z)))\<^sup>2" by (simp add: field_simps power2_eq_square) then have "Arccos(cos z) = - (\ * Ln ((exp (\ * z) + inverse (exp (\ * z))) / 2 + \ * csqrt (((\ - (exp (\ * z))\<^sup>2 * \) / (2 * exp (\ * z)))\<^sup>2)))" by (simp add: cos_exp_eq Arccos_def exp_minus power_divide) also have "... = - (\ * Ln ((exp (\ * z) + inverse (exp (\ * z))) / 2 + \ * ((\ - (exp (\ * z))\<^sup>2 * \) / (2 * exp (\ * z)))))" apply (subst csqrt_square) using assms Re_sin_pos [of z] Im_sin_nonneg [of z] Im_sin_nonneg2 [of z] by (auto simp: * Re_sin Im_sin) also have "... = - (\ * Ln (exp (\*z)))" by (simp add: field_simps power2_eq_square) also have "... = z" using assms by (subst Complex_Transcendental.Ln_exp, auto) finally show ?thesis . qed lemma Arccos_unique: "\cos z = w; 0 < Re z \ Re z < pi \ Re z = 0 \ 0 \ Im z \ Re z = pi \ Im z \ 0\ \ Arccos w = z" using Arccos_cos by blast lemma Arccos_0 [simp]: "Arccos 0 = pi/2" by (rule Arccos_unique) auto lemma Arccos_1 [simp]: "Arccos 1 = 0" by (rule Arccos_unique) auto lemma Arccos_minus1: "Arccos(-1) = pi" by (rule Arccos_unique) auto lemma has_field_derivative_Arccos: assumes "(Im z = 0 \ \Re z\ < 1)" shows "(Arccos has_field_derivative - inverse(sin(Arccos z))) (at z)" proof - have "x\<^sup>2 \ -1" for x::real by (sos "((R<1 + (([~1] * A=0) + (R<1 * (R<1 * [x__]^2)))))") with assms have "(cos (Arccos z))\<^sup>2 \ 1" by (auto simp: complex_eq_iff Re_power2 Im_power2 abs_square_eq_1) then have "- sin (Arccos z) \ 0" by (metis cos_squared_eq diff_0_right mult_zero_left neg_0_equal_iff_equal power2_eq_square) then have "(Arccos has_field_derivative inverse(- sin(Arccos z))) (at z)" by (rule has_field_derivative_inverse_basic [OF DERIV_cos _ _ open_ball [of z 1]]) (auto intro: isCont_Arccos assms) then show ?thesis by simp qed declare has_field_derivative_Arcsin [derivative_intros] declare has_field_derivative_Arcsin [THEN DERIV_chain2, derivative_intros] lemma field_differentiable_at_Arccos: "(Im z = 0 \ \Re z\ < 1) \ Arccos field_differentiable at z" using field_differentiable_def has_field_derivative_Arccos by blast lemma field_differentiable_within_Arccos: "(Im z = 0 \ \Re z\ < 1) \ Arccos field_differentiable (at z within s)" using field_differentiable_at_Arccos field_differentiable_within_subset by blast lemma continuous_within_Arccos: "(Im z = 0 \ \Re z\ < 1) \ continuous (at z within s) Arccos" using continuous_at_imp_continuous_within isCont_Arccos by blast lemma continuous_on_Arccos [continuous_intros]: "(\z. z \ s \ Im z = 0 \ \Re z\ < 1) \ continuous_on s Arccos" by (simp add: continuous_at_imp_continuous_on) lemma holomorphic_on_Arccos: "(\z. z \ s \ Im z = 0 \ \Re z\ < 1) \ Arccos holomorphic_on s" by (simp add: field_differentiable_within_Arccos holomorphic_on_def) subsection\<^marker>\tag unimportant\\Upper and Lower Bounds for Inverse Sine and Cosine\ lemma Arcsin_bounds: "\Re z\ < 1 \ \Re(Arcsin z)\ < pi/2" unfolding Re_Arcsin by (blast intro: Re_Ln_pos_lt_imp Arcsin_range_lemma) lemma Arccos_bounds: "\Re z\ < 1 \ 0 < Re(Arccos z) \ Re(Arccos z) < pi" unfolding Re_Arccos by (blast intro!: Im_Ln_pos_lt_imp Arccos_range_lemma) lemma Re_Arccos_bounds: "-pi < Re(Arccos z) \ Re(Arccos z) \ pi" unfolding Re_Arccos by (blast intro!: mpi_less_Im_Ln Im_Ln_le_pi Arccos_body_lemma) lemma Re_Arccos_bound: "\Re(Arccos z)\ \ pi" by (meson Re_Arccos_bounds abs_le_iff less_eq_real_def minus_less_iff) lemma Im_Arccos_bound: "\Im (Arccos w)\ \ cmod w" proof - have "(Im (Arccos w))\<^sup>2 \ (cmod (cos (Arccos w)))\<^sup>2 - (cos (Re (Arccos w)))\<^sup>2" using norm_cos_squared [of "Arccos w"] real_le_abs_sinh [of "Im (Arccos w)"] by (simp only: abs_le_square_iff) (simp add: field_split_simps) also have "... \ (cmod w)\<^sup>2" by (auto simp: cmod_power2) finally show ?thesis using abs_le_square_iff by force qed lemma Re_Arcsin_bounds: "-pi < Re(Arcsin z) & Re(Arcsin z) \ pi" unfolding Re_Arcsin by (blast intro!: mpi_less_Im_Ln Im_Ln_le_pi Arcsin_body_lemma) lemma Re_Arcsin_bound: "\Re(Arcsin z)\ \ pi" by (meson Re_Arcsin_bounds abs_le_iff less_eq_real_def minus_less_iff) lemma norm_Arccos_bounded: fixes w :: complex shows "norm (Arccos w) \ pi + norm w" proof - have Re: "(Re (Arccos w))\<^sup>2 \ pi\<^sup>2" "(Im (Arccos w))\<^sup>2 \ (cmod w)\<^sup>2" using Re_Arccos_bound [of w] Im_Arccos_bound [of w] abs_le_square_iff by force+ have "Arccos w \ Arccos w \ pi\<^sup>2 + (cmod w)\<^sup>2" using Re by (simp add: dot_square_norm cmod_power2 [of "Arccos w"]) then have "cmod (Arccos w) \ pi + cmod (cos (Arccos w))" apply (simp add: norm_le_square) by (metis dot_square_norm norm_ge_zero norm_le_square pi_ge_zero triangle_lemma) then show "cmod (Arccos w) \ pi + cmod w" by auto qed subsection\<^marker>\tag unimportant\\Interrelations between Arcsin and Arccos\ lemma cos_Arcsin_nonzero: assumes "z\<^sup>2 \ 1" shows "cos(Arcsin z) \ 0" proof - have eq: "(\ * z * (csqrt (1 - z\<^sup>2)))\<^sup>2 = z\<^sup>2 * (z\<^sup>2 - 1)" by (simp add: algebra_simps) have "\ * z * (csqrt (1 - z\<^sup>2)) \ z\<^sup>2 - 1" proof assume "\ * z * (csqrt (1 - z\<^sup>2)) = z\<^sup>2 - 1" then have "(\ * z * (csqrt (1 - z\<^sup>2)))\<^sup>2 = (z\<^sup>2 - 1)\<^sup>2" by simp then have "z\<^sup>2 * (z\<^sup>2 - 1) = (z\<^sup>2 - 1)*(z\<^sup>2 - 1)" using eq power2_eq_square by auto then show False using assms by simp qed then have "1 + \ * z * (csqrt (1 - z * z)) \ z\<^sup>2" by (metis add_minus_cancel power2_eq_square uminus_add_conv_diff) then have "2*(1 + \ * z * (csqrt (1 - z * z))) \ 2*z\<^sup>2" (*FIXME cancel_numeral_factor*) by (metis mult_cancel_left zero_neq_numeral) then have "(\ * z + csqrt (1 - z\<^sup>2))\<^sup>2 \ -1" using assms apply (simp add: power2_sum) apply (simp add: power2_eq_square algebra_simps) done then show ?thesis apply (simp add: cos_exp_eq Arcsin_def exp_minus) apply (simp add: divide_simps Arcsin_body_lemma) apply (metis add.commute minus_unique power2_eq_square) done qed lemma sin_Arccos_nonzero: assumes "z\<^sup>2 \ 1" shows "sin(Arccos z) \ 0" proof - have eq: "(\ * z * (csqrt (1 - z\<^sup>2)))\<^sup>2 = -(z\<^sup>2) * (1 - z\<^sup>2)" by (simp add: algebra_simps) have "\ * z * (csqrt (1 - z\<^sup>2)) \ 1 - z\<^sup>2" proof assume "\ * z * (csqrt (1 - z\<^sup>2)) = 1 - z\<^sup>2" then have "(\ * z * (csqrt (1 - z\<^sup>2)))\<^sup>2 = (1 - z\<^sup>2)\<^sup>2" by simp then have "-(z\<^sup>2) * (1 - z\<^sup>2) = (1 - z\<^sup>2)*(1 - z\<^sup>2)" using eq power2_eq_square by auto then have "-(z\<^sup>2) = (1 - z\<^sup>2)" using assms by (metis add.commute add.right_neutral diff_add_cancel mult_right_cancel) then show False using assms by simp qed then have "z\<^sup>2 + \ * z * (csqrt (1 - z\<^sup>2)) \ 1" by (simp add: algebra_simps) then have "2*(z\<^sup>2 + \ * z * (csqrt (1 - z\<^sup>2))) \ 2*1" by (metis mult_cancel_left2 zero_neq_numeral) (*FIXME cancel_numeral_factor*) then have "(z + \ * csqrt (1 - z\<^sup>2))\<^sup>2 \ 1" using assms by (metis Arccos_def add.commute add.left_neutral cancel_comm_monoid_add_class.diff_cancel cos_Arccos csqrt_0 mult_zero_right) then show ?thesis apply (simp add: sin_exp_eq Arccos_def exp_minus) apply (simp add: divide_simps Arccos_body_lemma) apply (simp add: power2_eq_square) done qed lemma cos_sin_csqrt: assumes "0 < cos(Re z) \ cos(Re z) = 0 \ Im z * sin(Re z) \ 0" shows "cos z = csqrt(1 - (sin z)\<^sup>2)" proof (rule csqrt_unique [THEN sym]) show "(cos z)\<^sup>2 = 1 - (sin z)\<^sup>2" by (simp add: cos_squared_eq) qed (use assms in \auto simp: Re_cos Im_cos add_pos_pos mult_le_0_iff zero_le_mult_iff\) lemma sin_cos_csqrt: assumes "0 < sin(Re z) \ sin(Re z) = 0 \ 0 \ Im z * cos(Re z)" shows "sin z = csqrt(1 - (cos z)\<^sup>2)" proof (rule csqrt_unique [THEN sym]) show "(sin z)\<^sup>2 = 1 - (cos z)\<^sup>2" by (simp add: sin_squared_eq) qed (use assms in \auto simp: Re_sin Im_sin add_pos_pos mult_le_0_iff zero_le_mult_iff\) lemma Arcsin_Arccos_csqrt_pos: "(0 < Re z | Re z = 0 & 0 \ Im z) \ Arcsin z = Arccos(csqrt(1 - z\<^sup>2))" by (simp add: Arcsin_def Arccos_def Complex.csqrt_square add.commute) lemma Arccos_Arcsin_csqrt_pos: "(0 < Re z | Re z = 0 & 0 \ Im z) \ Arccos z = Arcsin(csqrt(1 - z\<^sup>2))" by (simp add: Arcsin_def Arccos_def Complex.csqrt_square add.commute) lemma sin_Arccos: "0 < Re z | Re z = 0 & 0 \ Im z \ sin(Arccos z) = csqrt(1 - z\<^sup>2)" by (simp add: Arccos_Arcsin_csqrt_pos) lemma cos_Arcsin: "0 < Re z | Re z = 0 & 0 \ Im z \ cos(Arcsin z) = csqrt(1 - z\<^sup>2)" by (simp add: Arcsin_Arccos_csqrt_pos) subsection\<^marker>\tag unimportant\\Relationship with Arcsin on the Real Numbers\ lemma Im_Arcsin_of_real: assumes "\x\ \ 1" shows "Im (Arcsin (of_real x)) = 0" proof - have "csqrt (1 - (of_real x)\<^sup>2) = (if x^2 \ 1 then sqrt (1 - x^2) else \ * sqrt (x^2 - 1))" by (simp add: of_real_sqrt del: csqrt_of_real_nonneg) then have "cmod (\ * of_real x + csqrt (1 - (of_real x)\<^sup>2))^2 = 1" using assms abs_square_le_1 by (force simp add: Complex.cmod_power2) then have "cmod (\ * of_real x + csqrt (1 - (of_real x)\<^sup>2)) = 1" by (simp add: norm_complex_def) then show ?thesis by (simp add: Im_Arcsin exp_minus) qed corollary\<^marker>\tag unimportant\ Arcsin_in_Reals [simp]: "z \ \ \ \Re z\ \ 1 \ Arcsin z \ \" by (metis Im_Arcsin_of_real Re_complex_of_real Reals_cases complex_is_Real_iff) lemma arcsin_eq_Re_Arcsin: assumes "\x\ \ 1" shows "arcsin x = Re (Arcsin (of_real x))" unfolding arcsin_def proof (rule the_equality, safe) show "- (pi / 2) \ Re (Arcsin (complex_of_real x))" using Re_Ln_pos_le [OF Arcsin_body_lemma, of "of_real x"] by (auto simp: Complex.in_Reals_norm Re_Arcsin) next show "Re (Arcsin (complex_of_real x)) \ pi / 2" using Re_Ln_pos_le [OF Arcsin_body_lemma, of "of_real x"] by (auto simp: Complex.in_Reals_norm Re_Arcsin) next show "sin (Re (Arcsin (complex_of_real x))) = x" using Re_sin [of "Arcsin (of_real x)"] Arcsin_body_lemma [of "of_real x"] by (simp add: Im_Arcsin_of_real assms) next fix x' assume "- (pi / 2) \ x'" "x' \ pi / 2" "x = sin x'" then show "x' = Re (Arcsin (complex_of_real (sin x')))" unfolding sin_of_real [symmetric] by (subst Arcsin_sin) auto qed lemma of_real_arcsin: "\x\ \ 1 \ of_real(arcsin x) = Arcsin(of_real x)" by (metis Im_Arcsin_of_real add.right_neutral arcsin_eq_Re_Arcsin complex_eq mult_zero_right of_real_0) subsection\<^marker>\tag unimportant\\Relationship with Arccos on the Real Numbers\ lemma Im_Arccos_of_real: assumes "\x\ \ 1" shows "Im (Arccos (of_real x)) = 0" proof - have "csqrt (1 - (of_real x)\<^sup>2) = (if x^2 \ 1 then sqrt (1 - x^2) else \ * sqrt (x^2 - 1))" by (simp add: of_real_sqrt del: csqrt_of_real_nonneg) then have "cmod (of_real x + \ * csqrt (1 - (of_real x)\<^sup>2))^2 = 1" using assms abs_square_le_1 by (force simp add: Complex.cmod_power2) then have "cmod (of_real x + \ * csqrt (1 - (of_real x)\<^sup>2)) = 1" by (simp add: norm_complex_def) then show ?thesis by (simp add: Im_Arccos exp_minus) qed corollary\<^marker>\tag unimportant\ Arccos_in_Reals [simp]: "z \ \ \ \Re z\ \ 1 \ Arccos z \ \" by (metis Im_Arccos_of_real Re_complex_of_real Reals_cases complex_is_Real_iff) lemma arccos_eq_Re_Arccos: assumes "\x\ \ 1" shows "arccos x = Re (Arccos (of_real x))" unfolding arccos_def proof (rule the_equality, safe) show "0 \ Re (Arccos (complex_of_real x))" using Im_Ln_pos_le [OF Arccos_body_lemma, of "of_real x"] by (auto simp: Complex.in_Reals_norm Re_Arccos) next show "Re (Arccos (complex_of_real x)) \ pi" using Im_Ln_pos_le [OF Arccos_body_lemma, of "of_real x"] by (auto simp: Complex.in_Reals_norm Re_Arccos) next show "cos (Re (Arccos (complex_of_real x))) = x" using Re_cos [of "Arccos (of_real x)"] Arccos_body_lemma [of "of_real x"] by (simp add: Im_Arccos_of_real assms) next fix x' assume "0 \ x'" "x' \ pi" "x = cos x'" then show "x' = Re (Arccos (complex_of_real (cos x')))" unfolding cos_of_real [symmetric] by (subst Arccos_cos) auto qed lemma of_real_arccos: "\x\ \ 1 \ of_real(arccos x) = Arccos(of_real x)" by (metis Im_Arccos_of_real add.right_neutral arccos_eq_Re_Arccos complex_eq mult_zero_right of_real_0) subsection\<^marker>\tag unimportant\\Some interrelationships among the real inverse trig functions\ lemma arccos_arctan: assumes "-1 < x" "x < 1" shows "arccos x = pi/2 - arctan(x / sqrt(1 - x\<^sup>2))" proof - have "arctan(x / sqrt(1 - x\<^sup>2)) - (pi/2 - arccos x) = 0" proof (rule sin_eq_0_pi) show "- pi < arctan (x / sqrt (1 - x\<^sup>2)) - (pi / 2 - arccos x)" using arctan_lbound [of "x / sqrt(1 - x\<^sup>2)"] arccos_bounded [of x] assms by (simp add: algebra_simps) next show "arctan (x / sqrt (1 - x\<^sup>2)) - (pi / 2 - arccos x) < pi" using arctan_ubound [of "x / sqrt(1 - x\<^sup>2)"] arccos_bounded [of x] assms by (simp add: algebra_simps) next show "sin (arctan (x / sqrt (1 - x\<^sup>2)) - (pi / 2 - arccos x)) = 0" using assms by (simp add: algebra_simps sin_diff cos_add sin_arccos sin_arctan cos_arctan power2_eq_square square_eq_1_iff) qed then show ?thesis by simp qed lemma arcsin_plus_arccos: assumes "-1 \ x" "x \ 1" shows "arcsin x + arccos x = pi/2" proof - have "arcsin x = pi/2 - arccos x" apply (rule sin_inj_pi) using assms arcsin [OF assms] arccos [OF assms] by (auto simp: algebra_simps sin_diff) then show ?thesis by (simp add: algebra_simps) qed lemma arcsin_arccos_eq: "-1 \ x \ x \ 1 \ arcsin x = pi/2 - arccos x" using arcsin_plus_arccos by force lemma arccos_arcsin_eq: "-1 \ x \ x \ 1 \ arccos x = pi/2 - arcsin x" using arcsin_plus_arccos by force lemma arcsin_arctan: "-1 < x \ x < 1 \ arcsin x = arctan(x / sqrt(1 - x\<^sup>2))" by (simp add: arccos_arctan arcsin_arccos_eq) lemma csqrt_1_diff_eq: "csqrt (1 - (of_real x)\<^sup>2) = (if x^2 \ 1 then sqrt (1 - x^2) else \ * sqrt (x^2 - 1))" by ( simp add: of_real_sqrt del: csqrt_of_real_nonneg) lemma arcsin_arccos_sqrt_pos: "0 \ x \ x \ 1 \ arcsin x = arccos(sqrt(1 - x\<^sup>2))" apply (simp add: abs_square_le_1 arcsin_eq_Re_Arcsin arccos_eq_Re_Arccos) apply (subst Arcsin_Arccos_csqrt_pos) apply (auto simp: power_le_one csqrt_1_diff_eq) done lemma arcsin_arccos_sqrt_neg: "-1 \ x \ x \ 0 \ arcsin x = -arccos(sqrt(1 - x\<^sup>2))" using arcsin_arccos_sqrt_pos [of "-x"] by (simp add: arcsin_minus) lemma arccos_arcsin_sqrt_pos: "0 \ x \ x \ 1 \ arccos x = arcsin(sqrt(1 - x\<^sup>2))" apply (simp add: abs_square_le_1 arcsin_eq_Re_Arcsin arccos_eq_Re_Arccos) apply (subst Arccos_Arcsin_csqrt_pos) apply (auto simp: power_le_one csqrt_1_diff_eq) done lemma arccos_arcsin_sqrt_neg: "-1 \ x \ x \ 0 \ arccos x = pi - arcsin(sqrt(1 - x\<^sup>2))" using arccos_arcsin_sqrt_pos [of "-x"] by (simp add: arccos_minus) subsection\<^marker>\tag unimportant\\Continuity results for arcsin and arccos\ lemma continuous_on_Arcsin_real [continuous_intros]: "continuous_on {w \ \. \Re w\ \ 1} Arcsin" proof - have "continuous_on {w \ \. \Re w\ \ 1} (\x. complex_of_real (arcsin (Re x))) = continuous_on {w \ \. \Re w\ \ 1} (\x. complex_of_real (Re (Arcsin (of_real (Re x)))))" by (rule continuous_on_cong [OF refl]) (simp add: arcsin_eq_Re_Arcsin) also have "... = ?thesis" by (rule continuous_on_cong [OF refl]) simp finally show ?thesis using continuous_on_arcsin [OF continuous_on_Re [OF continuous_on_id], of "{w \ \. \Re w\ \ 1}"] continuous_on_of_real by fastforce qed lemma continuous_within_Arcsin_real: "continuous (at z within {w \ \. \Re w\ \ 1}) Arcsin" proof (cases "z \ {w \ \. \Re w\ \ 1}") case True then show ?thesis using continuous_on_Arcsin_real continuous_on_eq_continuous_within by blast next case False with closed_real_abs_le [of 1] show ?thesis by (rule continuous_within_closed_nontrivial) qed lemma continuous_on_Arccos_real: "continuous_on {w \ \. \Re w\ \ 1} Arccos" proof - have "continuous_on {w \ \. \Re w\ \ 1} (\x. complex_of_real (arccos (Re x))) = continuous_on {w \ \. \Re w\ \ 1} (\x. complex_of_real (Re (Arccos (of_real (Re x)))))" by (rule continuous_on_cong [OF refl]) (simp add: arccos_eq_Re_Arccos) also have "... = ?thesis" by (rule continuous_on_cong [OF refl]) simp finally show ?thesis using continuous_on_arccos [OF continuous_on_Re [OF continuous_on_id], of "{w \ \. \Re w\ \ 1}"] continuous_on_of_real by fastforce qed lemma continuous_within_Arccos_real: "continuous (at z within {w \ \. \Re w\ \ 1}) Arccos" proof (cases "z \ {w \ \. \Re w\ \ 1}") case True then show ?thesis using continuous_on_Arccos_real continuous_on_eq_continuous_within by blast next case False with closed_real_abs_le [of 1] show ?thesis by (rule continuous_within_closed_nontrivial) qed lemma sinh_ln_complex: "x \ 0 \ sinh (ln x :: complex) = (x - inverse x) / 2" by (simp add: sinh_def exp_minus scaleR_conv_of_real exp_of_real) lemma cosh_ln_complex: "x \ 0 \ cosh (ln x :: complex) = (x + inverse x) / 2" by (simp add: cosh_def exp_minus scaleR_conv_of_real) lemma tanh_ln_complex: "x \ 0 \ tanh (ln x :: complex) = (x ^ 2 - 1) / (x ^ 2 + 1)" by (simp add: tanh_def sinh_ln_complex cosh_ln_complex divide_simps power2_eq_square) subsection\Roots of unity\ theorem complex_root_unity: fixes j::nat assumes "n \ 0" shows "exp(2 * of_real pi * \ * of_nat j / of_nat n)^n = 1" proof - have *: "of_nat j * (complex_of_real pi * 2) = complex_of_real (2 * real j * pi)" by (simp) then show ?thesis apply (simp add: exp_of_nat_mult [symmetric] mult_ac exp_Euler) apply (simp only: * cos_of_real sin_of_real) apply simp done qed lemma complex_root_unity_eq: fixes j::nat and k::nat assumes "1 \ n" shows "(exp(2 * of_real pi * \ * of_nat j / of_nat n) = exp(2 * of_real pi * \ * of_nat k / of_nat n) \ j mod n = k mod n)" proof - have "(\z::int. \ * (of_nat j * (of_real pi * 2)) = \ * (of_nat k * (of_real pi * 2)) + \ * (of_int z * (of_nat n * (of_real pi * 2)))) \ (\z::int. of_nat j * (\ * (of_real pi * 2)) = (of_nat k + of_nat n * of_int z) * (\ * (of_real pi * 2)))" by (simp add: algebra_simps) also have "... \ (\z::int. of_nat j = of_nat k + of_nat n * (of_int z :: complex))" by simp also have "... \ (\z::int. of_nat j = of_nat k + of_nat n * z)" by (metis (mono_tags, opaque_lifting) of_int_add of_int_eq_iff of_int_mult of_int_of_nat_eq) also have "... \ int j mod int n = int k mod int n" by (auto simp: mod_eq_dvd_iff dvd_def algebra_simps) also have "... \ j mod n = k mod n" by (metis of_nat_eq_iff zmod_int) finally have "(\z. \ * (of_nat j * (of_real pi * 2)) = \ * (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 *) qed corollary bij_betw_roots_unity: "bij_betw (\j. exp(2 * of_real pi * \ * of_nat j / of_nat n)) {.. * of_nat j / of_nat n) | j. j < n}" by (auto simp: bij_betw_def inj_on_def complex_root_unity_eq) lemma complex_root_unity_eq_1: fixes j::nat and k::nat assumes "1 \ n" shows "exp(2 * of_real pi * \ * of_nat j / of_nat n) = 1 \ n dvd j" proof - have "1 = exp(2 * of_real pi * \ * (of_nat n / of_nat n))" using assms by simp then have "exp(2 * of_real pi * \ * (of_nat j / of_nat n)) = 1 \ j mod n = n mod n" using complex_root_unity_eq [of n j n] assms by simp then show ?thesis by auto qed lemma finite_complex_roots_unity_explicit: "finite {exp(2 * of_real pi * \ * of_nat j / of_nat n) | j::nat. j < n}" by simp lemma card_complex_roots_unity_explicit: "card {exp(2 * of_real pi * \ * of_nat j / of_nat n) | j::nat. j < n} = n" by (simp add: Finite_Set.bij_betw_same_card [OF bij_betw_roots_unity, symmetric]) lemma complex_roots_unity: assumes "1 \ n" shows "{z::complex. z^n = 1} = {exp(2 * of_real pi * \ * of_nat j / of_nat n) | j. j < n}" apply (rule Finite_Set.card_seteq [symmetric]) using assms apply (auto simp: card_complex_roots_unity_explicit finite_roots_unity complex_root_unity card_roots_unity) done lemma card_complex_roots_unity: "1 \ n \ card {z::complex. z^n = 1} = n" by (simp add: card_complex_roots_unity_explicit complex_roots_unity) lemma complex_not_root_unity: "1 \ n \ \u::complex. norm u = 1 \ u^n \ 1" apply (rule_tac x="exp (of_real pi * \ * of_real (1 / n))" in exI) apply (auto simp: Re_complex_div_eq_0 exp_of_nat_mult [symmetric] mult_ac exp_Euler) done end diff --git a/src/HOL/Analysis/Elementary_Topology.thy b/src/HOL/Analysis/Elementary_Topology.thy --- a/src/HOL/Analysis/Elementary_Topology.thy +++ b/src/HOL/Analysis/Elementary_Topology.thy @@ -1,2684 +1,2693 @@ (* Author: L C Paulson, University of Cambridge Author: Amine Chaieb, University of Cambridge Author: Robert Himmelmann, TU Muenchen Author: Brian Huffman, Portland State University *) chapter \Topology\ theory Elementary_Topology imports "HOL-Library.Set_Idioms" "HOL-Library.Disjoint_Sets" Product_Vector begin section \Elementary Topology\ subsubsection\<^marker>\tag unimportant\ \Affine transformations of intervals\ lemma real_affinity_le: "0 < m \ m * x + c \ y \ x \ inverse m * y + - (c / m)" for m :: "'a::linordered_field" by (simp add: field_simps) lemma real_le_affinity: "0 < m \ y \ m * x + c \ inverse m * y + - (c / m) \ x" for m :: "'a::linordered_field" by (simp add: field_simps) lemma real_affinity_lt: "0 < m \ m * x + c < y \ x < inverse m * y + - (c / m)" for m :: "'a::linordered_field" by (simp add: field_simps) lemma real_lt_affinity: "0 < m \ y < m * x + c \ inverse m * y + - (c / m) < x" for m :: "'a::linordered_field" by (simp add: field_simps) lemma real_affinity_eq: "m \ 0 \ m * x + c = y \ x = inverse m * y + - (c / m)" for m :: "'a::linordered_field" by (simp add: field_simps) lemma real_eq_affinity: "m \ 0 \ y = m * x + c \ inverse m * y + - (c / m) = x" for m :: "'a::linordered_field" by (simp add: field_simps) subsection \Topological Basis\ context topological_space begin definition\<^marker>\tag important\ "topological_basis B \ (\b\B. open b) \ (\x. open x \ (\B'. B' \ B \ \B' = x))" 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 lemma topological_basis_iff: assumes "\B'. B' \ B \ open B'" shows "topological_basis B \ (\O'. open O' \ (\x\O'. \B'\B. x \ B' \ B' \ O'))" (is "_ \ ?rhs") proof safe fix O' and x::'a assume H: "topological_basis B" "open O'" "x \ O'" then have "(\B'\B. \B' = O')" by (simp add: topological_basis_def) then obtain B' where "B' \ B" "O' = \B'" by auto then show "\B'\B. x \ B' \ B' \ O'" using H by auto next assume H: ?rhs show "topological_basis B" using assms unfolding topological_basis_def proof safe fix O' :: "'a set" assume "open O'" with H obtain f where "\x\O'. f x \ B \ x \ f x \ f x \ O'" by (force intro: bchoice simp: Bex_def) then show "\B'\B. \B' = O'" by (auto intro: exI[where x="{f x |x. x \ O'}"]) qed qed lemma topological_basisI: 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 lemma topological_basisE: fixes O' assumes "topological_basis B" 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 lemma topological_basis_open: assumes "topological_basis B" and "X \ B" shows "open X" using assms by (simp add: topological_basis_def) lemma topological_basis_imp_subbasis: assumes B: "topological_basis B" shows "open = generate_topology B" proof (intro ext iffI) fix S :: "'a set" assume "open S" with B obtain B' where "B' \ B" "S = \B'" unfolding topological_basis_def by blast then show "generate_topology B S" by (auto intro: generate_topology.intros dest: topological_basis_open) next fix S :: "'a set" assume "generate_topology B S" then show "open S" by induct (auto dest: topological_basis_open[OF B]) qed 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'" 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 end lemma topological_basis_prod: 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) subsection \Countable Basis\ locale\<^marker>\tag important\ countable_basis = topological_space p for p::"'a set \ bool" + fixes B :: "'a set set" assumes is_basis: "topological_basis B" and countable_basis: "countable B" begin lemma open_countable_basis_ex: assumes "p X" shows "\B' \ B. X = \B'" using assms countable_basis is_basis unfolding topological_basis_def by blast lemma open_countable_basisE: assumes "p X" obtains B' where "B' \ B" "X = \B'" using assms open_countable_basis_ex by atomize_elim simp lemma countable_dense_exists: "\D::'a set. countable D \ (\X. p X \ X \ {} \ (\d \ D. d \ X))" proof - let ?f = "(\B'. SOME x. x \ B')" have "countable (?f ` B)" using countable_basis by simp with basis_dense[OF is_basis, of ?f] show ?thesis by (intro exI[where x="?f ` B"]) (metis (mono_tags) all_not_in_conv imageI someI) qed lemma countable_dense_setE: obtains D :: "'a set" where "countable D" "\X. p X \ X \ {} \ \d \ D. d \ X" using countable_dense_exists by blast end lemma countable_basis_openI: "countable_basis open B" if "countable B" "topological_basis B" using that by unfold_locales (simp_all add: topological_basis topological_space.topological_basis topological_space_axioms) lemma (in first_countable_topology) first_countable_basisE: fixes x :: 'a obtains \ where "countable \" "\A. A \ \ \ x \ A" "\A. A \ \ \ open A" "\S. open S \ x \ S \ (\A\\. A \ S)" 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) qed lemma (in first_countable_topology) first_countable_basis_Int_stableE: obtains \ where "countable \" "\A. A \ \ \ x \ A" "\A. A \ \ \ open A" "\S. open S \ x \ S \ (\A\\. A \ S)" "\A B. A \ \ \ B \ \ \ A \ B \ \" proof atomize_elim obtain \ where \: "countable \" "\B. B \ \ \ x \ B" "\B. B \ \ \ open B" "\S. open S \ x \ S \ \B\\. B \ S" by (rule first_countable_basisE) blast define \ where [abs_def]: "\ = (\N. \((\n. from_nat_into \ n) ` N)) ` (Collect finite::nat set set)" then show "\\. countable \ \ (\A. A \ \ \ x \ A) \ (\A. A \ \ \ open A) \ (\S. open S \ x \ S \ (\A\\. A \ S)) \ (\A B. A \ \ \ B \ \ \ A \ B \ \)" proof (safe intro!: exI[where x=\]) show "countable \" unfolding \_def by (intro countable_image countable_Collect_finite) fix A assume "A \ \" then show "x \ A" "open A" using \(4)[OF open_UNIV] by (auto simp: \_def intro: \ from_nat_into) next let ?int = "\N. \(from_nat_into \ ` N)" fix A B assume "A \ \" "B \ \" then obtain N M where "A = ?int N" "B = ?int M" "finite (N \ M)" by (auto simp: \_def) then show "A \ B \ \" by (auto simp: \_def intro!: image_eqI[where x="N \ M"]) next 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}"]) qed qed lemma (in topological_space) first_countableI: assumes "countable \" and 1: "\A. A \ \ \ x \ A" "\A. A \ \ \ open A" and 2: "\S. open S \ x \ S \ \A\\. A \ S" shows "\\::nat \ 'a set. (\i. x \ \ i \ open (\ i)) \ (\S. open S \ x \ S \ (\i. \ i \ S))" proof (safe intro!: exI[of _ "from_nat_into \"]) fix i have "\ \ {}" using 2[of UNIV] by auto show "x \ from_nat_into \ i" "open (from_nat_into \ i)" 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 qed instance prod :: (first_countable_topology, first_countable_topology) first_countable_topology proof fix x :: "'a \ 'b" obtain \ where \: "countable \" "\a. a \ \ \ fst x \ a" "\a. a \ \ \ open a" "\S. open S \ fst x \ S \ \a\\. a \ S" by (rule first_countable_basisE[of "fst x"]) blast obtain B where B: "countable B" "\a. a \ B \ snd x \ a" "\a. a \ B \ open a" "\S. open S \ snd x \ S \ \a\B. a \ S" by (rule first_countable_basisE[of "snd x"]) blast show "\\::nat \ ('a \ 'b) set. (\i. x \ \ i \ open (\ i)) \ (\S. open S \ x \ S \ (\i. \ i \ S))" proof (rule first_countableI[of "(\(a, b). a \ b) ` (\ \ B)"], safe) fix a b assume x: "a \ \" "b \ B" show "x \ a \ b" by (simp add: \(2) B(2) mem_Times_iff x) show "open (a \ b)" by (simp add: \(3) B(3) open_Times x) next fix S assume "open S" "x \ S" 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 ultimately show "\a\(\(a, b). a \ b) ` (\ \ B). a \ S" by (auto intro!: bexI[of _ "a \ b"] bexI[of _ a] bexI[of _ b]) qed (simp add: \ B) qed class second_countable_topology = topological_space + assumes ex_countable_subbasis: "\B::'a set set. countable B \ open = generate_topology B" begin lemma ex_countable_basis: "\B::'a set set. countable B \ topological_basis B" proof - from ex_countable_subbasis obtain B where B: "countable B" "open = generate_topology B" by blast let ?B = "Inter ` {b. finite b \ b \ B }" show ?thesis proof (intro exI conjI) show "countable ?B" by (intro countable_image countable_Collect_finite_subset B) { fix S assume "open S" then have "\B'\{b. finite b \ b \ B}. (\b\B'. \b) = S" unfolding B proof induct case UNIV show ?case by (intro exI[of _ "{{}}"]) simp next case (Int a b) then obtain x y where x: "a = \(Inter ` x)" "\i. i \ x \ finite i \ i \ B" and y: "b = \(Inter ` y)" "\i. i \ y \ finite i \ i \ B" by blast show ?case unfolding x y Int_UN_distrib2 by (intro exI[of _ "{i \ j| i j. i \ x \ j \ y}"]) (auto dest: x(2) y(2)) next case (UN K) then have "\k\K. \B'\{b. finite b \ b \ B}. \ (Inter ` B') = k" by auto then obtain k where "\ka\K. k ka \ {b. finite b \ b \ B} \ \(Inter ` (k ka)) = ka" unfolding bchoice_iff .. then show "\B'\{b. finite b \ b \ B}. \ (Inter ` B') = \K" by (intro exI[of _ "\(k ` K)"]) auto next case (Basis S) then show ?case by (intro exI[of _ "{{S}}"]) auto qed then have "(\B'\Inter ` {b. finite b \ b \ B}. \B' = S)" unfolding subset_image_iff by blast } then show "topological_basis ?B" unfolding topological_basis_def by (safe intro!: open_Inter) (simp_all add: B generate_topology.Basis subset_eq) qed qed end lemma univ_second_countable: obtains \ :: "'a::second_countable_topology set set" where "countable \" "\C. C \ \ \ open C" "\S. open S \ \U. U \ \ \ S = \U" by (metis ex_countable_basis topological_basis_def) proposition Lindelof: fixes \ :: "'a::second_countable_topology set set" assumes \: "\S. S \ \ \ open S" obtains \' where "\' \ \" "countable \'" "\\' = \\" proof - obtain \ :: "'a set set" where "countable \" "\C. C \ \ \ open C" and \: "\S. open S \ \U. U \ \ \ S = \U" 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 have "\S. \U. S \ \ \ U \ \ \ S \ U" by (simp add: \_def) then obtain G where G: "\S. S \ \ \ G S \ \ \ S \ G S" by metis have "\\ \ \\" unfolding \_def by (blast dest: \ \) moreover have "\\ \ \\" using \_def by blast ultimately have eq1: "\\ = \\" .. have eq2: "\\ = \ (G ` \)" using G eq1 by auto show ?thesis apply (rule_tac \' = "G ` \" in that) using G \countable \\ by (auto simp: eq1 eq2) qed lemma countable_disjoint_open_subsets: fixes \ :: "'a::second_countable_topology set set" assumes "\S. S \ \ \ open S" and pw: "pairwise disjnt \" shows "countable \" proof - obtain \' where "\' \ \" "countable \'" "\\' = \\" by (meson assms Lindelof) with pw have "\ \ insert {} \'" by (fastforce simp add: pairwise_def disjnt_iff) then show ?thesis by (simp add: \countable \'\ countable_subset) qed sublocale second_countable_topology < countable_basis "open" "SOME B. countable B \ topological_basis B" using someI_ex[OF ex_countable_basis] by unfold_locales safe instance prod :: (second_countable_topology, second_countable_topology) second_countable_topology proof obtain A :: "'a set set" where "countable A" "topological_basis A" using ex_countable_basis by auto moreover obtain B :: "'b set set" where "countable B" "topological_basis B" using ex_countable_basis by auto ultimately show "\B::('a \ 'b) set set. countable B \ open = generate_topology B" by (auto intro!: exI[of _ "(\(a, b). a \ b) ` (A \ B)"] topological_basis_prod topological_basis_imp_subbasis) qed instance second_countable_topology \ first_countable_topology proof fix x :: 'a define B :: "'a set set" where "B = (SOME B. countable B \ topological_basis B)" then have B: "countable B" "topological_basis B" using countable_basis is_basis by (auto simp: countable_basis is_basis) then show "\A::nat \ 'a set. (\i. x \ A i \ open (A i)) \ (\S. open S \ x \ S \ (\i. A i \ S))" by (intro first_countableI[of "{b\B. x \ b}"]) (fastforce simp: topological_space_class.topological_basis_def)+ qed instance nat :: second_countable_topology proof show "\B::nat set set. countable B \ open = generate_topology B" by (intro exI[of _ "range lessThan \ range greaterThan"]) (auto simp: open_nat_def) qed lemma countable_separating_set_linorder1: shows "\B::('a::{linorder_topology, second_countable_topology} set). countable B \ (\x y. x < y \ (\b \ B. x < b \ b \ y))" proof - obtain A::"'a set set" where "countable A" "topological_basis A" using ex_countable_basis by auto define B1 where "B1 = {(LEAST x. x \ U)| U. U \ A}" then have "countable B1" using \countable A\ by (simp add: Setcompr_eq_image) define B2 where "B2 = {(SOME x. x \ U)| U. U \ A}" then have "countable B2" using \countable A\ by (simp add: Setcompr_eq_image) have "\b \ B1 \ B2. x < b \ b \ y" if "x < y" for x y proof (cases) assume "\z. x < z \ z < y" then obtain z where z: "x < z \ z < y" by auto define U where "U = {x<.. U" using z U_def by simp ultimately obtain V where "V \ A" "z \ V" "V \ U" using topological_basisE[OF \topological_basis A\] by auto define w where "w = (SOME x. x \ V)" then have "w \ V" using \z \ V\ by (metis someI2) then have "x < w \ w \ y" using \w \ V\ \V \ U\ U_def by fastforce moreover have "w \ B1 \ B2" using w_def B2_def \V \ A\ by auto ultimately show ?thesis by auto next assume "\(\z. x < z \ z < y)" then have *: "\z. z > x \ z \ y" by auto define U where "U = {x<..}" then have "open U" by simp moreover have "y \ U" using \x < y\ U_def by simp ultimately obtain "V" where "V \ A" "y \ V" "V \ U" using topological_basisE[OF \topological_basis A\] by auto have "U = {y..}" unfolding U_def using * \x < y\ by auto then have "V \ {y..}" using \V \ U\ by simp then have "(LEAST w. w \ V) = y" using \y \ V\ by (meson Least_equality atLeast_iff subsetCE) then have "y \ B1 \ B2" using \V \ A\ B1_def by auto moreover have "x < y \ y \ y" using \x < y\ by simp ultimately show ?thesis by auto qed moreover have "countable (B1 \ B2)" using \countable B1\ \countable B2\ by simp ultimately show ?thesis by auto qed lemma countable_separating_set_linorder2: shows "\B::('a::{linorder_topology, second_countable_topology} set). countable B \ (\x y. x < y \ (\b \ B. x \ b \ b < y))" proof - obtain A::"'a set set" where "countable A" "topological_basis A" using ex_countable_basis by auto define B1 where "B1 = {(GREATEST x. x \ U) | U. U \ A}" then have "countable B1" using \countable A\ by (simp add: Setcompr_eq_image) define B2 where "B2 = {(SOME x. x \ U)| U. U \ A}" then have "countable B2" using \countable A\ by (simp add: Setcompr_eq_image) have "\b \ B1 \ B2. x \ b \ b < y" if "x < y" for x y proof (cases) assume "\z. x < z \ z < y" then obtain z where z: "x < z \ z < y" by auto define U where "U = {x<.. U" using z U_def by simp ultimately obtain "V" where "V \ A" "z \ V" "V \ U" using topological_basisE[OF \topological_basis A\] by auto define w where "w = (SOME x. x \ V)" then have "w \ V" using \z \ V\ by (metis someI2) then have "x \ w \ w < y" using \w \ V\ \V \ U\ U_def by fastforce moreover have "w \ B1 \ B2" using w_def B2_def \V \ A\ by auto ultimately show ?thesis by auto next assume "\(\z. x < z \ z < y)" then have *: "\z. z < y \ z \ x" using leI by blast define U where "U = {.. U" using \x < y\ U_def by simp ultimately obtain "V" where "V \ A" "x \ V" "V \ U" using topological_basisE[OF \topological_basis A\] by auto have "U = {..x}" unfolding U_def using * \x < y\ by auto then have "V \ {..x}" using \V \ U\ by simp then have "(GREATEST x. x \ V) = x" using \x \ V\ by (meson Greatest_equality atMost_iff subsetCE) then have "x \ B1 \ B2" using \V \ A\ B1_def by auto moreover have "x \ x \ x < y" using \x < y\ by simp ultimately show ?thesis by auto qed moreover have "countable (B1 \ B2)" using \countable B1\ \countable B2\ by simp ultimately show ?thesis by auto qed lemma countable_separating_set_dense_linorder: shows "\B::('a::{linorder_topology, dense_linorder, second_countable_topology} set). countable B \ (\x y. x < y \ (\b \ B. x < b \ b < y))" proof - obtain B::"'a set" where B: "countable B" "\x y. x < y \ (\b \ B. x < b \ b \ y)" using countable_separating_set_linorder1 by auto have "\b \ B. x < b \ b < y" if "x < y" for x y proof - obtain z where "x < z" "z < y" using \x < y\ dense by blast then obtain b where "b \ B" "x < b \ b \ z" using B(2) by auto then have "x < b \ b < y" using \z < y\ by auto then show ?thesis using \b \ B\ by auto qed then show ?thesis using B(1) by auto qed subsection \Polish spaces\ text \Textbooks define Polish spaces as completely metrizable. We assume the topology to be complete for a given metric.\ class polish_space = complete_space + second_countable_topology subsection \Limit Points\ definition\<^marker>\tag important\ (in topological_space) islimpt:: "'a \ 'a set \ bool" (infixr "islimpt" 60) where "x islimpt S \ (\T. x\T \ open T \ (\y\S. y\T \ y\x))" lemma islimptI: assumes "\T. x \ T \ open T \ \y\S. y \ T \ y \ x" shows "x islimpt S" using assms unfolding islimpt_def by auto lemma islimptE: assumes "x islimpt S" and "x \ T" and "open T" obtains y where "y \ S" and "y \ T" and "y \ x" using assms unfolding islimpt_def by auto lemma islimpt_iff_eventually: "x islimpt S \ \ eventually (\y. y \ S) (at x)" unfolding islimpt_def eventually_at_topological by auto lemma islimpt_subset: "x islimpt S \ S \ T \ x islimpt T" unfolding islimpt_def by fast lemma islimpt_UNIV_iff: "x islimpt UNIV \ \ open {x}" unfolding islimpt_def by (safe, fast, case_tac "T = {x}", fast, fast) lemma islimpt_punctured: "x islimpt S = x islimpt (S-{x})" unfolding islimpt_def by blast text \A perfect space has no isolated points.\ lemma islimpt_UNIV [simp, intro]: "x islimpt UNIV" for x :: "'a::perfect_space" 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 lemma islimpt_EMPTY[simp]: "\ x islimpt {}" by (auto simp: islimpt_def) lemma islimpt_Un: "x islimpt (S \ T) \ x islimpt S \ x islimpt T" by (simp add: islimpt_iff_eventually eventually_conj_iff) lemma islimpt_insert: fixes x :: "'a::t1_space" shows "x islimpt (insert a s) \ x islimpt s" proof assume *: "x islimpt (insert a s)" show "x islimpt s" proof (rule islimptI) fix t assume t: "x \ t" "open t" show "\y\s. y \ t \ y \ x" proof (cases "x = a") case True obtain y where "y \ insert a s" "y \ t" "y \ x" using * t by (rule islimptE) with \x = a\ show ?thesis by auto next case False with t have t': "x \ t - {a}" "open (t - {a})" by (simp_all add: open_Diff) obtain y where "y \ insert a s" "y \ t - {a}" "y \ x" using * t' by (rule islimptE) then show ?thesis by auto qed qed next 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" 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" by (simp add: islimpt_Un islimpt_finite) lemma islimpt_eq_acc_point: fixes l :: "'a :: t1_space" shows "l islimpt S \ (\U. l\U \ open U \ infinite (U \ S))" proof (safe intro!: islimptI) fix U assume "l islimpt S" "l \ U" "open U" "finite (U \ S)" then have "l islimpt S" "l \ (U - (U \ S - {l}))" "open (U - (U \ S - {l}))" by (auto intro: finite_imp_closed) then show False by (rule islimptE) auto 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 then show "\y\S. y \ T \ y \ l" by auto qed lemma acc_point_range_imp_convergent_subsequence: fixes l :: "'a :: first_countable_topology" assumes l: "\U. l\U \ open U \ infinite (U \ range f)" shows "\r::nat\nat. strict_mono r \ (f \ r) \ l" proof - from countable_basis_at_decseq[of l] obtain A where A: "\i. open (A i)" "\i. l \ A i" "\S. open S \ l \ S \ eventually (\i. A i \ S) sequentially" by blast define s where "s n i = (SOME j. i < j \ f j \ A (Suc n))" for n i { fix n i 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 then have "\a. i < a \ f a \ A (Suc n)" by (auto simp: not_le) then have "i < s n i" "f (s n i) \ A (Suc n)" unfolding s_def by (auto intro: someI2_ex) } note s = this define r where "r = rec_nat (s 0 0) s" have "strict_mono r" by (auto simp: r_def s strict_mono_Suc_iff) moreover have "(\n. f (r n)) \ l" proof (rule topological_tendstoI) fix S assume "open S" "l \ S" with A(3) have "eventually (\i. A i \ S) sequentially" by auto moreover { fix i assume "Suc 0 \ i" then have "f (r i) \ A i" by (cases i) (simp_all add: r_def s) } then have "eventually (\i. f (r i) \ A i) sequentially" by (auto simp: eventually_sequentially) ultimately show "eventually (\i. f (r i) \ S) sequentially" by eventually_elim auto qed ultimately show "\r::nat\nat. strict_mono r \ (f \ r) \ l" by (auto simp: convergent_def comp_def) qed lemma islimpt_range_imp_convergent_subsequence: fixes l :: "'a :: {t1_space, first_countable_topology}" assumes l: "l islimpt (range f)" shows "\r::nat\nat. strict_mono r \ (f \ r) \ l" using l unfolding islimpt_eq_acc_point by (rule acc_point_range_imp_convergent_subsequence) lemma sequence_unique_limpt: fixes f :: "nat \ 'a::t2_space" assumes "(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 have "UNIV = {.. {N..}" by auto then have "l' islimpt (f ` ({.. {N..}))" using assms(2) by simp then have "l' islimpt (f ` {.. f ` {N..})" by (simp add: image_Un) then have "l' islimpt (f ` {N..})" by (simp add: islimpt_Un_finite) then obtain y where "y \ f ` {N..}" "y \ s" "y \ l'" 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 qed subsection \Interior of a Set\ definition\<^marker>\tag important\ interior :: "('a::topological_space) set \ 'a set" where "interior S = \{T. open T \ T \ S}" lemma interiorI [intro?]: assumes "open T" and "x \ T" and "T \ S" shows "x \ interior S" using assms unfolding interior_def by fast lemma interiorE [elim?]: assumes "x \ interior S" obtains T where "open T" and "x \ T" and "T \ S" using assms unfolding interior_def by fast lemma open_interior [simp, intro]: "open (interior S)" by (simp add: interior_def open_Union) lemma interior_subset: "interior S \ S" by (auto simp: interior_def) lemma interior_maximal: "T \ S \ open T \ T \ interior S" by (auto simp: interior_def) lemma interior_open: "open S \ interior S = S" by (intro equalityI interior_subset interior_maximal subset_refl) lemma interior_eq: "interior S = S \ open S" by (metis open_interior interior_open) lemma open_subset_interior: "open S \ S \ interior T \ S \ T" by (metis interior_maximal interior_subset subset_trans) lemma interior_empty [simp]: "interior {} = {}" using open_empty by (rule interior_open) lemma interior_UNIV [simp]: "interior UNIV = UNIV" using open_UNIV by (rule interior_open) lemma interior_interior [simp]: "interior (interior S) = interior S" using open_interior by (rule interior_open) lemma interior_mono: "S \ T \ interior S \ interior T" by (auto simp: interior_def) lemma interior_unique: assumes "T \ S" and "open T" assumes "\T'. T' \ S \ open T' \ T' \ T" shows "interior S = T" by (intro equalityI assms interior_subset open_interior interior_maximal) lemma interior_singleton [simp]: "interior {a} = {}" for a :: "'a::perfect_space" by (meson interior_eq interior_subset not_open_singleton subset_singletonD) lemma interior_Int [simp]: "interior (S \ T) = interior S \ interior T" by (meson Int_mono Int_subset_iff antisym_conv interior_maximal interior_subset open_Int open_interior) lemma eventually_nhds_in_nhd: "x \ interior s \ eventually (\y. y \ s) (nhds x)" using interior_subset[of s] by (subst eventually_nhds) blast lemma interior_limit_point [intro]: 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 lemma interior_closed_Un_empty_interior: assumes cS: "closed S" and iT: "interior T = {}" shows "interior (S \ T) = interior S" proof show "interior S \ interior (S \ T)" by (rule interior_mono) (rule Un_upper1) show "interior (S \ T) \ interior S" proof fix x assume "x \ interior (S \ T)" then obtain R where "open R" "x \ R" "R \ S \ T" .. show "x \ interior S" proof (rule ccontr) 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 qed qed qed lemma interior_Times: "interior (A \ B) = interior A \ interior B" proof (rule interior_unique) show "interior A \ interior B \ A \ B" by (intro Sigma_mono interior_subset) show "open (interior A \ interior B)" by (intro open_Times open_interior) fix T assume "T \ A \ B" and "open T" then show "T \ interior A \ interior B" proof safe fix x y assume "(x, y) \ T" then obtain C D where "open C" "open D" "C \ D \ T" "x \ C" "y \ D" using \open T\ unfolding open_prod_def by fast then have "open C" "open D" "C \ A" "D \ B" "x \ C" "y \ D" using \T \ A \ B\ by auto then show "x \ interior A" and "y \ interior B" by (auto intro: interiorI) qed qed lemma interior_Ici: fixes x :: "'a :: {dense_linorder,linorder_topology}" assumes "b < x" shows "interior {x ..} = {x <..}" proof (rule interior_unique) fix T assume "T \ {x ..}" "open T" moreover have "x \ T" proof assume "x \ T" obtain y where "y < x" "{y <.. x} \ T" using open_left[OF \open T\ \x \ T\ \b < x\] by auto with dense[OF \y < x\] obtain z where "z \ T" "z < x" by (auto simp: subset_eq Ball_def) with \T \ {x ..}\ show False by auto qed ultimately show "T \ {x <..}" by (auto simp: subset_eq less_le) qed auto lemma interior_Iic: fixes x :: "'a ::{dense_linorder,linorder_topology}" assumes "x < b" shows "interior {.. x} = {..< x}" proof (rule interior_unique) fix T assume "T \ {.. x}" "open T" moreover have "x \ T" proof assume "x \ T" obtain y where "x < y" "{x ..< y} \ T" using open_right[OF \open T\ \x \ T\ \x < b\] by auto with dense[OF \x < y\] obtain z where "z \ T" "x < z" by (auto simp: subset_eq Ball_def less_le) with \T \ {.. x}\ show False by auto qed ultimately show "T \ {..< x}" by (auto simp: subset_eq less_le) qed auto lemma countable_disjoint_nonempty_interior_subsets: fixes \ :: "'a::second_countable_topology set set" assumes pw: "pairwise disjnt \" and int: "\S. \S \ \; interior S = {}\ \ S = {}" shows "countable \" proof (rule countable_image_inj_on) have "disjoint (interior ` \)" using pw by (simp add: disjoint_image_subset interior_subset) then show "countable (interior ` \)" by (auto intro: countable_disjoint_open_subsets) show "inj_on interior \" using pw apply (clarsimp simp: inj_on_def pairwise_def) apply (metis disjnt_def disjnt_subset1 inf.orderE int interior_subset) done qed subsection \Closure of a Set\ definition\<^marker>\tag important\ closure :: "('a::topological_space) set \ 'a set" where "closure S = S \ {x . x islimpt S}" lemma interior_closure: "interior S = - (closure (- S))" by (auto simp: interior_def closure_def islimpt_def) lemma closure_interior: "closure S = - interior (- S)" by (simp add: interior_closure) lemma closed_closure[simp, intro]: "closed (closure S)" by (simp add: closure_interior closed_Compl) lemma closure_subset: "S \ closure S" by (simp add: closure_def) lemma closure_hull: "closure S = closed hull S" by (auto simp: hull_def closure_interior interior_def) lemma closure_eq: "closure S = S \ closed S" unfolding closure_hull using closed_Inter by (rule hull_eq) lemma closure_closed [simp]: "closed S \ closure S = S" by (simp only: closure_eq) lemma closure_closure [simp]: "closure (closure S) = closure S" unfolding closure_hull by (rule hull_hull) lemma closure_mono: "S \ T \ closure S \ closure T" unfolding closure_hull by (rule hull_mono) lemma closure_minimal: "S \ T \ closed T \ closure S \ T" unfolding closure_hull by (rule hull_minimal) lemma closure_unique: assumes "S \ T" and "closed T" and "\T'. S \ T' \ closed T' \ T \ T'" shows "closure S = T" using assms unfolding closure_hull by (rule hull_unique) lemma closure_empty [simp]: "closure {} = {}" using closed_empty by (rule closure_closed) lemma closure_UNIV [simp]: "closure UNIV = UNIV" using closed_UNIV by (rule closure_closed) lemma closure_Un [simp]: "closure (S \ T) = closure S \ closure T" by (simp add: closure_interior) lemma closure_eq_empty [iff]: "closure S = {} \ S = {}" using closure_empty closure_subset[of S] by blast lemma closure_subset_eq: "closure S \ S \ closed S" using closure_eq[of S] closure_subset[of S] by simp lemma open_Int_closure_eq_empty: "open S \ (S \ closure T) = {} \ S \ T = {}" using open_subset_interior[of S "- T"] using interior_subset[of "- T"] by (auto simp: closure_interior) lemma open_Int_closure_subset: "open S \ S \ closure T \ closure (S \ T)" 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 with * show "x \ closure (S \ T)" unfolding closure_def by blast qed lemma closure_complement: "closure (- S) = - interior S" by (simp add: closure_interior) lemma interior_complement: "interior (- S) = - closure S" by (simp add: closure_interior) lemma interior_diff: "interior(S - T) = interior S - closure T" by (simp add: Diff_eq interior_complement) lemma closure_Times: "closure (A \ B) = closure A \ closure B" proof (rule closure_unique) show "A \ B \ closure A \ closure B" by (intro Sigma_mono closure_subset) show "closed (closure A \ closure B)" by (intro closed_Times closed_closure) fix T assume "A \ B \ T" and "closed T" then show "closure A \ closure B \ T" apply (simp add: closed_def open_prod_def, clarify) apply (rule ccontr) apply (drule_tac x="(a, b)" in bspec, simp, clarify, rename_tac C D) apply (simp add: closure_interior interior_def) apply (drule_tac x=C in spec) apply (drule_tac x=D in spec, auto) done qed 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 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 islimpt_in_closure: "(x islimpt S) = (x\closure(S-{x}))" unfolding closure_def using islimpt_punctured by blast lemma connected_imp_connected_closure: "connected S \ connected (closure S)" by (rule connectedI) (meson closure_subset open_Int open_Int_closure_eq_empty subset_trans connectedD) lemma bdd_below_closure: fixes A :: "real set" assumes "bdd_below A" shows "bdd_below (closure A)" proof - from assms obtain m where "\x. x \ A \ m \ x" by (auto simp: bdd_below_def) then have "A \ {m..}" by auto then have "closure A \ {m..}" using closed_real_atLeast by (rule closure_minimal) then show ?thesis by (auto simp: bdd_below_def) qed subsection \Frontier (also known as boundary)\ definition\<^marker>\tag important\ frontier :: "('a::topological_space) set \ 'a set" where "frontier S = closure S - interior S" lemma frontier_closed [iff]: "closed (frontier S)" by (simp add: frontier_def closed_Diff) lemma frontier_closures: "frontier S = closure S \ closure (- S)" by (auto simp: frontier_def interior_closure) lemma frontier_Int: "frontier(S \ T) = closure(S \ T) \ (frontier S \ frontier T)" proof - have "closure (S \ T) \ closure S" "closure (S \ T) \ closure T" by (simp_all add: closure_mono) then show ?thesis by (auto simp: frontier_closures) qed lemma frontier_Int_subset: "frontier(S \ T) \ frontier S \ frontier T" by (auto simp: frontier_Int) 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 lemma frontier_subset_closed: "closed S \ frontier S \ S" by (metis frontier_def closure_closed Diff_subset) lemma frontier_empty [simp]: "frontier {} = {}" 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 lemma frontier_complement [simp]: "frontier (- S) = frontier S" by (auto simp: frontier_def closure_complement interior_complement) lemma frontier_Un_subset: "frontier(S \ T) \ frontier S \ frontier T" by (metis compl_sup frontier_Int_subset frontier_complement) lemma frontier_disjoint_eq: "frontier S \ S = {} \ open S" using frontier_complement frontier_subset_eq[of "- S"] unfolding open_closed by auto lemma frontier_UNIV [simp]: "frontier UNIV = {}" using frontier_complement frontier_empty by fastforce lemma frontier_interiors: "frontier s = - interior(s) - interior(-s)" by (simp add: Int_commute frontier_def interior_closure) lemma frontier_interior_subset: "frontier(interior S) \ frontier S" 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 subsection\<^marker>\tag unimportant\ \Filters and the ``eventually true'' quantifier\ 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 lemma trivial_limit_at_iff: "trivial_limit (at a) \ \ a islimpt UNIV" using trivial_limit_within [of a UNIV] by simp lemma trivial_limit_at: "\ trivial_limit (at a)" for a :: "'a::perfect_space" by (rule at_neq_bot) lemma not_trivial_limit_within: "\ trivial_limit (at x within S) = (x \ closure (S - {x}))" 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] 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 at_within_eq_bot_iff: "at c within A = bot \ c \ closure (A - {c})" using not_trivial_limit_within[of c A] by blast text \Some property holds "sufficiently close" to the limit point.\ lemma trivial_limit_eventually: "trivial_limit net \ eventually P net" by simp lemma trivial_limit_eq: "trivial_limit net \ (\P. eventually P net)" by (simp add: filter_eq_iff) lemma Lim_topological: "(f \ l) net \ trivial_limit net \ (\S. open S \ l \ S \ eventually (\x. f x \ S) net)" unfolding tendsto_def trivial_limit_eq by auto lemma eventually_within_Un: "eventually P (at x within (s \ t)) \ eventually P (at x within s) \ eventually P (at x within t)" unfolding eventually_at_filter by (auto elim!: eventually_rev_mp) lemma Lim_within_union: "(f \ l) (at x within (s \ t)) \ (f \ l) (at x within s) \ (f \ l) (at x within t)" unfolding tendsto_def by (auto simp: eventually_within_Un) subsection \Limits\ text \The expected monotonicity property.\ lemma Lim_Un: assumes "(f \ l) (at x within S)" "(f \ l) (at x within T)" shows "(f \ l) (at x within (S \ T))" using assms unfolding at_within_union by (rule filterlim_sup) lemma Lim_Un_univ: "(f \ l) (at x within S) \ (f \ l) (at x within T) \ S \ T = UNIV \ (f \ l) (at x)" by (metis Lim_Un) text \Interrelations between restricted and unrestricted limits.\ lemma Lim_at_imp_Lim_at_within: "(f \ l) (at x) \ (f \ l) (at x within S)" by (metis order_refl filterlim_mono subset_UNIV at_le) 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 lemma at_within_interior: "x \ interior S \ at x within S = at x" unfolding filter_eq_iff by (intro allI eventually_within_interior) lemma Lim_within_LIMSEQ: fixes a :: "'a::first_countable_topology" assumes "\S. (\n. S n \ a \ S n \ T) \ S \ a \ (\n. X (S n)) \ L" shows "(X \ L) (at a within T)" using assms unfolding tendsto_def [where l=L] by (simp add: sequentially_imp_eventually_within) lemma Lim_right_bound: fixes f :: "'a :: {linorder_topology, conditionally_complete_linorder, no_top} \ 'b::{linorder_topology, conditionally_complete_linorder}" assumes mono: "\a b. a \ I \ b \ I \ x < a \ a \ b \ f a \ f b" and bnd: "\a. a \ I \ x < a \ K \ f a" shows "(f \ Inf (f ` ({x<..} \ I))) (at x within ({x<..} \ I))" proof (cases "{x<..} \ I = {}") case True then show ?thesis by simp next case False show ?thesis proof (rule order_tendstoI) fix a assume a: "a < Inf (f ` ({x<..} \ I))" { fix y assume "y \ {x<..} \ I" with False bnd have "Inf (f ` ({x<..} \ I)) \ f y" by (auto intro!: cInf_lower bdd_belowI2) with a have "a < f y" by (blast intro: less_le_trans) } then show "eventually (\x. a < f x) (at x within ({x<..} \ I))" by (auto simp: eventually_at_filter intro: exI[of _ 1] zero_less_one) next fix a assume "Inf (f ` ({x<..} \ I)) < a" from cInf_lessD[OF _ this] False obtain y where y: "x < y" "y \ I" "f y < a" by auto then have "eventually (\x. x \ I \ f x < a) (at_right x)" unfolding eventually_at_right[OF \x < y\] by (metis less_imp_le le_less_trans mono) then show "eventually (\x. f x < a) (at x within ({x<..} \ I))" unfolding eventually_at_filter by eventually_elim simp qed qed (*could prove directly from islimpt_sequential_inj, but only for metric spaces*) lemma islimpt_sequential: fixes x :: "'a::first_countable_topology" shows "x islimpt S \ (\f. (\n::nat. f n \ S - {x}) \ (f \ x) sequentially)" (is "?lhs = ?rhs") proof assume ?lhs from countable_basis_at_decseq[of x] obtain A where A: "\i. open (A i)" "\i. x \ A i" "\S. open S \ x \ S \ eventually (\i. A i \ S) sequentially" by blast define f where "f n = (SOME y. y \ S \ y \ A n \ x \ y)" for n { fix n from \?lhs\ have "\y. y \ S \ y \ A n \ x \ y" unfolding islimpt_def using A(1,2)[of n] by auto then have "f n \ S \ f n \ A n \ x \ f n" unfolding f_def by (rule someI_ex) then have "f n \ S" "f n \ A n" "x \ f n" by auto } then have "\n. f n \ S - {x}" by auto moreover have "(\n. f n) \ x" proof (rule topological_tendstoI) fix S assume "open S" "x \ S" from A(3)[OF this] \\n. f n \ A n\ show "eventually (\x. f x \ S) sequentially" by (auto elim!: eventually_mono) qed ultimately show ?rhs by fast next assume ?rhs then obtain f :: "nat \ 'a" where f: "\n. f n \ S - {x}" and lim: "f \ x" by auto show ?lhs unfolding islimpt_def proof safe fix T assume "open T" "x \ T" from lim[THEN topological_tendstoD, OF this] f show "\y\S. y \ T \ y \ x" unfolding eventually_sequentially by auto qed qed text\These are special for limits out of the same topological space.\ lemma Lim_within_id: "(id \ a) (at a within s)" unfolding id_def by (rule tendsto_ident_at) lemma Lim_at_id: "(id \ a) (at a)" unfolding id_def by (rule tendsto_ident_at) text\It's also sometimes useful to extract the limit point from the filter.\ abbreviation netlimit :: "'a::t2_space filter \ 'a" where "netlimit F \ Lim F (\x. x)" lemma netlimit_at [simp]: fixes a :: "'a::{perfect_space,t2_space}" shows "netlimit (at a) = a" using Lim_ident_at [of a UNIV] by simp lemma lim_within_interior: "x \ interior S \ (f \ l) (at x within S) \ (f \ l) (at x)" by (metis at_within_interior) lemma netlimit_within_interior: fixes x :: "'a::{t2_space,perfect_space}" assumes "x \ interior S" shows "netlimit (at x within S) = x" using assms by (metis at_within_interior netlimit_at) text\Useful lemmas on closure and set of possible sequential limits.\ 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 lemma closed_sequential_limits: fixes S :: "'a::first_countable_topology set" shows "closed S \ (\x l. (\n. x n \ S) \ (x \ l) sequentially \ l \ S)" 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)" proof - have *: "(s \ t) \ {x. x \ s} = s" "(s \ t) \ {x. x \ s} = t - s" by auto 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)" by (rule filterlim_at_within_closure_implies_filterlim) (use \x \ _\ in \auto intro!: tendsto_within_subset[OF g] simp: closure_def intro: islimpt_subset\) ultimately show ?thesis by (intro filterlim_at_within_If) (simp_all only: *) qed subsection \Compactness\ 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 = {}") case True show thesis by (rule that [of 1]) (auto simp: True) next case False 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" 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) 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" 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}"]] using f by auto from g(1,3) have g':"\x\g. \xa \ s. x = f xa" 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)" 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 ultimately show False using g(2) using finite_subset by auto qed lemma sequence_infinite_lemma: fixes f :: "nat \ 'a::t1_space" assumes "\n. f n \ l" and "(f \ l) sequentially" shows "infinite (range f)" proof assume "finite (range f)" then have "l \ range f \ closed (range f)" using \finite (range f)\ assms(1) finite_imp_closed by blast then have "eventually (\n. f n \ - range f) sequentially" by (metis Compl_iff assms(2) open_Compl topological_tendstoD) then show False unfolding eventually_sequentially by auto 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" proof - { fix x l assume as: "\n::nat. x n \ s" "(x \ l) sequentially" then have "l \ s" proof (cases "\n. x n \ l") case False then show "l\s" using as(1) by auto next case True note cas = this 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 qed } then show ?thesis unfolding closed_sequential_limits by fast qed 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.\ lemma compact_Un [intro]: 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'" unfolding compact_eq_Heine_Borel by (auto elim!: allE[of _ f]) moreover 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'" by (auto intro!: exI[of _ "s' \ t'"]) qed lemma compact_Union [intro]: "finite S \ (\T. T \ S \ compact T) \ compact (\S)" by (induct set: finite) auto lemma compact_UN [intro]: "finite A \ (\x. x \ A \ compact (B x)) \ compact (\x\A. B x)" by (rule compact_Union) auto lemma closed_Int_compact [intro]: 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)" 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 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})" by (simp add: open_Diff) text\Compactness expressed with filters\ lemma closure_iff_nhds_not_empty: "x \ closure X \ (\A. \S\A. open S \ x \ S \ X \ A \ {})" proof safe assume x: "x \ closure X" fix 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) 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 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) qed lemma compact_filter: "compact U \ (\F. F \ bot \ eventually (\x. x \ U) F \ (\x\U. inf (nhds x) F \ bot))" proof (intro allI iffI impI compact_fip[THEN iffD2] notI) fix F assume "compact U" assume F: "F \ bot" "eventually (\x. x \ U) F" then have "U \ {}" by (auto simp: eventually_False) define Z where "Z = closure ` {A. eventually (\x. x \ A) F}" then have "\z\Z. closed z" by auto moreover have ev_Z: "\z. z \ Z \ eventually (\x. x \ z) F" unfolding Z_def by (auto elim: eventually_mono intro: subsetD[OF closure_subset]) have "(\B \ Z. finite B \ U \ \B \ {})" proof (intro allI impI) fix B assume "finite B" "B \ Z" with \finite B\ ev_Z F(2) have "eventually (\x. x \ U \ (\B)) F" by (auto simp: eventually_ball_finite_distrib eventually_conj_iff) with F show "U \ \B \ {}" by (intro notI) (simp add: eventually_False) qed ultimately have "U \ \Z \ {}" using \compact U\ unfolding compact_fip by blast then obtain x where "x \ U" and x: "\z. z \ Z \ x \ z" by auto have "\P. eventually P (inf (nhds x) F) \ P \ bot" unfolding eventually_inf eventually_nhds 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}"] 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) qed with \x \ U\ show "\x\U. inf (nhds x) F \ bot" by (metis eventually_bot) next fix A assume A: "\a\A. closed a" "\B\A. finite B \ U \ \B \ {}" "U \ \A = {}" define F where "F = (INF a\insert U A. principal a)" have "F \ bot" unfolding F_def proof (rule INF_filter_not_bot) fix X assume X: "X \ insert U A" "finite X" with A(2)[THEN spec, of "X - {U}"] have "U \ \(X - {U}) \ {}" by auto with X show "(INF a\X. principal a) \ bot" by (auto simp: INF_principal_finite principal_eq_bot_iff) qed moreover have "F \ principal U" unfolding F_def by auto then have "eventually (\x. x \ U) F" by (auto simp: le_filter_def eventually_principal) moreover assume "\F. F \ bot \ eventually (\x. x \ U) F \ (\x\U. inf (nhds x) F \ bot)" ultimately obtain x where "x \ U" and x: "inf (nhds x) F \ bot" by auto { fix V assume "V \ A" then have "F \ principal V" unfolding F_def by (intro INF_lower2[of V]) auto then have V: "eventually (\x. x \ V) F" by (auto simp: le_filter_def eventually_principal) have "x \ closure V" unfolding closure_iff_nhds_not_empty proof (intro impI allI) fix S A assume "open S" "x \ S" "S \ A" then have "eventually (\x. x \ A) (nhds x)" by (auto simp: eventually_nhds) with V have "eventually (\x. x \ V \ A) (inf (nhds x) F)" by (auto simp: eventually_inf) with x show "V \ A \ {}" by (auto simp del: Int_iff simp add: trivial_limit_def) qed then have "x \ V" using \V \ A\ A(1) by simp } with \x\U\ have "x \ U \ \A" by auto with \U \ \A = {}\ show False by auto qed definition\<^marker>\tag important\ countably_compact :: "('a::topological_space) set \ bool" where "countably_compact U \ (\A. countable A \ (\a\A. open a) \ U \ \A \ (\T\A. finite T \ U \ \T))" lemma countably_compactE: assumes "countably_compact s" and "\t\C. open t" and "s \ \C" "countable C" obtains C' where "C' \ C" and "finite C'" and "s \ \C'" using assms unfolding countably_compact_def by metis lemma countably_compactI: assumes "\C. \t\C. open t \ s \ \C \ countable C \ (\C'\C. finite C' \ s \ \C')" shows "countably_compact s" using assms unfolding countably_compact_def by metis lemma compact_imp_countably_compact: "compact U \ countably_compact U" by (auto simp: compact_eq_Heine_Borel countably_compact_def) lemma countably_compact_imp_compact: assumes "countably_compact U" and ccover: "countable B" "\b\B. open b" and basis: "\T x. open T \ x \ T \ x \ U \ \b\B. x \ b \ b \ U \ T" shows "compact U" using \countably_compact U\ unfolding compact_eq_Heine_Borel countably_compact_def proof safe fix A assume A: "\a\A. open a" "U \ \A" assume *: "\A. countable A \ (\a\A. open a) \ U \ \A \ (\T\A. finite T \ U \ \T)" moreover define C where "C = {b\B. \a\A. b \ U \ a}" ultimately have "countable C" "\a\C. open a" unfolding C_def using ccover by auto moreover have "\A \ U \ \C" proof safe fix x a assume "x \ U" "x \ a" "a \ A" with basis[of a x] A obtain b where "b \ B" "x \ b" "b \ U \ a" by blast with \a \ A\ show "x \ \C" unfolding C_def by auto qed then have "U \ \C" using \U \ \A\ by auto ultimately obtain T where T: "T\C" "finite T" "U \ \T" using * by metis then have "\t\T. \a\A. t \ U \ a" by (auto simp: C_def) then obtain f where "\t\T. f t \ A \ t \ U \ f t" unfolding bchoice_iff Bex_def .. with T show "\T\A. finite T \ U \ \T" unfolding C_def by (intro exI[of _ "f`T"]) fastforce qed proposition countably_compact_imp_compact_second_countable: "countably_compact U \ compact (U :: 'a :: second_countable_topology set)" proof (rule countably_compact_imp_compact) fix T and x :: 'a assume "open T" "x \ T" from topological_basisE[OF is_basis this] obtain b where "b \ (SOME B. countable B \ topological_basis B)" "x \ b" "b \ T" . then show "\b\SOME B. countable B \ topological_basis B. x \ b \ b \ U \ T" by blast qed (insert countable_basis topological_basis_open[OF is_basis], auto) lemma countably_compact_eq_compact: "countably_compact U \ compact (U :: 'a :: second_countable_topology set)" using countably_compact_imp_compact_second_countable compact_imp_countably_compact by blast 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))" lemma seq_compactI: assumes "\f. \n. f n \ S \ \l\S. \r::nat\nat. strict_mono r \ ((f \ r) \ l) sequentially" shows "seq_compact S" unfolding seq_compact_def using assms by fast lemma seq_compactE: assumes "seq_compact S" "\n. f n \ S" obtains l r where "l \ S" "strict_mono (r :: nat \ nat)" "((f \ r) \ l) sequentially" 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 lemma seq_compact_Int_closed: 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" 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" by (rule seq_compactE) 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 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) lemma seq_compact_imp_countably_compact: fixes U :: "'a :: first_countable_topology set" assumes "seq_compact U" shows "countably_compact U" proof (safe intro!: countably_compactI) fix A assume A: "\a\A. open a" "U \ \A" "countable A" have subseq: "\X. range X \ U \ \r x. x \ U \ strict_mono (r :: nat \ nat) \ (X \ r) \ x" using \seq_compact U\ by (fastforce simp: seq_compact_def subset_eq) show "\T\A. finite T \ U \ \T" proof cases assume "finite A" with A show ?thesis by auto next assume "infinite A" then have "A \ {}" by auto show ?thesis proof (rule ccontr) assume "\ (\T\A. finite T \ U \ \T)" then have "\T. \x. T \ A \ finite T \ (x \ U - \T)" by auto then obtain X' where T: "\T. T \ A \ finite T \ X' T \ U - \T" by metis define X where "X n = X' (from_nat_into A ` {.. n})" for n have X: "\n. X n \ U - (\i\n. from_nat_into A i)" using \A \ {}\ unfolding X_def by (intro T) (auto intro: from_nat_into) then have "range X \ U" by auto with subseq[of X] obtain r x where "x \ U" and r: "strict_mono r" "(X \ r) \ x" 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] have "eventually (\i. X (r i) \ from_nat_into A n) sequentially" unfolding tendsto_def by (auto simp: comp_def) 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" by auto moreover from \strict_mono r\[THEN seq_suble, of "max n N"] have "\i. n \ r i \ N \ i" by (auto intro!: exI[of _ "max n N"]) ultimately show False by auto qed qed qed lemma compact_imp_seq_compact: fixes U :: "'a :: first_countable_topology set" assumes "compact U" shows "seq_compact U" unfolding seq_compact_def proof safe fix X :: "nat \ 'a" assume "\n. X n \ U" then have "eventually (\x. x \ U) (filtermap X sequentially)" by (auto simp: eventually_filtermap) moreover have "filtermap X sequentially \ bot" by (simp add: trivial_limit_def eventually_filtermap) ultimately obtain x where "x \ U" and x: "inf (nhds x) (filtermap X sequentially) \ bot" (is "?F \ _") using \compact U\ by (auto simp: compact_filter) from countable_basis_at_decseq[of x] obtain A where A: "\i. open (A i)" "\i. x \ A i" "\S. open S \ x \ S \ eventually (\i. A i \ S) sequentially" by blast define s where "s n i = (SOME j. i < j \ X j \ A (Suc n))" for n i { fix n i have "\a. i < a \ X a \ A (Suc n)" proof (rule ccontr) assume "\ (\a>i. X a \ A (Suc n))" then have "\a. Suc i \ a \ X a \ A (Suc n)" by auto then have "eventually (\x. x \ A (Suc n)) (filtermap X sequentially)" by (auto simp: eventually_filtermap eventually_sequentially) moreover have "eventually (\x. x \ A (Suc n)) (nhds x)" using A(1,2)[of "Suc n"] by (auto simp: eventually_nhds) ultimately have "eventually (\x. False) ?F" by (auto simp: eventually_inf) with x show False by (simp add: eventually_False) qed then have "i < s n i" "X (s n i) \ A (Suc n)" unfolding s_def by (auto intro: someI2_ex) } note s = this define r where "r = rec_nat (s 0 0) s" have "strict_mono r" by (auto simp: r_def s strict_mono_Suc_iff) moreover have "(\n. X (r n)) \ x" proof (rule topological_tendstoI) fix S assume "open S" "x \ S" with A(3) have "eventually (\i. A i \ S) sequentially" by auto moreover { fix i assume "Suc 0 \ i" then have "X (r i) \ A i" by (cases i) (simp_all add: r_def s) } then have "eventually (\i. X (r i) \ A i) sequentially" by (auto simp: eventually_sequentially) ultimately show "eventually (\i. X (r i) \ S) sequentially" by eventually_elim auto qed ultimately show "\x \ U. \r. strict_mono r \ (X \ r) \ x" using \x \ U\ by (auto simp: convergent_def comp_def) 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)" 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" 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\ unfolding C_def 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" unfolding C_def by (auto intro: countable_Collect_finite_subset) ultimately 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)))" by (metis (lifting) finite_subset_image C_def) 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\ 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 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 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 lemma seq_compact_eq_compact: fixes U :: "'a :: second_countable_topology set" shows "seq_compact U \ compact U" 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" 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)" 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 lemma compact_Times: 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)" 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" 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)" by metis 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)" 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)" 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)" 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'" by (intro exI[of _ "(\x\e. d x)"]) (auto simp: subset_eq) qed lemma tube_lemma: assumes "compact K" assumes "open W" assumes "{x0} \ K \ W" shows "\X0. x0 \ X0 \ open X0 \ X0 \ K \ W" proof - { fix y assume "y \ K" then have "(x0, y) \ W" using assms by auto with \open W\ have "\X0 Y. open X0 \ open Y \ x0 \ X0 \ y \ Y \ X0 \ Y \ W" by (rule open_prod_elim) blast } then obtain X0 Y where *: "\y \ K. open (X0 y) \ open (Y y) \ x0 \ X0 y \ y \ Y y \ X0 y \ Y y \ W" by metis from * have "\t\Y ` K. open t" "K \ \(Y ` K)" by auto with \compact K\ obtain CC where CC: "CC \ Y ` K" "finite CC" "K \ \CC" by (meson compactE) then obtain c where c: "\C. C \ CC \ c C \ K \ C = Y (c C)" by (force intro!: choice) with * CC show ?thesis by (force intro!: exI[where x="\C\CC. X0 (c C)"]) (* SLOW *) qed lemma continuous_on_prod_compactE: fixes fx::"'a::topological_space \ 'b::topological_space \ 'c::metric_space" and e::real assumes cont_fx: "continuous_on (U \ C) fx" assumes "compact C" assumes [intro]: "x0 \ U" notes [continuous_intros] = continuous_on_compose2[OF cont_fx] assumes "e > 0" obtains X0 where "x0 \ X0" "open X0" "\x\X0 \ U. \t \ C. dist (fx (x, t)) (fx (x0, t)) \ e" proof - define psi where "psi = (\(x, t). dist (fx (x, t)) (fx (x0, t)))" define W0 where "W0 = {(x, t) \ U \ C. psi (x, t) < e}" have W0_eq: "W0 = psi -` {.. U \ C" by (auto simp: vimage_def W0_def) 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 have "{x0} \ C \ W \ U \ C" unfolding W by (auto simp: W0_def psi_def \0 < e\) then have "{x0} \ C \ W" by blast from tube_lemma[OF \compact C\ \open W\ this] obtain X0 where X0: "x0 \ X0" "open X0" "X0 \ C \ W" by blast have "\x\X0 \ U. \t \ C. dist (fx (x, t)) (fx (x0, t)) \ e" proof safe fix x assume x: "x \ X0" "x \ U" 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 qed from X0(1,2) this show ?thesis .. qed subsection \Continuity\ lemma continuous_at_imp_continuous_within: "continuous (at x) f \ continuous (at x within s) f" unfolding continuous_within continuous_at using Lim_at_imp_Lim_at_within by auto lemma Lim_trivial_limit: "trivial_limit net \ (f \ l) net" by simp 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" unfolding continuous_within by(metis tendsto_within_subset) lemma continuous_on_interior: "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" unfolding continuous_on_def tendsto_def eventually_at_topological by simp text \Characterization of various kinds of continuity in terms of sequences.\ 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" 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 " 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) show ?thesis by (auto simp: assms(1) continuous_within[symmetric] tendsto_at_within_iff_tendsto_nhds[symmetric] intro!: filterlim_compose[OF _ *]) 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" "(x \ a) F " shows "((\n. f (x n)) \ f a) F" by (auto intro!: continuous_within_tendsto_compose[OF assms(1)] simp add: assms) 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 \ ((f \ x) \ f a) sequentially)" 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: fixes f :: "'a::{first_countable_topology, t2_space} \ 'b::topological_space" assumes "\u. u \ a \ (\n. f (u n)) \ f a" shows "continuous (at a) f" using continuous_within_sequentiallyI[of a UNIV f] assms by auto lemma continuous_at_sequentially: fixes f :: "'a::metric_space \ 'b::topological_space" shows "continuous (at a) f \ (\x. (x \ a) sequentially --> ((f \ x) \ f a) sequentially)" using continuous_within_sequentially[of a UNIV f] by simp 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" using assms unfolding continuous_on_eq_continuous_within 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 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 lemma continuous_imp_tendsto: assumes "continuous (at x0) f" and "x \ x0" shows "(f \ x) \ (f x0)" proof (rule topological_tendstoI) fix S assume "open S" "f x0 \ S" then obtain T where T_def: "open T" "x0 \ T" "\x\T. f x \ S" using assms continuous_at_open by metis then have "eventually (\n. x n \ T) sequentially" using assms T_def by (auto simp: tendsto_def) then show "eventually (\n. (f \ x) n \ S) sequentially" using T_def by (auto elim!: eventually_mono) qed 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" 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" 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) lemma homeomorphism_ident: "homeomorphism T T (\a. a) (\a. a)" by (rule homeomorphismI) auto lemma homeomorphism_compose: assumes "homeomorphism S T f g" "homeomorphism T U h k" shows "homeomorphism S U (h o f) (g o k)" using assms unfolding homeomorphism_def by (intro conjI ballI continuous_on_compose) (auto simp: image_iff) lemma homeomorphism_cong: "homeomorphism X' Y' f' g'" if "homeomorphism X Y f g" "X' = X" "Y' = Y" "\x. x \ X \ f' x = f x" "\y. y \ Y \ g' y = g y" using that by (auto simp add: homeomorphism_def) lemma homeomorphism_empty [simp]: "homeomorphism {} {} f g" unfolding homeomorphism_def by auto lemma homeomorphism_symD: "homeomorphism S t f g \ homeomorphism t S g f" by (simp add: homeomorphism_def) lemma homeomorphism_sym: "homeomorphism S t f g = homeomorphism t S g f" by (force simp: homeomorphism_def) definition\<^marker>\tag important\ homeomorphic :: "'a::topological_space set \ 'b::topological_space set \ bool" (infixr "homeomorphic" 60) where "s homeomorphic t \ (\f g. homeomorphism s t f g)" lemma homeomorphic_empty [iff]: "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_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) 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 lemma homeomorphicI [intro?]: "\f ` S = T; g ` T = S; continuous_on S f; continuous_on T g; \x. x \ S \ g(f(x)) = x; \y. y \ T \ f(g(y)) = y\ \ S homeomorphic T" unfolding homeomorphic_def homeomorphism_def by metis 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) lemma homeomorphism_apply1: "\homeomorphism S T f g; x \ S\ \ g(f x) = x" by (simp add: homeomorphism_def) lemma homeomorphism_apply2: "\homeomorphism S T f g; x \ T\ \ f(g x) = x" by (simp add: homeomorphism_def) lemma homeomorphism_image1: "homeomorphism S T f g \ f ` S = T" by (simp add: homeomorphism_def) lemma homeomorphism_image2: "homeomorphism S T f g \ g ` T = S" by (simp add: homeomorphism_def) lemma homeomorphism_cont1: "homeomorphism S T f g \ continuous_on S f" by (simp add: homeomorphism_def) lemma homeomorphism_cont2: "homeomorphism S T f g \ continuous_on T g" by (simp add: homeomorphism_def) lemma continuous_on_no_limpt: "(\x. \ x islimpt S) \ continuous_on S f" unfolding continuous_on_def by (metis UNIV_I empty_iff eventually_at_topological islimptE open_UNIV tendsto_def trivial_limit_within) lemma continuous_on_finite: fixes S :: "'a::t1_space set" shows "finite S \ continuous_on S f" by (metis continuous_on_no_limpt islimpt_finite) lemma homeomorphic_finite: fixes S :: "'a::t1_space set" and T :: "'b::t1_space set" assumes "finite T" shows "S homeomorphic T \ finite S \ finite T \ card S = card T" (is "?lhs = ?rhs") 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) 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 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" 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 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" unfolding homeomorphic_def by (metis homeomorphism_compact) text\Preservation of topological properties.\ lemma homeomorphic_compactness: "s homeomorphic t \ (compact s \ compact t)" unfolding homeomorphic_def homeomorphism_def by (metis compact_continuous_image) subsection\<^marker>\tag unimportant\ \On Linorder Topologies\ lemma islimpt_greaterThanLessThan1: fixes a b::"'a::{linorder_topology, dense_order}" assumes "a < b" shows "a islimpt {a<.. T" from open_right[OF this \a < b\] obtain c where c: "a < c" "{a.. T" by auto with assms dense[of a "min c b"] show "\y\{a<.. T \ y \ a" by (metis atLeastLessThan_iff greaterThanLessThan_iff min_less_iff_conj not_le order.strict_implies_order subset_eq) qed lemma islimpt_greaterThanLessThan2: fixes a b::"'a::{linorder_topology, dense_order}" assumes "a < b" shows "b islimpt {a<.. T" from open_left[OF this \a < b\] obtain c where c: "c < b" "{c<..b} \ T" by auto with assms dense[of "max a c" b] show "\y\{a<.. T \ y \ b" by (metis greaterThanAtMost_iff greaterThanLessThan_iff max_less_iff_conj not_le order.strict_implies_order subset_eq) qed lemma closure_greaterThanLessThan[simp]: fixes a b::"'a::{linorder_topology, dense_order}" shows "a < b \ closure {a <..< b} = {a .. b}" (is "_ \ ?l = ?r") proof have "?l \ closure ?r" by (rule closure_mono) auto thus "closure {a<.. {a..b}" by simp qed (auto simp: closure_def order.order_iff_strict islimpt_greaterThanLessThan1 islimpt_greaterThanLessThan2) lemma closure_greaterThan[simp]: fixes a b::"'a::{no_top, linorder_topology, dense_order}" shows "closure {a<..} = {a..}" proof - from gt_ex obtain b where "a < b" by auto hence "{a<..} = {a<.. {b..}" by auto also have "closure \ = {a..}" using \a < b\ unfolding closure_Un by auto finally show ?thesis . qed lemma closure_lessThan[simp]: fixes b::"'a::{no_bot, linorder_topology, dense_order}" shows "closure {.. {..a}" by auto also have "closure \ = {..b}" using \a < b\ unfolding closure_Un by auto finally show ?thesis . qed lemma closure_atLeastLessThan[simp]: fixes a b::"'a::{linorder_topology, dense_order}" assumes "a < b" shows "closure {a ..< b} = {a .. b}" proof - from assms have "{a ..< b} = {a} \ {a <..< b}" by auto also have "closure \ = {a .. b}" unfolding closure_Un by (auto simp: assms less_imp_le) finally show ?thesis . qed lemma closure_greaterThanAtMost[simp]: fixes a b::"'a::{linorder_topology, dense_order}" assumes "a < b" shows "closure {a <.. b} = {a .. b}" proof - from assms have "{a <.. b} = {b} \ {a <..< b}" by auto also have "closure \ = {a .. b}" unfolding closure_Un by (auto simp: assms less_imp_le) finally show ?thesis . qed end \ No newline at end of file diff --git a/src/HOL/Analysis/Polytope.thy b/src/HOL/Analysis/Polytope.thy --- a/src/HOL/Analysis/Polytope.thy +++ b/src/HOL/Analysis/Polytope.thy @@ -1,3913 +1,3912 @@ section \Faces, Extreme Points, Polytopes, Polyhedra etc\ text\Ported from HOL Light by L C Paulson\ theory Polytope imports Cartesian_Euclidean_Space Path_Connected begin subsection \Faces of a (usually convex) set\ definition\<^marker>\tag important\ face_of :: "['a::real_vector set, 'a set] \ bool" (infixr "(face'_of)" 50) where "T face_of S \ T \ S \ convex T \ (\a \ S. \b \ S. \x \ T. x \ open_segment a b \ a \ T \ b \ T)" lemma face_ofD: "\T face_of S; x \ open_segment a b; a \ S; b \ S; x \ T\ \ a \ T \ b \ T" unfolding face_of_def by blast lemma face_of_translation_eq [simp]: "((+) a ` T face_of (+) a ` S) \ T face_of S" proof - have *: "\a T S. T face_of S \ ((+) a ` T face_of (+) a ` S)" by (simp add: face_of_def) show ?thesis by (force simp: image_comp o_def dest: * [where a = "-a"] intro: *) qed lemma face_of_linear_image: assumes "linear f" "inj f" shows "(f ` c face_of f ` S) \ c face_of S" by (simp add: face_of_def inj_image_subset_iff inj_image_mem_iff open_segment_linear_image assms) lemma face_of_refl: "convex S \ S face_of S" by (auto simp: face_of_def) lemma face_of_refl_eq: "S face_of S \ convex S" by (auto simp: face_of_def) lemma empty_face_of [iff]: "{} face_of S" by (simp add: face_of_def) lemma face_of_empty [simp]: "S face_of {} \ S = {}" by (meson empty_face_of face_of_def subset_empty) lemma face_of_trans [trans]: "\S face_of T; T face_of u\ \ S face_of u" unfolding face_of_def by (safe; blast) lemma face_of_face: "T face_of S \ (f face_of T \ f face_of S \ f \ T)" unfolding face_of_def by (safe; blast) lemma face_of_subset: "\F face_of S; F \ T; T \ S\ \ F face_of T" unfolding face_of_def by (safe; blast) lemma face_of_slice: "\F face_of S; convex T\ \ (F \ T) face_of (S \ T)" unfolding face_of_def by (blast intro: convex_Int) lemma face_of_Int: "\t1 face_of S; t2 face_of S\ \ (t1 \ t2) face_of S" unfolding face_of_def by (blast intro: convex_Int) lemma face_of_Inter: "\A \ {}; \T. T \ A \ T face_of S\ \ (\ A) face_of S" unfolding face_of_def by (blast intro: convex_Inter) lemma face_of_Int_Int: "\F face_of T; F' face_of t'\ \ (F \ F') face_of (T \ t')" unfolding face_of_def by (blast intro: convex_Int) lemma face_of_imp_subset: "T face_of S \ T \ S" unfolding face_of_def by blast proposition face_of_imp_eq_affine_Int: fixes S :: "'a::euclidean_space set" assumes S: "convex S" and T: "T face_of S" shows "T = (affine hull T) \ S" proof - have "convex T" using T by (simp add: face_of_def) have *: False if x: "x \ affine hull T" and "x \ S" "x \ T" and y: "y \ rel_interior T" for x y proof - obtain e where "e>0" and e: "cball y e \ affine hull T \ T" using y by (auto simp: rel_interior_cball) have "y \ x" "y \ S" "y \ T" using face_of_imp_subset rel_interior_subset T that by blast+ then have zne: "\u. \u \ {0<..<1}; (1 - u) *\<^sub>R y + u *\<^sub>R x \ T\ \ False" using \x \ S\ \x \ T\ \T face_of S\ unfolding face_of_def by (meson greaterThanLessThan_iff in_segment(2)) have in01: "min (1/2) (e / norm (x - y)) \ {0<..<1}" using \y \ x\ \e > 0\ by simp have \
: "norm (min (1/2) (e / norm (x - y)) *\<^sub>R y - min (1/2) (e / norm (x - y)) *\<^sub>R x) \ e" using \e > 0\ by (simp add: scaleR_diff_right [symmetric] norm_minus_commute min_mult_distrib_right) show False apply (rule zne [OF in01 e [THEN subsetD]]) using \y \ T\ apply (simp add: hull_inc mem_affine x) by (simp add: dist_norm algebra_simps \
) qed show ?thesis proof (rule subset_antisym) show "T \ affine hull T \ S" using assms by (simp add: hull_subset face_of_imp_subset) show "affine hull T \ S \ T" using "*" \convex T\ rel_interior_eq_empty by fastforce qed qed lemma face_of_imp_closed: fixes S :: "'a::euclidean_space set" assumes "convex S" "closed S" "T face_of S" shows "closed T" by (metis affine_affine_hull affine_closed closed_Int face_of_imp_eq_affine_Int assms) lemma face_of_Int_supporting_hyperplane_le_strong: assumes "convex(S \ {x. a \ x = b})" and aleb: "\x. x \ S \ a \ x \ b" shows "(S \ {x. a \ x = b}) face_of S" proof - have *: "a \ u = a \ x" if "x \ open_segment u v" "u \ S" "v \ S" and b: "b = a \ x" for u v x proof (rule antisym) show "a \ u \ a \ x" using aleb \u \ S\ \b = a \ x\ by blast next obtain \ where "b = a \ ((1 - \) *\<^sub>R u + \ *\<^sub>R v)" "0 < \" "\ < 1" using \b = a \ x\ \x \ open_segment u v\ in_segment by (auto simp: open_segment_image_interval split: if_split_asm) then have "b + \ * (a \ u) \ a \ u + \ * b" using aleb [OF \v \ S\] by (simp add: algebra_simps) then have "(1 - \) * b \ (1 - \) * (a \ u)" by (simp add: algebra_simps) then have "b \ a \ u" using \\ < 1\ by auto with b show "a \ x \ a \ u" by simp qed show ?thesis using "*" open_segment_commute by (fastforce simp add: face_of_def assms) qed lemma face_of_Int_supporting_hyperplane_ge_strong: "\convex(S \ {x. a \ x = b}); \x. x \ S \ a \ x \ b\ \ (S \ {x. a \ x = b}) face_of S" using face_of_Int_supporting_hyperplane_le_strong [of S "-a" "-b"] by simp lemma face_of_Int_supporting_hyperplane_le: "\convex S; \x. x \ S \ a \ x \ b\ \ (S \ {x. a \ x = b}) face_of S" by (simp add: convex_Int convex_hyperplane face_of_Int_supporting_hyperplane_le_strong) lemma face_of_Int_supporting_hyperplane_ge: "\convex S; \x. x \ S \ a \ x \ b\ \ (S \ {x. a \ x = b}) face_of S" by (simp add: convex_Int convex_hyperplane face_of_Int_supporting_hyperplane_ge_strong) lemma face_of_imp_convex: "T face_of S \ convex T" using face_of_def by blast lemma face_of_imp_compact: fixes S :: "'a::euclidean_space set" shows "\convex S; compact S; T face_of S\ \ compact T" by (meson bounded_subset compact_eq_bounded_closed face_of_imp_closed face_of_imp_subset) lemma face_of_Int_subface: "\A \ B face_of A; A \ B face_of B; C face_of A; D face_of B\ \ (C \ D) face_of C \ (C \ D) face_of D" by (meson face_of_Int_Int face_of_face inf_le1 inf_le2) lemma subset_of_face_of: fixes S :: "'a::real_normed_vector set" assumes "T face_of S" "u \ S" "T \ (rel_interior u) \ {}" shows "u \ T" proof fix c assume "c \ u" obtain b where "b \ T" "b \ rel_interior u" using assms by auto then obtain e where "e>0" "b \ u" and e: "cball b e \ affine hull u \ u" by (auto simp: rel_interior_cball) show "c \ T" proof (cases "b=c") case True with \b \ T\ show ?thesis by blast next case False define d where "d = b + (e / norm(b - c)) *\<^sub>R (b - c)" have "d \ cball b e \ affine hull u" using \e > 0\ \b \ u\ \c \ u\ by (simp add: d_def dist_norm hull_inc mem_affine_3_minus False) with e have "d \ u" by blast have nbc: "norm (b - c) + e > 0" using \e > 0\ by (metis add.commute le_less_trans less_add_same_cancel2 norm_ge_zero) then have [simp]: "d \ c" using False scaleR_cancel_left [of "1 + (e / norm (b - c))" b c] by (simp add: algebra_simps d_def) (simp add: field_split_simps) have [simp]: "((e - e * e / (e + norm (b - c))) / norm (b - c)) = (e / (e + norm (b - c)))" using False nbc by (simp add: divide_simps) (simp add: algebra_simps) have "b \ open_segment d c" apply (simp add: open_segment_image_interval) apply (simp add: d_def algebra_simps image_def) apply (rule_tac x="e / (e + norm (b - c))" in bexI) using False nbc \0 < e\ by (auto simp: algebra_simps) then have "d \ T \ c \ T" by (meson \b \ T\ \c \ u\ \d \ u\ assms face_ofD subset_iff) then show ?thesis .. qed qed lemma face_of_eq: fixes S :: "'a::real_normed_vector set" assumes "T face_of S" "U face_of S" "(rel_interior T) \ (rel_interior U) \ {}" shows "T = U" using assms unfolding disjoint_iff_not_equal by (metis IntI empty_iff face_of_imp_subset mem_rel_interior_ball subset_antisym subset_of_face_of) lemma face_of_disjoint_rel_interior: fixes S :: "'a::real_normed_vector set" assumes "T face_of S" "T \ S" shows "T \ rel_interior S = {}" by (meson assms subset_of_face_of face_of_imp_subset order_refl subset_antisym) lemma face_of_disjoint_interior: fixes S :: "'a::real_normed_vector set" assumes "T face_of S" "T \ S" shows "T \ interior S = {}" proof - have "T \ interior S \ rel_interior S" by (meson inf_sup_ord(2) interior_subset_rel_interior order.trans) thus ?thesis by (metis (no_types) Int_greatest assms face_of_disjoint_rel_interior inf_sup_ord(1) subset_empty) qed lemma face_of_subset_rel_boundary: fixes S :: "'a::real_normed_vector set" assumes "T face_of S" "T \ S" shows "T \ (S - rel_interior S)" by (meson DiffI assms disjoint_iff_not_equal face_of_disjoint_rel_interior face_of_imp_subset rev_subsetD subsetI) lemma face_of_subset_rel_frontier: fixes S :: "'a::real_normed_vector set" assumes "T face_of S" "T \ S" shows "T \ rel_frontier S" using assms closure_subset face_of_disjoint_rel_interior face_of_imp_subset rel_frontier_def by fastforce lemma face_of_aff_dim_lt: fixes S :: "'a::euclidean_space set" assumes "convex S" "T face_of S" "T \ S" shows "aff_dim T < aff_dim S" proof - have "aff_dim T \ aff_dim S" by (simp add: face_of_imp_subset aff_dim_subset assms) moreover have "aff_dim T \ aff_dim S" proof (cases "T = {}") case True then show ?thesis by (metis aff_dim_empty \T \ S\) next case False then show ?thesis by (metis Set.set_insert assms convex_rel_frontier_aff_dim dual_order.irrefl face_of_imp_convex face_of_subset_rel_frontier insert_not_empty subsetI) qed ultimately show ?thesis by simp qed lemma subset_of_face_of_affine_hull: fixes S :: "'a::euclidean_space set" assumes T: "T face_of S" and "convex S" "U \ S" and dis: "\ disjnt (affine hull T) (rel_interior U)" shows "U \ T" proof (rule subset_of_face_of [OF T \U \ S\]) show "T \ rel_interior U \ {}" using face_of_imp_eq_affine_Int [OF \convex S\ T] rel_interior_subset [of U] dis \U \ S\ disjnt_def by fastforce qed lemma affine_hull_face_of_disjoint_rel_interior: fixes S :: "'a::euclidean_space set" assumes "convex S" "F face_of S" "F \ S" shows "affine hull F \ rel_interior S = {}" by (metis assms disjnt_def face_of_imp_subset order_refl subset_antisym subset_of_face_of_affine_hull) lemma affine_diff_divide: assumes "affine S" "k \ 0" "k \ 1" and xy: "x \ S" "y /\<^sub>R (1 - k) \ S" shows "(x - y) /\<^sub>R k \ S" proof - have "inverse(k) *\<^sub>R (x - y) = (1 - inverse k) *\<^sub>R inverse(1 - k) *\<^sub>R y + inverse(k) *\<^sub>R x" using assms by (simp add: algebra_simps) (simp add: scaleR_left_distrib [symmetric] field_split_simps) then show ?thesis using \affine S\ xy by (auto simp: affine_alt) qed proposition face_of_convex_hulls: assumes S: "finite S" "T \ S" and disj: "affine hull T \ convex hull (S - T) = {}" shows "(convex hull T) face_of (convex hull S)" proof - have fin: "finite T" "finite (S - T)" using assms by (auto simp: finite_subset) have *: "x \ convex hull T" if x: "x \ convex hull S" and y: "y \ convex hull S" and w: "w \ convex hull T" "w \ open_segment x y" for x y w proof - have waff: "w \ affine hull T" using convex_hull_subset_affine_hull w by blast obtain a b where a: "\i. i \ S \ 0 \ a i" and asum: "sum a S = 1" and aeqx: "(\i\S. a i *\<^sub>R i) = x" and b: "\i. i \ S \ 0 \ b i" and bsum: "sum b S = 1" and beqy: "(\i\S. b i *\<^sub>R i) = y" using x y by (auto simp: assms convex_hull_finite) obtain u where "(1 - u) *\<^sub>R x + u *\<^sub>R y \ convex hull T" "x \ y" and weq: "w = (1 - u) *\<^sub>R x + u *\<^sub>R y" and u01: "0 < u" "u < 1" using w by (auto simp: open_segment_image_interval split: if_split_asm) define c where "c i = (1 - u) * a i + u * b i" for i have cge0: "\i. i \ S \ 0 \ c i" using a b u01 by (simp add: c_def) have sumc1: "sum c S = 1" by (simp add: c_def sum.distrib sum_distrib_left [symmetric] asum bsum) have sumci_xy: "(\i\S. c i *\<^sub>R i) = (1 - u) *\<^sub>R x + u *\<^sub>R y" apply (simp add: c_def sum.distrib scaleR_left_distrib) by (simp only: scaleR_scaleR [symmetric] Real_Vector_Spaces.scaleR_right.sum [symmetric] aeqx beqy) show ?thesis proof (cases "sum c (S - T) = 0") case True have ci0: "\i. i \ (S - T) \ c i = 0" using True cge0 fin(2) sum_nonneg_eq_0_iff by auto have a0: "a i = 0" if "i \ (S - T)" for i using ci0 [OF that] u01 a [of i] b [of i] that by (simp add: c_def Groups.ordered_comm_monoid_add_class.add_nonneg_eq_0_iff) have [simp]: "sum a T = 1" using assms by (metis sum.mono_neutral_cong_right a0 asum) show ?thesis apply (simp add: convex_hull_finite \finite T\) apply (rule_tac x=a in exI) using a0 assms apply (auto simp: cge0 a aeqx [symmetric] sum.mono_neutral_right) done next case False define k where "k = sum c (S - T)" have "k > 0" using False unfolding k_def by (metis DiffD1 antisym_conv cge0 sum_nonneg not_less) have weq_sumsum: "w = sum (\x. c x *\<^sub>R x) T + sum (\x. c x *\<^sub>R x) (S - T)" by (metis (no_types) add.commute S(1) S(2) sum.subset_diff sumci_xy weq) show ?thesis proof (cases "k = 1") case True then have "sum c T = 0" by (simp add: S k_def sum_diff sumc1) then have [simp]: "sum c (S - T) = 1" by (simp add: S sum_diff sumc1) have ci0: "\i. i \ T \ c i = 0" by (meson \finite T\ \sum c T = 0\ \T \ S\ cge0 sum_nonneg_eq_0_iff subsetCE) then have [simp]: "(\i\S-T. c i *\<^sub>R i) = w" by (simp add: weq_sumsum) have "w \ convex hull (S - T)" apply (simp add: convex_hull_finite fin) apply (rule_tac x=c in exI) apply (auto simp: cge0 weq True k_def) done then show ?thesis using disj waff by blast next case False then have sumcf: "sum c T = 1 - k" by (simp add: S k_def sum_diff sumc1) have ge0: "\x. x \ T \ 0 \ inverse (1 - k) * c x" by (metis \T \ S\ cge0 inverse_nonnegative_iff_nonnegative mult_nonneg_nonneg subsetD sum_nonneg sumcf) have eq1: "(\x\T. inverse (1 - k) * c x) = 1" by (metis False eq_iff_diff_eq_0 mult.commute right_inverse sum_distrib_left sumcf) have "(\i\T. c i *\<^sub>R i) /\<^sub>R (1 - k) \ convex hull T" apply (simp add: convex_hull_finite fin) apply (rule_tac x="\i. inverse (1-k) * c i" in exI) by (metis (mono_tags, lifting) eq1 ge0 scaleR_scaleR scale_sum_right sum.cong) with \0 < k\ have "inverse(k) *\<^sub>R (w - sum (\i. c i *\<^sub>R i) T) \ affine hull T" by (simp add: affine_diff_divide [OF affine_affine_hull] False waff convex_hull_subset_affine_hull [THEN subsetD]) moreover have "inverse(k) *\<^sub>R (w - sum (\x. c x *\<^sub>R x) T) \ convex hull (S - T)" apply (simp add: weq_sumsum convex_hull_finite fin) apply (rule_tac x="\i. inverse k * c i" in exI) using \k > 0\ cge0 apply (auto simp: scaleR_right.sum sum_distrib_left [symmetric] k_def [symmetric]) done ultimately show ?thesis using disj by blast qed qed qed have [simp]: "convex hull T \ convex hull S" by (simp add: \T \ S\ hull_mono) show ?thesis using open_segment_commute by (auto simp: face_of_def intro: *) qed proposition face_of_convex_hull_insert: assumes "finite S" "a \ affine hull S" and T: "T face_of convex hull S" shows "T face_of convex hull insert a S" proof - have "convex hull S face_of convex hull insert a S" by (simp add: assms face_of_convex_hulls insert_Diff_if subset_insertI) then show ?thesis using T face_of_trans by blast qed proposition face_of_affine_trivial: assumes "affine S" "T face_of S" shows "T = {} \ T = S" proof (rule ccontr, clarsimp) assume "T \ {}" "T \ S" then obtain a where "a \ T" by auto then have "a \ S" using \T face_of S\ face_of_imp_subset by blast have "S \ T" proof fix b assume "b \ S" show "b \ T" proof (cases "a = b") case True with \a \ T\ show ?thesis by auto next case False then have "a \ 2 *\<^sub>R a - b" by (simp add: scaleR_2) with False have "a \ open_segment (2 *\<^sub>R a - b) b" apply (clarsimp simp: open_segment_def closed_segment_def) apply (rule_tac x="1/2" in exI) by (simp add: algebra_simps) moreover have "2 *\<^sub>R a - b \ S" by (rule mem_affine [OF \affine S\ \a \ S\ \b \ S\, of 2 "-1", simplified]) moreover note \b \ S\ \a \ T\ ultimately show ?thesis by (rule face_ofD [OF \T face_of S\, THEN conjunct2]) qed qed then show False using \T \ S\ \T face_of S\ face_of_imp_subset by blast qed lemma face_of_affine_eq: "affine S \ (T face_of S \ T = {} \ T = S)" using affine_imp_convex face_of_affine_trivial face_of_refl by auto proposition Inter_faces_finite_altbound: fixes T :: "'a::euclidean_space set set" assumes cfaI: "\c. c \ T \ c face_of S" shows "\F'. finite F' \ F' \ T \ card F' \ DIM('a) + 2 \ \F' = \T" proof (cases "\F'. finite F' \ F' \ T \ card F' \ DIM('a) + 2 \ (\c. c \ T \ c \ (\F') \ (\F'))") case True 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 using that proof (induction n) case 0 then show ?case by (simp add: c) next case (Suc n) then show ?case by (auto simp: c card_insert_if) qed have aff_dim_le: "aff_dim(\(d n)) \ DIM('a) - int n" if "n \ DIM('a) + 2" for n using that proof (induction n) case 0 then show ?case by (simp add: aff_dim_le_DIM) next case (Suc n) have fs: "\(d (Suc n)) face_of S" by (meson Suc.prems cfaI dn_le_Suc dn_notempty face_of_Inter subsetCE) have condn: "convex (\(d n))" using Suc.prems nat_le_linear not_less_eq_eq by (blast intro: face_of_imp_convex cfaI convex_Inter dest: dn_le_Suc) have fdn: "\(d (Suc n)) face_of \(d n)" by (metis (no_types, lifting) Inter_anti_mono Suc.prems dSuc cfaI dn_le_Suc dn_notempty face_of_Inter face_of_imp_subset face_of_subset subset_iff subset_insertI) have ne: "\(d (Suc n)) \ \(d n)" by (metis (no_types, lifting) Suc.prems Suc_leD c complete_lattice_class.Inf_insert dSuc dn_le_Suc less_irrefl order.trans) have *: "\m::int. \d. \d'::int. d < d' \ d' \ m - n \ d \ m - of_nat(n+1)" by arith have "aff_dim (\(d (Suc n))) < aff_dim (\(d n))" by (rule face_of_aff_dim_lt [OF condn fdn ne]) moreover have "aff_dim (\(d n)) \ int (DIM('a)) - int n" using Suc by auto ultimately have "aff_dim (\(d (Suc n))) \ int (DIM('a)) - (n+1)" by arith then show ?case by linarith qed have "aff_dim (\(d (DIM('a) + 2))) \ -2" using aff_dim_le [OF order_refl] by simp with aff_dim_geq [of "\(d (DIM('a) + 2))"] show ?thesis using order.trans by fastforce next case False then show ?thesis apply simp apply (erule ex_forward) by blast qed lemma faces_of_translation: "{F. F face_of image (\x. a + x) S} = image (image (\x. a + x)) {F. F face_of S}" proof - have "\F. F face_of (+) a ` S \ \G. G face_of S \ F = (+) a ` G" by (metis face_of_imp_subset face_of_translation_eq subset_imageE) then show ?thesis by (auto simp: image_iff) qed proposition face_of_Times: assumes "F face_of S" and "F' face_of S'" shows "(F \ F') face_of (S \ S')" proof - have "F \ F' \ S \ S'" using assms [unfolded face_of_def] by blast moreover have "convex (F \ F')" using assms [unfolded face_of_def] by (blast intro: convex_Times) moreover have "a \ F \ a' \ F' \ b \ F \ b' \ F'" if "a \ S" "b \ S" "a' \ S'" "b' \ S'" "x \ F \ F'" "x \ open_segment (a,a') (b,b')" for a b a' b' x proof (cases "b=a \ b'=a'") case True with that show ?thesis using assms by (force simp: in_segment dest: face_ofD) next case False with assms [unfolded face_of_def] that show ?thesis by (blast dest!: open_segment_PairD) qed ultimately show ?thesis unfolding face_of_def by blast qed corollary face_of_Times_decomp: fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set" shows "C face_of (S \ S') \ (\F F'. F face_of S \ F' face_of S' \ C = F \ F')" (is "?lhs = ?rhs") proof assume C: ?lhs show ?rhs proof (cases "C = {}") case True then show ?thesis by auto next case False have 1: "fst ` C \ S" "snd ` C \ S'" using C face_of_imp_subset by fastforce+ have "convex C" using C by (metis face_of_imp_convex) have conv: "convex (fst ` C)" "convex (snd ` C)" by (simp_all add: \convex C\ convex_linear_image linear_fst linear_snd) have fstab: "a \ fst ` C \ b \ fst ` C" if "a \ S" "b \ S" "x \ open_segment a b" "(x,x') \ C" for a b x x' proof - have *: "(x,x') \ open_segment (a,x') (b,x')" using that by (auto simp: in_segment) show ?thesis using face_ofD [OF C *] that face_of_imp_subset [OF C] by force qed have fst: "fst ` C face_of S" by (force simp: face_of_def 1 conv fstab) have sndab: "a' \ snd ` C \ b' \ snd ` C" if "a' \ S'" "b' \ S'" "x' \ open_segment a' b'" "(x,x') \ C" for a' b' x x' proof - have *: "(x,x') \ open_segment (x,a') (x,b')" using that by (auto simp: in_segment) show ?thesis using face_ofD [OF C *] that face_of_imp_subset [OF C] by force qed have snd: "snd ` C face_of S'" by (force simp: face_of_def 1 conv sndab) have cc: "rel_interior C \ rel_interior (fst ` C) \ rel_interior (snd ` C)" by (force simp: face_of_Times rel_interior_Times conv fst snd \convex C\ linear_fst linear_snd rel_interior_convex_linear_image [symmetric]) have "C = fst ` C \ snd ` C" proof (rule face_of_eq [OF C]) show "fst ` C \ snd ` C face_of S \ S'" by (simp add: face_of_Times rel_interior_Times conv fst snd) show "rel_interior C \ rel_interior (fst ` C \ snd ` C) \ {}" using False rel_interior_eq_empty \convex C\ cc by (auto simp: face_of_Times rel_interior_Times conv fst) qed with fst snd show ?thesis by metis qed next assume ?rhs with face_of_Times show ?lhs by auto qed lemma face_of_Times_eq: fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set" shows "(F \ F') face_of (S \ S') \ F = {} \ F' = {} \ F face_of S \ F' face_of S'" by (auto simp: face_of_Times_decomp times_eq_iff) lemma hyperplane_face_of_halfspace_le: "{x. a \ x = b} face_of {x. a \ x \ b}" proof - have "{x. a \ x \ b} \ {x. a \ x = b} = {x. a \ x = b}" by auto with face_of_Int_supporting_hyperplane_le [OF convex_halfspace_le [of a b], of a b] show ?thesis by auto qed lemma hyperplane_face_of_halfspace_ge: "{x. a \ x = b} face_of {x. a \ x \ b}" proof - have "{x. a \ x \ b} \ {x. a \ x = b} = {x. a \ x = b}" by auto with face_of_Int_supporting_hyperplane_ge [OF convex_halfspace_ge [of b a], of b a] show ?thesis by auto qed lemma face_of_halfspace_le: fixes a :: "'n::euclidean_space" shows "F face_of {x. a \ x \ b} \ F = {} \ F = {x. a \ x = b} \ F = {x. a \ x \ b}" (is "?lhs = ?rhs") proof (cases "a = 0") case True then show ?thesis using face_of_affine_eq affine_UNIV by auto next case False then have ine: "interior {x. a \ x \ b} \ {}" using halfspace_eq_empty_lt interior_halfspace_le by blast show ?thesis proof assume L: ?lhs have "F face_of {x. a \ x = b}" if "F \ {x. a \ x \ b}" proof - have "F face_of rel_frontier {x. a \ x \ b}" proof (rule face_of_subset [OF L]) show "F \ rel_frontier {x. a \ x \ b}" by (simp add: L face_of_subset_rel_frontier that) qed (force simp: rel_frontier_def closed_halfspace_le) then show ?thesis using False by (simp add: frontier_halfspace_le rel_frontier_nonempty_interior [OF ine]) qed with L show ?rhs using affine_hyperplane face_of_affine_eq by blast next assume ?rhs then show ?lhs by (metis convex_halfspace_le empty_face_of face_of_refl hyperplane_face_of_halfspace_le) qed qed lemma face_of_halfspace_ge: fixes a :: "'n::euclidean_space" shows "F face_of {x. a \ x \ b} \ F = {} \ F = {x. a \ x = b} \ F = {x. a \ x \ b}" using face_of_halfspace_le [of F "-a" "-b"] by simp subsection\Exposed faces\ text\That is, faces that are intersection with supporting hyperplane\ definition\<^marker>\tag important\ exposed_face_of :: "['a::euclidean_space set, 'a set] \ bool" (infixr "(exposed'_face'_of)" 50) where "T exposed_face_of S \ T face_of S \ (\a b. S \ {x. a \ x \ b} \ T = S \ {x. a \ x = b})" lemma empty_exposed_face_of [iff]: "{} exposed_face_of S" apply (simp add: exposed_face_of_def) apply (rule_tac x=0 in exI) apply (rule_tac x=1 in exI, force) done lemma exposed_face_of_refl_eq [simp]: "S exposed_face_of S \ convex S" proof assume S: "convex S" have "S \ {x. 0 \ x \ 0} \ S = S \ {x. 0 \ x = 0}" by auto with S show "S exposed_face_of S" using exposed_face_of_def face_of_refl_eq by blast qed (simp add: exposed_face_of_def face_of_refl_eq) lemma exposed_face_of_refl: "convex S \ S exposed_face_of S" by simp lemma exposed_face_of: "T exposed_face_of S \ T face_of S \ (T = {} \ T = S \ (\a b. a \ 0 \ S \ {x. a \ x \ b} \ T = S \ {x. a \ x = b}))" proof (cases "T = {}") case True then show ?thesis by simp next case False show ?thesis proof (cases "T = S") case True then show ?thesis by (simp add: face_of_refl_eq) next case False with \T \ {}\ show ?thesis apply (auto simp: exposed_face_of_def) apply (metis inner_zero_left) done qed qed lemma exposed_face_of_Int_supporting_hyperplane_le: "\convex S; \x. x \ S \ a \ x \ b\ \ (S \ {x. a \ x = b}) exposed_face_of S" by (force simp: exposed_face_of_def face_of_Int_supporting_hyperplane_le) lemma exposed_face_of_Int_supporting_hyperplane_ge: "\convex S; \x. x \ S \ a \ x \ b\ \ (S \ {x. a \ x = b}) exposed_face_of S" using exposed_face_of_Int_supporting_hyperplane_le [of S "-a" "-b"] by simp proposition exposed_face_of_Int: assumes "T exposed_face_of S" and "u exposed_face_of S" shows "(T \ u) exposed_face_of S" proof - obtain a b where T: "S \ {x. a \ x = b} face_of S" and S: "S \ {x. a \ x \ b}" and teq: "T = S \ {x. a \ x = b}" using assms by (auto simp: exposed_face_of_def) obtain a' b' where u: "S \ {x. a' \ x = b'} face_of S" and s': "S \ {x. a' \ x \ b'}" and ueq: "u = S \ {x. a' \ x = b'}" using assms by (auto simp: exposed_face_of_def) have tu: "T \ u face_of S" using T teq u ueq by (simp add: face_of_Int) have ss: "S \ {x. (a + a') \ x \ b + b'}" using S s' by (force simp: inner_left_distrib) show ?thesis apply (simp add: exposed_face_of_def tu) apply (rule_tac x="a+a'" in exI) apply (rule_tac x="b+b'" in exI) using S s' apply (fastforce simp: ss inner_left_distrib teq ueq) done qed proposition exposed_face_of_Inter: fixes P :: "'a::euclidean_space set set" assumes "P \ {}" and "\T. T \ P \ T exposed_face_of S" shows "\P exposed_face_of S" proof - obtain Q where "finite Q" and QsubP: "Q \ P" "card Q \ DIM('a) + 2" and IntQ: "\Q = \P" using Inter_faces_finite_altbound [of P S] assms [unfolded exposed_face_of] by force show ?thesis proof (cases "Q = {}") case True then show ?thesis by (metis IntQ Inter_UNIV_conv(2) assms(1) assms(2) ex_in_conv) next case False have "Q \ {T. T exposed_face_of S}" using QsubP assms by blast moreover have "Q \ {T. T exposed_face_of S} \ \Q exposed_face_of S" using \finite Q\ False by (induction Q rule: finite_induct; use exposed_face_of_Int in fastforce) ultimately show ?thesis by (simp add: IntQ) qed qed proposition exposed_face_of_sums: assumes "convex S" and "convex T" and "F exposed_face_of {x + y | x y. x \ S \ y \ T}" (is "F exposed_face_of ?ST") obtains k l where "k exposed_face_of S" "l exposed_face_of T" "F = {x + y | x y. x \ k \ y \ l}" proof (cases "F = {}") case True then show ?thesis using that by blast next case False show ?thesis proof (cases "F = ?ST") case True then show ?thesis using assms exposed_face_of_refl_eq that by blast next case False obtain p where "p \ F" using \F \ {}\ by blast moreover obtain u z where T: "?ST \ {x. u \ x = z} face_of ?ST" and S: "?ST \ {x. u \ x \ z}" and feq: "F = ?ST \ {x. u \ x = z}" using assms by (auto simp: exposed_face_of_def) ultimately obtain a0 b0 where p: "p = a0 + b0" and "a0 \ S" "b0 \ T" and z: "u \ p = z" by auto have lez: "u \ (x + y) \ z" if "x \ S" "y \ T" for x y using S that by auto have sef: "S \ {x. u \ x = u \ a0} exposed_face_of S" proof (rule exposed_face_of_Int_supporting_hyperplane_le [OF \convex S\]) show "\x. x \ S \ u \ x \ u \ a0" by (metis p z add_le_cancel_right inner_right_distrib lez [OF _ \b0 \ T\]) qed have tef: "T \ {x. u \ x = u \ b0} exposed_face_of T" proof (rule exposed_face_of_Int_supporting_hyperplane_le [OF \convex T\]) show "\x. x \ T \ u \ x \ u \ b0" by (metis p z add.commute add_le_cancel_right inner_right_distrib lez [OF \a0 \ S\]) qed have "{x + y |x y. x \ S \ u \ x = u \ a0 \ y \ T \ u \ y = u \ b0} \ F" by (auto simp: feq) (metis inner_right_distrib p z) moreover have "F \ {x + y |x y. x \ S \ u \ x = u \ a0 \ y \ T \ u \ y = u \ b0}" proof - have "\x y. \z = u \ (x + y); x \ S; y \ T\ \ u \ x = u \ a0 \ u \ y = u \ b0" using z p \a0 \ S\ \b0 \ T\ apply (simp add: inner_right_distrib) apply (metis add_le_cancel_right antisym lez [unfolded inner_right_distrib] add.commute) done then show ?thesis using feq by blast qed ultimately have "F = {x + y |x y. x \ S \ {x. u \ x = u \ a0} \ y \ T \ {x. u \ x = u \ b0}}" by blast then show ?thesis by (rule that [OF sef tef]) qed qed proposition exposed_face_of_parallel: "T exposed_face_of S \ T face_of S \ (\a b. S \ {x. a \ x \ b} \ T = S \ {x. a \ x = b} \ (T \ {} \ T \ S \ a \ 0) \ (T \ S \ (\w \ affine hull S. (w + a) \ affine hull S)))" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs proof (clarsimp simp: exposed_face_of_def) fix a b assume faceS: "S \ {x. a \ x = b} face_of S" and Ssub: "S \ {x. a \ x \ b}" show "\c d. S \ {x. c \ x \ d} \ S \ {x. a \ x = b} = S \ {x. c \ x = d} \ (S \ {x. a \ x = b} \ {} \ S \ {x. a \ x = b} \ S \ c \ 0) \ (S \ {x. a \ x = b} \ S \ (\w \ affine hull S. w + c \ affine hull S))" proof (cases "affine hull S \ {x. -a \ x \ -b} = {} \ affine hull S \ {x. - a \ x \ - b}") case True then show ?thesis proof assume "affine hull S \ {x. - a \ x \ - b} = {}" then show ?thesis apply (rule_tac x="0" in exI) apply (rule_tac x="1" in exI) using hull_subset by fastforce next assume "affine hull S \ {x. - a \ x \ - b}" then show ?thesis apply (rule_tac x="0" in exI) apply (rule_tac x="0" in exI) using Ssub hull_subset by fastforce qed next case False then obtain a' b' where "a' \ 0" and le: "affine hull S \ {x. a' \ x \ b'} = affine hull S \ {x. - a \ x \ - b}" and eq: "affine hull S \ {x. a' \ x = b'} = affine hull S \ {x. - a \ x = - b}" and mem: "\w. w \ affine hull S \ w + a' \ affine hull S" using affine_parallel_slice affine_affine_hull by metis show ?thesis proof (intro conjI impI allI ballI exI) have *: "S \ - (affine hull S \ {x. P x}) \ affine hull S \ {x. Q x} \ S \ {x. \ P x \ Q x}" for P Q using hull_subset by fastforce have "S \ {x. \ (a' \ x \ b') \ a' \ x = b'}" by (rule *) (use le eq Ssub in auto) then show "S \ {x. - a' \ x \ - b'}" by auto show "S \ {x. a \ x = b} = S \ {x. - a' \ x = - b'}" using eq hull_subset [of S affine] by force show "\S \ {x. a \ x = b} \ {}; S \ {x. a \ x = b} \ S\ \ - a' \ 0" using \a' \ 0\ by auto show "w + - a' \ affine hull S" if "S \ {x. a \ x = b} \ S" "w \ affine hull S" for w proof - have "w + 1 *\<^sub>R (w - (w + a')) \ affine hull S" using affine_affine_hull mem mem_affine_3_minus that(2) by blast then show ?thesis by simp qed qed qed qed next assume ?rhs then show ?lhs unfolding exposed_face_of_def by blast qed subsection\Extreme points of a set: its singleton faces\ definition\<^marker>\tag important\ extreme_point_of :: "['a::real_vector, 'a set] \ bool" (infixr "(extreme'_point'_of)" 50) where "x extreme_point_of S \ x \ S \ (\a \ S. \b \ S. x \ open_segment a b)" lemma extreme_point_of_stillconvex: "convex S \ (x extreme_point_of S \ x \ S \ convex(S - {x}))" by (fastforce simp add: convex_contains_segment extreme_point_of_def open_segment_def) lemma face_of_singleton: "{x} face_of S \ x extreme_point_of S" by (fastforce simp add: extreme_point_of_def face_of_def) lemma extreme_point_not_in_REL_INTERIOR: fixes S :: "'a::real_normed_vector set" shows "\x extreme_point_of S; S \ {x}\ \ x \ rel_interior S" by (metis disjoint_iff face_of_disjoint_rel_interior face_of_singleton insertI1) lemma extreme_point_not_in_interior: fixes S :: "'a::{real_normed_vector, perfect_space} set" assumes "x extreme_point_of S" shows "x \ interior S" proof (cases "S = {x}") case False then show ?thesis by (meson assms subsetD extreme_point_not_in_REL_INTERIOR interior_subset_rel_interior) qed (simp add: empty_interior_finite) lemma extreme_point_of_face: "F face_of S \ v extreme_point_of F \ v extreme_point_of S \ v \ F" by (meson empty_subsetI face_of_face face_of_singleton insert_subset) lemma extreme_point_of_convex_hull: "x extreme_point_of (convex hull S) \ x \ S" using hull_minimal [of S "(convex hull S) - {x}" convex] using hull_subset [of S convex] by (force simp add: extreme_point_of_stillconvex) proposition extreme_points_of_convex_hull: "{x. x extreme_point_of (convex hull S)} \ S" using extreme_point_of_convex_hull by auto lemma extreme_point_of_empty [simp]: "\ (x extreme_point_of {})" by (simp add: extreme_point_of_def) lemma extreme_point_of_singleton [iff]: "x extreme_point_of {a} \ x = a" using extreme_point_of_stillconvex by auto lemma extreme_point_of_translation_eq: "(a + x) extreme_point_of (image (\x. a + x) S) \ x extreme_point_of S" by (auto simp: extreme_point_of_def) lemma extreme_points_of_translation: "{x. x extreme_point_of (image (\x. a + x) S)} = (\x. a + x) ` {x. x extreme_point_of S}" using extreme_point_of_translation_eq by auto (metis (no_types, lifting) image_iff mem_Collect_eq minus_add_cancel) lemma extreme_points_of_translation_subtract: "{x. x extreme_point_of (image (\x. x - a) S)} = (\x. x - a) ` {x. x extreme_point_of S}" using extreme_points_of_translation [of "- a" S] by simp lemma extreme_point_of_Int: "\x extreme_point_of S; x extreme_point_of T\ \ x extreme_point_of (S \ T)" by (simp add: extreme_point_of_def) lemma extreme_point_of_Int_supporting_hyperplane_le: "\S \ {x. a \ x = b} = {c}; \x. x \ S \ a \ x \ b\ \ c extreme_point_of S" by (metis convex_singleton face_of_Int_supporting_hyperplane_le_strong face_of_singleton) lemma extreme_point_of_Int_supporting_hyperplane_ge: "\S \ {x. a \ x = b} = {c}; \x. x \ S \ a \ x \ b\ \ c extreme_point_of S" using extreme_point_of_Int_supporting_hyperplane_le [of S "-a" "-b" c] by simp lemma exposed_point_of_Int_supporting_hyperplane_le: "\S \ {x. a \ x = b} = {c}; \x. x \ S \ a \ x \ b\ \ {c} exposed_face_of S" unfolding exposed_face_of_def by (force simp: face_of_singleton extreme_point_of_Int_supporting_hyperplane_le) lemma exposed_point_of_Int_supporting_hyperplane_ge: "\S \ {x. a \ x = b} = {c}; \x. x \ S \ a \ x \ b\ \ {c} exposed_face_of S" using exposed_point_of_Int_supporting_hyperplane_le [of S "-a" "-b" c] by simp lemma extreme_point_of_convex_hull_insert: assumes "finite S" "a \ convex hull S" shows "a extreme_point_of (convex hull (insert a S))" proof (cases "a \ S") case False then show ?thesis using face_of_convex_hulls [of "insert a S" "{a}"] assms by (auto simp: face_of_singleton hull_same) qed (use assms in \simp add: hull_inc\) subsection\Facets\ definition\<^marker>\tag important\ facet_of :: "['a::euclidean_space set, 'a set] \ bool" (infixr "(facet'_of)" 50) where "F facet_of S \ F face_of S \ F \ {} \ aff_dim F = aff_dim S - 1" lemma facet_of_empty [simp]: "\ S facet_of {}" by (simp add: facet_of_def) lemma facet_of_irrefl [simp]: "\ S facet_of S " by (simp add: facet_of_def) lemma facet_of_imp_face_of: "F facet_of S \ F face_of S" by (simp add: facet_of_def) lemma facet_of_imp_subset: "F facet_of S \ F \ S" by (simp add: face_of_imp_subset facet_of_def) lemma hyperplane_facet_of_halfspace_le: "a \ 0 \ {x. a \ x = b} facet_of {x. a \ x \ b}" unfolding facet_of_def hyperplane_eq_empty by (auto simp: hyperplane_face_of_halfspace_ge hyperplane_face_of_halfspace_le Suc_leI of_nat_diff aff_dim_halfspace_le) lemma hyperplane_facet_of_halfspace_ge: "a \ 0 \ {x. a \ x = b} facet_of {x. a \ x \ b}" unfolding facet_of_def hyperplane_eq_empty by (auto simp: hyperplane_face_of_halfspace_le hyperplane_face_of_halfspace_ge Suc_leI of_nat_diff aff_dim_halfspace_ge) lemma facet_of_halfspace_le: "F facet_of {x. a \ x \ b} \ a \ 0 \ F = {x. a \ x = b}" (is "?lhs = ?rhs") proof assume c: ?lhs with c facet_of_irrefl show ?rhs by (force simp: aff_dim_halfspace_le facet_of_def face_of_halfspace_le cong: conj_cong split: if_split_asm) next assume ?rhs then show ?lhs by (simp add: hyperplane_facet_of_halfspace_le) qed lemma facet_of_halfspace_ge: "F facet_of {x. a \ x \ b} \ a \ 0 \ F = {x. a \ x = b}" using facet_of_halfspace_le [of F "-a" "-b"] by simp subsection \Edges: faces of affine dimension 1\ (*FIXME too small subsection, rearrange? *) definition\<^marker>\tag important\ edge_of :: "['a::euclidean_space set, 'a set] \ bool" (infixr "(edge'_of)" 50) where "e edge_of S \ e face_of S \ aff_dim e = 1" lemma edge_of_imp_subset: "S edge_of T \ S \ T" by (simp add: edge_of_def face_of_imp_subset) subsection\Existence of extreme points\ proposition different_norm_3_collinear_points: fixes a :: "'a::euclidean_space" assumes "x \ open_segment a b" "norm(a) = norm(b)" "norm(x) = norm(b)" shows False proof - obtain u where "norm ((1 - u) *\<^sub>R a + u *\<^sub>R b) = norm b" and "a \ b" and u01: "0 < u" "u < 1" using assms by (auto simp: open_segment_image_interval if_splits) then have "(1 - u) *\<^sub>R a \ (1 - u) *\<^sub>R a + ((1 - u) * 2) *\<^sub>R a \ u *\<^sub>R b = (1 - u * u) *\<^sub>R (a \ a)" using assms by (simp add: norm_eq algebra_simps inner_commute) then have "(1 - u) *\<^sub>R ((1 - u) *\<^sub>R a \ a + (2 * u) *\<^sub>R a \ b) = (1 - u) *\<^sub>R ((1 + u) *\<^sub>R (a \ a))" by (simp add: algebra_simps) then have "(1 - u) *\<^sub>R (a \ a) + (2 * u) *\<^sub>R (a \ b) = (1 + u) *\<^sub>R (a \ a)" using u01 by auto then have "a \ b = a \ a" using u01 by (simp add: algebra_simps) then have "a = b" using \norm(a) = norm(b)\ norm_eq vector_eq by fastforce then show ?thesis using \a \ b\ by force qed proposition extreme_point_exists_convex: fixes S :: "'a::euclidean_space set" assumes "compact S" "convex S" "S \ {}" obtains x where "x extreme_point_of S" proof - obtain x where "x \ S" and xsup: "\y. y \ S \ norm y \ norm x" using distance_attains_sup [of S 0] assms by auto have False if "a \ S" "b \ S" and x: "x \ open_segment a b" for a b proof - have noax: "norm a \ norm x" and nobx: "norm b \ norm x" using xsup that by auto have "a \ b" using empty_iff open_segment_idem x by auto show False by (metis dist_0_norm dist_decreases_open_segment noax nobx not_le x) qed then show ?thesis by (meson \x \ S\ extreme_point_of_def that) qed subsection\Krein-Milman, the weaker form\ proposition Krein_Milman: fixes S :: "'a::euclidean_space set" assumes "compact S" "convex S" shows "S = closure(convex hull {x. x extreme_point_of S})" proof (cases "S = {}") case True then show ?thesis by simp next case False have "closed S" by (simp add: \compact S\ compact_imp_closed) have "closure (convex hull {x. x extreme_point_of S}) \ S" by (simp add: \closed S\ assms closure_minimal extreme_point_of_def hull_minimal) moreover have "u \ closure (convex hull {x. x extreme_point_of S})" if "u \ S" for u proof (rule ccontr) assume unot: "u \ closure(convex hull {x. x extreme_point_of S})" then obtain a b where "a \ u < b" and ab: "\x. x \ closure(convex hull {x. x extreme_point_of S}) \ b < a \ x" using separating_hyperplane_closed_point [of "closure(convex hull {x. x extreme_point_of S})"] by blast have "continuous_on S ((\) a)" by (rule continuous_intros)+ then obtain m where "m \ S" and m: "\y. y \ S \ a \ m \ a \ y" using continuous_attains_inf [of S "\x. a \ x"] \compact S\ \u \ S\ by auto define T where "T = S \ {x. a \ x = a \ m}" have "m \ T" by (simp add: T_def \m \ S\) moreover have "compact T" by (simp add: T_def compact_Int_closed [OF \compact S\ closed_hyperplane]) moreover have "convex T" by (simp add: T_def convex_Int [OF \convex S\ convex_hyperplane]) ultimately obtain v where v: "v extreme_point_of T" using extreme_point_exists_convex [of T] by auto then have "{v} face_of T" by (simp add: face_of_singleton) also have "T face_of S" by (simp add: T_def m face_of_Int_supporting_hyperplane_ge [OF \convex S\]) finally have "v extreme_point_of S" by (simp add: face_of_singleton) then have "b < a \ v" using closure_subset by (simp add: closure_hull hull_inc ab) then show False using \a \ u < b\ \{v} face_of T\ face_of_imp_subset m T_def that by fastforce qed ultimately show ?thesis by blast qed text\Now the sharper form.\ lemma Krein_Milman_Minkowski_aux: fixes S :: "'a::euclidean_space set" assumes n: "dim S = n" and S: "compact S" "convex S" "0 \ S" shows "0 \ convex hull {x. x extreme_point_of S}" using n S proof (induction n arbitrary: S rule: less_induct) case (less n S) show ?case proof (cases "0 \ rel_interior S") case True with Krein_Milman less.prems show ?thesis by (metis subsetD convex_convex_hull convex_rel_interior_closure rel_interior_subset) next case False have "rel_interior S \ {}" by (simp add: rel_interior_convex_nonempty_aux less) then obtain c where c: "c \ rel_interior S" by blast obtain a where "a \ 0" and le_ay: "\y. y \ S \ a \ 0 \ a \ y" and less_ay: "\y. y \ rel_interior S \ a \ 0 < a \ y" by (blast intro: supporting_hyperplane_rel_boundary intro!: less False) have face: "S \ {x. a \ x = 0} face_of S" using face_of_Int_supporting_hyperplane_ge le_ay \convex S\ by auto then have co: "compact (S \ {x. a \ x = 0})" "convex (S \ {x. a \ x = 0})" using less.prems by (blast intro: face_of_imp_compact face_of_imp_convex)+ have "a \ y = 0" if "y \ span (S \ {x. a \ x = 0})" for y proof - have "y \ span {x. a \ x = 0}" by (metis inf.cobounded2 span_mono subsetCE that) then show ?thesis by (blast intro: span_induct [OF _ subspace_hyperplane]) qed then have "dim (S \ {x. a \ x = 0}) < n" by (metis (no_types) less_ay c subsetD dim_eq_span inf.strict_order_iff inf_le1 \dim S = n\ not_le rel_interior_subset span_0 span_base) then have "0 \ convex hull {x. x extreme_point_of (S \ {x. a \ x = 0})}" by (rule less.IH) (auto simp: co less.prems) then show ?thesis by (metis (mono_tags, lifting) Collect_mono_iff face extreme_point_of_face hull_mono subset_iff) qed qed theorem Krein_Milman_Minkowski: fixes S :: "'a::euclidean_space set" assumes "compact S" "convex S" shows "S = convex hull {x. x extreme_point_of S}" proof show "S \ convex hull {x. x extreme_point_of S}" proof fix a assume [simp]: "a \ S" have 1: "compact ((+) (- a) ` S)" by (simp add: \compact S\ compact_translation_subtract cong: image_cong_simp) have 2: "convex ((+) (- a) ` S)" by (simp add: \convex S\ compact_translation_subtract) show a_invex: "a \ convex hull {x. x extreme_point_of S}" using Krein_Milman_Minkowski_aux [OF refl 1 2] convex_hull_translation [of "-a"] by (auto simp: extreme_points_of_translation_subtract translation_assoc cong: image_cong_simp) qed next show "convex hull {x. x extreme_point_of S} \ S" proof - have "{a. a extreme_point_of S} \ S" using extreme_point_of_def by blast then show ?thesis by (simp add: \convex S\ hull_minimal) qed qed subsection\Applying it to convex hulls of explicitly indicated finite sets\ corollary Krein_Milman_polytope: fixes S :: "'a::euclidean_space set" shows "finite S \ convex hull S = convex hull {x. x extreme_point_of (convex hull S)}" by (simp add: Krein_Milman_Minkowski finite_imp_compact_convex_hull) lemma extreme_points_of_convex_hull_eq: fixes S :: "'a::euclidean_space set" shows "\compact S; \T. T \ S \ convex hull T \ convex hull S\ \ {x. x extreme_point_of (convex hull S)} = S" by (metis (full_types) Krein_Milman_Minkowski compact_convex_hull convex_convex_hull extreme_points_of_convex_hull psubsetI) lemma extreme_point_of_convex_hull_eq: fixes S :: "'a::euclidean_space set" shows "\compact S; \T. T \ S \ convex hull T \ convex hull S\ \ (x extreme_point_of (convex hull S) \ x \ S)" using extreme_points_of_convex_hull_eq by auto lemma extreme_point_of_convex_hull_convex_independent: fixes S :: "'a::euclidean_space set" assumes "compact S" and S: "\a. a \ S \ a \ convex hull (S - {a})" shows "(x extreme_point_of (convex hull S) \ x \ S)" proof - have "convex hull T \ convex hull S" if "T \ S" for T proof - obtain a where "T \ S" "a \ S" "a \ T" using \T \ S\ by blast then show ?thesis by (metis (full_types) Diff_eq_empty_iff Diff_insert0 S hull_mono hull_subset insert_Diff_single subsetCE) qed then show ?thesis by (rule extreme_point_of_convex_hull_eq [OF \compact S\]) qed lemma extreme_point_of_convex_hull_affine_independent: fixes S :: "'a::euclidean_space set" shows "\ affine_dependent S \ (x extreme_point_of (convex hull S) \ x \ S)" by (metis aff_independent_finite affine_dependent_def affine_hull_convex_hull extreme_point_of_convex_hull_convex_independent finite_imp_compact hull_inc) text\Elementary proofs exist, not requiring Euclidean spaces and all this development\ lemma extreme_point_of_convex_hull_2: fixes x :: "'a::euclidean_space" shows "x extreme_point_of (convex hull {a,b}) \ x = a \ x = b" proof - have "x extreme_point_of (convex hull {a,b}) \ x \ {a,b}" by (intro extreme_point_of_convex_hull_affine_independent affine_independent_2) then show ?thesis by simp qed lemma extreme_point_of_segment: fixes x :: "'a::euclidean_space" shows "x extreme_point_of closed_segment a b \ x = a \ x = b" by (simp add: extreme_point_of_convex_hull_2 segment_convex_hull) lemma face_of_convex_hull_subset: fixes S :: "'a::euclidean_space set" assumes "compact S" and T: "T face_of (convex hull S)" obtains s' where "s' \ S" "T = convex hull s'" proof show "{x. x extreme_point_of T} \ S" using T extreme_point_of_convex_hull extreme_point_of_face by blast show "T = convex hull {x. x extreme_point_of T}" proof (rule Krein_Milman_Minkowski) show "compact T" using T assms compact_convex_hull face_of_imp_compact by auto show "convex T" using T face_of_imp_convex by blast qed qed lemma face_of_convex_hull_aux: assumes eq: "x *\<^sub>R p = u *\<^sub>R a + v *\<^sub>R b + w *\<^sub>R c" and x: "u + v + w = x" "x \ 0" and S: "affine S" "a \ S" "b \ S" "c \ S" shows "p \ S" proof - have "p = (u *\<^sub>R a + v *\<^sub>R b + w *\<^sub>R c) /\<^sub>R x" by (metis \x \ 0\ eq mult.commute right_inverse scaleR_one scaleR_scaleR) moreover have "affine hull {a,b,c} \ S" by (simp add: S hull_minimal) moreover have "(u *\<^sub>R a + v *\<^sub>R b + w *\<^sub>R c) /\<^sub>R x \ affine hull {a,b,c}" apply (simp add: affine_hull_3) apply (rule_tac x="u/x" in exI) apply (rule_tac x="v/x" in exI) apply (rule_tac x="w/x" in exI) using x apply (auto simp: field_split_simps) done ultimately show ?thesis by force qed proposition face_of_convex_hull_insert_eq: fixes a :: "'a :: euclidean_space" assumes "finite S" and a: "a \ affine hull S" shows "(F face_of (convex hull (insert a S)) \ F face_of (convex hull S) \ (\F'. F' face_of (convex hull S) \ F = convex hull (insert a F')))" (is "F face_of ?CAS \ _") proof safe assume F: "F face_of ?CAS" and *: "\F'. F' face_of convex hull S \ F = convex hull insert a F'" obtain T where T: "T \ insert a S" and FeqT: "F = convex hull T" by (metis F \finite S\ compact_insert finite_imp_compact face_of_convex_hull_subset) show "F face_of convex hull S" proof (cases "a \ T") case True have "F = convex hull insert a (convex hull T \ convex hull S)" proof have "T \ insert a (convex hull T \ convex hull S)" using T hull_subset by fastforce then show "F \ convex hull insert a (convex hull T \ convex hull S)" by (simp add: FeqT hull_mono) show "convex hull insert a (convex hull T \ convex hull S) \ F" by (simp add: FeqT True hull_inc hull_minimal) qed moreover have "convex hull T \ convex hull S face_of convex hull S" by (metis F FeqT convex_convex_hull face_of_slice hull_mono inf.absorb_iff2 subset_insertI) ultimately show ?thesis using * by force next case False then show ?thesis by (metis FeqT F T face_of_subset hull_mono subset_insert subset_insertI) qed next assume "F face_of convex hull S" show "F face_of ?CAS" by (simp add: \F face_of convex hull S\ a face_of_convex_hull_insert \finite S\) next fix F assume F: "F face_of convex hull S" show "convex hull insert a F face_of ?CAS" proof (cases "S = {}") case True then show ?thesis using F face_of_affine_eq by auto next case False have anotc: "a \ convex hull S" by (metis (no_types) a affine_hull_convex_hull hull_inc) show ?thesis proof (cases "F = {}") case True show ?thesis using anotc by (simp add: \F = {}\ \finite S\ extreme_point_of_convex_hull_insert face_of_singleton) next case False have "convex hull insert a F \ ?CAS" by (simp add: F a \finite S\ convex_hull_subset face_of_convex_hull_insert face_of_imp_subset hull_inc) moreover have "(\y v. (1 - ub) *\<^sub>R a + ub *\<^sub>R b = (1 - v) *\<^sub>R a + v *\<^sub>R y \ 0 \ v \ v \ 1 \ y \ F) \ (\x u. (1 - uc) *\<^sub>R a + uc *\<^sub>R c = (1 - u) *\<^sub>R a + u *\<^sub>R x \ 0 \ u \ u \ 1 \ x \ F)" if *: "(1 - ux) *\<^sub>R a + ux *\<^sub>R x \ open_segment ((1 - ub) *\<^sub>R a + ub *\<^sub>R b) ((1 - uc) *\<^sub>R a + uc *\<^sub>R c)" and "0 \ ub" "ub \ 1" "0 \ uc" "uc \ 1" "0 \ ux" "ux \ 1" and b: "b \ convex hull S" and c: "c \ convex hull S" and "x \ F" for b c ub uc ux x proof - have xah: "x \ affine hull S" using F convex_hull_subset_affine_hull face_of_imp_subset \x \ F\ by blast have ah: "b \ affine hull S" "c \ affine hull S" using b c convex_hull_subset_affine_hull by blast+ obtain v where ne: "(1 - ub) *\<^sub>R a + ub *\<^sub>R b \ (1 - uc) *\<^sub>R a + uc *\<^sub>R c" and eq: "(1 - ux) *\<^sub>R a + ux *\<^sub>R x = (1 - v) *\<^sub>R ((1 - ub) *\<^sub>R a + ub *\<^sub>R b) + v *\<^sub>R ((1 - uc) *\<^sub>R a + uc *\<^sub>R c)" and "0 < v" "v < 1" using * by (auto simp: in_segment) then have 0: "((1 - ux) - ((1 - v) * (1 - ub) + v * (1 - uc))) *\<^sub>R a + (ux *\<^sub>R x - (((1 - v) * ub) *\<^sub>R b + (v * uc) *\<^sub>R c)) = 0" by (auto simp: algebra_simps) then have "((1 - ux) - ((1 - v) * (1 - ub) + v * (1 - uc))) *\<^sub>R a = ((1 - v) * ub) *\<^sub>R b + (v * uc) *\<^sub>R c + (-ux) *\<^sub>R x" by (auto simp: algebra_simps) then have "a \ affine hull S" if "1 - ux - ((1 - v) * (1 - ub) + v * (1 - uc)) \ 0" by (rule face_of_convex_hull_aux) (use b c xah ah that in \auto simp: algebra_simps\) then have "1 - ux - ((1 - v) * (1 - ub) + v * (1 - uc)) = 0" using a by blast with 0 have equx: "(1 - v) * ub + v * uc = ux" and uxx: "ux *\<^sub>R x = (((1 - v) * ub) *\<^sub>R b + (v * uc) *\<^sub>R c)" by auto (auto simp: algebra_simps) show ?thesis proof (cases "uc = 0") case True then show ?thesis using equx \0 \ ub\ \ub \ 1\ \v < 1\ uxx \x \ F\ by force next case False show ?thesis proof (cases "ub = 0") case True then show ?thesis using equx \0 \ uc\ \uc \ 1\ \0 < v\ uxx \x \ F\ by force next case False then have "0 < ub" "0 < uc" using \uc \ 0\ \0 \ ub\ \0 \ uc\ by auto then have "(1 - v) * ub > 0" "v * uc > 0" by (simp_all add: \0 < uc\ \0 < v\ \v < 1\) then have "ux \ 0" using equx \0 < v\ by auto have "b \ F \ c \ F" proof (cases "b = c") case True then show ?thesis by (metis \ux \ 0\ equx real_vector.scale_cancel_left scaleR_add_left uxx \x \ F\) next 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" 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" . then have "x \ open_segment b c" apply (simp add: in_segment \b \ c\) apply (rule_tac x="(v * uc) / ux" in exI) using \0 \ ux\ \ux \ 0\ \0 < uc\ \0 < v\ \0 < ub\ \v < 1\ equx apply (force simp: field_split_simps) done then show ?thesis by (rule face_ofD [OF F _ b c \x \ F\]) qed with \0 \ ub\ \ub \ 1\ \0 \ uc\ \uc \ 1\ show ?thesis by blast qed qed qed moreover have "convex hull F = F" by (meson F convex_hull_eq face_of_imp_convex) ultimately show ?thesis unfolding face_of_def by (fastforce simp: convex_hull_insert_alt \S \ {}\ \F \ {}\) qed qed qed lemma face_of_convex_hull_insert2: fixes a :: "'a :: euclidean_space" assumes S: "finite S" and a: "a \ affine hull S" and F: "F face_of convex hull S" shows "convex hull (insert a F) face_of convex hull (insert a S)" by (metis F face_of_convex_hull_insert_eq [OF S a]) proposition face_of_convex_hull_affine_independent: fixes S :: "'a::euclidean_space set" assumes "\ affine_dependent S" shows "(T face_of (convex hull S) \ (\c. c \ S \ T = convex hull c))" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (meson \T face_of convex hull S\ aff_independent_finite assms face_of_convex_hull_subset finite_imp_compact) next assume ?rhs then obtain c where "c \ S" and T: "T = convex hull c" by blast have "affine hull c \ affine hull (S - c) = {}" by (intro disjoint_affine_hull [OF assms \c \ S\], auto) then have "affine hull c \ convex hull (S - c) = {}" using convex_hull_subset_affine_hull by fastforce then show ?lhs by (metis face_of_convex_hulls \c \ S\ aff_independent_finite assms T) qed lemma facet_of_convex_hull_affine_independent: fixes S :: "'a::euclidean_space set" assumes "\ affine_dependent S" shows "T facet_of (convex hull S) \ T \ {} \ (\u. u \ S \ T = convex hull (S - {u}))" (is "?lhs = ?rhs") proof assume ?lhs then have "T face_of (convex hull S)" "T \ {}" and afft: "aff_dim T = aff_dim (convex hull S) - 1" by (auto simp: facet_of_def) then obtain c where "c \ S" and c: "T = convex hull c" by (auto simp: face_of_convex_hull_affine_independent [OF assms]) then have affs: "aff_dim S = aff_dim c + 1" by (metis aff_dim_convex_hull afft eq_diff_eq) have "\ affine_dependent c" using \c \ S\ affine_dependent_subset assms by blast with affs have "card (S - c) = 1" apply (simp add: aff_dim_affine_independent [symmetric] aff_dim_convex_hull) by (metis aff_dim_affine_independent aff_independent_finite One_nat_def \c \ S\ add.commute add_diff_cancel_right' assms card_Diff_subset card_mono of_nat_1 of_nat_diff of_nat_eq_iff) then obtain u where u: "u \ S - c" by (metis DiffI \c \ S\ aff_independent_finite assms cancel_comm_monoid_add_class.diff_cancel card_Diff_subset subsetI subset_antisym zero_neq_one) then have u: "S = insert u c" by (metis Diff_subset \c \ S\ \card (S - c) = 1\ card_1_singletonE double_diff insert_Diff insert_subset singletonD) have "T = convex hull (c - {u})" by (metis Diff_empty Diff_insert0 \T facet_of convex hull S\ c facet_of_irrefl insert_absorb u) with \T \ {}\ show ?rhs using c u by auto next assume ?rhs then obtain u where "T \ {}" "u \ S" and u: "T = convex hull (S - {u})" by (force simp: facet_of_def) then have "\ S \ {u}" using \T \ {}\ u by auto have "aff_dim (S - {u}) = aff_dim S - 1" using assms \u \ S\ unfolding affine_dependent_def by (metis add_diff_cancel_right' aff_dim_insert insert_Diff [of u S]) then have "aff_dim (convex hull (S - {u})) = aff_dim (convex hull S) - 1" by (simp add: aff_dim_convex_hull) then show ?lhs by (metis Diff_subset \T \ {}\ assms face_of_convex_hull_affine_independent facet_of_def u) qed lemma facet_of_convex_hull_affine_independent_alt: fixes S :: "'a::euclidean_space set" assumes "\ affine_dependent S" shows "(T facet_of (convex hull S) \ 2 \ card S \ (\u. u \ S \ T = convex hull (S - {u})))" (is "?lhs = ?rhs") proof assume L: ?lhs then obtain x where "x \ S" and x: "T = convex hull (S - {x})" and "finite S" using assms facet_of_convex_hull_affine_independent aff_independent_finite by blast moreover have "Suc (Suc 0) \ card S" using L x \x \ S\ \finite S\ by (metis Suc_leI assms card.remove convex_hull_eq_empty card_gt_0_iff facet_of_convex_hull_affine_independent finite_Diff not_less_eq_eq) ultimately show ?rhs by auto next assume ?rhs then show ?lhs using assms by (auto simp: facet_of_convex_hull_affine_independent Set.subset_singleton_iff) qed lemma segment_face_of: assumes "(closed_segment a b) face_of S" shows "a extreme_point_of S" "b extreme_point_of S" proof - have as: "{a} face_of S" by (metis (no_types) assms convex_hull_singleton empty_iff extreme_point_of_convex_hull_insert face_of_face face_of_singleton finite.emptyI finite.insertI insert_absorb insert_iff segment_convex_hull) moreover have "{b} face_of S" proof - have "b \ convex hull {a} \ b extreme_point_of convex hull {b, a}" by (meson extreme_point_of_convex_hull_insert finite.emptyI finite.insertI) moreover have "closed_segment a b = convex hull {b, a}" using closed_segment_commute segment_convex_hull by blast ultimately show ?thesis by (metis as assms face_of_face convex_hull_singleton empty_iff face_of_singleton insertE) qed ultimately show "a extreme_point_of S" "b extreme_point_of S" using face_of_singleton by blast+ qed proposition Krein_Milman_frontier: fixes S :: "'a::euclidean_space set" assumes "convex S" "compact S" shows "S = convex hull (frontier S)" (is "?lhs = ?rhs") proof have "?lhs \ convex hull {x. x extreme_point_of S}" using Krein_Milman_Minkowski assms by blast also have "... \ ?rhs" proof (rule hull_mono) show "{x. x extreme_point_of S} \ frontier S" using closure_subset by (auto simp: frontier_def extreme_point_not_in_interior extreme_point_of_def) qed finally show "?lhs \ ?rhs" . 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" by (simp add: \convex S\ hull_same) finally show "?rhs \ ?lhs" . qed subsection\Polytopes\ definition\<^marker>\tag important\ polytope where "polytope S \ \v. finite v \ S = convex hull v" lemma polytope_translation_eq: "polytope (image (\x. a + x) S) \ polytope S" proof - have "\a A. polytope A \ polytope ((+) a ` A)" by (metis (no_types) convex_hull_translation finite_imageI polytope_def) then show ?thesis by (metis (no_types) add.left_inverse image_add_0 translation_assoc) qed lemma polytope_linear_image: "\linear f; polytope p\ \ polytope(image f p)" unfolding polytope_def using convex_hull_linear_image by blast lemma polytope_empty: "polytope {}" using convex_hull_empty polytope_def by blast lemma polytope_convex_hull: "finite S \ polytope(convex hull S)" using polytope_def by auto lemma polytope_Times: "\polytope S; polytope T\ \ polytope(S \ T)" unfolding polytope_def by (metis finite_cartesian_product convex_hull_Times) lemma face_of_polytope_polytope: fixes S :: "'a::euclidean_space set" shows "\polytope S; F face_of S\ \ polytope F" unfolding polytope_def by (meson face_of_convex_hull_subset finite_imp_compact finite_subset) lemma finite_polytope_faces: fixes S :: "'a::euclidean_space set" assumes "polytope S" shows "finite {F. F face_of S}" proof - obtain v where "finite v" "S = convex hull v" using assms polytope_def by auto have "finite ((hull) convex ` {T. T \ v})" by (simp add: \finite v\) moreover have "{F. F face_of S} \ ((hull) convex ` {T. T \ v})" by (metis (no_types, lifting) \finite v\ \S = convex hull v\ face_of_convex_hull_subset finite_imp_compact image_eqI mem_Collect_eq subsetI) ultimately show ?thesis by (blast intro: finite_subset) qed lemma finite_polytope_facets: assumes "polytope S" shows "finite {T. T facet_of S}" by (simp add: assms facet_of_def finite_polytope_faces) lemma polytope_scaling: assumes "polytope S" shows "polytope (image (\x. c *\<^sub>R x) S)" by (simp add: assms polytope_linear_image) lemma polytope_imp_compact: fixes S :: "'a::real_normed_vector set" shows "polytope S \ compact S" by (metis finite_imp_compact_convex_hull polytope_def) lemma polytope_imp_convex: "polytope S \ convex S" by (metis convex_convex_hull polytope_def) lemma polytope_imp_closed: fixes S :: "'a::real_normed_vector set" shows "polytope S \ closed S" by (simp add: compact_imp_closed polytope_imp_compact) lemma polytope_imp_bounded: fixes S :: "'a::real_normed_vector set" shows "polytope S \ bounded S" by (simp add: compact_imp_bounded polytope_imp_compact) lemma polytope_interval: "polytope(cbox a b)" unfolding polytope_def by (meson closed_interval_as_convex_hull) lemma polytope_sing: "polytope {a}" using polytope_def by force lemma face_of_polytope_insert: "\polytope S; a \ affine hull S; F face_of S\ \ F face_of convex hull (insert a S)" by (metis (no_types, lifting) affine_hull_convex_hull face_of_convex_hull_insert hull_insert polytope_def) proposition face_of_polytope_insert2: fixes a :: "'a :: euclidean_space" assumes "polytope S" "a \ affine hull S" "F face_of S" shows "convex hull (insert a F) face_of convex hull (insert a S)" proof - obtain V where "finite V" "S = convex hull V" using assms by (auto simp: polytope_def) then have "convex hull (insert a F) face_of convex hull (insert a V)" using affine_hull_convex_hull assms face_of_convex_hull_insert2 by blast then show ?thesis by (metis \S = convex hull V\ hull_insert) qed subsection\Polyhedra\ definition\<^marker>\tag important\ polyhedron where "polyhedron S \ \F. finite F \ S = \ F \ (\h \ F. \a b. a \ 0 \ h = {x. a \ x \ b})" lemma polyhedron_Int [intro,simp]: "\polyhedron S; polyhedron T\ \ polyhedron (S \ T)" apply (clarsimp simp add: polyhedron_def) subgoal for F G by (rule_tac x="F \ G" in exI, auto) done lemma polyhedron_UNIV [iff]: "polyhedron UNIV" unfolding polyhedron_def by (rule_tac x="{}" in exI) auto lemma polyhedron_Inter [intro,simp]: "\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)" proof - define i::'a where "(i \ SOME i. i \ Basis)" 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 ultimately show ?thesis unfolding polyhedron_def by (rule_tac x="{{x. i \ x \ -1}, {x. -i \ x \ -1}}" in exI) force qed lemma polyhedron_halfspace_le: fixes a :: "'a :: euclidean_space" shows "polyhedron {x. a \ x \ b}" proof (cases "a = 0") case True then show ?thesis by auto next case False then show ?thesis unfolding polyhedron_def by (rule_tac x="{{x. a \ x \ b}}" in exI) auto qed lemma polyhedron_halfspace_ge: fixes a :: "'a :: euclidean_space" shows "polyhedron {x. a \ x \ b}" using polyhedron_halfspace_le [of "-a" "-b"] by simp lemma polyhedron_hyperplane: fixes a :: "'a :: euclidean_space" shows "polyhedron {x. a \ x = b}" proof - have "{x. a \ x = b} = {x. a \ x \ b} \ {x. a \ x \ b}" by force then show ?thesis by (simp add: polyhedron_halfspace_ge polyhedron_halfspace_le) qed 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]) lemma polyhedron_imp_closed: fixes S :: "'a :: euclidean_space set" shows "polyhedron S \ closed S" by (metis closed_Inter closed_halfspace_le polyhedron_def) lemma polyhedron_imp_convex: fixes S :: "'a :: euclidean_space set" shows "polyhedron S \ convex S" by (metis convex_Inter convex_halfspace_le polyhedron_def) lemma polyhedron_affine_hull: fixes S :: "'a :: euclidean_space set" shows "polyhedron(affine hull S)" by (simp add: affine_imp_polyhedron) subsection\Canonical polyhedron representation making facial structure explicit\ proposition polyhedron_Int_affine: fixes S :: "'a :: euclidean_space set" 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 proposition rel_interior_polyhedron_explicit: assumes "finite F" and seq: "S = affine hull S \ \F" and faceq: "\h. h \ F \ a h \ 0 \ h = {x. a h \ x \ b h}" and psub: "\F'. F' \ F \ S \ affine hull S \ \F'" shows "rel_interior S = {x \ S. \h \ F. a h \ x < b h}" proof - have rels: "\x. x \ rel_interior S \ x \ S" by (meson IntE mem_rel_interior) moreover have "a i \ x < b i" if x: "x \ rel_interior S" and "i \ F" for x i proof - have fif: "F - {i} \ F" using \i \ F\ Diff_insert_absorb Diff_subset set_insert psubsetI by blast then have "S \ affine hull S \ \(F - {i})" by (rule psub) then obtain z where ssub: "S \ \(F - {i})" and zint: "z \ \(F - {i})" and "z \ S" and zaff: "z \ affine hull S" by auto have "z \ x" using \z \ S\ rels x by blast have "z \ affine hull S \ \F" using \z \ S\ seq by auto then have aiz: "a i \ z > b i" using faceq zint zaff by fastforce obtain e where "e > 0" "x \ S" and e: "ball x e \ affine hull S \ S" using x by (auto simp: mem_rel_interior_ball) then have ins: "\y. \norm (x - y) < e; y \ affine hull S\ \ y \ S" by (metis IntI subsetD dist_norm mem_ball) 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)" using \e > 0\ by (simp add: \_def) 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) 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" using \e > 0\ \z \ x\ by (rule_tac l = \ in that) (auto simp: \_def) then have i: "l *\<^sub>R z + (1 - l) *\<^sub>R x \ i" using seq \i \ F\ by auto have "b i * l + (a i \ x) * (1 - l) < a i \ (l *\<^sub>R z + (1 - l) *\<^sub>R x)" using l by (simp add: algebra_simps aiz) also have "\ \ b i" using i l using faceq mem_Collect_eq \i \ F\ by blast finally have "(a i \ x) * (1 - l) < b i * (1 - l)" by (simp add: algebra_simps) with l show ?thesis by simp qed moreover have "x \ rel_interior S" if "x \ S" and less: "\h. h \ F \ a h \ x < b h" for x proof - have 1: "\h. h \ F \ x \ interior h" by (metis interior_halfspace_le mem_Collect_eq less faceq) have 2: "\y. \\h\F. y \ interior h; y \ affine hull S\ \ y \ S" by (metis IntI Inter_iff subsetD interior_subset seq) show ?thesis apply (simp add: rel_interior \x \ S\) apply (rule_tac x="\h\F. interior h" in exI) apply (auto simp: \finite F\ open_INT 1 2) done qed ultimately show ?thesis by blast qed lemma polyhedron_Int_affine_parallel: fixes S :: "'a :: euclidean_space set" shows "polyhedron S \ (\F. finite F \ S = (affine hull S) \ (\F) \ (\h \ F. \a b. a \ 0 \ h = {x. a \ x \ b} \ (\x \ affine hull S. (x + a) \ affine hull S)))" (is "?lhs = ?rhs") proof assume ?lhs then obtain F where "finite F" and seq: "S = (affine hull S) \ \F" and faces: "\h. h \ F \ \a b. a \ 0 \ h = {x. a \ x \ b}" by (fastforce simp add: polyhedron_Int_affine) then obtain a b where ab: "\h. h \ F \ a h \ 0 \ h = {x. a h \ x \ b h}" by metis show ?rhs proof - have "\a' b'. a' \ 0 \ affine hull S \ {x. a' \ x \ b'} = affine hull S \ h \ (\w \ affine hull S. (w + a') \ affine hull S)" if "h \ F" "\(affine hull S \ h)" for h proof - 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)) 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 using affine_parallel_slice [of "affine hull S"] by (metis \h = {x. a h \ x \ b h}\ affine_affine_hull) qed then obtain a b where ab: "\h. \h \ F; \ (affine hull S \ h)\ \ a h \ 0 \ 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) 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 qed next assume ?rhs then show ?lhs by (metis polyhedron_Int_affine) qed proposition polyhedron_Int_affine_parallel_minimal: fixes S :: "'a :: euclidean_space set" shows "polyhedron S \ (\F. finite F \ S = (affine hull S) \ (\F) \ (\h \ F. \a b. a \ 0 \ h = {x. a \ x \ b} \ (\x \ affine hull S. (x + a) \ affine hull S)) \ (\F'. F' \ F \ S \ (affine hull S) \ (\F')))" (is "?lhs = ?rhs") proof assume ?lhs then obtain f0 where f0: "finite f0" "S = (affine hull S) \ (\f0)" (is "?P f0") "\h \ f0. \a b. a \ 0 \ h = {x. a \ x \ b} \ (\x \ affine hull S. (x + a) \ affine hull S)" (is "?Q f0") by (force simp: polyhedron_Int_affine_parallel) define n where "n = (LEAST n. \F. card F = n \ finite F \ ?P F \ ?Q F)" have nf: "\F. card F = n \ finite F \ ?P F \ ?Q F" apply (simp add: n_def) apply (rule LeastI [where k = "card f0"]) using f0 apply auto done then obtain F where F: "card F = n" "finite F" and seq: "?P F" and aff: "?Q F" by blast then have "\ (finite g \ ?P g \ ?Q g)" if "card g < n" for g using that by (auto simp: n_def dest!: not_less_Least) then have *: "\ (?P g \ ?Q g)" if "g \ F" for g using that \finite F\ psubset_card_mono \card F = n\ by (metis finite_Int inf.strict_order_iff) have 1: "\F'. F' \ F \ S \ affine hull S \ \F'" by (subst seq) blast have 2: "S \ affine hull S \ \F'" if "F' \ F" for F' using * [OF that] by (metis IntE aff inf.strict_order_iff that) show ?rhs by (metis \finite F\ seq aff psubsetI 1 2) next assume ?rhs then show ?lhs by (auto simp: polyhedron_Int_affine_parallel) qed lemma polyhedron_Int_affine_minimal: fixes S :: "'a :: euclidean_space set" shows "polyhedron S \ (\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) proposition facet_of_polyhedron_explicit: assumes "finite F" and seq: "S = affine hull S \ \F" and faceq: "\h. h \ F \ a h \ 0 \ h = {x. a h \ x \ b h}" and psub: "\F'. F' \ F \ S \ affine hull S \ \F'" shows "C facet_of S \ (\h. h \ F \ C = S \ {x. a h \ x = b h})" proof (cases "S = {}") case True with psub show ?thesis by force next case False have "polyhedron S" unfolding polyhedron_Int_affine by (metis \finite F\ faceq seq) then have "convex S" by (rule polyhedron_imp_convex) with False rel_interior_eq_empty have "rel_interior S \ {}" by blast then obtain x where "x \ rel_interior S" by auto then obtain T where "open T" "x \ T" "x \ S" "T \ affine hull S \ S" by (force simp: mem_rel_interior) then have xaff: "x \ affine hull S" and xint: "x \ \F" using seq hull_inc by auto have "rel_interior S = {x \ S. \h\F. a h \ x < b h}" by (rule rel_interior_polyhedron_explicit [OF \finite F\ seq faceq psub]) with \x \ rel_interior S\ have [simp]: "\h. h\F \ a h \ x < b h" by blast have *: "(S \ {x. a h \ x = b h}) facet_of S" if "h \ F" for h proof - have "S \ affine hull S \ \(F - {h})" using psub that by (metis Diff_disjoint Diff_subset insert_disjoint(2) psubsetI) then obtain z where zaff: "z \ affine hull S" and zint: "z \ \(F - {h})" and "z \ S" by force then have "z \ x" "z \ h" using seq \x \ S\ by auto 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 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" have "0 < l" "l < 1" using able xltz \b h < a h \ z\ \h \ F\ by (auto simp: l_def field_split_simps) have awlt: "a i \ w < b i" if "i \ F" "i \ h" for i proof - have "(1 - l) * (a i \ x) < (1 - l) * b i" by (simp add: \l < 1\ \i \ F\) 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) 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" proof (rule face_of_Int_supporting_hyperplane_le) show "\x. x \ S \ a h \ x \ b h" using faceq seq that by fastforce qed fact have "w \ affine hull S" by (simp add: w_def mem_affine xaff zaff) moreover have "w \ \F" 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}" 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 next show "affine hull S \ {x. a h \ x = b h} \ affine hull (S \ {x. a h \ x = b h})" proof fix y assume yaff: "y \ affine hull S \ {y. a h \ y = b h}" obtain T where "0 < T" and T: "\j. \j \ F; j \ h\ \ T * (a j \ y - a j \ w) \ b j - a j \ w" proof (cases "F - {h} = {}") case True then show ?thesis by (rule_tac T=1 in that) auto next case False then obtain h' where h': "h' \ F - {h}" by auto let ?body = "(\j. if 0 < a j \ y - a j \ w then (b j - a j \ w) / (a j \ y - a j \ w) else 1) ` (F - {h})" define inff where "inff = Inf ?body" from \finite F\ have "finite ?body" by blast moreover from h' have "?body \ {}" by blast moreover have "j > 0" if "j \ ?body" for j proof - from that obtain x where "x \ F" and "x \ h" and *: "j = (if 0 < a x \ y - a x \ w then (b x - a x \ w) / (a x \ y - a x \ w) else 1)" by blast with awlt [of x] have "a x \ w < b x" by simp with * show ?thesis by simp qed ultimately have "0 < inff" by (simp_all add: finite_less_Inf_iff inff_def) moreover have "inff * (a j \ y - a j \ w) \ b j - a j \ w" if "j \ F" "j \ h" for j proof (cases "a j \ w < a j \ y") case True then have "inff \ (b j - a j \ w) / (a j \ y - a j \ w)" unfolding inff_def using \finite F\ by (auto intro: cInf_le_finite simp add: that split: if_split_asm) then show ?thesis using \0 < inff\ awlt [OF that] mult_strict_left_mono by (fastforce simp add: field_split_simps split: if_split_asm) next 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" by (simp add: awlt that) finally show ?thesis by simp qed ultimately show ?thesis by (blast intro: that) qed define C where "C = (1 - T) *\<^sub>R w + T *\<^sub>R y" have "(1 - T) *\<^sub>R w + T *\<^sub>R y \ j" if "j \ F" for j proof (cases "j = h") case True have "(1 - T) *\<^sub>R w + T *\<^sub>R y \ {x. a h \ x \ b h}" using weq yaff by (auto simp: algebra_simps) with True faceq [OF that] show ?thesis by metis next case False with T that have "(1 - T) *\<^sub>R w + T *\<^sub>R y \ {x. a j \ x \ b j}" by (simp add: algebra_simps) with faceq [OF that] show ?thesis by simp qed moreover have "(1 - T) *\<^sub>R w + T *\<^sub>R y \ affine hull S" using yaff \w \ affine hull S\ affine_affine_hull affine_alt by blast ultimately have "C \ S" using seq by (force simp: C_def) moreover have "a h \ C = b h" using yaff by (force simp: C_def algebra_simps weq) ultimately have caff: "C \ affine hull (S \ {y. a h \ y = b h})" by (simp add: hull_inc) have waff: "w \ affine hull (S \ {y. a h \ y = b h})" using \w \ S\ weq by (blast intro: hull_inc) have yeq: "y = (1 - inverse T) *\<^sub>R w + C /\<^sub>R T" using \0 < T\ by (simp add: C_def algebra_simps) show "y \ affine hull (S \ {y. a h \ y = b h})" 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" 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) qed show ?thesis proof show "\h. h \ F \ C = S \ {x. a h \ x = b h} \ C facet_of S" using * by blast next assume "C facet_of S" then have "C face_of S" "convex C" "C \ {}" and affc: "aff_dim C = aff_dim S - 1" by (auto simp: facet_of_def face_of_imp_convex) then obtain x where x: "x \ rel_interior C" by (force simp: rel_interior_eq_empty) then have "x \ C" by (meson subsetD rel_interior_subset) then have "x \ S" using \C facet_of S\ facet_of_imp_subset by blast have rels: "rel_interior S = {x \ S. \h\F. a h \ x < b h}" by (rule rel_interior_polyhedron_explicit [OF assms]) have "C \ S" using \C facet_of S\ facet_of_irrefl by blast then have "x \ rel_interior S" by (metis IntI empty_iff \x \ C\ \C \ S\ \C face_of S\ face_of_disjoint_rel_interior) with rels \x \ S\ obtain i where "i \ F" and i: "a i \ x \ b i" by force have "x \ {u. a i \ u \ b i}" by (metis IntD2 InterE \i \ F\ \x \ S\ faceq seq) then have "a i \ x \ b i" by simp then have "a i \ x = b i" using i by auto have "C \ S \ {x. a i \ x = b i}" proof (rule subset_of_face_of [of _ S]) show "S \ {x. a i \ x = b i} face_of S" by (simp add: "*" \i \ F\ facet_of_imp_face_of) show "C \ S" by (simp add: \C face_of S\ face_of_imp_subset) show "S \ {x. a i \ x = b i} \ rel_interior C \ {}" using \a i \ x = b i\ \x \ S\ x by blast qed then have cface: "C face_of (S \ {x. a i \ x = b i})" by (meson \C face_of S\ face_of_subset inf_le1) have con: "convex (S \ {x. a i \ x = b i})" by (simp add: \convex S\ convex_Int convex_hyperplane) show "\h. h \ F \ C = S \ {x. a h \ x = b h}" apply (rule_tac x=i in exI) by (metis (no_types) * \i \ F\ affc facet_of_def less_irrefl face_of_aff_dim_lt [OF con cface]) qed qed lemma face_of_polyhedron_subset_explicit: fixes S :: "'a :: euclidean_space set" assumes "finite F" and seq: "S = affine hull S \ \F" and faceq: "\h. h \ F \ a h \ 0 \ h = {x. a h \ x \ b h}" and psub: "\F'. F' \ F \ S \ affine hull S \ \F'" and C: "C face_of S" and "C \ {}" "C \ S" obtains h where "h \ F" "C \ S \ {x. a h \ x = b h}" proof - have "C \ S" using \C face_of S\ by (simp add: face_of_imp_subset) have "polyhedron S" by (metis \finite F\ faceq polyhedron_Int polyhedron_Inter polyhedron_affine_hull polyhedron_halfspace_le seq) then have "convex S" by (simp add: polyhedron_imp_convex) then have *: "(S \ {x. a h \ x = b h}) face_of S" if "h \ F" for h using faceq seq face_of_Int_supporting_hyperplane_le that by fastforce have "rel_interior C \ {}" using C \C \ {}\ face_of_imp_convex rel_interior_eq_empty by blast then obtain x where "x \ rel_interior C" by auto have rels: "rel_interior S = {x \ S. \h\F. a h \ x < b h}" by (rule rel_interior_polyhedron_explicit [OF \finite F\ seq faceq psub]) then have xnot: "x \ rel_interior S" by (metis IntI \x \ rel_interior C\ C \C \ S\ contra_subsetD empty_iff face_of_disjoint_rel_interior rel_interior_subset) then have "x \ S" using \C \ S\ \x \ rel_interior C\ rel_interior_subset by auto then have xint: "x \ \F" using seq by blast have "F \ {}" using assms by (metis affine_Int affine_Inter affine_affine_hull ex_in_conv face_of_affine_trivial) then obtain i where "i \ F" "\ (a i \ x < b i)" using \x \ S\ rels xnot by auto with xint have "a i \ x = b i" by (metis eq_iff mem_Collect_eq not_le Inter_iff faceq) have face: "S \ {x. a i \ x = b i} face_of S" by (simp add: "*" \i \ F\) show ?thesis proof show "C \ S \ {x. a i \ x = b i}" using subset_of_face_of [OF face \C \ S\] \a i \ x = b i\ \x \ rel_interior C\ \x \ S\ by blast qed fact qed text\Initial part of proof duplicates that above\ proposition face_of_polyhedron_explicit: fixes S :: "'a :: euclidean_space set" assumes "finite F" and seq: "S = affine hull S \ \F" and faceq: "\h. h \ F \ a h \ 0 \ h = {x. a h \ x \ b h}" and psub: "\F'. F' \ F \ S \ affine hull S \ \F'" and C: "C face_of S" and "C \ {}" "C \ S" shows "C = \{S \ {x. a h \ x = b h} | h. h \ F \ C \ S \ {x. a h \ x = b h}}" proof - let ?ab = "\h. {x. a h \ x = b h}" have "C \ S" using \C face_of S\ by (simp add: face_of_imp_subset) have "polyhedron S" by (metis \finite F\ faceq polyhedron_Int polyhedron_Inter polyhedron_affine_hull polyhedron_halfspace_le seq) then have "convex S" by (simp add: polyhedron_imp_convex) then have *: "(S \ ?ab h) face_of S" if "h \ F" for h using faceq seq face_of_Int_supporting_hyperplane_le that by fastforce have "rel_interior C \ {}" using C \C \ {}\ face_of_imp_convex rel_interior_eq_empty by blast then obtain z where z: "z \ rel_interior C" by auto have rels: "rel_interior S = {z \ S. \h\F. a h \ z < b h}" by (rule rel_interior_polyhedron_explicit [OF \finite F\ seq faceq psub]) then have xnot: "z \ rel_interior S" by (metis IntI \z \ rel_interior C\ C \C \ S\ contra_subsetD empty_iff face_of_disjoint_rel_interior rel_interior_subset) then have "z \ S" using \C \ S\ \z \ rel_interior C\ rel_interior_subset by auto with seq have xint: "z \ \F" by blast have "open (\h\{h \ F. a h \ z < b h}. {w. a h \ w < b h})" by (auto simp: \finite F\ open_halfspace_lt open_INT) then obtain e where "0 < e" "ball z e \ (\h\{h \ F. a h \ z < b h}. {w. a h \ w < b h})" by (auto intro: openE [of _ z]) then have e: "\h. \h \ F; a h \ z < b h\ \ ball z e \ {w. a h \ w < b h}" by blast have "C \ (S \ ?ab h) \ z \ S \ ?ab h" if "h \ F" for h proof show "z \ S \ ?ab h \ C \ S \ ?ab h" by (metis "*" Collect_cong IntI \C \ S\ empty_iff subset_of_face_of that z) next show "C \ S \ ?ab h \ z \ S \ ?ab h" using \z \ rel_interior C\ rel_interior_subset by force qed then have **: "{S \ ?ab h | h. h \ F \ C \ S \ C \ ?ab h} = {S \ ?ab h |h. h \ F \ z \ S \ ?ab h}" by blast have bsub: "ball z e \ affine hull \{S \ ?ab h |h. h \ F \ a h \ z = b h} \ affine hull S \ \F \ \{?ab h |h. h \ F \ a h \ z = b h}" if "i \ F" and i: "a i \ z = b i" for i proof - have sub: "ball z e \ \{?ab h |h. h \ F \ a h \ z = b h} \ j" if "j \ F" for j proof - have "a j \ z \ b j" using faceq that xint by auto then consider "a j \ z < b j" | "a j \ z = b j" by linarith then have "\G. G \ {?ab h |h. h \ F \ a h \ z = b h} \ ball z e \ G \ j" proof cases assume "a j \ z < b j" then have "ball z e \ {x. a i \ x = b i} \ j" using e [OF \j \ F\] faceq that by (fastforce simp: ball_def) then show ?thesis by (rule_tac x="{x. a i \ x = b i}" in exI) (force simp: \i \ F\ i) next assume eq: "a j \ z = b j" with faceq that show ?thesis by (rule_tac x="{x. a j \ x = b j}" in exI) (fastforce simp add: \j \ F\) qed then show ?thesis by blast qed have 1: "affine hull \{S \ ?ab h |h. h \ F \ a h \ z = b h} \ affine hull S" using that \z \ S\ by (intro hull_mono) auto have 2: "affine hull \{S \ ?ab h |h. h \ F \ a h \ z = b h} \ \{?ab h |h. h \ F \ a h \ z = b h}" by (rule hull_minimal) (auto intro: affine_hyperplane) have 3: "ball z e \ \{?ab h |h. h \ F \ a h \ z = b h} \ \F" by (iprover intro: sub Inter_greatest) have *: "\A \ (B :: 'a set); A \ C; E \ C \ D\ \ E \ A \ (B \ D) \ C" for A B C D E by blast show ?thesis by (intro * 1 2 3) qed have "\h. h \ F \ C \ ?ab h" using assms by (metis face_of_polyhedron_subset_explicit [OF \finite F\ seq faceq psub] le_inf_iff) then have fac: "\{S \ ?ab h |h. h \ F \ C \ S \ ?ab h} face_of S" using * by (force simp: \C \ S\ intro: face_of_Inter) have red: "(\a. P a \ T \ S \ \{F X |X. P X}) \ T \ \{S \ F X |X::'a set. P X}" for P T F by blast have "ball z e \ affine hull \{S \ ?ab h |h. h \ F \ a h \ z = b h} \ \{S \ ?ab h |h. h \ F \ a h \ z = b h}" by (rule red) (metis seq bsub) with \0 < e\ have zinrel: "z \ rel_interior (\{S \ ?ab h |h. h \ F \ z \ S \ a h \ z = b h})" by (auto simp: mem_rel_interior_ball \z \ S\) show ?thesis using z zinrel by (intro face_of_eq [OF C fac]) (force simp: **) qed subsection\More general corollaries from the explicit representation\ corollary facet_of_polyhedron: assumes "polyhedron S" and "C facet_of S" obtains a b where "a \ 0" "S \ {x. a \ x \ b}" "C = S \ {x. a \ x = b}" proof - obtain F where "finite F" and seq: "S = affine hull S \ \F" and faces: "\h. h \ F \ \a b. a \ 0 \ h = {x. a \ x \ b}" and min: "\F'. F' \ F \ S \ (affine hull S) \ \F'" using assms by (simp add: polyhedron_Int_affine_minimal) meson then obtain a b where ab: "\h. h \ F \ a h \ 0 \ h = {x. a h \ x \ b h}" by metis obtain i where "i \ F" and C: "C = S \ {x. a i \ x = b i}" using facet_of_polyhedron_explicit [OF \finite F\ seq ab min] assms by force moreover have ssub: "S \ {x. a i \ x \ b i}" using \i \ F\ ab by (subst seq) auto ultimately show ?thesis by (rule_tac a = "a i" and b = "b i" in that) (simp_all add: ab) qed corollary face_of_polyhedron: assumes "polyhedron S" and "C face_of S" and "C \ {}" and "C \ S" shows "C = \{F. F facet_of S \ C \ F}" proof - obtain F where "finite F" and seq: "S = affine hull S \ \F" and faces: "\h. h \ F \ \a b. a \ 0 \ h = {x. a \ x \ b}" and min: "\F'. F' \ F \ S \ (affine hull S) \ \F'" using assms by (simp add: polyhedron_Int_affine_minimal) meson then obtain a b where ab: "\h. h \ F \ a h \ 0 \ h = {x. a h \ x \ b h}" by metis show ?thesis apply (subst face_of_polyhedron_explicit [OF \finite F\ seq ab min]) apply (auto simp: assms facet_of_polyhedron_explicit [OF \finite F\ seq ab min] cong: Collect_cong) done qed lemma face_of_polyhedron_subset_facet: assumes "polyhedron S" and "C face_of S" and "C \ {}" and "C \ S" obtains F where "F facet_of S" "C \ F" using face_of_polyhedron assms by (metis (no_types, lifting) Inf_greatest antisym_conv face_of_imp_subset mem_Collect_eq) lemma exposed_face_of_polyhedron: assumes "polyhedron S" shows "F exposed_face_of S \ F face_of S" proof show "F exposed_face_of S \ F face_of S" by (simp add: exposed_face_of_def) next assume "F face_of S" show "F exposed_face_of S" proof (cases "F = {} \ F = S") case True then show ?thesis using \F face_of S\ exposed_face_of by blast next case False then have "{g. g facet_of S \ F \ g} \ {}" by (metis Collect_empty_eq_bot \F face_of S\ assms empty_def face_of_polyhedron_subset_facet) moreover have "\T. \T facet_of S; F \ T\ \ T exposed_face_of S" by (metis assms exposed_face_of facet_of_imp_face_of facet_of_polyhedron) ultimately have "\{G. G facet_of S \ F \ G} exposed_face_of S" by (metis (no_types, lifting) mem_Collect_eq exposed_face_of_Inter) then show ?thesis using False \F face_of S\ assms face_of_polyhedron by fastforce qed qed lemma face_of_polyhedron_polyhedron: fixes S :: "'a :: euclidean_space set" assumes "polyhedron S" "c face_of S" shows "polyhedron c" by (metis assms face_of_imp_eq_affine_Int polyhedron_Int polyhedron_affine_hull polyhedron_imp_convex) lemma finite_polyhedron_faces: fixes S :: "'a :: euclidean_space set" assumes "polyhedron S" shows "finite {F. F face_of S}" proof - obtain F where "finite F" and seq: "S = affine hull S \ \F" and faces: "\h. h \ F \ \a b. a \ 0 \ h = {x. a \ x \ b}" and min: "\F'. F' \ F \ S \ (affine hull S) \ \F'" using assms by (simp add: polyhedron_Int_affine_minimal) meson then obtain a b where ab: "\h. h \ F \ a h \ 0 \ h = {x. a h \ x \ b h}" by metis have "finite {\{S \ {x. a h \ x = b h} |h. h \ F'}| F'. F' \ Pow F}" by (simp add: \finite F\) moreover have "{F. F face_of S} - {{}, S} \ {\{S \ {x. a h \ x = b h} |h. h \ F'}| F'. F' \ Pow F}" apply clarify apply (rename_tac c) apply (drule face_of_polyhedron_explicit [OF \finite F\ seq ab min, simplified], simp_all) apply (rule_tac x="{h \ F. c \ S \ {x. a h \ x = b h}}" in exI, auto) done ultimately show ?thesis by (meson finite.emptyI finite.insertI finite_Diff2 finite_subset) qed lemma finite_polyhedron_exposed_faces: "polyhedron S \ finite {F. F exposed_face_of S}" using exposed_face_of_polyhedron finite_polyhedron_faces by fastforce lemma finite_polyhedron_extreme_points: fixes S :: "'a :: euclidean_space set" assumes "polyhedron S" shows "finite {v. v extreme_point_of S}" proof - have "finite {v. {v} face_of S}" using assms by (intro finite_subset [OF _ finite_vimageI [OF finite_polyhedron_faces]], auto) then show ?thesis by (simp add: face_of_singleton) qed lemma finite_polyhedron_facets: fixes S :: "'a :: euclidean_space set" shows "polyhedron S \ finite {F. F facet_of S}" unfolding facet_of_def by (blast intro: finite_subset [OF _ finite_polyhedron_faces]) proposition rel_interior_of_polyhedron: fixes S :: "'a :: euclidean_space set" assumes "polyhedron S" shows "rel_interior S = S - \{F. F facet_of S}" proof - obtain F where "finite F" and seq: "S = affine hull S \ \F" and faces: "\h. h \ F \ \a b. a \ 0 \ h = {x. a \ x \ b}" and min: "\F'. F' \ F \ S \ (affine hull S) \ \F'" using assms by (simp add: polyhedron_Int_affine_minimal) meson then obtain a b where ab: "\h. h \ F \ a h \ 0 \ h = {x. a h \ x \ b h}" by metis have facet: "(c facet_of S) \ (\h. h \ F \ c = S \ {x. a h \ x = b h})" for c by (rule facet_of_polyhedron_explicit [OF \finite F\ seq ab min]) have rel: "rel_interior S = {x \ S. \h\F. a h \ x < b h}" by (rule rel_interior_polyhedron_explicit [OF \finite F\ seq ab min]) have "a h \ x < b h" if "x \ S" "h \ F" and xnot: "x \ \{F. F facet_of S}" for x h proof - have "x \ \F" using seq that by force with \h \ F\ ab have "a h \ x \ b h" by auto then consider "a h \ x < b h" | "a h \ x = b h" by linarith then show ?thesis proof cases case 1 then show ?thesis . next case 2 have "Collect ((\) x) \ Collect ((\) (\{A. A facet_of S}))" using xnot by fastforce then have "F \ Collect ((\) h)" using 2 \x \ S\ facet by blast with 2 that \x \ \F\ show ?thesis by blast qed qed moreover have "\h\F. a h \ x \ b h" if "x \ \{F. F facet_of S}" for x using that by (force simp: facet) ultimately show ?thesis by (force simp: rel) qed lemma rel_boundary_of_polyhedron: fixes S :: "'a :: euclidean_space set" assumes "polyhedron S" shows "S - rel_interior S = \ {F. F facet_of S}" using facet_of_imp_subset by (fastforce simp add: rel_interior_of_polyhedron assms) lemma rel_frontier_of_polyhedron: fixes S :: "'a :: euclidean_space set" assumes "polyhedron S" shows "rel_frontier S = \ {F. F facet_of S}" by (simp add: assms rel_frontier_def polyhedron_imp_closed rel_boundary_of_polyhedron) lemma rel_frontier_of_polyhedron_alt: fixes S :: "'a :: euclidean_space set" assumes "polyhedron S" shows "rel_frontier S = \ {F. F face_of S \ F \ S}" proof show "rel_frontier S \ \ {F. F face_of S \ F \ S}" by (force simp: rel_frontier_of_polyhedron facet_of_def assms) qed (use face_of_subset_rel_frontier in fastforce) text\A characterization of polyhedra as having finitely many faces\ proposition polyhedron_eq_finite_exposed_faces: fixes S :: "'a :: euclidean_space set" shows "polyhedron S \ closed S \ convex S \ finite {F. F exposed_face_of S}" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (auto simp: polyhedron_imp_closed polyhedron_imp_convex finite_polyhedron_exposed_faces) next assume ?rhs then have "closed S" "convex S" and fin: "finite {F. F exposed_face_of S}" by auto show ?lhs proof (cases "S = {}") case True then show ?thesis by auto next case False define F where "F = {h. h exposed_face_of S \ h \ {} \ h \ S}" have "finite F" by (simp add: fin F_def) have hface: "h face_of S" and "\a b. a \ 0 \ S \ {x. a \ x \ b} \ h = S \ {x. a \ x = b}" if "h \ F" for h using exposed_face_of F_def that by blast+ then obtain a b where ab: "\h. h \ F \ a h \ 0 \ S \ {x. a h \ x \ b h} \ h = S \ {x. a h \ x = b h}" by metis have *: "False" if paff: "p \ affine hull S" and "p \ S" and pint: "p \ \{{x. a h \ x \ b h} |h. h \ F}" for p proof - have "rel_interior S \ {}" by (simp add: \S \ {}\ \convex S\ rel_interior_eq_empty) then obtain c where c: "c \ rel_interior S" by auto with rel_interior_subset have "c \ S" by blast have ccp: "closed_segment c p \ affine hull S" by (meson affine_affine_hull affine_imp_convex c closed_segment_subset hull_subset paff rel_interior_subset subsetCE) have oS: "openin (top_of_set (closed_segment c p)) (closed_segment c p \ rel_interior S)" by (force simp: openin_rel_interior openin_Int intro: openin_subtopology_Int_subset [OF _ ccp]) obtain x where xcl: "x \ closed_segment c p" and "x \ S" and xnot: "x \ rel_interior S" using connected_openin [of "closed_segment c p"] apply simp apply (drule_tac x="closed_segment c p \ rel_interior S" in spec) apply (drule mp [OF _ oS]) apply (drule_tac x="closed_segment c p \ (- S)" in spec) using rel_interior_subset \closed S\ c \p \ S\ apply blast done then obtain \ where "0 \ \" "\ \ 1" and xeq: "x = (1 - \) *\<^sub>R c + \ *\<^sub>R p" by (auto simp: in_segment) show False proof (cases "\=0 \ \=1") case True with xeq c xnot \x \ S\ \p \ S\ show False by auto next case False then have xos: "x \ open_segment c p" using \x \ S\ c open_segment_def that(2) xcl xnot by auto have xclo: "x \ closure S" using \x \ S\ closure_subset by blast obtain d where "d \ 0" and dle: "\y. y \ closure S \ d \ x \ d \ y" and dless: "\y. y \ rel_interior S \ d \ x < d \ y" by (metis supporting_hyperplane_relative_frontier [OF \convex S\ xclo xnot]) have sex: "S \ {y. d \ y = d \ x} exposed_face_of S" by (simp add: \closed S\ dle exposed_face_of_Int_supporting_hyperplane_ge [OF \convex S\]) have sne: "S \ {y. d \ y = d \ x} \ {}" using \x \ S\ by blast have sns: "S \ {y. d \ y = d \ x} \ S" by (metis (mono_tags) Int_Collect c subsetD dless not_le order_refl rel_interior_subset) obtain h where "h \ F" "x \ h" using F_def \x \ S\ sex sns by blast have abface: "{y. a h \ y = b h} face_of {y. a h \ y \ b h}" using hyperplane_face_of_halfspace_le by blast then have "c \ h" using face_ofD [OF abface xos] \c \ S\ \h \ F\ ab pint \x \ h\ by blast with c have "h \ rel_interior S \ {}" by blast then show False using \h \ F\ F_def face_of_disjoint_rel_interior hface by auto qed qed have "S \ affine hull S \ \{{x. a h \ x \ b h} |h. h \ F}" using ab by (auto simp: hull_subset) moreover have "affine hull S \ \{{x. a h \ x \ b h} |h. h \ F} \ S" using * by blast ultimately have "S = affine hull S \ \ {{x. a h \ x \ b h} |h. h \ F}" .. then show ?thesis apply (rule ssubst) apply (force intro: polyhedron_affine_hull polyhedron_halfspace_le simp: \finite F\) done qed qed corollary polyhedron_eq_finite_faces: fixes S :: "'a :: euclidean_space set" shows "polyhedron S \ closed S \ convex S \ finite {F. F face_of S}" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (simp add: finite_polyhedron_faces polyhedron_imp_closed polyhedron_imp_convex) next assume ?rhs then show ?lhs by (force simp: polyhedron_eq_finite_exposed_faces exposed_face_of intro: finite_subset) qed lemma polyhedron_linear_image_eq: fixes h :: "'a :: euclidean_space \ 'b :: euclidean_space" assumes "linear h" "bij h" shows "polyhedron (h ` S) \ polyhedron S" proof - have *: "{f. P f} = (image h) ` {f. P (h ` f)}" for P apply safe apply (rule_tac x="inv h ` x" in image_eqI) apply (auto simp: \bij h\ bij_is_surj image_f_inv_f) done have "inj h" using bij_is_inj assms by blast then have injim: "inj_on ((`) h) A" for A by (simp add: inj_on_def inj_image_eq_iff) show ?thesis using \linear h\ \inj h\ apply (simp add: polyhedron_eq_finite_faces closed_injective_linear_image_eq) apply (simp add: * face_of_linear_image [of h _ S, symmetric] finite_image_iff injim) done qed lemma polyhedron_negations: fixes S :: "'a :: euclidean_space set" shows "polyhedron S \ polyhedron(image uminus S)" by (subst polyhedron_linear_image_eq) (auto simp: bij_uminus intro!: linear_uminus) subsection\Relation between polytopes and polyhedra\ proposition polytope_eq_bounded_polyhedron: fixes S :: "'a :: euclidean_space set" shows "polytope S \ polyhedron S \ bounded S" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (simp add: finite_polytope_faces polyhedron_eq_finite_faces polytope_imp_closed polytope_imp_convex polytope_imp_bounded) next assume R: ?rhs then have "finite {v. v extreme_point_of S}" by (simp add: finite_polyhedron_extreme_points) moreover have "S = convex hull {v. v extreme_point_of S}" using R by (simp add: Krein_Milman_Minkowski compact_eq_bounded_closed polyhedron_imp_closed polyhedron_imp_convex) ultimately show ?lhs unfolding polytope_def by blast qed lemma polytope_Int: fixes S :: "'a :: euclidean_space set" shows "\polytope S; polytope T\ \ polytope(S \ T)" by (simp add: polytope_eq_bounded_polyhedron bounded_Int) lemma polytope_Int_polyhedron: fixes S :: "'a :: euclidean_space set" shows "\polytope S; polyhedron T\ \ polytope(S \ T)" by (simp add: bounded_Int polytope_eq_bounded_polyhedron) lemma polyhedron_Int_polytope: fixes S :: "'a :: euclidean_space set" shows "\polyhedron S; polytope T\ \ polytope(S \ T)" by (simp add: bounded_Int polytope_eq_bounded_polyhedron) lemma polytope_imp_polyhedron: fixes S :: "'a :: euclidean_space set" shows "polytope S \ polyhedron S" by (simp add: polytope_eq_bounded_polyhedron) lemma polytope_facet_exists: fixes p :: "'a :: euclidean_space set" assumes "polytope p" "0 < aff_dim p" obtains F where "F facet_of p" proof (cases "p = {}") case True with assms show ?thesis by auto next case False then obtain v where "v extreme_point_of p" using extreme_point_exists_convex by (blast intro: \polytope p\ polytope_imp_compact polytope_imp_convex) then show ?thesis by (metis face_of_polyhedron_subset_facet polytope_imp_polyhedron aff_dim_sing all_not_in_conv assms face_of_singleton less_irrefl singletonI that) qed lemma polyhedron_interval [iff]: "polyhedron(cbox a b)" by (metis polytope_imp_polyhedron polytope_interval) lemma polyhedron_convex_hull: fixes S :: "'a :: euclidean_space set" shows "finite S \ polyhedron(convex hull S)" by (simp add: polytope_convex_hull polytope_imp_polyhedron) subsection\Relative and absolute frontier of a polytope\ lemma rel_boundary_of_convex_hull: fixes S :: "'a::euclidean_space set" assumes "\ affine_dependent S" shows "(convex hull S) - rel_interior(convex hull S) = (\a\S. convex hull (S - {a}))" proof - have "finite S" by (metis assms aff_independent_finite) then consider "card S = 0" | "card S = 1" | "2 \ card S" by arith then show ?thesis proof cases case 1 then have "S = {}" by (simp add: \finite S\) then show ?thesis by simp next case 2 show ?thesis by (auto intro: card_1_singletonE [OF \card S = 1\]) next case 3 with assms show ?thesis by (auto simp: polyhedron_convex_hull rel_boundary_of_polyhedron facet_of_convex_hull_affine_independent_alt \finite S\) qed qed proposition frontier_of_convex_hull: fixes S :: "'a::euclidean_space set" assumes "card S = Suc (DIM('a))" shows "frontier(convex hull S) = \ {convex hull (S - {a}) | a. a \ S}" proof (cases "affine_dependent S") case True have [iff]: "finite S" using assms using card.infinite by force then have ccs: "closed (convex hull S)" by (simp add: compact_imp_closed finite_imp_compact_convex_hull) { fix x T assume "int (card T) \ aff_dim S + 1" "finite T" "T \ S""x \ convex hull T" then have "S \ T" using True \finite S\ aff_dim_le_card affine_independent_iff_card by fastforce then obtain a where "a \ S" "a \ T" using \T \ S\ by blast then have "\y\S. x \ convex hull (S - {y})" using True affine_independent_iff_card [of S] by (metis (no_types, opaque_lifting) Diff_eq_empty_iff Diff_insert0 \a \ T\ \T \ S\ \x \ convex hull T\ hull_mono insert_Diff_single subsetCE) } note * = this have 1: "convex hull S \ (\ a\S. convex hull (S - {a}))" by (subst caratheodory_aff_dim) (blast dest: *) have 2: "\((\a. convex hull (S - {a})) ` S) \ convex hull S" by (rule Union_least) (metis (no_types, lifting) Diff_subset hull_mono imageE) show ?thesis using True apply (simp add: segment_convex_hull frontier_def) using interior_convex_hull_eq_empty [OF assms] apply (simp add: closure_closed [OF ccs]) using "1" "2" by auto next case False then have "frontier (convex hull S) = closure (convex hull S) - interior (convex hull S)" by (simp add: rel_boundary_of_convex_hull frontier_def) also have "\ = (convex hull S) - rel_interior(convex hull S)" by (metis False aff_independent_finite assms closure_convex_hull finite_imp_compact_convex_hull hull_hull interior_convex_hull_eq_empty rel_interior_nonempty_interior) also have "\ = \{convex hull (S - {a}) |a. a \ S}" proof - have "convex hull S - rel_interior (convex hull S) = rel_frontier (convex hull S)" by (simp add: False aff_independent_finite polyhedron_convex_hull rel_boundary_of_polyhedron rel_frontier_of_polyhedron) then show ?thesis by (simp add: False rel_frontier_convex_hull_cases) qed finally show ?thesis . qed subsection\Special case of a triangle\ proposition frontier_of_triangle: fixes a :: "'a::euclidean_space" assumes "DIM('a) = 2" shows "frontier(convex hull {a,b,c}) = closed_segment a b \ closed_segment b c \ closed_segment c a" (is "?lhs = ?rhs") proof (cases "b = a \ c = a \ c = b") case True then show ?thesis by (auto simp: assms segment_convex_hull frontier_def empty_interior_convex_hull insert_commute card_insert_le_m1 hull_inc insert_absorb) next case False then have [simp]: "card {a, b, c} = Suc (DIM('a))" by (simp add: card.insert_remove Set.insert_Diff_if assms) show ?thesis proof show "?lhs \ ?rhs" using False by (force simp: segment_convex_hull frontier_of_convex_hull insert_Diff_if insert_commute split: if_split_asm) show "?rhs \ ?lhs" using False apply (simp add: frontier_of_convex_hull segment_convex_hull) apply (intro conjI subsetI) apply (rule_tac X="convex hull {a,b}" in UnionI; force simp: Set.insert_Diff_if) apply (rule_tac X="convex hull {b,c}" in UnionI; force) apply (rule_tac X="convex hull {a,c}" in UnionI; force simp: insert_commute Set.insert_Diff_if) done qed qed corollary inside_of_triangle: fixes a :: "'a::euclidean_space" assumes "DIM('a) = 2" shows "inside (closed_segment a b \ closed_segment b c \ closed_segment c a) = interior(convex hull {a,b,c})" by (metis assms frontier_of_triangle bounded_empty bounded_insert convex_convex_hull inside_frontier_eq_interior bounded_convex_hull) corollary interior_of_triangle: fixes a :: "'a::euclidean_space" assumes "DIM('a) = 2" shows "interior(convex hull {a,b,c}) = convex hull {a,b,c} - (closed_segment a b \ closed_segment b c \ closed_segment c a)" using interior_subset by (force simp: frontier_of_triangle [OF assms, symmetric] frontier_def Diff_Diff_Int) subsection\Subdividing a cell complex\ lemma subdivide_interval: fixes x::real assumes "a < \x - y\" "0 < a" obtains n where "n \ \" "x < n * a \ n * a < y \ y < n * a \ n * a < x" proof - consider "a + x < y" | "a + y < x" using assms by linarith then show ?thesis proof cases case 1 let ?n = "of_int (floor (x/a)) + 1" have x: "x < ?n * a" by (meson \0 < a\ divide_less_eq floor_eq_iff) have "?n * a \ a + x" apply (simp add: algebra_simps) by (metis assms(2) floor_divide_lower mult.commute) also have "... < y" by (rule 1) finally have "?n * a < y" . with x show ?thesis using Ints_1 Ints_add Ints_of_int that by blast next case 2 let ?n = "of_int (floor (y/a)) + 1" have y: "y < ?n * a" by (meson \0 < a\ divide_less_eq floor_eq_iff) have "?n * a \ a + y" apply (simp add: algebra_simps) by (metis assms(2) floor_divide_lower mult.commute) also have "... < x" by (rule 2) finally have "?n * a < x" . then show ?thesis using Ints_1 Ints_add Ints_of_int that y by blast qed qed lemma cell_subdivision_lemma: assumes "finite \" and "\X. X \ \ \ polytope X" and "\X. X \ \ \ aff_dim X \ d" and "\X Y. \X \ \; Y \ \\ \ (X \ Y) face_of X" and "finite I" shows "\\. \\ = \\ \ finite \ \ (\C \ \. \D. D \ \ \ C \ D) \ (\C \ \. \x \ C. \D. D \ \ \ x \ D \ D \ C) \ (\X \ \. polytope X) \ (\X \ \. aff_dim X \ d) \ (\X \ \. \Y \ \. X \ Y face_of X) \ (\X \ \. \x \ X. \y \ X. \a b. (a,b) \ I \ a \ x \ b \ a \ y \ b \ a \ x \ b \ a \ y \ b)" using \finite I\ proof induction case empty then show ?case by (rule_tac x="\" in exI) (auto simp: assms) next case (insert ab I) then obtain \ where eq: "\\ = \\" and "finite \" and sub1: "\C. C \ \ \ \D. D \ \ \ C \ D" and sub2: "\C x. C \ \ \ x \ C \ \D. D \ \ \ x \ D \ D \ C" and poly: "\X. X \ \ \ polytope X" and aff: "\X. X \ \ \ aff_dim X \ d" and face: "\X Y. \X \ \; Y \ \\ \ X \ Y face_of X" and I: "\X x y a b. \X \ \; x \ X; y \ X; (a,b) \ I\ \ a \ x \ b \ a \ y \ b \ a \ x \ b \ a \ y \ b" by (auto simp: that) obtain a b where "ab = (a,b)" by fastforce let ?\ = "(\X. X \ {x. a \ x \ b}) ` \ \ (\X. X \ {x. a \ x \ b}) ` \" have eqInt: "(S \ Collect P) \ (T \ Collect Q) = (S \ T) \ (Collect P \ Collect Q)" for S T::"'a set" and P Q by blast show ?case proof (intro conjI exI) show "\?\ = \\" by (force simp: eq [symmetric]) show "finite ?\" using \finite \\ by force show "\X \ ?\. polytope X" by (force simp: poly polytope_Int_polyhedron polyhedron_halfspace_le polyhedron_halfspace_ge) show "\X \ ?\. aff_dim X \ d" by (auto; metis order_trans aff aff_dim_subset inf_le1) show "\X \ ?\. \x \ X. \y \ X. \a b. (a,b) \ insert ab I \ a \ x \ b \ a \ y \ b \ a \ x \ b \ a \ y \ b" using \ab = (a, b)\ I by fastforce show "\X \ ?\. \Y \ ?\. X \ Y face_of X" by (auto simp: eqInt halfspace_Int_eq face_of_Int_Int face face_of_halfspace_le face_of_halfspace_ge) show "\C \ ?\. \D. D \ \ \ C \ D" using sub1 by force show "\C\\. \x\C. \D. D \ ?\ \ x \ D \ D \ C" proof (intro ballI) fix C z assume "C \ \" "z \ C" with sub2 obtain D where D: "D \ \" "z \ D" "D \ C" by blast have "D \ \ \ z \ D \ {x. a \ x \ b} \ D \ {x. a \ x \ b} \ C \ D \ \ \ z \ D \ {x. a \ x \ b} \ D \ {x. a \ x \ b} \ C" using linorder_class.linear [of "a \ z" b] D by blast then show "\D. D \ ?\ \ z \ D \ D \ C" by blast qed qed qed proposition cell_complex_subdivision_exists: fixes \ :: "'a::euclidean_space set set" assumes "0 < e" "finite \" and poly: "\X. X \ \ \ polytope X" and aff: "\X. X \ \ \ aff_dim X \ d" and face: "\X Y. \X \ \; Y \ \\ \ X \ Y face_of X" obtains "\'" where "finite \'" "\\' = \\" "\X. X \ \' \ diameter X < e" "\X. X \ \' \ polytope X" "\X. X \ \' \ aff_dim X \ d" "\X Y. \X \ \'; Y \ \'\ \ X \ Y face_of X" "\C. C \ \' \ \D. D \ \ \ C \ D" "\C x. C \ \ \ x \ C \ \D. D \ \' \ x \ D \ D \ C" proof - have "bounded(\\)" by (simp add: \finite \\ poly bounded_Union polytope_imp_bounded) then obtain B where "B > 0" and B: "\x. x \ \\ \ norm x < B" by (meson bounded_pos_less) define C where "C \ {z \ \. \z * e / 2 / real DIM('a)\ \ B}" define I where "I \ \i \ Basis. \j \ C. { (i::'a, j * e / 2 / DIM('a)) }" have "C \ {x \ \. - B / (e / 2 / real DIM('a)) \ x \ x \ B / (e / 2 / real DIM('a))}" using \0 < e\ by (auto simp: field_split_simps C_def) then have "finite C" using finite_int_segment finite_subset by blast then have "finite I" by (simp add: I_def) obtain \' where eq: "\\' = \\" and "finite \'" and poly: "\X. X \ \' \ polytope X" and aff: "\X. X \ \' \ aff_dim X \ d" and face: "\X Y. \X \ \'; Y \ \'\ \ X \ Y face_of X" and I: "\X x y a b. \X \ \'; x \ X; y \ X; (a,b) \ I\ \ a \ x \ b \ a \ y \ b \ a \ x \ b \ a \ y \ b" and sub1: "\C. C \ \' \ \D. D \ \ \ C \ D" and sub2: "\C x. C \ \ \ x \ C \ \D. D \ \' \ x \ D \ D \ C" apply (rule exE [OF cell_subdivision_lemma]) using assms \finite I\ by auto show ?thesis proof (rule_tac \'="\'" in that) show "diameter X < e" if "X \ \'" for X proof - have "diameter X \ e/2" proof (rule diameter_le) show "norm (x - y) \ e / 2" if "x \ X" "y \ X" for x y proof - have "norm x < B" "norm y < B" 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))" proof (rule sum_bounded_above) fix i::'a assume "i \ Basis" then have I': "\z b. \z \ C; b = z * e / (2 * real DIM('a))\ \ i \ x \ b \ i \ y \ b \ i \ x \ b \ i \ y \ b" using I[of X x y] \X \ \'\ that unfolding I_def by auto show "\(x - y) \ i\ \ e / 2 / real DIM('a)" proof (rule ccontr) assume "\ \(x - y) \ i\ \ e / 2 / real DIM('a)" then have xyi: "\i \ x - i \ y\ > e / 2 / real DIM('a)" by (simp add: inner_commute inner_diff_right) obtain n where "n \ \" and n: "i \ x < n * (e / 2 / real DIM('a)) \ n * (e / 2 / real DIM('a)) < i \ y \ i \ y < n * (e / 2 / real DIM('a)) \ n * (e / 2 / real DIM('a)) < i \ x" using subdivide_interval [OF xyi] DIM_positive \0 < e\ by (auto simp: zero_less_divide_iff) have "\i \ x\ < B" by (metis \i \ Basis\ \norm x < B\ inner_commute norm_bound_Basis_lt) have "\i \ y\ < B" by (metis \i \ Basis\ \norm y < B\ inner_commute norm_bound_Basis_lt) have *: "\n * e\ \ B * (2 * real DIM('a))" if "\ix\ < B" "\iy\ < B" and ix: "ix * (2 * real DIM('a)) < n * e" and iy: "n * e < iy * (2 * real DIM('a))" for ix iy proof (rule abs_leI) have "iy * (2 * real DIM('a)) \ B * (2 * real DIM('a))" by (rule mult_right_mono) (use \\iy\ < B\ in linarith)+ then show "n * e \ B * (2 * real DIM('a))" using iy by linarith next have "- ix * (2 * real DIM('a)) \ B * (2 * real DIM('a))" by (rule mult_right_mono) (use \\ix\ < B\ in linarith)+ then show "- (n * e) \ B * (2 * real DIM('a))" using ix by linarith qed have "n \ C" using \n \ \\ n by (auto simp: C_def divide_simps intro: * \\i \ x\ < B\ \\i \ y\ < B\) show False using I' [OF \n \ C\ refl] n by auto qed qed also have "... = e / 2" by simp finally show ?thesis . qed qed (use \0 < e\ in force) also have "... < e" by (simp add: \0 < e\) finally show ?thesis . qed qed (auto simp: eq poly aff face sub1 sub2 \finite \'\) qed subsection\Simplexes\ text\The notion of n-simplex for integer \<^term>\n \ -1\\ definition\<^marker>\tag important\ simplex :: "int \ 'a::euclidean_space set \ bool" (infix "simplex" 50) where "n simplex S \ \C. \ affine_dependent C \ int(card C) = n + 1 \ S = convex hull C" lemma simplex: "n simplex S \ (\C. finite C \ \ affine_dependent C \ int(card C) = n + 1 \ S = convex hull C)" by (auto simp add: simplex_def intro: aff_independent_finite) lemma simplex_convex_hull: "\ affine_dependent C \ int(card C) = n + 1 \ n simplex (convex hull C)" by (auto simp add: simplex_def) lemma convex_simplex: "n simplex S \ convex S" by (metis convex_convex_hull simplex_def) lemma compact_simplex: "n simplex S \ compact S" unfolding simplex using finite_imp_compact_convex_hull by blast lemma closed_simplex: "n simplex S \ closed S" by (simp add: compact_imp_closed compact_simplex) lemma simplex_imp_polytope: "n simplex S \ polytope S" unfolding simplex_def polytope_def using aff_independent_finite by blast lemma simplex_imp_polyhedron: "n simplex S \ polyhedron S" by (simp add: polytope_imp_polyhedron simplex_imp_polytope) lemma simplex_dim_ge: "n simplex S \ -1 \ n" by (metis (no_types, opaque_lifting) aff_dim_geq affine_independent_iff_card diff_add_cancel diff_diff_eq2 simplex_def) lemma simplex_empty [simp]: "n simplex {} \ n = -1" proof assume "n simplex {}" then show "n = -1" unfolding simplex by (metis card.empty convex_hull_eq_empty diff_0 diff_eq_eq of_nat_0) next assume "n = -1" then show "n simplex {}" by (fastforce simp: simplex) qed lemma simplex_minus_1 [simp]: "-1 simplex S \ S = {}" by (metis simplex cancel_comm_monoid_add_class.diff_cancel card_0_eq diff_minus_eq_add of_nat_eq_0_iff simplex_empty) lemma aff_dim_simplex: "n simplex S \ aff_dim S = n" by (metis simplex add.commute add_diff_cancel_left' aff_dim_convex_hull affine_independent_iff_card) lemma zero_simplex_sing: "0 simplex {a}" apply (simp add: simplex_def) using affine_independent_1 card_1_singleton_iff convex_hull_singleton by blast lemma simplex_sing [simp]: "n simplex {a} \ n = 0" using aff_dim_simplex aff_dim_sing zero_simplex_sing by blast lemma simplex_zero: "0 simplex S \ (\a. S = {a})" by (metis aff_dim_eq_0 aff_dim_simplex simplex_sing) lemma one_simplex_segment: "a \ b \ 1 simplex closed_segment a b" unfolding simplex_def by (rule_tac x="{a,b}" in exI) (auto simp: segment_convex_hull) lemma simplex_segment_cases: "(if a = b then 0 else 1) simplex closed_segment a b" by (auto simp: one_simplex_segment) lemma simplex_segment: "\n. n simplex closed_segment a b" using simplex_segment_cases by metis lemma polytope_lowdim_imp_simplex: assumes "polytope P" "aff_dim P \ 1" obtains n where "n simplex P" proof (cases "P = {}") case True then show ?thesis by (simp add: that) next case False then show ?thesis by (metis assms compact_convex_collinear_segment collinear_aff_dim polytope_imp_compact polytope_imp_convex simplex_segment_cases that) qed lemma simplex_insert_dimplus1: fixes n::int assumes "n simplex S" and a: "a \ affine hull S" shows "(n+1) simplex (convex hull (insert a S))" proof - obtain C where C: "finite C" "\ affine_dependent C" "int(card C) = n+1" and S: "S = convex hull C" using assms unfolding simplex by force show ?thesis unfolding simplex proof (intro exI conjI) have "aff_dim S = n" using aff_dim_simplex assms(1) by blast moreover have "a \ affine hull C" using S a affine_hull_convex_hull by blast moreover have "a \ C" using S a hull_inc by fastforce ultimately show "\ affine_dependent (insert a C)" by (simp add: C S aff_dim_convex_hull aff_dim_insert affine_independent_iff_card) next have "a \ C" using S a hull_inc by fastforce then show "int (card (insert a C)) = n + 1 + 1" by (simp add: C) next show "convex hull insert a S = convex hull (insert a C)" by (simp add: S convex_hull_insert_segments) qed (use C in auto) qed subsection \Simplicial complexes and triangulations\ definition\<^marker>\tag important\ simplicial_complex where "simplicial_complex \ \ finite \ \ (\S \ \. \n. n simplex S) \ (\F S. S \ \ \ F face_of S \ F \ \) \ (\S S'. S \ \ \ S' \ \ \ (S \ S') face_of S)" definition\<^marker>\tag important\ triangulation where "triangulation \ \ finite \ \ (\T \ \. \n. n simplex T) \ (\T T'. T \ \ \ T' \ \ \ (T \ T') face_of T)" subsection\Refining a cell complex to a simplicial complex\ proposition convex_hull_insert_Int_eq: fixes z :: "'a :: euclidean_space" assumes z: "z \ rel_interior S" and T: "T \ rel_frontier S" and U: "U \ rel_frontier S" and "convex S" "convex T" "convex U" shows "convex hull (insert z T) \ convex hull (insert z U) = convex hull (insert z (T \ U))" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" proof (cases "T={} \ U={}") case True then show ?thesis by auto next case False then have "T \ {}" "U \ {}" by auto have TU: "convex (T \ U)" by (simp add: \convex T\ \convex U\ convex_Int) have "(\x\T. closed_segment z x) \ (\x\U. closed_segment z x) \ (if T \ U = {} then {z} else \((closed_segment z) ` (T \ U)))" (is "_ \ ?IF") proof clarify fix x t u assume xt: "x \ closed_segment z t" and xu: "x \ closed_segment z u" and "t \ T" "u \ U" then have ne: "t \ z" "u \ z" using T U z unfolding rel_frontier_def by blast+ show "x \ ?IF" proof (cases "x = z") case True then show ?thesis by auto next case False have t: "t \ closure S" using T \t \ T\ rel_frontier_def by auto have u: "u \ closure S" using U \u \ U\ rel_frontier_def by auto show ?thesis proof (cases "t = u") case True then show ?thesis using \t \ T\ \u \ U\ xt by auto next case False have tnot: "t \ closed_segment u z" proof - have "t \ closure S - rel_interior S" using T \t \ T\ rel_frontier_def by blast then have "t \ open_segment z u" by (meson DiffD2 rel_interior_closure_convex_segment [OF \convex S\ z u] subsetD) then show ?thesis by (simp add: \t \ u\ \t \ z\ open_segment_commute open_segment_def) qed moreover have "u \ closed_segment z t" using rel_interior_closure_convex_segment [OF \convex S\ z t] \u \ U\ \u \ z\ U [unfolded rel_frontier_def] tnot by (auto simp: closed_segment_eq_open) ultimately have "\(between (t,u) z | between (u,z) t | between (z,t) u)" if "x \ z" using that xt xu by (meson between_antisym between_mem_segment between_trans_2 ends_in_segment(2)) then have "\ collinear {t, z, u}" if "x \ z" by (auto simp: that collinear_between_cases between_commute) moreover have "collinear {t, z, x}" by (metis closed_segment_commute collinear_2 collinear_closed_segment collinear_triples ends_in_segment(1) insert_absorb insert_absorb2 xt) moreover have "collinear {z, x, u}" by (metis closed_segment_commute collinear_2 collinear_closed_segment collinear_triples ends_in_segment(1) insert_absorb insert_absorb2 xu) ultimately have False using collinear_3_trans [of t z x u] \x \ z\ by blast then show ?thesis by metis qed qed qed then show ?thesis using False \convex T\ \convex U\ TU by (simp add: convex_hull_insert_segments hull_same split: if_split_asm) qed show "?rhs \ ?lhs" by (metis inf_greatest hull_mono inf.cobounded1 inf.cobounded2 insert_mono) qed lemma simplicial_subdivision_aux: assumes "finite \" and "\C. C \ \ \ polytope C" and "\C. C \ \ \ aff_dim C \ of_nat n" and "\C F. \C \ \; F face_of C\ \ F \ \" and "\C1 C2. \C1 \ \; C2 \ \\ \ C1 \ C2 face_of C1" shows "\\. simplicial_complex \ \ (\K \ \. aff_dim K \ of_nat n) \ \\ = \\ \ (\C \ \. \F. finite F \ F \ \ \ C = \F) \ (\K \ \. \C. C \ \ \ K \ C)" using assms proof (induction n arbitrary: \ rule: less_induct) case (less n) then have poly\: "\C. C \ \ \ polytope C" and aff\: "\C. C \ \ \ aff_dim C \ of_nat n" and face\: "\C F. \C \ \; F face_of C\ \ F \ \" and intface\: "\C1 C2. \C1 \ \; C2 \ \\ \ C1 \ C2 face_of C1" by metis+ show ?case proof (cases "n \ 1") case True have "\s. \n \ 1; s \ \\ \ \m. m simplex s" using poly\ aff\ by (force intro: polytope_lowdim_imp_simplex) then show ?thesis unfolding simplicial_complex_def using True by (rule_tac x="\" in exI) (auto simp: less.prems) next case False define \ where "\ \ {C \ \. aff_dim C < n}" have "finite \" "\C. C \ \ \ polytope C" "\C. C \ \ \ aff_dim C \ int (n - 1)" "\C1 C2. \C1 \ \; C2 \ \\ \ C1 \ C2 face_of C1" using less.prems by (auto simp: \_def) moreover have \
: "\C F. \C \ \; F face_of C\ \ F \ \" using less.prems unfolding \_def by (metis (no_types, lifting) mem_Collect_eq aff_dim_subset face_of_imp_subset less_le not_le) ultimately obtain \ where "simplicial_complex \" and aff_dim\: "\K. K \ \ \ aff_dim K \ int (n - 1)" and "\\ = \\" and fin\: "\C. C \ \ \ \F. finite F \ F \ \ \ C = \F" and C\: "\K. K \ \ \ \C. C \ \ \ K \ C" using less.IH [of "n-1" \] False by auto then have "finite \" and simpl\: "\S. S \ \ \ \n. n simplex S" and face\: "\F S. \S \ \; F face_of S\ \ F \ \" and faceI\: "\S S'. \S \ \; S' \ \\ \ (S \ S') face_of S" by (auto simp: simplicial_complex_def) define \ where "\ \ {C \ \. aff_dim C = n}" have "finite \" by (simp add: \_def less.prems(1)) have poly\: "\C. C \ \ \ polytope C" and convex\: "\C. C \ \ \ convex C" and closed\: "\C. C \ \ \ closed C" by (auto simp: \_def poly\ polytope_imp_convex polytope_imp_closed) have in_rel_interior: "(SOME z. z \ rel_interior C) \ rel_interior C" if "C \ \" for C using that poly\ polytope_imp_convex rel_interior_aff_dim some_in_eq by (fastforce simp: \_def) have *: "\T. \ affine_dependent T \ card T \ n \ aff_dim K < n \ K = convex hull T" if "K \ \" for K proof - obtain r where r: "r simplex K" using \K \ \\ simpl\ by blast have "r = aff_dim K" using \r simplex K\ aff_dim_simplex by blast with r show ?thesis unfolding simplex_def using False \\K. K \ \ \ aff_dim K \ int (n - 1)\ that by fastforce qed have ahK_C_disjoint: "affine hull K \ rel_interior C = {}" if "C \ \" "K \ \" "K \ rel_frontier C" for C K proof - have "convex C" "closed C" by (auto simp: convex\ closed\ \C \ \\) obtain F where F: "F face_of C" and "F \ C" "K \ F" proof - obtain L where "L \ \" "K \ L" using \K \ \\ C\ by blast have "K \ rel_frontier C" by (simp add: \K \ rel_frontier C\) also have "... \ C" by (simp add: \closed C\ rel_frontier_def subset_iff) finally have "K \ C" . have "L \ C face_of C" using \_def \_def \C \ \\ \L \ \\ intface\ by (simp add: inf_commute) moreover have "L \ C \ C" using \C \ \\ \L \ \\ by (metis (mono_tags, lifting) \_def \_def intface\ mem_Collect_eq not_le order_refl \
) moreover have "K \ L \ C" using \C \ \\ \L \ \\ \K \ C\ \K \ L\ by (auto simp: \_def \_def) ultimately show ?thesis using that by metis qed have "affine hull F \ rel_interior C = {}" by (rule affine_hull_face_of_disjoint_rel_interior [OF \convex C\ F \F \ C\]) with hull_mono [OF \K \ F\] show "affine hull K \ rel_interior C = {}" by fastforce qed let ?\ = "(\C \ \. \K \ \ \ Pow (rel_frontier C). {convex hull (insert (SOME z. z \ rel_interior C) K)})" have "\\. simplicial_complex \ \ (\K \ \. aff_dim K \ of_nat n) \ (\C \ \. \F. F \ \ \ C = \F) \ (\K \ \. \C. C \ \ \ K \ C)" proof (rule exI, intro conjI ballI) show "simplicial_complex (\ \ ?\)" unfolding simplicial_complex_def proof (intro conjI impI ballI allI) show "finite (\ \ ?\)" using \finite \\ \finite \\ by simp show "\n. n simplex S" if "S \ \ \ ?\" for S using that ahK_C_disjoint in_rel_interior simpl\ simplex_insert_dimplus1 by fastforce show "F \ \ \ ?\" if S: "S \ \ \ ?\ \ F face_of S" for F S proof - have "F \ \" if "S \ \" using S face\ that by blast moreover have "F \ \ \ ?\" if "F face_of S" "C \ \" "K \ \" and "K \ rel_frontier C" and S: "S = convex hull insert (SOME z. z \ rel_interior C) K" for C K proof - let ?z = "SOME z. z \ rel_interior C" have "?z \ rel_interior C" by (simp add: in_rel_interior \C \ \\) moreover obtain I where "\ affine_dependent I" "card I \ n" "aff_dim K < int n" "K = convex hull I" using * [OF \K \ \\] by auto ultimately have "?z \ affine hull I" using ahK_C_disjoint affine_hull_convex_hull that by blast have "compact I" "finite I" 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" using face_of_convex_hull_subset [of "insert ?z I" F] by auto show ?thesis proof (cases "?z \ J") case True have "F \ (\K\\ \ Pow (rel_frontier C). {convex hull insert ?z K})" proof have "convex hull (J - {?z}) face_of K" by (metis True \J \ insert ?z I\ \K = convex hull I\ \\ affine_dependent I\ face_of_convex_hull_affine_independent subset_insert_iff) then have "convex hull (J - {?z}) \ \" by (rule face\ [OF \K \ \\]) moreover have "\x. x \ convex hull (J - {?z}) \ x \ rel_frontier C" by (metis True \J \ insert ?z I\ \K = convex hull I\ subsetD hull_mono subset_insert_iff that(4)) ultimately show "convex hull (J - {?z}) \ \ \ Pow (rel_frontier C)" by auto let ?F = "convex hull insert ?z (convex hull (J - {?z}))" have "F \ ?F" apply (clarsimp simp: \F = convex hull J\) by (metis True subsetD hull_mono hull_subset subset_insert_iff) moreover have "?F \ F" apply (clarsimp simp: \F = convex hull J\) by (metis (no_types, lifting) True convex_hull_eq_empty convex_hull_insert_segments hull_hull insert_Diff) ultimately show "F \ {?F}" by auto qed with \C\\\ show ?thesis by auto next 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 \ \\) then show "F \ \ \ ?\" by blast qed qed ultimately show ?thesis using that by auto qed have \
: "X \ Y face_of X \ X \ Y face_of Y" if XY: "X \ \" "Y \ ?\" for X Y proof - obtain C K where "C \ \" "K \ \" "K \ rel_frontier C" and Y: "Y = convex hull insert (SOME z. z \ rel_interior C) K" using XY by blast have "convex C" by (simp add: \C \ \\ convex\) have "K \ C" by (metis DiffE \C \ \\ \K \ rel_frontier C\ closed\ closure_closed rel_frontier_def subset_iff) let ?z = "(SOME z. z \ rel_interior C)" have z: "?z \ rel_interior C" using \C \ \\ in_rel_interior by blast obtain D where "D \ \" "X \ D" using C\ \X \ \\ by blast have "D \ rel_interior C = (C \ D) \ rel_interior C" using rel_interior_subset by blast also have "(C \ D) \ rel_interior C = {}" proof (rule face_of_disjoint_rel_interior) show "C \ D face_of C" using \_def \_def \C \ \\ \D \ \\ intface\ by blast show "C \ D \ C" by (metis (mono_tags, lifting) Int_lower2 \_def \_def \C \ \\ \D \ \\ aff_dim_subset mem_Collect_eq not_le) qed finally have DC: "D \ rel_interior C = {}" . have eq: "X \ convex hull (insert ?z K) = X \ convex hull K" proof (rule Int_convex_hull_insert_rel_exterior [OF \convex C\ \K \ C\ z]) show "disjnt X (rel_interior C)" using DC by (meson \X \ D\ disjnt_def disjnt_subset1) qed obtain I where I: "\ affine_dependent I" and Keq: "K = convex hull I" and [simp]: "convex hull K = K" using "*" \K \ \\ by force then have "?z \ affine hull I" 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" 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 by (simp add: XY(1) Y \K \ \\ eq faceI\) qed show "S \ S' face_of S" if "S \ \ \ ?\ \ S' \ \ \ ?\" for S S' using that proof (elim conjE UnE) fix X Y assume "X \ \" and "Y \ \" then show "X \ Y face_of X" by (simp add: faceI\) next fix X Y assume XY: "X \ \" "Y \ ?\" then show "X \ Y face_of X" "Y \ X face_of Y" using \
[OF XY] by (auto simp: Int_commute) next fix X Y assume XY: "X \ ?\" "Y \ ?\" show "X \ Y face_of X" proof - obtain C K D L where "C \ \" "K \ \" "K \ rel_frontier C" and X: "X = convex hull insert (SOME z. z \ rel_interior C) K" and "D \ \" "L \ \" "L \ rel_frontier D" and Y: "Y = convex hull insert (SOME z. z \ rel_interior D) L" using XY by blast let ?z = "(SOME z. z \ rel_interior C)" have z: "?z \ rel_interior C" using \C \ \\ in_rel_interior by blast have "convex C" by (simp add: \C \ \\ convex\) have "convex K" using "*" \K \ \\ by blast have "convex L" by (meson \L \ \\ convex_simplex simpl\) show ?thesis proof (cases "D=C") case True then have "L \ rel_frontier C" 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\) 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 case False have "convex D" by (simp add: \D \ \\ convex\) have "K \ C" by (metis DiffE \C \ \\ \K \ rel_frontier C\ closed\ closure_closed rel_frontier_def subset_eq) have "L \ D" by (metis DiffE \D \ \\ \L \ rel_frontier D\ closed\ closure_closed rel_frontier_def subset_eq) let ?w = "(SOME w. w \ rel_interior D)" have w: "?w \ rel_interior D" using \D \ \\ in_rel_interior by blast have "C \ rel_interior D = (D \ C) \ rel_interior D" using rel_interior_subset by blast also have "(D \ C) \ rel_interior D = {}" proof (rule face_of_disjoint_rel_interior) show "D \ C face_of D" using \_def \C \ \\ \D \ \\ intface\ by blast have "D \ \ \ aff_dim D = int n" using \_def \D \ \\ by blast moreover have "C \ \ \ aff_dim C = int n" using \_def \C \ \\ by blast ultimately show "D \ C \ D" by (metis Int_commute False face_of_aff_dim_lt inf.idem inf_le1 intface\ not_le poly\ polytope_imp_convex) 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) 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) = convex hull (insert ?z K) \ convex hull L" by (rule Int_convex_hull_insert_rel_exterior [OF \convex D\ \L \ D\ w]) have ch_id: "convex hull K = K" "convex hull L = L" using "*" \K \ \\ \L \ \\ hull_same by auto have "convex C" 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" 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) show "C \ D face_of C" using \_def \C \ \\ \D \ \\ intface\ by blast have "D \ \" "aff_dim D = int n" using \_def \D \ \\ by fastforce+ moreover have "C \ \" "aff_dim C = int n" using \_def \C \ \\ by fastforce+ ultimately have "aff_dim D + - 1 * aff_dim C \ 0" by fastforce then have "\ C face_of D" using False \convex D\ face_of_aff_dim_lt by fastforce show "C \ D \ C" by (metis inf_commute \C \ \\ \D \ \\ \\ C face_of D\ intface\) qed then have "D \ rel_interior C = {}" by (metis inf.absorb_iff2 inf_assoc inf_sup_aci(1) rel_interior_subset) then show "disjnt L (rel_interior C)" by (meson \L \ D\ disjnt_def disjnt_subset1) next show "L \ convex hull K = convex hull K \ L" by force qed 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" 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) 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" 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) qed finally have 2: "convex hull insert ?z K \ convex hull L face_of convex hull insert ?w L" . show ?thesis by (simp add: X Y eq 1 2) qed qed qed qed show "\F \ \ \ ?\. C = \F" if "C \ \" for C proof (cases "C \ \") case True then show ?thesis by (meson UnCI fin\ subsetD subsetI) next case False then have "C \ \" by (simp add: \_def \_def aff\ less_le that) let ?z = "SOME z. z \ rel_interior C" have z: "?z \ rel_interior C" using \C \ \\ in_rel_interior by blast let ?F = "\K \ \ \ Pow (rel_frontier C). {convex hull (insert ?z K)}" have "?F \ ?\" using \C \ \\ by blast moreover have "C \ \?F" proof fix x assume "x \ C" have "convex C" using \C \ \\ convex\ by blast have "bounded C" using \C \ \\ by (simp add: poly\ polytope_imp_bounded that) have "polytope C" using \C \ \\ poly\ by auto have "\ (?z = x \ C = {?z})" using \C \ \\ aff_dim_sing [of ?z] \\ n \ 1\ by (force simp: \_def) then obtain y where y: "y \ rel_frontier C" and xzy: "x \ closed_segment ?z y" and sub: "open_segment ?z y \ rel_interior C" by (blast intro: segment_to_rel_frontier [OF \convex C\ \bounded C\ z \x \ C\]) then obtain F where "y \ F" "F face_of C" "F \ C" by (auto simp: rel_frontier_of_polyhedron_alt [OF polytope_imp_polyhedron [OF \polytope C\]]) then obtain \ where "finite \" "\ \ \" "F = \\" by (metis (mono_tags, lifting) \_def \C \ \\ \convex C\ aff\ face\ face_of_aff_dim_lt fin\ le_less_trans mem_Collect_eq not_less) then obtain K where "y \ K" "K \ \" using \y \ F\ by blast moreover have x: "x \ convex hull {?z,y}" using segment_convex_hull xzy by auto moreover have "convex hull {?z,y} \ convex hull insert ?z K" by (metis (full_types) \y \ K\ hull_mono empty_subsetI insertCI insert_subset) moreover have "K \ \" using \K \ \\ \\ \ \\ by blast moreover have "K \ rel_frontier C" using \F = \\\ \F \ C\ \F face_of C\ \K \ \\ face_of_subset_rel_frontier by fastforce ultimately show "x \ \?F" by force qed moreover have "convex hull insert (SOME z. z \ rel_interior C) K \ C" if "K \ \" "K \ rel_frontier C" for K proof (rule hull_minimal) show "insert (SOME z. z \ rel_interior C) K \ C" using that \C \ \\ in_rel_interior rel_interior_subset by (force simp: closure_eq rel_frontier_def closed\) show "convex C" by (simp add: \C \ \\ convex\) qed then have "\?F \ C" by auto ultimately show ?thesis by blast qed have "(\C. C \ \ \ L \ C) \ aff_dim L \ int n" if "L \ \ \ ?\" for L using that proof assume "L \ \" then show ?thesis using C\ \_def "*" by fastforce next assume "L \ ?\" then obtain C K where "C \ \" and L: "L = convex hull insert (SOME z. z \ rel_interior C) K" and K: "K \ \" "K \ rel_frontier C" by auto then have "convex hull C = C" by (meson convex\ convex_hull_eq) then have "convex C" by (metis (no_types) convex_convex_hull) have "rel_frontier C \ C" by (metis DiffE closed\ \C \ \\ closure_closed rel_frontier_def subsetI) have "K \ C" using K \rel_frontier C \ C\ by blast have "C \ \" using \_def \C \ \\ by auto moreover have "L \ C" using K L \C \ \\ by (metis \K \ C\ \convex hull C = C\ contra_subsetD hull_mono in_rel_interior insert_subset rel_interior_subset) ultimately show ?thesis using \rel_frontier C \ C\ \L \ C\ aff\ aff_dim_subset \C \ \\ dual_order.trans by blast qed then show "\C. C \ \ \ L \ C" "aff_dim L \ int n" if "L \ \ \ ?\" for L using that by auto qed then show ?thesis apply (rule ex_forward, safe) apply (meson Union_iff subsetCE, fastforce) by (meson infinite_super simplicial_complex_def) qed qed lemma simplicial_subdivision_of_cell_complex_lowdim: assumes "finite \" and poly: "\C. C \ \ \ polytope C" and face: "\C1 C2. \C1 \ \; C2 \ \\ \ C1 \ C2 face_of C1" and aff: "\C. C \ \ \ aff_dim C \ d" obtains \ where "simplicial_complex \" "\K. K \ \ \ aff_dim K \ d" "\\ = \\" "\C. C \ \ \ \F. finite F \ F \ \ \ C = \F" "\K. K \ \ \ \C. C \ \ \ K \ C" proof (cases "d \ 0") case True then obtain n where n: "d = of_nat n" using zero_le_imp_eq_int by blast have "\\. simplicial_complex \ \ (\K\\. aff_dim K \ int n) \ \\ = \(\C\\. {F. F face_of C}) \ (\C\\C\\. {F. F face_of C}. \F. finite F \ F \ \ \ C = \F) \ (\K\\. \C. C \ (\C\\. {F. F face_of C}) \ K \ C)" proof (rule simplicial_subdivision_aux) show "finite (\C\\. {F. F face_of C})" using \finite \\ poly polyhedron_eq_finite_faces polytope_imp_polyhedron by fastforce show "polytope F" if "F \ (\C\\. {F. F face_of C})" for F using poly that face_of_polytope_polytope by blast show "aff_dim F \ int n" if "F \ (\C\\. {F. F face_of C})" for F using that by clarify (metis n aff_dim_subset aff face_of_imp_subset order_trans) show "F \ (\C\\. {F. F face_of C})" if "G \ (\C\\. {F. F face_of C})" and "F face_of G" for F G using that face_of_trans by blast next fix F1 F2 assume "F1 \ (\C\\. {F. F face_of C})" and "F2 \ (\C\\. {F. F face_of C})" then obtain C1 C2 where "C1 \ \" "C2 \ \" and F: "F1 face_of C1" "F2 face_of C2" by auto show "F1 \ F2 face_of F1" using face_of_Int_subface [OF _ _ F] by (metis \C1 \ \\ \C2 \ \\ face inf_commute) qed moreover have "\(\C\\. {F. F face_of C}) = \\" using face_of_imp_subset face by blast ultimately show ?thesis using face_of_imp_subset n by (fastforce intro!: that simp add: poly face_of_refl polytope_imp_convex) next case False then have m1: "\C. C \ \ \ aff_dim C = -1" by (metis aff aff_dim_empty_eq aff_dim_negative_iff dual_order.trans not_less) then have face\: "\F S. \S \ \; F face_of S\ \ F \ \" by (metis aff_dim_empty face_of_empty) show ?thesis proof have "\S. S \ \ \ \n. n simplex S" by (metis (no_types) m1 aff_dim_empty simplex_minus_1) then show "simplicial_complex \" by (auto simp: simplicial_complex_def \finite \\ face intro: face\) show "aff_dim K \ d" if "K \ \" for K by (simp add: that aff) show "\F. finite F \ F \ \ \ C = \F" if "C \ \" for C using \C \ \\ equals0I by auto show "\C. C \ \ \ K \ C" if "K \ \" for K using \K \ \\ by blast qed auto qed proposition simplicial_subdivision_of_cell_complex: assumes "finite \" and poly: "\C. C \ \ \ polytope C" and face: "\C1 C2. \C1 \ \; C2 \ \\ \ C1 \ C2 face_of C1" obtains \ where "simplicial_complex \" "\\ = \\" "\C. C \ \ \ \F. finite F \ F \ \ \ C = \F" "\K. K \ \ \ \C. C \ \ \ K \ C" by (blast intro: simplicial_subdivision_of_cell_complex_lowdim [OF assms aff_dim_le_DIM]) corollary fine_simplicial_subdivision_of_cell_complex: assumes "0 < e" "finite \" and poly: "\C. C \ \ \ polytope C" and face: "\C1 C2. \C1 \ \; C2 \ \\ \ C1 \ C2 face_of C1" obtains \ where "simplicial_complex \" "\K. K \ \ \ diameter K < e" "\\ = \\" "\C. C \ \ \ \F. finite F \ F \ \ \ C = \F" "\K. K \ \ \ \C. C \ \ \ K \ C" proof - obtain \ where \: "finite \" "\\ = \\" and diapoly: "\X. X \ \ \ diameter X < e" "\X. X \ \ \ polytope X" and "\X Y. \X \ \; Y \ \\ \ X \ Y face_of X" and \covers: "\C x. C \ \ \ x \ C \ \D. D \ \ \ x \ D \ D \ C" and \covered: "\C. C \ \ \ \D. D \ \ \ C \ D" by (blast intro: cell_complex_subdivision_exists [OF \0 < e\ \finite \\ poly aff_dim_le_DIM face]) then obtain \ where \: "simplicial_complex \" "\\ = \\" and \covers: "\C. C \ \ \ \F. finite F \ F \ \ \ C = \F" and \covered: "\K. K \ \ \ \C. C \ \ \ K \ C" using simplicial_subdivision_of_cell_complex [OF \finite \\] by metis show ?thesis proof show "simplicial_complex \" by (rule \) show "diameter K < e" if "K \ \" for K by (metis le_less_trans diapoly \covered diameter_subset polytope_imp_bounded that) show "\\ = \\" by (simp add: \(2) \\\ = \\\) show "\F. finite F \ F \ \ \ C = \F" if "C \ \" for C proof - { fix x assume "x \ C" then obtain D where "D \ \" "x \ D" "D \ C" using \covers \C \ \\ \covers by force then have "\X\\ \ Pow C. x \ X" using \D \ \\ \D \ C\ \x \ D\ by blast } moreover have "finite (\ \ Pow C)" using \simplicial_complex \\ simplicial_complex_def by auto ultimately show ?thesis by (rule_tac x="(\ \ Pow C)" in exI) auto qed show "\C. C \ \ \ K \ C" if "K \ \" for K by (meson \covered \covered order_trans that) qed qed subsection\Some results on cell division with full-dimensional cells only\ lemma convex_Union_fulldim_cells: assumes "finite \" and clo: "\C. C \ \ \ closed C" and con: "\C. C \ \ \ convex C" and eq: "\\ = U"and "convex U" shows "\{C \ \. aff_dim C = aff_dim U} = U" (is "?lhs = U") proof - have "closed U" using \finite \\ clo eq by blast have "?lhs \ U" using eq by blast moreover have "U \ ?lhs" proof (cases "\C \ \. aff_dim C = aff_dim U") case True then show ?thesis using eq by blast next case False have "closed ?lhs" by (simp add: \finite \\ clo closed_Union) moreover have "U \ closure ?lhs" proof - have "U \ closure(\{U - C |C. C \ \ \ aff_dim C < aff_dim U})" proof (rule Baire [OF \closed U\]) show "countable {U - C |C. C \ \ \ aff_dim C < aff_dim U}" using \finite \\ uncountable_infinite by fastforce have "\C. C \ \ \ openin (top_of_set U) (U-C)" by (metis Sup_upper clo closed_limpt closedin_limpt eq openin_diff openin_subtopology_self) then show "openin (top_of_set U) T \ U \ closure T" 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" 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" 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 then show ?thesis by (auto intro!: closure_mono) qed finally show ?thesis . qed ultimately show ?thesis using closure_subset_eq by blast qed ultimately show ?thesis by blast qed proposition fine_triangular_subdivision_of_cell_complex: assumes "0 < e" "finite \" and poly: "\C. C \ \ \ polytope C" and aff: "\C. C \ \ \ aff_dim C = d" and face: "\C1 C2. \C1 \ \; C2 \ \\ \ C1 \ C2 face_of C1" obtains \ where "triangulation \" "\k. k \ \ \ diameter k < e" "\k. k \ \ \ aff_dim k = d" "\\ = \\" "\C. C \ \ \ \f. finite f \ f \ \ \ C = \f" "\k. k \ \ \ \C. C \ \ \ k \ C" proof - obtain \ where "simplicial_complex \" and dia\: "\K. K \ \ \ diameter K < e" and "\\ = \\" and in\: "\C. C \ \ \ \F. finite F \ F \ \ \ C = \F" and in\: "\K. K \ \ \ \C. C \ \ \ K \ C" by (blast intro: fine_simplicial_subdivision_of_cell_complex [OF \e > 0\ \finite \\ poly face]) let ?\ = "{K \ \. aff_dim K = d}" show thesis proof show "triangulation ?\" using \simplicial_complex \\ by (auto simp: triangulation_def simplicial_complex_def) show "diameter L < e" if "L \ {K \ \. aff_dim K = d}" for L using that by (auto simp: dia\) show "aff_dim L = d" if "L \ {K \ \. aff_dim K = d}" for L using that by auto show "\F. finite F \ F \ {K \ \. aff_dim K = d} \ C = \F" if "C \ \" for C proof - obtain F where "finite F" "F \ \" "C = \F" using in\ [OF \C \ \\] by auto show ?thesis proof (intro exI conjI) show "finite {K \ F. aff_dim K = d}" by (simp add: \finite F\) show "{K \ F. aff_dim K = d} \ {K \ \. aff_dim K = d}" using \F \ \\ by blast have "d = aff_dim C" by (simp add: aff that) moreover have "\K. K \ F \ closed K \ convex K" using \simplicial_complex \\ \F \ \\ unfolding simplicial_complex_def by (metis subsetCE \F \ \\ closed_simplex convex_simplex) moreover have "convex (\F)" using \C = \F\ poly polytope_imp_convex that by blast ultimately show "C = \{K \ F. aff_dim K = d}" by (simp add: convex_Union_fulldim_cells \C = \F\ \finite F\) qed qed then show "\{K \ \. aff_dim K = d} = \\" by auto (meson in\ subsetCE) show "\C. C \ \ \ L \ C" if "L \ {K \ \. aff_dim K = d}" for L using that by (auto simp: in\) qed qed end diff --git a/src/HOL/Limits.thy b/src/HOL/Limits.thy --- a/src/HOL/Limits.thy +++ b/src/HOL/Limits.thy @@ -1,3079 +1,3078 @@ (* Title: HOL/Limits.thy Author: Brian Huffman Author: Jacques D. Fleuriot, University of Cambridge Author: Lawrence C Paulson Author: Jeremy Avigad *) section \Limits on Real Vector Spaces\ theory Limits imports Real_Vector_Spaces begin text \Lemmas related to shifting/scaling\ lemma range_add [simp]: fixes a::"'a::group_add" shows "range ((+) a) = UNIV" by (metis add_minus_cancel surjI) lemma range_diff [simp]: fixes a::"'a::group_add" shows "range ((-) a) = UNIV" by (metis (full_types) add_minus_cancel diff_minus_eq_add surj_def) lemma range_mult [simp]: fixes a::"real" shows "range ((*) a) = (if a=0 then {0} else UNIV)" by (simp add: surj_def) (meson dvdE dvd_field_iff) subsection \Filter going to infinity norm\ definition at_infinity :: "'a::real_normed_vector filter" where "at_infinity = (INF r. principal {x. r \ norm x})" lemma eventually_at_infinity: "eventually P at_infinity \ (\b. \x. b \ norm x \ P x)" unfolding at_infinity_def by (subst eventually_INF_base) (auto simp: subset_eq eventually_principal intro!: exI[of _ "max a b" for a b]) corollary eventually_at_infinity_pos: "eventually p at_infinity \ (\b. 0 < b \ (\x. norm x \ b \ p x))" unfolding eventually_at_infinity by (meson le_less_trans norm_ge_zero not_le zero_less_one) lemma at_infinity_eq_at_top_bot: "(at_infinity :: real filter) = sup at_top at_bot" proof - have 1: "\\n\u. A n; \n\v. A n\ \ \b. \x. b \ \x\ \ A x" for A and u v::real by (rule_tac x="max (- v) u" in exI) (auto simp: abs_real_def) have 2: "\x. u \ \x\ \ A x \ \N. \n\N. A n" for A and u::real by (meson abs_less_iff le_cases less_le_not_le) have 3: "\x. u \ \x\ \ A x \ \N. \n\N. A n" for A and u::real by (metis (full_types) abs_ge_self abs_minus_cancel le_minus_iff order_trans) show ?thesis by (auto simp: filter_eq_iff eventually_sup eventually_at_infinity eventually_at_top_linorder eventually_at_bot_linorder intro: 1 2 3) qed lemma at_top_le_at_infinity: "at_top \ (at_infinity :: real filter)" unfolding at_infinity_eq_at_top_bot by simp lemma at_bot_le_at_infinity: "at_bot \ (at_infinity :: real filter)" unfolding at_infinity_eq_at_top_bot by simp lemma filterlim_at_top_imp_at_infinity: "filterlim f at_top F \ filterlim f at_infinity F" for f :: "_ \ real" by (rule filterlim_mono[OF _ at_top_le_at_infinity order_refl]) lemma filterlim_real_at_infinity_sequentially: "filterlim real at_infinity sequentially" by (simp add: filterlim_at_top_imp_at_infinity filterlim_real_sequentially) lemma lim_infinity_imp_sequentially: "(f \ l) at_infinity \ ((\n. f(n)) \ l) sequentially" by (simp add: filterlim_at_top_imp_at_infinity filterlim_compose filterlim_real_sequentially) subsubsection \Boundedness\ definition Bfun :: "('a \ 'b::metric_space) \ 'a filter \ bool" where Bfun_metric_def: "Bfun f F = (\y. \K>0. eventually (\x. dist (f x) y \ K) F)" abbreviation Bseq :: "(nat \ 'a::metric_space) \ bool" where "Bseq X \ Bfun X sequentially" lemma Bseq_conv_Bfun: "Bseq X \ Bfun X sequentially" .. lemma Bseq_ignore_initial_segment: "Bseq X \ Bseq (\n. X (n + k))" unfolding Bfun_metric_def by (subst eventually_sequentially_seg) lemma Bseq_offset: "Bseq (\n. X (n + k)) \ Bseq X" unfolding Bfun_metric_def by (subst (asm) eventually_sequentially_seg) lemma Bfun_def: "Bfun f F \ (\K>0. eventually (\x. norm (f x) \ K) F)" unfolding Bfun_metric_def norm_conv_dist proof safe fix y K assume K: "0 < K" and *: "eventually (\x. dist (f x) y \ K) F" moreover have "eventually (\x. dist (f x) 0 \ dist (f x) y + dist 0 y) F" by (intro always_eventually) (metis dist_commute dist_triangle) with * have "eventually (\x. dist (f x) 0 \ K + dist 0 y) F" by eventually_elim auto with \0 < K\ show "\K>0. eventually (\x. dist (f x) 0 \ K) F" by (intro exI[of _ "K + dist 0 y"] add_pos_nonneg conjI zero_le_dist) auto qed (force simp del: norm_conv_dist [symmetric]) lemma BfunI: assumes K: "eventually (\x. norm (f x) \ K) F" shows "Bfun f F" unfolding Bfun_def proof (intro exI conjI allI) show "0 < max K 1" by simp show "eventually (\x. norm (f x) \ max K 1) F" using K by (rule eventually_mono) simp qed lemma BfunE: assumes "Bfun f F" obtains B where "0 < B" and "eventually (\x. norm (f x) \ B) F" using assms unfolding Bfun_def by blast lemma Cauchy_Bseq: assumes "Cauchy X" shows "Bseq X" proof - have "\y K. 0 < K \ (\N. \n\N. dist (X n) y \ K)" if "\m n. \m \ M; n \ M\ \ dist (X m) (X n) < 1" for M by (meson order.order_iff_strict that zero_less_one) with assms show ?thesis by (force simp: Cauchy_def Bfun_metric_def eventually_sequentially) qed subsubsection \Bounded Sequences\ lemma BseqI': "(\n. norm (X n) \ K) \ Bseq X" by (intro BfunI) (auto simp: eventually_sequentially) lemma Bseq_def: "Bseq X \ (\K>0. \n. norm (X n) \ K)" unfolding Bfun_def eventually_sequentially proof safe fix N K assume "0 < K" "\n\N. norm (X n) \ K" then show "\K>0. \n. norm (X n) \ K" by (intro exI[of _ "max (Max (norm ` X ` {..N})) K"] max.strict_coboundedI2) (auto intro!: imageI not_less[where 'a=nat, THEN iffD1] Max_ge simp: le_max_iff_disj) qed auto lemma BseqE: "Bseq X \ (\K. 0 < K \ \n. norm (X n) \ K \ Q) \ Q" unfolding Bseq_def by auto lemma BseqD: "Bseq X \ \K. 0 < K \ (\n. norm (X n) \ K)" by (simp add: Bseq_def) lemma BseqI: "0 < K \ \n. norm (X n) \ K \ Bseq X" by (auto simp: Bseq_def) lemma Bseq_bdd_above: "Bseq X \ bdd_above (range X)" for X :: "nat \ real" proof (elim BseqE, intro bdd_aboveI2) fix K n assume "0 < K" "\n. norm (X n) \ K" then show "X n \ K" by (auto elim!: allE[of _ n]) qed lemma Bseq_bdd_above': "Bseq X \ bdd_above (range (\n. norm (X n)))" for X :: "nat \ 'a :: real_normed_vector" proof (elim BseqE, intro bdd_aboveI2) fix K n assume "0 < K" "\n. norm (X n) \ K" then show "norm (X n) \ K" by (auto elim!: allE[of _ n]) qed lemma Bseq_bdd_below: "Bseq X \ bdd_below (range X)" for X :: "nat \ real" proof (elim BseqE, intro bdd_belowI2) fix K n assume "0 < K" "\n. norm (X n) \ K" then show "- K \ X n" by (auto elim!: allE[of _ n]) qed lemma Bseq_eventually_mono: assumes "eventually (\n. norm (f n) \ norm (g n)) sequentially" "Bseq g" shows "Bseq f" proof - from assms(2) obtain K where "0 < K" and "eventually (\n. norm (g n) \ K) sequentially" unfolding Bfun_def by fast with assms(1) have "eventually (\n. norm (f n) \ K) sequentially" by (fast elim: eventually_elim2 order_trans) with \0 < K\ show "Bseq f" unfolding Bfun_def by fast qed lemma lemma_NBseq_def: "(\K > 0. \n. norm (X n) \ K) \ (\N. \n. norm (X n) \ real(Suc N))" proof safe fix K :: real from reals_Archimedean2 obtain n :: nat where "K < real n" .. then have "K \ real (Suc n)" by auto moreover assume "\m. norm (X m) \ K" ultimately have "\m. norm (X m) \ real (Suc n)" by (blast intro: order_trans) then show "\N. \n. norm (X n) \ real (Suc N)" .. next show "\N. \n. norm (X n) \ real (Suc N) \ \K>0. \n. norm (X n) \ K" using of_nat_0_less_iff by blast qed text \Alternative definition for \Bseq\.\ lemma Bseq_iff: "Bseq X \ (\N. \n. norm (X n) \ real(Suc N))" by (simp add: Bseq_def) (simp add: lemma_NBseq_def) lemma lemma_NBseq_def2: "(\K > 0. \n. norm (X n) \ K) = (\N. \n. norm (X n) < real(Suc N))" proof - have *: "\N. \n. norm (X n) \ 1 + real N \ \N. \n. norm (X n) < 1 + real N" by (metis add.commute le_less_trans less_add_one of_nat_Suc) then show ?thesis unfolding lemma_NBseq_def by (metis less_le_not_le not_less_iff_gr_or_eq of_nat_Suc) qed text \Yet another definition for Bseq.\ lemma Bseq_iff1a: "Bseq X \ (\N. \n. norm (X n) < real (Suc N))" by (simp add: Bseq_def lemma_NBseq_def2) subsubsection \A Few More Equivalence Theorems for Boundedness\ text \Alternative formulation for boundedness.\ lemma Bseq_iff2: "Bseq X \ (\k > 0. \x. \n. norm (X n + - x) \ k)" by (metis BseqE BseqI' add.commute add_cancel_right_left add_uminus_conv_diff norm_add_leD norm_minus_cancel norm_minus_commute) text \Alternative formulation for boundedness.\ lemma Bseq_iff3: "Bseq X \ (\k>0. \N. \n. norm (X n + - X N) \ k)" (is "?P \ ?Q") proof assume ?P then obtain K where *: "0 < K" and **: "\n. norm (X n) \ K" by (auto simp: Bseq_def) from * have "0 < K + norm (X 0)" by (rule order_less_le_trans) simp from ** have "\n. norm (X n - X 0) \ K + norm (X 0)" by (auto intro: order_trans norm_triangle_ineq4) then have "\n. norm (X n + - X 0) \ K + norm (X 0)" by simp with \0 < K + norm (X 0)\ show ?Q by blast next assume ?Q then show ?P by (auto simp: Bseq_iff2) qed subsubsection \Upper Bounds and Lubs of Bounded Sequences\ lemma Bseq_minus_iff: "Bseq (\n. - (X n) :: 'a::real_normed_vector) \ Bseq X" by (simp add: Bseq_def) lemma Bseq_add: fixes f :: "nat \ 'a::real_normed_vector" assumes "Bseq f" shows "Bseq (\x. f x + c)" proof - from assms obtain K where K: "\x. norm (f x) \ K" unfolding Bseq_def by blast { fix x :: nat have "norm (f x + c) \ norm (f x) + norm c" by (rule norm_triangle_ineq) also have "norm (f x) \ K" by (rule K) finally have "norm (f x + c) \ K + norm c" by simp } then show ?thesis by (rule BseqI') qed lemma Bseq_add_iff: "Bseq (\x. f x + c) \ Bseq f" for f :: "nat \ 'a::real_normed_vector" using Bseq_add[of f c] Bseq_add[of "\x. f x + c" "-c"] by auto lemma Bseq_mult: fixes f g :: "nat \ 'a::real_normed_field" assumes "Bseq f" and "Bseq g" shows "Bseq (\x. f x * g x)" proof - from assms obtain K1 K2 where K: "norm (f x) \ K1" "K1 > 0" "norm (g x) \ K2" "K2 > 0" for x unfolding Bseq_def by blast then have "norm (f x * g x) \ K1 * K2" for x by (auto simp: norm_mult intro!: mult_mono) then show ?thesis by (rule BseqI') qed lemma Bfun_const [simp]: "Bfun (\_. c) F" unfolding Bfun_metric_def by (auto intro!: exI[of _ c] exI[of _ "1::real"]) lemma Bseq_cmult_iff: fixes c :: "'a::real_normed_field" assumes "c \ 0" shows "Bseq (\x. c * f x) \ Bseq f" proof assume "Bseq (\x. c * f x)" with Bfun_const have "Bseq (\x. inverse c * (c * f x))" by (rule Bseq_mult) with \c \ 0\ show "Bseq f" by (simp add: field_split_simps) qed (intro Bseq_mult Bfun_const) lemma Bseq_subseq: "Bseq f \ Bseq (\x. f (g x))" for f :: "nat \ 'a::real_normed_vector" unfolding Bseq_def by auto lemma Bseq_Suc_iff: "Bseq (\n. f (Suc n)) \ Bseq f" for f :: "nat \ 'a::real_normed_vector" using Bseq_offset[of f 1] by (auto intro: Bseq_subseq) lemma increasing_Bseq_subseq_iff: assumes "\x y. x \ y \ norm (f x :: 'a::real_normed_vector) \ norm (f y)" "strict_mono g" shows "Bseq (\x. f (g x)) \ Bseq f" proof assume "Bseq (\x. f (g x))" then obtain K where K: "\x. norm (f (g x)) \ K" unfolding Bseq_def by auto { fix x :: nat from filterlim_subseq[OF assms(2)] obtain y where "g y \ x" by (auto simp: filterlim_at_top eventually_at_top_linorder) then have "norm (f x) \ norm (f (g y))" using assms(1) by blast also have "norm (f (g y)) \ K" by (rule K) finally have "norm (f x) \ K" . } then show "Bseq f" by (rule BseqI') qed (use Bseq_subseq[of f g] in simp_all) lemma nonneg_incseq_Bseq_subseq_iff: fixes f :: "nat \ real" and g :: "nat \ nat" assumes "\x. f x \ 0" "incseq f" "strict_mono g" shows "Bseq (\x. f (g x)) \ Bseq f" using assms by (intro increasing_Bseq_subseq_iff) (auto simp: incseq_def) lemma Bseq_eq_bounded: "range f \ {a..b} \ Bseq f" for a b :: real proof (rule BseqI'[where K="max (norm a) (norm b)"]) fix n assume "range f \ {a..b}" then have "f n \ {a..b}" by blast then show "norm (f n) \ max (norm a) (norm b)" by auto qed lemma incseq_bounded: "incseq X \ \i. X i \ B \ Bseq X" for B :: real by (intro Bseq_eq_bounded[of X "X 0" B]) (auto simp: incseq_def) lemma decseq_bounded: "decseq X \ \i. B \ X i \ Bseq X" for B :: real by (intro Bseq_eq_bounded[of X B "X 0"]) (auto simp: decseq_def) subsubsection\<^marker>\tag unimportant\ \Polynomal function extremal theorem, from HOL Light\ lemma polyfun_extremal_lemma: fixes c :: "nat \ 'a::real_normed_div_algebra" assumes "0 < e" shows "\M. \z. M \ norm(z) \ norm (\i\n. c(i) * z^i) \ e * norm(z) ^ (Suc n)" proof (induct n) case 0 with assms show ?case apply (rule_tac x="norm (c 0) / e" in exI) apply (auto simp: field_simps) done next case (Suc n) obtain M where M: "\z. M \ norm z \ norm (\i\n. c i * z^i) \ e * norm z ^ Suc n" using Suc assms by blast show ?case proof (rule exI [where x= "max M (1 + norm(c(Suc n)) / e)"], clarsimp simp del: power_Suc) fix z::'a assume z1: "M \ norm z" and "1 + norm (c (Suc n)) / e \ norm z" then have z2: "e + norm (c (Suc n)) \ e * norm z" using assms by (simp add: field_simps) have "norm (\i\n. c i * z^i) \ e * norm z ^ Suc n" using M [OF z1] by simp then have "norm (\i\n. c i * z^i) + norm (c (Suc n) * z ^ Suc n) \ e * norm z ^ Suc n + norm (c (Suc n) * z ^ Suc n)" by simp then have "norm ((\i\n. c i * z^i) + c (Suc n) * z ^ Suc n) \ e * norm z ^ Suc n + norm (c (Suc n) * z ^ Suc n)" by (blast intro: norm_triangle_le elim: ) also have "... \ (e + norm (c (Suc n))) * norm z ^ Suc n" by (simp add: norm_power norm_mult algebra_simps) also have "... \ (e * norm z) * norm z ^ Suc n" by (metis z2 mult.commute mult_left_mono norm_ge_zero norm_power) finally show "norm ((\i\n. c i * z^i) + c (Suc n) * z ^ Suc n) \ e * norm z ^ Suc (Suc n)" by simp qed qed lemma polyfun_extremal: (*COMPLEX_POLYFUN_EXTREMAL in HOL Light*) fixes c :: "nat \ 'a::real_normed_div_algebra" assumes k: "c k \ 0" "1\k" and kn: "k\n" shows "eventually (\z. norm (\i\n. c(i) * z^i) \ B) at_infinity" using kn proof (induction n) case 0 then show ?case using k by simp next case (Suc m) show ?case proof (cases "c (Suc m) = 0") case True then show ?thesis using Suc k by auto (metis antisym_conv less_eq_Suc_le not_le) next case False then obtain M where M: "\z. M \ norm z \ norm (\i\m. c i * z^i) \ norm (c (Suc m)) / 2 * norm z ^ Suc m" using polyfun_extremal_lemma [of "norm(c (Suc m)) / 2" c m] Suc by auto have "\b. \z. b \ norm z \ B \ norm (\i\Suc m. c i * z^i)" proof (rule exI [where x="max M (max 1 (\B\ / (norm(c (Suc m)) / 2)))"], clarsimp simp del: power_Suc) fix z::'a assume z1: "M \ norm z" "1 \ norm z" and "\B\ * 2 / norm (c (Suc m)) \ norm z" then have z2: "\B\ \ norm (c (Suc m)) * norm z / 2" using False by (simp add: field_simps) have nz: "norm z \ norm z ^ Suc m" by (metis \1 \ norm z\ One_nat_def less_eq_Suc_le power_increasing power_one_right zero_less_Suc) have *: "\y x. norm (c (Suc m)) * norm z / 2 \ norm y - norm x \ B \ norm (x + y)" by (metis abs_le_iff add.commute norm_diff_ineq order_trans z2) have "norm z * norm (c (Suc m)) + 2 * norm (\i\m. c i * z^i) \ norm (c (Suc m)) * norm z + norm (c (Suc m)) * norm z ^ Suc m" using M [of z] Suc z1 by auto also have "... \ 2 * (norm (c (Suc m)) * norm z ^ Suc m)" using nz by (simp add: mult_mono del: power_Suc) finally show "B \ norm ((\i\m. c i * z^i) + c (Suc m) * z ^ Suc m)" using Suc.IH apply (auto simp: eventually_at_infinity) apply (rule *) apply (simp add: field_simps norm_mult norm_power) done qed then show ?thesis by (simp add: eventually_at_infinity) qed qed subsection \Convergence to Zero\ definition Zfun :: "('a \ 'b::real_normed_vector) \ 'a filter \ bool" where "Zfun f F = (\r>0. eventually (\x. norm (f x) < r) F)" lemma ZfunI: "(\r. 0 < r \ eventually (\x. norm (f x) < r) F) \ Zfun f F" by (simp add: Zfun_def) lemma ZfunD: "Zfun f F \ 0 < r \ eventually (\x. norm (f x) < r) F" by (simp add: Zfun_def) lemma Zfun_ssubst: "eventually (\x. f x = g x) F \ Zfun g F \ Zfun f F" unfolding Zfun_def by (auto elim!: eventually_rev_mp) lemma Zfun_zero: "Zfun (\x. 0) F" unfolding Zfun_def by simp lemma Zfun_norm_iff: "Zfun (\x. norm (f x)) F = Zfun (\x. f x) F" unfolding Zfun_def by simp lemma Zfun_imp_Zfun: assumes f: "Zfun f F" and g: "eventually (\x. norm (g x) \ norm (f x) * K) F" shows "Zfun (\x. g x) F" proof (cases "0 < K") case K: True show ?thesis proof (rule ZfunI) fix r :: real assume "0 < r" then have "0 < r / K" using K by simp then have "eventually (\x. norm (f x) < r / K) F" using ZfunD [OF f] by blast with g show "eventually (\x. norm (g x) < r) F" proof eventually_elim case (elim x) then have "norm (f x) * K < r" by (simp add: pos_less_divide_eq K) then show ?case by (simp add: order_le_less_trans [OF elim(1)]) qed qed next case False then have K: "K \ 0" by (simp only: not_less) show ?thesis proof (rule ZfunI) fix r :: real assume "0 < r" from g show "eventually (\x. norm (g x) < r) F" proof eventually_elim case (elim x) also have "norm (f x) * K \ norm (f x) * 0" using K norm_ge_zero by (rule mult_left_mono) finally show ?case using \0 < r\ by simp qed qed qed lemma Zfun_le: "Zfun g F \ \x. norm (f x) \ norm (g x) \ Zfun f F" by (erule Zfun_imp_Zfun [where K = 1]) simp lemma Zfun_add: assumes f: "Zfun f F" and g: "Zfun g F" shows "Zfun (\x. f x + g x) F" proof (rule ZfunI) fix r :: real assume "0 < r" then have r: "0 < r / 2" by simp have "eventually (\x. norm (f x) < r/2) F" using f r by (rule ZfunD) moreover have "eventually (\x. norm (g x) < r/2) F" using g r by (rule ZfunD) ultimately show "eventually (\x. norm (f x + g x) < r) F" proof eventually_elim case (elim x) have "norm (f x + g x) \ norm (f x) + norm (g x)" by (rule norm_triangle_ineq) also have "\ < r/2 + r/2" using elim by (rule add_strict_mono) finally show ?case by simp qed qed lemma Zfun_minus: "Zfun f F \ Zfun (\x. - f x) F" unfolding Zfun_def by simp lemma Zfun_diff: "Zfun f F \ Zfun g F \ Zfun (\x. f x - g x) F" using Zfun_add [of f F "\x. - g x"] by (simp add: Zfun_minus) lemma (in bounded_linear) Zfun: assumes g: "Zfun g F" shows "Zfun (\x. f (g x)) F" proof - obtain K where "norm (f x) \ norm x * K" for x using bounded by blast then have "eventually (\x. norm (f (g x)) \ norm (g x) * K) F" by simp with g show ?thesis by (rule Zfun_imp_Zfun) qed lemma (in bounded_bilinear) Zfun: assumes f: "Zfun f F" and g: "Zfun g F" shows "Zfun (\x. f x ** g x) F" proof (rule ZfunI) fix r :: real assume r: "0 < r" obtain K where K: "0 < K" and norm_le: "norm (x ** y) \ norm x * norm y * K" for x y using pos_bounded by blast from K have K': "0 < inverse K" by (rule positive_imp_inverse_positive) have "eventually (\x. norm (f x) < r) F" using f r by (rule ZfunD) moreover have "eventually (\x. norm (g x) < inverse K) F" using g K' by (rule ZfunD) ultimately show "eventually (\x. norm (f x ** g x) < r) F" proof eventually_elim case (elim x) have "norm (f x ** g x) \ norm (f x) * norm (g x) * K" by (rule norm_le) also have "norm (f x) * norm (g x) * K < r * inverse K * K" by (intro mult_strict_right_mono mult_strict_mono' norm_ge_zero elim K) also from K have "r * inverse K * K = r" by simp finally show ?case . qed qed lemma (in bounded_bilinear) Zfun_left: "Zfun f F \ Zfun (\x. f x ** a) F" by (rule bounded_linear_left [THEN bounded_linear.Zfun]) lemma (in bounded_bilinear) Zfun_right: "Zfun f F \ Zfun (\x. a ** f x) F" by (rule bounded_linear_right [THEN bounded_linear.Zfun]) lemmas Zfun_mult = bounded_bilinear.Zfun [OF bounded_bilinear_mult] lemmas Zfun_mult_right = bounded_bilinear.Zfun_right [OF bounded_bilinear_mult] lemmas Zfun_mult_left = bounded_bilinear.Zfun_left [OF bounded_bilinear_mult] lemma tendsto_Zfun_iff: "(f \ a) F = Zfun (\x. f x - a) F" by (simp only: tendsto_iff Zfun_def dist_norm) lemma tendsto_0_le: "(f \ 0) F \ eventually (\x. norm (g x) \ norm (f x) * K) F \ (g \ 0) F" by (simp add: Zfun_imp_Zfun tendsto_Zfun_iff) subsubsection \Distance and norms\ lemma tendsto_dist [tendsto_intros]: fixes l m :: "'a::metric_space" assumes f: "(f \ l) F" and g: "(g \ m) F" shows "((\x. dist (f x) (g x)) \ dist l m) F" proof (rule tendstoI) fix e :: real assume "0 < e" then have e2: "0 < e/2" by simp from tendstoD [OF f e2] tendstoD [OF g e2] show "eventually (\x. dist (dist (f x) (g x)) (dist l m) < e) F" proof (eventually_elim) case (elim x) then show "dist (dist (f x) (g x)) (dist l m) < e" unfolding dist_real_def using dist_triangle2 [of "f x" "g x" "l"] and dist_triangle2 [of "g x" "l" "m"] and dist_triangle3 [of "l" "m" "f x"] and dist_triangle [of "f x" "m" "g x"] by arith qed qed lemma continuous_dist[continuous_intros]: fixes f g :: "_ \ 'a :: metric_space" shows "continuous F f \ continuous F g \ continuous F (\x. dist (f x) (g x))" unfolding continuous_def by (rule tendsto_dist) lemma continuous_on_dist[continuous_intros]: fixes f g :: "_ \ 'a :: metric_space" shows "continuous_on s f \ continuous_on s g \ continuous_on s (\x. dist (f x) (g x))" unfolding continuous_on_def by (auto intro: tendsto_dist) lemma continuous_at_dist: "isCont (dist a) b" using continuous_on_dist [OF continuous_on_const continuous_on_id] continuous_on_eq_continuous_within by blast lemma tendsto_norm [tendsto_intros]: "(f \ a) F \ ((\x. norm (f x)) \ norm a) F" unfolding norm_conv_dist by (intro tendsto_intros) lemma continuous_norm [continuous_intros]: "continuous F f \ continuous F (\x. norm (f x))" unfolding continuous_def by (rule tendsto_norm) lemma continuous_on_norm [continuous_intros]: "continuous_on s f \ continuous_on s (\x. norm (f x))" unfolding continuous_on_def by (auto intro: tendsto_norm) lemma continuous_on_norm_id [continuous_intros]: "continuous_on S norm" by (intro continuous_on_id continuous_on_norm) lemma tendsto_norm_zero: "(f \ 0) F \ ((\x. norm (f x)) \ 0) F" by (drule tendsto_norm) simp lemma tendsto_norm_zero_cancel: "((\x. norm (f x)) \ 0) F \ (f \ 0) F" unfolding tendsto_iff dist_norm by simp lemma tendsto_norm_zero_iff: "((\x. norm (f x)) \ 0) F \ (f \ 0) F" unfolding tendsto_iff dist_norm by simp lemma tendsto_rabs [tendsto_intros]: "(f \ l) F \ ((\x. \f x\) \ \l\) F" for l :: real by (fold real_norm_def) (rule tendsto_norm) lemma continuous_rabs [continuous_intros]: "continuous F f \ continuous F (\x. \f x :: real\)" unfolding real_norm_def[symmetric] by (rule continuous_norm) lemma continuous_on_rabs [continuous_intros]: "continuous_on s f \ continuous_on s (\x. \f x :: real\)" unfolding real_norm_def[symmetric] by (rule continuous_on_norm) lemma tendsto_rabs_zero: "(f \ (0::real)) F \ ((\x. \f x\) \ 0) F" by (fold real_norm_def) (rule tendsto_norm_zero) lemma tendsto_rabs_zero_cancel: "((\x. \f x\) \ (0::real)) F \ (f \ 0) F" by (fold real_norm_def) (rule tendsto_norm_zero_cancel) lemma tendsto_rabs_zero_iff: "((\x. \f x\) \ (0::real)) F \ (f \ 0) F" by (fold real_norm_def) (rule tendsto_norm_zero_iff) subsection \Topological Monoid\ class topological_monoid_add = topological_space + monoid_add + assumes tendsto_add_Pair: "LIM x (nhds a \\<^sub>F nhds b). fst x + snd x :> nhds (a + b)" class topological_comm_monoid_add = topological_monoid_add + comm_monoid_add lemma tendsto_add [tendsto_intros]: fixes a b :: "'a::topological_monoid_add" shows "(f \ a) F \ (g \ b) F \ ((\x. f x + g x) \ a + b) F" using filterlim_compose[OF tendsto_add_Pair, of "\x. (f x, g x)" a b F] by (simp add: nhds_prod[symmetric] tendsto_Pair) lemma continuous_add [continuous_intros]: fixes f g :: "_ \ 'b::topological_monoid_add" shows "continuous F f \ continuous F g \ continuous F (\x. f x + g x)" unfolding continuous_def by (rule tendsto_add) lemma continuous_on_add [continuous_intros]: fixes f g :: "_ \ 'b::topological_monoid_add" shows "continuous_on s f \ continuous_on s g \ continuous_on s (\x. f x + g x)" unfolding continuous_on_def by (auto intro: tendsto_add) lemma tendsto_add_zero: fixes f g :: "_ \ 'b::topological_monoid_add" shows "(f \ 0) F \ (g \ 0) F \ ((\x. f x + g x) \ 0) F" by (drule (1) tendsto_add) simp lemma tendsto_sum [tendsto_intros]: fixes f :: "'a \ 'b \ 'c::topological_comm_monoid_add" shows "(\i. i \ I \ (f i \ a i) F) \ ((\x. \i\I. f i x) \ (\i\I. a i)) F" by (induct I rule: infinite_finite_induct) (simp_all add: tendsto_add) lemma tendsto_null_sum: fixes f :: "'a \ 'b \ 'c::topological_comm_monoid_add" assumes "\i. i \ I \ ((\x. f x i) \ 0) F" shows "((\i. sum (f i) I) \ 0) F" using tendsto_sum [of I "\x y. f y x" "\x. 0"] assms by simp lemma continuous_sum [continuous_intros]: fixes f :: "'a \ 'b::t2_space \ 'c::topological_comm_monoid_add" shows "(\i. i \ I \ continuous F (f i)) \ continuous F (\x. \i\I. f i x)" unfolding continuous_def by (rule tendsto_sum) lemma continuous_on_sum [continuous_intros]: fixes f :: "'a \ 'b::topological_space \ 'c::topological_comm_monoid_add" shows "(\i. i \ I \ continuous_on S (f i)) \ continuous_on S (\x. \i\I. f i x)" unfolding continuous_on_def by (auto intro: tendsto_sum) instance nat :: topological_comm_monoid_add by standard (simp add: nhds_discrete principal_prod_principal filterlim_principal eventually_principal) instance int :: topological_comm_monoid_add by standard (simp add: nhds_discrete principal_prod_principal filterlim_principal eventually_principal) subsubsection \Topological group\ class topological_group_add = topological_monoid_add + group_add + assumes tendsto_uminus_nhds: "(uminus \ - a) (nhds a)" begin lemma tendsto_minus [tendsto_intros]: "(f \ a) F \ ((\x. - f x) \ - a) F" by (rule filterlim_compose[OF tendsto_uminus_nhds]) end class topological_ab_group_add = topological_group_add + ab_group_add instance topological_ab_group_add < topological_comm_monoid_add .. lemma continuous_minus [continuous_intros]: "continuous F f \ continuous F (\x. - f x)" for f :: "'a::t2_space \ 'b::topological_group_add" unfolding continuous_def by (rule tendsto_minus) lemma continuous_on_minus [continuous_intros]: "continuous_on s f \ continuous_on s (\x. - f x)" for f :: "_ \ 'b::topological_group_add" unfolding continuous_on_def by (auto intro: tendsto_minus) lemma tendsto_minus_cancel: "((\x. - f x) \ - a) F \ (f \ a) F" for a :: "'a::topological_group_add" by (drule tendsto_minus) simp lemma tendsto_minus_cancel_left: "(f \ - (y::_::topological_group_add)) F \ ((\x. - f x) \ y) F" using tendsto_minus_cancel[of f "- y" F] tendsto_minus[of f "- y" F] by auto lemma tendsto_diff [tendsto_intros]: fixes a b :: "'a::topological_group_add" shows "(f \ a) F \ (g \ b) F \ ((\x. f x - g x) \ a - b) F" using tendsto_add [of f a F "\x. - g x" "- b"] by (simp add: tendsto_minus) lemma continuous_diff [continuous_intros]: fixes f g :: "'a::t2_space \ 'b::topological_group_add" shows "continuous F f \ continuous F g \ continuous F (\x. f x - g x)" unfolding continuous_def by (rule tendsto_diff) lemma continuous_on_diff [continuous_intros]: fixes f g :: "_ \ 'b::topological_group_add" shows "continuous_on s f \ continuous_on s g \ continuous_on s (\x. f x - g x)" unfolding continuous_on_def by (auto intro: tendsto_diff) lemma continuous_on_op_minus: "continuous_on (s::'a::topological_group_add set) ((-) x)" by (rule continuous_intros | simp)+ instance real_normed_vector < topological_ab_group_add proof fix a b :: 'a show "((\x. fst x + snd x) \ a + b) (nhds a \\<^sub>F nhds b)" unfolding tendsto_Zfun_iff add_diff_add using tendsto_fst[OF filterlim_ident, of "(a,b)"] tendsto_snd[OF filterlim_ident, of "(a,b)"] by (intro Zfun_add) (auto simp: tendsto_Zfun_iff[symmetric] nhds_prod[symmetric] intro!: tendsto_fst) show "(uminus \ - a) (nhds a)" unfolding tendsto_Zfun_iff minus_diff_minus using filterlim_ident[of "nhds a"] by (intro Zfun_minus) (simp add: tendsto_Zfun_iff) qed lemmas real_tendsto_sandwich = tendsto_sandwich[where 'a=real] subsubsection \Linear operators and multiplication\ lemma linear_times [simp]: "linear (\x. c * x)" for c :: "'a::real_algebra" by (auto simp: linearI distrib_left) lemma (in bounded_linear) tendsto: "(g \ a) F \ ((\x. f (g x)) \ f a) F" by (simp only: tendsto_Zfun_iff diff [symmetric] Zfun) lemma (in bounded_linear) continuous: "continuous F g \ continuous F (\x. f (g x))" using tendsto[of g _ F] by (auto simp: continuous_def) lemma (in bounded_linear) continuous_on: "continuous_on s g \ continuous_on s (\x. f (g x))" using tendsto[of g] by (auto simp: continuous_on_def) lemma (in bounded_linear) tendsto_zero: "(g \ 0) F \ ((\x. f (g x)) \ 0) F" by (drule tendsto) (simp only: zero) lemma (in bounded_bilinear) tendsto: "(f \ a) F \ (g \ b) F \ ((\x. f x ** g x) \ a ** b) F" by (simp only: tendsto_Zfun_iff prod_diff_prod Zfun_add Zfun Zfun_left Zfun_right) lemma (in bounded_bilinear) continuous: "continuous F f \ continuous F g \ continuous F (\x. f x ** g x)" using tendsto[of f _ F g] by (auto simp: continuous_def) lemma (in bounded_bilinear) continuous_on: "continuous_on s f \ continuous_on s g \ continuous_on s (\x. f x ** g x)" using tendsto[of f _ _ g] by (auto simp: continuous_on_def) lemma (in bounded_bilinear) tendsto_zero: assumes f: "(f \ 0) F" and g: "(g \ 0) F" shows "((\x. f x ** g x) \ 0) F" using tendsto [OF f g] by (simp add: zero_left) lemma (in bounded_bilinear) tendsto_left_zero: "(f \ 0) F \ ((\x. f x ** c) \ 0) F" by (rule bounded_linear.tendsto_zero [OF bounded_linear_left]) lemma (in bounded_bilinear) tendsto_right_zero: "(f \ 0) F \ ((\x. c ** f x) \ 0) F" by (rule bounded_linear.tendsto_zero [OF bounded_linear_right]) lemmas tendsto_of_real [tendsto_intros] = bounded_linear.tendsto [OF bounded_linear_of_real] lemmas tendsto_scaleR [tendsto_intros] = bounded_bilinear.tendsto [OF bounded_bilinear_scaleR] text\Analogous type class for multiplication\ class topological_semigroup_mult = topological_space + semigroup_mult + assumes tendsto_mult_Pair: "LIM x (nhds a \\<^sub>F nhds b). fst x * snd x :> nhds (a * b)" instance real_normed_algebra < topological_semigroup_mult proof fix a b :: 'a show "((\x. fst x * snd x) \ a * b) (nhds a \\<^sub>F nhds b)" unfolding nhds_prod[symmetric] using tendsto_fst[OF filterlim_ident, of "(a,b)"] tendsto_snd[OF filterlim_ident, of "(a,b)"] by (simp add: bounded_bilinear.tendsto [OF bounded_bilinear_mult]) qed lemma tendsto_mult [tendsto_intros]: fixes a b :: "'a::topological_semigroup_mult" shows "(f \ a) F \ (g \ b) F \ ((\x. f x * g x) \ a * b) F" using filterlim_compose[OF tendsto_mult_Pair, of "\x. (f x, g x)" a b F] by (simp add: nhds_prod[symmetric] tendsto_Pair) lemma tendsto_mult_left: "(f \ l) F \ ((\x. c * (f x)) \ c * l) F" for c :: "'a::topological_semigroup_mult" by (rule tendsto_mult [OF tendsto_const]) lemma tendsto_mult_right: "(f \ l) F \ ((\x. (f x) * c) \ l * c) F" for c :: "'a::topological_semigroup_mult" by (rule tendsto_mult [OF _ tendsto_const]) lemma tendsto_mult_left_iff [simp]: "c \ 0 \ tendsto(\x. c * f x) (c * l) F \ tendsto f l F" for c :: "'a::{topological_semigroup_mult,field}" by (auto simp: tendsto_mult_left dest: tendsto_mult_left [where c = "1/c"]) lemma tendsto_mult_right_iff [simp]: "c \ 0 \ tendsto(\x. f x * c) (l * c) F \ tendsto f l F" for c :: "'a::{topological_semigroup_mult,field}" by (auto simp: tendsto_mult_right dest: tendsto_mult_left [where c = "1/c"]) lemma tendsto_zero_mult_left_iff [simp]: fixes c::"'a::{topological_semigroup_mult,field}" assumes "c \ 0" shows "(\n. c * a n)\ 0 \ a \ 0" using assms tendsto_mult_left tendsto_mult_left_iff by fastforce lemma tendsto_zero_mult_right_iff [simp]: fixes c::"'a::{topological_semigroup_mult,field}" assumes "c \ 0" shows "(\n. a n * c)\ 0 \ a \ 0" using assms tendsto_mult_right tendsto_mult_right_iff by fastforce lemma tendsto_zero_divide_iff [simp]: fixes c::"'a::{topological_semigroup_mult,field}" assumes "c \ 0" shows "(\n. a n / c)\ 0 \ a \ 0" using tendsto_zero_mult_right_iff [of "1/c" a] assms by (simp add: field_simps) lemma lim_const_over_n [tendsto_intros]: fixes a :: "'a::real_normed_field" shows "(\n. a / of_nat n) \ 0" using tendsto_mult [OF tendsto_const [of a] lim_1_over_n] by simp lemmas continuous_of_real [continuous_intros] = bounded_linear.continuous [OF bounded_linear_of_real] lemmas continuous_scaleR [continuous_intros] = bounded_bilinear.continuous [OF bounded_bilinear_scaleR] lemmas continuous_mult [continuous_intros] = bounded_bilinear.continuous [OF bounded_bilinear_mult] lemmas continuous_on_of_real [continuous_intros] = bounded_linear.continuous_on [OF bounded_linear_of_real] lemmas continuous_on_scaleR [continuous_intros] = bounded_bilinear.continuous_on [OF bounded_bilinear_scaleR] lemmas continuous_on_mult [continuous_intros] = bounded_bilinear.continuous_on [OF bounded_bilinear_mult] lemmas tendsto_mult_zero = bounded_bilinear.tendsto_zero [OF bounded_bilinear_mult] lemmas tendsto_mult_left_zero = bounded_bilinear.tendsto_left_zero [OF bounded_bilinear_mult] lemmas tendsto_mult_right_zero = bounded_bilinear.tendsto_right_zero [OF bounded_bilinear_mult] lemma continuous_mult_left: fixes c::"'a::real_normed_algebra" shows "continuous F f \ continuous F (\x. c * f x)" by (rule continuous_mult [OF continuous_const]) lemma continuous_mult_right: fixes c::"'a::real_normed_algebra" shows "continuous F f \ continuous F (\x. f x * c)" by (rule continuous_mult [OF _ continuous_const]) lemma continuous_on_mult_left: fixes c::"'a::real_normed_algebra" shows "continuous_on s f \ continuous_on s (\x. c * f x)" by (rule continuous_on_mult [OF continuous_on_const]) lemma continuous_on_mult_right: fixes c::"'a::real_normed_algebra" shows "continuous_on s f \ continuous_on s (\x. f x * c)" by (rule continuous_on_mult [OF _ continuous_on_const]) lemma continuous_on_mult_const [simp]: fixes c::"'a::real_normed_algebra" shows "continuous_on s ((*) c)" by (intro continuous_on_mult_left continuous_on_id) lemma tendsto_divide_zero: fixes c :: "'a::real_normed_field" shows "(f \ 0) F \ ((\x. f x / c) \ 0) F" by (cases "c=0") (simp_all add: divide_inverse tendsto_mult_left_zero) lemma tendsto_power [tendsto_intros]: "(f \ a) F \ ((\x. f x ^ n) \ a ^ n) F" for f :: "'a \ 'b::{power,real_normed_algebra}" by (induct n) (simp_all add: tendsto_mult) lemma tendsto_null_power: "\(f \ 0) F; 0 < n\ \ ((\x. f x ^ n) \ 0) F" for f :: "'a \ 'b::{power,real_normed_algebra_1}" using tendsto_power [of f 0 F n] by (simp add: power_0_left) lemma continuous_power [continuous_intros]: "continuous F f \ continuous F (\x. (f x)^n)" for f :: "'a::t2_space \ 'b::{power,real_normed_algebra}" unfolding continuous_def by (rule tendsto_power) lemma continuous_on_power [continuous_intros]: fixes f :: "_ \ 'b::{power,real_normed_algebra}" shows "continuous_on s f \ continuous_on s (\x. (f x)^n)" unfolding continuous_on_def by (auto intro: tendsto_power) lemma tendsto_prod [tendsto_intros]: fixes f :: "'a \ 'b \ 'c::{real_normed_algebra,comm_ring_1}" shows "(\i. i \ S \ (f i \ L i) F) \ ((\x. \i\S. f i x) \ (\i\S. L i)) F" by (induct S rule: infinite_finite_induct) (simp_all add: tendsto_mult) lemma continuous_prod [continuous_intros]: fixes f :: "'a \ 'b::t2_space \ 'c::{real_normed_algebra,comm_ring_1}" shows "(\i. i \ S \ continuous F (f i)) \ continuous F (\x. \i\S. f i x)" unfolding continuous_def by (rule tendsto_prod) lemma continuous_on_prod [continuous_intros]: fixes f :: "'a \ _ \ 'c::{real_normed_algebra,comm_ring_1}" shows "(\i. i \ S \ continuous_on s (f i)) \ continuous_on s (\x. \i\S. f i x)" unfolding continuous_on_def by (auto intro: tendsto_prod) lemma tendsto_of_real_iff: "((\x. of_real (f x) :: 'a::real_normed_div_algebra) \ of_real c) F \ (f \ c) F" unfolding tendsto_iff by simp lemma tendsto_add_const_iff: "((\x. c + f x :: 'a::topological_group_add) \ c + d) F \ (f \ d) F" using tendsto_add[OF tendsto_const[of c], of f d] and tendsto_add[OF tendsto_const[of "-c"], of "\x. c + f x" "c + d"] by auto class topological_monoid_mult = topological_semigroup_mult + monoid_mult class topological_comm_monoid_mult = topological_monoid_mult + comm_monoid_mult lemma tendsto_power_strong [tendsto_intros]: fixes f :: "_ \ 'b :: topological_monoid_mult" assumes "(f \ a) F" "(g \ b) F" shows "((\x. f x ^ g x) \ a ^ b) F" proof - have "((\x. f x ^ b) \ a ^ b) F" by (induction b) (auto intro: tendsto_intros assms) also from assms(2) have "eventually (\x. g x = b) F" by (simp add: nhds_discrete filterlim_principal) hence "eventually (\x. f x ^ b = f x ^ g x) F" by eventually_elim simp hence "((\x. f x ^ b) \ a ^ b) F \ ((\x. f x ^ g x) \ a ^ b) F" by (intro filterlim_cong refl) finally show ?thesis . qed lemma continuous_mult' [continuous_intros]: fixes f g :: "_ \ 'b::topological_semigroup_mult" shows "continuous F f \ continuous F g \ continuous F (\x. f x * g x)" unfolding continuous_def by (rule tendsto_mult) lemma continuous_power' [continuous_intros]: fixes f :: "_ \ 'b::topological_monoid_mult" shows "continuous F f \ continuous F g \ continuous F (\x. f x ^ g x)" unfolding continuous_def by (rule tendsto_power_strong) auto lemma continuous_on_mult' [continuous_intros]: fixes f g :: "_ \ 'b::topological_semigroup_mult" shows "continuous_on A f \ continuous_on A g \ continuous_on A (\x. f x * g x)" unfolding continuous_on_def by (auto intro: tendsto_mult) lemma continuous_on_power' [continuous_intros]: fixes f :: "_ \ 'b::topological_monoid_mult" shows "continuous_on A f \ continuous_on A g \ continuous_on A (\x. f x ^ g x)" unfolding continuous_on_def by (auto intro: tendsto_power_strong) lemma tendsto_mult_one: fixes f g :: "_ \ 'b::topological_monoid_mult" shows "(f \ 1) F \ (g \ 1) F \ ((\x. f x * g x) \ 1) F" by (drule (1) tendsto_mult) simp lemma tendsto_prod' [tendsto_intros]: fixes f :: "'a \ 'b \ 'c::topological_comm_monoid_mult" shows "(\i. i \ I \ (f i \ a i) F) \ ((\x. \i\I. f i x) \ (\i\I. a i)) F" by (induct I rule: infinite_finite_induct) (simp_all add: tendsto_mult) lemma tendsto_one_prod': fixes f :: "'a \ 'b \ 'c::topological_comm_monoid_mult" assumes "\i. i \ I \ ((\x. f x i) \ 1) F" shows "((\i. prod (f i) I) \ 1) F" using tendsto_prod' [of I "\x y. f y x" "\x. 1"] assms by simp lemma continuous_prod' [continuous_intros]: fixes f :: "'a \ 'b::t2_space \ 'c::topological_comm_monoid_mult" shows "(\i. i \ I \ continuous F (f i)) \ continuous F (\x. \i\I. f i x)" unfolding continuous_def by (rule tendsto_prod') lemma continuous_on_prod' [continuous_intros]: fixes f :: "'a \ 'b::topological_space \ 'c::topological_comm_monoid_mult" shows "(\i. i \ I \ continuous_on S (f i)) \ continuous_on S (\x. \i\I. f i x)" unfolding continuous_on_def by (auto intro: tendsto_prod') instance nat :: topological_comm_monoid_mult by standard (simp add: nhds_discrete principal_prod_principal filterlim_principal eventually_principal) instance int :: topological_comm_monoid_mult by standard (simp add: nhds_discrete principal_prod_principal filterlim_principal eventually_principal) class comm_real_normed_algebra_1 = real_normed_algebra_1 + comm_monoid_mult context real_normed_field begin subclass comm_real_normed_algebra_1 proof from norm_mult[of "1 :: 'a" 1] show "norm 1 = 1" by simp qed (simp_all add: norm_mult) end subsubsection \Inverse and division\ lemma (in bounded_bilinear) Zfun_prod_Bfun: assumes f: "Zfun f F" and g: "Bfun g F" shows "Zfun (\x. f x ** g x) F" proof - obtain K where K: "0 \ K" and norm_le: "\x y. norm (x ** y) \ norm x * norm y * K" using nonneg_bounded by blast obtain B where B: "0 < B" and norm_g: "eventually (\x. norm (g x) \ B) F" using g by (rule BfunE) have "eventually (\x. norm (f x ** g x) \ norm (f x) * (B * K)) F" using norm_g proof eventually_elim case (elim x) have "norm (f x ** g x) \ norm (f x) * norm (g x) * K" by (rule norm_le) also have "\ \ norm (f x) * B * K" by (intro mult_mono' order_refl norm_g norm_ge_zero mult_nonneg_nonneg K elim) also have "\ = norm (f x) * (B * K)" by (rule mult.assoc) finally show "norm (f x ** g x) \ norm (f x) * (B * K)" . qed with f show ?thesis by (rule Zfun_imp_Zfun) qed lemma (in bounded_bilinear) Bfun_prod_Zfun: assumes f: "Bfun f F" and g: "Zfun g F" shows "Zfun (\x. f x ** g x) F" using flip g f by (rule bounded_bilinear.Zfun_prod_Bfun) lemma Bfun_inverse: fixes a :: "'a::real_normed_div_algebra" assumes f: "(f \ a) F" assumes a: "a \ 0" shows "Bfun (\x. inverse (f x)) F" proof - from a have "0 < norm a" by simp then have "\r>0. r < norm a" by (rule dense) then obtain r where r1: "0 < r" and r2: "r < norm a" by blast have "eventually (\x. dist (f x) a < r) F" using tendstoD [OF f r1] by blast then have "eventually (\x. norm (inverse (f x)) \ inverse (norm a - r)) F" proof eventually_elim case (elim x) then have 1: "norm (f x - a) < r" by (simp add: dist_norm) then have 2: "f x \ 0" using r2 by auto then have "norm (inverse (f x)) = inverse (norm (f x))" by (rule nonzero_norm_inverse) also have "\ \ inverse (norm a - r)" proof (rule le_imp_inverse_le) show "0 < norm a - r" using r2 by simp have "norm a - norm (f x) \ norm (a - f x)" by (rule norm_triangle_ineq2) also have "\ = norm (f x - a)" by (rule norm_minus_commute) also have "\ < r" using 1 . finally show "norm a - r \ norm (f x)" by simp qed finally show "norm (inverse (f x)) \ inverse (norm a - r)" . qed then show ?thesis by (rule BfunI) qed lemma tendsto_inverse [tendsto_intros]: fixes a :: "'a::real_normed_div_algebra" assumes f: "(f \ a) F" and a: "a \ 0" shows "((\x. inverse (f x)) \ inverse a) F" proof - from a have "0 < norm a" by simp with f have "eventually (\x. dist (f x) a < norm a) F" by (rule tendstoD) then have "eventually (\x. f x \ 0) F" unfolding dist_norm by (auto elim!: eventually_mono) with a have "eventually (\x. inverse (f x) - inverse a = - (inverse (f x) * (f x - a) * inverse a)) F" by (auto elim!: eventually_mono simp: inverse_diff_inverse) moreover have "Zfun (\x. - (inverse (f x) * (f x - a) * inverse a)) F" by (intro Zfun_minus Zfun_mult_left bounded_bilinear.Bfun_prod_Zfun [OF bounded_bilinear_mult] Bfun_inverse [OF f a] f [unfolded tendsto_Zfun_iff]) ultimately show ?thesis unfolding tendsto_Zfun_iff by (rule Zfun_ssubst) qed lemma continuous_inverse: fixes f :: "'a::t2_space \ 'b::real_normed_div_algebra" assumes "continuous F f" and "f (Lim F (\x. x)) \ 0" shows "continuous F (\x. inverse (f x))" using assms unfolding continuous_def by (rule tendsto_inverse) lemma continuous_at_within_inverse[continuous_intros]: fixes f :: "'a::t2_space \ 'b::real_normed_div_algebra" assumes "continuous (at a within s) f" and "f a \ 0" shows "continuous (at a within s) (\x. inverse (f x))" using assms unfolding continuous_within by (rule tendsto_inverse) lemma continuous_on_inverse[continuous_intros]: fixes f :: "'a::topological_space \ 'b::real_normed_div_algebra" assumes "continuous_on s f" and "\x\s. f x \ 0" shows "continuous_on s (\x. inverse (f x))" using assms unfolding continuous_on_def by (blast intro: tendsto_inverse) lemma tendsto_divide [tendsto_intros]: fixes a b :: "'a::real_normed_field" shows "(f \ a) F \ (g \ b) F \ b \ 0 \ ((\x. f x / g x) \ a / b) F" by (simp add: tendsto_mult tendsto_inverse divide_inverse) lemma continuous_divide: fixes f g :: "'a::t2_space \ 'b::real_normed_field" assumes "continuous F f" and "continuous F g" and "g (Lim F (\x. x)) \ 0" shows "continuous F (\x. (f x) / (g x))" using assms unfolding continuous_def by (rule tendsto_divide) lemma continuous_at_within_divide[continuous_intros]: fixes f g :: "'a::t2_space \ 'b::real_normed_field" assumes "continuous (at a within s) f" "continuous (at a within s) g" and "g a \ 0" shows "continuous (at a within s) (\x. (f x) / (g x))" using assms unfolding continuous_within by (rule tendsto_divide) lemma isCont_divide[continuous_intros, simp]: fixes f g :: "'a::t2_space \ 'b::real_normed_field" assumes "isCont f a" "isCont g a" "g a \ 0" shows "isCont (\x. (f x) / g x) a" using assms unfolding continuous_at by (rule tendsto_divide) lemma continuous_on_divide[continuous_intros]: fixes f :: "'a::topological_space \ 'b::real_normed_field" assumes "continuous_on s f" "continuous_on s g" and "\x\s. g x \ 0" shows "continuous_on s (\x. (f x) / (g x))" using assms unfolding continuous_on_def by (blast intro: tendsto_divide) lemma tendsto_power_int [tendsto_intros]: fixes a :: "'a::real_normed_div_algebra" assumes f: "(f \ a) F" and a: "a \ 0" shows "((\x. power_int (f x) n) \ power_int a n) F" using assms by (cases n rule: int_cases4) (auto intro!: tendsto_intros simp: power_int_minus) lemma continuous_power_int: fixes f :: "'a::t2_space \ 'b::real_normed_div_algebra" assumes "continuous F f" and "f (Lim F (\x. x)) \ 0" shows "continuous F (\x. power_int (f x) n)" using assms unfolding continuous_def by (rule tendsto_power_int) lemma continuous_at_within_power_int[continuous_intros]: fixes f :: "'a::t2_space \ 'b::real_normed_div_algebra" assumes "continuous (at a within s) f" and "f a \ 0" shows "continuous (at a within s) (\x. power_int (f x) n)" using assms unfolding continuous_within by (rule tendsto_power_int) lemma continuous_on_power_int [continuous_intros]: fixes f :: "'a::topological_space \ 'b::real_normed_div_algebra" assumes "continuous_on s f" and "\x\s. f x \ 0" shows "continuous_on s (\x. power_int (f x) n)" using assms unfolding continuous_on_def by (blast intro: tendsto_power_int) lemma tendsto_sgn [tendsto_intros]: "(f \ l) F \ l \ 0 \ ((\x. sgn (f x)) \ sgn l) F" for l :: "'a::real_normed_vector" unfolding sgn_div_norm by (simp add: tendsto_intros) lemma continuous_sgn: fixes f :: "'a::t2_space \ 'b::real_normed_vector" assumes "continuous F f" and "f (Lim F (\x. x)) \ 0" shows "continuous F (\x. sgn (f x))" using assms unfolding continuous_def by (rule tendsto_sgn) lemma continuous_at_within_sgn[continuous_intros]: fixes f :: "'a::t2_space \ 'b::real_normed_vector" assumes "continuous (at a within s) f" and "f a \ 0" shows "continuous (at a within s) (\x. sgn (f x))" using assms unfolding continuous_within by (rule tendsto_sgn) lemma isCont_sgn[continuous_intros]: fixes f :: "'a::t2_space \ 'b::real_normed_vector" assumes "isCont f a" and "f a \ 0" shows "isCont (\x. sgn (f x)) a" using assms unfolding continuous_at by (rule tendsto_sgn) lemma continuous_on_sgn[continuous_intros]: fixes f :: "'a::topological_space \ 'b::real_normed_vector" assumes "continuous_on s f" and "\x\s. f x \ 0" shows "continuous_on s (\x. sgn (f x))" using assms unfolding continuous_on_def by (blast intro: tendsto_sgn) lemma filterlim_at_infinity: fixes f :: "_ \ 'a::real_normed_vector" assumes "0 \ c" shows "(LIM x F. f x :> at_infinity) \ (\r>c. eventually (\x. r \ norm (f x)) F)" unfolding filterlim_iff eventually_at_infinity proof safe fix P :: "'a \ bool" fix b assume *: "\r>c. eventually (\x. r \ norm (f x)) F" assume P: "\x. b \ norm x \ P x" have "max b (c + 1) > c" by auto with * have "eventually (\x. max b (c + 1) \ norm (f x)) F" by auto then show "eventually (\x. P (f x)) F" proof eventually_elim case (elim x) with P show "P (f x)" by auto qed qed force lemma filterlim_at_infinity_imp_norm_at_top: fixes F assumes "filterlim f at_infinity F" shows "filterlim (\x. norm (f x)) at_top F" proof - { fix r :: real have "\\<^sub>F x in F. r \ norm (f x)" using filterlim_at_infinity[of 0 f F] assms by (cases "r > 0") (auto simp: not_less intro: always_eventually order.trans[OF _ norm_ge_zero]) } thus ?thesis by (auto simp: filterlim_at_top) qed lemma filterlim_norm_at_top_imp_at_infinity: fixes F assumes "filterlim (\x. norm (f x)) at_top F" shows "filterlim f at_infinity F" using filterlim_at_infinity[of 0 f F] assms by (auto simp: filterlim_at_top) lemma filterlim_norm_at_top: "filterlim norm at_top at_infinity" by (rule filterlim_at_infinity_imp_norm_at_top) (rule filterlim_ident) lemma filterlim_at_infinity_conv_norm_at_top: "filterlim f at_infinity G \ filterlim (\x. norm (f x)) at_top G" by (auto simp: filterlim_at_infinity[OF order.refl] filterlim_at_top_gt[of _ _ 0]) lemma eventually_not_equal_at_infinity: "eventually (\x. x \ (a :: 'a :: {real_normed_vector})) at_infinity" proof - from filterlim_norm_at_top[where 'a = 'a] have "\\<^sub>F x in at_infinity. norm a < norm (x::'a)" by (auto simp: filterlim_at_top_dense) thus ?thesis by eventually_elim auto qed lemma filterlim_int_of_nat_at_topD: fixes F assumes "filterlim (\x. f (int x)) F at_top" shows "filterlim f F at_top" proof - have "filterlim (\x. f (int (nat x))) F at_top" by (rule filterlim_compose[OF assms filterlim_nat_sequentially]) also have "?this \ filterlim f F at_top" by (intro filterlim_cong refl eventually_mono [OF eventually_ge_at_top[of "0::int"]]) auto finally show ?thesis . qed lemma filterlim_int_sequentially [tendsto_intros]: "filterlim int at_top sequentially" unfolding filterlim_at_top proof fix C :: int show "eventually (\n. int n \ C) at_top" using eventually_ge_at_top[of "nat \C\"] by eventually_elim linarith qed lemma filterlim_real_of_int_at_top [tendsto_intros]: "filterlim real_of_int at_top at_top" unfolding filterlim_at_top proof fix C :: real show "eventually (\n. real_of_int n \ C) at_top" using eventually_ge_at_top[of "\C\"] by eventually_elim linarith qed lemma filterlim_abs_real: "filterlim (abs::real \ real) at_top at_top" proof (subst filterlim_cong[OF refl refl]) from eventually_ge_at_top[of "0::real"] show "eventually (\x::real. \x\ = x) at_top" by eventually_elim simp qed (simp_all add: filterlim_ident) lemma filterlim_of_real_at_infinity [tendsto_intros]: "filterlim (of_real :: real \ 'a :: real_normed_algebra_1) at_infinity at_top" by (intro filterlim_norm_at_top_imp_at_infinity) (auto simp: filterlim_abs_real) lemma not_tendsto_and_filterlim_at_infinity: fixes c :: "'a::real_normed_vector" assumes "F \ bot" and "(f \ c) F" and "filterlim f at_infinity F" shows False proof - from tendstoD[OF assms(2), of "1/2"] have "eventually (\x. dist (f x) c < 1/2) F" by simp moreover from filterlim_at_infinity[of "norm c" f F] assms(3) have "eventually (\x. norm (f x) \ norm c + 1) F" by simp ultimately have "eventually (\x. False) F" proof eventually_elim fix x assume A: "dist (f x) c < 1/2" assume "norm (f x) \ norm c + 1" also have "norm (f x) = dist (f x) 0" by simp also have "\ \ dist (f x) c + dist c 0" by (rule dist_triangle) finally show False using A by simp qed with assms show False by simp qed lemma filterlim_at_infinity_imp_not_convergent: assumes "filterlim f at_infinity sequentially" shows "\ convergent f" by (rule notI, rule not_tendsto_and_filterlim_at_infinity[OF _ _ assms]) (simp_all add: convergent_LIMSEQ_iff) lemma filterlim_at_infinity_imp_eventually_ne: assumes "filterlim f at_infinity F" shows "eventually (\z. f z \ c) F" proof - have "norm c + 1 > 0" by (intro add_nonneg_pos) simp_all with filterlim_at_infinity[OF order.refl, of f F] assms have "eventually (\z. norm (f z) \ norm c + 1) F" by blast then show ?thesis by eventually_elim auto qed lemma tendsto_of_nat [tendsto_intros]: "filterlim (of_nat :: nat \ 'a::real_normed_algebra_1) at_infinity sequentially" proof (subst filterlim_at_infinity[OF order.refl], intro allI impI) fix r :: real assume r: "r > 0" define n where "n = nat \r\" from r have n: "\m\n. of_nat m \ r" unfolding n_def by linarith from eventually_ge_at_top[of n] show "eventually (\m. norm (of_nat m :: 'a) \ r) sequentially" by eventually_elim (use n in simp_all) qed subsection \Relate \<^const>\at\, \<^const>\at_left\ and \<^const>\at_right\\ text \ This lemmas are useful for conversion between \<^term>\at x\ to \<^term>\at_left x\ and \<^term>\at_right x\ and also \<^term>\at_right 0\. \ lemmas filterlim_split_at_real = filterlim_split_at[where 'a=real] lemma filtermap_nhds_shift: "filtermap (\x. x - d) (nhds a) = nhds (a - d)" for a d :: "'a::real_normed_vector" by (rule filtermap_fun_inverse[where g="\x. x + d"]) (auto intro!: tendsto_eq_intros filterlim_ident) lemma filtermap_nhds_minus: "filtermap (\x. - x) (nhds a) = nhds (- a)" for a :: "'a::real_normed_vector" by (rule filtermap_fun_inverse[where g=uminus]) (auto intro!: tendsto_eq_intros filterlim_ident) lemma filtermap_at_shift: "filtermap (\x. x - d) (at a) = at (a - d)" for a d :: "'a::real_normed_vector" by (simp add: filter_eq_iff eventually_filtermap eventually_at_filter filtermap_nhds_shift[symmetric]) lemma filtermap_at_right_shift: "filtermap (\x. x - d) (at_right a) = at_right (a - d)" for a d :: "real" by (simp add: filter_eq_iff eventually_filtermap eventually_at_filter filtermap_nhds_shift[symmetric]) lemma filterlim_shift: fixes d :: "'a::real_normed_vector" assumes "filterlim f F (at a)" shows "filterlim (f \ (+) d) F (at (a - d))" unfolding filterlim_iff proof (intro strip) fix P assume "eventually P F" then have "\\<^sub>F x in filtermap (\y. y - d) (at a). P (f (d + x))" using assms by (force simp add: filterlim_iff eventually_filtermap) then show "(\\<^sub>F x in at (a - d). P ((f \ (+) d) x))" by (force simp add: filtermap_at_shift) qed lemma filterlim_shift_iff: fixes d :: "'a::real_normed_vector" shows "filterlim (f \ (+) d) F (at (a - d)) = filterlim f F (at a)" (is "?lhs = ?rhs") proof assume L: ?lhs show ?rhs using filterlim_shift [OF L, of "-d"] by (simp add: filterlim_iff) qed (metis filterlim_shift) lemma at_right_to_0: "at_right a = filtermap (\x. x + a) (at_right 0)" for a :: real using filtermap_at_right_shift[of "-a" 0] by simp lemma filterlim_at_right_to_0: "filterlim f F (at_right a) \ filterlim (\x. f (x + a)) F (at_right 0)" for a :: real unfolding filterlim_def filtermap_filtermap at_right_to_0[of a] .. lemma eventually_at_right_to_0: "eventually P (at_right a) \ eventually (\x. P (x + a)) (at_right 0)" for a :: real unfolding at_right_to_0[of a] by (simp add: eventually_filtermap) lemma at_to_0: "at a = filtermap (\x. x + a) (at 0)" for a :: "'a::real_normed_vector" using filtermap_at_shift[of "-a" 0] by simp lemma filterlim_at_to_0: "filterlim f F (at a) \ filterlim (\x. f (x + a)) F (at 0)" for a :: "'a::real_normed_vector" unfolding filterlim_def filtermap_filtermap at_to_0[of a] .. lemma eventually_at_to_0: "eventually P (at a) \ eventually (\x. P (x + a)) (at 0)" for a :: "'a::real_normed_vector" unfolding at_to_0[of a] by (simp add: eventually_filtermap) lemma filtermap_at_minus: "filtermap (\x. - x) (at a) = at (- a)" for a :: "'a::real_normed_vector" by (simp add: filter_eq_iff eventually_filtermap eventually_at_filter filtermap_nhds_minus[symmetric]) lemma at_left_minus: "at_left a = filtermap (\x. - x) (at_right (- a))" for a :: real by (simp add: filter_eq_iff eventually_filtermap eventually_at_filter filtermap_nhds_minus[symmetric]) lemma at_right_minus: "at_right a = filtermap (\x. - x) (at_left (- a))" for a :: real by (simp add: filter_eq_iff eventually_filtermap eventually_at_filter filtermap_nhds_minus[symmetric]) lemma filterlim_at_left_to_right: "filterlim f F (at_left a) \ filterlim (\x. f (- x)) F (at_right (-a))" for a :: real unfolding filterlim_def filtermap_filtermap at_left_minus[of a] .. lemma eventually_at_left_to_right: "eventually P (at_left a) \ eventually (\x. P (- x)) (at_right (-a))" for a :: real unfolding at_left_minus[of a] by (simp add: eventually_filtermap) lemma filterlim_uminus_at_top_at_bot: "LIM x at_bot. - x :: real :> at_top" unfolding filterlim_at_top eventually_at_bot_dense by (metis leI minus_less_iff order_less_asym) lemma filterlim_uminus_at_bot_at_top: "LIM x at_top. - x :: real :> at_bot" unfolding filterlim_at_bot eventually_at_top_dense by (metis leI less_minus_iff order_less_asym) lemma at_bot_mirror : shows "(at_bot::('a::{ordered_ab_group_add,linorder} filter)) = filtermap uminus at_top" proof (rule filtermap_fun_inverse[symmetric]) show "filterlim uminus at_top (at_bot::'a filter)" using eventually_at_bot_linorder filterlim_at_top le_minus_iff by force show "filterlim uminus (at_bot::'a filter) at_top" by (simp add: filterlim_at_bot minus_le_iff) qed auto lemma at_top_mirror : shows "(at_top::('a::{ordered_ab_group_add,linorder} filter)) = filtermap uminus at_bot" apply (subst at_bot_mirror) by (auto simp: filtermap_filtermap) lemma filterlim_at_top_mirror: "(LIM x at_top. f x :> F) \ (LIM x at_bot. f (-x::real) :> F)" unfolding filterlim_def at_top_mirror filtermap_filtermap .. lemma filterlim_at_bot_mirror: "(LIM x at_bot. f x :> F) \ (LIM x at_top. f (-x::real) :> F)" unfolding filterlim_def at_bot_mirror filtermap_filtermap .. lemma filterlim_uminus_at_top: "(LIM x F. f x :> at_top) \ (LIM x F. - (f x) :: real :> at_bot)" using filterlim_compose[OF filterlim_uminus_at_bot_at_top, of f F] and filterlim_compose[OF filterlim_uminus_at_top_at_bot, of "\x. - f x" F] by auto lemma tendsto_at_botI_sequentially: fixes f :: "real \ 'b::first_countable_topology" assumes *: "\X. filterlim X at_bot sequentially \ (\n. f (X n)) \ y" shows "(f \ y) at_bot" unfolding filterlim_at_bot_mirror proof (rule tendsto_at_topI_sequentially) fix X :: "nat \ real" assume "filterlim X at_top sequentially" thus "(\n. f (-X n)) \ y" by (intro *) (auto simp: filterlim_uminus_at_top) qed lemma filterlim_at_infinity_imp_filterlim_at_top: assumes "filterlim (f :: 'a \ real) at_infinity F" assumes "eventually (\x. f x > 0) F" shows "filterlim f at_top F" proof - from assms(2) have *: "eventually (\x. norm (f x) = f x) F" by eventually_elim simp from assms(1) show ?thesis unfolding filterlim_at_infinity_conv_norm_at_top by (subst (asm) filterlim_cong[OF refl refl *]) qed lemma filterlim_at_infinity_imp_filterlim_at_bot: assumes "filterlim (f :: 'a \ real) at_infinity F" assumes "eventually (\x. f x < 0) F" shows "filterlim f at_bot F" proof - from assms(2) have *: "eventually (\x. norm (f x) = -f x) F" by eventually_elim simp from assms(1) have "filterlim (\x. - f x) at_top F" unfolding filterlim_at_infinity_conv_norm_at_top by (subst (asm) filterlim_cong[OF refl refl *]) thus ?thesis by (simp add: filterlim_uminus_at_top) qed lemma filterlim_uminus_at_bot: "(LIM x F. f x :> at_bot) \ (LIM x F. - (f x) :: real :> at_top)" unfolding filterlim_uminus_at_top by simp lemma filterlim_inverse_at_top_right: "LIM x at_right (0::real). inverse x :> at_top" unfolding filterlim_at_top_gt[where c=0] eventually_at_filter proof safe fix Z :: real assume [arith]: "0 < Z" then have "eventually (\x. x < inverse Z) (nhds 0)" by (auto simp: eventually_nhds_metric dist_real_def intro!: exI[of _ "\inverse Z\"]) then show "eventually (\x. x \ 0 \ x \ {0<..} \ Z \ inverse x) (nhds 0)" by (auto elim!: eventually_mono simp: inverse_eq_divide field_simps) qed lemma tendsto_inverse_0: fixes x :: "_ \ 'a::real_normed_div_algebra" shows "(inverse \ (0::'a)) at_infinity" unfolding tendsto_Zfun_iff diff_0_right Zfun_def eventually_at_infinity proof safe fix r :: real assume "0 < r" show "\b. \x. b \ norm x \ norm (inverse x :: 'a) < r" proof (intro exI[of _ "inverse (r / 2)"] allI impI) fix x :: 'a from \0 < r\ have "0 < inverse (r / 2)" by simp also assume *: "inverse (r / 2) \ norm x" finally show "norm (inverse x) < r" using * \0 < r\ by (subst nonzero_norm_inverse) (simp_all add: inverse_eq_divide field_simps) qed qed lemma tendsto_add_filterlim_at_infinity: fixes c :: "'b::real_normed_vector" and F :: "'a filter" assumes "(f \ c) F" and "filterlim g at_infinity F" shows "filterlim (\x. f x + g x) at_infinity F" proof (subst filterlim_at_infinity[OF order_refl], safe) fix r :: real assume r: "r > 0" from assms(1) have "((\x. norm (f x)) \ norm c) F" by (rule tendsto_norm) then have "eventually (\x. norm (f x) < norm c + 1) F" by (rule order_tendstoD) simp_all moreover from r have "r + norm c + 1 > 0" by (intro add_pos_nonneg) simp_all with assms(2) have "eventually (\x. norm (g x) \ r + norm c + 1) F" unfolding filterlim_at_infinity[OF order_refl] by (elim allE[of _ "r + norm c + 1"]) simp_all ultimately show "eventually (\x. norm (f x + g x) \ r) F" proof eventually_elim fix x :: 'a assume A: "norm (f x) < norm c + 1" and B: "r + norm c + 1 \ norm (g x)" from A B have "r \ norm (g x) - norm (f x)" by simp also have "norm (g x) - norm (f x) \ norm (g x + f x)" by (rule norm_diff_ineq) finally show "r \ norm (f x + g x)" by (simp add: add_ac) qed qed lemma tendsto_add_filterlim_at_infinity': fixes c :: "'b::real_normed_vector" and F :: "'a filter" assumes "filterlim f at_infinity F" and "(g \ c) F" shows "filterlim (\x. f x + g x) at_infinity F" by (subst add.commute) (rule tendsto_add_filterlim_at_infinity assms)+ lemma filterlim_inverse_at_right_top: "LIM x at_top. inverse x :> at_right (0::real)" unfolding filterlim_at by (auto simp: eventually_at_top_dense) (metis tendsto_inverse_0 filterlim_mono at_top_le_at_infinity order_refl) lemma filterlim_inverse_at_top: "(f \ (0 :: real)) F \ eventually (\x. 0 < f x) F \ LIM x F. inverse (f x) :> at_top" by (intro filterlim_compose[OF filterlim_inverse_at_top_right]) (simp add: filterlim_def eventually_filtermap eventually_mono at_within_def le_principal) lemma filterlim_inverse_at_bot_neg: "LIM x (at_left (0::real)). inverse x :> at_bot" by (simp add: filterlim_inverse_at_top_right filterlim_uminus_at_bot filterlim_at_left_to_right) lemma filterlim_inverse_at_bot: "(f \ (0 :: real)) F \ eventually (\x. f x < 0) F \ LIM x F. inverse (f x) :> at_bot" unfolding filterlim_uminus_at_bot inverse_minus_eq[symmetric] by (rule filterlim_inverse_at_top) (simp_all add: tendsto_minus_cancel_left[symmetric]) lemma at_right_to_top: "(at_right (0::real)) = filtermap inverse at_top" by (intro filtermap_fun_inverse[symmetric, where g=inverse]) (auto intro: filterlim_inverse_at_top_right filterlim_inverse_at_right_top) lemma eventually_at_right_to_top: "eventually P (at_right (0::real)) \ eventually (\x. P (inverse x)) at_top" unfolding at_right_to_top eventually_filtermap .. lemma filterlim_at_right_to_top: "filterlim f F (at_right (0::real)) \ (LIM x at_top. f (inverse x) :> F)" unfolding filterlim_def at_right_to_top filtermap_filtermap .. lemma at_top_to_right: "at_top = filtermap inverse (at_right (0::real))" unfolding at_right_to_top filtermap_filtermap inverse_inverse_eq filtermap_ident .. lemma eventually_at_top_to_right: "eventually P at_top \ eventually (\x. P (inverse x)) (at_right (0::real))" unfolding at_top_to_right eventually_filtermap .. lemma filterlim_at_top_to_right: "filterlim f F at_top \ (LIM x (at_right (0::real)). f (inverse x) :> F)" unfolding filterlim_def at_top_to_right filtermap_filtermap .. lemma filterlim_inverse_at_infinity: fixes x :: "_ \ 'a::{real_normed_div_algebra, division_ring}" shows "filterlim inverse at_infinity (at (0::'a))" unfolding filterlim_at_infinity[OF order_refl] proof safe fix r :: real assume "0 < r" then show "eventually (\x::'a. r \ norm (inverse x)) (at 0)" unfolding eventually_at norm_inverse by (intro exI[of _ "inverse r"]) (auto simp: norm_conv_dist[symmetric] field_simps inverse_eq_divide) qed lemma filterlim_inverse_at_iff: fixes g :: "'a \ 'b::{real_normed_div_algebra, division_ring}" shows "(LIM x F. inverse (g x) :> at 0) \ (LIM x F. g x :> at_infinity)" unfolding filterlim_def filtermap_filtermap[symmetric] proof assume "filtermap g F \ at_infinity" then have "filtermap inverse (filtermap g F) \ filtermap inverse at_infinity" by (rule filtermap_mono) also have "\ \ at 0" using tendsto_inverse_0[where 'a='b] by (auto intro!: exI[of _ 1] simp: le_principal eventually_filtermap filterlim_def at_within_def eventually_at_infinity) finally show "filtermap inverse (filtermap g F) \ at 0" . next assume "filtermap inverse (filtermap g F) \ at 0" then have "filtermap inverse (filtermap inverse (filtermap g F)) \ filtermap inverse (at 0)" by (rule filtermap_mono) with filterlim_inverse_at_infinity show "filtermap g F \ at_infinity" by (auto intro: order_trans simp: filterlim_def filtermap_filtermap) qed lemma tendsto_mult_filterlim_at_infinity: fixes c :: "'a::real_normed_field" assumes "(f \ c) F" "c \ 0" assumes "filterlim g at_infinity F" shows "filterlim (\x. f x * g x) at_infinity F" proof - have "((\x. inverse (f x) * inverse (g x)) \ inverse c * 0) F" by (intro tendsto_mult tendsto_inverse assms filterlim_compose[OF tendsto_inverse_0]) then have "filterlim (\x. inverse (f x) * inverse (g x)) (at (inverse c * 0)) F" unfolding filterlim_at using assms by (auto intro: filterlim_at_infinity_imp_eventually_ne tendsto_imp_eventually_ne eventually_conj) then show ?thesis by (subst filterlim_inverse_at_iff[symmetric]) simp_all qed lemma tendsto_inverse_0_at_top: "LIM x F. f x :> at_top \ ((\x. inverse (f x) :: real) \ 0) F" by (metis filterlim_at filterlim_mono[OF _ at_top_le_at_infinity order_refl] filterlim_inverse_at_iff) lemma filterlim_inverse_at_top_iff: "eventually (\x. 0 < f x) F \ (LIM x F. inverse (f x) :> at_top) \ (f \ (0 :: real)) F" by (auto dest: tendsto_inverse_0_at_top filterlim_inverse_at_top) lemma filterlim_at_top_iff_inverse_0: "eventually (\x. 0 < f x) F \ (LIM x F. f x :> at_top) \ ((inverse \ f) \ (0 :: real)) F" using filterlim_inverse_at_top_iff [of "inverse \ f"] by auto lemma real_tendsto_divide_at_top: fixes c::"real" assumes "(f \ c) F" assumes "filterlim g at_top F" shows "((\x. f x / g x) \ 0) F" by (auto simp: divide_inverse_commute intro!: tendsto_mult[THEN tendsto_eq_rhs] tendsto_inverse_0_at_top assms) lemma mult_nat_left_at_top: "c > 0 \ filterlim (\x. c * x) at_top sequentially" for c :: nat by (rule filterlim_subseq) (auto simp: strict_mono_def) lemma mult_nat_right_at_top: "c > 0 \ filterlim (\x. x * c) at_top sequentially" for c :: nat by (rule filterlim_subseq) (auto simp: strict_mono_def) lemma filterlim_times_pos: "LIM x F1. c * f x :> at_right l" if "filterlim f (at_right p) F1" "0 < c" "l = c * p" for c::"'a::{linordered_field, linorder_topology}" unfolding filterlim_iff proof safe fix P assume "\\<^sub>F x in at_right l. P x" then obtain d where "c * p < d" "\y. y > c * p \ y < d \ P y" unfolding \l = _ \ eventually_at_right_field by auto then have "\\<^sub>F a in at_right p. P (c * a)" by (auto simp: eventually_at_right_field \0 < c\ field_simps intro!: exI[where x="d/c"]) from that(1)[unfolded filterlim_iff, rule_format, OF this] show "\\<^sub>F x in F1. P (c * f x)" . qed lemma filtermap_nhds_times: "c \ 0 \ filtermap (times c) (nhds a) = nhds (c * a)" for a c :: "'a::real_normed_field" by (rule filtermap_fun_inverse[where g="\x. inverse c * x"]) (auto intro!: tendsto_eq_intros filterlim_ident) lemma filtermap_times_pos_at_right: fixes c::"'a::{linordered_field, linorder_topology}" assumes "c > 0" shows "filtermap (times c) (at_right p) = at_right (c * p)" using assms by (intro filtermap_fun_inverse[where g="\x. inverse c * x"]) (auto intro!: filterlim_ident filterlim_times_pos) lemma at_to_infinity: "(at (0::'a::{real_normed_field,field})) = filtermap inverse at_infinity" proof (rule antisym) have "(inverse \ (0::'a)) at_infinity" by (fact tendsto_inverse_0) then show "filtermap inverse at_infinity \ at (0::'a)" using filterlim_def filterlim_ident filterlim_inverse_at_iff by fastforce next have "filtermap inverse (filtermap inverse (at (0::'a))) \ filtermap inverse at_infinity" using filterlim_inverse_at_infinity unfolding filterlim_def by (rule filtermap_mono) then show "at (0::'a) \ filtermap inverse at_infinity" by (simp add: filtermap_ident filtermap_filtermap) qed lemma lim_at_infinity_0: fixes l :: "'a::{real_normed_field,field}" shows "(f \ l) at_infinity \ ((f \ inverse) \ l) (at (0::'a))" by (simp add: tendsto_compose_filtermap at_to_infinity filtermap_filtermap) lemma lim_zero_infinity: fixes l :: "'a::{real_normed_field,field}" shows "((\x. f(1 / x)) \ l) (at (0::'a)) \ (f \ l) at_infinity" by (simp add: inverse_eq_divide lim_at_infinity_0 comp_def) text \ We only show rules for multiplication and addition when the functions are either against a real value or against infinity. Further rules are easy to derive by using @{thm filterlim_uminus_at_top}. \ lemma filterlim_tendsto_pos_mult_at_top: assumes f: "(f \ c) F" and c: "0 < c" and g: "LIM x F. g x :> at_top" shows "LIM x F. (f x * g x :: real) :> at_top" unfolding filterlim_at_top_gt[where c=0] proof safe fix Z :: real assume "0 < Z" from f \0 < c\ have "eventually (\x. c / 2 < f x) F" by (auto dest!: tendstoD[where e="c / 2"] elim!: eventually_mono simp: dist_real_def abs_real_def split: if_split_asm) moreover from g have "eventually (\x. (Z / c * 2) \ g x) F" unfolding filterlim_at_top by auto ultimately show "eventually (\x. Z \ f x * g x) F" proof eventually_elim case (elim x) with \0 < Z\ \0 < c\ have "c / 2 * (Z / c * 2) \ f x * g x" by (intro mult_mono) (auto simp: zero_le_divide_iff) with \0 < c\ show "Z \ f x * g x" by simp qed qed lemma filterlim_at_top_mult_at_top: assumes f: "LIM x F. f x :> at_top" and g: "LIM x F. g x :> at_top" shows "LIM x F. (f x * g x :: real) :> at_top" unfolding filterlim_at_top_gt[where c=0] proof safe fix Z :: real assume "0 < Z" from f have "eventually (\x. 1 \ f x) F" unfolding filterlim_at_top by auto moreover from g have "eventually (\x. Z \ g x) F" unfolding filterlim_at_top by auto ultimately show "eventually (\x. Z \ f x * g x) F" proof eventually_elim case (elim x) with \0 < Z\ have "1 * Z \ f x * g x" by (intro mult_mono) (auto simp: zero_le_divide_iff) then show "Z \ f x * g x" by simp qed qed lemma filterlim_at_top_mult_tendsto_pos: assumes f: "(f \ c) F" and c: "0 < c" and g: "LIM x F. g x :> at_top" shows "LIM x F. (g x * f x:: real) :> at_top" by (auto simp: mult.commute intro!: filterlim_tendsto_pos_mult_at_top f c g) lemma filterlim_tendsto_pos_mult_at_bot: fixes c :: real assumes "(f \ c) F" "0 < c" "filterlim g at_bot F" shows "LIM x F. f x * g x :> at_bot" using filterlim_tendsto_pos_mult_at_top[OF assms(1,2), of "\x. - g x"] assms(3) unfolding filterlim_uminus_at_bot by simp lemma filterlim_tendsto_neg_mult_at_bot: fixes c :: real assumes c: "(f \ c) F" "c < 0" and g: "filterlim g at_top F" shows "LIM x F. f x * g x :> at_bot" using c filterlim_tendsto_pos_mult_at_top[of "\x. - f x" "- c" F, OF _ _ g] unfolding filterlim_uminus_at_bot tendsto_minus_cancel_left by simp lemma filterlim_pow_at_top: fixes f :: "'a \ real" assumes "0 < n" and f: "LIM x F. f x :> at_top" shows "LIM x F. (f x)^n :: real :> at_top" using \0 < n\ proof (induct n) case 0 then show ?case by simp next case (Suc n) with f show ?case by (cases "n = 0") (auto intro!: filterlim_at_top_mult_at_top) qed lemma filterlim_pow_at_bot_even: fixes f :: "real \ real" shows "0 < n \ LIM x F. f x :> at_bot \ even n \ LIM x F. (f x)^n :> at_top" using filterlim_pow_at_top[of n "\x. - f x" F] by (simp add: filterlim_uminus_at_top) lemma filterlim_pow_at_bot_odd: fixes f :: "real \ real" shows "0 < n \ LIM x F. f x :> at_bot \ odd n \ LIM x F. (f x)^n :> at_bot" using filterlim_pow_at_top[of n "\x. - f x" F] by (simp add: filterlim_uminus_at_bot) lemma filterlim_power_at_infinity [tendsto_intros]: fixes F and f :: "'a \ 'b :: real_normed_div_algebra" assumes "filterlim f at_infinity F" "n > 0" shows "filterlim (\x. f x ^ n) at_infinity F" by (rule filterlim_norm_at_top_imp_at_infinity) (auto simp: norm_power intro!: filterlim_pow_at_top assms intro: filterlim_at_infinity_imp_norm_at_top) lemma filterlim_tendsto_add_at_top: assumes f: "(f \ c) F" and g: "LIM x F. g x :> at_top" shows "LIM x F. (f x + g x :: real) :> at_top" unfolding filterlim_at_top_gt[where c=0] proof safe fix Z :: real assume "0 < Z" from f have "eventually (\x. c - 1 < f x) F" by (auto dest!: tendstoD[where e=1] elim!: eventually_mono simp: dist_real_def) moreover from g have "eventually (\x. Z - (c - 1) \ g x) F" unfolding filterlim_at_top by auto ultimately show "eventually (\x. Z \ f x + g x) F" by eventually_elim simp qed lemma LIM_at_top_divide: fixes f g :: "'a \ real" assumes f: "(f \ a) F" "0 < a" and g: "(g \ 0) F" "eventually (\x. 0 < g x) F" shows "LIM x F. f x / g x :> at_top" unfolding divide_inverse by (rule filterlim_tendsto_pos_mult_at_top[OF f]) (rule filterlim_inverse_at_top[OF g]) lemma filterlim_at_top_add_at_top: assumes f: "LIM x F. f x :> at_top" and g: "LIM x F. g x :> at_top" shows "LIM x F. (f x + g x :: real) :> at_top" unfolding filterlim_at_top_gt[where c=0] proof safe fix Z :: real assume "0 < Z" from f have "eventually (\x. 0 \ f x) F" unfolding filterlim_at_top by auto moreover from g have "eventually (\x. Z \ g x) F" unfolding filterlim_at_top by auto ultimately show "eventually (\x. Z \ f x + g x) F" by eventually_elim simp qed lemma tendsto_divide_0: fixes f :: "_ \ 'a::{real_normed_div_algebra, division_ring}" assumes f: "(f \ c) F" and g: "LIM x F. g x :> at_infinity" shows "((\x. f x / g x) \ 0) F" using tendsto_mult[OF f filterlim_compose[OF tendsto_inverse_0 g]] by (simp add: divide_inverse) lemma linear_plus_1_le_power: fixes x :: real assumes x: "0 \ x" shows "real n * x + 1 \ (x + 1) ^ n" proof (induct n) case 0 then show ?case by simp next case (Suc n) from x have "real (Suc n) * x + 1 \ (x + 1) * (real n * x + 1)" by (simp add: field_simps) also have "\ \ (x + 1)^Suc n" using Suc x by (simp add: mult_left_mono) finally show ?case . qed lemma filterlim_realpow_sequentially_gt1: fixes x :: "'a :: real_normed_div_algebra" assumes x[arith]: "1 < norm x" shows "LIM n sequentially. x ^ n :> at_infinity" proof (intro filterlim_at_infinity[THEN iffD2] allI impI) fix y :: real assume "0 < y" obtain N :: nat where "y < real N * (norm x - 1)" by (meson diff_gt_0_iff_gt reals_Archimedean3 x) also have "\ \ real N * (norm x - 1) + 1" by simp also have "\ \ (norm x - 1 + 1) ^ N" by (rule linear_plus_1_le_power) simp also have "\ = norm x ^ N" by simp finally have "\n\N. y \ norm x ^ n" by (metis order_less_le_trans power_increasing order_less_imp_le x) then show "eventually (\n. y \ norm (x ^ n)) sequentially" unfolding eventually_sequentially by (auto simp: norm_power) qed simp lemma filterlim_divide_at_infinity: fixes f g :: "'a \ 'a :: real_normed_field" assumes "filterlim f (nhds c) F" "filterlim g (at 0) F" "c \ 0" shows "filterlim (\x. f x / g x) at_infinity F" proof - have "filterlim (\x. f x * inverse (g x)) at_infinity F" by (intro tendsto_mult_filterlim_at_infinity[OF assms(1,3)] filterlim_compose [OF filterlim_inverse_at_infinity assms(2)]) thus ?thesis by (simp add: field_simps) qed subsection \Floor and Ceiling\ lemma eventually_floor_less: fixes f :: "'a \ 'b::{order_topology,floor_ceiling}" assumes f: "(f \ l) F" and l: "l \ \" shows "\\<^sub>F x in F. of_int (floor l) < f x" by (intro order_tendstoD[OF f]) (metis Ints_of_int antisym_conv2 floor_correct l) lemma eventually_less_ceiling: fixes f :: "'a \ 'b::{order_topology,floor_ceiling}" assumes f: "(f \ l) F" and l: "l \ \" shows "\\<^sub>F x in F. f x < of_int (ceiling l)" by (intro order_tendstoD[OF f]) (metis Ints_of_int l le_of_int_ceiling less_le) lemma eventually_floor_eq: fixes f::"'a \ 'b::{order_topology,floor_ceiling}" assumes f: "(f \ l) F" and l: "l \ \" shows "\\<^sub>F x in F. floor (f x) = floor l" using eventually_floor_less[OF assms] eventually_less_ceiling[OF assms] by eventually_elim (meson floor_less_iff less_ceiling_iff not_less_iff_gr_or_eq) lemma eventually_ceiling_eq: fixes f::"'a \ 'b::{order_topology,floor_ceiling}" assumes f: "(f \ l) F" and l: "l \ \" shows "\\<^sub>F x in F. ceiling (f x) = ceiling l" using eventually_floor_less[OF assms] eventually_less_ceiling[OF assms] by eventually_elim (meson floor_less_iff less_ceiling_iff not_less_iff_gr_or_eq) lemma tendsto_of_int_floor: fixes f::"'a \ 'b::{order_topology,floor_ceiling}" assumes "(f \ l) F" and "l \ \" shows "((\x. of_int (floor (f x)) :: 'c::{ring_1,topological_space}) \ of_int (floor l)) F" using eventually_floor_eq[OF assms] by (simp add: eventually_mono topological_tendstoI) lemma tendsto_of_int_ceiling: fixes f::"'a \ 'b::{order_topology,floor_ceiling}" assumes "(f \ l) F" and "l \ \" shows "((\x. of_int (ceiling (f x)):: 'c::{ring_1,topological_space}) \ of_int (ceiling l)) F" using eventually_ceiling_eq[OF assms] by (simp add: eventually_mono topological_tendstoI) lemma continuous_on_of_int_floor: "continuous_on (UNIV - \::'a::{order_topology, floor_ceiling} set) (\x. of_int (floor x)::'b::{ring_1, topological_space})" unfolding continuous_on_def by (auto intro!: tendsto_of_int_floor) lemma continuous_on_of_int_ceiling: "continuous_on (UNIV - \::'a::{order_topology, floor_ceiling} set) (\x. of_int (ceiling x)::'b::{ring_1, topological_space})" unfolding continuous_on_def by (auto intro!: tendsto_of_int_ceiling) subsection \Limits of Sequences\ lemma [trans]: "X = Y \ Y \ z \ X \ z" by simp lemma LIMSEQ_iff: fixes L :: "'a::real_normed_vector" shows "(X \ L) = (\r>0. \no. \n \ no. norm (X n - L) < r)" unfolding lim_sequentially dist_norm .. lemma LIMSEQ_I: "(\r. 0 < r \ \no. \n\no. norm (X n - L) < r) \ X \ L" for L :: "'a::real_normed_vector" by (simp add: LIMSEQ_iff) lemma LIMSEQ_D: "X \ L \ 0 < r \ \no. \n\no. norm (X n - L) < r" for L :: "'a::real_normed_vector" by (simp add: LIMSEQ_iff) lemma LIMSEQ_linear: "X \ x \ l > 0 \ (\ n. X (n * l)) \ x" unfolding tendsto_def eventually_sequentially by (metis div_le_dividend div_mult_self1_is_m le_trans mult.commute) text \Transformation of limit.\ lemma Lim_transform: "(g \ a) F \ ((\x. f x - g x) \ 0) F \ (f \ a) F" for a b :: "'a::real_normed_vector" using tendsto_add [of g a F "\x. f x - g x" 0] by simp lemma Lim_transform2: "(f \ a) F \ ((\x. f x - g x) \ 0) F \ (g \ a) F" for a b :: "'a::real_normed_vector" by (erule Lim_transform) (simp add: tendsto_minus_cancel) proposition Lim_transform_eq: "((\x. f x - g x) \ 0) F \ (f \ a) F \ (g \ a) F" for a :: "'a::real_normed_vector" using Lim_transform Lim_transform2 by blast lemma Lim_transform_eventually: "\(f \ l) F; eventually (\x. f x = g x) F\ \ (g \ l) F" using eventually_elim2 by (fastforce simp add: tendsto_def) lemma Lim_transform_within: assumes "(f \ l) (at x within S)" and "0 < d" and "\x'. x'\S \ 0 < dist x' x \ dist x' x < d \ f x' = g x'" shows "(g \ l) (at x within S)" proof (rule Lim_transform_eventually) show "eventually (\x. f x = g x) (at x within S)" using assms by (auto simp: eventually_at) show "(f \ l) (at x within S)" by fact qed lemma filterlim_transform_within: assumes "filterlim g G (at x within S)" assumes "G \ F" "0x'. x' \ S \ 0 < dist x' x \ dist x' x < d \ f x' = g x') " shows "filterlim f F (at x within S)" using assms apply (elim filterlim_mono_eventually) unfolding eventually_at by auto text \Common case assuming being away from some crucial point like 0.\ lemma Lim_transform_away_within: fixes a b :: "'a::t1_space" assumes "a \ b" and "\x\S. x \ a \ x \ b \ f x = g x" and "(f \ l) (at a within S)" shows "(g \ l) (at a within S)" proof (rule Lim_transform_eventually) show "(f \ l) (at a within S)" by fact show "eventually (\x. f x = g x) (at a within S)" unfolding eventually_at_topological by (rule exI [where x="- {b}"]) (simp add: open_Compl assms) qed lemma Lim_transform_away_at: fixes a b :: "'a::t1_space" assumes ab: "a \ b" and fg: "\x. x \ a \ x \ b \ f x = g x" and fl: "(f \ l) (at a)" shows "(g \ l) (at a)" using Lim_transform_away_within[OF ab, of UNIV f g l] fg fl by simp text \Alternatively, within an open set.\ lemma Lim_transform_within_open: assumes "(f \ l) (at a within T)" and "open s" and "a \ s" and "\x. x\s \ x \ a \ f x = g x" shows "(g \ l) (at a within T)" proof (rule Lim_transform_eventually) show "eventually (\x. f x = g x) (at a within T)" unfolding eventually_at_topological using assms by auto show "(f \ l) (at a within T)" by fact qed text \A congruence rule allowing us to transform limits assuming not at point.\ lemma Lim_cong_within: assumes "a = b" and "x = y" and "S = T" and "\x. x \ b \ x \ T \ f x = g x" shows "(f \ x) (at a within S) \ (g \ y) (at b within T)" unfolding tendsto_def eventually_at_topological using assms by simp text \An unbounded sequence's inverse tends to 0.\ lemma LIMSEQ_inverse_zero: assumes "\r::real. \N. \n\N. r < X n" shows "(\n. inverse (X n)) \ 0" apply (rule filterlim_compose[OF tendsto_inverse_0]) by (metis assms eventually_at_top_linorderI filterlim_at_top_dense filterlim_at_top_imp_at_infinity) text \The sequence \<^term>\1/n\ tends to 0 as \<^term>\n\ tends to infinity.\ lemma LIMSEQ_inverse_real_of_nat: "(\n. inverse (real (Suc n))) \ 0" by (metis filterlim_compose tendsto_inverse_0 filterlim_mono order_refl filterlim_Suc filterlim_compose[OF filterlim_real_sequentially] at_top_le_at_infinity) text \ The sequence \<^term>\r + 1/n\ tends to \<^term>\r\ as \<^term>\n\ tends to infinity is now easily proved. \ lemma LIMSEQ_inverse_real_of_nat_add: "(\n. r + inverse (real (Suc n))) \ r" using tendsto_add [OF tendsto_const LIMSEQ_inverse_real_of_nat] by auto lemma LIMSEQ_inverse_real_of_nat_add_minus: "(\n. r + -inverse (real (Suc n))) \ r" using tendsto_add [OF tendsto_const tendsto_minus [OF LIMSEQ_inverse_real_of_nat]] by auto lemma LIMSEQ_inverse_real_of_nat_add_minus_mult: "(\n. r * (1 + - inverse (real (Suc n)))) \ r" using tendsto_mult [OF tendsto_const LIMSEQ_inverse_real_of_nat_add_minus [of 1]] by auto lemma lim_inverse_n: "((\n. inverse(of_nat n)) \ (0::'a::real_normed_field)) sequentially" using lim_1_over_n by (simp add: inverse_eq_divide) lemma lim_inverse_n': "((\n. 1 / n) \ 0) sequentially" using lim_inverse_n by (simp add: inverse_eq_divide) lemma LIMSEQ_Suc_n_over_n: "(\n. of_nat (Suc n) / of_nat n :: 'a :: real_normed_field) \ 1" proof (rule Lim_transform_eventually) show "eventually (\n. 1 + inverse (of_nat n :: 'a) = of_nat (Suc n) / of_nat n) sequentially" using eventually_gt_at_top[of "0::nat"] by eventually_elim (simp add: field_simps) have "(\n. 1 + inverse (of_nat n) :: 'a) \ 1 + 0" by (intro tendsto_add tendsto_const lim_inverse_n) then show "(\n. 1 + inverse (of_nat n) :: 'a) \ 1" by simp qed lemma LIMSEQ_n_over_Suc_n: "(\n. of_nat n / of_nat (Suc n) :: 'a :: real_normed_field) \ 1" proof (rule Lim_transform_eventually) show "eventually (\n. inverse (of_nat (Suc n) / of_nat n :: 'a) = of_nat n / of_nat (Suc n)) sequentially" using eventually_gt_at_top[of "0::nat"] by eventually_elim (simp add: field_simps del: of_nat_Suc) have "(\n. inverse (of_nat (Suc n) / of_nat n :: 'a)) \ inverse 1" by (intro tendsto_inverse LIMSEQ_Suc_n_over_n) simp_all then show "(\n. inverse (of_nat (Suc n) / of_nat n :: 'a)) \ 1" by simp qed subsection \Convergence on sequences\ lemma convergent_cong: assumes "eventually (\x. f x = g x) sequentially" shows "convergent f \ convergent g" unfolding convergent_def by (subst filterlim_cong[OF refl refl assms]) (rule refl) lemma convergent_Suc_iff: "convergent (\n. f (Suc n)) \ convergent f" by (auto simp: convergent_def filterlim_sequentially_Suc) lemma convergent_ignore_initial_segment: "convergent (\n. f (n + m)) = convergent f" proof (induct m arbitrary: f) case 0 then show ?case by simp next case (Suc m) have "convergent (\n. f (n + Suc m)) \ convergent (\n. f (Suc n + m))" by simp also have "\ \ convergent (\n. f (n + m))" by (rule convergent_Suc_iff) also have "\ \ convergent f" by (rule Suc) finally show ?case . qed lemma convergent_add: fixes X Y :: "nat \ 'a::topological_monoid_add" assumes "convergent (\n. X n)" and "convergent (\n. Y n)" shows "convergent (\n. X n + Y n)" using assms unfolding convergent_def by (blast intro: tendsto_add) lemma convergent_sum: fixes X :: "'a \ nat \ 'b::topological_comm_monoid_add" shows "(\i. i \ A \ convergent (\n. X i n)) \ convergent (\n. \i\A. X i n)" by (induct A rule: infinite_finite_induct) (simp_all add: convergent_const convergent_add) lemma (in bounded_linear) convergent: assumes "convergent (\n. X n)" shows "convergent (\n. f (X n))" using assms unfolding convergent_def by (blast intro: tendsto) lemma (in bounded_bilinear) convergent: assumes "convergent (\n. X n)" and "convergent (\n. Y n)" shows "convergent (\n. X n ** Y n)" using assms unfolding convergent_def by (blast intro: tendsto) lemma convergent_minus_iff: fixes X :: "nat \ 'a::topological_group_add" shows "convergent X \ convergent (\n. - X n)" unfolding convergent_def by (force dest: tendsto_minus) lemma convergent_diff: fixes X Y :: "nat \ 'a::topological_group_add" assumes "convergent (\n. X n)" assumes "convergent (\n. Y n)" shows "convergent (\n. X n - Y n)" using assms unfolding convergent_def by (blast intro: tendsto_diff) lemma convergent_norm: assumes "convergent f" shows "convergent (\n. norm (f n))" proof - from assms have "f \ lim f" by (simp add: convergent_LIMSEQ_iff) then have "(\n. norm (f n)) \ norm (lim f)" by (rule tendsto_norm) then show ?thesis by (auto simp: convergent_def) qed lemma convergent_of_real: "convergent f \ convergent (\n. of_real (f n) :: 'a::real_normed_algebra_1)" unfolding convergent_def by (blast intro!: tendsto_of_real) lemma convergent_add_const_iff: "convergent (\n. c + f n :: 'a::topological_ab_group_add) \ convergent f" proof assume "convergent (\n. c + f n)" from convergent_diff[OF this convergent_const[of c]] show "convergent f" by simp next assume "convergent f" from convergent_add[OF convergent_const[of c] this] show "convergent (\n. c + f n)" by simp qed lemma convergent_add_const_right_iff: "convergent (\n. f n + c :: 'a::topological_ab_group_add) \ convergent f" using convergent_add_const_iff[of c f] by (simp add: add_ac) lemma convergent_diff_const_right_iff: "convergent (\n. f n - c :: 'a::topological_ab_group_add) \ convergent f" using convergent_add_const_right_iff[of f "-c"] by (simp add: add_ac) lemma convergent_mult: fixes X Y :: "nat \ 'a::topological_semigroup_mult" assumes "convergent (\n. X n)" and "convergent (\n. Y n)" shows "convergent (\n. X n * Y n)" using assms unfolding convergent_def by (blast intro: tendsto_mult) lemma convergent_mult_const_iff: assumes "c \ 0" shows "convergent (\n. c * f n :: 'a::{field,topological_semigroup_mult}) \ convergent f" proof assume "convergent (\n. c * f n)" from assms convergent_mult[OF this convergent_const[of "inverse c"]] show "convergent f" by (simp add: field_simps) next assume "convergent f" from convergent_mult[OF convergent_const[of c] this] show "convergent (\n. c * f n)" by simp qed lemma convergent_mult_const_right_iff: fixes c :: "'a::{field,topological_semigroup_mult}" assumes "c \ 0" shows "convergent (\n. f n * c) \ convergent f" using convergent_mult_const_iff[OF assms, of f] by (simp add: mult_ac) lemma convergent_imp_Bseq: "convergent f \ Bseq f" by (simp add: Cauchy_Bseq convergent_Cauchy) text \A monotone sequence converges to its least upper bound.\ lemma LIMSEQ_incseq_SUP: fixes X :: "nat \ 'a::{conditionally_complete_linorder,linorder_topology}" assumes u: "bdd_above (range X)" and X: "incseq X" shows "X \ (SUP i. X i)" by (rule order_tendstoI) (auto simp: eventually_sequentially u less_cSUP_iff intro: X[THEN incseqD] less_le_trans cSUP_lessD[OF u]) lemma LIMSEQ_decseq_INF: fixes X :: "nat \ 'a::{conditionally_complete_linorder, linorder_topology}" assumes u: "bdd_below (range X)" and X: "decseq X" shows "X \ (INF i. X i)" by (rule order_tendstoI) (auto simp: eventually_sequentially u cINF_less_iff intro: X[THEN decseqD] le_less_trans less_cINF_D[OF u]) text \Main monotonicity theorem.\ lemma Bseq_monoseq_convergent: "Bseq X \ monoseq X \ convergent X" for X :: "nat \ real" by (auto simp: monoseq_iff convergent_def intro: LIMSEQ_decseq_INF LIMSEQ_incseq_SUP dest: Bseq_bdd_above Bseq_bdd_below) lemma Bseq_mono_convergent: "Bseq X \ (\m n. m \ n \ X m \ X n) \ convergent X" for X :: "nat \ real" by (auto intro!: Bseq_monoseq_convergent incseq_imp_monoseq simp: incseq_def) lemma monoseq_imp_convergent_iff_Bseq: "monoseq f \ convergent f \ Bseq f" for f :: "nat \ real" using Bseq_monoseq_convergent[of f] convergent_imp_Bseq[of f] by blast lemma Bseq_monoseq_convergent'_inc: fixes f :: "nat \ real" shows "Bseq (\n. f (n + M)) \ (\m n. M \ m \ m \ n \ f m \ f n) \ convergent f" by (subst convergent_ignore_initial_segment [symmetric, of _ M]) (auto intro!: Bseq_monoseq_convergent simp: monoseq_def) lemma Bseq_monoseq_convergent'_dec: fixes f :: "nat \ real" shows "Bseq (\n. f (n + M)) \ (\m n. M \ m \ m \ n \ f m \ f n) \ convergent f" by (subst convergent_ignore_initial_segment [symmetric, of _ M]) (auto intro!: Bseq_monoseq_convergent simp: monoseq_def) lemma Cauchy_iff: "Cauchy X \ (\e>0. \M. \m\M. \n\M. norm (X m - X n) < e)" for X :: "nat \ 'a::real_normed_vector" unfolding Cauchy_def dist_norm .. lemma CauchyI: "(\e. 0 < e \ \M. \m\M. \n\M. norm (X m - X n) < e) \ Cauchy X" for X :: "nat \ 'a::real_normed_vector" by (simp add: Cauchy_iff) lemma CauchyD: "Cauchy X \ 0 < e \ \M. \m\M. \n\M. norm (X m - X n) < e" for X :: "nat \ 'a::real_normed_vector" by (simp add: Cauchy_iff) lemma incseq_convergent: fixes X :: "nat \ real" assumes "incseq X" and "\i. X i \ B" obtains L where "X \ L" "\i. X i \ L" proof atomize_elim from incseq_bounded[OF assms] \incseq X\ Bseq_monoseq_convergent[of X] obtain L where "X \ L" by (auto simp: convergent_def monoseq_def incseq_def) with \incseq X\ show "\L. X \ L \ (\i. X i \ L)" by (auto intro!: exI[of _ L] incseq_le) qed lemma decseq_convergent: fixes X :: "nat \ real" assumes "decseq X" and "\i. B \ X i" obtains L where "X \ L" "\i. L \ X i" proof atomize_elim from decseq_bounded[OF assms] \decseq X\ Bseq_monoseq_convergent[of X] obtain L where "X \ L" by (auto simp: convergent_def monoseq_def decseq_def) with \decseq X\ show "\L. X \ L \ (\i. L \ X i)" by (auto intro!: exI[of _ L] decseq_ge) qed lemma monoseq_convergent: fixes X :: "nat \ real" assumes X: "monoseq X" and B: "\i. \X i\ \ B" obtains L where "X \ L" using X unfolding monoseq_iff proof assume "incseq X" show thesis using abs_le_D1 [OF B] incseq_convergent [OF \incseq X\] that by meson next assume "decseq X" show thesis using decseq_convergent [OF \decseq X\] that by (metis B abs_le_iff add.inverse_inverse neg_le_iff_le) qed subsection \Power Sequences\ lemma Bseq_realpow: "0 \ x \ x \ 1 \ Bseq (\n. x ^ n)" for x :: real by (metis decseq_bounded decseq_def power_decreasing zero_le_power) lemma monoseq_realpow: "0 \ x \ x \ 1 \ monoseq (\n. x ^ n)" for x :: real using monoseq_def power_decreasing by blast lemma convergent_realpow: "0 \ x \ x \ 1 \ convergent (\n. x ^ n)" for x :: real by (blast intro!: Bseq_monoseq_convergent Bseq_realpow monoseq_realpow) lemma LIMSEQ_inverse_realpow_zero: "1 < x \ (\n. inverse (x ^ n)) \ 0" for x :: real by (rule filterlim_compose[OF tendsto_inverse_0 filterlim_realpow_sequentially_gt1]) simp lemma LIMSEQ_realpow_zero: fixes x :: real assumes "0 \ x" "x < 1" shows "(\n. x ^ n) \ 0" proof (cases "x = 0") case False with \0 \ x\ have x0: "0 < x" by simp then have "1 < inverse x" using \x < 1\ by (rule one_less_inverse) then have "(\n. inverse (inverse x ^ n)) \ 0" by (rule LIMSEQ_inverse_realpow_zero) then show ?thesis by (simp add: power_inverse) next case True show ?thesis by (rule LIMSEQ_imp_Suc) (simp add: True) qed lemma LIMSEQ_power_zero [tendsto_intros]: "norm x < 1 \ (\n. x ^ n) \ 0" for x :: "'a::real_normed_algebra_1" apply (drule LIMSEQ_realpow_zero [OF norm_ge_zero]) by (simp add: Zfun_le norm_power_ineq tendsto_Zfun_iff) lemma LIMSEQ_divide_realpow_zero: "1 < x \ (\n. a / (x ^ n) :: real) \ 0" by (rule tendsto_divide_0 [OF tendsto_const filterlim_realpow_sequentially_gt1]) simp lemma tendsto_power_zero: fixes x::"'a::real_normed_algebra_1" assumes "filterlim f at_top F" assumes "norm x < 1" shows "((\y. x ^ (f y)) \ 0) F" proof (rule tendstoI) fix e::real assume "0 < e" from tendstoD[OF LIMSEQ_power_zero[OF \norm x < 1\] \0 < e\] have "\\<^sub>F xa in sequentially. norm (x ^ xa) < e" by simp then obtain N where N: "norm (x ^ n) < e" if "n \ N" for n by (auto simp: eventually_sequentially) have "\\<^sub>F i in F. f i \ N" using \filterlim f sequentially F\ by (simp add: filterlim_at_top) then show "\\<^sub>F i in F. dist (x ^ f i) 0 < e" by eventually_elim (auto simp: N) qed text \Limit of \<^term>\c^n\ for \<^term>\\c\ < 1\.\ lemma LIMSEQ_abs_realpow_zero: "\c\ < 1 \ (\n. \c\ ^ n :: real) \ 0" by (rule LIMSEQ_realpow_zero [OF abs_ge_zero]) lemma LIMSEQ_abs_realpow_zero2: "\c\ < 1 \ (\n. c ^ n :: real) \ 0" by (rule LIMSEQ_power_zero) simp subsection \Limits of Functions\ lemma LIM_eq: "f \a\ L = (\r>0. \s>0. \x. x \ a \ norm (x - a) < s \ norm (f x - L) < r)" for a :: "'a::real_normed_vector" and L :: "'b::real_normed_vector" by (simp add: LIM_def dist_norm) lemma LIM_I: "(\r. 0 < r \ \s>0. \x. x \ a \ norm (x - a) < s \ norm (f x - L) < r) \ f \a\ L" for a :: "'a::real_normed_vector" and L :: "'b::real_normed_vector" by (simp add: LIM_eq) lemma LIM_D: "f \a\ L \ 0 < r \ \s>0.\x. x \ a \ norm (x - a) < s \ norm (f x - L) < r" for a :: "'a::real_normed_vector" and L :: "'b::real_normed_vector" by (simp add: LIM_eq) lemma LIM_offset: "f \a\ L \ (\x. f (x + k)) \(a - k)\ L" for a :: "'a::real_normed_vector" by (simp add: filtermap_at_shift[symmetric, of a k] filterlim_def filtermap_filtermap) lemma LIM_offset_zero: "f \a\ L \ (\h. f (a + h)) \0\ L" for a :: "'a::real_normed_vector" by (drule LIM_offset [where k = a]) (simp add: add.commute) lemma LIM_offset_zero_cancel: "(\h. f (a + h)) \0\ L \ f \a\ L" for a :: "'a::real_normed_vector" by (drule LIM_offset [where k = "- a"]) simp lemma LIM_offset_zero_iff: "NO_MATCH 0 a \ f \a\ L \ (\h. f (a + h)) \0\ L" for f :: "'a :: real_normed_vector \ _" using LIM_offset_zero_cancel[of f a L] LIM_offset_zero[of f L a] by auto lemma tendsto_offset_zero_iff: fixes f :: "'a :: real_normed_vector \ _" assumes " NO_MATCH 0 a" "a \ S" "open S" shows "(f \ L) (at a within S) \ ((\h. f (a + h)) \ L) (at 0)" using assms by (simp add: tendsto_within_open_NO_MATCH LIM_offset_zero_iff) lemma LIM_zero: "(f \ l) F \ ((\x. f x - l) \ 0) F" for f :: "'a \ 'b::real_normed_vector" unfolding tendsto_iff dist_norm by simp lemma LIM_zero_cancel: fixes f :: "'a \ 'b::real_normed_vector" shows "((\x. f x - l) \ 0) F \ (f \ l) F" unfolding tendsto_iff dist_norm by simp lemma LIM_zero_iff: "((\x. f x - l) \ 0) F = (f \ l) F" for f :: "'a \ 'b::real_normed_vector" unfolding tendsto_iff dist_norm by simp lemma LIM_imp_LIM: fixes f :: "'a::topological_space \ 'b::real_normed_vector" fixes g :: "'a::topological_space \ 'c::real_normed_vector" assumes f: "f \a\ l" and le: "\x. x \ a \ norm (g x - m) \ norm (f x - l)" shows "g \a\ m" by (rule metric_LIM_imp_LIM [OF f]) (simp add: dist_norm le) lemma LIM_equal2: fixes f g :: "'a::real_normed_vector \ 'b::topological_space" assumes "0 < R" and "\x. x \ a \ norm (x - a) < R \ f x = g x" shows "g \a\ l \ f \a\ l" by (rule metric_LIM_equal2 [OF _ assms]) (simp_all add: dist_norm) lemma LIM_compose2: fixes a :: "'a::real_normed_vector" assumes f: "f \a\ b" and g: "g \b\ c" and inj: "\d>0. \x. x \ a \ norm (x - a) < d \ f x \ b" shows "(\x. g (f x)) \a\ c" by (rule metric_LIM_compose2 [OF f g inj [folded dist_norm]]) lemma real_LIM_sandwich_zero: fixes f g :: "'a::topological_space \ real" assumes f: "f \a\ 0" and 1: "\x. x \ a \ 0 \ g x" and 2: "\x. x \ a \ g x \ f x" shows "g \a\ 0" proof (rule LIM_imp_LIM [OF f]) (* FIXME: use tendsto_sandwich *) fix x assume x: "x \ a" with 1 have "norm (g x - 0) = g x" by simp also have "g x \ f x" by (rule 2 [OF x]) also have "f x \ \f x\" by (rule abs_ge_self) also have "\f x\ = norm (f x - 0)" by simp finally show "norm (g x - 0) \ norm (f x - 0)" . qed subsection \Continuity\ lemma LIM_isCont_iff: "(f \a\ f a) = ((\h. f (a + h)) \0\ f a)" for f :: "'a::real_normed_vector \ 'b::topological_space" by (rule iffI [OF LIM_offset_zero LIM_offset_zero_cancel]) lemma isCont_iff: "isCont f x = (\h. f (x + h)) \0\ f x" for f :: "'a::real_normed_vector \ 'b::topological_space" by (simp add: isCont_def LIM_isCont_iff) lemma isCont_LIM_compose2: fixes a :: "'a::real_normed_vector" assumes f [unfolded isCont_def]: "isCont f a" and g: "g \f a\ l" and inj: "\d>0. \x. x \ a \ norm (x - a) < d \ f x \ f a" shows "(\x. g (f x)) \a\ l" by (rule LIM_compose2 [OF f g inj]) lemma isCont_norm [simp]: "isCont f a \ isCont (\x. norm (f x)) a" for f :: "'a::t2_space \ 'b::real_normed_vector" by (fact continuous_norm) lemma isCont_rabs [simp]: "isCont f a \ isCont (\x. \f x\) a" for f :: "'a::t2_space \ real" by (fact continuous_rabs) lemma isCont_add [simp]: "isCont f a \ isCont g a \ isCont (\x. f x + g x) a" for f :: "'a::t2_space \ 'b::topological_monoid_add" by (fact continuous_add) lemma isCont_minus [simp]: "isCont f a \ isCont (\x. - f x) a" for f :: "'a::t2_space \ 'b::real_normed_vector" by (fact continuous_minus) lemma isCont_diff [simp]: "isCont f a \ isCont g a \ isCont (\x. f x - g x) a" for f :: "'a::t2_space \ 'b::real_normed_vector" by (fact continuous_diff) lemma isCont_mult [simp]: "isCont f a \ isCont g a \ isCont (\x. f x * g x) a" for f g :: "'a::t2_space \ 'b::real_normed_algebra" by (fact continuous_mult) lemma (in bounded_linear) isCont: "isCont g a \ isCont (\x. f (g x)) a" by (fact continuous) lemma (in bounded_bilinear) isCont: "isCont f a \ isCont g a \ isCont (\x. f x ** g x) a" by (fact continuous) lemmas isCont_scaleR [simp] = bounded_bilinear.isCont [OF bounded_bilinear_scaleR] lemmas isCont_of_real [simp] = bounded_linear.isCont [OF bounded_linear_of_real] lemma isCont_power [simp]: "isCont f a \ isCont (\x. f x ^ n) a" for f :: "'a::t2_space \ 'b::{power,real_normed_algebra}" by (fact continuous_power) lemma isCont_sum [simp]: "\i\A. isCont (f i) a \ isCont (\x. \i\A. f i x) a" for f :: "'a \ 'b::t2_space \ 'c::topological_comm_monoid_add" by (auto intro: continuous_sum) subsection \Uniform Continuity\ lemma uniformly_continuous_on_def: fixes f :: "'a::metric_space \ 'b::metric_space" shows "uniformly_continuous_on s f \ (\e>0. \d>0. \x\s. \x'\s. dist x' x < d \ dist (f x') (f x) < e)" unfolding uniformly_continuous_on_uniformity uniformity_dist filterlim_INF filterlim_principal eventually_inf_principal by (force simp: Ball_def uniformity_dist[symmetric] eventually_uniformity_metric) abbreviation isUCont :: "['a::metric_space \ 'b::metric_space] \ bool" where "isUCont f \ uniformly_continuous_on UNIV f" lemma isUCont_def: "isUCont f \ (\r>0. \s>0. \x y. dist x y < s \ dist (f x) (f y) < r)" by (auto simp: uniformly_continuous_on_def dist_commute) lemma isUCont_isCont: "isUCont f \ isCont f x" by (drule uniformly_continuous_imp_continuous) (simp add: continuous_on_eq_continuous_at) lemma uniformly_continuous_on_Cauchy: fixes f :: "'a::metric_space \ 'b::metric_space" assumes "uniformly_continuous_on S f" "Cauchy X" "\n. X n \ S" shows "Cauchy (\n. f (X n))" using assms unfolding uniformly_continuous_on_def by (meson Cauchy_def) lemma isUCont_Cauchy: "isUCont f \ Cauchy X \ Cauchy (\n. f (X n))" by (rule uniformly_continuous_on_Cauchy[where S=UNIV and f=f]) simp_all lemma uniformly_continuous_imp_Cauchy_continuous: fixes f :: "'a::metric_space \ 'b::metric_space" shows "\uniformly_continuous_on S f; Cauchy \; \n. (\ n) \ S\ \ Cauchy(f \ \)" by (simp add: uniformly_continuous_on_def Cauchy_def) meson lemma (in bounded_linear) isUCont: "isUCont f" unfolding isUCont_def dist_norm proof (intro allI impI) fix r :: real assume r: "0 < r" obtain K where K: "0 < K" and norm_le: "norm (f x) \ norm x * K" for x using pos_bounded by blast show "\s>0. \x y. norm (x - y) < s \ norm (f x - f y) < r" proof (rule exI, safe) from r K show "0 < r / K" by simp next fix x y :: 'a assume xy: "norm (x - y) < r / K" have "norm (f x - f y) = norm (f (x - y))" by (simp only: diff) also have "\ \ norm (x - y) * K" by (rule norm_le) also from K xy have "\ < r" by (simp only: pos_less_divide_eq) finally show "norm (f x - f y) < r" . qed qed lemma (in bounded_linear) Cauchy: "Cauchy X \ Cauchy (\n. f (X n))" by (rule isUCont [THEN isUCont_Cauchy]) lemma LIM_less_bound: fixes f :: "real \ real" assumes ev: "b < x" "\ x' \ { b <..< x}. 0 \ f x'" and "isCont f x" shows "0 \ f x" proof (rule tendsto_lowerbound) show "(f \ f x) (at_left x)" using \isCont f x\ by (simp add: filterlim_at_split isCont_def) show "eventually (\x. 0 \ f x) (at_left x)" using ev by (auto simp: eventually_at dist_real_def intro!: exI[of _ "x - b"]) qed simp subsection \Nested Intervals and Bisection -- Needed for Compactness\ lemma nested_sequence_unique: assumes "\n. f n \ f (Suc n)" "\n. g (Suc n) \ g n" "\n. f n \ g n" "(\n. f n - g n) \ 0" shows "\l::real. ((\n. f n \ l) \ f \ l) \ ((\n. l \ g n) \ g \ l)" proof - have "incseq f" unfolding incseq_Suc_iff by fact have "decseq g" unfolding decseq_Suc_iff by fact have "f n \ g 0" for n proof - from \decseq g\ have "g n \ g 0" by (rule decseqD) simp with \\n. f n \ g n\[THEN spec, of n] show ?thesis by auto qed then obtain u where "f \ u" "\i. f i \ u" using incseq_convergent[OF \incseq f\] by auto moreover have "f 0 \ g n" for n proof - from \incseq f\ have "f 0 \ f n" by (rule incseqD) simp with \\n. f n \ g n\[THEN spec, of n] show ?thesis by simp qed then obtain l where "g \ l" "\i. l \ g i" using decseq_convergent[OF \decseq g\] by auto moreover note LIMSEQ_unique[OF assms(4) tendsto_diff[OF \f \ u\ \g \ l\]] ultimately show ?thesis by auto qed lemma Bolzano[consumes 1, case_names trans local]: fixes P :: "real \ real \ bool" assumes [arith]: "a \ b" and trans: "\a b c. P a b \ P b c \ a \ b \ b \ c \ P a c" 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) have [simp]: "l n \ u n" for n by (induct n) auto have "\x. ((\n. l n \ x) \ l \ x) \ ((\n. x \ u n) \ u \ x)" proof (safe intro!: nested_sequence_unique) show "l n \ l (Suc n)" "u (Suc n) \ u n" for n by (induct n) auto next have "l n - u n = (a - b) / 2^n" for n by (induct n) (auto simp: field_simps) then show "(\n. l n - u n) \ 0" by (simp add: LIMSEQ_divide_realpow_zero) qed fact then obtain x where x: "\n. l n \ x" "\n. x \ u n" and "l \ x" "u \ x" by auto obtain d where "0 < d" and d: "a \ x \ x \ b \ b - a < d \ P a b" for a b using \l 0 \ x\ \x \ u 0\ local[of x] by auto show "P a b" proof (rule ccontr) assume "\ P a b" have "\ P (l n) (u n)" for n proof (induct n) case 0 then show ?case by (simp add: \\ P a b\) next case (Suc n) with trans[of "l n" "(l n + u n) / 2" "u n"] show ?case by auto qed moreover { have "eventually (\n. x - d / 2 < l n) sequentially" using \0 < d\ \l \ x\ by (intro order_tendstoD[of _ x]) auto moreover have "eventually (\n. u n < x + d / 2) sequentially" using \0 < d\ \u \ x\ by (intro order_tendstoD[of _ x]) auto ultimately have "eventually (\n. P (l n) (u n)) sequentially" proof eventually_elim case (elim n) from add_strict_mono[OF this] have "u n - l n < d" by simp with x show "P (l n) (u n)" by (rule d) qed } ultimately show False by simp qed qed lemma compact_Icc[simp, intro]: "compact {a .. b::real}" proof (cases "a \ b", rule compactI) fix C assume C: "a \ b" "\t\C. open t" "{a..b} \ \C" define T where "T = {a .. b}" from C(1,3) show "\C'\C. finite C' \ {a..b} \ \C'" proof (induct rule: Bolzano) case (trans a b c) then have *: "{a..c} = {a..b} \ {b..c}" by auto with trans obtain C1 C2 where "C1\C" "finite C1" "{a..b} \ \C1" "C2\C" "finite C2" "{b..c} \ \C2" by auto with trans show ?case unfolding * by (intro exI[of _ "C1 \ C2"]) auto next case (local x) with C have "x \ \C" by auto with C(2) obtain c where "x \ c" "open c" "c \ C" by auto then obtain e where "0 < e" "{x - e <..< x + e} \ c" by (auto simp: open_dist dist_real_def subset_eq Ball_def abs_less_iff) with \c \ C\ show ?case by (safe intro!: exI[of _ "e/2"] exI[of _ "{c}"]) auto qed qed simp lemma continuous_image_closed_interval: fixes a b and f :: "real \ real" defines "S \ {a..b}" assumes "a \ b" and f: "continuous_on S f" shows "\c d. f`S = {c..d} \ c \ d" proof - have S: "compact S" "S \ {}" using \a \ b\ by (auto simp: S_def) obtain c where "c \ S" "\d\S. f d \ f c" using continuous_attains_sup[OF S f] by auto moreover obtain d where "d \ S" "\c\S. f d \ f c" using continuous_attains_inf[OF S f] by auto moreover have "connected (f`S)" using connected_continuous_image[OF f] connected_Icc by (auto simp: S_def) ultimately have "f ` S = {f d .. f c} \ f d \ f c" by (auto simp: connected_iff_interval) then show ?thesis by auto qed lemma open_Collect_positive: fixes f :: "'a::topological_space \ real" assumes f: "continuous_on s f" shows "\A. open A \ A \ s = {x\s. 0 < f x}" using continuous_on_open_invariant[THEN iffD1, OF f, rule_format, of "{0 <..}"] by (auto simp: Int_def field_simps) lemma open_Collect_less_Int: fixes f g :: "'a::topological_space \ real" assumes f: "continuous_on s f" and g: "continuous_on s g" shows "\A. open A \ A \ s = {x\s. f x < g x}" using open_Collect_positive[OF continuous_on_diff[OF g f]] by (simp add: field_simps) subsection \Boundedness of continuous functions\ text\By bisection, function continuous on closed interval is bounded above\ lemma isCont_eq_Ub: fixes f :: "real \ 'a::linorder_topology" shows "a \ b \ \x::real. a \ x \ x \ b \ isCont f x \ \M. (\x. a \ x \ x \ b \ f x \ M) \ (\x. a \ x \ x \ b \ f x = M)" using continuous_attains_sup[of "{a..b}" f] by (auto simp: continuous_at_imp_continuous_on Ball_def Bex_def) lemma isCont_eq_Lb: fixes f :: "real \ 'a::linorder_topology" shows "a \ b \ \x. a \ x \ x \ b \ isCont f x \ \M. (\x. a \ x \ x \ b \ M \ f x) \ (\x. a \ x \ x \ b \ f x = M)" using continuous_attains_inf[of "{a..b}" f] by (auto simp: continuous_at_imp_continuous_on Ball_def Bex_def) lemma isCont_bounded: fixes f :: "real \ 'a::linorder_topology" shows "a \ b \ \x. a \ x \ x \ b \ isCont f x \ \M. \x. a \ x \ x \ b \ f x \ M" using isCont_eq_Ub[of a b f] by auto lemma isCont_has_Ub: fixes f :: "real \ 'a::linorder_topology" shows "a \ b \ \x. a \ x \ x \ b \ isCont f x \ \M. (\x. a \ x \ x \ b \ f x \ M) \ (\N. N < M \ (\x. a \ x \ x \ b \ N < f x))" using isCont_eq_Ub[of a b f] by auto lemma isCont_Lb_Ub: fixes f :: "real \ real" assumes "a \ b" "\x. a \ x \ x \ b \ isCont f x" shows "\L M. (\x. a \ x \ x \ b \ L \ f x \ f x \ M) \ (\y. L \ y \ y \ M \ (\x. a \ x \ x \ b \ (f x = y)))" proof - obtain M where M: "a \ M" "M \ b" "\x. a \ x \ x \ b \ f x \ f M" using isCont_eq_Ub[OF assms] by auto obtain L where L: "a \ L" "L \ b" "\x. a \ x \ x \ b \ f L \ f x" using isCont_eq_Lb[OF assms] by auto have "(\x. a \ x \ x \ b \ f L \ f x \ f x \ f M)" using M L by simp moreover have "(\y. f L \ y \ y \ f M \ (\x\a. x \ b \ f x = y))" proof (cases "L \ M") case True then show ?thesis using IVT[of f L _ M] M L assms by (metis order.trans) next case False then show ?thesis using IVT2[of f L _ M] by (metis L(2) M(1) assms(2) le_cases order.trans) qed ultimately show ?thesis by blast qed text \Continuity of inverse function.\ lemma isCont_inverse_function: fixes f g :: "real \ real" assumes d: "0 < d" and inj: "\z. \z-x\ \ d \ g (f z) = z" and cont: "\z. \z-x\ \ d \ isCont f z" shows "isCont g (f x)" proof - let ?A = "f (x - d)" let ?B = "f (x + d)" let ?D = "{x - d..x + d}" have f: "continuous_on ?D f" using cont by (intro continuous_at_imp_continuous_on ballI) auto then have g: "continuous_on (f`?D) g" using inj by (intro continuous_on_inv) auto from d f have "{min ?A ?B <..< max ?A ?B} \ f ` ?D" by (intro connected_contains_Ioo connected_continuous_image) (auto split: split_min split_max) with g have "continuous_on {min ?A ?B <..< max ?A ?B} g" by (rule continuous_on_subset) moreover have "(?A < f x \ f x < ?B) \ (?B < f x \ f x < ?A)" using d inj by (intro continuous_inj_imp_mono[OF _ _ f] inj_on_imageI2[of g, OF inj_onI]) auto then have "f x \ {min ?A ?B <..< max ?A ?B}" by auto ultimately show ?thesis by (simp add: continuous_on_eq_continuous_at) qed lemma isCont_inverse_function2: fixes f g :: "real \ real" shows "\a < x; x < b; \z. \a \ z; z \ b\ \ g (f z) = z; \z. \a \ z; z \ b\ \ isCont f z\ \ isCont g (f x)" apply (rule isCont_inverse_function [where f=f and d="min (x - a) (b - x)"]) apply (simp_all add: abs_le_iff) done text \Bartle/Sherbert: Introduction to Real Analysis, Theorem 4.2.9, p. 110.\ lemma LIM_fun_gt_zero: "f \c\ l \ 0 < l \ \r. 0 < r \ (\x. x \ c \ \c - x\ < r \ 0 < f x)" for f :: "real \ real" by (force simp: dest: LIM_D) lemma LIM_fun_less_zero: "f \c\ l \ l < 0 \ \r. 0 < r \ (\x. x \ c \ \c - x\ < r \ f x < 0)" for f :: "real \ real" by (drule LIM_D [where r="-l"]) force+ lemma LIM_fun_not_zero: "f \c\ l \ l \ 0 \ \r. 0 < r \ (\x. x \ c \ \c - x\ < r \ f x \ 0)" for f :: "real \ real" using LIM_fun_gt_zero[of f l c] LIM_fun_less_zero[of f l c] by (auto simp: neq_iff) end