diff --git a/src/HOL/Algebra/QuotRing.thy b/src/HOL/Algebra/QuotRing.thy --- a/src/HOL/Algebra/QuotRing.thy +++ b/src/HOL/Algebra/QuotRing.thy @@ -1,1088 +1,1086 @@ (* Title: HOL/Algebra/QuotRing.thy Author: Stephan Hohe Author: Paulo Emílio de Vilhena *) theory QuotRing imports RingHom begin section \Quotient Rings\ subsection \Multiplication on Cosets\ definition rcoset_mult :: "[('a, _) ring_scheme, 'a set, 'a set, 'a set] \ 'a set" ("[mod _:] _ \\ _" [81,81,81] 80) where "rcoset_mult R I A B = (\a\A. \b\B. I +>\<^bsub>R\<^esub> (a \\<^bsub>R\<^esub> b))" text \\<^const>\rcoset_mult\ fulfils the properties required by congruences\ lemma (in ideal) rcoset_mult_add: assumes "x \ carrier R" "y \ carrier R" shows "[mod I:] (I +> x) \ (I +> y) = I +> (x \ y)" proof - have 1: "z \ I +> x \ y" if x'rcos: "x' \ I +> x" and y'rcos: "y' \ I +> y" and zrcos: "z \ I +> x' \ y'" for z x' y' proof - from that obtain hx hy hz where hxI: "hx \ I" and x': "x' = hx \ x" and hyI: "hy \ I" and y': "y' = hy \ y" and hzI: "hz \ I" and z: "z = hz \ (x' \ y')" by (auto simp: a_r_coset_def r_coset_def) note carr = assms hxI[THEN a_Hcarr] hyI[THEN a_Hcarr] hzI[THEN a_Hcarr] from z x' y' have "z = hz \ ((hx \ x) \ (hy \ y))" by simp also from carr have "\ = (hz \ (hx \ (hy \ y)) \ x \ hy) \ x \ y" by algebra finally have z2: "z = (hz \ (hx \ (hy \ y)) \ x \ hy) \ x \ y" . from hxI hyI hzI carr have "hz \ (hx \ (hy \ y)) \ x \ hy \ I" by (simp add: I_l_closed I_r_closed) with z2 show ?thesis by (auto simp add: a_r_coset_def r_coset_def) qed have 2: "\a\I +> x. \b\I +> y. z \ I +> a \ b" if "z \ I +> x \ y" for z using assms a_rcos_self that by blast show ?thesis unfolding rcoset_mult_def using assms by (auto simp: intro!: 1 2) qed subsection \Quotient Ring Definition\ definition FactRing :: "[('a,'b) ring_scheme, 'a set] \ ('a set) ring" (infixl "Quot" 65) where "FactRing R I = \carrier = a_rcosets\<^bsub>R\<^esub> I, mult = rcoset_mult R I, one = (I +>\<^bsub>R\<^esub> \\<^bsub>R\<^esub>), zero = I, add = set_add R\" lemmas FactRing_simps = FactRing_def A_RCOSETS_defs a_r_coset_def[symmetric] subsection \Factorization over General Ideals\ text \The quotient is a ring\ lemma (in ideal) quotient_is_ring: "ring (R Quot I)" proof (rule ringI) show "abelian_group (R Quot I)" apply (rule comm_group_abelian_groupI) apply (simp add: FactRing_def a_factorgroup_is_comm_group[unfolded A_FactGroup_def']) done show "Group.monoid (R Quot I)" by (rule monoidI) (auto simp add: FactRing_simps rcoset_mult_add m_assoc) qed (auto simp: FactRing_simps rcoset_mult_add a_rcos_sum l_distr r_distr) text \This is a ring homomorphism\ lemma (in ideal) rcos_ring_hom: "((+>) I) \ ring_hom R (R Quot I)" by (simp add: ring_hom_memI FactRing_def a_rcosetsI[OF a_subset] rcoset_mult_add a_rcos_sum) lemma (in ideal) rcos_ring_hom_ring: "ring_hom_ring R (R Quot I) ((+>) I)" by (simp add: local.ring_axioms quotient_is_ring rcos_ring_hom ring_hom_ringI2) text \The quotient of a cring is also commutative\ lemma (in ideal) quotient_is_cring: assumes "cring R" shows "cring (R Quot I)" proof - interpret cring R by fact show ?thesis apply (intro cring.intro comm_monoid.intro comm_monoid_axioms.intro quotient_is_ring) apply (rule ring.axioms[OF quotient_is_ring]) apply (auto simp add: FactRing_simps rcoset_mult_add m_comm) done qed text \Cosets as a ring homomorphism on crings\ lemma (in ideal) rcos_ring_hom_cring: assumes "cring R" shows "ring_hom_cring R (R Quot I) ((+>) I)" proof - interpret cring R by fact show ?thesis apply (rule ring_hom_cringI) apply (rule rcos_ring_hom_ring) apply (rule is_cring) apply (rule quotient_is_cring) apply (rule is_cring) done qed subsection \Factorization over Prime Ideals\ text \The quotient ring generated by a prime ideal is a domain\ lemma (in primeideal) quotient_is_domain: "domain (R Quot I)" proof - have 1: "I +> \ = I \ False" using I_notcarr a_rcos_self one_imp_carrier by blast have 2: "I +> x = I" if carr: "x \ carrier R" "y \ carrier R" and a: "I +> x \ y = I" and b: "I +> y \ I" for x y by (metis I_prime a a_rcos_const a_rcos_self b m_closed that) show ?thesis apply (intro domain.intro quotient_is_cring is_cring domain_axioms.intro) apply (metis "1" FactRing_def monoid.simps(2) ring.simps(1)) apply (simp add: FactRing_simps) by (metis "2" rcoset_mult_add) qed text \Generating right cosets of a prime ideal is a homomorphism on commutative rings\ lemma (in primeideal) rcos_ring_hom_cring: "ring_hom_cring R (R Quot I) ((+>) I)" by (rule rcos_ring_hom_cring) (rule is_cring) subsection \Factorization over Maximal Ideals\ text \In a commutative ring, the quotient ring over a maximal ideal is a field. The proof follows ``W. Adkins, S. Weintraub: Algebra -- An Approach via Module Theory''\ proposition (in maximalideal) quotient_is_field: assumes "cring R" shows "field (R Quot I)" proof - interpret cring R by fact have 1: "\\<^bsub>R Quot I\<^esub> \ \\<^bsub>R Quot I\<^esub>" \ \Quotient is not empty\ proof assume "\\<^bsub>R Quot I\<^esub> = \\<^bsub>R Quot I\<^esub>" then have II1: "I = I +> \" by (simp add: FactRing_def) then have "I = carrier R" using a_rcos_self one_imp_carrier by blast with I_notcarr show False by simp qed have 2: "\y\carrier R. I +> a \ y = I +> \" if IanI: "I +> a \ I" and acarr: "a \ carrier R" for a \ \Existence of Inverse\ proof - \ \Helper ideal \J\\ define J :: "'a set" where "J = (carrier R #> a) <+> I" have idealJ: "ideal J R" using J_def acarr add_ideals cgenideal_eq_rcos cgenideal_ideal is_ideal by auto have IinJ: "I \ J" proof (clarsimp simp: J_def r_coset_def set_add_defs) fix x assume xI: "x \ I" have "x = \ \ a \ x" by (simp add: acarr xI) with xI show "\xa\carrier R. \k\I. x = xa \ a \ k" by fast qed have JnI: "J \ I" proof - have "a \ I" using IanI a_rcos_const by blast moreover have "a \ J" proof (simp add: J_def r_coset_def set_add_defs) from acarr have "a = \ \ a \ \" by algebra with one_closed and additive_subgroup.zero_closed[OF is_additive_subgroup] show "\x\carrier R. \k\I. a = x \ a \ k" by fast qed ultimately show ?thesis by blast qed then have Jcarr: "J = carrier R" using I_maximal IinJ additive_subgroup.a_subset idealJ ideal_def by blast \ \Calculating an inverse for \<^term>\a\\ from one_closed[folded Jcarr] obtain r i where rcarr: "r \ carrier R" and iI: "i \ I" and one: "\ = r \ a \ i" by (auto simp add: J_def r_coset_def set_add_defs) from one and rcarr and acarr and iI[THEN a_Hcarr] have rai1: "a \ r = \i \ \" by algebra \ \Lifting to cosets\ from iI have "\i \ \ \ I +> \" by (intro a_rcosI, simp, intro a_subset, simp) with rai1 have "a \ r \ I +> \" by simp then have "I +> \ = I +> a \ r" by (rule a_repr_independence, simp) (rule a_subgroup) from rcarr and this[symmetric] show "\r\carrier R. I +> a \ r = I +> \" by fast qed show ?thesis apply (intro cring.cring_fieldI2 quotient_is_cring is_cring 1) apply (clarsimp simp add: FactRing_simps rcoset_mult_add 2) done qed lemma (in ring_hom_ring) trivial_hom_iff: "(h ` (carrier R) = { \\<^bsub>S\<^esub> }) = (a_kernel R S h = carrier R)" using group_hom.trivial_hom_iff[OF a_group_hom] by (simp add: a_kernel_def) lemma (in ring_hom_ring) trivial_ker_imp_inj: assumes "a_kernel R S h = { \ }" shows "inj_on h (carrier R)" using group_hom.trivial_ker_imp_inj[OF a_group_hom] assms a_kernel_def[of R S h] by simp (* NEW ========================================================================== *) lemma (in ring_hom_ring) inj_iff_trivial_ker: shows "inj_on h (carrier R) \ a_kernel R S h = { \ }" using group_hom.inj_iff_trivial_ker[OF a_group_hom] a_kernel_def[of R S h] by simp (* NEW ========================================================================== *) corollary ring_hom_in_hom: assumes "h \ ring_hom R S" shows "h \ hom R S" and "h \ hom (add_monoid R) (add_monoid S)" using assms unfolding ring_hom_def hom_def by auto (* NEW ========================================================================== *) corollary set_add_hom: assumes "h \ ring_hom R S" "I \ carrier R" and "J \ carrier R" shows "h ` (I <+>\<^bsub>R\<^esub> J) = h ` I <+>\<^bsub>S\<^esub> h ` J" using set_mult_hom[OF ring_hom_in_hom(2)[OF assms(1)]] assms(2-3) unfolding a_kernel_def[of R S h] set_add_def by simp (* NEW ========================================================================== *) corollary a_coset_hom: assumes "h \ ring_hom R S" "I \ carrier R" "a \ carrier R" shows "h ` (a <+\<^bsub>R\<^esub> I) = h a <+\<^bsub>S\<^esub> (h ` I)" and "h ` (I +>\<^bsub>R\<^esub> a) = (h ` I) +>\<^bsub>S\<^esub> h a" using assms coset_hom[OF ring_hom_in_hom(2)[OF assms(1)], of I a] unfolding a_l_coset_def l_coset_eq_set_mult a_r_coset_def r_coset_eq_set_mult by simp_all (* NEW ========================================================================== *) corollary (in ring_hom_ring) set_add_ker_hom: assumes "I \ carrier R" shows "h ` (I <+> (a_kernel R S h)) = h ` I" and "h ` ((a_kernel R S h) <+> I) = h ` I" using group_hom.set_mult_ker_hom[OF a_group_hom] assms unfolding a_kernel_def[of R S h] set_add_def by simp+ lemma (in ring_hom_ring) non_trivial_field_hom_imp_inj: assumes "field R" shows "h ` (carrier R) \ { \\<^bsub>S\<^esub> } \ inj_on h (carrier R)" proof - assume "h ` (carrier R) \ { \\<^bsub>S\<^esub> }" hence "a_kernel R S h \ carrier R" using trivial_hom_iff by linarith hence "a_kernel R S h = { \ }" using field.all_ideals[OF assms] kernel_is_ideal by blast thus "inj_on h (carrier R)" using trivial_ker_imp_inj by blast qed lemma "field R \ cring R" using fieldE(1) by blast lemma non_trivial_field_hom_is_inj: assumes "h \ ring_hom R S" and "field R" and "field S" shows "inj_on h (carrier R)" proof - interpret ring_hom_cring R S h using assms(1) ring_hom_cring.intro[OF assms(2-3)[THEN fieldE(1)]] unfolding symmetric[OF ring_hom_cring_axioms_def] by simp have "h \\<^bsub>R\<^esub> = \\<^bsub>S\<^esub>" and "\\<^bsub>S\<^esub> \ \\<^bsub>S\<^esub>" using domain.one_not_zero[OF field.axioms(1)[OF assms(3)]] by auto hence "h ` (carrier R) \ { \\<^bsub>S\<^esub> }" using ring.kernel_zero ring.trivial_hom_iff by fastforce thus ?thesis using ring.non_trivial_field_hom_imp_inj[OF assms(2)] by simp qed lemma (in ring_hom_ring) img_is_add_subgroup: assumes "subgroup H (add_monoid R)" shows "subgroup (h ` H) (add_monoid S)" proof - have "group ((add_monoid R) \ carrier := H \)" using assms R.add.subgroup_imp_group by blast moreover have "H \ carrier R" by (simp add: R.add.subgroupE(1) assms) hence "h \ hom ((add_monoid R) \ carrier := H \) (add_monoid S)" unfolding hom_def by (auto simp add: subsetD) ultimately have "subgroup (h ` carrier ((add_monoid R) \ carrier := H \)) (add_monoid S)" using group_hom.img_is_subgroup[of "(add_monoid R) \ carrier := H \" "add_monoid S" h] using a_group_hom group_hom_axioms.intro group_hom_def by blast thus "subgroup (h ` H) (add_monoid S)" by simp qed lemma (in ring) ring_ideal_imp_quot_ideal: assumes "ideal I R" shows "ideal J R \ ideal ((+>) I ` J) (R Quot I)" proof - assume A: "ideal J R" show "ideal (((+>) I) ` J) (R Quot I)" proof (rule idealI) show "ring (R Quot I)" by (simp add: assms(1) ideal.quotient_is_ring) next have "subgroup J (add_monoid R)" by (simp add: additive_subgroup.a_subgroup A ideal.axioms(1)) moreover have "((+>) I) \ ring_hom R (R Quot I)" by (simp add: assms(1) ideal.rcos_ring_hom) ultimately show "subgroup ((+>) I ` J) (add_monoid (R Quot I))" using assms(1) ideal.rcos_ring_hom_ring ring_hom_ring.img_is_add_subgroup by blast next fix a x assume "a \ (+>) I ` J" "x \ carrier (R Quot I)" then obtain i j where i: "i \ carrier R" "x = I +> i" and j: "j \ J" "a = I +> j" unfolding FactRing_def using A_RCOSETS_def'[of R I] by auto hence "a \\<^bsub>R Quot I\<^esub> x = [mod I:] (I +> j) \ (I +> i)" unfolding FactRing_def by simp hence "a \\<^bsub>R Quot I\<^esub> x = I +> (j \ i)" using ideal.rcoset_mult_add[OF assms(1), of j i] i(1) j(1) A ideal.Icarr by force thus "a \\<^bsub>R Quot I\<^esub> x \ (+>) I ` J" using A i(1) j(1) by (simp add: ideal.I_r_closed) have "x \\<^bsub>R Quot I\<^esub> a = [mod I:] (I +> i) \ (I +> j)" unfolding FactRing_def i j by simp hence "x \\<^bsub>R Quot I\<^esub> a = I +> (i \ j)" using ideal.rcoset_mult_add[OF assms(1), of i j] i(1) j(1) A ideal.Icarr by force thus "x \\<^bsub>R Quot I\<^esub> a \ (+>) I ` J" using A i(1) j(1) by (simp add: ideal.I_l_closed) qed qed lemma (in ring_hom_ring) ideal_vimage: assumes "ideal I S" shows "ideal { r \ carrier R. h r \ I } R" (* or (carrier R) \ (h -` I) *) proof show "{ r \ carrier R. h r \ I } \ carrier (add_monoid R)" by auto next show "\\<^bsub>add_monoid R\<^esub> \ { r \ carrier R. h r \ I }" by (simp add: additive_subgroup.zero_closed assms ideal.axioms(1)) next fix a b assume "a \ { r \ carrier R. h r \ I }" and "b \ { r \ carrier R. h r \ I }" hence a: "a \ carrier R" "h a \ I" and b: "b \ carrier R" "h b \ I" by auto hence "h (a \ b) = (h a) \\<^bsub>S\<^esub> (h b)" using hom_add by blast moreover have "(h a) \\<^bsub>S\<^esub> (h b) \ I" using a b assms by (simp add: additive_subgroup.a_closed ideal.axioms(1)) ultimately show "a \\<^bsub>add_monoid R\<^esub> b \ { r \ carrier R. h r \ I }" using a(1) b (1) by auto have "h (\ a) = \\<^bsub>S\<^esub> (h a)" by (simp add: a) moreover have "\\<^bsub>S\<^esub> (h a) \ I" by (simp add: a(2) additive_subgroup.a_inv_closed assms ideal.axioms(1)) ultimately show "inv\<^bsub>add_monoid R\<^esub> a \ { r \ carrier R. h r \ I }" using a by (simp add: a_inv_def) next fix a r assume "a \ { r \ carrier R. h r \ I }" and r: "r \ carrier R" hence a: "a \ carrier R" "h a \ I" by auto have "h a \\<^bsub>S\<^esub> h r \ I" using assms a r by (simp add: ideal.I_r_closed) thus "a \ r \ { r \ carrier R. h r \ I }" by (simp add: a(1) r) have "h r \\<^bsub>S\<^esub> h a \ I" using assms a r by (simp add: ideal.I_l_closed) thus "r \ a \ { r \ carrier R. h r \ I }" by (simp add: a(1) r) qed lemma (in ring) canonical_proj_vimage_in_carrier: assumes "ideal I R" shows "J \ carrier (R Quot I) \ \ J \ carrier R" proof - assume A: "J \ carrier (R Quot I)" show "\ J \ carrier R" proof fix j assume j: "j \ \ J" then obtain j' where j': "j' \ J" "j \ j'" by blast then obtain r where r: "r \ carrier R" "j' = I +> r" using A j' unfolding FactRing_def using A_RCOSETS_def'[of R I] by auto thus "j \ carrier R" using j' assms by (meson a_r_coset_subset_G additive_subgroup.a_subset contra_subsetD ideal.axioms(1)) qed qed lemma (in ring) canonical_proj_vimage_mem_iff: assumes "ideal I R" "J \ carrier (R Quot I)" shows "\a. a \ carrier R \ (a \ (\ J)) = (I +> a \ J)" proof - fix a assume a: "a \ carrier R" show "(a \ (\ J)) = (I +> a \ J)" proof assume "a \ \ J" then obtain j where j: "j \ J" "a \ j" by blast then obtain r where r: "r \ carrier R" "j = I +> r" using assms j unfolding FactRing_def using A_RCOSETS_def'[of R I] by auto hence "I +> r = I +> a" using add.repr_independence[of a I r] j r by (metis a_r_coset_def additive_subgroup.a_subgroup assms(1) ideal.axioms(1)) thus "I +> a \ J" using r j by simp next assume "I +> a \ J" hence "\ \ a \ I +> a" using additive_subgroup.zero_closed[OF ideal.axioms(1)[OF assms(1)]] a_r_coset_def'[of R I a] by blast thus "a \ \ J" using a \I +> a \ J\ by auto qed qed corollary (in ring) quot_ideal_imp_ring_ideal: assumes "ideal I R" shows "ideal J (R Quot I) \ ideal (\ J) R" proof - assume A: "ideal J (R Quot I)" have "\ J = { r \ carrier R. I +> r \ J }" using canonical_proj_vimage_in_carrier[OF assms, of J] canonical_proj_vimage_mem_iff[OF assms, of J] additive_subgroup.a_subset[OF ideal.axioms(1)[OF A]] by blast thus "ideal (\ J) R" using ring_hom_ring.ideal_vimage[OF ideal.rcos_ring_hom_ring[OF assms] A] by simp qed lemma (in ring) ideal_incl_iff: assumes "ideal I R" "ideal J R" shows "(I \ J) = (J = (\ j \ J. I +> j))" proof assume A: "J = (\ j \ J. I +> j)" hence "I +> \ \ J" using additive_subgroup.zero_closed[OF ideal.axioms(1)[OF assms(2)]] by blast thus "I \ J" using additive_subgroup.a_subset[OF ideal.axioms(1)[OF assms(1)]] by simp next assume A: "I \ J" show "J = (\j\J. I +> j)" proof show "J \ (\ j \ J. I +> j)" proof fix j assume j: "j \ J" have "\ \ I" by (simp add: additive_subgroup.zero_closed assms(1) ideal.axioms(1)) hence "\ \ j \ I +> j" using a_r_coset_def'[of R I j] by blast thus "j \ (\j\J. I +> j)" using assms(2) j additive_subgroup.a_Hcarr ideal.axioms(1) by fastforce qed next show "(\ j \ J. I +> j) \ J" proof fix x assume "x \ (\ j \ J. I +> j)" then obtain j where j: "j \ J" "x \ I +> j" by blast then obtain i where i: "i \ I" "x = i \ j" using a_r_coset_def'[of R I j] by blast thus "x \ J" using assms(2) j A additive_subgroup.a_closed[OF ideal.axioms(1)[OF assms(2)]] by blast qed qed qed theorem (in ring) quot_ideal_correspondence: assumes "ideal I R" shows "bij_betw (\J. (+>) I ` J) { J. ideal J R \ I \ J } { J . ideal J (R Quot I) }" proof (rule bij_betw_byWitness[where ?f' = "\X. \ X"]) show "\J \ { J. ideal J R \ I \ J }. (\X. \ X) ((+>) I ` J) = J" using assms ideal_incl_iff by blast next show "(\J. (+>) I ` J) ` { J. ideal J R \ I \ J } \ { J. ideal J (R Quot I) }" using assms ring_ideal_imp_quot_ideal by auto next show "(\X. \ X) ` { J. ideal J (R Quot I) } \ { J. ideal J R \ I \ J }" proof fix J assume "J \ ((\X. \ X) ` { J. ideal J (R Quot I) })" then obtain J' where J': "ideal J' (R Quot I)" "J = \ J'" by blast hence "ideal J R" using assms quot_ideal_imp_ring_ideal by auto moreover have "I \ J'" using additive_subgroup.zero_closed[OF ideal.axioms(1)[OF J'(1)]] unfolding FactRing_def by simp ultimately show "J \ { J. ideal J R \ I \ J }" using J'(2) by auto qed next show "\J' \ { J. ideal J (R Quot I) }. ((+>) I ` (\ J')) = J'" proof fix J' assume "J' \ { J. ideal J (R Quot I) }" hence subset: "J' \ carrier (R Quot I) \ ideal J' (R Quot I)" using additive_subgroup.a_subset ideal_def by blast hence "((+>) I ` (\ J')) \ J'" using canonical_proj_vimage_in_carrier canonical_proj_vimage_mem_iff by (meson assms contra_subsetD image_subsetI) moreover have "J' \ ((+>) I ` (\ J'))" proof fix x assume "x \ J'" then obtain r where r: "r \ carrier R" "x = I +> r" using subset unfolding FactRing_def A_RCOSETS_def'[of R I] by auto hence "r \ (\ J')" using \x \ J'\ assms canonical_proj_vimage_mem_iff subset by blast thus "x \ ((+>) I ` (\ J'))" using r(2) by blast qed ultimately show "((+>) I ` (\ J')) = J'" by blast qed qed lemma (in cring) quot_domain_imp_primeideal: assumes "ideal P R" shows "domain (R Quot P) \ primeideal P R" proof - assume A: "domain (R Quot P)" show "primeideal P R" proof (rule primeidealI) show "ideal P R" using assms . show "cring R" using is_cring . next show "carrier R \ P" proof (rule ccontr) assume "\ carrier R \ P" hence "carrier R = P" by simp hence "\I. I \ carrier (R Quot P) \ I = P" unfolding FactRing_def A_RCOSETS_def' apply simp using a_coset_join2 additive_subgroup.a_subgroup assms ideal.axioms(1) by blast hence "\\<^bsub>(R Quot P)\<^esub> = \\<^bsub>(R Quot P)\<^esub>" by (metis assms ideal.quotient_is_ring ring.ring_simprules(2) ring.ring_simprules(6)) thus False using domain.one_not_zero[OF A] by simp qed next fix a b assume a: "a \ carrier R" and b: "b \ carrier R" and ab: "a \ b \ P" hence "P +> (a \ b) = \\<^bsub>(R Quot P)\<^esub>" unfolding FactRing_def by (simp add: a_coset_join2 additive_subgroup.a_subgroup assms ideal.axioms(1)) moreover have "(P +> a) \\<^bsub>(R Quot P)\<^esub> (P +> b) = P +> (a \ b)" unfolding FactRing_def using a b by (simp add: assms ideal.rcoset_mult_add) moreover have "P +> a \ carrier (R Quot P) \ P +> b \ carrier (R Quot P)" by (simp add: a b FactRing_def a_rcosetsI additive_subgroup.a_subset assms ideal.axioms(1)) ultimately have "P +> a = \\<^bsub>(R Quot P)\<^esub> \ P +> b = \\<^bsub>(R Quot P)\<^esub>" using domain.integral[OF A, of "P +> a" "P +> b"] by auto thus "a \ P \ b \ P" unfolding FactRing_def apply simp using a b assms a_coset_join1 additive_subgroup.a_subgroup ideal.axioms(1) by blast qed qed lemma (in cring) quot_domain_iff_primeideal: assumes "ideal P R" shows "domain (R Quot P) = primeideal P R" using quot_domain_imp_primeideal[OF assms] primeideal.quotient_is_domain[of P R] by auto subsection \Isomorphism\ definition ring_iso :: "_ \ _ \ ('a \ 'b) set" where "ring_iso R S = { h. h \ ring_hom R S \ bij_betw h (carrier R) (carrier S) }" definition is_ring_iso :: "_ \ _ \ bool" (infixr "\" 60) where "R \ S = (ring_iso R S \ {})" definition morphic_prop :: "_ \ ('a \ bool) \ bool" where "morphic_prop R P = ((P \\<^bsub>R\<^esub>) \ (\r \ carrier R. P r) \ (\r1 \ carrier R. \r2 \ carrier R. P (r1 \\<^bsub>R\<^esub> r2)) \ (\r1 \ carrier R. \r2 \ carrier R. P (r1 \\<^bsub>R\<^esub> r2)))" lemma ring_iso_memI: fixes R (structure) and S (structure) assumes "\x. x \ carrier R \ h x \ carrier S" and "\x y. \ x \ carrier R; y \ carrier R \ \ h (x \ y) = h x \\<^bsub>S\<^esub> h y" and "\x y. \ x \ carrier R; y \ carrier R \ \ h (x \ y) = h x \\<^bsub>S\<^esub> h y" and "h \ = \\<^bsub>S\<^esub>" and "bij_betw h (carrier R) (carrier S)" shows "h \ ring_iso R S" by (auto simp add: ring_hom_memI assms ring_iso_def) lemma ring_iso_memE: fixes R (structure) and S (structure) assumes "h \ ring_iso R S" shows "\x. x \ carrier R \ h x \ carrier S" and "\x y. \ x \ carrier R; y \ carrier R \ \ h (x \ y) = h x \\<^bsub>S\<^esub> h y" and "\x y. \ x \ carrier R; y \ carrier R \ \ h (x \ y) = h x \\<^bsub>S\<^esub> h y" and "h \ = \\<^bsub>S\<^esub>" and "bij_betw h (carrier R) (carrier S)" using assms unfolding ring_iso_def ring_hom_def by auto lemma morphic_propI: fixes R (structure) assumes "P \" and "\r. r \ carrier R \ P r" and "\r1 r2. \ r1 \ carrier R; r2 \ carrier R \ \ P (r1 \ r2)" and "\r1 r2. \ r1 \ carrier R; r2 \ carrier R \ \ P (r1 \ r2)" shows "morphic_prop R P" unfolding morphic_prop_def using assms by auto lemma morphic_propE: fixes R (structure) assumes "morphic_prop R P" shows "P \" and "\r. r \ carrier R \ P r" and "\r1 r2. \ r1 \ carrier R; r2 \ carrier R \ \ P (r1 \ r2)" and "\r1 r2. \ r1 \ carrier R; r2 \ carrier R \ \ P (r1 \ r2)" using assms unfolding morphic_prop_def by auto (* NEW ============================================================================ *) lemma (in ring) ring_hom_restrict: assumes "f \ ring_hom R S" and "\r. r \ carrier R \ f r = g r" shows "g \ ring_hom R S" using assms(2) ring_hom_memE[OF assms(1)] by (auto intro: ring_hom_memI) (* PROOF ========================================================================== *) lemma (in ring) ring_iso_restrict: assumes "f \ ring_iso R S" and "\r. r \ carrier R \ f r = g r" shows "g \ ring_iso R S" proof - have hom: "g \ ring_hom R S" using ring_hom_restrict assms unfolding ring_iso_def by auto have "bij_betw g (carrier R) (carrier S)" using bij_betw_cong[of "carrier R" f g] ring_iso_memE(5)[OF assms(1)] assms(2) by simp thus ?thesis using ring_hom_memE[OF hom] by (auto intro!: ring_iso_memI) qed lemma ring_iso_morphic_prop: assumes "f \ ring_iso R S" and "morphic_prop R P" and "\r. P r \ f r = g r" shows "g \ ring_iso R S" proof - have eq0: "\r. r \ carrier R \ f r = g r" and eq1: "f \\<^bsub>R\<^esub> = g \\<^bsub>R\<^esub>" and eq2: "\r1 r2. \ r1 \ carrier R; r2 \ carrier R \ \ f (r1 \\<^bsub>R\<^esub> r2) = g (r1 \\<^bsub>R\<^esub> r2)" and eq3: "\r1 r2. \ r1 \ carrier R; r2 \ carrier R \ \ f (r1 \\<^bsub>R\<^esub> r2) = g (r1 \\<^bsub>R\<^esub> r2)" using assms(2-3) unfolding morphic_prop_def by auto show ?thesis apply (rule ring_iso_memI) using assms(1) eq0 ring_iso_memE(1) apply fastforce apply (metis assms(1) eq0 eq2 ring_iso_memE(2)) apply (metis assms(1) eq0 eq3 ring_iso_memE(3)) using assms(1) eq1 ring_iso_memE(4) apply fastforce using assms(1) bij_betw_cong eq0 ring_iso_memE(5) by blast qed lemma (in ring) ring_hom_imp_img_ring: assumes "h \ ring_hom R S" shows "ring (S \ carrier := h ` (carrier R), zero := h \ \)" (is "ring ?h_img") proof - have "h \ hom (add_monoid R) (add_monoid S)" using assms unfolding hom_def ring_hom_def by auto hence "comm_group ((add_monoid S) \ carrier := h ` (carrier R), one := h \ \)" using add.hom_imp_img_comm_group[of h "add_monoid S"] by simp hence comm_group: "comm_group (add_monoid ?h_img)" by (auto intro: comm_monoidI simp add: monoid.defs) moreover have "h \ hom R S" using assms unfolding ring_hom_def hom_def by auto hence "monoid (S \ carrier := h ` (carrier R), one := h \ \)" using hom_imp_img_monoid[of h S] by simp hence monoid: "monoid ?h_img" using ring_hom_memE(4)[OF assms] unfolding monoid_def by (simp add: monoid.defs) show ?thesis proof (rule ringI, simp_all add: comm_group_abelian_groupI[OF comm_group] monoid) fix x y z assume "x \ h ` carrier R" "y \ h ` carrier R" "z \ h ` carrier R" then obtain r1 r2 r3 where r1: "r1 \ carrier R" "x = h r1" and r2: "r2 \ carrier R" "y = h r2" and r3: "r3 \ carrier R" "z = h r3" by blast hence "(x \\<^bsub>S\<^esub> y) \\<^bsub>S\<^esub> z = h ((r1 \ r2) \ r3)" using ring_hom_memE[OF assms] by auto also have " ... = h ((r1 \ r3) \ (r2 \ r3))" using l_distr[OF r1(1) r2(1) r3(1)] by simp also have " ... = (x \\<^bsub>S\<^esub> z) \\<^bsub>S\<^esub> (y \\<^bsub>S\<^esub> z)" using ring_hom_memE[OF assms] r1 r2 r3 by auto finally show "(x \\<^bsub>S\<^esub> y) \\<^bsub>S\<^esub> z = (x \\<^bsub>S\<^esub> z) \\<^bsub>S\<^esub> (y \\<^bsub>S\<^esub> z)" . have "z \\<^bsub>S\<^esub> (x \\<^bsub>S\<^esub> y) = h (r3 \ (r1 \ r2))" using ring_hom_memE[OF assms] r1 r2 r3 by auto also have " ... = h ((r3 \ r1) \ (r3 \ r2))" using r_distr[OF r1(1) r2(1) r3(1)] by simp also have " ... = (z \\<^bsub>S\<^esub> x) \\<^bsub>S\<^esub> (z \\<^bsub>S\<^esub> y)" using ring_hom_memE[OF assms] r1 r2 r3 by auto finally show "z \\<^bsub>S\<^esub> (x \\<^bsub>S\<^esub> y) = (z \\<^bsub>S\<^esub> x) \\<^bsub>S\<^esub> (z \\<^bsub>S\<^esub> y)" . qed qed lemma (in ring) ring_iso_imp_img_ring: assumes "h \ ring_iso R S" shows "ring (S \ zero := h \ \)" proof - have "ring (S \ carrier := h ` (carrier R), zero := h \ \)" using ring_hom_imp_img_ring[of h S] assms unfolding ring_iso_def by auto moreover have "h ` (carrier R) = carrier S" using assms unfolding ring_iso_def bij_betw_def by auto ultimately show ?thesis by simp qed lemma (in cring) ring_iso_imp_img_cring: assumes "h \ ring_iso R S" shows "cring (S \ zero := h \ \)" (is "cring ?h_img") proof - note m_comm interpret h_img?: ring ?h_img using ring_iso_imp_img_ring[OF assms] . show ?thesis proof (unfold_locales) fix x y assume "x \ carrier ?h_img" "y \ carrier ?h_img" then obtain r1 r2 where r1: "r1 \ carrier R" "x = h r1" and r2: "r2 \ carrier R" "y = h r2" using assms image_iff[where ?f = h and ?A = "carrier R"] unfolding ring_iso_def bij_betw_def by auto have "x \\<^bsub>(?h_img)\<^esub> y = h (r1 \ r2)" using assms r1 r2 unfolding ring_iso_def ring_hom_def by auto also have " ... = h (r2 \ r1)" using m_comm[OF r1(1) r2(1)] by simp also have " ... = y \\<^bsub>(?h_img)\<^esub> x" using assms r1 r2 unfolding ring_iso_def ring_hom_def by auto finally show "x \\<^bsub>(?h_img)\<^esub> y = y \\<^bsub>(?h_img)\<^esub> x" . qed qed lemma (in domain) ring_iso_imp_img_domain: assumes "h \ ring_iso R S" shows "domain (S \ zero := h \ \)" (is "domain ?h_img") proof - note aux = m_closed integral one_not_zero one_closed zero_closed interpret h_img?: cring ?h_img using ring_iso_imp_img_cring[OF assms] . show ?thesis proof (unfold_locales) have "\\<^bsub>?h_img\<^esub> = \\<^bsub>?h_img\<^esub> \ h \ = h \" using ring_iso_memE(4)[OF assms] by simp moreover have "h \ \ h \" using ring_iso_memE(5)[OF assms] aux(3-4) unfolding bij_betw_def inj_on_def by force ultimately show "\\<^bsub>?h_img\<^esub> \ \\<^bsub>?h_img\<^esub>" by auto next fix a b assume A: "a \\<^bsub>?h_img\<^esub> b = \\<^bsub>?h_img\<^esub>" "a \ carrier ?h_img" "b \ carrier ?h_img" then obtain r1 r2 where r1: "r1 \ carrier R" "a = h r1" and r2: "r2 \ carrier R" "b = h r2" using assms image_iff[where ?f = h and ?A = "carrier R"] unfolding ring_iso_def bij_betw_def by auto hence "a \\<^bsub>?h_img\<^esub> b = h (r1 \ r2)" using assms r1 r2 unfolding ring_iso_def ring_hom_def by auto hence "h (r1 \ r2) = h \" using A(1) by simp hence "r1 \ r2 = \" using ring_iso_memE(5)[OF assms] aux(1)[OF r1(1) r2(1)] aux(5) unfolding bij_betw_def inj_on_def by force hence "r1 = \ \ r2 = \" using aux(2)[OF _ r1(1) r2(1)] by simp thus "a = \\<^bsub>?h_img\<^esub> \ b = \\<^bsub>?h_img\<^esub>" unfolding r1 r2 by auto qed qed lemma (in field) ring_iso_imp_img_field: assumes "h \ ring_iso R S" shows "field (S \ zero := h \ \)" (is "field ?h_img") proof - interpret h_img?: domain ?h_img using ring_iso_imp_img_domain[OF assms] . show ?thesis proof (unfold_locales, auto simp add: Units_def) interpret field R using field_axioms . fix a assume a: "a \ carrier S" "a \\<^bsub>S\<^esub> h \ = \\<^bsub>S\<^esub>" then obtain r where r: "r \ carrier R" "a = h r" using assms image_iff[where ?f = h and ?A = "carrier R"] unfolding ring_iso_def bij_betw_def by auto have "a \\<^bsub>S\<^esub> h \ = h (r \ \)" unfolding r(2) using ring_iso_memE(2)[OF assms r(1)] by simp hence "h \ = h \" using ring_iso_memE(4)[OF assms] r(1) a(2) by simp thus False using ring_iso_memE(5)[OF assms] unfolding bij_betw_def inj_on_def by force next interpret field R using field_axioms . fix s assume s: "s \ carrier S" "s \ h \" then obtain r where r: "r \ carrier R" "s = h r" using assms image_iff[where ?f = h and ?A = "carrier R"] unfolding ring_iso_def bij_betw_def by auto hence "r \ \" using s(2) by auto hence inv_r: "inv r \ carrier R" "inv r \ \" "r \ inv r = \" "inv r \ r = \" using field_Units r(1) by auto have "h (inv r) \\<^bsub>S\<^esub> h r = h \" and "h r \\<^bsub>S\<^esub> h (inv r) = h \" using ring_iso_memE(2)[OF assms inv_r(1) r(1)] inv_r(3-4) ring_iso_memE(2)[OF assms r(1) inv_r(1)] by auto thus "\s' \ carrier S. s' \\<^bsub>S\<^esub> s = \\<^bsub>S\<^esub> \ s \\<^bsub>S\<^esub> s' = \\<^bsub>S\<^esub>" using ring_iso_memE(1,4)[OF assms] inv_r(1) r(2) by auto qed qed lemma ring_iso_same_card: "R \ S \ card (carrier R) = card (carrier S)" using bij_betw_same_card unfolding is_ring_iso_def ring_iso_def by auto (* ========================================================================== *) lemma ring_iso_set_refl: "id \ ring_iso R R" by (rule ring_iso_memI) (auto) corollary ring_iso_refl: "R \ R" using is_ring_iso_def ring_iso_set_refl by auto lemma ring_iso_set_trans: "\ f \ ring_iso R S; g \ ring_iso S Q \ \ (g \ f) \ ring_iso R Q" unfolding ring_iso_def using bij_betw_trans ring_hom_trans by fastforce corollary ring_iso_trans: "\ R \ S; S \ Q \ \ R \ Q" using ring_iso_set_trans unfolding is_ring_iso_def by blast lemma ring_iso_set_sym: assumes "ring R" and h: "h \ ring_iso R S" shows "(inv_into (carrier R) h) \ ring_iso S R" proof - have h_hom: "h \ ring_hom R S" and h_surj: "h ` (carrier R) = (carrier S)" and h_inj: "\ x1 x2. \ x1 \ carrier R; x2 \ carrier R \ \ h x1 = h x2 \ x1 = x2" using h unfolding ring_iso_def bij_betw_def inj_on_def by auto have h_inv_bij: "bij_betw (inv_into (carrier R) h) (carrier S) (carrier R)" - using bij_betw_inv_into h ring_iso_def by fastforce + by (simp add: bij_betw_inv_into h ring_iso_memE(5)) - show "inv_into (carrier R) h \ ring_iso S R" - apply (rule ring_iso_memI) - apply (simp add: h_surj inv_into_into) - apply (auto simp add: h_inv_bij) - using ring_iso_memE [OF h] bij_betwE [OF h_inv_bij] - apply (simp_all add: \ring R\ bij_betw_def bij_betw_inv_into_right inv_into_f_eq ring.ring_simprules(5)) - using ring_iso_memE [OF h] bij_betw_inv_into_right [of h "carrier R" "carrier S"] - apply (simp add: \ring R\ inv_into_f_eq ring.ring_simprules(1)) - by (simp add: \ring R\ inv_into_f_eq ring.ring_simprules(6)) + have "inv_into (carrier R) h \ ring_hom S R" + using ring_iso_memE [OF h] bij_betwE [OF h_inv_bij] \ring R\ + by (simp add: bij_betw_imp_inj_on bij_betw_inv_into_right inv_into_f_eq ring.ring_simprules ring_hom_memI) + moreover have "bij_betw (inv_into (carrier R) h) (carrier S) (carrier R)" + using h_inv_bij by force + ultimately show "inv_into (carrier R) h \ ring_iso S R" + by (simp add: ring_iso_def) qed corollary ring_iso_sym: assumes "ring R" shows "R \ S \ S \ R" using assms ring_iso_set_sym unfolding is_ring_iso_def by auto lemma (in ring_hom_ring) the_elem_simp [simp]: "\x. x \ carrier R \ the_elem (h ` ((a_kernel R S h) +> x)) = h x" proof - fix x assume x: "x \ carrier R" hence "h x \ h ` ((a_kernel R S h) +> x)" using homeq_imp_rcos by blast thus "the_elem (h ` ((a_kernel R S h) +> x)) = h x" by (metis (no_types, lifting) x empty_iff homeq_imp_rcos rcos_imp_homeq the_elem_image_unique) qed lemma (in ring_hom_ring) the_elem_inj: "\X Y. \ X \ carrier (R Quot (a_kernel R S h)); Y \ carrier (R Quot (a_kernel R S h)) \ \ the_elem (h ` X) = the_elem (h ` Y) \ X = Y" proof - fix X Y assume "X \ carrier (R Quot (a_kernel R S h))" and "Y \ carrier (R Quot (a_kernel R S h))" and Eq: "the_elem (h ` X) = the_elem (h ` Y)" then obtain x y where x: "x \ carrier R" "X = (a_kernel R S h) +> x" and y: "y \ carrier R" "Y = (a_kernel R S h) +> y" unfolding FactRing_def A_RCOSETS_def' by auto hence "h x = h y" using Eq by simp hence "x \ y \ (a_kernel R S h)" by (simp add: a_minus_def abelian_subgroup.a_rcos_module_imp abelian_subgroup_a_kernel homeq_imp_rcos x(1) y(1)) thus "X = Y" by (metis R.a_coset_add_inv1 R.minus_eq abelian_subgroup.a_rcos_const abelian_subgroup_a_kernel additive_subgroup.a_subset additive_subgroup_a_kernel x y) qed lemma (in ring_hom_ring) quot_mem: "\X. X \ carrier (R Quot (a_kernel R S h)) \ \x \ carrier R. X = (a_kernel R S h) +> x" proof - fix X assume "X \ carrier (R Quot (a_kernel R S h))" thus "\x \ carrier R. X = (a_kernel R S h) +> x" unfolding FactRing_simps by (simp add: a_r_coset_def) qed lemma (in ring_hom_ring) the_elem_wf: "\X. X \ carrier (R Quot (a_kernel R S h)) \ \y \ carrier S. (h ` X) = { y }" proof - fix X assume "X \ carrier (R Quot (a_kernel R S h))" then obtain x where x: "x \ carrier R" and X: "X = (a_kernel R S h) +> x" using quot_mem by blast hence "\x'. x' \ X \ h x' = h x" proof - fix x' assume "x' \ X" hence "x' \ (a_kernel R S h) +> x" using X by simp then obtain k where k: "k \ a_kernel R S h" "x' = k \ x" by (metis R.add.inv_closed R.add.m_assoc R.l_neg R.r_zero abelian_subgroup.a_elemrcos_carrier abelian_subgroup.a_rcos_module_imp abelian_subgroup_a_kernel x) hence "h x' = h k \\<^bsub>S\<^esub> h x" by (meson additive_subgroup.a_Hcarr additive_subgroup_a_kernel hom_add x) also have " ... = h x" using k by (auto simp add: x) finally show "h x' = h x" . qed moreover have "h x \ h ` X" by (simp add: X homeq_imp_rcos x) ultimately have "(h ` X) = { h x }" by blast thus "\y \ carrier S. (h ` X) = { y }" using x by simp qed corollary (in ring_hom_ring) the_elem_wf': "\X. X \ carrier (R Quot (a_kernel R S h)) \ \r \ carrier R. (h ` X) = { h r }" using the_elem_wf by (metis quot_mem the_elem_eq the_elem_simp) lemma (in ring_hom_ring) the_elem_hom: "(\X. the_elem (h ` X)) \ ring_hom (R Quot (a_kernel R S h)) S" proof (rule ring_hom_memI) show "\x. x \ carrier (R Quot a_kernel R S h) \ the_elem (h ` x) \ carrier S" using the_elem_wf by fastforce show "the_elem (h ` \\<^bsub>R Quot a_kernel R S h\<^esub>) = \\<^bsub>S\<^esub>" unfolding FactRing_def using the_elem_simp[of "\\<^bsub>R\<^esub>"] by simp fix X Y assume "X \ carrier (R Quot a_kernel R S h)" and "Y \ carrier (R Quot a_kernel R S h)" then obtain x y where x: "x \ carrier R" "X = (a_kernel R S h) +> x" and y: "y \ carrier R" "Y = (a_kernel R S h) +> y" using quot_mem by blast have "X \\<^bsub>R Quot a_kernel R S h\<^esub> Y = (a_kernel R S h) +> (x \ y)" by (simp add: FactRing_def ideal.rcoset_mult_add kernel_is_ideal x y) thus "the_elem (h ` (X \\<^bsub>R Quot a_kernel R S h\<^esub> Y)) = the_elem (h ` X) \\<^bsub>S\<^esub> the_elem (h ` Y)" by (simp add: x y) have "X \\<^bsub>R Quot a_kernel R S h\<^esub> Y = (a_kernel R S h) +> (x \ y)" using ideal.rcos_ring_hom kernel_is_ideal ring_hom_add x y by fastforce thus "the_elem (h ` (X \\<^bsub>R Quot a_kernel R S h\<^esub> Y)) = the_elem (h ` X) \\<^bsub>S\<^esub> the_elem (h ` Y)" by (simp add: x y) qed lemma (in ring_hom_ring) the_elem_surj: "(\X. (the_elem (h ` X))) ` carrier (R Quot (a_kernel R S h)) = (h ` (carrier R))" proof show "(\X. the_elem (h ` X)) ` carrier (R Quot a_kernel R S h) \ h ` carrier R" using the_elem_wf' by fastforce next show "h ` carrier R \ (\X. the_elem (h ` X)) ` carrier (R Quot a_kernel R S h)" proof fix y assume "y \ h ` carrier R" then obtain x where x: "x \ carrier R" "h x = y" by (metis image_iff) hence "the_elem (h ` ((a_kernel R S h) +> x)) = y" by simp moreover have "(a_kernel R S h) +> x \ carrier (R Quot (a_kernel R S h))" unfolding FactRing_simps by (auto simp add: x a_r_coset_def) ultimately show "y \ (\X. (the_elem (h ` X))) ` carrier (R Quot (a_kernel R S h))" by blast qed qed proposition (in ring_hom_ring) FactRing_iso_set_aux: "(\X. the_elem (h ` X)) \ ring_iso (R Quot (a_kernel R S h)) (S \ carrier := h ` (carrier R) \)" proof - have "bij_betw (\X. the_elem (h ` X)) (carrier (R Quot a_kernel R S h)) (h ` (carrier R))" unfolding bij_betw_def inj_on_def using the_elem_surj the_elem_inj by simp moreover have "(\X. the_elem (h ` X)): carrier (R Quot (a_kernel R S h)) \ h ` (carrier R)" using the_elem_wf' by fastforce hence "(\X. the_elem (h ` X)) \ ring_hom (R Quot (a_kernel R S h)) (S \ carrier := h ` (carrier R) \)" using the_elem_hom the_elem_wf' unfolding ring_hom_def by simp ultimately show ?thesis unfolding ring_iso_def using the_elem_hom by simp qed theorem (in ring_hom_ring) FactRing_iso_set: assumes "h ` carrier R = carrier S" shows "(\X. the_elem (h ` X)) \ ring_iso (R Quot (a_kernel R S h)) S" using FactRing_iso_set_aux assms by auto corollary (in ring_hom_ring) FactRing_iso: assumes "h ` carrier R = carrier S" shows "R Quot (a_kernel R S h) \ S" using FactRing_iso_set assms is_ring_iso_def by auto corollary (in ring) FactRing_zeroideal: shows "R Quot { \ } \ R" and "R \ R Quot { \ }" proof - have "ring_hom_ring R R id" using ring_axioms by (auto intro: ring_hom_ringI) moreover have "a_kernel R R id = { \ }" unfolding a_kernel_def' by auto ultimately show "R Quot { \ } \ R" and "R \ R Quot { \ }" using ring_hom_ring.FactRing_iso[of R R id] ring_iso_sym[OF ideal.quotient_is_ring[OF zeroideal], of R] by auto qed lemma (in ring_hom_ring) img_is_ring: "ring (S \ carrier := h ` (carrier R) \)" proof - let ?the_elem = "\X. the_elem (h ` X)" have FactRing_is_ring: "ring (R Quot (a_kernel R S h))" by (simp add: ideal.quotient_is_ring kernel_is_ideal) have "ring ((S \ carrier := ?the_elem ` (carrier (R Quot (a_kernel R S h))) \) \ zero := ?the_elem \\<^bsub>(R Quot (a_kernel R S h))\<^esub> \)" using ring.ring_iso_imp_img_ring[OF FactRing_is_ring, of ?the_elem "S \ carrier := ?the_elem ` (carrier (R Quot (a_kernel R S h))) \"] FactRing_iso_set_aux the_elem_surj by auto moreover have "\ \ (a_kernel R S h)" using a_kernel_def'[of R S h] by auto hence "\ \ (a_kernel R S h) +> \" using a_r_coset_def'[of R "a_kernel R S h" \] by force hence "\\<^bsub>S\<^esub> \ (h ` ((a_kernel R S h) +> \))" using hom_one by force hence "?the_elem \\<^bsub>(R Quot (a_kernel R S h))\<^esub> = \\<^bsub>S\<^esub>" using the_elem_wf[of "(a_kernel R S h) +> \"] by (simp add: FactRing_def) moreover have "\\<^bsub>S\<^esub> \ (h ` (a_kernel R S h))" using a_kernel_def'[of R S h] hom_zero by force hence "\\<^bsub>S\<^esub> \ (h ` \\<^bsub>(R Quot (a_kernel R S h))\<^esub>)" by (simp add: FactRing_def) hence "?the_elem \\<^bsub>(R Quot (a_kernel R S h))\<^esub> = \\<^bsub>S\<^esub>" using the_elem_wf[OF ring.ring_simprules(2)[OF FactRing_is_ring]] by (metis singletonD the_elem_eq) ultimately have "ring ((S \ carrier := h ` (carrier R) \) \ one := \\<^bsub>S\<^esub>, zero := \\<^bsub>S\<^esub> \)" using the_elem_surj by simp thus ?thesis by auto qed lemma (in ring_hom_ring) img_is_cring: assumes "cring S" shows "cring (S \ carrier := h ` (carrier R) \)" proof - interpret ring "S \ carrier := h ` (carrier R) \" using img_is_ring . show ?thesis apply unfold_locales using assms unfolding cring_def comm_monoid_def comm_monoid_axioms_def by auto qed lemma (in ring_hom_ring) img_is_domain: assumes "domain S" shows "domain (S \ carrier := h ` (carrier R) \)" proof - interpret cring "S \ carrier := h ` (carrier R) \" using img_is_cring assms unfolding domain_def by simp show ?thesis apply unfold_locales using assms unfolding domain_def domain_axioms_def apply auto using hom_closed by blast qed proposition (in ring_hom_ring) primeideal_vimage: assumes "cring R" shows "primeideal P S \ primeideal { r \ carrier R. h r \ P } R" proof - assume A: "primeideal P S" hence is_ideal: "ideal P S" unfolding primeideal_def by simp have "ring_hom_ring R (S Quot P) (((+>\<^bsub>S\<^esub>) P) \ h)" (is "ring_hom_ring ?A ?B ?h") using ring_hom_trans[OF homh, of "(+>\<^bsub>S\<^esub>) P" "S Quot P"] ideal.rcos_ring_hom_ring[OF is_ideal] assms unfolding ring_hom_ring_def ring_hom_ring_axioms_def cring_def by simp then interpret hom: ring_hom_ring R "S Quot P" "((+>\<^bsub>S\<^esub>) P) \ h" by simp have "inj_on (\X. the_elem (?h ` X)) (carrier (R Quot (a_kernel R (S Quot P) ?h)))" using hom.the_elem_inj unfolding inj_on_def by simp moreover have "ideal (a_kernel R (S Quot P) ?h) R" using hom.kernel_is_ideal by auto have hom': "ring_hom_ring (R Quot (a_kernel R (S Quot P) ?h)) (S Quot P) (\X. the_elem (?h ` X))" using hom.the_elem_hom hom.kernel_is_ideal by (meson hom.ring_hom_ring_axioms ideal.rcos_ring_hom_ring ring_hom_ring_axioms_def ring_hom_ring_def) ultimately have "primeideal (a_kernel R (S Quot P) ?h) R" using ring_hom_ring.inj_on_domain[OF hom'] primeideal.quotient_is_domain[OF A] cring.quot_domain_imp_primeideal[OF assms hom.kernel_is_ideal] by simp moreover have "a_kernel R (S Quot P) ?h = { r \ carrier R. h r \ P }" proof show "a_kernel R (S Quot P) ?h \ { r \ carrier R. h r \ P }" proof fix r assume "r \ a_kernel R (S Quot P) ?h" hence r: "r \ carrier R" "P +>\<^bsub>S\<^esub> (h r) = P" unfolding a_kernel_def kernel_def FactRing_def by auto hence "h r \ P" using S.a_rcosI R.l_zero S.l_zero additive_subgroup.a_subset[OF ideal.axioms(1)[OF is_ideal]] additive_subgroup.zero_closed[OF ideal.axioms(1)[OF is_ideal]] hom_closed by metis thus "r \ { r \ carrier R. h r \ P }" using r by simp qed next show "{ r \ carrier R. h r \ P } \ a_kernel R (S Quot P) ?h" proof fix r assume "r \ { r \ carrier R. h r \ P }" hence r: "r \ carrier R" "h r \ P" by simp_all hence "?h r = P" by (simp add: S.a_coset_join2 additive_subgroup.a_subgroup ideal.axioms(1) is_ideal) thus "r \ a_kernel R (S Quot P) ?h" unfolding a_kernel_def kernel_def FactRing_def using r(1) by auto qed qed ultimately show "primeideal { r \ carrier R. h r \ P } R" by simp qed end diff --git a/src/HOL/Analysis/Abstract_Topology_2.thy b/src/HOL/Analysis/Abstract_Topology_2.thy --- a/src/HOL/Analysis/Abstract_Topology_2.thy +++ b/src/HOL/Analysis/Abstract_Topology_2.thy @@ -1,1580 +1,1594 @@ (* Author: L C Paulson, University of Cambridge Author: Amine Chaieb, University of Cambridge Author: Robert Himmelmann, TU Muenchen Author: Brian Huffman, Portland State University *) section \Abstract Topology 2\ theory Abstract_Topology_2 imports Elementary_Topology Abstract_Topology "HOL-Library.Indicator_Function" begin text \Combination of Elementary and Abstract Topology\ lemma approachable_lt_le2: "(\(d::real) > 0. \x. Q x \ f x < d \ P x) \ (\d>0. \x. f x \ d \ Q x \ P x)" by (meson dense less_eq_real_def order_le_less_trans) lemma triangle_lemma: fixes x y z :: real assumes x: "0 \ x" and y: "0 \ y" and z: "0 \ z" and xy: "x\<^sup>2 \ y\<^sup>2 + z\<^sup>2" shows "x \ y + z" proof - have "y\<^sup>2 + z\<^sup>2 \ y\<^sup>2 + 2 * y * z + z\<^sup>2" using z y by simp with xy have th: "x\<^sup>2 \ (y + z)\<^sup>2" by (simp add: power2_eq_square field_simps) from y z have yz: "y + z \ 0" by arith from power2_le_imp_le[OF th yz] show ?thesis . qed lemma isCont_indicator: fixes x :: "'a::t2_space" shows "isCont (indicator A :: 'a \ real) x = (x \ frontier A)" proof auto fix x assume cts_at: "isCont (indicator A :: 'a \ real) x" and fr: "x \ frontier A" with continuous_at_open have 1: "\V::real set. open V \ indicator A x \ V \ (\U::'a set. open U \ x \ U \ (\y\U. indicator A y \ V))" by auto show False proof (cases "x \ A") assume x: "x \ A" hence "indicator A x \ ({0<..<2} :: real set)" by simp with 1 obtain U where U: "open U" "x \ U" "\y\U. indicator A y \ ({0<..<2} :: real set)" using open_greaterThanLessThan by metis hence "\y\U. indicator A y > (0::real)" unfolding greaterThanLessThan_def by auto hence "U \ A" using indicator_eq_0_iff by force hence "x \ interior A" using U interiorI by auto thus ?thesis using fr unfolding frontier_def by simp next assume x: "x \ A" hence "indicator A x \ ({-1<..<1} :: real set)" by simp with 1 obtain U where U: "open U" "x \ U" "\y\U. indicator A y \ ({-1<..<1} :: real set)" using 1 open_greaterThanLessThan by metis hence "\y\U. indicator A y < (1::real)" unfolding greaterThanLessThan_def by auto hence "U \ -A" by auto hence "x \ interior (-A)" using U interiorI by auto thus ?thesis using fr interior_complement unfolding frontier_def by auto qed next assume nfr: "x \ frontier A" hence "x \ interior A \ x \ interior (-A)" by (auto simp: frontier_def closure_interior) thus "isCont ((indicator A)::'a \ real) x" proof assume int: "x \ interior A" then obtain U where U: "open U" "x \ U" "U \ A" unfolding interior_def by auto hence "\y\U. indicator A y = (1::real)" unfolding indicator_def by auto hence "continuous_on U (indicator A)" by (simp add: indicator_eq_1_iff) thus ?thesis using U continuous_on_eq_continuous_at by auto next assume ext: "x \ interior (-A)" then obtain U where U: "open U" "x \ U" "U \ -A" unfolding interior_def by auto then have "continuous_on U (indicator A)" using continuous_on_topological by (auto simp: subset_iff) thus ?thesis using U continuous_on_eq_continuous_at by auto qed qed lemma islimpt_closure: "\S \ T; \x. \x islimpt S; x \ T\ \ x \ S\ \ S = T \ closure S" using closure_def by fastforce lemma closedin_limpt: "closedin (top_of_set T) S \ S \ T \ (\x. x islimpt S \ x \ T \ x \ S)" proof - have "\U x. \closed U; S = T \ U; x islimpt S; x \ T\ \ x \ S" by (meson IntI closed_limpt inf_le2 islimpt_subset) then show ?thesis by (metis closed_closure closedin_closed closedin_imp_subset islimpt_closure) qed lemma closedin_closed_eq: "closed S \ closedin (top_of_set S) T \ closed T \ T \ S" by (meson closedin_limpt closed_subset closedin_closed_trans) lemma connected_closed_set: "closed S \ connected S \ (\A B. closed A \ closed B \ A \ {} \ B \ {} \ A \ B = S \ A \ B = {})" unfolding connected_closedin_eq closedin_closed_eq connected_closedin_eq by blast text \If a connnected set is written as the union of two nonempty closed sets, then these sets have to intersect.\ lemma connected_as_closed_union: assumes "connected C" "C = A \ B" "closed A" "closed B" "A \ {}" "B \ {}" shows "A \ B \ {}" by (metis assms closed_Un connected_closed_set) lemma closedin_subset_trans: "closedin (top_of_set U) S \ S \ T \ T \ U \ closedin (top_of_set T) S" by (meson closedin_limpt subset_iff) lemma openin_subset_trans: "openin (top_of_set U) S \ S \ T \ T \ U \ openin (top_of_set T) S" by (auto simp: openin_open) lemma closedin_compact: "\compact S; closedin (top_of_set S) T\ \ compact T" by (metis closedin_closed compact_Int_closed) lemma closedin_compact_eq: fixes S :: "'a::t2_space set" shows "compact S \ (closedin (top_of_set S) T \ compact T \ T \ S)" by (metis closedin_imp_subset closedin_compact closed_subset compact_imp_closed) subsection \Closure\ lemma euclidean_closure_of [simp]: "euclidean closure_of S = closure S" by (auto simp: closure_of_def closure_def islimpt_def) lemma closure_openin_Int_closure: assumes ope: "openin (top_of_set U) S" and "T \ U" shows "closure(S \ closure T) = closure(S \ T)" proof obtain V where "open V" and S: "S = U \ V" using ope using openin_open by metis show "closure (S \ closure T) \ closure (S \ T)" proof (clarsimp simp: S) fix x assume "x \ closure (U \ V \ closure T)" then have "V \ closure T \ A \ x \ closure A" for A by (metis closure_mono subsetD inf.coboundedI2 inf_assoc) then have "x \ closure (T \ V)" by (metis \open V\ closure_closure inf_commute open_Int_closure_subset) then show "x \ closure (U \ V \ T)" by (metis \T \ U\ inf.absorb_iff2 inf_assoc inf_commute) qed next show "closure (S \ T) \ closure (S \ closure T)" by (meson Int_mono closure_mono closure_subset order_refl) qed corollary infinite_openin: fixes S :: "'a :: t1_space set" shows "\openin (top_of_set U) S; x \ S; x islimpt U\ \ infinite S" by (clarsimp simp add: openin_open islimpt_eq_acc_point inf_commute) lemma closure_Int_ballI: assumes "\U. \openin (top_of_set S) U; U \ {}\ \ T \ U \ {}" shows "S \ closure T" proof (clarsimp simp: closure_iff_nhds_not_empty) fix x and A and V assume "x \ S" "V \ A" "open V" "x \ V" "T \ A = {}" then have "openin (top_of_set S) (A \ V \ S)" by (simp add: inf_absorb2 openin_subtopology_Int) moreover have "A \ V \ S \ {}" using \x \ V\ \V \ A\ \x \ S\ by auto ultimately show False using \T \ A = {}\ assms by fastforce qed subsection \Frontier\ lemma euclidean_interior_of [simp]: "euclidean interior_of S = interior S" by (auto simp: interior_of_def interior_def) lemma euclidean_frontier_of [simp]: "euclidean frontier_of S = frontier S" by (auto simp: frontier_of_def frontier_def) lemma connected_Int_frontier: "\connected S; S \ T \ {}; S - T \ {}\ \ S \ frontier T \ {}" apply (simp add: frontier_interiors connected_openin, safe) apply (drule_tac x="S \ interior T" in spec, safe) apply (drule_tac [2] x="S \ interior (-T)" in spec) apply (auto simp: disjoint_eq_subset_Compl dest: interior_subset [THEN subsetD]) done subsection \Compactness\ lemma openin_delete: fixes a :: "'a :: t1_space" shows "openin (top_of_set u) S \ openin (top_of_set u) (S - {a})" by (metis Int_Diff open_delete openin_open) lemma compact_eq_openin_cover: "compact S \ (\C. (\c\C. openin (top_of_set S) c) \ S \ \C \ (\D\C. finite D \ S \ \D))" proof safe fix C assume "compact S" and "\c\C. openin (top_of_set S) c" and "S \ \C" then have "\c\{T. open T \ S \ T \ C}. open c" and "S \ \{T. open T \ S \ T \ C}" unfolding openin_open by force+ with \compact S\ obtain D where "D \ {T. open T \ S \ T \ C}" and "finite D" and "S \ \D" by (meson compactE) then have "image (\T. S \ T) D \ C \ finite (image (\T. S \ T) D) \ S \ \(image (\T. S \ T) D)" by auto then show "\D\C. finite D \ S \ \D" .. next assume 1: "\C. (\c\C. openin (top_of_set S) c) \ S \ \C \ (\D\C. finite D \ S \ \D)" show "compact S" proof (rule compactI) fix C let ?C = "image (\T. S \ T) C" assume "\t\C. open t" and "S \ \C" then have "(\c\?C. openin (top_of_set S) c) \ S \ \?C" unfolding openin_open by auto with 1 obtain D where "D \ ?C" and "finite D" and "S \ \D" by metis let ?D = "inv_into C (\T. S \ T) ` D" have "?D \ C \ finite ?D \ S \ \?D" proof (intro conjI) from \D \ ?C\ show "?D \ C" by (fast intro: inv_into_into) from \finite D\ show "finite ?D" by (rule finite_imageI) from \S \ \D\ show "S \ \?D" by (metis \D \ (\) S ` C\ image_inv_into_cancel inf_Sup le_infE) qed then show "\D\C. finite D \ S \ \D" .. qed qed subsection \Continuity\ lemma interior_image_subset: assumes "inj f" "\x. continuous (at x) f" shows "interior (f ` S) \ f ` (interior S)" proof fix x assume "x \ interior (f ` S)" then obtain T where as: "open T" "x \ T" "T \ f ` S" .. then have "x \ f ` S" by auto then obtain y where y: "y \ S" "x = f y" by auto have "open (f -` T)" using assms \open T\ by (simp add: continuous_at_imp_continuous_on open_vimage) moreover have "y \ vimage f T" using \x = f y\ \x \ T\ by simp moreover have "vimage f T \ S" using \T \ image f S\ \inj f\ unfolding inj_on_def subset_eq by auto ultimately have "y \ interior S" .. with \x = f y\ show "x \ f ` interior S" .. qed subsection\<^marker>\tag unimportant\ \Equality of continuous functions on closure and related results\ lemma continuous_closedin_preimage_constant: fixes f :: "_ \ 'b::t1_space" shows "continuous_on S f \ closedin (top_of_set S) {x \ S. f x = a}" using continuous_closedin_preimage[of S f "{a}"] by (simp add: vimage_def Collect_conj_eq) lemma continuous_closed_preimage_constant: fixes f :: "_ \ 'b::t1_space" shows "continuous_on S f \ closed S \ closed {x \ S. f x = a}" using continuous_closed_preimage[of S f "{a}"] by (simp add: vimage_def Collect_conj_eq) lemma continuous_constant_on_closure: fixes f :: "_ \ 'b::t1_space" assumes "continuous_on (closure S) f" and "\x. x \ S \ f x = a" and "x \ closure S" shows "f x = a" using continuous_closed_preimage_constant[of "closure S" f a] assms closure_minimal[of S "{x \ closure S. f x = a}"] closure_subset unfolding subset_eq by auto lemma image_closure_subset: assumes contf: "continuous_on (closure S) f" and "closed T" and "(f ` S) \ T" shows "f ` (closure S) \ T" proof - have "S \ {x \ closure S. f x \ T}" using assms(3) closure_subset by auto moreover have "closed (closure S \ f -` T)" using continuous_closed_preimage[OF contf] \closed T\ by auto ultimately have "closure S = (closure S \ f -` T)" using closure_minimal[of S "(closure S \ f -` T)"] by auto then show ?thesis by auto qed subsection\<^marker>\tag unimportant\ \A function constant on a set\ definition constant_on (infixl "(constant'_on)" 50) where "f constant_on A \ \y. \x\A. f x = y" lemma constant_on_subset: "\f constant_on A; B \ A\ \ f constant_on B" unfolding constant_on_def by blast lemma injective_not_constant: fixes S :: "'a::{perfect_space} set" shows "\open S; inj_on f S; f constant_on S\ \ S = {}" unfolding constant_on_def by (metis equals0I inj_on_contraD islimpt_UNIV islimpt_def) +lemma constant_on_compose: + assumes "f constant_on A" + shows "g \ f constant_on A" + using assms by (auto simp: constant_on_def) + +lemma not_constant_onI: + "f x \ f y \ x \ A \ y \ A \ \f constant_on A" + unfolding constant_on_def by metis + +lemma constant_onE: + assumes "f constant_on S" and "\x. x\S \ f x = g x" + shows "g constant_on S" + using assms unfolding constant_on_def by simp + lemma constant_on_closureI: fixes f :: "_ \ 'b::t1_space" assumes cof: "f constant_on S" and contf: "continuous_on (closure S) f" - shows "f constant_on (closure S)" -using continuous_constant_on_closure [OF contf] cof unfolding constant_on_def -by metis + shows "f constant_on (closure S)" + using continuous_constant_on_closure [OF contf] cof unfolding constant_on_def + by metis subsection\<^marker>\tag unimportant\ \Continuity relative to a union.\ lemma continuous_on_Un_local: "\closedin (top_of_set (S \ T)) S; closedin (top_of_set (S \ T)) T; continuous_on S f; continuous_on T f\ \ continuous_on (S \ T) f" unfolding continuous_on closedin_limpt by (metis Lim_trivial_limit Lim_within_union Un_iff trivial_limit_within) lemma continuous_on_cases_local: "\closedin (top_of_set (S \ T)) S; closedin (top_of_set (S \ T)) T; continuous_on S f; continuous_on T g; \x. \x \ S \ \P x \ x \ T \ P x\ \ f x = g x\ \ continuous_on (S \ T) (\x. if P x then f x else g x)" by (rule continuous_on_Un_local) (auto intro: continuous_on_eq) lemma continuous_on_cases_le: fixes h :: "'a :: topological_space \ real" assumes "continuous_on {x \ S. h x \ a} f" and "continuous_on {x \ S. a \ h x} g" and h: "continuous_on S h" and "\x. \x \ S; h x = a\ \ f x = g x" shows "continuous_on S (\x. if h x \ a then f(x) else g(x))" proof - have S: "S = (S \ h -` atMost a) \ (S \ h -` atLeast a)" by force have 1: "closedin (top_of_set S) (S \ h -` atMost a)" by (rule continuous_closedin_preimage [OF h closed_atMost]) have 2: "closedin (top_of_set S) (S \ h -` atLeast a)" by (rule continuous_closedin_preimage [OF h closed_atLeast]) have [simp]: "S \ h -` {..a} = {x \ S. h x \ a}" "S \ h -` {a..} = {x \ S. a \ h x}" by auto have "continuous_on (S \ h -` {..a} \ S \ h -` {a..}) (\x. if h x \ a then f x else g x)" by (intro continuous_on_cases_local) (use 1 2 S assms in auto) then show ?thesis using S by force qed lemma continuous_on_cases_1: fixes S :: "real set" assumes "continuous_on {t \ S. t \ a} f" and "continuous_on {t \ S. a \ t} g" and "a \ S \ f a = g a" shows "continuous_on S (\t. if t \ a then f(t) else g(t))" using assms by (auto intro: continuous_on_cases_le [where h = id, simplified]) subsection\<^marker>\tag unimportant\\Inverse function property for open/closed maps\ lemma continuous_on_inverse_open_map: assumes contf: "continuous_on S f" and imf: "f ` S = T" and injf: "\x. x \ S \ g (f x) = x" and oo: "\U. openin (top_of_set S) U \ openin (top_of_set T) (f ` U)" shows "continuous_on T g" proof - from imf injf have gTS: "g ` T = S" by force from imf injf have fU: "U \ S \ (f ` U) = T \ g -` U" for U by force show ?thesis by (simp add: continuous_on_open [of T g] gTS) (metis openin_imp_subset fU oo) qed lemma continuous_on_inverse_closed_map: assumes contf: "continuous_on S f" and imf: "f ` S = T" and injf: "\x. x \ S \ g(f x) = x" and oo: "\U. closedin (top_of_set S) U \ closedin (top_of_set T) (f ` U)" shows "continuous_on T g" proof - from imf injf have gTS: "g ` T = S" by force from imf injf have fU: "U \ S \ (f ` U) = T \ g -` U" for U by force show ?thesis by (simp add: continuous_on_closed [of T g] gTS) (metis closedin_imp_subset fU oo) qed lemma homeomorphism_injective_open_map: assumes contf: "continuous_on S f" and imf: "f ` S = T" and injf: "inj_on f S" and oo: "\U. openin (top_of_set S) U \ openin (top_of_set T) (f ` U)" obtains g where "homeomorphism S T f g" proof have "continuous_on T (inv_into S f)" by (metis contf continuous_on_inverse_open_map imf injf inv_into_f_f oo) with imf injf contf show "homeomorphism S T f (inv_into S f)" by (auto simp: homeomorphism_def) qed lemma homeomorphism_injective_closed_map: assumes contf: "continuous_on S f" and imf: "f ` S = T" and injf: "inj_on f S" and oo: "\U. closedin (top_of_set S) U \ closedin (top_of_set T) (f ` U)" obtains g where "homeomorphism S T f g" proof have "continuous_on T (inv_into S f)" by (metis contf continuous_on_inverse_closed_map imf injf inv_into_f_f oo) with imf injf contf show "homeomorphism S T f (inv_into S f)" by (auto simp: homeomorphism_def) qed lemma homeomorphism_imp_open_map: assumes hom: "homeomorphism S T f g" and oo: "openin (top_of_set S) U" shows "openin (top_of_set T) (f ` U)" proof - from hom oo have [simp]: "f ` U = T \ g -` U" using openin_subset by (fastforce simp: homeomorphism_def rev_image_eqI) from hom have "continuous_on T g" unfolding homeomorphism_def by blast moreover have "g ` T = S" by (metis hom homeomorphism_def) ultimately show ?thesis by (simp add: continuous_on_open oo) qed lemma homeomorphism_imp_closed_map: assumes hom: "homeomorphism S T f g" and oo: "closedin (top_of_set S) U" shows "closedin (top_of_set T) (f ` U)" proof - from hom oo have [simp]: "f ` U = T \ g -` U" using closedin_subset by (fastforce simp: homeomorphism_def rev_image_eqI) from hom have "continuous_on T g" unfolding homeomorphism_def by blast moreover have "g ` T = S" by (metis hom homeomorphism_def) ultimately show ?thesis by (simp add: continuous_on_closed oo) qed subsection\<^marker>\tag unimportant\ \Seperability\ lemma subset_second_countable: obtains \ :: "'a:: second_countable_topology set set" where "countable \" "{} \ \" "\C. C \ \ \ openin(top_of_set S) C" "\T. openin(top_of_set S) T \ \\. \ \ \ \ T = \\" proof - obtain \ :: "'a set set" where "countable \" and opeB: "\C. C \ \ \ openin(top_of_set S) C" and \: "\T. openin(top_of_set S) T \ \\. \ \ \ \ T = \\" proof - obtain \ :: "'a set set" where "countable \" and ope: "\C. C \ \ \ open C" and \: "\S. open S \ \U. U \ \ \ S = \U" by (metis univ_second_countable that) show ?thesis proof show "countable ((\C. S \ C) ` \)" by (simp add: \countable \\) show "\C. C \ (\) S ` \ \ openin (top_of_set S) C" using ope by auto show "\T. openin (top_of_set S) T \ \\\(\) S ` \. T = \\" by (metis \ image_mono inf_Sup openin_open) qed qed show ?thesis proof show "countable (\ - {{}})" using \countable \\ by blast show "\C. \C \ \ - {{}}\ \ openin (top_of_set S) C" by (simp add: \\C. C \ \ \ openin (top_of_set S) C\) show "\\\\ - {{}}. T = \\" if "openin (top_of_set S) T" for T using \ [OF that] apply clarify by (rule_tac x="\ - {{}}" in exI, auto) qed auto qed lemma Lindelof_openin: fixes \ :: "'a::second_countable_topology set set" assumes "\S. S \ \ \ openin (top_of_set U) S" obtains \' where "\' \ \" "countable \'" "\\' = \\" proof - have "\S. S \ \ \ \T. open T \ S = U \ T" using assms by (simp add: openin_open) then obtain tf where tf: "\S. S \ \ \ open (tf S) \ (S = U \ tf S)" by metis have [simp]: "\\'. \' \ \ \ \\' = U \ \(tf ` \')" using tf by fastforce obtain \ where "countable \ \ \ \ tf ` \" "\\ = \(tf ` \)" using tf by (force intro: Lindelof [of "tf ` \"]) then obtain \' where \': "\' \ \" "countable \'" "\\' = \\" by (clarsimp simp add: countable_subset_image) then show ?thesis .. qed subsection\<^marker>\tag unimportant\\Closed Maps\ lemma continuous_imp_closed_map: fixes f :: "'a::t2_space \ 'b::t2_space" assumes "closedin (top_of_set S) U" "continuous_on S f" "f ` S = T" "compact S" shows "closedin (top_of_set T) (f ` U)" by (metis assms closedin_compact_eq compact_continuous_image continuous_on_subset subset_image_iff) lemma closed_map_restrict: assumes cloU: "closedin (top_of_set (S \ f -` T')) U" and cc: "\U. closedin (top_of_set S) U \ closedin (top_of_set T) (f ` U)" and "T' \ T" shows "closedin (top_of_set T') (f ` U)" proof - obtain V where "closed V" "U = S \ f -` T' \ V" using cloU by (auto simp: closedin_closed) with cc [of "S \ V"] \T' \ T\ show ?thesis by (fastforce simp add: closedin_closed) qed subsection\<^marker>\tag unimportant\\Open Maps\ lemma open_map_restrict: assumes opeU: "openin (top_of_set (S \ f -` T')) U" and oo: "\U. openin (top_of_set S) U \ openin (top_of_set T) (f ` U)" and "T' \ T" shows "openin (top_of_set T') (f ` U)" proof - obtain V where "open V" "U = S \ f -` T' \ V" using opeU by (auto simp: openin_open) with oo [of "S \ V"] \T' \ T\ show ?thesis by (fastforce simp add: openin_open) qed subsection\<^marker>\tag unimportant\\Quotient maps\ lemma quotient_map_imp_continuous_open: assumes T: "f ` S \ T" and ope: "\U. U \ T \ (openin (top_of_set S) (S \ f -` U) \ openin (top_of_set T) U)" shows "continuous_on S f" proof - have [simp]: "S \ f -` f ` S = S" by auto show ?thesis by (meson T continuous_on_open_gen ope openin_imp_subset) qed lemma quotient_map_imp_continuous_closed: assumes T: "f ` S \ T" and ope: "\U. U \ T \ (closedin (top_of_set S) (S \ f -` U) \ closedin (top_of_set T) U)" shows "continuous_on S f" proof - have [simp]: "S \ f -` f ` S = S" by auto show ?thesis by (meson T closedin_imp_subset continuous_on_closed_gen ope) qed lemma open_map_imp_quotient_map: assumes contf: "continuous_on S f" and T: "T \ f ` S" and ope: "\T. openin (top_of_set S) T \ openin (top_of_set (f ` S)) (f ` T)" shows "openin (top_of_set S) (S \ f -` T) = openin (top_of_set (f ` S)) T" proof - have "T = f ` (S \ f -` T)" using T by blast then show ?thesis using "ope" contf continuous_on_open by metis qed lemma closed_map_imp_quotient_map: assumes contf: "continuous_on S f" and T: "T \ f ` S" and ope: "\T. closedin (top_of_set S) T \ closedin (top_of_set (f ` S)) (f ` T)" shows "openin (top_of_set S) (S \ f -` T) \ openin (top_of_set (f ` S)) T" (is "?lhs = ?rhs") proof assume ?lhs then have *: "closedin (top_of_set S) (S - (S \ f -` T))" using closedin_diff by fastforce have [simp]: "(f ` S - f ` (S - (S \ f -` T))) = T" using T by blast show ?rhs using ope [OF *, unfolded closedin_def] by auto next assume ?rhs with contf show ?lhs by (auto simp: continuous_on_open) qed lemma continuous_right_inverse_imp_quotient_map: assumes contf: "continuous_on S f" and imf: "f ` S \ T" and contg: "continuous_on T g" and img: "g ` T \ S" and fg [simp]: "\y. y \ T \ f(g y) = y" and U: "U \ T" shows "openin (top_of_set S) (S \ f -` U) \ openin (top_of_set T) U" (is "?lhs = ?rhs") proof - have f: "\Z. openin (top_of_set (f ` S)) Z \ openin (top_of_set S) (S \ f -` Z)" and g: "\Z. openin (top_of_set (g ` T)) Z \ openin (top_of_set T) (T \ g -` Z)" using contf contg by (auto simp: continuous_on_open) show ?thesis proof have "T \ g -` (g ` T \ (S \ f -` U)) = {x \ T. f (g x) \ U}" using imf img by blast also have "... = U" using U by auto finally have eq: "T \ g -` (g ` T \ (S \ f -` U)) = U" . assume ?lhs then have *: "openin (top_of_set (g ` T)) (g ` T \ (S \ f -` U))" by (meson img openin_Int openin_subtopology_Int_subset openin_subtopology_self) show ?rhs using g [OF *] eq by auto next assume rhs: ?rhs show ?lhs by (metis f fg image_eqI image_subset_iff imf img openin_subopen openin_subtopology_self openin_trans rhs) qed qed lemma continuous_left_inverse_imp_quotient_map: assumes "continuous_on S f" and "continuous_on (f ` S) g" and "\x. x \ S \ g(f x) = x" and "U \ f ` S" shows "openin (top_of_set S) (S \ f -` U) \ openin (top_of_set (f ` S)) U" using assms by (intro continuous_right_inverse_imp_quotient_map) auto lemma continuous_imp_quotient_map: fixes f :: "'a::t2_space \ 'b::t2_space" assumes "continuous_on S f" "f ` S = T" "compact S" "U \ T" shows "openin (top_of_set S) (S \ f -` U) \ openin (top_of_set T) U" by (simp add: assms closed_map_imp_quotient_map continuous_imp_closed_map) subsection\<^marker>\tag unimportant\\Pasting lemmas for functions, for of casewise definitions\ subsubsection\on open sets\ lemma pasting_lemma: assumes ope: "\i. i \ I \ openin X (T i)" and cont: "\i. i \ I \ continuous_map(subtopology X (T i)) Y (f i)" and f: "\i j x. \i \ I; j \ I; x \ topspace X \ T i \ T j\ \ f i x = f j x" and g: "\x. x \ topspace X \ \j. j \ I \ x \ T j \ g x = f j x" shows "continuous_map X Y g" unfolding continuous_map_openin_preimage_eq proof (intro conjI allI impI) show "g ` topspace X \ topspace Y" using g cont continuous_map_image_subset_topspace by fastforce next fix U assume Y: "openin Y U" have T: "T i \ topspace X" if "i \ I" for i using ope by (simp add: openin_subset that) have *: "topspace X \ g -` U = (\i \ I. T i \ f i -` U)" using f g T by fastforce have "\i. i \ I \ openin X (T i \ f i -` U)" using cont unfolding continuous_map_openin_preimage_eq by (metis Y T inf.commute inf_absorb1 ope topspace_subtopology openin_trans_full) then show "openin X (topspace X \ g -` U)" by (auto simp: *) qed lemma pasting_lemma_exists: assumes X: "topspace X \ (\i \ I. T i)" and ope: "\i. i \ I \ openin X (T i)" and cont: "\i. i \ I \ continuous_map (subtopology X (T i)) Y (f i)" and f: "\i j x. \i \ I; j \ I; x \ topspace X \ T i \ T j\ \ f i x = f j x" obtains g where "continuous_map X Y g" "\x i. \i \ I; x \ topspace X \ T i\ \ g x = f i x" proof let ?h = "\x. f (SOME i. i \ I \ x \ T i) x" show "continuous_map X Y ?h" apply (rule pasting_lemma [OF ope cont]) apply (blast intro: f)+ by (metis (no_types, lifting) UN_E X subsetD someI_ex) show "f (SOME i. i \ I \ x \ T i) x = f i x" if "i \ I" "x \ topspace X \ T i" for i x by (metis (no_types, lifting) IntD2 IntI f someI_ex that) qed lemma pasting_lemma_locally_finite: assumes fin: "\x. x \ topspace X \ \V. openin X V \ x \ V \ finite {i \ I. T i \ V \ {}}" and clo: "\i. i \ I \ closedin X (T i)" and cont: "\i. i \ I \ continuous_map(subtopology X (T i)) Y (f i)" and f: "\i j x. \i \ I; j \ I; x \ topspace X \ T i \ T j\ \ f i x = f j x" and g: "\x. x \ topspace X \ \j. j \ I \ x \ T j \ g x = f j x" shows "continuous_map X Y g" unfolding continuous_map_closedin_preimage_eq proof (intro conjI allI impI) show "g ` topspace X \ topspace Y" using g cont continuous_map_image_subset_topspace by fastforce next fix U assume Y: "closedin Y U" have T: "T i \ topspace X" if "i \ I" for i using clo by (simp add: closedin_subset that) have *: "topspace X \ g -` U = (\i \ I. T i \ f i -` U)" using f g T by fastforce have cTf: "\i. i \ I \ closedin X (T i \ f i -` U)" using cont unfolding continuous_map_closedin_preimage_eq topspace_subtopology by (simp add: Int_absorb1 T Y clo closedin_closed_subtopology) have sub: "{Z \ (\i. T i \ f i -` U) ` I. Z \ V \ {}} \ (\i. T i \ f i -` U) ` {i \ I. T i \ V \ {}}" for V by auto have 1: "(\i\I. T i \ f i -` U) \ topspace X" using T by blast then have "locally_finite_in X ((\i. T i \ f i -` U) ` I)" unfolding locally_finite_in_def using finite_subset [OF sub] fin by force then show "closedin X (topspace X \ g -` U)" by (smt (verit, best) * cTf closedin_locally_finite_Union image_iff) qed subsubsection\Likewise on closed sets, with a finiteness assumption\ lemma pasting_lemma_closed: assumes fin: "finite I" and clo: "\i. i \ I \ closedin X (T i)" and cont: "\i. i \ I \ continuous_map(subtopology X (T i)) Y (f i)" and f: "\i j x. \i \ I; j \ I; x \ topspace X \ T i \ T j\ \ f i x = f j x" and g: "\x. x \ topspace X \ \j. j \ I \ x \ T j \ g x = f j x" shows "continuous_map X Y g" using pasting_lemma_locally_finite [OF _ clo cont f g] fin by auto lemma pasting_lemma_exists_locally_finite: assumes fin: "\x. x \ topspace X \ \V. openin X V \ x \ V \ finite {i \ I. T i \ V \ {}}" and X: "topspace X \ \(T ` I)" and clo: "\i. i \ I \ closedin X (T i)" and cont: "\i. i \ I \ continuous_map(subtopology X (T i)) Y (f i)" and f: "\i j x. \i \ I; j \ I; x \ topspace X \ T i \ T j\ \ f i x = f j x" and g: "\x. x \ topspace X \ \j. j \ I \ x \ T j \ g x = f j x" obtains g where "continuous_map X Y g" "\x i. \i \ I; x \ topspace X \ T i\ \ g x = f i x" proof show "continuous_map X Y (\x. f(@i. i \ I \ x \ T i) x)" apply (rule pasting_lemma_locally_finite [OF fin]) apply (blast intro: assms)+ by (metis (no_types, lifting) UN_E X set_rev_mp someI_ex) next fix x i assume "i \ I" and "x \ topspace X \ T i" then show "f (SOME i. i \ I \ x \ T i) x = f i x" by (metis (mono_tags, lifting) IntE IntI f someI2) qed lemma pasting_lemma_exists_closed: assumes fin: "finite I" and X: "topspace X \ \(T ` I)" and clo: "\i. i \ I \ closedin X (T i)" and cont: "\i. i \ I \ continuous_map(subtopology X (T i)) Y (f i)" and f: "\i j x. \i \ I; j \ I; x \ topspace X \ T i \ T j\ \ f i x = f j x" obtains g where "continuous_map X Y g" "\x i. \i \ I; x \ topspace X \ T i\ \ g x = f i x" proof show "continuous_map X Y (\x. f (SOME i. i \ I \ x \ T i) x)" apply (rule pasting_lemma_closed [OF \finite I\ clo cont]) apply (blast intro: f)+ by (metis (mono_tags, lifting) UN_iff X someI_ex subset_iff) next fix x i assume "i \ I" "x \ topspace X \ T i" then show "f (SOME i. i \ I \ x \ T i) x = f i x" by (metis (no_types, lifting) IntD2 IntI f someI_ex) qed lemma continuous_map_cases: assumes f: "continuous_map (subtopology X (X closure_of {x. P x})) Y f" and g: "continuous_map (subtopology X (X closure_of {x. \ P x})) Y g" and fg: "\x. x \ X frontier_of {x. P x} \ f x = g x" shows "continuous_map X Y (\x. if P x then f x else g x)" proof (rule pasting_lemma_closed) let ?f = "\b. if b then f else g" let ?g = "\x. if P x then f x else g x" let ?T = "\b. if b then X closure_of {x. P x} else X closure_of {x. ~P x}" show "finite {True,False}" by auto have eq: "topspace X - Collect P = topspace X \ {x. \ P x}" by blast show "?f i x = ?f j x" if "i \ {True,False}" "j \ {True,False}" and x: "x \ topspace X \ ?T i \ ?T j" for i j x proof - have "f x = g x" if "i" "\ j" by (smt (verit, best) Diff_Diff_Int closure_of_interior_of closure_of_restrict eq fg frontier_of_closures interior_of_complement that x) moreover have "g x = f x" if "x \ X closure_of {x. \ P x}" "x \ X closure_of Collect P" "\ i" "j" for x by (metis IntI closure_of_restrict eq fg frontier_of_closures that) ultimately show ?thesis using that by (auto simp flip: closure_of_restrict) qed show "\j. j \ {True,False} \ x \ ?T j \ (if P x then f x else g x) = ?f j x" if "x \ topspace X" for x by simp (metis in_closure_of mem_Collect_eq that) qed (auto simp: f g) lemma continuous_map_cases_alt: assumes f: "continuous_map (subtopology X (X closure_of {x \ topspace X. P x})) Y f" and g: "continuous_map (subtopology X (X closure_of {x \ topspace X. ~P x})) Y g" and fg: "\x. x \ X frontier_of {x \ topspace X. P x} \ f x = g x" shows "continuous_map X Y (\x. if P x then f x else g x)" apply (rule continuous_map_cases) using assms apply (simp_all add: Collect_conj_eq closure_of_restrict [symmetric] frontier_of_restrict [symmetric]) done lemma continuous_map_cases_function: assumes contp: "continuous_map X Z p" and contf: "continuous_map (subtopology X {x \ topspace X. p x \ Z closure_of U}) Y f" and contg: "continuous_map (subtopology X {x \ topspace X. p x \ Z closure_of (topspace Z - U)}) Y g" and fg: "\x. \x \ topspace X; p x \ Z frontier_of U\ \ f x = g x" shows "continuous_map X Y (\x. if p x \ U then f x else g x)" proof (rule continuous_map_cases_alt) show "continuous_map (subtopology X (X closure_of {x \ topspace X. p x \ U})) Y f" proof (rule continuous_map_from_subtopology_mono) let ?T = "{x \ topspace X. p x \ Z closure_of U}" show "continuous_map (subtopology X ?T) Y f" by (simp add: contf) show "X closure_of {x \ topspace X. p x \ U} \ ?T" by (rule continuous_map_closure_preimage_subset [OF contp]) qed show "continuous_map (subtopology X (X closure_of {x \ topspace X. p x \ U})) Y g" proof (rule continuous_map_from_subtopology_mono) let ?T = "{x \ topspace X. p x \ Z closure_of (topspace Z - U)}" show "continuous_map (subtopology X ?T) Y g" by (simp add: contg) have "X closure_of {x \ topspace X. p x \ U} \ X closure_of {x \ topspace X. p x \ topspace Z - U}" by (smt (verit, del_insts) DiffI mem_Collect_eq subset_iff closure_of_mono continuous_map_closedin contp) then show "X closure_of {x \ topspace X. p x \ U} \ ?T" by (rule order_trans [OF _ continuous_map_closure_preimage_subset [OF contp]]) qed next show "f x = g x" if "x \ X frontier_of {x \ topspace X. p x \ U}" for x using that continuous_map_frontier_frontier_preimage_subset [OF contp, of U] fg by blast qed subsection \Retractions\ definition\<^marker>\tag important\ retraction :: "('a::topological_space) set \ 'a set \ ('a \ 'a) \ bool" where "retraction S T r \ T \ S \ continuous_on S r \ r ` S \ T \ (\x\T. r x = x)" definition\<^marker>\tag important\ retract_of (infixl "retract'_of" 50) where "T retract_of S \ (\r. retraction S T r)" lemma retraction_idempotent: "retraction S T r \ x \ S \ r (r x) = r x" unfolding retraction_def by auto text \Preservation of fixpoints under (more general notion of) retraction\ lemma invertible_fixpoint_property: fixes S :: "'a::topological_space set" and T :: "'b::topological_space set" assumes contt: "continuous_on T i" and "i ` T \ S" and contr: "continuous_on S r" and "r ` S \ T" and ri: "\y. y \ T \ r (i y) = y" and FP: "\f. \continuous_on S f; f ` S \ S\ \ \x\S. f x = x" and contg: "continuous_on T g" and "g ` T \ T" obtains y where "y \ T" and "g y = y" proof - have "\x\S. (i \ g \ r) x = x" proof (rule FP) show "continuous_on S (i \ g \ r)" by (meson contt contr assms(4) contg assms(8) continuous_on_compose continuous_on_subset) show "(i \ g \ r) ` S \ S" using assms(2,4,8) by force qed then obtain x where x: "x \ S" "(i \ g \ r) x = x" .. then have *: "g (r x) \ T" using assms(4,8) by auto have "r ((i \ g \ r) x) = r x" using x by auto then show ?thesis using "*" ri that by auto qed lemma homeomorphic_fixpoint_property: fixes S :: "'a::topological_space set" and T :: "'b::topological_space set" assumes "S homeomorphic T" shows "(\f. continuous_on S f \ f ` S \ S \ (\x\S. f x = x)) \ (\g. continuous_on T g \ g ` T \ T \ (\y\T. g y = y))" (is "?lhs = ?rhs") proof - obtain r i where r: "\x\S. i (r x) = x" "r ` S = T" "continuous_on S r" "\y\T. r (i y) = y" "i ` T = S" "continuous_on T i" using assms unfolding homeomorphic_def homeomorphism_def by blast show ?thesis proof assume ?lhs with r show ?rhs by (metis invertible_fixpoint_property[of T i S r] order_refl) next assume ?rhs with r show ?lhs by (metis invertible_fixpoint_property[of S r T i] order_refl) qed qed lemma retract_fixpoint_property: fixes f :: "'a::topological_space \ 'b::topological_space" and S :: "'a set" assumes "T retract_of S" and FP: "\f. \continuous_on S f; f ` S \ S\ \ \x\S. f x = x" and contg: "continuous_on T g" and "g ` T \ T" obtains y where "y \ T" and "g y = y" proof - obtain h where "retraction S T h" using assms(1) unfolding retract_of_def .. then show ?thesis unfolding retraction_def using invertible_fixpoint_property[OF continuous_on_id _ _ _ _ FP] by (metis assms(4) contg image_ident that) qed lemma retraction: "retraction S T r \ T \ S \ continuous_on S r \ r ` S = T \ (\x \ T. r x = x)" by (force simp: retraction_def) lemma retractionE: \ \yields properties normalized wrt. simp -- less likely to loop\ assumes "retraction S T r" obtains "T = r ` S" "r ` S \ S" "continuous_on S r" "\x. x \ S \ r (r x) = r x" proof (rule that) from retraction [of S T r] assms have "T \ S" "continuous_on S r" "r ` S = T" and "\x \ T. r x = x" by simp_all then show "T = r ` S" "r ` S \ S" "continuous_on S r" by simp_all from \\x \ T. r x = x\ have "r x = x" if "x \ T" for x using that by simp with \r ` S = T\ show "r (r x) = r x" if "x \ S" for x using that by auto qed lemma retract_ofE: \ \yields properties normalized wrt. simp -- less likely to loop\ assumes "T retract_of S" obtains r where "T = r ` S" "r ` S \ S" "continuous_on S r" "\x. x \ S \ r (r x) = r x" proof - from assms obtain r where "retraction S T r" by (auto simp add: retract_of_def) with that show thesis by (auto elim: retractionE) qed lemma retract_of_imp_extensible: assumes "S retract_of T" and "continuous_on S f" and "f ` S \ U" obtains g where "continuous_on T g" "g ` T \ U" "\x. x \ S \ g x = f x" proof - from \S retract_of T\ obtain r where "retraction T S r" by (auto simp add: retract_of_def) show thesis by (rule that [of "f \ r"]) (use \continuous_on S f\ \f ` S \ U\ \retraction T S r\ in \auto simp: continuous_on_compose2 retraction\) qed lemma idempotent_imp_retraction: assumes "continuous_on S f" and "f ` S \ S" and "\x. x \ S \ f(f x) = f x" shows "retraction S (f ` S) f" by (simp add: assms retraction) lemma retraction_subset: assumes "retraction S T r" and "T \ s'" and "s' \ S" shows "retraction s' T r" unfolding retraction_def by (metis assms continuous_on_subset image_mono retraction) lemma retract_of_subset: assumes "T retract_of S" and "T \ s'" and "s' \ S" shows "T retract_of s'" by (meson assms retract_of_def retraction_subset) lemma retraction_refl [simp]: "retraction S S (\x. x)" by (simp add: retraction) lemma retract_of_refl [iff]: "S retract_of S" unfolding retract_of_def retraction_def using continuous_on_id by blast lemma retract_of_imp_subset: "S retract_of T \ S \ T" by (simp add: retract_of_def retraction_def) lemma retract_of_empty [simp]: "({} retract_of S) \ S = {}" "(S retract_of {}) \ S = {}" by (auto simp: retract_of_def retraction_def) lemma retract_of_singleton [iff]: "({x} retract_of S) \ x \ S" unfolding retract_of_def retraction_def by force lemma retraction_comp: "\retraction S T f; retraction T U g\ \ retraction S U (g \ f)" by (smt (verit, best) comp_apply continuous_on_compose image_comp retraction subset_iff) lemma retract_of_trans [trans]: assumes "S retract_of T" and "T retract_of U" shows "S retract_of U" using assms by (auto simp: retract_of_def intro: retraction_comp) lemma closedin_retract: fixes S :: "'a :: t2_space set" assumes "S retract_of T" shows "closedin (top_of_set T) S" proof - obtain r where r: "S \ T" "continuous_on T r" "r ` T \ S" "\x. x \ S \ r x = x" using assms by (auto simp: retract_of_def retraction_def) have "S = {x\T. x = r x}" using r by auto also have "\ = T \ ((\x. (x, r x)) -` ({y. \x. y = (x, x)}))" unfolding vimage_def mem_Times_iff fst_conv snd_conv using r by auto also have "closedin (top_of_set T) \" by (rule continuous_closedin_preimage) (auto intro!: closed_diagonal continuous_on_Pair r) finally show ?thesis . qed lemma closedin_self [simp]: "closedin (top_of_set S) S" by simp lemma retract_of_closed: fixes S :: "'a :: t2_space set" shows "\closed T; S retract_of T\ \ closed S" by (metis closedin_retract closedin_closed_eq) lemma retract_of_compact: "\compact T; S retract_of T\ \ compact S" by (metis compact_continuous_image retract_of_def retraction) lemma retract_of_connected: "\connected T; S retract_of T\ \ connected S" by (metis Topological_Spaces.connected_continuous_image retract_of_def retraction) lemma retraction_openin_vimage_iff: assumes r: "retraction S T r" and "U \ T" shows "openin (top_of_set S) (S \ r -` U) \ openin (top_of_set T) U" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" using r retraction_def retractionE by (smt (verit, best) continuous_right_inverse_imp_quotient_map retraction_subset \U \ T\) show "?rhs \ ?lhs" by (meson continuous_openin_preimage r retraction_def) qed lemma retract_of_Times: "\S retract_of s'; T retract_of t'\ \ (S \ T) retract_of (s' \ t')" apply (simp add: retract_of_def retraction_def Sigma_mono, clarify) apply (rename_tac f g) apply (rule_tac x="\z. ((f \ fst) z, (g \ snd) z)" in exI) apply (rule conjI continuous_intros | erule continuous_on_subset | force)+ done subsection\Retractions on a topological space\ definition retract_of_space :: "'a set \ 'a topology \ bool" (infix "retract'_of'_space" 50) where "S retract_of_space X \ S \ topspace X \ (\r. continuous_map X (subtopology X S) r \ (\x \ S. r x = x))" lemma retract_of_space_retraction_maps: "S retract_of_space X \ S \ topspace X \ (\r. retraction_maps X (subtopology X S) r id)" by (auto simp: retract_of_space_def retraction_maps_def) lemma retract_of_space_section_map: "S retract_of_space X \ S \ topspace X \ section_map (subtopology X S) X id" unfolding retract_of_space_def retraction_maps_def section_map_def by (auto simp: continuous_map_from_subtopology) lemma retract_of_space_imp_subset: "S retract_of_space X \ S \ topspace X" by (simp add: retract_of_space_def) lemma retract_of_space_topspace: "topspace X retract_of_space X" using retract_of_space_def by force lemma retract_of_space_empty [simp]: "{} retract_of_space X \ topspace X = {}" by (auto simp: continuous_map_def retract_of_space_def) lemma retract_of_space_singleton [simp]: "{a} retract_of_space X \ a \ topspace X" proof - have "continuous_map X (subtopology X {a}) (\x. a) \ (\x. a) a = a" if "a \ topspace X" using that by simp then show ?thesis by (force simp: retract_of_space_def) qed lemma retract_of_space_clopen: assumes "openin X S" "closedin X S" "S = {} \ topspace X = {}" shows "S retract_of_space X" proof (cases "S = {}") case False then obtain a where "a \ S" by blast show ?thesis unfolding retract_of_space_def proof (intro exI conjI) show "S \ topspace X" by (simp add: assms closedin_subset) have "continuous_map X X (\x. if x \ S then x else a)" proof (rule continuous_map_cases) show "continuous_map (subtopology X (X closure_of {x. x \ S})) X (\x. x)" by (simp add: continuous_map_from_subtopology) show "continuous_map (subtopology X (X closure_of {x. x \ S})) X (\x. a)" using \S \ topspace X\ \a \ S\ by force show "x = a" if "x \ X frontier_of {x. x \ S}" for x using assms that clopenin_eq_frontier_of by fastforce qed then show "continuous_map X (subtopology X S) (\x. if x \ S then x else a)" using \S \ topspace X\ \a \ S\ by (auto simp: continuous_map_in_subtopology) qed auto qed (use assms in auto) lemma retract_of_space_disjoint_union: assumes "openin X S" "openin X T" and ST: "disjnt S T" "S \ T = topspace X" and "S = {} \ topspace X = {}" shows "S retract_of_space X" proof (rule retract_of_space_clopen) have "S \ T = {}" by (meson ST disjnt_def) then have "S = topspace X - T" using ST by auto then show "closedin X S" using \openin X T\ by blast qed (auto simp: assms) lemma retraction_maps_section_image1: assumes "retraction_maps X Y r s" shows "s ` (topspace Y) retract_of_space X" unfolding retract_of_space_section_map proof show "s ` topspace Y \ topspace X" using assms continuous_map_image_subset_topspace retraction_maps_def by blast show "section_map (subtopology X (s ` topspace Y)) X id" unfolding section_map_def using assms retraction_maps_to_retract_maps by blast qed lemma retraction_maps_section_image2: "retraction_maps X Y r s \ subtopology X (s ` (topspace Y)) homeomorphic_space Y" using embedding_map_imp_homeomorphic_space homeomorphic_space_sym section_imp_embedding_map section_map_def by blast subsection\Paths and path-connectedness\ definition pathin :: "'a topology \ (real \ 'a) \ bool" where "pathin X g \ continuous_map (subtopology euclideanreal {0..1}) X g" lemma pathin_compose: "\pathin X g; continuous_map X Y f\ \ pathin Y (f \ g)" by (simp add: continuous_map_compose pathin_def) lemma pathin_subtopology: "pathin (subtopology X S) g \ pathin X g \ (\x \ {0..1}. g x \ S)" by (auto simp: pathin_def continuous_map_in_subtopology) lemma pathin_const: "pathin X (\x. a) \ a \ topspace X" by (simp add: pathin_def) lemma path_start_in_topspace: "pathin X g \ g 0 \ topspace X" by (force simp: pathin_def continuous_map) lemma path_finish_in_topspace: "pathin X g \ g 1 \ topspace X" by (force simp: pathin_def continuous_map) lemma path_image_subset_topspace: "pathin X g \ g ` ({0..1}) \ topspace X" by (force simp: pathin_def continuous_map) definition path_connected_space :: "'a topology \ bool" where "path_connected_space X \ \x \ topspace X. \ y \ topspace X. \g. pathin X g \ g 0 = x \ g 1 = y" definition path_connectedin :: "'a topology \ 'a set \ bool" where "path_connectedin X S \ S \ topspace X \ path_connected_space(subtopology X S)" lemma path_connectedin_absolute [simp]: "path_connectedin (subtopology X S) S \ path_connectedin X S" by (simp add: path_connectedin_def subtopology_subtopology) lemma path_connectedin_subset_topspace: "path_connectedin X S \ S \ topspace X" by (simp add: path_connectedin_def) lemma path_connectedin_subtopology: "path_connectedin (subtopology X S) T \ path_connectedin X T \ T \ S" by (auto simp: path_connectedin_def subtopology_subtopology inf.absorb2) lemma path_connectedin: "path_connectedin X S \ S \ topspace X \ (\x \ S. \y \ S. \g. pathin X g \ g ` {0..1} \ S \ g 0 = x \ g 1 = y)" unfolding path_connectedin_def path_connected_space_def pathin_def continuous_map_in_subtopology by (intro conj_cong refl ball_cong) (simp_all add: inf.absorb_iff2) lemma path_connectedin_topspace: "path_connectedin X (topspace X) \ path_connected_space X" by (simp add: path_connectedin_def) lemma path_connected_imp_connected_space: assumes "path_connected_space X" shows "connected_space X" proof - have *: "\S. connectedin X S \ g 0 \ S \ g 1 \ S" if "pathin X g" for g proof (intro exI conjI) have "continuous_map (subtopology euclideanreal {0..1}) X g" using connectedin_absolute that by (simp add: pathin_def) then show "connectedin X (g ` {0..1})" by (rule connectedin_continuous_map_image) auto qed auto show ?thesis using assms by (auto intro: * simp add: path_connected_space_def connected_space_subconnected Ball_def) qed lemma path_connectedin_imp_connectedin: "path_connectedin X S \ connectedin X S" by (simp add: connectedin_def path_connected_imp_connected_space path_connectedin_def) lemma path_connected_space_topspace_empty: "topspace X = {} \ path_connected_space X" by (simp add: path_connected_space_def) lemma path_connectedin_empty [simp]: "path_connectedin X {}" by (simp add: path_connectedin) lemma path_connectedin_singleton [simp]: "path_connectedin X {a} \ a \ topspace X" proof show "path_connectedin X {a} \ a \ topspace X" by (simp add: path_connectedin) show "a \ topspace X \ path_connectedin X {a}" unfolding path_connectedin using pathin_const by fastforce qed lemma path_connectedin_continuous_map_image: assumes f: "continuous_map X Y f" and S: "path_connectedin X S" shows "path_connectedin Y (f ` S)" proof - have fX: "f ` (topspace X) \ topspace Y" by (metis f continuous_map_image_subset_topspace) show ?thesis unfolding path_connectedin proof (intro conjI ballI; clarify?) fix x assume "x \ S" show "f x \ topspace Y" by (meson S fX \x \ S\ image_subset_iff path_connectedin_subset_topspace set_mp) next fix x y assume "x \ S" and "y \ S" then obtain g where g: "pathin X g" "g ` {0..1} \ S" "g 0 = x" "g 1 = y" using S by (force simp: path_connectedin) show "\g. pathin Y g \ g ` {0..1} \ f ` S \ g 0 = f x \ g 1 = f y" proof (intro exI conjI) show "pathin Y (f \ g)" using \pathin X g\ f pathin_compose by auto qed (use g in auto) qed qed lemma path_connectedin_discrete_topology: "path_connectedin (discrete_topology U) S \ S \ U \ (\a. S \ {a})" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" by (meson connectedin_discrete_topology path_connectedin_imp_connectedin) show "?rhs \ ?lhs" using subset_singletonD by fastforce qed lemma path_connected_space_discrete_topology: "path_connected_space (discrete_topology U) \ (\a. U \ {a})" by (metis path_connectedin_discrete_topology path_connectedin_topspace path_connected_space_topspace_empty subset_singletonD topspace_discrete_topology) lemma homeomorphic_path_connected_space_imp: "\path_connected_space X; X homeomorphic_space Y\ \ path_connected_space Y" unfolding homeomorphic_space_def homeomorphic_maps_def by (metis (no_types, opaque_lifting) continuous_map_closedin continuous_map_image_subset_topspace imageI order_class.order.antisym path_connectedin_continuous_map_image path_connectedin_topspace subsetI) lemma homeomorphic_path_connected_space: "X homeomorphic_space Y \ path_connected_space X \ path_connected_space Y" by (meson homeomorphic_path_connected_space_imp homeomorphic_space_sym) lemma homeomorphic_map_path_connectedness: assumes "homeomorphic_map X Y f" "U \ topspace X" shows "path_connectedin Y (f ` U) \ path_connectedin X U" unfolding path_connectedin_def proof (intro conj_cong homeomorphic_path_connected_space) show "(f ` U \ topspace Y) = (U \ topspace X)" using assms homeomorphic_imp_surjective_map by blast next assume "U \ topspace X" show "subtopology Y (f ` U) homeomorphic_space subtopology X U" using assms unfolding homeomorphic_eq_everything_map by (metis (no_types, opaque_lifting) assms homeomorphic_map_subtopologies homeomorphic_space homeomorphic_space_sym image_mono inf.absorb_iff2) qed lemma homeomorphic_map_path_connectedness_eq: "homeomorphic_map X Y f \ path_connectedin X U \ U \ topspace X \ path_connectedin Y (f ` U)" by (meson homeomorphic_map_path_connectedness path_connectedin_def) subsection\Connected components\ definition connected_component_of :: "'a topology \ 'a \ 'a \ bool" where "connected_component_of X x y \ \T. connectedin X T \ x \ T \ y \ T" abbreviation connected_component_of_set where "connected_component_of_set X x \ Collect (connected_component_of X x)" definition connected_components_of :: "'a topology \ ('a set) set" where "connected_components_of X \ connected_component_of_set X ` topspace X" lemma connected_component_in_topspace: "connected_component_of X x y \ x \ topspace X \ y \ topspace X" by (meson connected_component_of_def connectedin_subset_topspace in_mono) lemma connected_component_of_refl: "connected_component_of X x x \ x \ topspace X" by (meson connected_component_in_topspace connected_component_of_def connectedin_sing insertI1) lemma connected_component_of_sym: "connected_component_of X x y \ connected_component_of X y x" by (meson connected_component_of_def) lemma connected_component_of_trans: "\connected_component_of X x y; connected_component_of X y z\ \ connected_component_of X x z" unfolding connected_component_of_def using connectedin_Un by blast lemma connected_component_of_mono: "\connected_component_of (subtopology X S) x y; S \ T\ \ connected_component_of (subtopology X T) x y" by (metis connected_component_of_def connectedin_subtopology inf.absorb_iff2 subtopology_subtopology) lemma connected_component_of_set: "connected_component_of_set X x = {y. \T. connectedin X T \ x \ T \ y \ T}" by (meson connected_component_of_def) lemma connected_component_of_subset_topspace: "connected_component_of_set X x \ topspace X" using connected_component_in_topspace by force lemma connected_component_of_eq_empty: "connected_component_of_set X x = {} \ (x \ topspace X)" using connected_component_in_topspace connected_component_of_refl by fastforce lemma connected_space_iff_connected_component: "connected_space X \ (\x \ topspace X. \y \ topspace X. connected_component_of X x y)" by (simp add: connected_component_of_def connected_space_subconnected) lemma connected_space_imp_connected_component_of: "\connected_space X; a \ topspace X; b \ topspace X\ \ connected_component_of X a b" by (simp add: connected_space_iff_connected_component) lemma connected_space_connected_component_set: "connected_space X \ (\x \ topspace X. connected_component_of_set X x = topspace X)" using connected_component_of_subset_topspace connected_space_iff_connected_component by fastforce lemma connected_component_of_maximal: "\connectedin X S; x \ S\ \ S \ connected_component_of_set X x" by (meson Ball_Collect connected_component_of_def) lemma connected_component_of_equiv: "connected_component_of X x y \ x \ topspace X \ y \ topspace X \ connected_component_of X x = connected_component_of X y" apply (simp add: connected_component_in_topspace fun_eq_iff) by (meson connected_component_of_refl connected_component_of_sym connected_component_of_trans) lemma connected_component_of_disjoint: "disjnt (connected_component_of_set X x) (connected_component_of_set X y) \ ~(connected_component_of X x y)" using connected_component_of_equiv unfolding disjnt_iff by force lemma connected_component_of_eq: "connected_component_of X x = connected_component_of X y \ (x \ topspace X) \ (y \ topspace X) \ x \ topspace X \ y \ topspace X \ connected_component_of X x y" by (metis Collect_empty_eq_bot connected_component_of_eq_empty connected_component_of_equiv) lemma connectedin_connected_component_of: "connectedin X (connected_component_of_set X x)" proof - have "connected_component_of_set X x = \ {T. connectedin X T \ x \ T}" by (auto simp: connected_component_of_def) then show ?thesis by (metis (no_types, lifting) InterI connectedin_Union emptyE mem_Collect_eq) qed lemma Union_connected_components_of: "\(connected_components_of X) = topspace X" unfolding connected_components_of_def using connected_component_in_topspace connected_component_of_refl by fastforce lemma connected_components_of_maximal: "\C \ connected_components_of X; connectedin X S; ~disjnt C S\ \ S \ C" unfolding connected_components_of_def disjnt_def apply clarify by (metis Int_emptyI connected_component_of_def connected_component_of_trans mem_Collect_eq) lemma pairwise_disjoint_connected_components_of: "pairwise disjnt (connected_components_of X)" unfolding connected_components_of_def pairwise_def by (smt (verit, best) connected_component_of_disjoint connected_component_of_eq imageE) lemma complement_connected_components_of_Union: "C \ connected_components_of X \ topspace X - C = \ (connected_components_of X - {C})" by (metis Union_connected_components_of bot.extremum ccpo_Sup_singleton diff_Union_pairwise_disjoint insert_subset pairwise_disjoint_connected_components_of) lemma nonempty_connected_components_of: "C \ connected_components_of X \ C \ {}" unfolding connected_components_of_def by (metis (no_types, lifting) connected_component_of_eq_empty imageE) lemma connected_components_of_subset: "C \ connected_components_of X \ C \ topspace X" using Union_connected_components_of by fastforce lemma connectedin_connected_components_of: assumes "C \ connected_components_of X" shows "connectedin X C" proof - have "C \ connected_component_of_set X ` topspace X" using assms connected_components_of_def by blast then show ?thesis using connectedin_connected_component_of by fastforce qed lemma connected_component_in_connected_components_of: "connected_component_of_set X a \ connected_components_of X \ a \ topspace X" by (metis (no_types, lifting) connected_component_of_eq_empty connected_components_of_def image_iff) lemma connected_space_iff_components_eq: "connected_space X \ (\C \ connected_components_of X. \C' \ connected_components_of X. C = C')" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" by (simp add: connected_components_of_def connected_space_connected_component_set) show "?rhs \ ?lhs" by (metis Union_connected_components_of Union_iff connected_space_subconnected connectedin_connected_components_of) qed lemma connected_components_of_eq_empty: "connected_components_of X = {} \ topspace X = {}" by (simp add: connected_components_of_def) lemma connected_components_of_empty_space: "topspace X = {} \ connected_components_of X = {}" by (simp add: connected_components_of_eq_empty) lemma connected_components_of_subset_sing: "connected_components_of X \ {S} \ connected_space X \ (topspace X = {} \ topspace X = S)" proof (cases "topspace X = {}") case True then show ?thesis by (simp add: connected_components_of_empty_space connected_space_topspace_empty) next case False then have "\connected_components_of X \ {S}\ \ topspace X = S" by (metis Sup_empty Union_connected_components_of ccpo_Sup_singleton subset_singleton_iff) with False show ?thesis unfolding connected_components_of_def by (metis connected_space_connected_component_set empty_iff image_subset_iff insert_iff) qed lemma connected_space_iff_components_subset_singleton: "connected_space X \ (\a. connected_components_of X \ {a})" by (simp add: connected_components_of_subset_sing) lemma connected_components_of_eq_singleton: "connected_components_of X = {S} \ connected_space X \ topspace X \ {} \ S = topspace X" by (metis ccpo_Sup_singleton connected_components_of_subset_sing insert_not_empty subset_singleton_iff) lemma connected_components_of_connected_space: "connected_space X \ connected_components_of X = (if topspace X = {} then {} else {topspace X})" by (simp add: connected_components_of_eq_empty connected_components_of_eq_singleton) lemma exists_connected_component_of_superset: assumes "connectedin X S" and ne: "topspace X \ {}" shows "\C. C \ connected_components_of X \ S \ C" proof (cases "S = {}") case True then show ?thesis using ne connected_components_of_def by blast next case False then show ?thesis by (meson all_not_in_conv assms(1) connected_component_in_connected_components_of connected_component_of_maximal connectedin_subset_topspace in_mono) qed lemma closedin_connected_components_of: assumes "C \ connected_components_of X" shows "closedin X C" proof - obtain x where "x \ topspace X" and x: "C = connected_component_of_set X x" using assms by (auto simp: connected_components_of_def) have "connected_component_of_set X x \ topspace X" by (simp add: connected_component_of_subset_topspace) moreover have "X closure_of connected_component_of_set X x \ connected_component_of_set X x" proof (rule connected_component_of_maximal) show "connectedin X (X closure_of connected_component_of_set X x)" by (simp add: connectedin_closure_of connectedin_connected_component_of) show "x \ X closure_of connected_component_of_set X x" by (simp add: \x \ topspace X\ closure_of connected_component_of_refl) qed ultimately show ?thesis using closure_of_subset_eq x by auto qed lemma closedin_connected_component_of: "closedin X (connected_component_of_set X x)" by (metis closedin_connected_components_of closedin_empty connected_component_in_connected_components_of connected_component_of_eq_empty) lemma connected_component_of_eq_overlap: "connected_component_of_set X x = connected_component_of_set X y \ (x \ topspace X) \ (y \ topspace X) \ ~(connected_component_of_set X x \ connected_component_of_set X y = {})" using connected_component_of_equiv by fastforce lemma connected_component_of_nonoverlap: "connected_component_of_set X x \ connected_component_of_set X y = {} \ (x \ topspace X) \ (y \ topspace X) \ ~(connected_component_of_set X x = connected_component_of_set X y)" by (metis connected_component_of_eq_empty connected_component_of_eq_overlap inf.idem) lemma connected_component_of_overlap: "~(connected_component_of_set X x \ connected_component_of_set X y = {}) \ x \ topspace X \ y \ topspace X \ connected_component_of_set X x = connected_component_of_set X y" by (meson connected_component_of_nonoverlap) end \ No newline at end of file diff --git a/src/HOL/Complex.thy b/src/HOL/Complex.thy --- a/src/HOL/Complex.thy +++ b/src/HOL/Complex.thy @@ -1,1354 +1,1377 @@ (* Title: HOL/Complex.thy Author: Jacques D. Fleuriot, 2001 University of Edinburgh Author: Lawrence C Paulson, 2003/4 *) section \Complex Numbers: Rectangular and Polar Representations\ theory Complex imports Transcendental Real_Vector_Spaces begin text \ We use the \<^theory_text>\codatatype\ command to define the type of complex numbers. This allows us to use \<^theory_text>\primcorec\ to define complex functions by defining their real and imaginary result separately. \ codatatype complex = Complex (Re: real) (Im: real) lemma complex_surj: "Complex (Re z) (Im z) = z" by (rule complex.collapse) lemma complex_eqI [intro?]: "Re x = Re y \ Im x = Im y \ x = y" by (rule complex.expand) simp lemma complex_eq_iff: "x = y \ Re x = Re y \ Im x = Im y" by (auto intro: complex.expand) subsection \Addition and Subtraction\ instantiation complex :: ab_group_add begin primcorec zero_complex where "Re 0 = 0" | "Im 0 = 0" primcorec plus_complex where "Re (x + y) = Re x + Re y" | "Im (x + y) = Im x + Im y" primcorec uminus_complex where "Re (- x) = - Re x" | "Im (- x) = - Im x" primcorec minus_complex where "Re (x - y) = Re x - Re y" | "Im (x - y) = Im x - Im y" instance by standard (simp_all add: complex_eq_iff) end subsection \Multiplication and Division\ instantiation complex :: field begin primcorec one_complex where "Re 1 = 1" | "Im 1 = 0" primcorec times_complex where "Re (x * y) = Re x * Re y - Im x * Im y" | "Im (x * y) = Re x * Im y + Im x * Re y" primcorec inverse_complex where "Re (inverse x) = Re x / ((Re x)\<^sup>2 + (Im x)\<^sup>2)" | "Im (inverse x) = - Im x / ((Re x)\<^sup>2 + (Im x)\<^sup>2)" definition "x div y = x * inverse y" for x y :: complex instance by standard (simp_all add: complex_eq_iff divide_complex_def distrib_left distrib_right right_diff_distrib left_diff_distrib power2_eq_square add_divide_distrib [symmetric]) end lemma Re_divide: "Re (x / y) = (Re x * Re y + Im x * Im y) / ((Re y)\<^sup>2 + (Im y)\<^sup>2)" by (simp add: divide_complex_def add_divide_distrib) lemma Im_divide: "Im (x / y) = (Im x * Re y - Re x * Im y) / ((Re y)\<^sup>2 + (Im y)\<^sup>2)" by (simp add: divide_complex_def diff_divide_distrib) lemma Complex_divide: "(x / y) = Complex ((Re x * Re y + Im x * Im y) / ((Re y)\<^sup>2 + (Im y)\<^sup>2)) ((Im x * Re y - Re x * Im y) / ((Re y)\<^sup>2 + (Im y)\<^sup>2))" by (metis Im_divide Re_divide complex_surj) lemma Re_power2: "Re (x ^ 2) = (Re x)^2 - (Im x)^2" by (simp add: power2_eq_square) lemma Im_power2: "Im (x ^ 2) = 2 * Re x * Im x" by (simp add: power2_eq_square) lemma Re_power_real [simp]: "Im x = 0 \ Re (x ^ n) = Re x ^ n " by (induct n) simp_all lemma Im_power_real [simp]: "Im x = 0 \ Im (x ^ n) = 0" by (induct n) simp_all subsection \Scalar Multiplication\ instantiation complex :: real_field begin primcorec scaleR_complex where "Re (scaleR r x) = r * Re x" | "Im (scaleR r x) = r * Im x" instance proof fix a b :: real and x y :: complex show "scaleR a (x + y) = scaleR a x + scaleR a y" by (simp add: complex_eq_iff distrib_left) show "scaleR (a + b) x = scaleR a x + scaleR b x" by (simp add: complex_eq_iff distrib_right) show "scaleR a (scaleR b x) = scaleR (a * b) x" by (simp add: complex_eq_iff mult.assoc) show "scaleR 1 x = x" by (simp add: complex_eq_iff) show "scaleR a x * y = scaleR a (x * y)" by (simp add: complex_eq_iff algebra_simps) show "x * scaleR a y = scaleR a (x * y)" by (simp add: complex_eq_iff algebra_simps) qed end subsection \Numerals, Arithmetic, and Embedding from R\ declare [[coercion "of_real :: real \ complex"]] declare [[coercion "of_rat :: rat \ complex"]] declare [[coercion "of_int :: int \ complex"]] declare [[coercion "of_nat :: nat \ complex"]] abbreviation complex_of_nat::"nat \ complex" where "complex_of_nat \ of_nat" abbreviation complex_of_int::"int \ complex" where "complex_of_int \ of_int" abbreviation complex_of_rat::"rat \ complex" where "complex_of_rat \ of_rat" abbreviation complex_of_real :: "real \ complex" where "complex_of_real \ of_real" lemma complex_Re_of_nat [simp]: "Re (of_nat n) = of_nat n" by (induct n) simp_all lemma complex_Im_of_nat [simp]: "Im (of_nat n) = 0" by (induct n) simp_all lemma complex_Re_of_int [simp]: "Re (of_int z) = of_int z" by (cases z rule: int_diff_cases) simp lemma complex_Im_of_int [simp]: "Im (of_int z) = 0" by (cases z rule: int_diff_cases) simp lemma complex_Re_numeral [simp]: "Re (numeral v) = numeral v" using complex_Re_of_int [of "numeral v"] by simp lemma complex_Im_numeral [simp]: "Im (numeral v) = 0" using complex_Im_of_int [of "numeral v"] by simp lemma Re_complex_of_real [simp]: "Re (complex_of_real z) = z" by (simp add: of_real_def) lemma Im_complex_of_real [simp]: "Im (complex_of_real z) = 0" by (simp add: of_real_def) lemma Re_divide_numeral [simp]: "Re (z / numeral w) = Re z / numeral w" by (simp add: Re_divide sqr_conv_mult) lemma Im_divide_numeral [simp]: "Im (z / numeral w) = Im z / numeral w" by (simp add: Im_divide sqr_conv_mult) lemma Re_divide_of_nat [simp]: "Re (z / of_nat n) = Re z / of_nat n" by (cases n) (simp_all add: Re_divide field_split_simps power2_eq_square del: of_nat_Suc) lemma Im_divide_of_nat [simp]: "Im (z / of_nat n) = Im z / of_nat n" by (cases n) (simp_all add: Im_divide field_split_simps power2_eq_square del: of_nat_Suc) lemma Re_inverse [simp]: "r \ \ \ Re (inverse r) = inverse (Re r)" by (metis Re_complex_of_real Reals_cases of_real_inverse) lemma Im_inverse [simp]: "r \ \ \ Im (inverse r) = 0" by (metis Im_complex_of_real Reals_cases of_real_inverse) lemma of_real_Re [simp]: "z \ \ \ of_real (Re z) = z" by (auto simp: Reals_def) lemma complex_Re_fact [simp]: "Re (fact n) = fact n" proof - have "(fact n :: complex) = of_real (fact n)" by simp also have "Re \ = fact n" by (subst Re_complex_of_real) simp_all finally show ?thesis . qed lemma complex_Im_fact [simp]: "Im (fact n) = 0" by (subst of_nat_fact [symmetric]) (simp only: complex_Im_of_nat) lemma Re_prod_Reals: "(\x. x \ A \ f x \ \) \ Re (prod f A) = prod (\x. Re (f x)) A" proof (induction A rule: infinite_finite_induct) case (insert x A) hence "Re (prod f (insert x A)) = Re (f x) * Re (prod f A) - Im (f x) * Im (prod f A)" by simp also from insert.prems have "f x \ \" by simp hence "Im (f x) = 0" by (auto elim!: Reals_cases) also have "Re (prod f A) = (\x\A. Re (f x))" by (intro insert.IH insert.prems) auto finally show ?case using insert.hyps by simp qed auto subsection \The Complex Number $i$\ primcorec imaginary_unit :: complex ("\") where "Re \ = 0" | "Im \ = 1" lemma Complex_eq: "Complex a b = a + \ * b" by (simp add: complex_eq_iff) lemma complex_eq: "a = Re a + \ * Im a" by (simp add: complex_eq_iff) lemma fun_complex_eq: "f = (\x. Re (f x) + \ * Im (f x))" by (simp add: fun_eq_iff complex_eq) lemma i_squared [simp]: "\ * \ = -1" by (simp add: complex_eq_iff) lemma power2_i [simp]: "\\<^sup>2 = -1" by (simp add: power2_eq_square) lemma inverse_i [simp]: "inverse \ = - \" by (rule inverse_unique) simp lemma divide_i [simp]: "x / \ = - \ * x" by (simp add: divide_complex_def) lemma complex_i_mult_minus [simp]: "\ * (\ * x) = - x" by (simp add: mult.assoc [symmetric]) lemma complex_i_not_zero [simp]: "\ \ 0" by (simp add: complex_eq_iff) lemma complex_i_not_one [simp]: "\ \ 1" by (simp add: complex_eq_iff) lemma complex_i_not_numeral [simp]: "\ \ numeral w" by (simp add: complex_eq_iff) lemma complex_i_not_neg_numeral [simp]: "\ \ - numeral w" by (simp add: complex_eq_iff) lemma complex_split_polar: "\r a. z = complex_of_real r * (cos a + \ * sin a)" by (simp add: complex_eq_iff polar_Ex) lemma i_even_power [simp]: "\ ^ (n * 2) = (-1) ^ n" by (metis mult.commute power2_i power_mult) lemma Re_i_times [simp]: "Re (\ * z) = - Im z" by simp lemma Im_i_times [simp]: "Im (\ * z) = Re z" by simp lemma i_times_eq_iff: "\ * w = z \ w = - (\ * z)" by auto lemma divide_numeral_i [simp]: "z / (numeral n * \) = - (\ * z) / numeral n" by (metis divide_divide_eq_left divide_i mult.commute mult_minus_right) lemma imaginary_eq_real_iff [simp]: assumes "y \ Reals" "x \ Reals" shows "\ * y = x \ x=0 \ y=0" by (metis Im_complex_of_real Im_i_times assms mult_zero_right of_real_0 of_real_Re) lemma real_eq_imaginary_iff [simp]: assumes "y \ Reals" "x \ Reals" shows "x = \ * y \ x=0 \ y=0" using assms imaginary_eq_real_iff by fastforce subsection \Vector Norm\ instantiation complex :: real_normed_field begin definition "norm z = sqrt ((Re z)\<^sup>2 + (Im z)\<^sup>2)" abbreviation cmod :: "complex \ real" where "cmod \ norm" definition complex_sgn_def: "sgn x = x /\<^sub>R cmod x" definition dist_complex_def: "dist x y = cmod (x - y)" definition uniformity_complex_def [code del]: "(uniformity :: (complex \ complex) filter) = (INF e\{0 <..}. principal {(x, y). dist x y < e})" definition open_complex_def [code del]: "open (U :: complex set) \ (\x\U. eventually (\(x', y). x' = x \ y \ U) uniformity)" instance proof fix r :: real and x y :: complex and S :: "complex set" show "(norm x = 0) = (x = 0)" by (simp add: norm_complex_def complex_eq_iff) show "norm (x + y) \ norm x + norm y" by (simp add: norm_complex_def complex_eq_iff real_sqrt_sum_squares_triangle_ineq) show "norm (scaleR r x) = \r\ * norm x" by (simp add: norm_complex_def complex_eq_iff power_mult_distrib distrib_left [symmetric] real_sqrt_mult) show "norm (x * y) = norm x * norm y" by (simp add: norm_complex_def complex_eq_iff real_sqrt_mult [symmetric] power2_eq_square algebra_simps) qed (rule complex_sgn_def dist_complex_def open_complex_def uniformity_complex_def)+ end declare uniformity_Abort[where 'a = complex, code] lemma norm_ii [simp]: "norm \ = 1" by (simp add: norm_complex_def) lemma cmod_unit_one: "cmod (cos a + \ * sin a) = 1" by (simp add: norm_complex_def) lemma cmod_complex_polar: "cmod (r * (cos a + \ * sin a)) = \r\" by (simp add: norm_mult cmod_unit_one) lemma complex_Re_le_cmod: "Re x \ cmod x" unfolding norm_complex_def by (rule real_sqrt_sum_squares_ge1) lemma complex_mod_minus_le_complex_mod: "- cmod x \ cmod x" by (rule order_trans [OF _ norm_ge_zero]) simp lemma complex_mod_triangle_ineq2: "cmod (b + a) - cmod b \ cmod a" by (rule ord_le_eq_trans [OF norm_triangle_ineq2]) simp lemma abs_Re_le_cmod: "\Re x\ \ cmod x" by (simp add: norm_complex_def) lemma abs_Im_le_cmod: "\Im x\ \ cmod x" by (simp add: norm_complex_def) lemma cmod_le: "cmod z \ \Re z\ + \Im z\" using norm_complex_def sqrt_sum_squares_le_sum_abs by presburger lemma cmod_eq_Re: "Im z = 0 \ cmod z = \Re z\" by (simp add: norm_complex_def) lemma cmod_eq_Im: "Re z = 0 \ cmod z = \Im z\" by (simp add: norm_complex_def) lemma cmod_power2: "(cmod z)\<^sup>2 = (Re z)\<^sup>2 + (Im z)\<^sup>2" by (simp add: norm_complex_def) lemma cmod_plus_Re_le_0_iff: "cmod z + Re z \ 0 \ Re z = - cmod z" using abs_Re_le_cmod[of z] by auto lemma cmod_Re_le_iff: "Im x = Im y \ cmod x \ cmod y \ \Re x\ \ \Re y\" by (metis add.commute add_le_cancel_left norm_complex_def real_sqrt_abs real_sqrt_le_iff) lemma cmod_Im_le_iff: "Re x = Re y \ cmod x \ cmod y \ \Im x\ \ \Im y\" by (metis add_le_cancel_left norm_complex_def real_sqrt_abs real_sqrt_le_iff) lemma Im_eq_0: "\Re z\ = cmod z \ Im z = 0" by (subst (asm) power_eq_iff_eq_base[symmetric, where n=2]) (auto simp add: norm_complex_def) lemma abs_sqrt_wlog: "(\x. x \ 0 \ P x (x\<^sup>2)) \ P \x\ (x\<^sup>2)" for x::"'a::linordered_idom" by (metis abs_ge_zero power2_abs) lemma complex_abs_le_norm: "\Re z\ + \Im z\ \ sqrt 2 * norm z" unfolding norm_complex_def apply (rule abs_sqrt_wlog [where x="Re z"]) apply (rule abs_sqrt_wlog [where x="Im z"]) apply (rule power2_le_imp_le) apply (simp_all add: power2_sum add.commute sum_squares_bound real_sqrt_mult [symmetric]) done lemma complex_unit_circle: "z \ 0 \ (Re z / cmod z)\<^sup>2 + (Im z / cmod z)\<^sup>2 = 1" by (simp add: norm_complex_def complex_eq_iff power2_eq_square add_divide_distrib [symmetric]) text \Properties of complex signum.\ lemma sgn_eq: "sgn z = z / complex_of_real (cmod z)" by (simp add: sgn_div_norm divide_inverse scaleR_conv_of_real mult.commute) lemma Re_sgn [simp]: "Re(sgn z) = Re(z)/cmod z" by (simp add: complex_sgn_def divide_inverse) lemma Im_sgn [simp]: "Im(sgn z) = Im(z)/cmod z" by (simp add: complex_sgn_def divide_inverse) subsection \Absolute value\ instantiation complex :: field_abs_sgn begin definition abs_complex :: "complex \ complex" where "abs_complex = of_real \ norm" instance proof qed (auto simp add: abs_complex_def complex_sgn_def norm_divide norm_mult scaleR_conv_of_real field_simps) end subsection \Completeness of the Complexes\ lemma bounded_linear_Re: "bounded_linear Re" by (rule bounded_linear_intro [where K=1]) (simp_all add: norm_complex_def) lemma bounded_linear_Im: "bounded_linear Im" by (rule bounded_linear_intro [where K=1]) (simp_all add: norm_complex_def) lemmas Cauchy_Re = bounded_linear.Cauchy [OF bounded_linear_Re] lemmas Cauchy_Im = bounded_linear.Cauchy [OF bounded_linear_Im] lemmas tendsto_Re [tendsto_intros] = bounded_linear.tendsto [OF bounded_linear_Re] lemmas tendsto_Im [tendsto_intros] = bounded_linear.tendsto [OF bounded_linear_Im] lemmas isCont_Re [simp] = bounded_linear.isCont [OF bounded_linear_Re] lemmas isCont_Im [simp] = bounded_linear.isCont [OF bounded_linear_Im] lemmas continuous_Re [simp] = bounded_linear.continuous [OF bounded_linear_Re] lemmas continuous_Im [simp] = bounded_linear.continuous [OF bounded_linear_Im] lemmas continuous_on_Re [continuous_intros] = bounded_linear.continuous_on[OF bounded_linear_Re] lemmas continuous_on_Im [continuous_intros] = bounded_linear.continuous_on[OF bounded_linear_Im] lemmas has_derivative_Re [derivative_intros] = bounded_linear.has_derivative[OF bounded_linear_Re] lemmas has_derivative_Im [derivative_intros] = bounded_linear.has_derivative[OF bounded_linear_Im] lemmas sums_Re = bounded_linear.sums [OF bounded_linear_Re] lemmas sums_Im = bounded_linear.sums [OF bounded_linear_Im] lemmas Re_suminf = bounded_linear.suminf[OF bounded_linear_Re] lemmas Im_suminf = bounded_linear.suminf[OF bounded_linear_Im] lemma tendsto_Complex [tendsto_intros]: "(f \ a) F \ (g \ b) F \ ((\x. Complex (f x) (g x)) \ Complex a b) F" unfolding Complex_eq by (auto intro!: tendsto_intros) lemma tendsto_complex_iff: "(f \ x) F \ (((\x. Re (f x)) \ Re x) F \ ((\x. Im (f x)) \ Im x) F)" proof safe assume "((\x. Re (f x)) \ Re x) F" "((\x. Im (f x)) \ Im x) F" from tendsto_Complex[OF this] show "(f \ x) F" unfolding complex.collapse . qed (auto intro: tendsto_intros) lemma continuous_complex_iff: "continuous F f \ continuous F (\x. Re (f x)) \ continuous F (\x. Im (f x))" by (simp only: continuous_def tendsto_complex_iff) lemma continuous_on_of_real_o_iff [simp]: "continuous_on S (\x. complex_of_real (g x)) = continuous_on S g" using continuous_on_Re continuous_on_of_real by fastforce lemma continuous_on_of_real_id [simp]: "continuous_on S (of_real :: real \ 'a::real_normed_algebra_1)" by (rule continuous_on_of_real [OF continuous_on_id]) lemma has_vector_derivative_complex_iff: "(f has_vector_derivative x) F \ ((\x. Re (f x)) has_field_derivative (Re x)) F \ ((\x. Im (f x)) has_field_derivative (Im x)) F" by (simp add: has_vector_derivative_def has_field_derivative_def has_derivative_def tendsto_complex_iff algebra_simps bounded_linear_scaleR_left bounded_linear_mult_right) lemma has_field_derivative_Re[derivative_intros]: "(f has_vector_derivative D) F \ ((\x. Re (f x)) has_field_derivative (Re D)) F" unfolding has_vector_derivative_complex_iff by safe lemma has_field_derivative_Im[derivative_intros]: "(f has_vector_derivative D) F \ ((\x. Im (f x)) has_field_derivative (Im D)) F" unfolding has_vector_derivative_complex_iff by safe instance complex :: banach proof fix X :: "nat \ complex" assume X: "Cauchy X" then have "(\n. Complex (Re (X n)) (Im (X n))) \ Complex (lim (\n. Re (X n))) (lim (\n. Im (X n)))" by (intro tendsto_Complex convergent_LIMSEQ_iff[THEN iffD1] Cauchy_convergent_iff[THEN iffD1] Cauchy_Re Cauchy_Im) then show "convergent X" unfolding complex.collapse by (rule convergentI) qed declare DERIV_power[where 'a=complex, unfolded of_nat_def[symmetric], derivative_intros] subsection \Complex Conjugation\ primcorec cnj :: "complex \ complex" where "Re (cnj z) = Re z" | "Im (cnj z) = - Im z" lemma complex_cnj_cancel_iff [simp]: "cnj x = cnj y \ x = y" by (simp add: complex_eq_iff) lemma complex_cnj_cnj [simp]: "cnj (cnj z) = z" by (simp add: complex_eq_iff) +lemma in_image_cnj_iff: "z \ cnj ` A \ cnj z \ A" + by (metis complex_cnj_cnj image_iff) + +lemma image_cnj_conv_vimage_cnj: "cnj ` A = cnj -` A" + using in_image_cnj_iff by blast + lemma complex_cnj_zero [simp]: "cnj 0 = 0" by (simp add: complex_eq_iff) lemma complex_cnj_zero_iff [iff]: "cnj z = 0 \ z = 0" by (simp add: complex_eq_iff) lemma complex_cnj_one_iff [simp]: "cnj z = 1 \ z = 1" by (simp add: complex_eq_iff) lemma complex_cnj_add [simp]: "cnj (x + y) = cnj x + cnj y" by (simp add: complex_eq_iff) lemma cnj_sum [simp]: "cnj (sum f s) = (\x\s. cnj (f x))" by (induct s rule: infinite_finite_induct) auto lemma complex_cnj_diff [simp]: "cnj (x - y) = cnj x - cnj y" by (simp add: complex_eq_iff) lemma complex_cnj_minus [simp]: "cnj (- x) = - cnj x" by (simp add: complex_eq_iff) lemma complex_cnj_one [simp]: "cnj 1 = 1" by (simp add: complex_eq_iff) lemma complex_cnj_mult [simp]: "cnj (x * y) = cnj x * cnj y" by (simp add: complex_eq_iff) lemma cnj_prod [simp]: "cnj (prod f s) = (\x\s. cnj (f x))" by (induct s rule: infinite_finite_induct) auto lemma complex_cnj_inverse [simp]: "cnj (inverse x) = inverse (cnj x)" by (simp add: complex_eq_iff) lemma complex_cnj_divide [simp]: "cnj (x / y) = cnj x / cnj y" by (simp add: divide_complex_def) lemma complex_cnj_power [simp]: "cnj (x ^ n) = cnj x ^ n" by (induct n) simp_all lemma complex_cnj_of_nat [simp]: "cnj (of_nat n) = of_nat n" by (simp add: complex_eq_iff) lemma complex_cnj_of_int [simp]: "cnj (of_int z) = of_int z" by (simp add: complex_eq_iff) lemma complex_cnj_numeral [simp]: "cnj (numeral w) = numeral w" by (simp add: complex_eq_iff) lemma complex_cnj_neg_numeral [simp]: "cnj (- numeral w) = - numeral w" by (simp add: complex_eq_iff) lemma complex_cnj_scaleR [simp]: "cnj (scaleR r x) = scaleR r (cnj x)" by (simp add: complex_eq_iff) lemma complex_mod_cnj [simp]: "cmod (cnj z) = cmod z" by (simp add: norm_complex_def) lemma complex_cnj_complex_of_real [simp]: "cnj (of_real x) = of_real x" by (simp add: complex_eq_iff) lemma complex_cnj_i [simp]: "cnj \ = - \" by (simp add: complex_eq_iff) lemma complex_add_cnj: "z + cnj z = complex_of_real (2 * Re z)" by (simp add: complex_eq_iff) lemma complex_diff_cnj: "z - cnj z = complex_of_real (2 * Im z) * \" by (simp add: complex_eq_iff) lemma Ints_cnj [intro]: "x \ \ \ cnj x \ \" by (auto elim!: Ints_cases) lemma cnj_in_Ints_iff [simp]: "cnj x \ \ \ x \ \" using Ints_cnj[of x] Ints_cnj[of "cnj x"] by auto lemma complex_mult_cnj: "z * cnj z = complex_of_real ((Re z)\<^sup>2 + (Im z)\<^sup>2)" by (simp add: complex_eq_iff power2_eq_square) lemma cnj_add_mult_eq_Re: "z * cnj w + cnj z * w = 2 * Re (z * cnj w)" by (rule complex_eqI) auto lemma complex_mod_mult_cnj: "cmod (z * cnj z) = (cmod z)\<^sup>2" by (simp add: norm_mult power2_eq_square) lemma complex_mod_sqrt_Re_mult_cnj: "cmod z = sqrt (Re (z * cnj z))" by (simp add: norm_complex_def power2_eq_square) lemma complex_In_mult_cnj_zero [simp]: "Im (z * cnj z) = 0" by simp lemma complex_cnj_fact [simp]: "cnj (fact n) = fact n" by (subst of_nat_fact [symmetric], subst complex_cnj_of_nat) simp lemma complex_cnj_pochhammer [simp]: "cnj (pochhammer z n) = pochhammer (cnj z) n" by (induct n arbitrary: z) (simp_all add: pochhammer_rec) lemma bounded_linear_cnj: "bounded_linear cnj" using complex_cnj_add complex_cnj_scaleR by (rule bounded_linear_intro [where K=1]) simp lemma linear_cnj: "linear cnj" using bounded_linear.linear[OF bounded_linear_cnj] . lemmas tendsto_cnj [tendsto_intros] = bounded_linear.tendsto [OF bounded_linear_cnj] and isCont_cnj [simp] = bounded_linear.isCont [OF bounded_linear_cnj] and continuous_cnj [simp, continuous_intros] = bounded_linear.continuous [OF bounded_linear_cnj] and continuous_on_cnj [simp, continuous_intros] = bounded_linear.continuous_on [OF bounded_linear_cnj] and has_derivative_cnj [simp, derivative_intros] = bounded_linear.has_derivative [OF bounded_linear_cnj] lemma lim_cnj: "((\x. cnj(f x)) \ cnj l) F \ (f \ l) F" by (simp add: tendsto_iff dist_complex_def complex_cnj_diff [symmetric] del: complex_cnj_diff) lemma sums_cnj: "((\x. cnj(f x)) sums cnj l) \ (f sums l)" by (simp add: sums_def lim_cnj cnj_sum [symmetric] del: cnj_sum) lemma differentiable_cnj_iff: "(\z. cnj (f z)) differentiable at x within A \ f differentiable at x within A" proof assume "(\z. cnj (f z)) differentiable at x within A" then obtain D where "((\z. cnj (f z)) has_derivative D) (at x within A)" by (auto simp: differentiable_def) from has_derivative_cnj[OF this] show "f differentiable at x within A" by (auto simp: differentiable_def) next assume "f differentiable at x within A" then obtain D where "(f has_derivative D) (at x within A)" by (auto simp: differentiable_def) from has_derivative_cnj[OF this] show "(\z. cnj (f z)) differentiable at x within A" by (auto simp: differentiable_def) qed lemma has_vector_derivative_cnj [derivative_intros]: assumes "(f has_vector_derivative f') (at z within A)" shows "((\z. cnj (f z)) has_vector_derivative cnj f') (at z within A)" using assms by (auto simp: has_vector_derivative_complex_iff intro: derivative_intros) subsection \Basic Lemmas\ lemma complex_of_real_code[code_unfold]: "of_real = (\x. Complex x 0)" by (intro ext, auto simp: complex_eq_iff) lemma complex_eq_0: "z=0 \ (Re z)\<^sup>2 + (Im z)\<^sup>2 = 0" by (metis zero_complex.sel complex_eqI sum_power2_eq_zero_iff) lemma complex_neq_0: "z\0 \ (Re z)\<^sup>2 + (Im z)\<^sup>2 > 0" by (metis complex_eq_0 less_numeral_extra(3) sum_power2_gt_zero_iff) lemma complex_norm_square: "of_real ((norm z)\<^sup>2) = z * cnj z" by (cases z) (auto simp: complex_eq_iff norm_complex_def power2_eq_square[symmetric] of_real_power[symmetric] simp del: of_real_power) lemma complex_div_cnj: "a / b = (a * cnj b) / (norm b)\<^sup>2" using complex_norm_square by auto lemma Re_complex_div_eq_0: "Re (a / b) = 0 \ Re (a * cnj b) = 0" by (auto simp add: Re_divide) lemma Im_complex_div_eq_0: "Im (a / b) = 0 \ Im (a * cnj b) = 0" by (auto simp add: Im_divide) lemma complex_div_gt_0: "(Re (a / b) > 0 \ Re (a * cnj b) > 0) \ (Im (a / b) > 0 \ Im (a * cnj b) > 0)" proof (cases "b = 0") case True then show ?thesis by auto next case False then have "0 < (Re b)\<^sup>2 + (Im b)\<^sup>2" by (simp add: complex_eq_iff sum_power2_gt_zero_iff) then show ?thesis by (simp add: Re_divide Im_divide zero_less_divide_iff) qed lemma Re_complex_div_gt_0: "Re (a / b) > 0 \ Re (a * cnj b) > 0" and Im_complex_div_gt_0: "Im (a / b) > 0 \ Im (a * cnj b) > 0" using complex_div_gt_0 by auto lemma Re_complex_div_ge_0: "Re (a / b) \ 0 \ Re (a * cnj b) \ 0" by (metis le_less Re_complex_div_eq_0 Re_complex_div_gt_0) lemma Im_complex_div_ge_0: "Im (a / b) \ 0 \ Im (a * cnj b) \ 0" by (metis Im_complex_div_eq_0 Im_complex_div_gt_0 le_less) lemma Re_complex_div_lt_0: "Re (a / b) < 0 \ Re (a * cnj b) < 0" by (metis less_asym neq_iff Re_complex_div_eq_0 Re_complex_div_gt_0) lemma Im_complex_div_lt_0: "Im (a / b) < 0 \ Im (a * cnj b) < 0" by (metis Im_complex_div_eq_0 Im_complex_div_gt_0 less_asym neq_iff) lemma Re_complex_div_le_0: "Re (a / b) \ 0 \ Re (a * cnj b) \ 0" by (metis not_le Re_complex_div_gt_0) lemma Im_complex_div_le_0: "Im (a / b) \ 0 \ Im (a * cnj b) \ 0" by (metis Im_complex_div_gt_0 not_le) lemma Re_divide_of_real [simp]: "Re (z / of_real r) = Re z / r" by (simp add: Re_divide power2_eq_square) lemma Im_divide_of_real [simp]: "Im (z / of_real r) = Im z / r" by (simp add: Im_divide power2_eq_square) lemma Re_divide_Reals [simp]: "r \ \ \ Re (z / r) = Re z / Re r" by (metis Re_divide_of_real of_real_Re) lemma Im_divide_Reals [simp]: "r \ \ \ Im (z / r) = Im z / Re r" by (metis Im_divide_of_real of_real_Re) lemma Re_sum[simp]: "Re (sum f s) = (\x\s. Re (f x))" by (induct s rule: infinite_finite_induct) auto lemma Im_sum[simp]: "Im (sum f s) = (\x\s. Im(f x))" by (induct s rule: infinite_finite_induct) auto lemma sum_Re_le_cmod: "(\i\I. Re (z i)) \ cmod (\i\I. z i)" by (metis Re_sum complex_Re_le_cmod) lemma sum_Im_le_cmod: "(\i\I. Im (z i)) \ cmod (\i\I. z i)" by (smt (verit, best) Im_sum abs_Im_le_cmod sum.cong) lemma sums_complex_iff: "f sums x \ ((\x. Re (f x)) sums Re x) \ ((\x. Im (f x)) sums Im x)" unfolding sums_def tendsto_complex_iff Im_sum Re_sum .. lemma summable_complex_iff: "summable f \ summable (\x. Re (f x)) \ summable (\x. Im (f x))" unfolding summable_def sums_complex_iff[abs_def] by (metis complex.sel) lemma summable_complex_of_real [simp]: "summable (\n. complex_of_real (f n)) \ summable f" unfolding summable_complex_iff by simp lemma summable_Re: "summable f \ summable (\x. Re (f x))" unfolding summable_complex_iff by blast lemma summable_Im: "summable f \ summable (\x. Im (f x))" unfolding summable_complex_iff by blast lemma complex_is_Nat_iff: "z \ \ \ Im z = 0 \ (\i. Re z = of_nat i)" by (auto simp: Nats_def complex_eq_iff) lemma complex_is_Int_iff: "z \ \ \ Im z = 0 \ (\i. Re z = of_int i)" by (auto simp: Ints_def complex_eq_iff) lemma complex_is_Real_iff: "z \ \ \ Im z = 0" by (auto simp: Reals_def complex_eq_iff) lemma Reals_cnj_iff: "z \ \ \ cnj z = z" by (auto simp: complex_is_Real_iff complex_eq_iff) lemma in_Reals_norm: "z \ \ \ norm z = \Re z\" by (simp add: complex_is_Real_iff norm_complex_def) lemma Re_Reals_divide: "r \ \ \ Re (r / z) = Re r * Re z / (norm z)\<^sup>2" by (simp add: Re_divide complex_is_Real_iff cmod_power2) lemma Im_Reals_divide: "r \ \ \ Im (r / z) = -Re r * Im z / (norm z)\<^sup>2" by (simp add: Im_divide complex_is_Real_iff cmod_power2) lemma series_comparison_complex: fixes f:: "nat \ 'a::banach" assumes sg: "summable g" and "\n. g n \ \" "\n. Re (g n) \ 0" and fg: "\n. n \ N \ norm(f n) \ norm(g n)" shows "summable f" proof - have g: "\n. cmod (g n) = Re (g n)" using assms by (metis abs_of_nonneg in_Reals_norm) show ?thesis by (metis fg g sg summable_comparison_test summable_complex_iff) qed subsection \Polar Form for Complex Numbers\ lemma complex_unimodular_polar: assumes "norm z = 1" obtains t where "0 \ t" "t < 2 * pi" "z = Complex (cos t) (sin t)" by (metis cmod_power2 one_power2 complex_surj sincos_total_2pi [of "Re z" "Im z"] assms) subsubsection \$\cos \theta + i \sin \theta$\ primcorec cis :: "real \ complex" where "Re (cis a) = cos a" | "Im (cis a) = sin a" lemma cis_zero [simp]: "cis 0 = 1" by (simp add: complex_eq_iff) lemma norm_cis [simp]: "norm (cis a) = 1" by (simp add: norm_complex_def) lemma sgn_cis [simp]: "sgn (cis a) = cis a" by (simp add: sgn_div_norm) lemma cis_2pi [simp]: "cis (2 * pi) = 1" by (simp add: cis.ctr complex_eq_iff) lemma cis_neq_zero [simp]: "cis a \ 0" by (metis norm_cis norm_zero zero_neq_one) lemma cis_cnj: "cnj (cis t) = cis (-t)" by (simp add: complex_eq_iff) lemma cis_mult: "cis a * cis b = cis (a + b)" by (simp add: complex_eq_iff cos_add sin_add) lemma DeMoivre: "(cis a) ^ n = cis (real n * a)" by (induct n) (simp_all add: algebra_simps cis_mult) lemma cis_inverse [simp]: "inverse (cis a) = cis (- a)" by (simp add: complex_eq_iff) lemma cis_divide: "cis a / cis b = cis (a - b)" by (simp add: divide_complex_def cis_mult) +lemma divide_conv_cnj: "norm z = 1 \ x / z = x * cnj z" + by (metis complex_div_cnj div_by_1 mult_1 of_real_1 power2_eq_square) + +lemma i_not_in_Reals [simp, intro]: "\ \ \" + by (auto simp: complex_is_Real_iff) + +lemma powr_power_complex: "z \ 0 \ n \ 0 \ (z powr u :: complex) ^ n = z powr (of_nat n * u)" + by (induction n) (auto simp: algebra_simps powr_add) + lemma cos_n_Re_cis_pow_n: "cos (real n * a) = Re (cis a ^ n)" by (auto simp add: DeMoivre) lemma sin_n_Im_cis_pow_n: "sin (real n * a) = Im (cis a ^ n)" by (auto simp add: DeMoivre) lemma cis_pi [simp]: "cis pi = -1" by (simp add: complex_eq_iff) lemma cis_pi_half[simp]: "cis (pi / 2) = \" by (simp add: cis.ctr complex_eq_iff) lemma cis_minus_pi_half[simp]: "cis (-(pi / 2)) = -\" by (simp add: cis.ctr complex_eq_iff) lemma cis_multiple_2pi[simp]: "n \ \ \ cis (2 * pi * n) = 1" by (auto elim!: Ints_cases simp: cis.ctr one_complex.ctr) +lemma minus_cis: "-cis x = cis (x + pi)" + by (simp flip: cis_mult) + +lemma minus_cis': "-cis x = cis (x - pi)" + by (simp flip: cis_divide) subsubsection \$r(\cos \theta + i \sin \theta)$\ definition rcis :: "real \ real \ complex" where "rcis r a = complex_of_real r * cis a" lemma Re_rcis [simp]: "Re(rcis r a) = r * cos a" by (simp add: rcis_def) lemma Im_rcis [simp]: "Im(rcis r a) = r * sin a" by (simp add: rcis_def) lemma rcis_Ex: "\r a. z = rcis r a" by (simp add: complex_eq_iff polar_Ex) lemma complex_mod_rcis [simp]: "cmod (rcis r a) = \r\" by (simp add: rcis_def norm_mult) lemma cis_rcis_eq: "cis a = rcis 1 a" by (simp add: rcis_def) lemma rcis_mult: "rcis r1 a * rcis r2 b = rcis (r1 * r2) (a + b)" by (simp add: rcis_def cis_mult) lemma rcis_zero_mod [simp]: "rcis 0 a = 0" by (simp add: rcis_def) lemma rcis_zero_arg [simp]: "rcis r 0 = complex_of_real r" by (simp add: rcis_def) lemma rcis_eq_zero_iff [simp]: "rcis r a = 0 \ r = 0" by (simp add: rcis_def) lemma DeMoivre2: "(rcis r a) ^ n = rcis (r ^ n) (real n * a)" by (simp add: rcis_def power_mult_distrib DeMoivre) lemma rcis_inverse: "inverse(rcis r a) = rcis (1 / r) (- a)" by (simp add: divide_inverse rcis_def) lemma rcis_divide: "rcis r1 a / rcis r2 b = rcis (r1 / r2) (a - b)" by (simp add: rcis_def cis_divide [symmetric]) subsubsection \Complex exponential\ lemma exp_Reals_eq: assumes "z \ \" shows "exp z = of_real (exp (Re z))" using assms by (auto elim!: Reals_cases simp: exp_of_real) lemma cis_conv_exp: "cis b = exp (\ * b)" proof - have "(\ * complex_of_real b) ^ n /\<^sub>R fact n = of_real (cos_coeff n * b^n) + \ * of_real (sin_coeff n * b^n)" for n :: nat proof - have "\ ^ n = fact n *\<^sub>R (cos_coeff n + \ * sin_coeff n)" by (induct n) (simp_all add: sin_coeff_Suc cos_coeff_Suc complex_eq_iff Re_divide Im_divide field_simps power2_eq_square add_nonneg_eq_0_iff) then show ?thesis by (simp add: field_simps) qed then show ?thesis using sin_converges [of b] cos_converges [of b] by (auto simp add: Complex_eq cis.ctr exp_def simp del: of_real_mult intro!: sums_unique sums_add sums_mult sums_of_real) qed lemma exp_eq_polar: "exp z = exp (Re z) * cis (Im z)" unfolding cis_conv_exp exp_of_real [symmetric] mult_exp_exp by (cases z) (simp add: Complex_eq) lemma Re_exp: "Re (exp z) = exp (Re z) * cos (Im z)" unfolding exp_eq_polar by simp lemma Im_exp: "Im (exp z) = exp (Re z) * sin (Im z)" unfolding exp_eq_polar by simp lemma norm_cos_sin [simp]: "norm (Complex (cos t) (sin t)) = 1" by (simp add: norm_complex_def) lemma norm_exp_eq_Re [simp]: "norm (exp z) = exp (Re z)" by (simp add: cis.code cmod_complex_polar exp_eq_polar Complex_eq) lemma complex_exp_exists: "\a r. z = complex_of_real r * exp a" using cis_conv_exp rcis_Ex rcis_def by force lemma exp_pi_i [simp]: "exp (of_real pi * \) = -1" by (metis cis_conv_exp cis_pi mult.commute) lemma exp_pi_i' [simp]: "exp (\ * of_real pi) = -1" using cis_conv_exp cis_pi by auto lemma exp_two_pi_i [simp]: "exp (2 * of_real pi * \) = 1" by (simp add: exp_eq_polar complex_eq_iff) lemma exp_two_pi_i' [simp]: "exp (\ * (of_real pi * 2)) = 1" by (metis exp_two_pi_i mult.commute) lemma continuous_on_cis [continuous_intros]: "continuous_on A f \ continuous_on A (\x. cis (f x))" by (auto simp: cis_conv_exp intro!: continuous_intros) subsubsection \Complex argument\ definition Arg :: "complex \ real" where "Arg z = (if z = 0 then 0 else (SOME a. sgn z = cis a \ - pi < a \ a \ pi))" lemma Arg_zero: "Arg 0 = 0" by (simp add: Arg_def) lemma cis_Arg_unique: assumes "sgn z = cis x" and "-pi < x" and "x \ pi" shows "Arg z = x" proof - from assms have "z \ 0" by auto have "(SOME a. sgn z = cis a \ -pi < a \ a \ pi) = x" proof fix a define d where "d = a - x" assume a: "sgn z = cis a \ - pi < a \ a \ pi" from a assms have "- (2*pi) < d \ d < 2*pi" unfolding d_def by simp moreover from a assms have "cos a = cos x" and "sin a = sin x" by (simp_all add: complex_eq_iff) then have cos: "cos d = 1" by (simp add: d_def cos_diff) moreover from cos have "sin d = 0" by (rule cos_one_sin_zero) ultimately have "d = 0" by (auto simp: sin_zero_iff elim!: evenE dest!: less_2_cases) then show "a = x" by (simp add: d_def) qed (simp add: assms del: Re_sgn Im_sgn) with \z \ 0\ show "Arg z = x" by (simp add: Arg_def) qed lemma Arg_correct: assumes "z \ 0" shows "sgn z = cis (Arg z) \ -pi < Arg z \ Arg z \ pi" proof (simp add: Arg_def assms, rule someI_ex) obtain r a where z: "z = rcis r a" using rcis_Ex by fast with assms have "r \ 0" by auto define b where "b = (if 0 < r then a else a + pi)" have b: "sgn z = cis b" using \r \ 0\ by (simp add: z b_def rcis_def of_real_def sgn_scaleR sgn_if complex_eq_iff) have cis_2pi_nat: "cis (2 * pi * real_of_nat n) = 1" for n by (induct n) (simp_all add: distrib_left cis_mult [symmetric] complex_eq_iff) have cis_2pi_int: "cis (2 * pi * real_of_int x) = 1" for x by (cases x rule: int_diff_cases) (simp add: right_diff_distrib cis_divide [symmetric] cis_2pi_nat) define c where "c = b - 2 * pi * of_int \(b - pi) / (2 * pi)\" have "sgn z = cis c" by (simp add: b c_def cis_divide [symmetric] cis_2pi_int) moreover have "- pi < c \ c \ pi" using ceiling_correct [of "(b - pi) / (2*pi)"] by (simp add: c_def less_divide_eq divide_le_eq algebra_simps del: le_of_int_ceiling) ultimately show "\a. sgn z = cis a \ -pi < a \ a \ pi" by fast qed lemma Arg_bounded: "- pi < Arg z \ Arg z \ pi" by (cases "z = 0") (simp_all add: Arg_zero Arg_correct) lemma cis_Arg: "z \ 0 \ cis (Arg z) = sgn z" by (simp add: Arg_correct) lemma rcis_cmod_Arg: "rcis (cmod z) (Arg z) = z" by (cases "z = 0") (simp_all add: rcis_def cis_Arg sgn_div_norm of_real_def) lemma rcis_cnj: shows "cnj a = rcis (cmod a) (- Arg a)" by (metis cis_cnj complex_cnj_complex_of_real complex_cnj_mult rcis_cmod_Arg rcis_def) lemma cos_Arg_i_mult_zero [simp]: "y \ 0 \ Re y = 0 \ cos (Arg y) = 0" using cis_Arg [of y] by (simp add: complex_eq_iff) lemma Arg_ii [simp]: "Arg \ = pi/2" by (rule cis_Arg_unique; simp add: sgn_eq) lemma Arg_minus_ii [simp]: "Arg (-\) = -pi/2" proof (rule cis_Arg_unique) show "sgn (- \) = cis (- pi / 2)" by (simp add: sgn_eq) show "- pi / 2 \ pi" using pi_not_less_zero by linarith qed auto subsection \Complex n-th roots\ lemma bij_betw_roots_unity: assumes "n > 0" shows "bij_betw (\k. cis (2 * pi * real k / real n)) {.. = cis (2*pi*(real k - real l)/n)" using assms by (simp add: field_simps cis_divide) finally have "cos (2*pi*(real k - real l) / n) = 1" by (simp add: complex_eq_iff) then obtain m :: int where "2 * pi * (real k - real l) / real n = real_of_int m * 2 * pi" by (subst (asm) cos_one_2pi_int) blast hence "real_of_int (int k - int l) = real_of_int (m * int n)" unfolding of_int_diff of_int_mult using assms by (simp add: nonzero_divide_eq_eq) also note of_int_eq_iff finally have *: "abs m * n = abs (int k - int l)" by (simp add: abs_mult) also have "\ < int n" using kl by linarith finally have "m = 0" using assms by simp with * show "k = l" by simp qed have subset: "?f ` {.. {z. z ^ n = 1}" proof safe fix k :: nat have "cis (2 * pi * real k / real n) ^ n = cis (2 * pi) ^ k" using assms by (simp add: DeMoivre mult_ac) also have "cis (2 * pi) = 1" by (simp add: complex_eq_iff) finally show "?f k ^ n = 1" by simp qed have "n = card {.. \ card {z::complex. z ^ n = 1}" by (intro card_inj_on_le[OF inj]) (auto simp: finite_roots_unity) finally have card: "card {z::complex. z ^ n = 1} = n" using assms by (intro antisym card_roots_unity) auto have "card (?f ` {.. 0" shows "card {z::complex. z ^ n = 1} = n" using bij_betw_same_card [OF bij_betw_roots_unity [OF assms]] by simp lemma bij_betw_nth_root_unity: fixes c :: complex and n :: nat assumes c: "c \ 0" and n: "n > 0" defines "c' \ root n (norm c) * cis (Arg c / n)" shows "bij_betw (\z. c' * z) {z. z ^ n = 1} {z. z ^ n = c}" proof - have "c' ^ n = of_real (root n (norm c) ^ n) * cis (Arg c)" unfolding of_real_power using n by (simp add: c'_def power_mult_distrib DeMoivre) also from n have "root n (norm c) ^ n = norm c" by simp also from c have "of_real \ * cis (Arg c) = c" by (simp add: cis_Arg Complex.sgn_eq) finally have [simp]: "c' ^ n = c" . show ?thesis unfolding bij_betw_def inj_on_def proof safe fix z :: complex assume "z ^ n = 1" hence "(c' * z) ^ n = c' ^ n" by (simp add: power_mult_distrib) also have "c' ^ n = of_real (root n (norm c) ^ n) * cis (Arg c)" unfolding of_real_power using n by (simp add: c'_def power_mult_distrib DeMoivre) also from n have "root n (norm c) ^ n = norm c" by simp also from c have "\ * cis (Arg c) = c" by (simp add: cis_Arg Complex.sgn_eq) finally show "(c' * z) ^ n = c" . next fix z assume z: "c = z ^ n" define z' where "z' = z / c'" from c and n have "c' \ 0" by (auto simp: c'_def) with n c have "z = c' * z'" and "z' ^ n = 1" by (auto simp: z'_def power_divide z) thus "z \ (\z. c' * z) ` {z. z ^ n = 1}" by blast qed (insert c n, auto simp: c'_def) qed lemma finite_nth_roots [intro]: assumes "n > 0" shows "finite {z::complex. z ^ n = c}" proof (cases "c = 0") case True with assms have "{z::complex. z ^ n = c} = {0}" by auto thus ?thesis by simp next case False from assms have "finite {z::complex. z ^ n = 1}" by (intro finite_roots_unity) simp_all also have "?this \ ?thesis" by (rule bij_betw_finite, rule bij_betw_nth_root_unity) fact+ finally show ?thesis . qed lemma card_nth_roots: assumes "c \ 0" "n > 0" shows "card {z::complex. z ^ n = c} = n" proof - have "card {z. z ^ n = c} = card {z::complex. z ^ n = 1}" by (rule sym, rule bij_betw_same_card, rule bij_betw_nth_root_unity) fact+ also have "\ = n" by (rule card_roots_unity_eq) fact+ finally show ?thesis . qed lemma sum_roots_unity: assumes "n > 1" shows "\{z::complex. z ^ n = 1} = 0" proof - define \ where "\ = cis (2 * pi / real n)" have [simp]: "\ \ 1" proof assume "\ = 1" with assms obtain k :: int where "2 * pi / real n = 2 * pi * of_int k" by (auto simp: \_def complex_eq_iff cos_one_2pi_int) with assms have "real n * of_int k = 1" by (simp add: field_simps) also have "real n * of_int k = of_int (int n * k)" by simp also have "1 = (of_int 1 :: real)" by simp also note of_int_eq_iff finally show False using assms by (auto simp: zmult_eq_1_iff) qed have "(\z | z ^ n = 1. z :: complex) = (\k = (\k ^ k)" by (intro sum.cong refl) (auto simp: \_def DeMoivre mult_ac) also have "\ = (\ ^ n - 1) / (\ - 1)" by (subst geometric_sum) auto also have "\ ^ n - 1 = cis (2 * pi) - 1" using assms by (auto simp: \_def DeMoivre) also have "\ = 0" by (simp add: complex_eq_iff) finally show ?thesis by simp qed lemma sum_nth_roots: assumes "n > 1" shows "\{z::complex. z ^ n = c} = 0" proof (cases "c = 0") case True with assms have "{z::complex. z ^ n = c} = {0}" by auto also have "\\ = 0" by simp finally show ?thesis . next case False define c' where "c' = root n (norm c) * cis (Arg c / n)" from False and assms have "(\{z. z ^ n = c}) = (\z | z ^ n = 1. c' * z)" by (subst sum.reindex_bij_betw [OF bij_betw_nth_root_unity, symmetric]) (auto simp: sum_distrib_left finite_roots_unity c'_def) also from assms have "\ = 0" by (simp add: sum_distrib_left [symmetric] sum_roots_unity) finally show ?thesis . qed subsection \Square root of complex numbers\ primcorec csqrt :: "complex \ complex" where "Re (csqrt z) = sqrt ((cmod z + Re z) / 2)" | "Im (csqrt z) = (if Im z = 0 then 1 else sgn (Im z)) * sqrt ((cmod z - Re z) / 2)" lemma csqrt_of_real_nonneg [simp]: "Im x = 0 \ Re x \ 0 \ csqrt x = sqrt (Re x)" by (simp add: complex_eq_iff norm_complex_def) lemma csqrt_of_real_nonpos [simp]: "Im x = 0 \ Re x \ 0 \ csqrt x = \ * sqrt \Re x\" by (simp add: complex_eq_iff norm_complex_def) lemma of_real_sqrt: "x \ 0 \ of_real (sqrt x) = csqrt (of_real x)" by (simp add: complex_eq_iff norm_complex_def) lemma csqrt_0 [simp]: "csqrt 0 = 0" by simp lemma csqrt_1 [simp]: "csqrt 1 = 1" by simp lemma csqrt_ii [simp]: "csqrt \ = (1 + \) / sqrt 2" by (simp add: complex_eq_iff Re_divide Im_divide real_sqrt_divide real_div_sqrt) lemma power2_csqrt[simp,algebra]: "(csqrt z)\<^sup>2 = z" proof (cases "Im z = 0") case True then show ?thesis using real_sqrt_pow2[of "Re z"] real_sqrt_pow2[of "- Re z"] by (cases "0::real" "Re z" rule: linorder_cases) (simp_all add: complex_eq_iff Re_power2 Im_power2 power2_eq_square cmod_eq_Re) next case False moreover have "cmod z * cmod z - Re z * Re z = Im z * Im z" by (simp add: norm_complex_def power2_eq_square) moreover have "\Re z\ \ cmod z" by (simp add: norm_complex_def) ultimately show ?thesis by (simp add: Re_power2 Im_power2 complex_eq_iff real_sgn_eq field_simps real_sqrt_mult[symmetric] real_sqrt_divide) qed +lemma norm_csqrt [simp]: "norm (csqrt z) = sqrt (norm z)" + by (metis abs_of_nonneg norm_ge_zero norm_mult power2_csqrt power2_eq_square real_sqrt_abs) + lemma csqrt_eq_0 [simp]: "csqrt z = 0 \ z = 0" by auto (metis power2_csqrt power_eq_0_iff) lemma csqrt_eq_1 [simp]: "csqrt z = 1 \ z = 1" by auto (metis power2_csqrt power2_eq_1_iff) lemma csqrt_principal: "0 < Re (csqrt z) \ Re (csqrt z) = 0 \ 0 \ Im (csqrt z)" by (auto simp add: not_less cmod_plus_Re_le_0_iff Im_eq_0) lemma Re_csqrt: "0 \ Re (csqrt z)" by (metis csqrt_principal le_less) lemma csqrt_square: assumes "0 < Re b \ (Re b = 0 \ 0 \ Im b)" shows "csqrt (b^2) = b" proof - have "csqrt (b^2) = b \ csqrt (b^2) = - b" by (simp add: power2_eq_iff[symmetric]) moreover have "csqrt (b^2) \ -b \ b = 0" using csqrt_principal[of "b ^ 2"] assms by (intro disjCI notI) (auto simp: complex_eq_iff) ultimately show ?thesis by auto qed lemma csqrt_unique: "w\<^sup>2 = z \ 0 < Re w \ Re w = 0 \ 0 \ Im w \ csqrt z = w" by (auto simp: csqrt_square) lemma csqrt_minus [simp]: assumes "Im x < 0 \ (Im x = 0 \ 0 \ Re x)" shows "csqrt (- x) = \ * csqrt x" proof - have "csqrt ((\ * csqrt x)^2) = \ * csqrt x" proof (rule csqrt_square) have "Im (csqrt x) \ 0" using assms by (auto simp add: cmod_eq_Re mult_le_0_iff field_simps complex_Re_le_cmod) then show "0 < Re (\ * csqrt x) \ Re (\ * csqrt x) = 0 \ 0 \ Im (\ * csqrt x)" by (auto simp add: Re_csqrt simp del: csqrt.simps) qed also have "(\ * csqrt x)^2 = - x" by (simp add: power_mult_distrib) finally show ?thesis . qed text \Legacy theorem names\ lemmas cmod_def = norm_complex_def lemma legacy_Complex_simps: shows Complex_eq_0: "Complex a b = 0 \ a = 0 \ b = 0" and complex_add: "Complex a b + Complex c d = Complex (a + c) (b + d)" and complex_minus: "- (Complex a b) = Complex (- a) (- b)" and complex_diff: "Complex a b - Complex c d = Complex (a - c) (b - d)" and Complex_eq_1: "Complex a b = 1 \ a = 1 \ b = 0" and Complex_eq_neg_1: "Complex a b = - 1 \ a = - 1 \ b = 0" and complex_mult: "Complex a b * Complex c d = Complex (a * c - b * d) (a * d + b * c)" and complex_inverse: "inverse (Complex a b) = Complex (a / (a\<^sup>2 + b\<^sup>2)) (- b / (a\<^sup>2 + b\<^sup>2))" and Complex_eq_numeral: "Complex a b = numeral w \ a = numeral w \ b = 0" and Complex_eq_neg_numeral: "Complex a b = - numeral w \ a = - numeral w \ b = 0" and complex_scaleR: "scaleR r (Complex a b) = Complex (r * a) (r * b)" and Complex_eq_i: "Complex x y = \ \ x = 0 \ y = 1" and i_mult_Complex: "\ * Complex a b = Complex (- b) a" and Complex_mult_i: "Complex a b * \ = Complex (- b) a" and i_complex_of_real: "\ * complex_of_real r = Complex 0 r" and complex_of_real_i: "complex_of_real r * \ = Complex 0 r" and Complex_add_complex_of_real: "Complex x y + complex_of_real r = Complex (x+r) y" and complex_of_real_add_Complex: "complex_of_real r + Complex x y = Complex (r+x) y" and Complex_mult_complex_of_real: "Complex x y * complex_of_real r = Complex (x*r) (y*r)" and complex_of_real_mult_Complex: "complex_of_real r * Complex x y = Complex (r*x) (r*y)" and complex_eq_cancel_iff2: "(Complex x y = complex_of_real xa) = (x = xa \ y = 0)" and complex_cnj: "cnj (Complex a b) = Complex a (- b)" and Complex_sum': "sum (\x. Complex (f x) 0) s = Complex (sum f s) 0" and Complex_sum: "Complex (sum f s) 0 = sum (\x. Complex (f x) 0) s" and complex_of_real_def: "complex_of_real r = Complex r 0" and complex_norm: "cmod (Complex x y) = sqrt (x\<^sup>2 + y\<^sup>2)" by (simp_all add: norm_complex_def field_simps complex_eq_iff Re_divide Im_divide) lemma Complex_in_Reals: "Complex x 0 \ \" by (metis Reals_of_real complex_of_real_def) text \Express a complex number as a linear combination of two others, not collinear with the origin\ lemma complex_axes: assumes "Im (y/x) \ 0" obtains a b where "z = of_real a * x + of_real b * y" proof - define dd where "dd \ Re y * Im x - Im y * Re x" define a where "a = (Im z * Re y - Re z * Im y) / dd" define b where "b = (Re z * Im x - Im z * Re x) / dd" have "dd \ 0" using assms by (auto simp: dd_def Im_complex_div_eq_0) have "a * Re x + b * Re y = Re z" using \dd \ 0\ apply (simp add: a_def b_def field_simps) by (metis dd_def diff_add_cancel distrib_right mult.assoc mult.commute) moreover have "a * Im x + b * Im y = Im z" using \dd \ 0\ apply (simp add: a_def b_def field_simps) by (metis (no_types) dd_def diff_add_cancel distrib_right mult.assoc mult.commute) ultimately have "z = of_real a * x + of_real b * y" by (simp add: complex_eqI) then show ?thesis using that by simp qed end diff --git a/src/HOL/Deriv.thy b/src/HOL/Deriv.thy --- a/src/HOL/Deriv.thy +++ b/src/HOL/Deriv.thy @@ -1,2427 +1,2430 @@ (* Title: HOL/Deriv.thy Author: Jacques D. Fleuriot, University of Cambridge, 1998 Author: Brian Huffman Author: Lawrence C Paulson, 2004 Author: Benjamin Porter, 2005 *) section \Differentiation\ theory Deriv imports Limits begin subsection \Frechet derivative\ definition has_derivative :: "('a::real_normed_vector \ 'b::real_normed_vector) \ ('a \ 'b) \ 'a filter \ bool" (infix "(has'_derivative)" 50) where "(f has_derivative f') F \ bounded_linear f' \ ((\y. ((f y - f (Lim F (\x. x))) - f' (y - Lim F (\x. x))) /\<^sub>R norm (y - Lim F (\x. x))) \ 0) F" text \ Usually the filter \<^term>\F\ is \<^term>\at x within s\. \<^term>\(f has_derivative D) (at x within s)\ means: \<^term>\D\ is the derivative of function \<^term>\f\ at point \<^term>\x\ within the set \<^term>\s\. Where \<^term>\s\ is used to express left or right sided derivatives. In most cases \<^term>\s\ is either a variable or \<^term>\UNIV\. \ text \These are the only cases we'll care about, probably.\ lemma has_derivative_within: "(f has_derivative f') (at x within s) \ bounded_linear f' \ ((\y. (1 / norm(y - x)) *\<^sub>R (f y - (f x + f' (y - x)))) \ 0) (at x within s)" unfolding has_derivative_def tendsto_iff by (subst eventually_Lim_ident_at) (auto simp add: field_simps) lemma has_derivative_eq_rhs: "(f has_derivative f') F \ f' = g' \ (f has_derivative g') F" by simp definition has_field_derivative :: "('a::real_normed_field \ 'a) \ 'a \ 'a filter \ bool" (infix "(has'_field'_derivative)" 50) where "(f has_field_derivative D) F \ (f has_derivative (*) D) F" lemma DERIV_cong: "(f has_field_derivative X) F \ X = Y \ (f has_field_derivative Y) F" by simp definition has_vector_derivative :: "(real \ 'b::real_normed_vector) \ 'b \ real filter \ bool" (infix "has'_vector'_derivative" 50) where "(f has_vector_derivative f') net \ (f has_derivative (\x. x *\<^sub>R f')) net" lemma has_vector_derivative_eq_rhs: "(f has_vector_derivative X) F \ X = Y \ (f has_vector_derivative Y) F" by simp named_theorems derivative_intros "structural introduction rules for derivatives" setup \ let val eq_thms = @{thms has_derivative_eq_rhs DERIV_cong has_vector_derivative_eq_rhs} fun eq_rule thm = get_first (try (fn eq_thm => eq_thm OF [thm])) eq_thms in Global_Theory.add_thms_dynamic (\<^binding>\derivative_eq_intros\, fn context => Named_Theorems.get (Context.proof_of context) \<^named_theorems>\derivative_intros\ |> map_filter eq_rule) end \ text \ The following syntax is only used as a legacy syntax. \ abbreviation (input) FDERIV :: "('a::real_normed_vector \ 'b::real_normed_vector) \ 'a \ ('a \ 'b) \ bool" ("(FDERIV (_)/ (_)/ :> (_))" [1000, 1000, 60] 60) where "FDERIV f x :> f' \ (f has_derivative f') (at x)" lemma has_derivative_bounded_linear: "(f has_derivative f') F \ bounded_linear f'" by (simp add: has_derivative_def) lemma has_derivative_linear: "(f has_derivative f') F \ linear f'" using bounded_linear.linear[OF has_derivative_bounded_linear] . lemma has_derivative_ident[derivative_intros, simp]: "((\x. x) has_derivative (\x. x)) F" by (simp add: has_derivative_def) -lemma has_derivative_id [derivative_intros, simp]: "(id has_derivative id) (at a)" +lemma has_derivative_id [derivative_intros, simp]: "(id has_derivative id) F" by (metis eq_id_iff has_derivative_ident) +lemma shift_has_derivative_id: "((+) d has_derivative (\x. x)) F" + using has_derivative_def by fastforce + lemma has_derivative_const[derivative_intros, simp]: "((\x. c) has_derivative (\x. 0)) F" by (simp add: has_derivative_def) lemma (in bounded_linear) bounded_linear: "bounded_linear f" .. lemma (in bounded_linear) has_derivative: "(g has_derivative g') F \ ((\x. f (g x)) has_derivative (\x. f (g' x))) F" unfolding has_derivative_def by (auto simp add: bounded_linear_compose [OF bounded_linear] scaleR diff dest: tendsto) lemmas has_derivative_scaleR_right [derivative_intros] = bounded_linear.has_derivative [OF bounded_linear_scaleR_right] lemmas has_derivative_scaleR_left [derivative_intros] = bounded_linear.has_derivative [OF bounded_linear_scaleR_left] lemmas has_derivative_mult_right [derivative_intros] = bounded_linear.has_derivative [OF bounded_linear_mult_right] lemmas has_derivative_mult_left [derivative_intros] = bounded_linear.has_derivative [OF bounded_linear_mult_left] lemmas has_derivative_of_real[derivative_intros, simp] = bounded_linear.has_derivative[OF bounded_linear_of_real] lemma has_derivative_add[simp, derivative_intros]: assumes f: "(f has_derivative f') F" and g: "(g has_derivative g') F" shows "((\x. f x + g x) has_derivative (\x. f' x + g' x)) F" unfolding has_derivative_def proof safe let ?x = "Lim F (\x. x)" let ?D = "\f f' y. ((f y - f ?x) - f' (y - ?x)) /\<^sub>R norm (y - ?x)" have "((\x. ?D f f' x + ?D g g' x) \ (0 + 0)) F" using f g by (intro tendsto_add) (auto simp: has_derivative_def) then show "(?D (\x. f x + g x) (\x. f' x + g' x) \ 0) F" by (simp add: field_simps scaleR_add_right scaleR_diff_right) qed (blast intro: bounded_linear_add f g has_derivative_bounded_linear) lemma has_derivative_sum[simp, derivative_intros]: "(\i. i \ I \ (f i has_derivative f' i) F) \ ((\x. \i\I. f i x) has_derivative (\x. \i\I. f' i x)) F" by (induct I rule: infinite_finite_induct) simp_all lemma has_derivative_minus[simp, derivative_intros]: "(f has_derivative f') F \ ((\x. - f x) has_derivative (\x. - f' x)) F" using has_derivative_scaleR_right[of f f' F "-1"] by simp lemma has_derivative_diff[simp, derivative_intros]: "(f has_derivative f') F \ (g has_derivative g') F \ ((\x. f x - g x) has_derivative (\x. f' x - g' x)) F" by (simp only: diff_conv_add_uminus has_derivative_add has_derivative_minus) lemma has_derivative_at_within: "(f has_derivative f') (at x within s) \ (bounded_linear f' \ ((\y. ((f y - f x) - f' (y - x)) /\<^sub>R norm (y - x)) \ 0) (at x within s))" proof (cases "at x within s = bot") case True then show ?thesis by (metis (no_types, lifting) has_derivative_within tendsto_bot) next case False then show ?thesis by (simp add: Lim_ident_at has_derivative_def) qed lemma has_derivative_iff_norm: "(f has_derivative f') (at x within s) \ bounded_linear f' \ ((\y. norm ((f y - f x) - f' (y - x)) / norm (y - x)) \ 0) (at x within s)" using tendsto_norm_zero_iff[of _ "at x within s", where 'b="'b", symmetric] by (simp add: has_derivative_at_within divide_inverse ac_simps) lemma has_derivative_at: "(f has_derivative D) (at x) \ (bounded_linear D \ (\h. norm (f (x + h) - f x - D h) / norm h) \0\ 0)" by (simp add: has_derivative_iff_norm LIM_offset_zero_iff) lemma field_has_derivative_at: fixes x :: "'a::real_normed_field" shows "(f has_derivative (*) D) (at x) \ (\h. (f (x + h) - f x) / h) \0\ D" (is "?lhs = ?rhs") proof - have "?lhs = (\h. norm (f (x + h) - f x - D * h) / norm h) \0 \ 0" by (simp add: bounded_linear_mult_right has_derivative_at) also have "... = (\y. norm ((f (x + y) - f x - D * y) / y)) \0\ 0" by (simp cong: LIM_cong flip: nonzero_norm_divide) also have "... = (\y. norm ((f (x + y) - f x) / y - D / y * y)) \0\ 0" by (simp only: diff_divide_distrib times_divide_eq_left [symmetric]) also have "... = ?rhs" by (simp add: tendsto_norm_zero_iff LIM_zero_iff cong: LIM_cong) finally show ?thesis . qed lemma has_derivative_iff_Ex: "(f has_derivative f') (at x) \ bounded_linear f' \ (\e. (\h. f (x+h) = f x + f' h + e h) \ ((\h. norm (e h) / norm h) \ 0) (at 0))" unfolding has_derivative_at by force lemma has_derivative_at_within_iff_Ex: assumes "x \ S" "open S" shows "(f has_derivative f') (at x within S) \ bounded_linear f' \ (\e. (\h. x+h \ S \ f (x+h) = f x + f' h + e h) \ ((\h. norm (e h) / norm h) \ 0) (at 0))" (is "?lhs = ?rhs") proof safe show "bounded_linear f'" if "(f has_derivative f') (at x within S)" using has_derivative_bounded_linear that by blast show "\e. (\h. x + h \ S \ f (x + h) = f x + f' h + e h) \ (\h. norm (e h) / norm h) \0\ 0" if "(f has_derivative f') (at x within S)" by (metis (full_types) assms that has_derivative_iff_Ex at_within_open) show "(f has_derivative f') (at x within S)" if "bounded_linear f'" and eq [rule_format]: "\h. x + h \ S \ f (x + h) = f x + f' h + e h" and 0: "(\h. norm (e (h::'a)::'b) / norm h) \0\ 0" for e proof - have 1: "f y - f x = f' (y-x) + e (y-x)" if "y \ S" for y using eq [of "y-x"] that by simp have 2: "((\y. norm (e (y-x)) / norm (y - x)) \ 0) (at x within S)" by (simp add: "0" assms tendsto_offset_zero_iff) have "((\y. norm (f y - f x - f' (y - x)) / norm (y - x)) \ 0) (at x within S)" by (simp add: Lim_cong_within 1 2) then show ?thesis by (simp add: has_derivative_iff_norm \bounded_linear f'\) qed qed lemma has_derivativeI: "bounded_linear f' \ ((\y. ((f y - f x) - f' (y - x)) /\<^sub>R norm (y - x)) \ 0) (at x within s) \ (f has_derivative f') (at x within s)" by (simp add: has_derivative_at_within) lemma has_derivativeI_sandwich: assumes e: "0 < e" and bounded: "bounded_linear f'" and sandwich: "(\y. y \ s \ y \ x \ dist y x < e \ norm ((f y - f x) - f' (y - x)) / norm (y - x) \ H y)" and "(H \ 0) (at x within s)" shows "(f has_derivative f') (at x within s)" unfolding has_derivative_iff_norm proof safe show "((\y. norm (f y - f x - f' (y - x)) / norm (y - x)) \ 0) (at x within s)" proof (rule tendsto_sandwich[where f="\x. 0"]) show "(H \ 0) (at x within s)" by fact show "eventually (\n. norm (f n - f x - f' (n - x)) / norm (n - x) \ H n) (at x within s)" unfolding eventually_at using e sandwich by auto qed (auto simp: le_divide_eq) qed fact lemma has_derivative_subset: "(f has_derivative f') (at x within s) \ t \ s \ (f has_derivative f') (at x within t)" by (auto simp add: has_derivative_iff_norm intro: tendsto_within_subset) lemma has_derivative_within_singleton_iff: "(f has_derivative g) (at x within {x}) \ bounded_linear g" by (auto intro!: has_derivativeI_sandwich[where e=1] has_derivative_bounded_linear) subsubsection \Limit transformation for derivatives\ lemma has_derivative_transform_within: assumes "(f has_derivative f') (at x within s)" and "0 < d" and "x \ s" and "\x'. \x' \ s; dist x' x < d\ \ f x' = g x'" shows "(g has_derivative f') (at x within s)" using assms unfolding has_derivative_within by (force simp add: intro: Lim_transform_within) lemma has_derivative_transform_within_open: assumes "(f has_derivative f') (at x within t)" and "open s" and "x \ s" and "\x. x\s \ f x = g x" shows "(g has_derivative f') (at x within t)" using assms unfolding has_derivative_within by (force simp add: intro: Lim_transform_within_open) lemma has_derivative_transform: assumes "x \ s" "\x. x \ s \ g x = f x" assumes "(f has_derivative f') (at x within s)" shows "(g has_derivative f') (at x within s)" using assms by (intro has_derivative_transform_within[OF _ zero_less_one, where g=g]) auto lemma has_derivative_transform_eventually: assumes "(f has_derivative f') (at x within s)" "(\\<^sub>F x' in at x within s. f x' = g x')" assumes "f x = g x" "x \ s" shows "(g has_derivative f') (at x within s)" using assms proof - from assms(2,3) obtain d where "d > 0" "\x'. x' \ s \ dist x' x < d \ f x' = g x'" by (force simp: eventually_at) from has_derivative_transform_within[OF assms(1) this(1) assms(4) this(2)] show ?thesis . qed lemma has_field_derivative_transform_within: assumes "(f has_field_derivative f') (at a within S)" and "0 < d" and "a \ S" and "\x. \x \ S; dist x a < d\ \ f x = g x" shows "(g has_field_derivative f') (at a within S)" using assms unfolding has_field_derivative_def by (metis has_derivative_transform_within) lemma has_field_derivative_transform_within_open: assumes "(f has_field_derivative f') (at a)" and "open S" "a \ S" and "\x. x \ S \ f x = g x" shows "(g has_field_derivative f') (at a)" using assms unfolding has_field_derivative_def by (metis has_derivative_transform_within_open) subsection \Continuity\ lemma has_derivative_continuous: assumes f: "(f has_derivative f') (at x within s)" shows "continuous (at x within s) f" proof - from f interpret F: bounded_linear f' by (rule has_derivative_bounded_linear) note F.tendsto[tendsto_intros] let ?L = "\f. (f \ 0) (at x within s)" have "?L (\y. norm ((f y - f x) - f' (y - x)) / norm (y - x))" using f unfolding has_derivative_iff_norm by blast then have "?L (\y. norm ((f y - f x) - f' (y - x)) / norm (y - x) * norm (y - x))" (is ?m) by (rule tendsto_mult_zero) (auto intro!: tendsto_eq_intros) also have "?m \ ?L (\y. norm ((f y - f x) - f' (y - x)))" by (intro filterlim_cong) (simp_all add: eventually_at_filter) finally have "?L (\y. (f y - f x) - f' (y - x))" by (rule tendsto_norm_zero_cancel) then have "?L (\y. ((f y - f x) - f' (y - x)) + f' (y - x))" by (rule tendsto_eq_intros) (auto intro!: tendsto_eq_intros simp: F.zero) then have "?L (\y. f y - f x)" by simp from tendsto_add[OF this tendsto_const, of "f x"] show ?thesis by (simp add: continuous_within) qed subsection \Composition\ lemma tendsto_at_iff_tendsto_nhds_within: "f x = y \ (f \ y) (at x within s) \ (f \ y) (inf (nhds x) (principal s))" unfolding tendsto_def eventually_inf_principal eventually_at_filter by (intro ext all_cong imp_cong) (auto elim!: eventually_mono) lemma has_derivative_in_compose: assumes f: "(f has_derivative f') (at x within s)" and g: "(g has_derivative g') (at (f x) within (f`s))" shows "((\x. g (f x)) has_derivative (\x. g' (f' x))) (at x within s)" proof - from f interpret F: bounded_linear f' by (rule has_derivative_bounded_linear) from g interpret G: bounded_linear g' by (rule has_derivative_bounded_linear) from F.bounded obtain kF where kF: "\x. norm (f' x) \ norm x * kF" by fast from G.bounded obtain kG where kG: "\x. norm (g' x) \ norm x * kG" by fast note G.tendsto[tendsto_intros] let ?L = "\f. (f \ 0) (at x within s)" let ?D = "\f f' x y. (f y - f x) - f' (y - x)" let ?N = "\f f' x y. norm (?D f f' x y) / norm (y - x)" let ?gf = "\x. g (f x)" and ?gf' = "\x. g' (f' x)" define Nf where "Nf = ?N f f' x" define Ng where [abs_def]: "Ng y = ?N g g' (f x) (f y)" for y show ?thesis proof (rule has_derivativeI_sandwich[of 1]) show "bounded_linear (\x. g' (f' x))" using f g by (blast intro: bounded_linear_compose has_derivative_bounded_linear) next fix y :: 'a assume neq: "y \ x" have "?N ?gf ?gf' x y = norm (g' (?D f f' x y) + ?D g g' (f x) (f y)) / norm (y - x)" by (simp add: G.diff G.add field_simps) also have "\ \ norm (g' (?D f f' x y)) / norm (y - x) + Ng y * (norm (f y - f x) / norm (y - x))" by (simp add: add_divide_distrib[symmetric] divide_right_mono norm_triangle_ineq G.zero Ng_def) also have "\ \ Nf y * kG + Ng y * (Nf y + kF)" proof (intro add_mono mult_left_mono) have "norm (f y - f x) = norm (?D f f' x y + f' (y - x))" by simp also have "\ \ norm (?D f f' x y) + norm (f' (y - x))" by (rule norm_triangle_ineq) also have "\ \ norm (?D f f' x y) + norm (y - x) * kF" using kF by (intro add_mono) simp finally show "norm (f y - f x) / norm (y - x) \ Nf y + kF" by (simp add: neq Nf_def field_simps) qed (use kG in \simp_all add: Ng_def Nf_def neq zero_le_divide_iff field_simps\) finally show "?N ?gf ?gf' x y \ Nf y * kG + Ng y * (Nf y + kF)" . next have [tendsto_intros]: "?L Nf" using f unfolding has_derivative_iff_norm Nf_def .. from f have "(f \ f x) (at x within s)" by (blast intro: has_derivative_continuous continuous_within[THEN iffD1]) then have f': "LIM x at x within s. f x :> inf (nhds (f x)) (principal (f`s))" unfolding filterlim_def by (simp add: eventually_filtermap eventually_at_filter le_principal) have "((?N g g' (f x)) \ 0) (at (f x) within f`s)" using g unfolding has_derivative_iff_norm .. then have g': "((?N g g' (f x)) \ 0) (inf (nhds (f x)) (principal (f`s)))" by (rule tendsto_at_iff_tendsto_nhds_within[THEN iffD1, rotated]) simp have [tendsto_intros]: "?L Ng" unfolding Ng_def by (rule filterlim_compose[OF g' f']) show "((\y. Nf y * kG + Ng y * (Nf y + kF)) \ 0) (at x within s)" by (intro tendsto_eq_intros) auto qed simp qed lemma has_derivative_compose: "(f has_derivative f') (at x within s) \ (g has_derivative g') (at (f x)) \ ((\x. g (f x)) has_derivative (\x. g' (f' x))) (at x within s)" by (blast intro: has_derivative_in_compose has_derivative_subset) lemma has_derivative_in_compose2: assumes "\x. x \ t \ (g has_derivative g' x) (at x within t)" assumes "f ` s \ t" "x \ s" assumes "(f has_derivative f') (at x within s)" shows "((\x. g (f x)) has_derivative (\y. g' (f x) (f' y))) (at x within s)" using assms by (auto intro: has_derivative_subset intro!: has_derivative_in_compose[of f f' x s g]) lemma (in bounded_bilinear) FDERIV: assumes f: "(f has_derivative f') (at x within s)" and g: "(g has_derivative g') (at x within s)" shows "((\x. f x ** g x) has_derivative (\h. f x ** g' h + f' h ** g x)) (at x within s)" proof - from bounded_linear.bounded [OF has_derivative_bounded_linear [OF f]] obtain KF where norm_F: "\x. norm (f' x) \ norm x * KF" by fast from pos_bounded obtain K where K: "0 < K" and norm_prod: "\a b. norm (a ** b) \ norm a * norm b * K" by fast let ?D = "\f f' y. f y - f x - f' (y - x)" let ?N = "\f f' y. norm (?D f f' y) / norm (y - x)" define Ng where "Ng = ?N g g'" define Nf where "Nf = ?N f f'" let ?fun1 = "\y. norm (f y ** g y - f x ** g x - (f x ** g' (y - x) + f' (y - x) ** g x)) / norm (y - x)" let ?fun2 = "\y. norm (f x) * Ng y * K + Nf y * norm (g y) * K + KF * norm (g y - g x) * K" let ?F = "at x within s" show ?thesis proof (rule has_derivativeI_sandwich[of 1]) show "bounded_linear (\h. f x ** g' h + f' h ** g x)" by (intro bounded_linear_add bounded_linear_compose [OF bounded_linear_right] bounded_linear_compose [OF bounded_linear_left] has_derivative_bounded_linear [OF g] has_derivative_bounded_linear [OF f]) next from g have "(g \ g x) ?F" by (intro continuous_within[THEN iffD1] has_derivative_continuous) moreover from f g have "(Nf \ 0) ?F" "(Ng \ 0) ?F" by (simp_all add: has_derivative_iff_norm Ng_def Nf_def) ultimately have "(?fun2 \ norm (f x) * 0 * K + 0 * norm (g x) * K + KF * norm (0::'b) * K) ?F" by (intro tendsto_intros) (simp_all add: LIM_zero_iff) then show "(?fun2 \ 0) ?F" by simp next fix y :: 'd assume "y \ x" have "?fun1 y = norm (f x ** ?D g g' y + ?D f f' y ** g y + f' (y - x) ** (g y - g x)) / norm (y - x)" by (simp add: diff_left diff_right add_left add_right field_simps) also have "\ \ (norm (f x) * norm (?D g g' y) * K + norm (?D f f' y) * norm (g y) * K + norm (y - x) * KF * norm (g y - g x) * K) / norm (y - x)" by (intro divide_right_mono mult_mono' order_trans [OF norm_triangle_ineq add_mono] order_trans [OF norm_prod mult_right_mono] mult_nonneg_nonneg order_refl norm_ge_zero norm_F K [THEN order_less_imp_le]) also have "\ = ?fun2 y" by (simp add: add_divide_distrib Ng_def Nf_def) finally show "?fun1 y \ ?fun2 y" . qed simp qed lemmas has_derivative_mult[simp, derivative_intros] = bounded_bilinear.FDERIV[OF bounded_bilinear_mult] lemmas has_derivative_scaleR[simp, derivative_intros] = bounded_bilinear.FDERIV[OF bounded_bilinear_scaleR] lemma has_derivative_prod[simp, derivative_intros]: fixes f :: "'i \ 'a::real_normed_vector \ 'b::real_normed_field" shows "(\i. i \ I \ (f i has_derivative f' i) (at x within S)) \ ((\x. \i\I. f i x) has_derivative (\y. \i\I. f' i y * (\j\I - {i}. f j x))) (at x within S)" proof (induct I rule: infinite_finite_induct) case infinite then show ?case by simp next case empty then show ?case by simp next case (insert i I) let ?P = "\y. f i x * (\i\I. f' i y * (\j\I - {i}. f j x)) + (f' i y) * (\i\I. f i x)" have "((\x. f i x * (\i\I. f i x)) has_derivative ?P) (at x within S)" using insert by (intro has_derivative_mult) auto also have "?P = (\y. \i'\insert i I. f' i' y * (\j\insert i I - {i'}. f j x))" using insert(1,2) by (auto simp add: sum_distrib_left insert_Diff_if intro!: ext sum.cong) finally show ?case using insert by simp qed lemma has_derivative_power[simp, derivative_intros]: fixes f :: "'a :: real_normed_vector \ 'b :: real_normed_field" assumes f: "(f has_derivative f') (at x within S)" shows "((\x. f x^n) has_derivative (\y. of_nat n * f' y * f x^(n - 1))) (at x within S)" using has_derivative_prod[OF f, of "{..< n}"] by (simp add: prod_constant ac_simps) lemma has_derivative_inverse': fixes x :: "'a::real_normed_div_algebra" assumes x: "x \ 0" shows "(inverse has_derivative (\h. - (inverse x * h * inverse x))) (at x within S)" (is "(_ has_derivative ?f) _") proof (rule has_derivativeI_sandwich) show "bounded_linear (\h. - (inverse x * h * inverse x))" by (simp add: bounded_linear_minus bounded_linear_mult_const bounded_linear_mult_right) show "0 < norm x" using x by simp have "(inverse \ inverse x) (at x within S)" using tendsto_inverse tendsto_ident_at x by auto then show "((\y. norm (inverse y - inverse x) * norm (inverse x)) \ 0) (at x within S)" by (simp add: LIM_zero_iff tendsto_mult_left_zero tendsto_norm_zero) next fix y :: 'a assume h: "y \ x" "dist y x < norm x" then have "y \ 0" by auto have "norm (inverse y - inverse x - ?f (y -x)) / norm (y - x) = norm (- (inverse y * (y - x) * inverse x - inverse x * (y - x) * inverse x)) / norm (y - x)" by (simp add: \y \ 0\ inverse_diff_inverse x) also have "... = norm ((inverse y - inverse x) * (y - x) * inverse x) / norm (y - x)" by (simp add: left_diff_distrib norm_minus_commute) also have "\ \ norm (inverse y - inverse x) * norm (y - x) * norm (inverse x) / norm (y - x)" by (simp add: norm_mult) also have "\ = norm (inverse y - inverse x) * norm (inverse x)" by simp finally show "norm (inverse y - inverse x - ?f (y -x)) / norm (y - x) \ norm (inverse y - inverse x) * norm (inverse x)" . qed lemma has_derivative_inverse[simp, derivative_intros]: fixes f :: "_ \ 'a::real_normed_div_algebra" assumes x: "f x \ 0" and f: "(f has_derivative f') (at x within S)" shows "((\x. inverse (f x)) has_derivative (\h. - (inverse (f x) * f' h * inverse (f x)))) (at x within S)" using has_derivative_compose[OF f has_derivative_inverse', OF x] . lemma has_derivative_divide[simp, derivative_intros]: fixes f :: "_ \ 'a::real_normed_div_algebra" assumes f: "(f has_derivative f') (at x within S)" and g: "(g has_derivative g') (at x within S)" assumes x: "g x \ 0" shows "((\x. f x / g x) has_derivative (\h. - f x * (inverse (g x) * g' h * inverse (g x)) + f' h / g x)) (at x within S)" using has_derivative_mult[OF f has_derivative_inverse[OF x g]] by (simp add: field_simps) lemma has_derivative_power_int': fixes x :: "'a::real_normed_field" assumes x: "x \ 0" shows "((\x. power_int x n) has_derivative (\y. y * (of_int n * power_int x (n - 1)))) (at x within S)" proof (cases n rule: int_cases4) case (nonneg n) thus ?thesis using x by (cases "n = 0") (auto intro!: derivative_eq_intros simp: field_simps power_int_diff fun_eq_iff simp flip: power_Suc) next case (neg n) thus ?thesis using x by (auto intro!: derivative_eq_intros simp: field_simps power_int_diff power_int_minus simp flip: power_Suc power_Suc2 power_add) qed lemma has_derivative_power_int[simp, derivative_intros]: fixes f :: "_ \ 'a::real_normed_field" assumes x: "f x \ 0" and f: "(f has_derivative f') (at x within S)" shows "((\x. power_int (f x) n) has_derivative (\h. f' h * (of_int n * power_int (f x) (n - 1)))) (at x within S)" using has_derivative_compose[OF f has_derivative_power_int', OF x] . text \Conventional form requires mult-AC laws. Types real and complex only.\ lemma has_derivative_divide'[derivative_intros]: fixes f :: "_ \ 'a::real_normed_field" assumes f: "(f has_derivative f') (at x within S)" and g: "(g has_derivative g') (at x within S)" and x: "g x \ 0" shows "((\x. f x / g x) has_derivative (\h. (f' h * g x - f x * g' h) / (g x * g x))) (at x within S)" proof - have "f' h / g x - f x * (inverse (g x) * g' h * inverse (g x)) = (f' h * g x - f x * g' h) / (g x * g x)" for h by (simp add: field_simps x) then show ?thesis using has_derivative_divide [OF f g] x by simp qed subsection \Uniqueness\ text \ This can not generally shown for \<^const>\has_derivative\, as we need to approach the point from all directions. There is a proof in \Analysis\ for \euclidean_space\. \ lemma has_derivative_at2: "(f has_derivative f') (at x) \ bounded_linear f' \ ((\y. (1 / (norm(y - x))) *\<^sub>R (f y - (f x + f' (y - x)))) \ 0) (at x)" using has_derivative_within [of f f' x UNIV] by simp lemma has_derivative_zero_unique: assumes "((\x. 0) has_derivative F) (at x)" shows "F = (\h. 0)" proof - interpret F: bounded_linear F using assms by (rule has_derivative_bounded_linear) let ?r = "\h. norm (F h) / norm h" have *: "?r \0\ 0" using assms unfolding has_derivative_at by simp show "F = (\h. 0)" proof show "F h = 0" for h proof (rule ccontr) assume **: "\ ?thesis" then have h: "h \ 0" by (auto simp add: F.zero) with ** have "0 < ?r h" by simp from LIM_D [OF * this] obtain S where S: "0 < S" and r: "\x. x \ 0 \ norm x < S \ ?r x < ?r h" by auto from dense [OF S] obtain t where t: "0 < t \ t < S" .. let ?x = "scaleR (t / norm h) h" have "?x \ 0" and "norm ?x < S" using t h by simp_all then have "?r ?x < ?r h" by (rule r) then show False using t h by (simp add: F.scaleR) qed qed qed lemma has_derivative_unique: assumes "(f has_derivative F) (at x)" and "(f has_derivative F') (at x)" shows "F = F'" proof - have "((\x. 0) has_derivative (\h. F h - F' h)) (at x)" using has_derivative_diff [OF assms] by simp then have "(\h. F h - F' h) = (\h. 0)" by (rule has_derivative_zero_unique) then show "F = F'" unfolding fun_eq_iff right_minus_eq . qed lemma has_derivative_Uniq: "\\<^sub>\\<^sub>1F. (f has_derivative F) (at x)" by (simp add: Uniq_def has_derivative_unique) subsection \Differentiability predicate\ definition differentiable :: "('a::real_normed_vector \ 'b::real_normed_vector) \ 'a filter \ bool" (infix "differentiable" 50) where "f differentiable F \ (\D. (f has_derivative D) F)" lemma differentiable_subset: "f differentiable (at x within s) \ t \ s \ f differentiable (at x within t)" unfolding differentiable_def by (blast intro: has_derivative_subset) lemmas differentiable_within_subset = differentiable_subset lemma differentiable_ident [simp, derivative_intros]: "(\x. x) differentiable F" unfolding differentiable_def by (blast intro: has_derivative_ident) lemma differentiable_const [simp, derivative_intros]: "(\z. a) differentiable F" unfolding differentiable_def by (blast intro: has_derivative_const) lemma differentiable_in_compose: "f differentiable (at (g x) within (g`s)) \ g differentiable (at x within s) \ (\x. f (g x)) differentiable (at x within s)" unfolding differentiable_def by (blast intro: has_derivative_in_compose) lemma differentiable_compose: "f differentiable (at (g x)) \ g differentiable (at x within s) \ (\x. f (g x)) differentiable (at x within s)" by (blast intro: differentiable_in_compose differentiable_subset) lemma differentiable_add [simp, derivative_intros]: "f differentiable F \ g differentiable F \ (\x. f x + g x) differentiable F" unfolding differentiable_def by (blast intro: has_derivative_add) lemma differentiable_sum[simp, derivative_intros]: assumes "finite s" "\a\s. (f a) differentiable net" shows "(\x. sum (\a. f a x) s) differentiable net" proof - from bchoice[OF assms(2)[unfolded differentiable_def]] show ?thesis by (auto intro!: has_derivative_sum simp: differentiable_def) qed lemma differentiable_minus [simp, derivative_intros]: "f differentiable F \ (\x. - f x) differentiable F" unfolding differentiable_def by (blast intro: has_derivative_minus) lemma differentiable_diff [simp, derivative_intros]: "f differentiable F \ g differentiable F \ (\x. f x - g x) differentiable F" unfolding differentiable_def by (blast intro: has_derivative_diff) lemma differentiable_mult [simp, derivative_intros]: fixes f g :: "'a::real_normed_vector \ 'b::real_normed_algebra" shows "f differentiable (at x within s) \ g differentiable (at x within s) \ (\x. f x * g x) differentiable (at x within s)" unfolding differentiable_def by (blast intro: has_derivative_mult) lemma differentiable_cmult_left_iff [simp]: fixes c::"'a::real_normed_field" shows "(\t. c * q t) differentiable at t \ c = 0 \ (\t. q t) differentiable at t" (is "?lhs = ?rhs") proof assume L: ?lhs {assume "c \ 0" then have "q differentiable at t" using differentiable_mult [OF differentiable_const L, of concl: "1/c"] by auto } then show ?rhs by auto qed auto lemma differentiable_cmult_right_iff [simp]: fixes c::"'a::real_normed_field" shows "(\t. q t * c) differentiable at t \ c = 0 \ (\t. q t) differentiable at t" (is "?lhs = ?rhs") by (simp add: mult.commute flip: differentiable_cmult_left_iff) lemma differentiable_inverse [simp, derivative_intros]: fixes f :: "'a::real_normed_vector \ 'b::real_normed_field" shows "f differentiable (at x within s) \ f x \ 0 \ (\x. inverse (f x)) differentiable (at x within s)" unfolding differentiable_def by (blast intro: has_derivative_inverse) lemma differentiable_divide [simp, derivative_intros]: fixes f g :: "'a::real_normed_vector \ 'b::real_normed_field" shows "f differentiable (at x within s) \ g differentiable (at x within s) \ g x \ 0 \ (\x. f x / g x) differentiable (at x within s)" unfolding divide_inverse by simp lemma differentiable_power [simp, derivative_intros]: fixes f g :: "'a::real_normed_vector \ 'b::real_normed_field" shows "f differentiable (at x within s) \ (\x. f x ^ n) differentiable (at x within s)" unfolding differentiable_def by (blast intro: has_derivative_power) lemma differentiable_power_int [simp, derivative_intros]: fixes f :: "'a::real_normed_vector \ 'b::real_normed_field" shows "f differentiable (at x within s) \ f x \ 0 \ (\x. power_int (f x) n) differentiable (at x within s)" unfolding differentiable_def by (blast intro: has_derivative_power_int) lemma differentiable_scaleR [simp, derivative_intros]: "f differentiable (at x within s) \ g differentiable (at x within s) \ (\x. f x *\<^sub>R g x) differentiable (at x within s)" unfolding differentiable_def by (blast intro: has_derivative_scaleR) lemma has_derivative_imp_has_field_derivative: "(f has_derivative D) F \ (\x. x * D' = D x) \ (f has_field_derivative D') F" unfolding has_field_derivative_def by (rule has_derivative_eq_rhs[of f D]) (simp_all add: fun_eq_iff mult.commute) lemma has_field_derivative_imp_has_derivative: "(f has_field_derivative D) F \ (f has_derivative (*) D) F" by (simp add: has_field_derivative_def) lemma DERIV_subset: "(f has_field_derivative f') (at x within s) \ t \ s \ (f has_field_derivative f') (at x within t)" by (simp add: has_field_derivative_def has_derivative_subset) lemma has_field_derivative_at_within: "(f has_field_derivative f') (at x) \ (f has_field_derivative f') (at x within s)" using DERIV_subset by blast abbreviation (input) DERIV :: "('a::real_normed_field \ 'a) \ 'a \ 'a \ bool" ("(DERIV (_)/ (_)/ :> (_))" [1000, 1000, 60] 60) where "DERIV f x :> D \ (f has_field_derivative D) (at x)" abbreviation has_real_derivative :: "(real \ real) \ real \ real filter \ bool" (infix "(has'_real'_derivative)" 50) where "(f has_real_derivative D) F \ (f has_field_derivative D) F" lemma real_differentiable_def: "f differentiable at x within s \ (\D. (f has_real_derivative D) (at x within s))" proof safe assume "f differentiable at x within s" then obtain f' where *: "(f has_derivative f') (at x within s)" unfolding differentiable_def by auto then obtain c where "f' = ((*) c)" by (metis real_bounded_linear has_derivative_bounded_linear mult.commute fun_eq_iff) with * show "\D. (f has_real_derivative D) (at x within s)" unfolding has_field_derivative_def by auto qed (auto simp: differentiable_def has_field_derivative_def) lemma real_differentiableE [elim?]: assumes f: "f differentiable (at x within s)" obtains df where "(f has_real_derivative df) (at x within s)" using assms by (auto simp: real_differentiable_def) lemma has_field_derivative_iff: "(f has_field_derivative D) (at x within S) \ ((\y. (f y - f x) / (y - x)) \ D) (at x within S)" proof - have "((\y. norm (f y - f x - D * (y - x)) / norm (y - x)) \ 0) (at x within S) = ((\y. (f y - f x) / (y - x) - D) \ 0) (at x within S)" by (smt (verit, best) Lim_cong_within divide_diff_eq_iff norm_divide right_minus_eq tendsto_norm_zero_iff) then show ?thesis by (simp add: has_field_derivative_def has_derivative_iff_norm bounded_linear_mult_right LIM_zero_iff) qed lemma DERIV_def: "DERIV f x :> D \ (\h. (f (x + h) - f x) / h) \0\ D" unfolding field_has_derivative_at has_field_derivative_def has_field_derivative_iff .. text \due to Christian Pardillo Laursen, replacing a proper epsilon-delta horror\ lemma field_derivative_lim_unique: assumes f: "(f has_field_derivative df) (at z)" and s: "s \ 0" "\n. s n \ 0" and a: "(\n. (f (z + s n) - f z) / s n) \ a" shows "df = a" proof - have "((\k. (f (z + k) - f z) / k) \ df) (at 0)" using f by (simp add: DERIV_def) with s have "((\n. (f (z + s n) - f z) / s n) \ df)" by (simp flip: LIMSEQ_SEQ_conv) then show ?thesis using a by (rule LIMSEQ_unique) qed lemma mult_commute_abs: "(\x. x * c) = (*) c" for c :: "'a::ab_semigroup_mult" by (simp add: fun_eq_iff mult.commute) lemma DERIV_compose_FDERIV: fixes f::"real\real" assumes "DERIV f (g x) :> f'" assumes "(g has_derivative g') (at x within s)" shows "((\x. f (g x)) has_derivative (\x. g' x * f')) (at x within s)" using assms has_derivative_compose[of g g' x s f "(*) f'"] by (auto simp: has_field_derivative_def ac_simps) subsection \Vector derivative\ text \It's for real derivatives only, and not obviously generalisable to field derivatives\ lemma has_real_derivative_iff_has_vector_derivative: "(f has_real_derivative y) F \ (f has_vector_derivative y) F" unfolding has_vector_derivative_def has_field_derivative_def real_scaleR_def mult_commute_abs .. lemma has_field_derivative_subset: "(f has_field_derivative y) (at x within s) \ t \ s \ (f has_field_derivative y) (at x within t)" by (fact DERIV_subset) lemma has_vector_derivative_const[simp, derivative_intros]: "((\x. c) has_vector_derivative 0) net" by (auto simp: has_vector_derivative_def) lemma has_vector_derivative_id[simp, derivative_intros]: "((\x. x) has_vector_derivative 1) net" by (auto simp: has_vector_derivative_def) lemma has_vector_derivative_minus[derivative_intros]: "(f has_vector_derivative f') net \ ((\x. - f x) has_vector_derivative (- f')) net" by (auto simp: has_vector_derivative_def) lemma has_vector_derivative_add[derivative_intros]: "(f has_vector_derivative f') net \ (g has_vector_derivative g') net \ ((\x. f x + g x) has_vector_derivative (f' + g')) net" by (auto simp: has_vector_derivative_def scaleR_right_distrib) lemma has_vector_derivative_sum[derivative_intros]: "(\i. i \ I \ (f i has_vector_derivative f' i) net) \ ((\x. \i\I. f i x) has_vector_derivative (\i\I. f' i)) net" by (auto simp: has_vector_derivative_def fun_eq_iff scaleR_sum_right intro!: derivative_eq_intros) lemma has_vector_derivative_diff[derivative_intros]: "(f has_vector_derivative f') net \ (g has_vector_derivative g') net \ ((\x. f x - g x) has_vector_derivative (f' - g')) net" by (auto simp: has_vector_derivative_def scaleR_diff_right) lemma has_vector_derivative_add_const: "((\t. g t + z) has_vector_derivative f') net = ((\t. g t) has_vector_derivative f') net" apply (intro iffI) apply (force dest: has_vector_derivative_diff [where g = "\t. z", OF _ has_vector_derivative_const]) apply (force dest: has_vector_derivative_add [OF _ has_vector_derivative_const]) done lemma has_vector_derivative_diff_const: "((\t. g t - z) has_vector_derivative f') net = ((\t. g t) has_vector_derivative f') net" using has_vector_derivative_add_const [where z = "-z"] by simp lemma (in bounded_linear) has_vector_derivative: assumes "(g has_vector_derivative g') F" shows "((\x. f (g x)) has_vector_derivative f g') F" using has_derivative[OF assms[unfolded has_vector_derivative_def]] by (simp add: has_vector_derivative_def scaleR) lemma (in bounded_bilinear) has_vector_derivative: assumes "(f has_vector_derivative f') (at x within s)" and "(g has_vector_derivative g') (at x within s)" shows "((\x. f x ** g x) has_vector_derivative (f x ** g' + f' ** g x)) (at x within s)" using FDERIV[OF assms(1-2)[unfolded has_vector_derivative_def]] by (simp add: has_vector_derivative_def scaleR_right scaleR_left scaleR_right_distrib) lemma has_vector_derivative_scaleR[derivative_intros]: "(f has_field_derivative f') (at x within s) \ (g has_vector_derivative g') (at x within s) \ ((\x. f x *\<^sub>R g x) has_vector_derivative (f x *\<^sub>R g' + f' *\<^sub>R g x)) (at x within s)" unfolding has_real_derivative_iff_has_vector_derivative by (rule bounded_bilinear.has_vector_derivative[OF bounded_bilinear_scaleR]) lemma has_vector_derivative_mult[derivative_intros]: "(f has_vector_derivative f') (at x within s) \ (g has_vector_derivative g') (at x within s) \ ((\x. f x * g x) has_vector_derivative (f x * g' + f' * g x)) (at x within s)" for f g :: "real \ 'a::real_normed_algebra" by (rule bounded_bilinear.has_vector_derivative[OF bounded_bilinear_mult]) lemma has_vector_derivative_of_real[derivative_intros]: "(f has_field_derivative D) F \ ((\x. of_real (f x)) has_vector_derivative (of_real D)) F" by (rule bounded_linear.has_vector_derivative[OF bounded_linear_of_real]) (simp add: has_real_derivative_iff_has_vector_derivative) lemma has_vector_derivative_real_field: "(f has_field_derivative f') (at (of_real a)) \ ((\x. f (of_real x)) has_vector_derivative f') (at a within s)" using has_derivative_compose[of of_real of_real a _ f "(*) f'"] by (simp add: scaleR_conv_of_real ac_simps has_vector_derivative_def has_field_derivative_def) lemma has_vector_derivative_continuous: "(f has_vector_derivative D) (at x within s) \ continuous (at x within s) f" by (auto intro: has_derivative_continuous simp: has_vector_derivative_def) lemma continuous_on_vector_derivative: "(\x. x \ S \ (f has_vector_derivative f' x) (at x within S)) \ continuous_on S f" by (auto simp: continuous_on_eq_continuous_within intro!: has_vector_derivative_continuous) lemma has_vector_derivative_mult_right[derivative_intros]: fixes a :: "'a::real_normed_algebra" shows "(f has_vector_derivative x) F \ ((\x. a * f x) has_vector_derivative (a * x)) F" by (rule bounded_linear.has_vector_derivative[OF bounded_linear_mult_right]) lemma has_vector_derivative_mult_left[derivative_intros]: fixes a :: "'a::real_normed_algebra" shows "(f has_vector_derivative x) F \ ((\x. f x * a) has_vector_derivative (x * a)) F" by (rule bounded_linear.has_vector_derivative[OF bounded_linear_mult_left]) lemma has_vector_derivative_divide[derivative_intros]: fixes a :: "'a::real_normed_field" shows "(f has_vector_derivative x) F \ ((\x. f x / a) has_vector_derivative (x / a)) F" using has_vector_derivative_mult_left [of f x F "inverse a"] by (simp add: field_class.field_divide_inverse) subsection \Derivatives\ lemma DERIV_D: "DERIV f x :> D \ (\h. (f (x + h) - f x) / h) \0\ D" by (simp add: DERIV_def) lemma has_field_derivativeD: "(f has_field_derivative D) (at x within S) \ ((\y. (f y - f x) / (y - x)) \ D) (at x within S)" by (simp add: has_field_derivative_iff) lemma DERIV_const [simp, derivative_intros]: "((\x. k) has_field_derivative 0) F" by (rule has_derivative_imp_has_field_derivative[OF has_derivative_const]) auto lemma DERIV_ident [simp, derivative_intros]: "((\x. x) has_field_derivative 1) F" by (rule has_derivative_imp_has_field_derivative[OF has_derivative_ident]) auto lemma field_differentiable_add[derivative_intros]: "(f has_field_derivative f') F \ (g has_field_derivative g') F \ ((\z. f z + g z) has_field_derivative f' + g') F" by (rule has_derivative_imp_has_field_derivative[OF has_derivative_add]) (auto simp: has_field_derivative_def field_simps mult_commute_abs) corollary DERIV_add: "(f has_field_derivative D) (at x within s) \ (g has_field_derivative E) (at x within s) \ ((\x. f x + g x) has_field_derivative D + E) (at x within s)" by (rule field_differentiable_add) lemma field_differentiable_minus[derivative_intros]: "(f has_field_derivative f') F \ ((\z. - (f z)) has_field_derivative -f') F" by (rule has_derivative_imp_has_field_derivative[OF has_derivative_minus]) (auto simp: has_field_derivative_def field_simps mult_commute_abs) corollary DERIV_minus: "(f has_field_derivative D) (at x within s) \ ((\x. - f x) has_field_derivative -D) (at x within s)" by (rule field_differentiable_minus) lemma field_differentiable_diff[derivative_intros]: "(f has_field_derivative f') F \ (g has_field_derivative g') F \ ((\z. f z - g z) has_field_derivative f' - g') F" by (simp only: diff_conv_add_uminus field_differentiable_add field_differentiable_minus) corollary DERIV_diff: "(f has_field_derivative D) (at x within s) \ (g has_field_derivative E) (at x within s) \ ((\x. f x - g x) has_field_derivative D - E) (at x within s)" by (rule field_differentiable_diff) lemma DERIV_continuous: "(f has_field_derivative D) (at x within s) \ continuous (at x within s) f" by (drule has_derivative_continuous[OF has_field_derivative_imp_has_derivative]) simp corollary DERIV_isCont: "DERIV f x :> D \ isCont f x" by (rule DERIV_continuous) lemma DERIV_atLeastAtMost_imp_continuous_on: assumes "\x. \a \ x; x \ b\ \ \y. DERIV f x :> y" shows "continuous_on {a..b} f" by (meson DERIV_isCont assms atLeastAtMost_iff continuous_at_imp_continuous_at_within continuous_on_eq_continuous_within) lemma DERIV_continuous_on: "(\x. x \ s \ (f has_field_derivative (D x)) (at x within s)) \ continuous_on s f" unfolding continuous_on_eq_continuous_within by (intro continuous_at_imp_continuous_on ballI DERIV_continuous) lemma DERIV_mult': "(f has_field_derivative D) (at x within s) \ (g has_field_derivative E) (at x within s) \ ((\x. f x * g x) has_field_derivative f x * E + D * g x) (at x within s)" by (rule has_derivative_imp_has_field_derivative[OF has_derivative_mult]) (auto simp: field_simps mult_commute_abs dest: has_field_derivative_imp_has_derivative) lemma DERIV_mult[derivative_intros]: "(f has_field_derivative Da) (at x within s) \ (g has_field_derivative Db) (at x within s) \ ((\x. f x * g x) has_field_derivative Da * g x + Db * f x) (at x within s)" by (rule has_derivative_imp_has_field_derivative[OF has_derivative_mult]) (auto simp: field_simps dest: has_field_derivative_imp_has_derivative) text \Derivative of linear multiplication\ lemma DERIV_cmult: "(f has_field_derivative D) (at x within s) \ ((\x. c * f x) has_field_derivative c * D) (at x within s)" by (drule DERIV_mult' [OF DERIV_const]) simp lemma DERIV_cmult_right: "(f has_field_derivative D) (at x within s) \ ((\x. f x * c) has_field_derivative D * c) (at x within s)" using DERIV_cmult by (auto simp add: ac_simps) lemma DERIV_cmult_Id [simp]: "((*) c has_field_derivative c) (at x within s)" using DERIV_ident [THEN DERIV_cmult, where c = c and x = x] by simp lemma DERIV_cdivide: "(f has_field_derivative D) (at x within s) \ ((\x. f x / c) has_field_derivative D / c) (at x within s)" using DERIV_cmult_right[of f D x s "1 / c"] by simp lemma DERIV_unique: "DERIV f x :> D \ DERIV f x :> E \ D = E" unfolding DERIV_def by (rule LIM_unique) lemma DERIV_Uniq: "\\<^sub>\\<^sub>1D. DERIV f x :> D" by (simp add: DERIV_unique Uniq_def) lemma DERIV_sum[derivative_intros]: "(\ n. n \ S \ ((\x. f x n) has_field_derivative (f' x n)) F) \ ((\x. sum (f x) S) has_field_derivative sum (f' x) S) F" by (rule has_derivative_imp_has_field_derivative [OF has_derivative_sum]) (auto simp: sum_distrib_left mult_commute_abs dest: has_field_derivative_imp_has_derivative) lemma DERIV_inverse'[derivative_intros]: assumes "(f has_field_derivative D) (at x within s)" and "f x \ 0" shows "((\x. inverse (f x)) has_field_derivative - (inverse (f x) * D * inverse (f x))) (at x within s)" proof - have "(f has_derivative (\x. x * D)) = (f has_derivative (*) D)" by (rule arg_cong [of "\x. x * D"]) (simp add: fun_eq_iff) with assms have "(f has_derivative (\x. x * D)) (at x within s)" by (auto dest!: has_field_derivative_imp_has_derivative) then show ?thesis using \f x \ 0\ by (auto intro: has_derivative_imp_has_field_derivative has_derivative_inverse) qed text \Power of \-1\\ lemma DERIV_inverse: "x \ 0 \ ((\x. inverse(x)) has_field_derivative - (inverse x ^ Suc (Suc 0))) (at x within s)" by (drule DERIV_inverse' [OF DERIV_ident]) simp text \Derivative of inverse\ lemma DERIV_inverse_fun: "(f has_field_derivative d) (at x within s) \ f x \ 0 \ ((\x. inverse (f x)) has_field_derivative (- (d * inverse(f x ^ Suc (Suc 0))))) (at x within s)" by (drule (1) DERIV_inverse') (simp add: ac_simps nonzero_inverse_mult_distrib) text \Derivative of quotient\ lemma DERIV_divide[derivative_intros]: "(f has_field_derivative D) (at x within s) \ (g has_field_derivative E) (at x within s) \ g x \ 0 \ ((\x. f x / g x) has_field_derivative (D * g x - f x * E) / (g x * g x)) (at x within s)" by (rule has_derivative_imp_has_field_derivative[OF has_derivative_divide]) (auto dest: has_field_derivative_imp_has_derivative simp: field_simps) lemma DERIV_quotient: "(f has_field_derivative d) (at x within s) \ (g has_field_derivative e) (at x within s)\ g x \ 0 \ ((\y. f y / g y) has_field_derivative (d * g x - (e * f x)) / (g x ^ Suc (Suc 0))) (at x within s)" by (drule (2) DERIV_divide) (simp add: mult.commute) lemma DERIV_power_Suc: "(f has_field_derivative D) (at x within s) \ ((\x. f x ^ Suc n) has_field_derivative (1 + of_nat n) * (D * f x ^ n)) (at x within s)" by (rule has_derivative_imp_has_field_derivative[OF has_derivative_power]) (auto simp: has_field_derivative_def) lemma DERIV_power[derivative_intros]: "(f has_field_derivative D) (at x within s) \ ((\x. f x ^ n) has_field_derivative of_nat n * (D * f x ^ (n - Suc 0))) (at x within s)" by (rule has_derivative_imp_has_field_derivative[OF has_derivative_power]) (auto simp: has_field_derivative_def) lemma DERIV_pow: "((\x. x ^ n) has_field_derivative real n * (x ^ (n - Suc 0))) (at x within s)" using DERIV_power [OF DERIV_ident] by simp lemma DERIV_power_int [derivative_intros]: assumes [derivative_intros]: "(f has_field_derivative d) (at x within s)" and [simp]: "f x \ 0" shows "((\x. power_int (f x) n) has_field_derivative (of_int n * power_int (f x) (n - 1) * d)) (at x within s)" proof (cases n rule: int_cases4) case (nonneg n) thus ?thesis by (cases "n = 0") (auto intro!: derivative_eq_intros simp: field_simps power_int_diff simp flip: power_Suc power_Suc2 power_add) next case (neg n) thus ?thesis by (auto intro!: derivative_eq_intros simp: field_simps power_int_diff power_int_minus simp flip: power_Suc power_Suc2 power_add) qed lemma DERIV_chain': "(f has_field_derivative D) (at x within s) \ DERIV g (f x) :> E \ ((\x. g (f x)) has_field_derivative E * D) (at x within s)" using has_derivative_compose[of f "(*) D" x s g "(*) E"] by (simp only: has_field_derivative_def mult_commute_abs ac_simps) corollary DERIV_chain2: "DERIV f (g x) :> Da \ (g has_field_derivative Db) (at x within s) \ ((\x. f (g x)) has_field_derivative Da * Db) (at x within s)" by (rule DERIV_chain') text \Standard version\ lemma DERIV_chain: "DERIV f (g x) :> Da \ (g has_field_derivative Db) (at x within s) \ (f \ g has_field_derivative Da * Db) (at x within s)" by (drule (1) DERIV_chain', simp add: o_def mult.commute) lemma DERIV_image_chain: "(f has_field_derivative Da) (at (g x) within (g ` s)) \ (g has_field_derivative Db) (at x within s) \ (f \ g has_field_derivative Da * Db) (at x within s)" using has_derivative_in_compose [of g "(*) Db" x s f "(*) Da "] by (simp add: has_field_derivative_def o_def mult_commute_abs ac_simps) (*These two are from HOL Light: HAS_COMPLEX_DERIVATIVE_CHAIN*) lemma DERIV_chain_s: assumes "(\x. x \ s \ DERIV g x :> g'(x))" and "DERIV f x :> f'" and "f x \ s" shows "DERIV (\x. g(f x)) x :> f' * g'(f x)" by (metis (full_types) DERIV_chain' mult.commute assms) lemma DERIV_chain3: (*HAS_COMPLEX_DERIVATIVE_CHAIN_UNIV*) assumes "(\x. DERIV g x :> g'(x))" and "DERIV f x :> f'" shows "DERIV (\x. g(f x)) x :> f' * g'(f x)" by (metis UNIV_I DERIV_chain_s [of UNIV] assms) text \Alternative definition for differentiability\ lemma DERIV_LIM_iff: fixes f :: "'a::{real_normed_vector,inverse} \ 'a" shows "((\h. (f (a + h) - f a) / h) \0\ D) = ((\x. (f x - f a) / (x - a)) \a\ D)" (is "?lhs = ?rhs") proof assume ?lhs then have "(\x. (f (a + (x + - a)) - f a) / (x + - a)) \0 - - a\ D" by (rule LIM_offset) then show ?rhs by simp next assume ?rhs then have "(\x. (f (x+a) - f a) / ((x+a) - a)) \a-a\ D" by (rule LIM_offset) then show ?lhs by (simp add: add.commute) qed lemma has_field_derivative_cong_ev: assumes "x = y" and *: "eventually (\x. x \ S \ f x = g x) (nhds x)" and "u = v" "S = t" "x \ S" shows "(f has_field_derivative u) (at x within S) = (g has_field_derivative v) (at y within t)" unfolding has_field_derivative_iff proof (rule filterlim_cong) from assms have "f y = g y" by (auto simp: eventually_nhds) with * show "\\<^sub>F z in at x within S. (f z - f x) / (z - x) = (g z - g y) / (z - y)" unfolding eventually_at_filter by eventually_elim (auto simp: assms \f y = g y\) qed (simp_all add: assms) lemma has_field_derivative_cong_eventually: assumes "eventually (\x. f x = g x) (at x within S)" "f x = g x" shows "(f has_field_derivative u) (at x within S) = (g has_field_derivative u) (at x within S)" unfolding has_field_derivative_iff proof (rule tendsto_cong) show "\\<^sub>F y in at x within S. (f y - f x) / (y - x) = (g y - g x) / (y - x)" using assms by (auto elim: eventually_mono) qed lemma DERIV_cong_ev: "x = y \ eventually (\x. f x = g x) (nhds x) \ u = v \ DERIV f x :> u \ DERIV g y :> v" by (rule has_field_derivative_cong_ev) simp_all lemma DERIV_mirror: "(DERIV f (- x) :> y) \ (DERIV (\x. f (- x)) x :> - y)" for f :: "real \ real" and x y :: real by (simp add: DERIV_def filterlim_at_split filterlim_at_left_to_right tendsto_minus_cancel_left field_simps conj_commute) lemma DERIV_shift: "(f has_field_derivative y) (at (x + z)) = ((\x. f (x + z)) has_field_derivative y) (at x)" by (simp add: DERIV_def field_simps) lemma DERIV_at_within_shift_lemma: assumes "(f has_field_derivative y) (at (z+x) within (+) z ` S)" shows "(f \ (+)z has_field_derivative y) (at x within S)" proof - have "((+)z has_field_derivative 1) (at x within S)" by (rule derivative_eq_intros | simp)+ with assms DERIV_image_chain show ?thesis by (metis mult.right_neutral) qed lemma DERIV_at_within_shift: "(f has_field_derivative y) (at (z+x) within (+) z ` S) \ ((\x. f (z+x)) has_field_derivative y) (at x within S)" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs using DERIV_at_within_shift_lemma unfolding o_def by blast next have [simp]: "(\x. x - z) ` (+) z ` S = S" by force assume R: ?rhs have "(f \ (+) z \ (+) (- z) has_field_derivative y) (at (z + x) within (+) z ` S)" by (rule DERIV_at_within_shift_lemma) (use R in \simp add: o_def\) then show ?lhs by (simp add: o_def) qed lemma floor_has_real_derivative: fixes f :: "real \ 'a::{floor_ceiling,order_topology}" assumes "isCont f x" and "f x \ \" shows "((\x. floor (f x)) has_real_derivative 0) (at x)" proof (subst DERIV_cong_ev[OF refl _ refl]) show "((\_. floor (f x)) has_real_derivative 0) (at x)" by simp have "\\<^sub>F y in at x. \f y\ = \f x\" by (rule eventually_floor_eq[OF assms[unfolded continuous_at]]) then show "\\<^sub>F y in nhds x. real_of_int \f y\ = real_of_int \f x\" unfolding eventually_at_filter by eventually_elim auto qed lemmas has_derivative_floor[derivative_intros] = floor_has_real_derivative[THEN DERIV_compose_FDERIV] lemma continuous_floor: fixes x::real shows "x \ \ \ continuous (at x) (real_of_int \ floor)" using floor_has_real_derivative [where f=id] by (auto simp: o_def has_field_derivative_def intro: has_derivative_continuous) lemma continuous_frac: fixes x::real assumes "x \ \" shows "continuous (at x) frac" proof - have "isCont (\x. real_of_int \x\) x" using continuous_floor [OF assms] by (simp add: o_def) then have *: "continuous (at x) (\x. x - real_of_int \x\)" by (intro continuous_intros) moreover have "\\<^sub>F x in nhds x. frac x = x - real_of_int \x\" by (simp add: frac_def) ultimately show ?thesis by (simp add: LIM_imp_LIM frac_def isCont_def) qed text \Caratheodory formulation of derivative at a point\ lemma CARAT_DERIV: "(DERIV f x :> l) \ (\g. (\z. f z - f x = g z * (z - x)) \ isCont g x \ g x = l)" (is "?lhs = ?rhs") proof assume ?lhs show "\g. (\z. f z - f x = g z * (z - x)) \ isCont g x \ g x = l" proof (intro exI conjI) let ?g = "(\z. if z = x then l else (f z - f x) / (z-x))" show "\z. f z - f x = ?g z * (z - x)" by simp show "isCont ?g x" using \?lhs\ by (simp add: isCont_iff DERIV_def cong: LIM_equal [rule_format]) show "?g x = l" by simp qed next assume ?rhs then show ?lhs by (auto simp add: isCont_iff DERIV_def cong: LIM_cong) qed subsection \Local extrema\ text \If \<^term>\0 < f' x\ then \<^term>\x\ is Locally Strictly Increasing At The Right.\ lemma has_real_derivative_pos_inc_right: fixes f :: "real \ real" assumes der: "(f has_real_derivative l) (at x within S)" and l: "0 < l" shows "\d > 0. \h > 0. x + h \ S \ h < d \ f x < f (x + h)" using assms proof - from der [THEN has_field_derivativeD, THEN tendstoD, OF l, unfolded eventually_at] obtain s where s: "0 < s" and all: "\xa. xa\S \ xa \ x \ dist xa x < s \ \(f xa - f x) / (xa - x) - l\ < l" by (auto simp: dist_real_def) then show ?thesis proof (intro exI conjI strip) show "0 < s" by (rule s) next fix h :: real assume "0 < h" "h < s" "x + h \ S" with all [of "x + h"] show "f x < f (x+h)" proof (simp add: abs_if dist_real_def pos_less_divide_eq split: if_split_asm) assume "\ (f (x + h) - f x) / h < l" and h: "0 < h" with l have "0 < (f (x + h) - f x) / h" by arith then show "f x < f (x + h)" by (simp add: pos_less_divide_eq h) qed qed qed lemma DERIV_pos_inc_right: fixes f :: "real \ real" assumes der: "DERIV f x :> l" and l: "0 < l" shows "\d > 0. \h > 0. h < d \ f x < f (x + h)" using has_real_derivative_pos_inc_right[OF assms] by auto lemma has_real_derivative_neg_dec_left: fixes f :: "real \ real" assumes der: "(f has_real_derivative l) (at x within S)" and "l < 0" shows "\d > 0. \h > 0. x - h \ S \ h < d \ f x < f (x - h)" proof - from \l < 0\ have l: "- l > 0" by simp from der [THEN has_field_derivativeD, THEN tendstoD, OF l, unfolded eventually_at] obtain s where s: "0 < s" and all: "\xa. xa\S \ xa \ x \ dist xa x < s \ \(f xa - f x) / (xa - x) - l\ < - l" by (auto simp: dist_real_def) then show ?thesis proof (intro exI conjI strip) show "0 < s" by (rule s) next fix h :: real assume "0 < h" "h < s" "x - h \ S" with all [of "x - h"] show "f x < f (x-h)" proof (simp add: abs_if pos_less_divide_eq dist_real_def split: if_split_asm) assume "- ((f (x-h) - f x) / h) < l" and h: "0 < h" with l have "0 < (f (x-h) - f x) / h" by arith then show "f x < f (x - h)" by (simp add: pos_less_divide_eq h) qed qed qed lemma DERIV_neg_dec_left: fixes f :: "real \ real" assumes der: "DERIV f x :> l" and l: "l < 0" shows "\d > 0. \h > 0. h < d \ f x < f (x - h)" using has_real_derivative_neg_dec_left[OF assms] by auto lemma has_real_derivative_pos_inc_left: fixes f :: "real \ real" shows "(f has_real_derivative l) (at x within S) \ 0 < l \ \d>0. \h>0. x - h \ S \ h < d \ f (x - h) < f x" by (rule has_real_derivative_neg_dec_left [of "\x. - f x" "-l" x S, simplified]) (auto simp add: DERIV_minus) lemma DERIV_pos_inc_left: fixes f :: "real \ real" shows "DERIV f x :> l \ 0 < l \ \d > 0. \h > 0. h < d \ f (x - h) < f x" using has_real_derivative_pos_inc_left by blast lemma has_real_derivative_neg_dec_right: fixes f :: "real \ real" shows "(f has_real_derivative l) (at x within S) \ l < 0 \ \d > 0. \h > 0. x + h \ S \ h < d \ f x > f (x + h)" by (rule has_real_derivative_pos_inc_right [of "\x. - f x" "-l" x S, simplified]) (auto simp add: DERIV_minus) lemma DERIV_neg_dec_right: fixes f :: "real \ real" shows "DERIV f x :> l \ l < 0 \ \d > 0. \h > 0. h < d \ f x > f (x + h)" using has_real_derivative_neg_dec_right by blast lemma DERIV_local_max: fixes f :: "real \ real" assumes der: "DERIV f x :> l" and d: "0 < d" and le: "\y. \x - y\ < d \ f y \ f x" shows "l = 0" proof (cases rule: linorder_cases [of l 0]) case equal then show ?thesis . next case less from DERIV_neg_dec_left [OF der less] obtain d' where d': "0 < d'" and lt: "\h > 0. h < d' \ f x < f (x - h)" by blast obtain e where "0 < e \ e < d \ e < d'" using field_lbound_gt_zero [OF d d'] .. with lt le [THEN spec [where x="x - e"]] show ?thesis by (auto simp add: abs_if) next case greater from DERIV_pos_inc_right [OF der greater] obtain d' where d': "0 < d'" and lt: "\h > 0. h < d' \ f x < f (x + h)" by blast obtain e where "0 < e \ e < d \ e < d'" using field_lbound_gt_zero [OF d d'] .. with lt le [THEN spec [where x="x + e"]] show ?thesis by (auto simp add: abs_if) qed text \Similar theorem for a local minimum\ lemma DERIV_local_min: fixes f :: "real \ real" shows "DERIV f x :> l \ 0 < d \ \y. \x - y\ < d \ f x \ f y \ l = 0" by (drule DERIV_minus [THEN DERIV_local_max]) auto text\In particular, if a function is locally flat\ lemma DERIV_local_const: fixes f :: "real \ real" shows "DERIV f x :> l \ 0 < d \ \y. \x - y\ < d \ f x = f y \ l = 0" by (auto dest!: DERIV_local_max) subsection \Rolle's Theorem\ text \Lemma about introducing open ball in open interval\ lemma lemma_interval_lt: fixes a b x :: real assumes "a < x" "x < b" shows "\d. 0 < d \ (\y. \x - y\ < d \ a < y \ y < b)" using linorder_linear [of "x - a" "b - x"] proof assume "x - a \ b - x" with assms show ?thesis by (rule_tac x = "x - a" in exI) auto next assume "b - x \ x - a" with assms show ?thesis by (rule_tac x = "b - x" in exI) auto qed lemma lemma_interval: "a < x \ x < b \ \d. 0 < d \ (\y. \x - y\ < d \ a \ y \ y \ b)" for a b x :: real by (force dest: lemma_interval_lt) text \Rolle's Theorem. If \<^term>\f\ is defined and continuous on the closed interval \[a,b]\ and differentiable on the open interval \(a,b)\, and \<^term>\f a = f b\, then there exists \x0 \ (a,b)\ such that \<^term>\f' x0 = 0\\ theorem Rolle_deriv: fixes f :: "real \ real" assumes "a < b" and fab: "f a = f b" and contf: "continuous_on {a..b} f" and derf: "\x. \a < x; x < b\ \ (f has_derivative f' x) (at x)" shows "\z. a < z \ z < b \ f' z = (\v. 0)" proof - have le: "a \ b" using \a < b\ by simp have "(a + b) / 2 \ {a..b}" using assms(1) by auto then have *: "{a..b} \ {}" by auto obtain x where x_max: "\z. a \ z \ z \ b \ f z \ f x" and "a \ x" "x \ b" using continuous_attains_sup[OF compact_Icc * contf] by (meson atLeastAtMost_iff) obtain x' where x'_min: "\z. a \ z \ z \ b \ f x' \ f z" and "a \ x'" "x' \ b" using continuous_attains_inf[OF compact_Icc * contf] by (meson atLeastAtMost_iff) consider "a < x" "x < b" | "x = a \ x = b" using \a \ x\ \x \ b\ by arith then show ?thesis proof cases case 1 \ \\<^term>\f\ attains its maximum within the interval\ then obtain l where der: "DERIV f x :> l" using derf differentiable_def real_differentiable_def by blast obtain d where d: "0 < d" and bound: "\y. \x - y\ < d \ a \ y \ y \ b" using lemma_interval [OF 1] by blast then have bound': "\y. \x - y\ < d \ f y \ f x" using x_max by blast \ \the derivative at a local maximum is zero\ have "l = 0" by (rule DERIV_local_max [OF der d bound']) with 1 der derf [of x] show ?thesis by (metis has_derivative_unique has_field_derivative_def mult_zero_left) next case 2 then have fx: "f b = f x" by (auto simp add: fab) consider "a < x'" "x' < b" | "x' = a \ x' = b" using \a \ x'\ \x' \ b\ by arith then show ?thesis proof cases case 1 \ \\<^term>\f\ attains its minimum within the interval\ then obtain l where der: "DERIV f x' :> l" using derf differentiable_def real_differentiable_def by blast from lemma_interval [OF 1] obtain d where d: "0y. \x'-y\ < d \ a \ y \ y \ b" by blast then have bound': "\y. \x' - y\ < d \ f x' \ f y" using x'_min by blast have "l = 0" by (rule DERIV_local_min [OF der d bound']) \ \the derivative at a local minimum is zero\ then show ?thesis using 1 der derf [of x'] by (metis has_derivative_unique has_field_derivative_def mult_zero_left) next case 2 \ \\<^term>\f\ is constant throughout the interval\ then have fx': "f b = f x'" by (auto simp: fab) from dense [OF \a < b\] obtain r where r: "a < r" "r < b" by blast obtain d where d: "0 < d" and bound: "\y. \r - y\ < d \ a \ y \ y \ b" using lemma_interval [OF r] by blast have eq_fb: "f z = f b" if "a \ z" and "z \ b" for z proof (rule order_antisym) show "f z \ f b" by (simp add: fx x_max that) show "f b \ f z" by (simp add: fx' x'_min that) qed have bound': "\y. \r - y\ < d \ f r = f y" proof (intro strip) fix y :: real assume lt: "\r - y\ < d" then have "f y = f b" by (simp add: eq_fb bound) then show "f r = f y" by (simp add: eq_fb r order_less_imp_le) qed obtain l where der: "DERIV f r :> l" using derf differentiable_def r(1) r(2) real_differentiable_def by blast have "l = 0" by (rule DERIV_local_const [OF der d bound']) \ \the derivative of a constant function is zero\ with r der derf [of r] show ?thesis by (metis has_derivative_unique has_field_derivative_def mult_zero_left) qed qed qed corollary Rolle: fixes a b :: real assumes ab: "a < b" "f a = f b" "continuous_on {a..b} f" and dif [rule_format]: "\x. \a < x; x < b\ \ f differentiable (at x)" shows "\z. a < z \ z < b \ DERIV f z :> 0" proof - obtain f' where f': "\x. \a < x; x < b\ \ (f has_derivative f' x) (at x)" using dif unfolding differentiable_def by metis then have "\z. a < z \ z < b \ f' z = (\v. 0)" by (metis Rolle_deriv [OF ab]) then show ?thesis using f' has_derivative_imp_has_field_derivative by fastforce qed subsection \Mean Value Theorem\ theorem mvt: fixes f :: "real \ real" assumes "a < b" and contf: "continuous_on {a..b} f" and derf: "\x. \a < x; x < b\ \ (f has_derivative f' x) (at x)" obtains \ where "a < \" "\ < b" "f b - f a = (f' \) (b - a)" proof - have "\\. a < \ \ \ < b \ (\y. f' \ y - (f b - f a) / (b - a) * y) = (\v. 0)" proof (intro Rolle_deriv[OF \a < b\]) fix x assume x: "a < x" "x < b" show "((\x. f x - (f b - f a) / (b - a) * x) has_derivative (\y. f' x y - (f b - f a) / (b - a) * y)) (at x)" by (intro derivative_intros derf[OF x]) qed (use assms in \auto intro!: continuous_intros simp: field_simps\) then show ?thesis by (smt (verit, ccfv_SIG) pos_le_divide_eq pos_less_divide_eq that) qed theorem MVT: fixes a b :: real assumes lt: "a < b" and contf: "continuous_on {a..b} f" and dif: "\x. \a < x; x < b\ \ f differentiable (at x)" shows "\l z. a < z \ z < b \ DERIV f z :> l \ f b - f a = (b - a) * l" proof - obtain f' :: "real \ real \ real" where derf: "\x. a < x \ x < b \ (f has_derivative f' x) (at x)" using dif unfolding differentiable_def by metis then obtain z where "a < z" "z < b" "f b - f a = (f' z) (b - a)" using mvt [OF lt contf] by blast then show ?thesis by (simp add: ac_simps) (metis derf dif has_derivative_unique has_field_derivative_imp_has_derivative real_differentiable_def) qed corollary MVT2: assumes "a < b" and der: "\x. \a \ x; x \ b\ \ DERIV f x :> f' x" shows "\z::real. a < z \ z < b \ (f b - f a = (b - a) * f' z)" proof - have "\l z. a < z \ z < b \ (f has_real_derivative l) (at z) \ f b - f a = (b - a) * l" proof (rule MVT [OF \a < b\]) show "continuous_on {a..b} f" by (meson DERIV_continuous atLeastAtMost_iff continuous_at_imp_continuous_on der) show "\x. \a < x; x < b\ \ f differentiable (at x)" using assms by (force dest: order_less_imp_le simp add: real_differentiable_def) qed with assms show ?thesis by (blast dest: DERIV_unique order_less_imp_le) qed lemma pos_deriv_imp_strict_mono: assumes "\x. (f has_real_derivative f' x) (at x)" assumes "\x. f' x > 0" shows "strict_mono f" proof (rule strict_monoI) fix x y :: real assume xy: "x < y" from assms and xy have "\z>x. z < y \ f y - f x = (y - x) * f' z" by (intro MVT2) (auto dest: connectedD_interval) then obtain z where z: "z > x" "z < y" "f y - f x = (y - x) * f' z" by blast note \f y - f x = (y - x) * f' z\ also have "(y - x) * f' z > 0" using xy assms by (intro mult_pos_pos) auto finally show "f x < f y" by simp qed proposition deriv_nonneg_imp_mono: assumes deriv: "\x. x \ {a..b} \ (g has_real_derivative g' x) (at x)" assumes nonneg: "\x. x \ {a..b} \ g' x \ 0" assumes ab: "a \ b" shows "g a \ g b" proof (cases "a < b") assume "a < b" from deriv have "\x. \x \ a; x \ b\ \ (g has_real_derivative g' x) (at x)" by simp with MVT2[OF \a < b\] and deriv obtain \ where \_ab: "\ > a" "\ < b" and g_ab: "g b - g a = (b - a) * g' \" by blast from \_ab ab nonneg have "(b - a) * g' \ \ 0" by simp with g_ab show ?thesis by simp qed (insert ab, simp) subsubsection \A function is constant if its derivative is 0 over an interval.\ lemma DERIV_isconst_end: fixes f :: "real \ real" assumes "a < b" and contf: "continuous_on {a..b} f" and 0: "\x. \a < x; x < b\ \ DERIV f x :> 0" shows "f b = f a" using MVT [OF \a < b\] "0" DERIV_unique contf real_differentiable_def by (fastforce simp: algebra_simps) lemma DERIV_isconst2: fixes f :: "real \ real" assumes "a < b" and contf: "continuous_on {a..b} f" and derf: "\x. \a < x; x < b\ \ DERIV f x :> 0" and "a \ x" "x \ b" shows "f x = f a" proof (cases "a < x") case True have *: "continuous_on {a..x} f" using \x \ b\ contf continuous_on_subset by fastforce show ?thesis by (rule DERIV_isconst_end [OF True *]) (use \x \ b\ derf in auto) qed (use \a \ x\ in auto) lemma DERIV_isconst3: fixes a b x y :: real assumes "a < b" and "x \ {a <..< b}" and "y \ {a <..< b}" and derivable: "\x. x \ {a <..< b} \ DERIV f x :> 0" shows "f x = f y" proof (cases "x = y") case False let ?a = "min x y" let ?b = "max x y" have *: "DERIV f z :> 0" if "?a \ z" "z \ ?b" for z proof - have "a < z" and "z < b" using that \x \ {a <..< b}\ and \y \ {a <..< b}\ by auto then have "z \ {a<.. 0" by (rule derivable) qed have isCont: "continuous_on {?a..?b} f" by (meson * DERIV_continuous_on atLeastAtMost_iff has_field_derivative_at_within) have DERIV: "\z. \?a < z; z < ?b\ \ DERIV f z :> 0" using * by auto have "?a < ?b" using \x \ y\ by auto from DERIV_isconst2[OF this isCont DERIV, of x] and DERIV_isconst2[OF this isCont DERIV, of y] show ?thesis by auto qed auto lemma DERIV_isconst_all: fixes f :: "real \ real" shows "\x. DERIV f x :> 0 \ f x = f y" apply (rule linorder_cases [of x y]) apply (metis DERIV_continuous DERIV_isconst_end continuous_at_imp_continuous_on)+ done lemma DERIV_const_ratio_const: fixes f :: "real \ real" assumes "a \ b" and df: "\x. DERIV f x :> k" shows "f b - f a = (b - a) * k" proof (cases a b rule: linorder_cases) case less show ?thesis using MVT [OF less] df by (metis DERIV_continuous DERIV_unique continuous_at_imp_continuous_on real_differentiable_def) next case greater have "f a - f b = (a - b) * k" using MVT [OF greater] df by (metis DERIV_continuous DERIV_unique continuous_at_imp_continuous_on real_differentiable_def) then show ?thesis by (simp add: algebra_simps) qed auto lemma DERIV_const_ratio_const2: fixes f :: "real \ real" assumes "a \ b" and df: "\x. DERIV f x :> k" shows "(f b - f a) / (b - a) = k" using DERIV_const_ratio_const [OF assms] \a \ b\ by auto lemma real_average_minus_first [simp]: "(a + b) / 2 - a = (b - a) / 2" for a b :: real by simp lemma real_average_minus_second [simp]: "(b + a) / 2 - a = (b - a) / 2" for a b :: real by simp text \Gallileo's "trick": average velocity = av. of end velocities.\ lemma DERIV_const_average: fixes v :: "real \ real" and a b :: real assumes neq: "a \ b" and der: "\x. DERIV v x :> k" shows "v ((a + b) / 2) = (v a + v b) / 2" proof (cases rule: linorder_cases [of a b]) case equal with neq show ?thesis by simp next case less have "(v b - v a) / (b - a) = k" by (rule DERIV_const_ratio_const2 [OF neq der]) then have "(b - a) * ((v b - v a) / (b - a)) = (b - a) * k" by simp moreover have "(v ((a + b) / 2) - v a) / ((a + b) / 2 - a) = k" by (rule DERIV_const_ratio_const2 [OF _ der]) (simp add: neq) ultimately show ?thesis using neq by force next case greater have "(v b - v a) / (b - a) = k" by (rule DERIV_const_ratio_const2 [OF neq der]) then have "(b - a) * ((v b - v a) / (b - a)) = (b - a) * k" by simp moreover have " (v ((b + a) / 2) - v a) / ((b + a) / 2 - a) = k" by (rule DERIV_const_ratio_const2 [OF _ der]) (simp add: neq) ultimately show ?thesis using neq by (force simp add: add.commute) qed subsubsection\A function with positive derivative is increasing\ text \A simple proof using the MVT, by Jeremy Avigad. And variants.\ lemma DERIV_pos_imp_increasing_open: fixes a b :: real and f :: "real \ real" assumes "a < b" and "\x. a < x \ x < b \ (\y. DERIV f x :> y \ y > 0)" and con: "continuous_on {a..b} f" shows "f a < f b" proof (rule ccontr) assume f: "\ ?thesis" have "\l z. a < z \ z < b \ DERIV f z :> l \ f b - f a = (b - a) * l" by (rule MVT) (use assms real_differentiable_def in \force+\) then obtain l z where z: "a < z" "z < b" "DERIV f z :> l" and "f b - f a = (b - a) * l" by auto with assms f have "\ l > 0" by (metis linorder_not_le mult_le_0_iff diff_le_0_iff_le) with assms z show False by (metis DERIV_unique) qed lemma DERIV_pos_imp_increasing: fixes a b :: real and f :: "real \ real" assumes "a < b" and der: "\x. \a \ x; x \ b\ \ \y. DERIV f x :> y \ y > 0" shows "f a < f b" by (metis less_le_not_le DERIV_atLeastAtMost_imp_continuous_on DERIV_pos_imp_increasing_open [OF \a < b\] der) lemma DERIV_nonneg_imp_nondecreasing: fixes a b :: real and f :: "real \ real" assumes "a \ b" and "\x. \a \ x; x \ b\ \ \y. DERIV f x :> y \ y \ 0" shows "f a \ f b" proof (rule ccontr, cases "a = b") assume "\ ?thesis" and "a = b" then show False by auto next assume *: "\ ?thesis" assume "a \ b" with \a \ b\ have "a < b" by linarith moreover have "continuous_on {a..b} f" by (meson DERIV_isCont assms(2) atLeastAtMost_iff continuous_at_imp_continuous_on) ultimately have "\l z. a < z \ z < b \ DERIV f z :> l \ f b - f a = (b - a) * l" using assms MVT [OF \a < b\, of f] real_differentiable_def less_eq_real_def by blast then obtain l z where lz: "a < z" "z < b" "DERIV f z :> l" and **: "f b - f a = (b - a) * l" by auto with * have "a < b" "f b < f a" by auto with ** have "\ l \ 0" by (auto simp add: not_le algebra_simps) (metis * add_le_cancel_right assms(1) less_eq_real_def mult_right_mono add_left_mono linear order_refl) with assms lz show False by (metis DERIV_unique order_less_imp_le) qed lemma DERIV_neg_imp_decreasing_open: fixes a b :: real and f :: "real \ real" assumes "a < b" and "\x. a < x \ x < b \ \y. DERIV f x :> y \ y < 0" and con: "continuous_on {a..b} f" shows "f a > f b" proof - have "(\x. -f x) a < (\x. -f x) b" proof (rule DERIV_pos_imp_increasing_open [of a b]) show "\x. \a < x; x < b\ \ \y. ((\x. - f x) has_real_derivative y) (at x) \ 0 < y" using assms by simp (metis field_differentiable_minus neg_0_less_iff_less) show "continuous_on {a..b} (\x. - f x)" using con continuous_on_minus by blast qed (use assms in auto) then show ?thesis by simp qed lemma DERIV_neg_imp_decreasing: fixes a b :: real and f :: "real \ real" assumes "a < b" and der: "\x. \a \ x; x \ b\ \ \y. DERIV f x :> y \ y < 0" shows "f a > f b" by (metis less_le_not_le DERIV_atLeastAtMost_imp_continuous_on DERIV_neg_imp_decreasing_open [OF \a < b\] der) lemma DERIV_nonpos_imp_nonincreasing: fixes a b :: real and f :: "real \ real" assumes "a \ b" and "\x. \a \ x; x \ b\ \ \y. DERIV f x :> y \ y \ 0" shows "f a \ f b" proof - have "(\x. -f x) a \ (\x. -f x) b" using DERIV_nonneg_imp_nondecreasing [of a b "\x. -f x"] assms DERIV_minus by fastforce then show ?thesis by simp qed lemma DERIV_pos_imp_increasing_at_bot: fixes f :: "real \ real" assumes "\x. x \ b \ (\y. DERIV f x :> y \ y > 0)" and lim: "(f \ flim) at_bot" shows "flim < f b" proof - have "\N. \n\N. f n \ f (b - 1)" by (rule_tac x="b - 2" in exI) (force intro: order.strict_implies_order DERIV_pos_imp_increasing assms) then have "flim \ f (b - 1)" by (auto simp: eventually_at_bot_linorder tendsto_upperbound [OF lim]) also have "\ < f b" by (force intro: DERIV_pos_imp_increasing [where f=f] assms) finally show ?thesis . qed lemma DERIV_neg_imp_decreasing_at_top: fixes f :: "real \ real" assumes der: "\x. x \ b \ \y. DERIV f x :> y \ y < 0" and lim: "(f \ flim) at_top" shows "flim < f b" apply (rule DERIV_pos_imp_increasing_at_bot [where f = "\i. f (-i)" and b = "-b", simplified]) apply (metis DERIV_mirror der le_minus_iff neg_0_less_iff_less) apply (metis filterlim_at_top_mirror lim) done text \Derivative of inverse function\ lemma DERIV_inverse_function: fixes f g :: "real \ real" assumes der: "DERIV f (g x) :> D" and neq: "D \ 0" and x: "a < x" "x < b" and inj: "\y. \a < y; y < b\ \ f (g y) = y" and cont: "isCont g x" shows "DERIV g x :> inverse D" unfolding has_field_derivative_iff proof (rule LIM_equal2) show "0 < min (x - a) (b - x)" using x by arith next fix y assume "norm (y - x) < min (x - a) (b - x)" then have "a < y" and "y < b" by (simp_all add: abs_less_iff) then show "(g y - g x) / (y - x) = inverse ((f (g y) - x) / (g y - g x))" by (simp add: inj) next have "(\z. (f z - f (g x)) / (z - g x)) \g x\ D" by (rule der [unfolded has_field_derivative_iff]) then have 1: "(\z. (f z - x) / (z - g x)) \g x\ D" using inj x by simp have 2: "\d>0. \y. y \ x \ norm (y - x) < d \ g y \ g x" proof (rule exI, safe) show "0 < min (x - a) (b - x)" using x by simp next fix y assume "norm (y - x) < min (x - a) (b - x)" then have y: "a < y" "y < b" by (simp_all add: abs_less_iff) assume "g y = g x" then have "f (g y) = f (g x)" by simp then have "y = x" using inj y x by simp also assume "y \ x" finally show False by simp qed have "(\y. (f (g y) - x) / (g y - g x)) \x\ D" using cont 1 2 by (rule isCont_LIM_compose2) then show "(\y. inverse ((f (g y) - x) / (g y - g x))) \x\ inverse D" using neq by (rule tendsto_inverse) qed subsection \Generalized Mean Value Theorem\ theorem GMVT: fixes a b :: real assumes alb: "a < b" and fc: "\x. a \ x \ x \ b \ isCont f x" and fd: "\x. a < x \ x < b \ f differentiable (at x)" and gc: "\x. a \ x \ x \ b \ isCont g x" and gd: "\x. a < x \ x < b \ g differentiable (at x)" shows "\g'c f'c c. DERIV g c :> g'c \ DERIV f c :> f'c \ a < c \ c < b \ (f b - f a) * g'c = (g b - g a) * f'c" proof - let ?h = "\x. (f b - f a) * g x - (g b - g a) * f x" have "\l z. a < z \ z < b \ DERIV ?h z :> l \ ?h b - ?h a = (b - a) * l" proof (rule MVT) from assms show "a < b" by simp show "continuous_on {a..b} ?h" by (simp add: continuous_at_imp_continuous_on fc gc) show "\x. \a < x; x < b\ \ ?h differentiable (at x)" using fd gd by simp qed then obtain l where l: "\z. a < z \ z < b \ DERIV ?h z :> l \ ?h b - ?h a = (b - a) * l" .. then obtain c where c: "a < c \ c < b \ DERIV ?h c :> l \ ?h b - ?h a = (b - a) * l" .. from c have cint: "a < c \ c < b" by auto then obtain g'c where g'c: "DERIV g c :> g'c" using gd real_differentiable_def by blast from c have "a < c \ c < b" by auto then obtain f'c where f'c: "DERIV f c :> f'c" using fd real_differentiable_def by blast from c have "DERIV ?h c :> l" by auto moreover have "DERIV ?h c :> g'c * (f b - f a) - f'c * (g b - g a)" using g'c f'c by (auto intro!: derivative_eq_intros) ultimately have leq: "l = g'c * (f b - f a) - f'c * (g b - g a)" by (rule DERIV_unique) have "?h b - ?h a = (b - a) * (g'c * (f b - f a) - f'c * (g b - g a))" proof - from c have "?h b - ?h a = (b - a) * l" by auto also from leq have "\ = (b - a) * (g'c * (f b - f a) - f'c * (g b - g a))" by simp finally show ?thesis by simp qed moreover have "?h b - ?h a = 0" proof - have "?h b - ?h a = ((f b)*(g b) - (f a)*(g b) - (g b)*(f b) + (g a)*(f b)) - ((f b)*(g a) - (f a)*(g a) - (g b)*(f a) + (g a)*(f a))" by (simp add: algebra_simps) then show ?thesis by auto qed ultimately have "(b - a) * (g'c * (f b - f a) - f'c * (g b - g a)) = 0" by auto with alb have "g'c * (f b - f a) - f'c * (g b - g a) = 0" by simp then have "g'c * (f b - f a) = f'c * (g b - g a)" by simp then have "(f b - f a) * g'c = (g b - g a) * f'c" by (simp add: ac_simps) with g'c f'c cint show ?thesis by auto qed lemma GMVT': fixes f g :: "real \ real" assumes "a < b" and isCont_f: "\z. a \ z \ z \ b \ isCont f z" and isCont_g: "\z. a \ z \ z \ b \ isCont g z" and DERIV_g: "\z. a < z \ z < b \ DERIV g z :> (g' z)" and DERIV_f: "\z. a < z \ z < b \ DERIV f z :> (f' z)" shows "\c. a < c \ c < b \ (f b - f a) * g' c = (g b - g a) * f' c" proof - have "\g'c f'c c. DERIV g c :> g'c \ DERIV f c :> f'c \ a < c \ c < b \ (f b - f a) * g'c = (g b - g a) * f'c" using assms by (intro GMVT) (force simp: real_differentiable_def)+ then obtain c where "a < c" "c < b" "(f b - f a) * g' c = (g b - g a) * f' c" using DERIV_f DERIV_g by (force dest: DERIV_unique) then show ?thesis by auto qed subsection \L'Hopitals rule\ lemma isCont_If_ge: fixes a :: "'a :: linorder_topology" assumes "continuous (at_left a) g" and f: "(f \ g a) (at_right a)" shows "isCont (\x. if x \ a then g x else f x) a" (is "isCont ?gf a") proof - have g: "(g \ g a) (at_left a)" using assms continuous_within by blast show ?thesis unfolding isCont_def continuous_within proof (intro filterlim_split_at; simp) show "(?gf \ g a) (at_left a)" by (subst filterlim_cong[OF refl refl, where g=g]) (simp_all add: eventually_at_filter less_le g) show "(?gf \ g a) (at_right a)" by (subst filterlim_cong[OF refl refl, where g=f]) (simp_all add: eventually_at_filter less_le f) qed qed lemma lhopital_right_0: fixes f0 g0 :: "real \ real" assumes f_0: "(f0 \ 0) (at_right 0)" and g_0: "(g0 \ 0) (at_right 0)" and ev: "eventually (\x. g0 x \ 0) (at_right 0)" "eventually (\x. g' x \ 0) (at_right 0)" "eventually (\x. DERIV f0 x :> f' x) (at_right 0)" "eventually (\x. DERIV g0 x :> g' x) (at_right 0)" and lim: "filterlim (\ x. (f' x / g' x)) F (at_right 0)" shows "filterlim (\ x. f0 x / g0 x) F (at_right 0)" proof - define f where [abs_def]: "f x = (if x \ 0 then 0 else f0 x)" for x then have "f 0 = 0" by simp define g where [abs_def]: "g x = (if x \ 0 then 0 else g0 x)" for x then have "g 0 = 0" by simp have "eventually (\x. g0 x \ 0 \ g' x \ 0 \ DERIV f0 x :> (f' x) \ DERIV g0 x :> (g' x)) (at_right 0)" using ev by eventually_elim auto then obtain a where [arith]: "0 < a" and g0_neq_0: "\x. 0 < x \ x < a \ g0 x \ 0" and g'_neq_0: "\x. 0 < x \ x < a \ g' x \ 0" and f0: "\x. 0 < x \ x < a \ DERIV f0 x :> (f' x)" and g0: "\x. 0 < x \ x < a \ DERIV g0 x :> (g' x)" unfolding eventually_at by (auto simp: dist_real_def) have g_neq_0: "\x. 0 < x \ x < a \ g x \ 0" using g0_neq_0 by (simp add: g_def) have f: "DERIV f x :> (f' x)" if x: "0 < x" "x < a" for x using that by (intro DERIV_cong_ev[THEN iffD1, OF _ _ _ f0[OF x]]) (auto simp: f_def eventually_nhds_metric dist_real_def intro!: exI[of _ x]) have g: "DERIV g x :> (g' x)" if x: "0 < x" "x < a" for x using that by (intro DERIV_cong_ev[THEN iffD1, OF _ _ _ g0[OF x]]) (auto simp: g_def eventually_nhds_metric dist_real_def intro!: exI[of _ x]) have "isCont f 0" unfolding f_def by (intro isCont_If_ge f_0 continuous_const) have "isCont g 0" unfolding g_def by (intro isCont_If_ge g_0 continuous_const) have "\\. \x\{0 <..< a}. 0 < \ x \ \ x < x \ f x / g x = f' (\ x) / g' (\ x)" proof (rule bchoice, rule ballI) fix x assume "x \ {0 <..< a}" then have x[arith]: "0 < x" "x < a" by auto with g'_neq_0 g_neq_0 \g 0 = 0\ have g': "\x. 0 < x \ x < a \ 0 \ g' x" "g 0 \ g x" by auto have "\x. 0 \ x \ x < a \ isCont f x" using \isCont f 0\ f by (auto intro: DERIV_isCont simp: le_less) moreover have "\x. 0 \ x \ x < a \ isCont g x" using \isCont g 0\ g by (auto intro: DERIV_isCont simp: le_less) ultimately have "\c. 0 < c \ c < x \ (f x - f 0) * g' c = (g x - g 0) * f' c" using f g \x < a\ by (intro GMVT') auto then obtain c where *: "0 < c" "c < x" "(f x - f 0) * g' c = (g x - g 0) * f' c" by blast moreover from * g'(1)[of c] g'(2) have "(f x - f 0) / (g x - g 0) = f' c / g' c" by (simp add: field_simps) ultimately show "\y. 0 < y \ y < x \ f x / g x = f' y / g' y" using \f 0 = 0\ \g 0 = 0\ by (auto intro!: exI[of _ c]) qed then obtain \ where "\x\{0 <..< a}. 0 < \ x \ \ x < x \ f x / g x = f' (\ x) / g' (\ x)" .. then have \: "eventually (\x. 0 < \ x \ \ x < x \ f x / g x = f' (\ x) / g' (\ x)) (at_right 0)" unfolding eventually_at by (intro exI[of _ a]) (auto simp: dist_real_def) moreover from \ have "eventually (\x. norm (\ x) \ x) (at_right 0)" by eventually_elim auto then have "((\x. norm (\ x)) \ 0) (at_right 0)" by (rule_tac real_tendsto_sandwich[where f="\x. 0" and h="\x. x"]) auto then have "(\ \ 0) (at_right 0)" by (rule tendsto_norm_zero_cancel) with \ have "filterlim \ (at_right 0) (at_right 0)" by (auto elim!: eventually_mono simp: filterlim_at) from this lim have "filterlim (\t. f' (\ t) / g' (\ t)) F (at_right 0)" by (rule_tac filterlim_compose[of _ _ _ \]) ultimately have "filterlim (\t. f t / g t) F (at_right 0)" (is ?P) by (rule_tac filterlim_cong[THEN iffD1, OF refl refl]) (auto elim: eventually_mono) also have "?P \ ?thesis" by (rule filterlim_cong) (auto simp: f_def g_def eventually_at_filter) finally show ?thesis . qed lemma lhopital_right: "(f \ 0) (at_right x) \ (g \ 0) (at_right x) \ eventually (\x. g x \ 0) (at_right x) \ eventually (\x. g' x \ 0) (at_right x) \ eventually (\x. DERIV f x :> f' x) (at_right x) \ eventually (\x. DERIV g x :> g' x) (at_right x) \ filterlim (\ x. (f' x / g' x)) F (at_right x) \ filterlim (\ x. f x / g x) F (at_right x)" for x :: real unfolding eventually_at_right_to_0[of _ x] filterlim_at_right_to_0[of _ _ x] DERIV_shift by (rule lhopital_right_0) lemma lhopital_left: "(f \ 0) (at_left x) \ (g \ 0) (at_left x) \ eventually (\x. g x \ 0) (at_left x) \ eventually (\x. g' x \ 0) (at_left x) \ eventually (\x. DERIV f x :> f' x) (at_left x) \ eventually (\x. DERIV g x :> g' x) (at_left x) \ filterlim (\ x. (f' x / g' x)) F (at_left x) \ filterlim (\ x. f x / g x) F (at_left x)" for x :: real unfolding eventually_at_left_to_right filterlim_at_left_to_right DERIV_mirror by (rule lhopital_right[where f'="\x. - f' (- x)"]) (auto simp: DERIV_mirror) lemma lhopital: "(f \ 0) (at x) \ (g \ 0) (at x) \ eventually (\x. g x \ 0) (at x) \ eventually (\x. g' x \ 0) (at x) \ eventually (\x. DERIV f x :> f' x) (at x) \ eventually (\x. DERIV g x :> g' x) (at x) \ filterlim (\ x. (f' x / g' x)) F (at x) \ filterlim (\ x. f x / g x) F (at x)" for x :: real unfolding eventually_at_split filterlim_at_split by (auto intro!: lhopital_right[of f x g g' f'] lhopital_left[of f x g g' f']) lemma lhopital_right_0_at_top: fixes f g :: "real \ real" assumes g_0: "LIM x at_right 0. g x :> at_top" and ev: "eventually (\x. g' x \ 0) (at_right 0)" "eventually (\x. DERIV f x :> f' x) (at_right 0)" "eventually (\x. DERIV g x :> g' x) (at_right 0)" and lim: "((\ x. (f' x / g' x)) \ x) (at_right 0)" shows "((\ x. f x / g x) \ x) (at_right 0)" unfolding tendsto_iff proof safe fix e :: real assume "0 < e" with lim[unfolded tendsto_iff, rule_format, of "e / 4"] have "eventually (\t. dist (f' t / g' t) x < e / 4) (at_right 0)" by simp from eventually_conj[OF eventually_conj[OF ev(1) ev(2)] eventually_conj[OF ev(3) this]] obtain a where [arith]: "0 < a" and g'_neq_0: "\x. 0 < x \ x < a \ g' x \ 0" and f0: "\x. 0 < x \ x \ a \ DERIV f x :> (f' x)" and g0: "\x. 0 < x \ x \ a \ DERIV g x :> (g' x)" and Df: "\t. 0 < t \ t < a \ dist (f' t / g' t) x < e / 4" unfolding eventually_at_le by (auto simp: dist_real_def) from Df have "eventually (\t. t < a) (at_right 0)" "eventually (\t::real. 0 < t) (at_right 0)" unfolding eventually_at by (auto intro!: exI[of _ a] simp: dist_real_def) moreover have "eventually (\t. 0 < g t) (at_right 0)" "eventually (\t. g a < g t) (at_right 0)" using g_0 by (auto elim: eventually_mono simp: filterlim_at_top_dense) moreover have inv_g: "((\x. inverse (g x)) \ 0) (at_right 0)" using tendsto_inverse_0 filterlim_mono[OF g_0 at_top_le_at_infinity order_refl] by (rule filterlim_compose) then have "((\x. norm (1 - g a * inverse (g x))) \ norm (1 - g a * 0)) (at_right 0)" by (intro tendsto_intros) then have "((\x. norm (1 - g a / g x)) \ 1) (at_right 0)" by (simp add: inverse_eq_divide) from this[unfolded tendsto_iff, rule_format, of 1] have "eventually (\x. norm (1 - g a / g x) < 2) (at_right 0)" by (auto elim!: eventually_mono simp: dist_real_def) moreover from inv_g have "((\t. norm ((f a - x * g a) * inverse (g t))) \ norm ((f a - x * g a) * 0)) (at_right 0)" by (intro tendsto_intros) then have "((\t. norm (f a - x * g a) / norm (g t)) \ 0) (at_right 0)" by (simp add: inverse_eq_divide) from this[unfolded tendsto_iff, rule_format, of "e / 2"] \0 < e\ have "eventually (\t. norm (f a - x * g a) / norm (g t) < e / 2) (at_right 0)" by (auto simp: dist_real_def) ultimately show "eventually (\t. dist (f t / g t) x < e) (at_right 0)" proof eventually_elim fix t assume t[arith]: "0 < t" "t < a" "g a < g t" "0 < g t" assume ineq: "norm (1 - g a / g t) < 2" "norm (f a - x * g a) / norm (g t) < e / 2" have "\y. t < y \ y < a \ (g a - g t) * f' y = (f a - f t) * g' y" using f0 g0 t(1,2) by (intro GMVT') (force intro!: DERIV_isCont)+ then obtain y where [arith]: "t < y" "y < a" and D_eq0: "(g a - g t) * f' y = (f a - f t) * g' y" by blast from D_eq0 have D_eq: "(f t - f a) / (g t - g a) = f' y / g' y" using \g a < g t\ g'_neq_0[of y] by (auto simp add: field_simps) have *: "f t / g t - x = ((f t - f a) / (g t - g a) - x) * (1 - g a / g t) + (f a - x * g a) / g t" by (simp add: field_simps) have "norm (f t / g t - x) \ norm (((f t - f a) / (g t - g a) - x) * (1 - g a / g t)) + norm ((f a - x * g a) / g t)" unfolding * by (rule norm_triangle_ineq) also have "\ = dist (f' y / g' y) x * norm (1 - g a / g t) + norm (f a - x * g a) / norm (g t)" by (simp add: abs_mult D_eq dist_real_def) also have "\ < (e / 4) * 2 + e / 2" using ineq Df[of y] \0 < e\ by (intro add_le_less_mono mult_mono) auto finally show "dist (f t / g t) x < e" by (simp add: dist_real_def) qed qed lemma lhopital_right_at_top: "LIM x at_right x. (g::real \ real) x :> at_top \ eventually (\x. g' x \ 0) (at_right x) \ eventually (\x. DERIV f x :> f' x) (at_right x) \ eventually (\x. DERIV g x :> g' x) (at_right x) \ ((\ x. (f' x / g' x)) \ y) (at_right x) \ ((\ x. f x / g x) \ y) (at_right x)" unfolding eventually_at_right_to_0[of _ x] filterlim_at_right_to_0[of _ _ x] DERIV_shift by (rule lhopital_right_0_at_top) lemma lhopital_left_at_top: "LIM x at_left x. g x :> at_top \ eventually (\x. g' x \ 0) (at_left x) \ eventually (\x. DERIV f x :> f' x) (at_left x) \ eventually (\x. DERIV g x :> g' x) (at_left x) \ ((\ x. (f' x / g' x)) \ y) (at_left x) \ ((\ x. f x / g x) \ y) (at_left x)" for x :: real unfolding eventually_at_left_to_right filterlim_at_left_to_right DERIV_mirror by (rule lhopital_right_at_top[where f'="\x. - f' (- x)"]) (auto simp: DERIV_mirror) lemma lhopital_at_top: "LIM x at x. (g::real \ real) x :> at_top \ eventually (\x. g' x \ 0) (at x) \ eventually (\x. DERIV f x :> f' x) (at x) \ eventually (\x. DERIV g x :> g' x) (at x) \ ((\ x. (f' x / g' x)) \ y) (at x) \ ((\ x. f x / g x) \ y) (at x)" unfolding eventually_at_split filterlim_at_split by (auto intro!: lhopital_right_at_top[of g x g' f f'] lhopital_left_at_top[of g x g' f f']) lemma lhospital_at_top_at_top: fixes f g :: "real \ real" assumes g_0: "LIM x at_top. g x :> at_top" and g': "eventually (\x. g' x \ 0) at_top" and Df: "eventually (\x. DERIV f x :> f' x) at_top" and Dg: "eventually (\x. DERIV g x :> g' x) at_top" and lim: "((\ x. (f' x / g' x)) \ x) at_top" shows "((\ x. f x / g x) \ x) at_top" unfolding filterlim_at_top_to_right proof (rule lhopital_right_0_at_top) let ?F = "\x. f (inverse x)" let ?G = "\x. g (inverse x)" let ?R = "at_right (0::real)" let ?D = "\f' x. f' (inverse x) * - (inverse x ^ Suc (Suc 0))" show "LIM x ?R. ?G x :> at_top" using g_0 unfolding filterlim_at_top_to_right . show "eventually (\x. DERIV ?G x :> ?D g' x) ?R" unfolding eventually_at_right_to_top using Dg eventually_ge_at_top[where c=1] by eventually_elim (rule derivative_eq_intros DERIV_chain'[where f=inverse] | simp)+ show "eventually (\x. DERIV ?F x :> ?D f' x) ?R" unfolding eventually_at_right_to_top using Df eventually_ge_at_top[where c=1] by eventually_elim (rule derivative_eq_intros DERIV_chain'[where f=inverse] | simp)+ show "eventually (\x. ?D g' x \ 0) ?R" unfolding eventually_at_right_to_top using g' eventually_ge_at_top[where c=1] by eventually_elim auto show "((\x. ?D f' x / ?D g' x) \ x) ?R" unfolding filterlim_at_right_to_top apply (intro filterlim_cong[THEN iffD2, OF refl refl _ lim]) using eventually_ge_at_top[where c=1] by eventually_elim simp qed lemma lhopital_right_at_top_at_top: fixes f g :: "real \ real" assumes f_0: "LIM x at_right a. f x :> at_top" assumes g_0: "LIM x at_right a. g x :> at_top" and ev: "eventually (\x. DERIV f x :> f' x) (at_right a)" "eventually (\x. DERIV g x :> g' x) (at_right a)" and lim: "filterlim (\ x. (f' x / g' x)) at_top (at_right a)" shows "filterlim (\ x. f x / g x) at_top (at_right a)" proof - from lim have pos: "eventually (\x. f' x / g' x > 0) (at_right a)" unfolding filterlim_at_top_dense by blast have "((\x. g x / f x) \ 0) (at_right a)" proof (rule lhopital_right_at_top) from pos show "eventually (\x. f' x \ 0) (at_right a)" by eventually_elim auto from tendsto_inverse_0_at_top[OF lim] show "((\x. g' x / f' x) \ 0) (at_right a)" by simp qed fact+ moreover from f_0 g_0 have "eventually (\x. f x > 0) (at_right a)" "eventually (\x. g x > 0) (at_right a)" unfolding filterlim_at_top_dense by blast+ hence "eventually (\x. g x / f x > 0) (at_right a)" by eventually_elim simp ultimately have "filterlim (\x. inverse (g x / f x)) at_top (at_right a)" by (rule filterlim_inverse_at_top) thus ?thesis by simp qed lemma lhopital_right_at_top_at_bot: fixes f g :: "real \ real" assumes f_0: "LIM x at_right a. f x :> at_top" assumes g_0: "LIM x at_right a. g x :> at_bot" and ev: "eventually (\x. DERIV f x :> f' x) (at_right a)" "eventually (\x. DERIV g x :> g' x) (at_right a)" and lim: "filterlim (\ x. (f' x / g' x)) at_bot (at_right a)" shows "filterlim (\ x. f x / g x) at_bot (at_right a)" proof - from ev(2) have ev': "eventually (\x. DERIV (\x. -g x) x :> -g' x) (at_right a)" by eventually_elim (auto intro: derivative_intros) have "filterlim (\x. f x / (-g x)) at_top (at_right a)" by (rule lhopital_right_at_top_at_top[where f' = f' and g' = "\x. -g' x"]) (insert assms ev', auto simp: filterlim_uminus_at_bot) hence "filterlim (\x. -(f x / g x)) at_top (at_right a)" by simp thus ?thesis by (simp add: filterlim_uminus_at_bot) qed lemma lhopital_left_at_top_at_top: fixes f g :: "real \ real" assumes f_0: "LIM x at_left a. f x :> at_top" assumes g_0: "LIM x at_left a. g x :> at_top" and ev: "eventually (\x. DERIV f x :> f' x) (at_left a)" "eventually (\x. DERIV g x :> g' x) (at_left a)" and lim: "filterlim (\ x. (f' x / g' x)) at_top (at_left a)" shows "filterlim (\ x. f x / g x) at_top (at_left a)" by (insert assms, unfold eventually_at_left_to_right filterlim_at_left_to_right DERIV_mirror, rule lhopital_right_at_top_at_top[where f'="\x. - f' (- x)"]) (insert assms, auto simp: DERIV_mirror) lemma lhopital_left_at_top_at_bot: fixes f g :: "real \ real" assumes f_0: "LIM x at_left a. f x :> at_top" assumes g_0: "LIM x at_left a. g x :> at_bot" and ev: "eventually (\x. DERIV f x :> f' x) (at_left a)" "eventually (\x. DERIV g x :> g' x) (at_left a)" and lim: "filterlim (\ x. (f' x / g' x)) at_bot (at_left a)" shows "filterlim (\ x. f x / g x) at_bot (at_left a)" by (insert assms, unfold eventually_at_left_to_right filterlim_at_left_to_right DERIV_mirror, rule lhopital_right_at_top_at_bot[where f'="\x. - f' (- x)"]) (insert assms, auto simp: DERIV_mirror) lemma lhopital_at_top_at_top: fixes f g :: "real \ real" assumes f_0: "LIM x at a. f x :> at_top" assumes g_0: "LIM x at a. g x :> at_top" and ev: "eventually (\x. DERIV f x :> f' x) (at a)" "eventually (\x. DERIV g x :> g' x) (at a)" and lim: "filterlim (\ x. (f' x / g' x)) at_top (at a)" shows "filterlim (\ x. f x / g x) at_top (at a)" using assms unfolding eventually_at_split filterlim_at_split by (auto intro!: lhopital_right_at_top_at_top[of f a g f' g'] lhopital_left_at_top_at_top[of f a g f' g']) lemma lhopital_at_top_at_bot: fixes f g :: "real \ real" assumes f_0: "LIM x at a. f x :> at_top" assumes g_0: "LIM x at a. g x :> at_bot" and ev: "eventually (\x. DERIV f x :> f' x) (at a)" "eventually (\x. DERIV g x :> g' x) (at a)" and lim: "filterlim (\ x. (f' x / g' x)) at_bot (at a)" shows "filterlim (\ x. f x / g x) at_bot (at a)" using assms unfolding eventually_at_split filterlim_at_split by (auto intro!: lhopital_right_at_top_at_bot[of f a g f' g'] lhopital_left_at_top_at_bot[of f a g f' g']) end diff --git a/src/HOL/Fun.thy b/src/HOL/Fun.thy --- a/src/HOL/Fun.thy +++ b/src/HOL/Fun.thy @@ -1,1367 +1,1378 @@ (* Title: HOL/Fun.thy Author: Tobias Nipkow, Cambridge University Computer Laboratory Author: Andrei Popescu, TU Muenchen Copyright 1994, 2012 *) section \Notions about functions\ theory Fun imports Set keywords "functor" :: thy_goal_defn begin lemma apply_inverse: "f x = u \ (\x. P x \ g (f x) = x) \ P x \ x = g u" by auto text \Uniqueness, so NOT the axiom of choice.\ lemma uniq_choice: "\x. \!y. Q x y \ \f. \x. Q x (f x)" by (force intro: theI') lemma b_uniq_choice: "\x\S. \!y. Q x y \ \f. \x\S. Q x (f x)" by (force intro: theI') subsection \The Identity Function \id\\ definition id :: "'a \ 'a" where "id = (\x. x)" lemma id_apply [simp]: "id x = x" by (simp add: id_def) lemma image_id [simp]: "image id = id" by (simp add: id_def fun_eq_iff) lemma vimage_id [simp]: "vimage id = id" by (simp add: id_def fun_eq_iff) lemma eq_id_iff: "(\x. f x = x) \ f = id" by auto code_printing constant id \ (Haskell) "id" subsection \The Composition Operator \f \ g\\ definition comp :: "('b \ 'c) \ ('a \ 'b) \ 'a \ 'c" (infixl "\" 55) where "f \ g = (\x. f (g x))" notation (ASCII) comp (infixl "o" 55) lemma comp_apply [simp]: "(f \ g) x = f (g x)" by (simp add: comp_def) lemma comp_assoc: "(f \ g) \ h = f \ (g \ h)" by (simp add: fun_eq_iff) lemma id_comp [simp]: "id \ g = g" by (simp add: fun_eq_iff) lemma comp_id [simp]: "f \ id = f" by (simp add: fun_eq_iff) lemma comp_eq_dest: "a \ b = c \ d \ a (b v) = c (d v)" by (simp add: fun_eq_iff) lemma comp_eq_elim: "a \ b = c \ d \ ((\v. a (b v) = c (d v)) \ R) \ R" by (simp add: fun_eq_iff) lemma comp_eq_dest_lhs: "a \ b = c \ a (b v) = c v" by clarsimp lemma comp_eq_id_dest: "a \ b = id \ c \ a (b v) = c v" by clarsimp lemma image_comp: "f ` (g ` r) = (f \ g) ` r" by auto lemma vimage_comp: "f -` (g -` x) = (g \ f) -` x" by auto lemma image_eq_imp_comp: "f ` A = g ` B \ (h \ f) ` A = (h \ g) ` B" by (auto simp: comp_def elim!: equalityE) lemma image_bind: "f ` (Set.bind A g) = Set.bind A ((`) f \ g)" by (auto simp add: Set.bind_def) lemma bind_image: "Set.bind (f ` A) g = Set.bind A (g \ f)" by (auto simp add: Set.bind_def) lemma (in group_add) minus_comp_minus [simp]: "uminus \ uminus = id" by (simp add: fun_eq_iff) lemma (in boolean_algebra) minus_comp_minus [simp]: "uminus \ uminus = id" by (simp add: fun_eq_iff) code_printing constant comp \ (SML) infixl 5 "o" and (Haskell) infixr 9 "." subsection \The Forward Composition Operator \fcomp\\ definition fcomp :: "('a \ 'b) \ ('b \ 'c) \ 'a \ 'c" (infixl "\>" 60) where "f \> g = (\x. g (f x))" lemma fcomp_apply [simp]: "(f \> g) x = g (f x)" by (simp add: fcomp_def) lemma fcomp_assoc: "(f \> g) \> h = f \> (g \> h)" by (simp add: fcomp_def) lemma id_fcomp [simp]: "id \> g = g" by (simp add: fcomp_def) lemma fcomp_id [simp]: "f \> id = f" by (simp add: fcomp_def) lemma fcomp_comp: "fcomp f g = comp g f" by (simp add: ext) code_printing constant fcomp \ (Eval) infixl 1 "#>" no_notation fcomp (infixl "\>" 60) subsection \Mapping functions\ definition map_fun :: "('c \ 'a) \ ('b \ 'd) \ ('a \ 'b) \ 'c \ 'd" where "map_fun f g h = g \ h \ f" lemma map_fun_apply [simp]: "map_fun f g h x = g (h (f x))" by (simp add: map_fun_def) subsection \Injectivity and Bijectivity\ definition inj_on :: "('a \ 'b) \ 'a set \ bool" \ \injective\ where "inj_on f A \ (\x\A. \y\A. f x = f y \ x = y)" definition bij_betw :: "('a \ 'b) \ 'a set \ 'b set \ bool" \ \bijective\ where "bij_betw f A B \ inj_on f A \ f ` A = B" text \ A common special case: functions injective, surjective or bijective over the entire domain type. \ abbreviation inj :: "('a \ 'b) \ bool" where "inj f \ inj_on f UNIV" abbreviation surj :: "('a \ 'b) \ bool" where "surj f \ range f = UNIV" translations \ \The negated case:\ "\ CONST surj f" \ "CONST range f \ CONST UNIV" abbreviation bij :: "('a \ 'b) \ bool" where "bij f \ bij_betw f UNIV UNIV" lemma inj_def: "inj f \ (\x y. f x = f y \ x = y)" unfolding inj_on_def by blast lemma injI: "(\x y. f x = f y \ x = y) \ inj f" unfolding inj_def by blast theorem range_ex1_eq: "inj f \ b \ range f \ (\!x. b = f x)" unfolding inj_def by blast lemma injD: "inj f \ f x = f y \ x = y" by (simp add: inj_def) lemma inj_on_eq_iff: "inj_on f A \ x \ A \ y \ A \ f x = f y \ x = y" by (auto simp: inj_on_def) lemma inj_on_cong: "(\a. a \ A \ f a = g a) \ inj_on f A \ inj_on g A" by (auto simp: inj_on_def) lemma image_strict_mono: "inj_on f B \ A \ B \ f ` A \ f ` B" unfolding inj_on_def by blast lemma inj_compose: "inj f \ inj g \ inj (f \ g)" by (simp add: inj_def) lemma inj_fun: "inj f \ inj (\x y. f x)" by (simp add: inj_def fun_eq_iff) lemma inj_eq: "inj f \ f x = f y \ x = y" by (simp add: inj_on_eq_iff) lemma inj_on_iff_Uniq: "inj_on f A \ (\x\A. \\<^sub>\\<^sub>1y. y\A \ f x = f y)" by (auto simp: Uniq_def inj_on_def) lemma inj_on_id[simp]: "inj_on id A" by (simp add: inj_on_def) lemma inj_on_id2[simp]: "inj_on (\x. x) A" by (simp add: inj_on_def) lemma inj_on_Int: "inj_on f A \ inj_on f B \ inj_on f (A \ B)" unfolding inj_on_def by blast lemma surj_id: "surj id" by simp lemma bij_id[simp]: "bij id" by (simp add: bij_betw_def) lemma bij_uminus: "bij (uminus :: 'a \ 'a::group_add)" unfolding bij_betw_def inj_on_def by (force intro: minus_minus [symmetric]) lemma bij_betwE: "bij_betw f A B \ \a\A. f a \ B" unfolding bij_betw_def by auto lemma inj_onI [intro?]: "(\x y. x \ A \ y \ A \ f x = f y \ x = y) \ inj_on f A" by (simp add: inj_on_def) lemma inj_on_inverseI: "(\x. x \ A \ g (f x) = x) \ inj_on f A" by (auto dest: arg_cong [of concl: g] simp add: inj_on_def) lemma inj_onD: "inj_on f A \ f x = f y \ x \ A \ y \ A \ x = y" unfolding inj_on_def by blast lemma inj_on_subset: assumes "inj_on f A" and "B \ A" shows "inj_on f B" proof (rule inj_onI) fix a b assume "a \ B" and "b \ B" with assms have "a \ A" and "b \ A" by auto moreover assume "f a = f b" ultimately show "a = b" using assms by (auto dest: inj_onD) qed lemma comp_inj_on: "inj_on f A \ inj_on g (f ` A) \ inj_on (g \ f) A" by (simp add: comp_def inj_on_def) lemma inj_on_imageI: "inj_on (g \ f) A \ inj_on g (f ` A)" by (auto simp add: inj_on_def) lemma inj_on_image_iff: "\x\A. \y\A. g (f x) = g (f y) \ g x = g y \ inj_on f A \ inj_on g (f ` A) \ inj_on g A" unfolding inj_on_def by blast lemma inj_on_contraD: "inj_on f A \ x \ y \ x \ A \ y \ A \ f x \ f y" unfolding inj_on_def by blast lemma inj_singleton [simp]: "inj_on (\x. {x}) A" by (simp add: inj_on_def) lemma inj_on_empty[iff]: "inj_on f {}" by (simp add: inj_on_def) lemma subset_inj_on: "inj_on f B \ A \ B \ inj_on f A" unfolding inj_on_def by blast lemma inj_on_Un: "inj_on f (A \ B) \ inj_on f A \ inj_on f B \ f ` (A - B) \ f ` (B - A) = {}" unfolding inj_on_def by (blast intro: sym) lemma inj_on_insert [iff]: "inj_on f (insert a A) \ inj_on f A \ f a \ f ` (A - {a})" unfolding inj_on_def by (blast intro: sym) lemma inj_on_diff: "inj_on f A \ inj_on f (A - B)" unfolding inj_on_def by blast lemma comp_inj_on_iff: "inj_on f A \ inj_on f' (f ` A) \ inj_on (f' \ f) A" by (auto simp: comp_inj_on inj_on_def) lemma inj_on_imageI2: "inj_on (f' \ f) A \ inj_on f A" by (auto simp: comp_inj_on inj_on_def) lemma inj_img_insertE: assumes "inj_on f A" assumes "x \ B" and "insert x B = f ` A" obtains x' A' where "x' \ A'" and "A = insert x' A'" and "x = f x'" and "B = f ` A'" proof - from assms have "x \ f ` A" by auto then obtain x' where *: "x' \ A" "x = f x'" by auto then have A: "A = insert x' (A - {x'})" by auto with assms * have B: "B = f ` (A - {x'})" by (auto dest: inj_on_contraD) have "x' \ A - {x'}" by simp from this A \x = f x'\ B show ?thesis .. qed lemma linorder_inj_onI: fixes A :: "'a::order set" assumes ne: "\x y. \x < y; x\A; y\A\ \ f x \ f y" and lin: "\x y. \x\A; y\A\ \ x\y \ y\x" shows "inj_on f A" proof (rule inj_onI) fix x y assume eq: "f x = f y" and "x\A" "y\A" then show "x = y" using lin [of x y] ne by (force simp: dual_order.order_iff_strict) qed lemma linorder_inj_onI': fixes A :: "'a :: linorder set" assumes "\i j. i \ A \ j \ A \ i < j \ f i \ f j" shows "inj_on f A" by (intro linorder_inj_onI) (auto simp add: assms) lemma linorder_injI: assumes "\x y::'a::linorder. x < y \ f x \ f y" shows "inj f" \ \Courtesy of Stephan Merz\ using assms by (simp add: linorder_inj_onI') lemma inj_on_image_Pow: "inj_on f A \inj_on (image f) (Pow A)" unfolding Pow_def inj_on_def by blast lemma bij_betw_image_Pow: "bij_betw f A B \ bij_betw (image f) (Pow A) (Pow B)" by (auto simp add: bij_betw_def inj_on_image_Pow image_Pow_surj) lemma surj_def: "surj f \ (\y. \x. y = f x)" by auto lemma surjI: assumes "\x. g (f x) = x" shows "surj g" using assms [symmetric] by auto lemma surjD: "surj f \ \x. y = f x" by (simp add: surj_def) lemma surjE: "surj f \ (\x. y = f x \ C) \ C" by (simp add: surj_def) blast lemma comp_surj: "surj f \ surj g \ surj (g \ f)" using image_comp [of g f UNIV] by simp lemma bij_betw_imageI: "inj_on f A \ f ` A = B \ bij_betw f A B" unfolding bij_betw_def by clarify lemma bij_betw_imp_surj_on: "bij_betw f A B \ f ` A = B" unfolding bij_betw_def by clarify lemma bij_betw_imp_surj: "bij_betw f A UNIV \ surj f" unfolding bij_betw_def by auto lemma bij_betw_empty1: "bij_betw f {} A \ A = {}" unfolding bij_betw_def by blast lemma bij_betw_empty2: "bij_betw f A {} \ A = {}" unfolding bij_betw_def by blast lemma inj_on_imp_bij_betw: "inj_on f A \ bij_betw f A (f ` A)" unfolding bij_betw_def by simp +lemma bij_betw_DiffI: + assumes "bij_betw f A B" "bij_betw f C D" "C \ A" "D \ B" + shows "bij_betw f (A - C) (B - D)" + using assms unfolding bij_betw_def inj_on_def by auto + +lemma bij_betw_singleton_iff [simp]: "bij_betw f {x} {y} \ f x = y" + by (auto simp: bij_betw_def) + +lemma bij_betw_singletonI [intro]: "f x = y \ bij_betw f {x} {y}" + by auto + lemma bij_betw_apply: "\bij_betw f A B; a \ A\ \ f a \ B" unfolding bij_betw_def by auto lemma bij_def: "bij f \ inj f \ surj f" by (rule bij_betw_def) lemma bijI: "inj f \ surj f \ bij f" by (rule bij_betw_imageI) lemma bij_is_inj: "bij f \ inj f" by (simp add: bij_def) lemma bij_is_surj: "bij f \ surj f" by (simp add: bij_def) lemma bij_betw_imp_inj_on: "bij_betw f A B \ inj_on f A" by (simp add: bij_betw_def) lemma bij_betw_trans: "bij_betw f A B \ bij_betw g B C \ bij_betw (g \ f) A C" by (auto simp add:bij_betw_def comp_inj_on) lemma bij_comp: "bij f \ bij g \ bij (g \ f)" by (rule bij_betw_trans) lemma bij_betw_comp_iff: "bij_betw f A A' \ bij_betw f' A' A'' \ bij_betw (f' \ f) A A''" by (auto simp add: bij_betw_def inj_on_def) lemma bij_betw_comp_iff2: assumes bij: "bij_betw f' A' A''" and img: "f ` A \ A'" shows "bij_betw f A A' \ bij_betw (f' \ f) A A''" (is "?L \ ?R") proof assume "?L" then show "?R" using assms by (auto simp add: bij_betw_comp_iff) next assume *: "?R" have "inj_on (f' \ f) A \ inj_on f A" using inj_on_imageI2 by blast moreover have "A' \ f ` A" proof fix a' assume **: "a' \ A'" with bij have "f' a' \ A''" unfolding bij_betw_def by auto with * obtain a where 1: "a \ A \ f' (f a) = f' a'" unfolding bij_betw_def by force with img have "f a \ A'" by auto with bij ** 1 have "f a = a'" unfolding bij_betw_def inj_on_def by auto with 1 show "a' \ f ` A" by auto qed ultimately show "?L" using img * by (auto simp add: bij_betw_def) qed lemma bij_betw_inv: assumes "bij_betw f A B" shows "\g. bij_betw g B A" proof - have i: "inj_on f A" and s: "f ` A = B" using assms by (auto simp: bij_betw_def) let ?P = "\b a. a \ A \ f a = b" let ?g = "\b. The (?P b)" have g: "?g b = a" if P: "?P b a" for a b proof - from that s have ex1: "\a. ?P b a" by blast then have uex1: "\!a. ?P b a" by (blast dest:inj_onD[OF i]) then show ?thesis using the1_equality[OF uex1, OF P] P by simp qed have "inj_on ?g B" proof (rule inj_onI) fix x y assume "x \ B" "y \ B" "?g x = ?g y" from s \x \ B\ obtain a1 where a1: "?P x a1" by blast from s \y \ B\ obtain a2 where a2: "?P y a2" by blast from g [OF a1] a1 g [OF a2] a2 \?g x = ?g y\ show "x = y" by simp qed moreover have "?g ` B = A" proof safe fix b assume "b \ B" with s obtain a where P: "?P b a" by blast with g[OF P] show "?g b \ A" by auto next fix a assume "a \ A" with s obtain b where P: "?P b a" by blast with s have "b \ B" by blast with g[OF P] have "\b\B. a = ?g b" by blast then show "a \ ?g ` B" by auto qed ultimately show ?thesis by (auto simp: bij_betw_def) qed lemma bij_betw_cong: "(\a. a \ A \ f a = g a) \ bij_betw f A A' = bij_betw g A A'" unfolding bij_betw_def inj_on_def by safe force+ (* somewhat slow *) lemma bij_betw_id[intro, simp]: "bij_betw id A A" unfolding bij_betw_def id_def by auto lemma bij_betw_id_iff: "bij_betw id A B \ A = B" by (auto simp add: bij_betw_def) lemma bij_betw_combine: "bij_betw f A B \ bij_betw f C D \ B \ D = {} \ bij_betw f (A \ C) (B \ D)" unfolding bij_betw_def inj_on_Un image_Un by auto lemma bij_betw_subset: "bij_betw f A A' \ B \ A \ f ` B = B' \ bij_betw f B B'" by (auto simp add: bij_betw_def inj_on_def) lemma bij_betw_ball: "bij_betw f A B \ (\b \ B. phi b) = (\a \ A. phi (f a))" unfolding bij_betw_def inj_on_def by blast lemma bij_pointE: assumes "bij f" obtains x where "y = f x" and "\x'. y = f x' \ x' = x" proof - from assms have "inj f" by (rule bij_is_inj) moreover from assms have "surj f" by (rule bij_is_surj) then have "y \ range f" by simp ultimately have "\!x. y = f x" by (simp add: range_ex1_eq) with that show thesis by blast qed lemma bij_iff: \<^marker>\contributor \Amine Chaieb\\ \bij f \ (\x. \!y. f y = x)\ (is \?P \ ?Q\) proof assume ?P then have \inj f\ \surj f\ by (simp_all add: bij_def) show ?Q proof fix y from \surj f\ obtain x where \y = f x\ by (auto simp add: surj_def) with \inj f\ show \\!x. f x = y\ by (auto simp add: inj_def) qed next assume ?Q then have \inj f\ by (auto simp add: inj_def) moreover have \\x. y = f x\ for y proof - from \?Q\ obtain x where \f x = y\ by blast then have \y = f x\ by simp then show ?thesis .. qed then have \surj f\ by (auto simp add: surj_def) ultimately show ?P by (rule bijI) qed lemma bij_betw_partition: \bij_betw f A B\ if \bij_betw f (A \ C) (B \ D)\ \bij_betw f C D\ \A \ C = {}\ \B \ D = {}\ proof - from that have \inj_on f (A \ C)\ \inj_on f C\ \f ` (A \ C) = B \ D\ \f ` C = D\ by (simp_all add: bij_betw_def) then have \inj_on f A\ and \f ` (A - C) \ f ` (C - A) = {}\ by (simp_all add: inj_on_Un) with \A \ C = {}\ have \f ` A \ f ` C = {}\ by auto with \f ` (A \ C) = B \ D\ \f ` C = D\ \B \ D = {}\ have \f ` A = B\ by blast with \inj_on f A\ show ?thesis by (simp add: bij_betw_def) qed lemma surj_image_vimage_eq: "surj f \ f ` (f -` A) = A" by simp lemma surj_vimage_empty: assumes "surj f" shows "f -` A = {} \ A = {}" using surj_image_vimage_eq [OF \surj f\, of A] by (intro iffI) fastforce+ lemma inj_vimage_image_eq: "inj f \ f -` (f ` A) = A" unfolding inj_def by blast lemma vimage_subsetD: "surj f \ f -` B \ A \ B \ f ` A" by (blast intro: sym) lemma vimage_subsetI: "inj f \ B \ f ` A \ f -` B \ A" unfolding inj_def by blast lemma vimage_subset_eq: "bij f \ f -` B \ A \ B \ f ` A" unfolding bij_def by (blast del: subsetI intro: vimage_subsetI vimage_subsetD) lemma inj_on_image_eq_iff: "inj_on f C \ A \ C \ B \ C \ f ` A = f ` B \ A = B" by (fastforce simp: inj_on_def) lemma inj_on_Un_image_eq_iff: "inj_on f (A \ B) \ f ` A = f ` B \ A = B" by (erule inj_on_image_eq_iff) simp_all lemma inj_on_image_Int: "inj_on f C \ A \ C \ B \ C \ f ` (A \ B) = f ` A \ f ` B" unfolding inj_on_def by blast lemma inj_on_image_set_diff: "inj_on f C \ A - B \ C \ B \ C \ f ` (A - B) = f ` A - f ` B" unfolding inj_on_def by blast lemma image_Int: "inj f \ f ` (A \ B) = f ` A \ f ` B" unfolding inj_def by blast lemma image_set_diff: "inj f \ f ` (A - B) = f ` A - f ` B" unfolding inj_def by blast lemma inj_on_image_mem_iff: "inj_on f B \ a \ B \ A \ B \ f a \ f ` A \ a \ A" by (auto simp: inj_on_def) lemma inj_image_mem_iff: "inj f \ f a \ f ` A \ a \ A" by (blast dest: injD) lemma inj_image_subset_iff: "inj f \ f ` A \ f ` B \ A \ B" by (blast dest: injD) lemma inj_image_eq_iff: "inj f \ f ` A = f ` B \ A = B" by (blast dest: injD) lemma surj_Compl_image_subset: "surj f \ - (f ` A) \ f ` (- A)" by auto lemma inj_image_Compl_subset: "inj f \ f ` (- A) \ - (f ` A)" by (auto simp: inj_def) lemma bij_image_Compl_eq: "bij f \ f ` (- A) = - (f ` A)" by (simp add: bij_def inj_image_Compl_subset surj_Compl_image_subset equalityI) lemma inj_vimage_singleton: "inj f \ f -` {a} \ {THE x. f x = a}" \ \The inverse image of a singleton under an injective function is included in a singleton.\ by (simp add: inj_def) (blast intro: the_equality [symmetric]) lemma inj_on_vimage_singleton: "inj_on f A \ f -` {a} \ A \ {THE x. x \ A \ f x = a}" by (auto simp add: inj_on_def intro: the_equality [symmetric]) lemma bij_betw_byWitness: assumes left: "\a \ A. f' (f a) = a" and right: "\a' \ A'. f (f' a') = a'" and "f ` A \ A'" and img2: "f' ` A' \ A" shows "bij_betw f A A'" using assms unfolding bij_betw_def inj_on_def proof safe fix a b assume "a \ A" "b \ A" with left have "a = f' (f a) \ b = f' (f b)" by simp moreover assume "f a = f b" ultimately show "a = b" by simp next fix a' assume *: "a' \ A'" with img2 have "f' a' \ A" by blast moreover from * right have "a' = f (f' a')" by simp ultimately show "a' \ f ` A" by blast qed corollary notIn_Un_bij_betw: assumes "b \ A" and "f b \ A'" and "bij_betw f A A'" shows "bij_betw f (A \ {b}) (A' \ {f b})" proof - have "bij_betw f {b} {f b}" unfolding bij_betw_def inj_on_def by simp with assms show ?thesis using bij_betw_combine[of f A A' "{b}" "{f b}"] by blast qed lemma notIn_Un_bij_betw3: assumes "b \ A" and "f b \ A'" shows "bij_betw f A A' = bij_betw f (A \ {b}) (A' \ {f b})" proof assume "bij_betw f A A'" then show "bij_betw f (A \ {b}) (A' \ {f b})" using assms notIn_Un_bij_betw [of b A f A'] by blast next assume *: "bij_betw f (A \ {b}) (A' \ {f b})" have "f ` A = A'" proof safe fix a assume **: "a \ A" then have "f a \ A' \ {f b}" using * unfolding bij_betw_def by blast moreover have False if "f a = f b" proof - have "a = b" using * ** that unfolding bij_betw_def inj_on_def by blast with \b \ A\ ** show ?thesis by blast qed ultimately show "f a \ A'" by blast next fix a' assume **: "a' \ A'" then have "a' \ f ` (A \ {b})" using * by (auto simp add: bij_betw_def) then obtain a where 1: "a \ A \ {b} \ f a = a'" by blast moreover have False if "a = b" using 1 ** \f b \ A'\ that by blast ultimately have "a \ A" by blast with 1 show "a' \ f ` A" by blast qed then show "bij_betw f A A'" using * bij_betw_subset[of f "A \ {b}" _ A] by blast qed lemma inj_on_disjoint_Un: assumes "inj_on f A" and "inj_on g B" and "f ` A \ g ` B = {}" shows "inj_on (\x. if x \ A then f x else g x) (A \ B)" using assms by (simp add: inj_on_def disjoint_iff) (blast) lemma bij_betw_disjoint_Un: assumes "bij_betw f A C" and "bij_betw g B D" and "A \ B = {}" and "C \ D = {}" shows "bij_betw (\x. if x \ A then f x else g x) (A \ B) (C \ D)" using assms by (auto simp: inj_on_disjoint_Un bij_betw_def) lemma involuntory_imp_bij: \bij f\ if \\x. f (f x) = x\ proof (rule bijI) from that show \surj f\ by (rule surjI) show \inj f\ proof (rule injI) fix x y assume \f x = f y\ then have \f (f x) = f (f y)\ by simp then show \x = y\ by (simp add: that) qed qed subsubsection \Inj/surj/bij of Algebraic Operations\ context cancel_semigroup_add begin lemma inj_on_add [simp]: "inj_on ((+) a) A" by (rule inj_onI) simp lemma inj_on_add' [simp]: "inj_on (\b. b + a) A" by (rule inj_onI) simp lemma bij_betw_add [simp]: "bij_betw ((+) a) A B \ (+) a ` A = B" by (simp add: bij_betw_def) end context group_add begin lemma diff_left_imp_eq: "a - b = a - c \ b = c" unfolding add_uminus_conv_diff[symmetric] by(drule local.add_left_imp_eq) simp lemma inj_uminus[simp, intro]: "inj_on uminus A" by (auto intro!: inj_onI) lemma surj_uminus[simp]: "surj uminus" using surjI minus_minus by blast lemma surj_plus [simp]: "surj ((+) a)" proof (standard, simp, standard, simp) fix x have "x = a + (-a + x)" by (simp add: add.assoc) thus "x \ range ((+) a)" by blast qed lemma surj_plus_right [simp]: "surj (\b. b+a)" proof (standard, simp, standard, simp) fix b show "b \ range (\b. b+a)" using diff_add_cancel[of b a, symmetric] by blast qed lemma inj_on_diff_left [simp]: \inj_on ((-) a) A\ by (auto intro: inj_onI dest!: diff_left_imp_eq) lemma inj_on_diff_right [simp]: \inj_on (\b. b - a) A\ by (auto intro: inj_onI simp add: algebra_simps) lemma surj_diff [simp]: "surj ((-) a)" proof (standard, simp, standard, simp) fix x have "x = a - (- x + a)" by (simp add: algebra_simps) thus "x \ range ((-) a)" by blast qed lemma surj_diff_right [simp]: "surj (\x. x - a)" proof (standard, simp, standard, simp) fix x have "x = x + a - a" by simp thus "x \ range (\x. x - a)" by fast qed lemma shows bij_plus: "bij ((+) a)" and bij_plus_right: "bij (\x. x + a)" and bij_uminus: "bij uminus" and bij_diff: "bij ((-) a)" and bij_diff_right: "bij (\x. x - a)" by(simp_all add: bij_def) lemma translation_subtract_Compl: "(\x. x - a) ` (- t) = - ((\x. x - a) ` t)" by(rule bij_image_Compl_eq) (auto simp add: bij_def surj_def inj_def diff_eq_eq intro!: add_diff_cancel[symmetric]) lemma translation_diff: "(+) a ` (s - t) = ((+) a ` s) - ((+) a ` t)" by auto lemma translation_subtract_diff: "(\x. x - a) ` (s - t) = ((\x. x - a) ` s) - ((\x. x - a) ` t)" by(rule image_set_diff)(simp add: inj_on_def diff_eq_eq) lemma translation_Int: "(+) a ` (s \ t) = ((+) a ` s) \ ((+) a ` t)" by auto lemma translation_subtract_Int: "(\x. x - a) ` (s \ t) = ((\x. x - a) ` s) \ ((\x. x - a) ` t)" by(rule image_Int)(simp add: inj_on_def diff_eq_eq) end (* TODO: prove in group_add *) context ab_group_add begin lemma translation_Compl: "(+) a ` (- t) = - ((+) a ` t)" proof (rule set_eqI) fix b show "b \ (+) a ` (- t) \ b \ - (+) a ` t" by (auto simp: image_iff algebra_simps intro!: bexI [of _ "b - a"]) qed end subsection \Function Updating\ definition fun_upd :: "('a \ 'b) \ 'a \ 'b \ ('a \ 'b)" where "fun_upd f a b = (\x. if x = a then b else f x)" nonterminal updbinds and updbind syntax "_updbind" :: "'a \ 'a \ updbind" ("(2_ :=/ _)") "" :: "updbind \ updbinds" ("_") "_updbinds":: "updbind \ updbinds \ updbinds" ("_,/ _") "_Update" :: "'a \ updbinds \ 'a" ("_/'((_)')" [1000, 0] 900) translations "_Update f (_updbinds b bs)" \ "_Update (_Update f b) bs" "f(x:=y)" \ "CONST fun_upd f x y" (* Hint: to define the sum of two functions (or maps), use case_sum. A nice infix syntax could be defined by notation case_sum (infixr "'(+')"80) *) lemma fun_upd_idem_iff: "f(x:=y) = f \ f x = y" unfolding fun_upd_def apply safe apply (erule subst) apply auto done lemma fun_upd_idem: "f x = y \ f(x := y) = f" by (simp only: fun_upd_idem_iff) lemma fun_upd_triv [iff]: "f(x := f x) = f" by (simp only: fun_upd_idem) lemma fun_upd_apply [simp]: "(f(x := y)) z = (if z = x then y else f z)" by (simp add: fun_upd_def) (* fun_upd_apply supersedes these two, but they are useful if fun_upd_apply is intentionally removed from the simpset *) lemma fun_upd_same: "(f(x := y)) x = y" by simp lemma fun_upd_other: "z \ x \ (f(x := y)) z = f z" by simp lemma fun_upd_upd [simp]: "f(x := y, x := z) = f(x := z)" by (simp add: fun_eq_iff) lemma fun_upd_twist: "a \ c \ (m(a := b))(c := d) = (m(c := d))(a := b)" by auto lemma inj_on_fun_updI: "inj_on f A \ y \ f ` A \ inj_on (f(x := y)) A" by (auto simp: inj_on_def) lemma fun_upd_image: "f(x := y) ` A = (if x \ A then insert y (f ` (A - {x})) else f ` A)" by auto lemma fun_upd_comp: "f \ (g(x := y)) = (f \ g)(x := f y)" by auto lemma fun_upd_eqD: "f(x := y) = g(x := z) \ y = z" by (simp add: fun_eq_iff split: if_split_asm) subsection \\override_on\\ definition override_on :: "('a \ 'b) \ ('a \ 'b) \ 'a set \ 'a \ 'b" where "override_on f g A = (\a. if a \ A then g a else f a)" lemma override_on_emptyset[simp]: "override_on f g {} = f" by (simp add: override_on_def) lemma override_on_apply_notin[simp]: "a \ A \ (override_on f g A) a = f a" by (simp add: override_on_def) lemma override_on_apply_in[simp]: "a \ A \ (override_on f g A) a = g a" by (simp add: override_on_def) lemma override_on_insert: "override_on f g (insert x X) = (override_on f g X)(x:=g x)" by (simp add: override_on_def fun_eq_iff) lemma override_on_insert': "override_on f g (insert x X) = (override_on (f(x:=g x)) g X)" by (simp add: override_on_def fun_eq_iff) subsection \Inversion of injective functions\ definition the_inv_into :: "'a set \ ('a \ 'b) \ ('b \ 'a)" where "the_inv_into A f = (\x. THE y. y \ A \ f y = x)" lemma the_inv_into_f_f: "inj_on f A \ x \ A \ the_inv_into A f (f x) = x" unfolding the_inv_into_def inj_on_def by blast lemma f_the_inv_into_f: "inj_on f A \ y \ f ` A \ f (the_inv_into A f y) = y" unfolding the_inv_into_def by (rule the1I2; blast dest: inj_onD) lemma f_the_inv_into_f_bij_betw: "bij_betw f A B \ (bij_betw f A B \ x \ B) \ f (the_inv_into A f x) = x" unfolding bij_betw_def by (blast intro: f_the_inv_into_f) lemma the_inv_into_into: "inj_on f A \ x \ f ` A \ A \ B \ the_inv_into A f x \ B" unfolding the_inv_into_def by (rule the1I2; blast dest: inj_onD) lemma the_inv_into_onto [simp]: "inj_on f A \ the_inv_into A f ` (f ` A) = A" by (fast intro: the_inv_into_into the_inv_into_f_f [symmetric]) lemma the_inv_into_f_eq: "inj_on f A \ f x = y \ x \ A \ the_inv_into A f y = x" by (force simp add: the_inv_into_f_f) lemma the_inv_into_comp: "inj_on f (g ` A) \ inj_on g A \ x \ f ` g ` A \ the_inv_into A (f \ g) x = (the_inv_into A g \ the_inv_into (g ` A) f) x" apply (rule the_inv_into_f_eq) apply (fast intro: comp_inj_on) apply (simp add: f_the_inv_into_f the_inv_into_into) apply (simp add: the_inv_into_into) done lemma inj_on_the_inv_into: "inj_on f A \ inj_on (the_inv_into A f) (f ` A)" by (auto intro: inj_onI simp: the_inv_into_f_f) lemma bij_betw_the_inv_into: "bij_betw f A B \ bij_betw (the_inv_into A f) B A" by (auto simp add: bij_betw_def inj_on_the_inv_into the_inv_into_into) lemma bij_betw_iff_bijections: "bij_betw f A B \ (\g. (\x \ A. f x \ B \ g(f x) = x) \ (\y \ B. g y \ A \ f(g y) = y))" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" by (auto simp: bij_betw_def f_the_inv_into_f the_inv_into_f_f the_inv_into_into exI[where ?x="the_inv_into A f"]) next show "?rhs \ ?lhs" by (force intro: bij_betw_byWitness) qed abbreviation the_inv :: "('a \ 'b) \ ('b \ 'a)" where "the_inv f \ the_inv_into UNIV f" lemma the_inv_f_f: "the_inv f (f x) = x" if "inj f" using that UNIV_I by (rule the_inv_into_f_f) subsection \Monotonicity\ definition monotone_on :: "'a set \ ('a \ 'a \ bool) \ ('b \ 'b \ bool) \ ('a \ 'b) \ bool" where "monotone_on A orda ordb f \ (\x\A. \y\A. orda x y \ ordb (f x) (f y))" abbreviation monotone :: "('a \ 'a \ bool) \ ('b \ 'b \ bool) \ ('a \ 'b) \ bool" where "monotone \ monotone_on UNIV" lemma monotone_def[no_atp]: "monotone orda ordb f \ (\x y. orda x y \ ordb (f x) (f y))" by (simp add: monotone_on_def) text \Lemma @{thm [source] monotone_def} is provided for backward compatibility.\ lemma monotone_onI: "(\x y. x \ A \ y \ A \ orda x y \ ordb (f x) (f y)) \ monotone_on A orda ordb f" by (simp add: monotone_on_def) lemma monotoneI[intro?]: "(\x y. orda x y \ ordb (f x) (f y)) \ monotone orda ordb f" by (rule monotone_onI) lemma monotone_onD: "monotone_on A orda ordb f \ x \ A \ y \ A \ orda x y \ ordb (f x) (f y)" by (simp add: monotone_on_def) lemma monotoneD[dest?]: "monotone orda ordb f \ orda x y \ ordb (f x) (f y)" by (rule monotone_onD[of UNIV, simplified]) lemma monotone_on_subset: "monotone_on A orda ordb f \ B \ A \ monotone_on B orda ordb f" by (auto intro: monotone_onI dest: monotone_onD) lemma monotone_on_empty[simp]: "monotone_on {} orda ordb f" by (auto intro: monotone_onI dest: monotone_onD) lemma monotone_on_o: assumes mono_f: "monotone_on A orda ordb f" and mono_g: "monotone_on B ordc orda g" and "g ` B \ A" shows "monotone_on B ordc ordb (f \ g)" proof (rule monotone_onI) fix x y assume "x \ B" and "y \ B" and "ordc x y" hence "orda (g x) (g y)" by (rule mono_g[THEN monotone_onD]) moreover from \g ` B \ A\ \x \ B\ \y \ B\ have "g x \ A" and "g y \ A" unfolding image_subset_iff by simp_all ultimately show "ordb ((f \ g) x) ((f \ g) y)" using mono_f[THEN monotone_onD] by simp qed subsubsection \Specializations For @{class ord} Type Class And More\ context ord begin abbreviation mono_on :: "'a set \ ('a \ 'b :: ord) \ bool" where "mono_on A \ monotone_on A (\) (\)" abbreviation strict_mono_on :: "'a set \ ('a \ 'b :: ord) \ bool" where "strict_mono_on A \ monotone_on A (<) (<)" lemma mono_on_def[no_atp]: "mono_on A f \ (\r s. r \ A \ s \ A \ r \ s \ f r \ f s)" by (auto simp add: monotone_on_def) lemma strict_mono_on_def[no_atp]: "strict_mono_on A f \ (\r s. r \ A \ s \ A \ r < s \ f r < f s)" by (auto simp add: monotone_on_def) text \Lemmas @{thm [source] mono_on_def} and @{thm [source] strict_mono_on_def} are provided for backward compatibility.\ lemma mono_onI: "(\r s. r \ A \ s \ A \ r \ s \ f r \ f s) \ mono_on A f" by (rule monotone_onI) lemma strict_mono_onI: "(\r s. r \ A \ s \ A \ r < s \ f r < f s) \ strict_mono_on A f" by (rule monotone_onI) lemma mono_onD: "\mono_on A f; r \ A; s \ A; r \ s\ \ f r \ f s" by (rule monotone_onD) lemma strict_mono_onD: "\strict_mono_on A f; r \ A; s \ A; r < s\ \ f r < f s" by (rule monotone_onD) lemma mono_on_subset: "mono_on A f \ B \ A \ mono_on B f" by (rule monotone_on_subset) end lemma mono_on_greaterD: assumes "mono_on A g" "x \ A" "y \ A" "g x > (g (y::_::linorder) :: _ :: linorder)" shows "x > y" proof (rule ccontr) assume "\x > y" hence "x \ y" by (simp add: not_less) from assms(1-3) and this have "g x \ g y" by (rule mono_onD) with assms(4) show False by simp qed context order begin abbreviation mono :: "('a \ 'b::order) \ bool" where "mono \ mono_on UNIV" abbreviation strict_mono :: "('a \ 'b::order) \ bool" where "strict_mono \ strict_mono_on UNIV" abbreviation antimono :: "('a \ 'b::order) \ bool" where "antimono \ monotone (\) (\x y. y \ x)" lemma mono_def[no_atp]: "mono f \ (\x y. x \ y \ f x \ f y)" by (simp add: monotone_on_def) lemma strict_mono_def[no_atp]: "strict_mono f \ (\x y. x < y \ f x < f y)" by (simp add: monotone_on_def) lemma antimono_def[no_atp]: "antimono f \ (\x y. x \ y \ f x \ f y)" by (simp add: monotone_on_def) text \Lemmas @{thm [source] mono_def}, @{thm [source] strict_mono_def}, and @{thm [source] antimono_def} are provided for backward compatibility.\ lemma monoI [intro?]: "(\x y. x \ y \ f x \ f y) \ mono f" by (rule monotoneI) lemma strict_monoI [intro?]: "(\x y. x < y \ f x < f y) \ strict_mono f" by (rule monotoneI) lemma antimonoI [intro?]: "(\x y. x \ y \ f x \ f y) \ antimono f" by (rule monotoneI) lemma monoD [dest?]: "mono f \ x \ y \ f x \ f y" by (rule monotoneD) lemma strict_monoD [dest?]: "strict_mono f \ x < y \ f x < f y" by (rule monotoneD) lemma antimonoD [dest?]: "antimono f \ x \ y \ f x \ f y" by (rule monotoneD) lemma monoE: assumes "mono f" assumes "x \ y" obtains "f x \ f y" proof from assms show "f x \ f y" by (simp add: mono_def) qed lemma antimonoE: fixes f :: "'a \ 'b::order" assumes "antimono f" assumes "x \ y" obtains "f x \ f y" proof from assms show "f x \ f y" by (simp add: antimono_def) qed lemma mono_imp_mono_on: "mono f \ mono_on A f" by (rule monotone_on_subset[OF _ subset_UNIV]) lemma strict_mono_mono [dest?]: assumes "strict_mono f" shows "mono f" proof (rule monoI) fix x y assume "x \ y" show "f x \ f y" proof (cases "x = y") case True then show ?thesis by simp next case False with \x \ y\ have "x < y" by simp with assms strict_monoD have "f x < f y" by auto then show ?thesis by simp qed qed end context linorder begin lemma mono_invE: fixes f :: "'a \ 'b::order" assumes "mono f" assumes "f x < f y" obtains "x \ y" proof show "x \ y" proof (rule ccontr) assume "\ x \ y" then have "y \ x" by simp with \mono f\ obtain "f y \ f x" by (rule monoE) with \f x < f y\ show False by simp qed qed lemma mono_strict_invE: fixes f :: "'a \ 'b::order" assumes "mono f" assumes "f x < f y" obtains "x < y" proof show "x < y" proof (rule ccontr) assume "\ x < y" then have "y \ x" by simp with \mono f\ obtain "f y \ f x" by (rule monoE) with \f x < f y\ show False by simp qed qed lemma strict_mono_eq: assumes "strict_mono f" shows "f x = f y \ x = y" proof assume "f x = f y" show "x = y" proof (cases x y rule: linorder_cases) case less with assms strict_monoD have "f x < f y" by auto with \f x = f y\ show ?thesis by simp next case equal then show ?thesis . next case greater with assms strict_monoD have "f y < f x" by auto with \f x = f y\ show ?thesis by simp qed qed simp lemma strict_mono_less_eq: assumes "strict_mono f" shows "f x \ f y \ x \ y" proof assume "x \ y" with assms strict_mono_mono monoD show "f x \ f y" by auto next assume "f x \ f y" show "x \ y" proof (rule ccontr) assume "\ x \ y" then have "y < x" by simp with assms strict_monoD have "f y < f x" by auto with \f x \ f y\ show False by simp qed qed lemma strict_mono_less: assumes "strict_mono f" shows "f x < f y \ x < y" using assms by (auto simp add: less_le Orderings.less_le strict_mono_eq strict_mono_less_eq) end lemma strict_mono_inv: fixes f :: "('a::linorder) \ ('b::linorder)" assumes "strict_mono f" and "surj f" and inv: "\x. g (f x) = x" shows "strict_mono g" proof fix x y :: 'b assume "x < y" from \surj f\ obtain x' y' where [simp]: "x = f x'" "y = f y'" by blast with \x < y\ and \strict_mono f\ have "x' < y'" by (simp add: strict_mono_less) with inv show "g x < g y" by simp qed lemma strict_mono_on_imp_inj_on: assumes "strict_mono_on A (f :: (_ :: linorder) \ (_ :: preorder))" shows "inj_on f A" proof (rule inj_onI) fix x y assume "x \ A" "y \ A" "f x = f y" thus "x = y" by (cases x y rule: linorder_cases) (auto dest: strict_mono_onD[OF assms, of x y] strict_mono_onD[OF assms, of y x]) qed lemma strict_mono_on_leD: assumes "strict_mono_on A (f :: (_ :: linorder) \ _ :: preorder)" "x \ A" "y \ A" "x \ y" shows "f x \ f y" proof (cases "x = y") case True then show ?thesis by simp next case False with assms have "f x < f y" using strict_mono_onD[OF assms(1)] by simp then show ?thesis by (rule less_imp_le) qed lemma strict_mono_on_eqD: fixes f :: "(_ :: linorder) \ (_ :: preorder)" assumes "strict_mono_on A f" "f x = f y" "x \ A" "y \ A" shows "y = x" using assms by (cases rule: linorder_cases) (auto dest: strict_mono_onD) lemma strict_mono_on_imp_mono_on: "strict_mono_on A (f :: (_ :: linorder) \ _ :: preorder) \ mono_on A f" by (rule mono_onI, rule strict_mono_on_leD) lemma mono_compose: "mono Q \ mono (\i x. Q i (f x))" unfolding mono_def le_fun_def by auto lemma mono_add: fixes a :: "'a::ordered_ab_semigroup_add" shows "mono ((+) a)" by (simp add: add_left_mono monoI) lemma (in semilattice_inf) mono_inf: "mono f \ f (A \ B) \ f A \ f B" for f :: "'a \ 'b::semilattice_inf" by (auto simp add: mono_def intro: Lattices.inf_greatest) lemma (in semilattice_sup) mono_sup: "mono f \ f A \ f B \ f (A \ B)" for f :: "'a \ 'b::semilattice_sup" by (auto simp add: mono_def intro: Lattices.sup_least) lemma (in linorder) min_of_mono: "mono f \ min (f m) (f n) = f (min m n)" by (auto simp: mono_def Orderings.min_def min_def intro: Orderings.antisym) lemma (in linorder) max_of_mono: "mono f \ max (f m) (f n) = f (max m n)" by (auto simp: mono_def Orderings.max_def max_def intro: Orderings.antisym) lemma (in linorder) max_of_antimono: "antimono f \ max (f x) (f y) = f (min x y)" and min_of_antimono: "antimono f \ min (f x) (f y) = f (max x y)" by (auto simp: antimono_def Orderings.max_def max_def Orderings.min_def min_def intro!: antisym) lemma (in linorder) strict_mono_imp_inj_on: "strict_mono f \ inj_on f A" by (auto intro!: inj_onI dest: strict_mono_eq) lemma mono_Int: "mono f \ f (A \ B) \ f A \ f B" by (fact mono_inf) lemma mono_Un: "mono f \ f A \ f B \ f (A \ B)" by (fact mono_sup) subsubsection \Least value operator\ lemma Least_mono: "mono f \ \x\S. \y\S. x \ y \ (LEAST y. y \ f ` S) = f (LEAST x. x \ S)" for f :: "'a::order \ 'b::order" \ \Courtesy of Stephan Merz\ apply clarify apply (erule_tac P = "\x. x \ S" in LeastI2_order) apply fast apply (rule LeastI2_order) apply (auto elim: monoD intro!: order_antisym) done subsection \Setup\ subsubsection \Proof tools\ text \Simplify terms of the form \f(\,x:=y,\,x:=z,\)\ to \f(\,x:=z,\)\\ simproc_setup fun_upd2 ("f(v := w, x := y)") = \fn _ => let fun gen_fun_upd NONE T _ _ = NONE | gen_fun_upd (SOME f) T x y = SOME (Const (\<^const_name>\fun_upd\, T) $ f $ x $ y) fun dest_fun_T1 (Type (_, T :: Ts)) = T fun find_double (t as Const (\<^const_name>\fun_upd\,T) $ f $ x $ y) = let fun find (Const (\<^const_name>\fun_upd\,T) $ g $ v $ w) = if v aconv x then SOME g else gen_fun_upd (find g) T v w | find t = NONE in (dest_fun_T1 T, gen_fun_upd (find f) T x y) end val ss = simpset_of \<^context> fun proc ctxt ct = let val t = Thm.term_of ct in (case find_double t of (T, NONE) => NONE | (T, SOME rhs) => SOME (Goal.prove ctxt [] [] (Logic.mk_equals (t, rhs)) (fn _ => resolve_tac ctxt [eq_reflection] 1 THEN resolve_tac ctxt @{thms ext} 1 THEN simp_tac (put_simpset ss ctxt) 1))) end in proc end \ subsubsection \Functorial structure of types\ ML_file \Tools/functor.ML\ functor map_fun: map_fun by (simp_all add: fun_eq_iff) functor vimage by (simp_all add: fun_eq_iff vimage_comp) text \Legacy theorem names\ lemmas o_def = comp_def lemmas o_apply = comp_apply lemmas o_assoc = comp_assoc [symmetric] lemmas id_o = id_comp lemmas o_id = comp_id lemmas o_eq_dest = comp_eq_dest lemmas o_eq_elim = comp_eq_elim lemmas o_eq_dest_lhs = comp_eq_dest_lhs lemmas o_eq_id_dest = comp_eq_id_dest end diff --git a/src/HOL/Library/Countable_Set.thy b/src/HOL/Library/Countable_Set.thy --- a/src/HOL/Library/Countable_Set.thy +++ b/src/HOL/Library/Countable_Set.thy @@ -1,470 +1,470 @@ (* Title: HOL/Library/Countable_Set.thy Author: Johannes Hölzl Author: Andrei Popescu *) section \Countable sets\ theory Countable_Set imports Countable Infinite_Set begin subsection \Predicate for countable sets\ definition countable :: "'a set \ bool" where "countable S \ (\f::'a \ nat. inj_on f S)" lemma countableE: assumes S: "countable S" obtains f :: "'a \ nat" where "inj_on f S" using S by (auto simp: countable_def) lemma countableI: "inj_on (f::'a \ nat) S \ countable S" by (auto simp: countable_def) lemma countableI': "inj_on (f::'a \ 'b::countable) S \ countable S" using comp_inj_on[of f S to_nat] by (auto intro: countableI) lemma countableE_bij: assumes S: "countable S" obtains f :: "nat \ 'a" and C :: "nat set" where "bij_betw f C S" using S by (blast elim: countableE dest: inj_on_imp_bij_betw bij_betw_inv) lemma countableI_bij: "bij_betw f (C::nat set) S \ countable S" by (blast intro: countableI bij_betw_inv_into bij_betw_imp_inj_on) lemma countable_finite: "finite S \ countable S" by (blast dest: finite_imp_inj_to_nat_seg countableI) lemma countableI_bij1: "bij_betw f A B \ countable A \ countable B" by (blast elim: countableE_bij intro: bij_betw_trans countableI_bij) lemma countableI_bij2: "bij_betw f B A \ countable A \ countable B" by (blast elim: countableE_bij intro: bij_betw_trans bij_betw_inv_into countableI_bij) lemma countable_iff_bij[simp]: "bij_betw f A B \ countable A \ countable B" by (blast intro: countableI_bij1 countableI_bij2) lemma countable_subset: "A \ B \ countable B \ countable A" by (auto simp: countable_def intro: subset_inj_on) lemma countableI_type[intro, simp]: "countable (A:: 'a :: countable set)" using countableI[of to_nat A] by auto subsection \Enumerate a countable set\ lemma countableE_infinite: assumes "countable S" "infinite S" obtains e :: "'a \ nat" where "bij_betw e S UNIV" proof - obtain f :: "'a \ nat" where "inj_on f S" using \countable S\ by (rule countableE) then have "bij_betw f S (f`S)" unfolding bij_betw_def by simp moreover from \inj_on f S\ \infinite S\ have inf_fS: "infinite (f`S)" by (auto dest: finite_imageD) then have "bij_betw (the_inv_into UNIV (enumerate (f`S))) (f`S) UNIV" by (intro bij_betw_the_inv_into bij_enumerate) ultimately have "bij_betw (the_inv_into UNIV (enumerate (f`S)) \ f) S UNIV" by (rule bij_betw_trans) then show thesis .. qed lemma countable_infiniteE': assumes "countable A" "infinite A" obtains g where "bij_betw g (UNIV :: nat set) A" - using bij_betw_inv[OF countableE_infinite[OF assms]] that by blast + by (meson assms bij_betw_inv countableE_infinite) lemma countable_enum_cases: assumes "countable S" obtains (finite) f :: "'a \ nat" where "finite S" "bij_betw f S {.. nat" where "infinite S" "bij_betw f S UNIV" using ex_bij_betw_finite_nat[of S] countableE_infinite \countable S\ by (cases "finite S") (auto simp add: atLeast0LessThan) definition to_nat_on :: "'a set \ 'a \ nat" where "to_nat_on S = (SOME f. if finite S then bij_betw f S {..< card S} else bij_betw f S UNIV)" definition from_nat_into :: "'a set \ nat \ 'a" where "from_nat_into S n = (if n \ to_nat_on S ` S then inv_into S (to_nat_on S) n else SOME s. s\S)" lemma to_nat_on_finite: "finite S \ bij_betw (to_nat_on S) S {..< card S}" using ex_bij_betw_finite_nat unfolding to_nat_on_def by (intro someI2_ex[where Q="\f. bij_betw f S {.. infinite S \ bij_betw (to_nat_on S) S UNIV" using countableE_infinite unfolding to_nat_on_def by (intro someI2_ex[where Q="\f. bij_betw f S UNIV"]) auto lemma bij_betw_from_nat_into_finite: "finite S \ bij_betw (from_nat_into S) {..< card S} S" unfolding from_nat_into_def[abs_def] using to_nat_on_finite[of S] apply (subst bij_betw_cong) apply (split if_split) apply (simp add: bij_betw_def) apply (auto cong: bij_betw_cong intro: bij_betw_inv_into to_nat_on_finite) done lemma bij_betw_from_nat_into: "countable S \ infinite S \ bij_betw (from_nat_into S) UNIV S" unfolding from_nat_into_def[abs_def] using to_nat_on_infinite[of S, unfolded bij_betw_def] by (auto cong: bij_betw_cong intro: bij_betw_inv_into to_nat_on_infinite) text \ The sum/product over the enumeration of a finite set equals simply the sum/product over the set \ context comm_monoid_set begin lemma card_from_nat_into: "F (\i. h (from_nat_into A i)) {..i. h (from_nat_into A i)) {.. 'a" where "A = range f" "inj f" by (metis bij_betw_def bij_betw_from_nat_into [OF assms]) lemma inj_on_to_nat_on[intro]: "countable A \ inj_on (to_nat_on A) A" using to_nat_on_infinite[of A] to_nat_on_finite[of A] by (cases "finite A") (auto simp: bij_betw_def) lemma to_nat_on_inj[simp]: "countable A \ a \ A \ b \ A \ to_nat_on A a = to_nat_on A b \ a = b" using inj_on_to_nat_on[of A] by (auto dest: inj_onD) lemma from_nat_into_to_nat_on[simp]: "countable A \ a \ A \ from_nat_into A (to_nat_on A a) = a" by (auto simp: from_nat_into_def intro!: inv_into_f_f) lemma subset_range_from_nat_into: "countable A \ A \ range (from_nat_into A)" by (auto intro: from_nat_into_to_nat_on[symmetric]) lemma from_nat_into: "A \ {} \ from_nat_into A n \ A" unfolding from_nat_into_def by (metis equals0I inv_into_into someI_ex) lemma range_from_nat_into_subset: "A \ {} \ range (from_nat_into A) \ A" using from_nat_into[of A] by auto lemma range_from_nat_into[simp]: "A \ {} \ countable A \ range (from_nat_into A) = A" by (metis equalityI range_from_nat_into_subset subset_range_from_nat_into) lemma image_to_nat_on: "countable A \ infinite A \ to_nat_on A ` A = UNIV" using to_nat_on_infinite[of A] by (simp add: bij_betw_def) lemma to_nat_on_surj: "countable A \ infinite A \ \a\A. to_nat_on A a = n" by (metis (no_types) image_iff iso_tuple_UNIV_I image_to_nat_on) lemma to_nat_on_from_nat_into[simp]: "n \ to_nat_on A ` A \ to_nat_on A (from_nat_into A n) = n" by (simp add: f_inv_into_f from_nat_into_def) lemma to_nat_on_from_nat_into_infinite[simp]: "countable A \ infinite A \ to_nat_on A (from_nat_into A n) = n" by (metis image_iff to_nat_on_surj to_nat_on_from_nat_into) lemma from_nat_into_inj: "countable A \ m \ to_nat_on A ` A \ n \ to_nat_on A ` A \ from_nat_into A m = from_nat_into A n \ m = n" by (subst to_nat_on_inj[symmetric, of A]) auto lemma from_nat_into_inj_infinite[simp]: "countable A \ infinite A \ from_nat_into A m = from_nat_into A n \ m = n" using image_to_nat_on[of A] from_nat_into_inj[of A m n] by simp lemma eq_from_nat_into_iff: "countable A \ x \ A \ i \ to_nat_on A ` A \ x = from_nat_into A i \ i = to_nat_on A x" by auto lemma from_nat_into_surj: "countable A \ a \ A \ \n. from_nat_into A n = a" by (rule exI[of _ "to_nat_on A a"]) simp lemma from_nat_into_inject[simp]: "A \ {} \ countable A \ B \ {} \ countable B \ from_nat_into A = from_nat_into B \ A = B" by (metis range_from_nat_into) lemma inj_on_from_nat_into: "inj_on from_nat_into ({A. A \ {} \ countable A})" unfolding inj_on_def by auto subsection \Closure properties of countability\ lemma countable_SIGMA[intro, simp]: "countable I \ (\i. i \ I \ countable (A i)) \ countable (SIGMA i : I. A i)" by (intro countableI'[of "\(i, a). (to_nat_on I i, to_nat_on (A i) a)"]) (auto simp: inj_on_def) lemma countable_image[intro, simp]: assumes "countable A" shows "countable (f`A)" proof - obtain g :: "'a \ nat" where "inj_on g A" using assms by (rule countableE) moreover have "inj_on (inv_into A f) (f`A)" "inv_into A f ` f ` A \ A" by (auto intro: inj_on_inv_into inv_into_into) ultimately show ?thesis by (blast dest: comp_inj_on subset_inj_on intro: countableI) qed lemma countable_image_inj_on: "countable (f ` A) \ inj_on f A \ countable A" by (metis countable_image the_inv_into_onto) lemma countable_image_inj_Int_vimage: "\inj_on f S; countable A\ \ countable (S \ f -` A)" by (meson countable_image_inj_on countable_subset image_subset_iff_subset_vimage inf_le2 inj_on_Int) lemma countable_image_inj_gen: "\inj_on f S; countable A\ \ countable {x \ S. f x \ A}" using countable_image_inj_Int_vimage by (auto simp: vimage_def Collect_conj_eq) lemma countable_image_inj_eq: "inj_on f S \ countable(f ` S) \ countable S" using countable_image_inj_on by blast lemma countable_image_inj: "\countable A; inj f\ \ countable {x. f x \ A}" by (metis (mono_tags, lifting) countable_image_inj_eq countable_subset image_Collect_subsetI inj_on_inverseI the_inv_f_f) lemma countable_UN[intro, simp]: fixes I :: "'i set" and A :: "'i => 'a set" assumes I: "countable I" assumes A: "\i. i \ I \ countable (A i)" shows "countable (\i\I. A i)" proof - have "(\i\I. A i) = snd ` (SIGMA i : I. A i)" by (auto simp: image_iff) then show ?thesis by (simp add: assms) qed lemma countable_Un[intro]: "countable A \ countable B \ countable (A \ B)" by (rule countable_UN[of "{True, False}" "\True \ A | False \ B", simplified]) (simp split: bool.split) lemma countable_Un_iff[simp]: "countable (A \ B) \ countable A \ countable B" by (metis countable_Un countable_subset inf_sup_ord(3,4)) lemma countable_Plus[intro, simp]: "countable A \ countable B \ countable (A <+> B)" by (simp add: Plus_def) lemma countable_empty[intro, simp]: "countable {}" by (blast intro: countable_finite) lemma countable_insert[intro, simp]: "countable A \ countable (insert a A)" using countable_Un[of "{a}" A] by (auto simp: countable_finite) lemma countable_Int1[intro, simp]: "countable A \ countable (A \ B)" by (force intro: countable_subset) lemma countable_Int2[intro, simp]: "countable B \ countable (A \ B)" by (blast intro: countable_subset) lemma countable_INT[intro, simp]: "i \ I \ countable (A i) \ countable (\i\I. A i)" by (blast intro: countable_subset) lemma countable_Diff[intro, simp]: "countable A \ countable (A - B)" by (blast intro: countable_subset) lemma countable_insert_eq [simp]: "countable (insert x A) = countable A" by auto (metis Diff_insert_absorb countable_Diff insert_absorb) lemma countable_vimage: "B \ range f \ countable (f -` B) \ countable B" by (metis Int_absorb2 countable_image image_vimage_eq) lemma surj_countable_vimage: "surj f \ countable (f -` B) \ countable B" by (metis countable_vimage top_greatest) lemma countable_Collect[simp]: "countable A \ countable {a \ A. \ a}" by (metis Collect_conj_eq Int_absorb Int_commute Int_def countable_Int1) lemma countable_Image: assumes "\y. y \ Y \ countable (X `` {y})" assumes "countable Y" shows "countable (X `` Y)" proof - have "countable (X `` (\y\Y. {y}))" unfolding Image_UN by (intro countable_UN assms) then show ?thesis by simp qed lemma countable_relpow: fixes X :: "'a rel" assumes Image_X: "\Y. countable Y \ countable (X `` Y)" assumes Y: "countable Y" shows "countable ((X ^^ i) `` Y)" using Y by (induct i arbitrary: Y) (auto simp: relcomp_Image Image_X) lemma countable_funpow: fixes f :: "'a set \ 'a set" assumes "\A. countable A \ countable (f A)" and "countable A" shows "countable ((f ^^ n) A)" by(induction n)(simp_all add: assms) lemma countable_rtrancl: "(\Y. countable Y \ countable (X `` Y)) \ countable Y \ countable (X\<^sup>* `` Y)" unfolding rtrancl_is_UN_relpow UN_Image by (intro countable_UN countableI_type countable_relpow) lemma countable_lists[intro, simp]: assumes A: "countable A" shows "countable (lists A)" proof - have "countable (lists (range (from_nat_into A)))" by (auto simp: lists_image) with A show ?thesis by (auto dest: subset_range_from_nat_into countable_subset lists_mono) qed lemma Collect_finite_eq_lists: "Collect finite = set ` lists UNIV" using finite_list by auto lemma countable_Collect_finite: "countable (Collect (finite::'a::countable set\bool))" by (simp add: Collect_finite_eq_lists) lemma countable_int: "countable \" unfolding Ints_def by auto lemma countable_rat: "countable \" unfolding Rats_def by auto lemma Collect_finite_subset_eq_lists: "{A. finite A \ A \ T} = set ` lists T" using finite_list by (auto simp: lists_eq_set) lemma countable_Collect_finite_subset: "countable T \ countable {A. finite A \ A \ T}" unfolding Collect_finite_subset_eq_lists by auto lemma countable_Fpow: "countable S \ countable (Fpow S)" using countable_Collect_finite_subset by (force simp add: Fpow_def conj_commute) lemma countable_set_option [simp]: "countable (set_option x)" by (cases x) auto subsection \Misc lemmas\ lemma countable_subset_image: "countable B \ B \ (f ` A) \ (\A'. countable A' \ A' \ A \ (B = f ` A'))" (is "?lhs = ?rhs") proof assume ?lhs show ?rhs by (rule exI [where x="inv_into A f ` B"]) (use \?lhs\ in \auto simp: f_inv_into_f subset_iff image_inv_into_cancel inv_into_into\) next assume ?rhs then show ?lhs by force qed lemma ex_subset_image_inj: "(\T. T \ f ` S \ P T) \ (\T. T \ S \ inj_on f T \ P (f ` T))" by (auto simp: subset_image_inj) lemma all_subset_image_inj: "(\T. T \ f ` S \ P T) \ (\T. T \ S \ inj_on f T \ P(f ` T))" by (metis subset_image_inj) lemma ex_countable_subset_image_inj: "(\T. countable T \ T \ f ` S \ P T) \ (\T. countable T \ T \ S \ inj_on f T \ P (f ` T))" by (metis countable_image_inj_eq subset_image_inj) lemma all_countable_subset_image_inj: "(\T. countable T \ T \ f ` S \ P T) \ (\T. countable T \ T \ S \ inj_on f T \P(f ` T))" by (metis countable_image_inj_eq subset_image_inj) lemma ex_countable_subset_image: "(\T. countable T \ T \ f ` S \ P T) \ (\T. countable T \ T \ S \ P (f ` T))" by (metis countable_subset_image) lemma all_countable_subset_image: "(\T. countable T \ T \ f ` S \ P T) \ (\T. countable T \ T \ S \ P(f ` T))" by (metis countable_subset_image) lemma countable_image_eq: "countable(f ` S) \ (\T. countable T \ T \ S \ f ` S = f ` T)" by (metis countable_image countable_image_inj_eq order_refl subset_image_inj) lemma countable_image_eq_inj: "countable(f ` S) \ (\T. countable T \ T \ S \ f ` S = f ` T \ inj_on f T)" by (metis countable_image_inj_eq order_refl subset_image_inj) lemma infinite_countable_subset': assumes X: "infinite X" shows "\C\X. countable C \ infinite C" proof - obtain f :: "nat \ 'a" where "inj f" "range f \ X" using infinite_countable_subset [OF X] by blast then show ?thesis by (intro exI[of _ "range f"]) (auto simp: range_inj_infinite) qed lemma countable_all: assumes S: "countable S" shows "(\s\S. P s) \ (\n::nat. from_nat_into S n \ S \ P (from_nat_into S n))" using S[THEN subset_range_from_nat_into] by auto lemma finite_sequence_to_countable_set: assumes "countable X" obtains F where "\i. F i \ X" "\i. F i \ F (Suc i)" "\i. finite (F i)" "(\i. F i) = X" proof - show thesis apply (rule that[of "\i. if X = {} then {} else from_nat_into X ` {..i}"]) apply (auto simp add: image_iff intro: from_nat_into split: if_splits) using assms from_nat_into_surj by (fastforce cong: image_cong) qed lemma transfer_countable[transfer_rule]: "bi_unique R \ rel_fun (rel_set R) (=) countable countable" by (rule rel_funI, erule (1) bi_unique_rel_set_lemma) (auto dest: countable_image_inj_on) subsection \Uncountable\ abbreviation uncountable where "uncountable A \ \ countable A" lemma uncountable_def: "uncountable A \ A \ {} \ \ (\f::(nat \ 'a). range f = A)" by (auto intro: inj_on_inv_into simp: countable_def) (metis all_not_in_conv inj_on_iff_surj subset_UNIV) lemma uncountable_bij_betw: "bij_betw f A B \ uncountable B \ uncountable A" unfolding bij_betw_def by (metis countable_image) lemma uncountable_infinite: "uncountable A \ infinite A" by (metis countable_finite) lemma uncountable_minus_countable: "uncountable A \ countable B \ uncountable (A - B)" using countable_Un[of B "A - B"] by auto lemma countable_Diff_eq [simp]: "countable (A - {x}) = countable A" by (meson countable_Diff countable_empty countable_insert uncountable_minus_countable) text \Every infinite set can be covered by a pairwise disjoint family of infinite sets. This version doesn't achieve equality, as it only covers a countable subset\ lemma infinite_infinite_partition: assumes "infinite A" obtains C :: "nat \ 'a set" where "pairwise (\i j. disjnt (C i) (C j)) UNIV" "(\i. C i) \ A" "\i. infinite (C i)" proof - obtain f :: "nat\'a" where "range f \ A" "inj f" using assms infinite_countable_subset by blast let ?C = "\i. range (\j. f (prod_encode (i,j)))" show thesis proof show "pairwise (\i j. disjnt (?C i) (?C j)) UNIV" by (auto simp: pairwise_def disjnt_def inj_on_eq_iff [OF \inj f\] inj_on_eq_iff [OF inj_prod_encode, of _ UNIV]) show "(\i. ?C i) \ A" using \range f \ A\ by blast have "infinite (range (\j. f (prod_encode (i, j))))" for i by (rule range_inj_infinite) (meson Pair_inject \inj f\ inj_def prod_encode_eq) then show "\i. infinite (?C i)" using that by auto qed qed end diff --git a/src/HOL/Library/Product_Order.thy b/src/HOL/Library/Product_Order.thy --- a/src/HOL/Library/Product_Order.thy +++ b/src/HOL/Library/Product_Order.thy @@ -1,310 +1,313 @@ (* Title: HOL/Library/Product_Order.thy Author: Brian Huffman *) section \Pointwise order on product types\ theory Product_Order imports Product_Plus begin subsection \Pointwise ordering\ instantiation prod :: (ord, ord) ord begin definition "x \ y \ fst x \ fst y \ snd x \ snd y" definition "(x::'a \ 'b) < y \ x \ y \ \ y \ x" instance .. end lemma fst_mono: "x \ y \ fst x \ fst y" unfolding less_eq_prod_def by simp lemma snd_mono: "x \ y \ snd x \ snd y" unfolding less_eq_prod_def by simp lemma Pair_mono: "x \ x' \ y \ y' \ (x, y) \ (x', y')" unfolding less_eq_prod_def by simp lemma Pair_le [simp]: "(a, b) \ (c, d) \ a \ c \ b \ d" unfolding less_eq_prod_def by simp +lemma atLeastAtMost_prod_eq: "{a..b} = {fst a..fst b} \ {snd a..snd b}" + by (auto simp: less_eq_prod_def) + instance prod :: (preorder, preorder) preorder proof fix x y z :: "'a \ 'b" show "x < y \ x \ y \ \ y \ x" by (rule less_prod_def) show "x \ x" unfolding less_eq_prod_def by fast assume "x \ y" and "y \ z" thus "x \ z" unfolding less_eq_prod_def by (fast elim: order_trans) qed instance prod :: (order, order) order by standard auto subsection \Binary infimum and supremum\ instantiation prod :: (inf, inf) inf begin definition "inf x y = (inf (fst x) (fst y), inf (snd x) (snd y))" lemma inf_Pair_Pair [simp]: "inf (a, b) (c, d) = (inf a c, inf b d)" unfolding inf_prod_def by simp lemma fst_inf [simp]: "fst (inf x y) = inf (fst x) (fst y)" unfolding inf_prod_def by simp lemma snd_inf [simp]: "snd (inf x y) = inf (snd x) (snd y)" unfolding inf_prod_def by simp instance .. end instance prod :: (semilattice_inf, semilattice_inf) semilattice_inf by standard auto instantiation prod :: (sup, sup) sup begin definition "sup x y = (sup (fst x) (fst y), sup (snd x) (snd y))" lemma sup_Pair_Pair [simp]: "sup (a, b) (c, d) = (sup a c, sup b d)" unfolding sup_prod_def by simp lemma fst_sup [simp]: "fst (sup x y) = sup (fst x) (fst y)" unfolding sup_prod_def by simp lemma snd_sup [simp]: "snd (sup x y) = sup (snd x) (snd y)" unfolding sup_prod_def by simp instance .. end instance prod :: (semilattice_sup, semilattice_sup) semilattice_sup by standard auto instance prod :: (lattice, lattice) lattice .. instance prod :: (distrib_lattice, distrib_lattice) distrib_lattice by standard (auto simp add: sup_inf_distrib1) subsection \Top and bottom elements\ instantiation prod :: (top, top) top begin definition "top = (top, top)" instance .. end lemma fst_top [simp]: "fst top = top" unfolding top_prod_def by simp lemma snd_top [simp]: "snd top = top" unfolding top_prod_def by simp lemma Pair_top_top: "(top, top) = top" unfolding top_prod_def by simp instance prod :: (order_top, order_top) order_top by standard (auto simp add: top_prod_def) instantiation prod :: (bot, bot) bot begin definition "bot = (bot, bot)" instance .. end lemma fst_bot [simp]: "fst bot = bot" unfolding bot_prod_def by simp lemma snd_bot [simp]: "snd bot = bot" unfolding bot_prod_def by simp lemma Pair_bot_bot: "(bot, bot) = bot" unfolding bot_prod_def by simp instance prod :: (order_bot, order_bot) order_bot by standard (auto simp add: bot_prod_def) instance prod :: (bounded_lattice, bounded_lattice) bounded_lattice .. instance prod :: (boolean_algebra, boolean_algebra) boolean_algebra by standard (auto simp add: prod_eqI diff_eq) subsection \Complete lattice operations\ instantiation prod :: (Inf, Inf) Inf begin definition "Inf A = (INF x\A. fst x, INF x\A. snd x)" instance .. end instantiation prod :: (Sup, Sup) Sup begin definition "Sup A = (SUP x\A. fst x, SUP x\A. snd x)" instance .. end instance prod :: (conditionally_complete_lattice, conditionally_complete_lattice) conditionally_complete_lattice by standard (force simp: less_eq_prod_def Inf_prod_def Sup_prod_def bdd_below_def bdd_above_def intro!: cInf_lower cSup_upper cInf_greatest cSup_least)+ instance prod :: (complete_lattice, complete_lattice) complete_lattice by standard (simp_all add: less_eq_prod_def Inf_prod_def Sup_prod_def INF_lower SUP_upper le_INF_iff SUP_le_iff bot_prod_def top_prod_def) lemma fst_Inf: "fst (Inf A) = (INF x\A. fst x)" by (simp add: Inf_prod_def) lemma fst_INF: "fst (INF x\A. f x) = (INF x\A. fst (f x))" by (simp add: fst_Inf image_image) lemma fst_Sup: "fst (Sup A) = (SUP x\A. fst x)" by (simp add: Sup_prod_def) lemma fst_SUP: "fst (SUP x\A. f x) = (SUP x\A. fst (f x))" by (simp add: fst_Sup image_image) lemma snd_Inf: "snd (Inf A) = (INF x\A. snd x)" by (simp add: Inf_prod_def) lemma snd_INF: "snd (INF x\A. f x) = (INF x\A. snd (f x))" by (simp add: snd_Inf image_image) lemma snd_Sup: "snd (Sup A) = (SUP x\A. snd x)" by (simp add: Sup_prod_def) lemma snd_SUP: "snd (SUP x\A. f x) = (SUP x\A. snd (f x))" by (simp add: snd_Sup image_image) lemma INF_Pair: "(INF x\A. (f x, g x)) = (INF x\A. f x, INF x\A. g x)" by (simp add: Inf_prod_def image_image) lemma SUP_Pair: "(SUP x\A. (f x, g x)) = (SUP x\A. f x, SUP x\A. g x)" by (simp add: Sup_prod_def image_image) text \Alternative formulations for set infima and suprema over the product of two complete lattices:\ lemma INF_prod_alt_def: \<^marker>\contributor \Alessandro Coglio\\ "Inf (f ` A) = (Inf ((fst \ f) ` A), Inf ((snd \ f) ` A))" by (simp add: Inf_prod_def image_image) lemma SUP_prod_alt_def: \<^marker>\contributor \Alessandro Coglio\\ "Sup (f ` A) = (Sup ((fst \ f) ` A), Sup((snd \ f) ` A))" by (simp add: Sup_prod_def image_image) subsection \Complete distributive lattices\ instance prod :: (complete_distrib_lattice, complete_distrib_lattice) complete_distrib_lattice \<^marker>\contributor \Alessandro Coglio\\ proof fix A::"('a\'b) set set" show "Inf (Sup ` A) \ Sup (Inf ` {f ` A |f. \Y\A. f Y \ Y})" by (simp add: Inf_prod_def Sup_prod_def INF_SUP_set image_image) qed subsection \Bekic's Theorem\ text \ Simultaneous fixed points over pairs can be written in terms of separate fixed points. Transliterated from HOLCF.Fix by Peter Gammie \ lemma lfp_prod: fixes F :: "'a::complete_lattice \ 'b::complete_lattice \ 'a \ 'b" assumes "mono F" shows "lfp F = (lfp (\x. fst (F (x, lfp (\y. snd (F (x, y)))))), (lfp (\y. snd (F (lfp (\x. fst (F (x, lfp (\y. snd (F (x, y)))))), y)))))" (is "lfp F = (?x, ?y)") proof(rule lfp_eqI[OF assms]) have 1: "fst (F (?x, ?y)) = ?x" by (rule trans [symmetric, OF lfp_unfold]) (blast intro!: monoI monoD[OF assms(1)] fst_mono snd_mono Pair_mono lfp_mono)+ have 2: "snd (F (?x, ?y)) = ?y" by (rule trans [symmetric, OF lfp_unfold]) (blast intro!: monoI monoD[OF assms(1)] fst_mono snd_mono Pair_mono lfp_mono)+ from 1 2 show "F (?x, ?y) = (?x, ?y)" by (simp add: prod_eq_iff) next fix z assume F_z: "F z = z" obtain x y where z: "z = (x, y)" by (rule prod.exhaust) from F_z z have F_x: "fst (F (x, y)) = x" by simp from F_z z have F_y: "snd (F (x, y)) = y" by simp let ?y1 = "lfp (\y. snd (F (x, y)))" have "?y1 \ y" by (rule lfp_lowerbound, simp add: F_y) hence "fst (F (x, ?y1)) \ fst (F (x, y))" by (simp add: assms fst_mono monoD) hence "fst (F (x, ?y1)) \ x" using F_x by simp hence 1: "?x \ x" by (simp add: lfp_lowerbound) hence "snd (F (?x, y)) \ snd (F (x, y))" by (simp add: assms snd_mono monoD) hence "snd (F (?x, y)) \ y" using F_y by simp hence 2: "?y \ y" by (simp add: lfp_lowerbound) show "(?x, ?y) \ z" using z 1 2 by simp qed lemma gfp_prod: fixes F :: "'a::complete_lattice \ 'b::complete_lattice \ 'a \ 'b" assumes "mono F" shows "gfp F = (gfp (\x. fst (F (x, gfp (\y. snd (F (x, y)))))), (gfp (\y. snd (F (gfp (\x. fst (F (x, gfp (\y. snd (F (x, y)))))), y)))))" (is "gfp F = (?x, ?y)") proof(rule gfp_eqI[OF assms]) have 1: "fst (F (?x, ?y)) = ?x" by (rule trans [symmetric, OF gfp_unfold]) (blast intro!: monoI monoD[OF assms(1)] fst_mono snd_mono Pair_mono gfp_mono)+ have 2: "snd (F (?x, ?y)) = ?y" by (rule trans [symmetric, OF gfp_unfold]) (blast intro!: monoI monoD[OF assms(1)] fst_mono snd_mono Pair_mono gfp_mono)+ from 1 2 show "F (?x, ?y) = (?x, ?y)" by (simp add: prod_eq_iff) next fix z assume F_z: "F z = z" obtain x y where z: "z = (x, y)" by (rule prod.exhaust) from F_z z have F_x: "fst (F (x, y)) = x" by simp from F_z z have F_y: "snd (F (x, y)) = y" by simp let ?y1 = "gfp (\y. snd (F (x, y)))" have "y \ ?y1" by (rule gfp_upperbound, simp add: F_y) hence "fst (F (x, y)) \ fst (F (x, ?y1))" by (simp add: assms fst_mono monoD) hence "x \ fst (F (x, ?y1))" using F_x by simp hence 1: "x \ ?x" by (simp add: gfp_upperbound) hence "snd (F (x, y)) \ snd (F (?x, y))" by (simp add: assms snd_mono monoD) hence "y \ snd (F (?x, y))" using F_y by simp hence 2: "y \ ?y" by (simp add: gfp_upperbound) show "z \ (?x, ?y)" using z 1 2 by simp qed end diff --git a/src/HOL/Power.thy b/src/HOL/Power.thy --- a/src/HOL/Power.thy +++ b/src/HOL/Power.thy @@ -1,1016 +1,1027 @@ (* Title: HOL/Power.thy Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1997 University of Cambridge *) section \Exponentiation\ theory Power imports Num begin subsection \Powers for Arbitrary Monoids\ class power = one + times begin primrec power :: "'a \ nat \ 'a" (infixr "^" 80) where power_0: "a ^ 0 = 1" | power_Suc: "a ^ Suc n = a * a ^ n" notation (latex output) power ("(_\<^bsup>_\<^esup>)" [1000] 1000) text \Special syntax for squares.\ abbreviation power2 :: "'a \ 'a" ("(_\<^sup>2)" [1000] 999) where "x\<^sup>2 \ x ^ 2" end context includes lifting_syntax begin lemma power_transfer [transfer_rule]: \(R ===> (=) ===> R) (^) (^)\ if [transfer_rule]: \R 1 1\ \(R ===> R ===> R) (*) (*)\ for R :: \'a::power \ 'b::power \ bool\ by (simp only: power_def [abs_def]) transfer_prover end context monoid_mult begin subclass power . lemma power_one [simp]: "1 ^ n = 1" by (induct n) simp_all lemma power_one_right [simp]: "a ^ 1 = a" by simp lemma power_Suc0_right [simp]: "a ^ Suc 0 = a" by simp lemma power_commutes: "a ^ n * a = a * a ^ n" by (induct n) (simp_all add: mult.assoc) lemma power_Suc2: "a ^ Suc n = a ^ n * a" by (simp add: power_commutes) lemma power_add: "a ^ (m + n) = a ^ m * a ^ n" by (induct m) (simp_all add: algebra_simps) lemma power_mult: "a ^ (m * n) = (a ^ m) ^ n" by (induct n) (simp_all add: power_add) lemma power_even_eq: "a ^ (2 * n) = (a ^ n)\<^sup>2" by (subst mult.commute) (simp add: power_mult) lemma power_odd_eq: "a ^ Suc (2*n) = a * (a ^ n)\<^sup>2" by (simp add: power_even_eq) lemma power_numeral_even: "z ^ numeral (Num.Bit0 w) = (let w = z ^ (numeral w) in w * w)" by (simp only: numeral_Bit0 power_add Let_def) lemma power_numeral_odd: "z ^ numeral (Num.Bit1 w) = (let w = z ^ (numeral w) in z * w * w)" by (simp only: numeral_Bit1 One_nat_def add_Suc_right add_0_right power_Suc power_add Let_def mult.assoc) lemma power2_eq_square: "a\<^sup>2 = a * a" by (simp add: numeral_2_eq_2) lemma power3_eq_cube: "a ^ 3 = a * a * a" by (simp add: numeral_3_eq_3 mult.assoc) lemma power4_eq_xxxx: "x^4 = x * x * x * x" by (simp add: mult.assoc power_numeral_even) lemma funpow_times_power: "(times x ^^ f x) = times (x ^ f x)" proof (induct "f x" arbitrary: f) case 0 then show ?case by (simp add: fun_eq_iff) next case (Suc n) define g where "g x = f x - 1" for x with Suc have "n = g x" by simp with Suc have "times x ^^ g x = times (x ^ g x)" by simp moreover from Suc g_def have "f x = g x + 1" by simp ultimately show ?case by (simp add: power_add funpow_add fun_eq_iff mult.assoc) qed lemma power_commuting_commutes: assumes "x * y = y * x" shows "x ^ n * y = y * x ^n" proof (induct n) case 0 then show ?case by simp next case (Suc n) have "x ^ Suc n * y = x ^ n * y * x" by (subst power_Suc2) (simp add: assms ac_simps) also have "\ = y * x ^ Suc n" by (simp only: Suc power_Suc2) (simp add: ac_simps) finally show ?case . qed lemma power_minus_mult: "0 < n \ a ^ (n - 1) * a = a ^ n" by (simp add: power_commutes split: nat_diff_split) lemma left_right_inverse_power: assumes "x * y = 1" shows "x ^ n * y ^ n = 1" proof (induct n) case (Suc n) moreover have "x ^ Suc n * y ^ Suc n = x^n * (x * y) * y^n" by (simp add: power_Suc2[symmetric] mult.assoc[symmetric]) ultimately show ?case by (simp add: assms) qed simp end context comm_monoid_mult begin lemma power_mult_distrib [algebra_simps, algebra_split_simps, field_simps, field_split_simps, divide_simps]: "(a * b) ^ n = (a ^ n) * (b ^ n)" by (induction n) (simp_all add: ac_simps) end text \Extract constant factors from powers.\ declare power_mult_distrib [where a = "numeral w" for w, simp] declare power_mult_distrib [where b = "numeral w" for w, simp] lemma power_add_numeral [simp]: "a^numeral m * a^numeral n = a^numeral (m + n)" for a :: "'a::monoid_mult" by (simp add: power_add [symmetric]) lemma power_add_numeral2 [simp]: "a^numeral m * (a^numeral n * b) = a^numeral (m + n) * b" for a :: "'a::monoid_mult" by (simp add: mult.assoc [symmetric]) lemma power_mult_numeral [simp]: "(a^numeral m)^numeral n = a^numeral (m * n)" for a :: "'a::monoid_mult" by (simp only: numeral_mult power_mult) context semiring_numeral begin lemma numeral_sqr: "numeral (Num.sqr k) = numeral k * numeral k" by (simp only: sqr_conv_mult numeral_mult) lemma numeral_pow: "numeral (Num.pow k l) = numeral k ^ numeral l" by (induct l) (simp_all only: numeral_class.numeral.simps pow.simps numeral_sqr numeral_mult power_add power_one_right) lemma power_numeral [simp]: "numeral k ^ numeral l = numeral (Num.pow k l)" by (rule numeral_pow [symmetric]) end context semiring_1 begin lemma of_nat_power [simp]: "of_nat (m ^ n) = of_nat m ^ n" by (induct n) simp_all lemma zero_power: "0 < n \ 0 ^ n = 0" by (cases n) simp_all lemma power_zero_numeral [simp]: "0 ^ numeral k = 0" by (simp add: numeral_eq_Suc) lemma zero_power2: "0\<^sup>2 = 0" (* delete? *) by (rule power_zero_numeral) lemma one_power2: "1\<^sup>2 = 1" (* delete? *) by (rule power_one) lemma power_0_Suc [simp]: "0 ^ Suc n = 0" by simp text \It looks plausible as a simprule, but its effect can be strange.\ lemma power_0_left: "0 ^ n = (if n = 0 then 1 else 0)" by (cases n) simp_all end context semiring_char_0 begin lemma numeral_power_eq_of_nat_cancel_iff [simp]: "numeral x ^ n = of_nat y \ numeral x ^ n = y" using of_nat_eq_iff by fastforce lemma real_of_nat_eq_numeral_power_cancel_iff [simp]: "of_nat y = numeral x ^ n \ y = numeral x ^ n" using numeral_power_eq_of_nat_cancel_iff [of x n y] by (metis (mono_tags)) lemma of_nat_eq_of_nat_power_cancel_iff[simp]: "(of_nat b) ^ w = of_nat x \ b ^ w = x" by (metis of_nat_power of_nat_eq_iff) lemma of_nat_power_eq_of_nat_cancel_iff[simp]: "of_nat x = (of_nat b) ^ w \ x = b ^ w" by (metis of_nat_eq_of_nat_power_cancel_iff) end context comm_semiring_1 begin text \The divides relation.\ lemma le_imp_power_dvd: assumes "m \ n" shows "a ^ m dvd a ^ n" proof from assms have "a ^ n = a ^ (m + (n - m))" by simp also have "\ = a ^ m * a ^ (n - m)" by (rule power_add) finally show "a ^ n = a ^ m * a ^ (n - m)" . qed lemma power_le_dvd: "a ^ n dvd b \ m \ n \ a ^ m dvd b" by (rule dvd_trans [OF le_imp_power_dvd]) lemma dvd_power_same: "x dvd y \ x ^ n dvd y ^ n" by (induct n) (auto simp add: mult_dvd_mono) lemma dvd_power_le: "x dvd y \ m \ n \ x ^ n dvd y ^ m" by (rule power_le_dvd [OF dvd_power_same]) lemma dvd_power [simp]: fixes n :: nat assumes "n > 0 \ x = 1" shows "x dvd (x ^ n)" using assms proof assume "0 < n" then have "x ^ n = x ^ Suc (n - 1)" by simp then show "x dvd (x ^ n)" by simp next assume "x = 1" then show "x dvd (x ^ n)" by simp qed end context semiring_1_no_zero_divisors begin subclass power . lemma power_eq_0_iff [simp]: "a ^ n = 0 \ a = 0 \ n > 0" by (induct n) auto lemma power_not_zero: "a \ 0 \ a ^ n \ 0" by (induct n) auto lemma zero_eq_power2 [simp]: "a\<^sup>2 = 0 \ a = 0" unfolding power2_eq_square by simp end context ring_1 begin lemma power_minus: "(- a) ^ n = (- 1) ^ n * a ^ n" proof (induct n) case 0 show ?case by simp next case (Suc n) then show ?case by (simp del: power_Suc add: power_Suc2 mult.assoc) qed lemma power_minus': "NO_MATCH 1 x \ (-x) ^ n = (-1)^n * x ^ n" by (rule power_minus) lemma power_minus_Bit0: "(- x) ^ numeral (Num.Bit0 k) = x ^ numeral (Num.Bit0 k)" by (induct k, simp_all only: numeral_class.numeral.simps power_add power_one_right mult_minus_left mult_minus_right minus_minus) lemma power_minus_Bit1: "(- x) ^ numeral (Num.Bit1 k) = - (x ^ numeral (Num.Bit1 k))" by (simp only: eval_nat_numeral(3) power_Suc power_minus_Bit0 mult_minus_left) lemma power2_minus [simp]: "(- a)\<^sup>2 = a\<^sup>2" by (fact power_minus_Bit0) lemma power_minus1_even [simp]: "(- 1) ^ (2*n) = 1" proof (induct n) case 0 show ?case by simp next case (Suc n) then show ?case by (simp add: power_add power2_eq_square) qed lemma power_minus1_odd: "(- 1) ^ Suc (2*n) = -1" by simp lemma power_minus_even [simp]: "(-a) ^ (2*n) = a ^ (2*n)" by (simp add: power_minus [of a]) end context ring_1_no_zero_divisors begin lemma power2_eq_1_iff: "a\<^sup>2 = 1 \ a = 1 \ a = - 1" using square_eq_1_iff [of a] by (simp add: power2_eq_square) end context idom begin lemma power2_eq_iff: "x\<^sup>2 = y\<^sup>2 \ x = y \ x = - y" unfolding power2_eq_square by (rule square_eq_iff) end context semidom_divide begin lemma power_diff: "a ^ (m - n) = (a ^ m) div (a ^ n)" if "a \ 0" and "n \ m" proof - define q where "q = m - n" with \n \ m\ have "m = q + n" by simp with \a \ 0\ q_def show ?thesis by (simp add: power_add) qed end context algebraic_semidom begin lemma div_power: "b dvd a \ (a div b) ^ n = a ^ n div b ^ n" by (induct n) (simp_all add: div_mult_div_if_dvd dvd_power_same) lemma is_unit_power_iff: "is_unit (a ^ n) \ is_unit a \ n = 0" by (induct n) (auto simp add: is_unit_mult_iff) lemma dvd_power_iff: assumes "x \ 0" shows "x ^ m dvd x ^ n \ is_unit x \ m \ n" proof assume *: "x ^ m dvd x ^ n" { assume "m > n" note * also have "x ^ n = x ^ n * 1" by simp also from \m > n\ have "m = n + (m - n)" by simp also have "x ^ \ = x ^ n * x ^ (m - n)" by (rule power_add) finally have "x ^ (m - n) dvd 1" using assms by (subst (asm) dvd_times_left_cancel_iff) simp_all with \m > n\ have "is_unit x" by (simp add: is_unit_power_iff) } thus "is_unit x \ m \ n" by force qed (auto intro: unit_imp_dvd simp: is_unit_power_iff le_imp_power_dvd) end context normalization_semidom_multiplicative begin lemma normalize_power: "normalize (a ^ n) = normalize a ^ n" by (induct n) (simp_all add: normalize_mult) lemma unit_factor_power: "unit_factor (a ^ n) = unit_factor a ^ n" by (induct n) (simp_all add: unit_factor_mult) end context division_ring begin text \Perhaps these should be simprules.\ lemma power_inverse [field_simps, field_split_simps, divide_simps]: "inverse a ^ n = inverse (a ^ n)" proof (cases "a = 0") case True then show ?thesis by (simp add: power_0_left) next case False then have "inverse (a ^ n) = inverse a ^ n" by (induct n) (simp_all add: nonzero_inverse_mult_distrib power_commutes) then show ?thesis by simp qed lemma power_one_over [field_simps, field_split_simps, divide_simps]: "(1 / a) ^ n = 1 / a ^ n" using power_inverse [of a] by (simp add: divide_inverse) end context field begin lemma power_divide [field_simps, field_split_simps, divide_simps]: "(a / b) ^ n = a ^ n / b ^ n" by (induct n) simp_all end subsection \Exponentiation on ordered types\ context linordered_semidom begin lemma zero_less_power [simp]: "0 < a \ 0 < a ^ n" by (induct n) simp_all lemma zero_le_power [simp]: "0 \ a \ 0 \ a ^ n" by (induct n) simp_all lemma power_mono: "a \ b \ 0 \ a \ a ^ n \ b ^ n" by (induct n) (auto intro: mult_mono order_trans [of 0 a b]) lemma one_le_power [simp]: "1 \ a \ 1 \ a ^ n" using power_mono [of 1 a n] by simp lemma power_le_one: "0 \ a \ a \ 1 \ a ^ n \ 1" using power_mono [of a 1 n] by simp lemma power_gt1_lemma: assumes gt1: "1 < a" shows "1 < a * a ^ n" proof - from gt1 have "0 \ a" by (fact order_trans [OF zero_le_one less_imp_le]) from gt1 have "1 * 1 < a * 1" by simp also from gt1 have "\ \ a * a ^ n" by (simp only: mult_mono \0 \ a\ one_le_power order_less_imp_le zero_le_one order_refl) finally show ?thesis by simp qed lemma power_gt1: "1 < a \ 1 < a ^ Suc n" by (simp add: power_gt1_lemma) lemma one_less_power [simp]: "1 < a \ 0 < n \ 1 < a ^ n" by (cases n) (simp_all add: power_gt1_lemma) lemma power_le_imp_le_exp: assumes gt1: "1 < a" shows "a ^ m \ a ^ n \ m \ n" proof (induct m arbitrary: n) case 0 show ?case by simp next case (Suc m) show ?case proof (cases n) case 0 with Suc have "a * a ^ m \ 1" by simp with gt1 show ?thesis by (force simp only: power_gt1_lemma not_less [symmetric]) next case (Suc n) with Suc.prems Suc.hyps show ?thesis by (force dest: mult_left_le_imp_le simp add: less_trans [OF zero_less_one gt1]) qed qed lemma of_nat_zero_less_power_iff [simp]: "of_nat x ^ n > 0 \ x > 0 \ n = 0" by (induct n) auto text \Surely we can strengthen this? It holds for \0 too.\ lemma power_inject_exp [simp]: \a ^ m = a ^ n \ m = n\ if \1 < a\ using that by (force simp add: order_class.order.antisym power_le_imp_le_exp) text \ Can relax the first premise to \<^term>\0 in the case of the natural numbers. \ lemma power_less_imp_less_exp: "1 < a \ a ^ m < a ^ n \ m < n" by (simp add: order_less_le [of m n] less_le [of "a^m" "a^n"] power_le_imp_le_exp) lemma power_strict_mono: "a < b \ 0 \ a \ 0 < n \ a ^ n < b ^ n" proof (induct n) case 0 then show ?case by simp next case (Suc n) then show ?case by (cases "n = 0") (auto simp: mult_strict_mono le_less_trans [of 0 a b]) qed lemma power_mono_iff [simp]: shows "\a \ 0; b \ 0; n>0\ \ a ^ n \ b ^ n \ a \ b" using power_mono [of a b] power_strict_mono [of b a] not_le by auto text\Lemma for \power_strict_decreasing\\ lemma power_Suc_less: "0 < a \ a < 1 \ a * a ^ n < a ^ n" by (induct n) (auto simp: mult_strict_left_mono) lemma power_strict_decreasing: "n < N \ 0 < a \ a < 1 \ a ^ N < a ^ n" proof (induction N) case 0 then show ?case by simp next case (Suc N) then show ?case using mult_strict_mono[of a 1 "a ^ N" "a ^ n"] by (auto simp add: power_Suc_less less_Suc_eq) qed text \Proof resembles that of \power_strict_decreasing\.\ lemma power_decreasing: "n \ N \ 0 \ a \ a \ 1 \ a ^ N \ a ^ n" proof (induction N) case 0 then show ?case by simp next case (Suc N) then show ?case using mult_mono[of a 1 "a^N" "a ^ n"] by (auto simp add: le_Suc_eq) qed lemma power_decreasing_iff [simp]: "\0 < b; b < 1\ \ b ^ m \ b ^ n \ n \ m" using power_strict_decreasing [of m n b] by (auto intro: power_decreasing ccontr) lemma power_strict_decreasing_iff [simp]: "\0 < b; b < 1\ \ b ^ m < b ^ n \ n < m" using power_decreasing_iff [of b m n] unfolding le_less by (auto dest: power_strict_decreasing le_neq_implies_less) lemma power_Suc_less_one: "0 < a \ a < 1 \ a ^ Suc n < 1" using power_strict_decreasing [of 0 "Suc n" a] by simp text \Proof again resembles that of \power_strict_decreasing\.\ lemma power_increasing: "n \ N \ 1 \ a \ a ^ n \ a ^ N" proof (induct N) case 0 then show ?case by simp next case (Suc N) then show ?case using mult_mono[of 1 a "a ^ n" "a ^ N"] by (auto simp add: le_Suc_eq order_trans [OF zero_le_one]) qed text \Lemma for \power_strict_increasing\.\ lemma power_less_power_Suc: "1 < a \ a ^ n < a * a ^ n" by (induct n) (auto simp: mult_strict_left_mono less_trans [OF zero_less_one]) lemma power_strict_increasing: "n < N \ 1 < a \ a ^ n < a ^ N" proof (induct N) case 0 then show ?case by simp next case (Suc N) then show ?case using mult_strict_mono[of 1 a "a^n" "a^N"] by (auto simp add: power_less_power_Suc less_Suc_eq less_trans [OF zero_less_one] less_imp_le) qed lemma power_increasing_iff [simp]: "1 < b \ b ^ x \ b ^ y \ x \ y" by (blast intro: power_le_imp_le_exp power_increasing less_imp_le) lemma power_strict_increasing_iff [simp]: "1 < b \ b ^ x < b ^ y \ x < y" by (blast intro: power_less_imp_less_exp power_strict_increasing) lemma power_le_imp_le_base: assumes le: "a ^ Suc n \ b ^ Suc n" and "0 \ b" shows "a \ b" proof (rule ccontr) assume "\ ?thesis" then have "b < a" by (simp only: linorder_not_le) then have "b ^ Suc n < a ^ Suc n" by (simp only: assms(2) power_strict_mono) with le show False by (simp add: linorder_not_less [symmetric]) qed lemma power_less_imp_less_base: assumes less: "a ^ n < b ^ n" assumes nonneg: "0 \ b" shows "a < b" proof (rule contrapos_pp [OF less]) assume "\ ?thesis" then have "b \ a" by (simp only: linorder_not_less) from this nonneg have "b ^ n \ a ^ n" by (rule power_mono) then show "\ a ^ n < b ^ n" by (simp only: linorder_not_less) qed lemma power_inject_base: "a ^ Suc n = b ^ Suc n \ 0 \ a \ 0 \ b \ a = b" by (blast intro: power_le_imp_le_base order.antisym eq_refl sym) lemma power_eq_imp_eq_base: "a ^ n = b ^ n \ 0 \ a \ 0 \ b \ 0 < n \ a = b" by (cases n) (simp_all del: power_Suc, rule power_inject_base) lemma power_eq_iff_eq_base: "0 < n \ 0 \ a \ 0 \ b \ a ^ n = b ^ n \ a = b" using power_eq_imp_eq_base [of a n b] by auto lemma power2_le_imp_le: "x\<^sup>2 \ y\<^sup>2 \ 0 \ y \ x \ y" unfolding numeral_2_eq_2 by (rule power_le_imp_le_base) lemma power2_less_imp_less: "x\<^sup>2 < y\<^sup>2 \ 0 \ y \ x < y" by (rule power_less_imp_less_base) lemma power2_eq_imp_eq: "x\<^sup>2 = y\<^sup>2 \ 0 \ x \ 0 \ y \ x = y" unfolding numeral_2_eq_2 by (erule (2) power_eq_imp_eq_base) simp lemma power_Suc_le_self: "0 \ a \ a \ 1 \ a ^ Suc n \ a" using power_decreasing [of 1 "Suc n" a] by simp lemma power2_eq_iff_nonneg [simp]: assumes "0 \ x" "0 \ y" shows "(x ^ 2 = y ^ 2) \ x = y" using assms power2_eq_imp_eq by blast lemma of_nat_less_numeral_power_cancel_iff[simp]: "of_nat x < numeral i ^ n \ x < numeral i ^ n" using of_nat_less_iff[of x "numeral i ^ n", unfolded of_nat_numeral of_nat_power] . lemma of_nat_le_numeral_power_cancel_iff[simp]: "of_nat x \ numeral i ^ n \ x \ numeral i ^ n" using of_nat_le_iff[of x "numeral i ^ n", unfolded of_nat_numeral of_nat_power] . lemma numeral_power_less_of_nat_cancel_iff[simp]: "numeral i ^ n < of_nat x \ numeral i ^ n < x" using of_nat_less_iff[of "numeral i ^ n" x, unfolded of_nat_numeral of_nat_power] . lemma numeral_power_le_of_nat_cancel_iff[simp]: "numeral i ^ n \ of_nat x \ numeral i ^ n \ x" using of_nat_le_iff[of "numeral i ^ n" x, unfolded of_nat_numeral of_nat_power] . lemma of_nat_le_of_nat_power_cancel_iff[simp]: "(of_nat b) ^ w \ of_nat x \ b ^ w \ x" by (metis of_nat_le_iff of_nat_power) lemma of_nat_power_le_of_nat_cancel_iff[simp]: "of_nat x \ (of_nat b) ^ w \ x \ b ^ w" by (metis of_nat_le_iff of_nat_power) lemma of_nat_less_of_nat_power_cancel_iff[simp]: "(of_nat b) ^ w < of_nat x \ b ^ w < x" by (metis of_nat_less_iff of_nat_power) lemma of_nat_power_less_of_nat_cancel_iff[simp]: "of_nat x < (of_nat b) ^ w \ x < b ^ w" by (metis of_nat_less_iff of_nat_power) +lemma power2_nonneg_ge_1_iff: + assumes "x \ 0" + shows "x ^ 2 \ 1 \ x \ 1" + using assms by (auto intro: power2_le_imp_le) + +lemma power2_nonneg_gt_1_iff: + assumes "x \ 0" + shows "x ^ 2 > 1 \ x > 1" + using assms by (auto intro: power_less_imp_less_base) + end - text \Some @{typ nat}-specific lemmas:\ lemma mono_ge2_power_minus_self: assumes "k \ 2" shows "mono (\m. k ^ m - m)" unfolding mono_iff_le_Suc proof fix n have "k ^ n < k ^ Suc n" using power_strict_increasing_iff[of k "n" "Suc n"] assms by linarith thus "k ^ n - n \ k ^ Suc n - Suc n" by linarith qed lemma self_le_ge2_pow[simp]: assumes "k \ 2" shows "m \ k ^ m" proof (induction m) case 0 show ?case by simp next case (Suc m) hence "Suc m \ Suc (k ^ m)" by simp also have "... \ k^m + k^m" using one_le_power[of k m] assms by linarith also have "... \ k * k^m" by (metis mult_2 mult_le_mono1[OF assms]) finally show ?case by simp qed lemma diff_le_diff_pow[simp]: assumes "k \ 2" shows "m - n \ k ^ m - k ^ n" proof (cases "n \ m") case True thus ?thesis using monoD[OF mono_ge2_power_minus_self[OF assms] True] self_le_ge2_pow[OF assms, of m] by (simp add: le_diff_conv le_diff_conv2) qed auto context linordered_ring_strict begin lemma sum_squares_eq_zero_iff: "x * x + y * y = 0 \ x = 0 \ y = 0" by (simp add: add_nonneg_eq_0_iff) lemma sum_squares_le_zero_iff: "x * x + y * y \ 0 \ x = 0 \ y = 0" by (simp add: le_less not_sum_squares_lt_zero sum_squares_eq_zero_iff) lemma sum_squares_gt_zero_iff: "0 < x * x + y * y \ x \ 0 \ y \ 0" by (simp add: not_le [symmetric] sum_squares_le_zero_iff) end context linordered_idom begin lemma zero_le_power2 [simp]: "0 \ a\<^sup>2" by (simp add: power2_eq_square) lemma zero_less_power2 [simp]: "0 < a\<^sup>2 \ a \ 0" by (force simp add: power2_eq_square zero_less_mult_iff linorder_neq_iff) lemma power2_less_0 [simp]: "\ a\<^sup>2 < 0" by (force simp add: power2_eq_square mult_less_0_iff) lemma power_abs: "\a ^ n\ = \a\ ^ n" \ \FIXME simp?\ by (induct n) (simp_all add: abs_mult) lemma power_sgn [simp]: "sgn (a ^ n) = sgn a ^ n" by (induct n) (simp_all add: sgn_mult) lemma abs_power_minus [simp]: "\(- a) ^ n\ = \a ^ n\" by (simp add: power_abs) lemma zero_less_power_abs_iff [simp]: "0 < \a\ ^ n \ a \ 0 \ n = 0" proof (induct n) case 0 show ?case by simp next case Suc then show ?case by (auto simp: zero_less_mult_iff) qed lemma zero_le_power_abs [simp]: "0 \ \a\ ^ n" by (rule zero_le_power [OF abs_ge_zero]) lemma power2_less_eq_zero_iff [simp]: "a\<^sup>2 \ 0 \ a = 0" by (simp add: le_less) lemma abs_power2 [simp]: "\a\<^sup>2\ = a\<^sup>2" by (simp add: power2_eq_square) lemma power2_abs [simp]: "\a\\<^sup>2 = a\<^sup>2" by (simp add: power2_eq_square) lemma odd_power_less_zero: "a < 0 \ a ^ Suc (2 * n) < 0" proof (induct n) case 0 then show ?case by simp next case (Suc n) have "a ^ Suc (2 * Suc n) = (a*a) * a ^ Suc(2*n)" by (simp add: ac_simps power_add power2_eq_square) then show ?case by (simp del: power_Suc add: Suc mult_less_0_iff mult_neg_neg) qed lemma odd_0_le_power_imp_0_le: "0 \ a ^ Suc (2 * n) \ 0 \ a" using odd_power_less_zero [of a n] by (force simp add: linorder_not_less [symmetric]) lemma zero_le_even_power'[simp]: "0 \ a ^ (2 * n)" proof (induct n) case 0 show ?case by simp next case (Suc n) have "a ^ (2 * Suc n) = (a*a) * a ^ (2*n)" by (simp add: ac_simps power_add power2_eq_square) then show ?case by (simp add: Suc zero_le_mult_iff) qed lemma sum_power2_ge_zero: "0 \ x\<^sup>2 + y\<^sup>2" by (intro add_nonneg_nonneg zero_le_power2) lemma not_sum_power2_lt_zero: "\ x\<^sup>2 + y\<^sup>2 < 0" unfolding not_less by (rule sum_power2_ge_zero) lemma sum_power2_eq_zero_iff: "x\<^sup>2 + y\<^sup>2 = 0 \ x = 0 \ y = 0" unfolding power2_eq_square by (simp add: add_nonneg_eq_0_iff) lemma sum_power2_le_zero_iff: "x\<^sup>2 + y\<^sup>2 \ 0 \ x = 0 \ y = 0" by (simp add: le_less sum_power2_eq_zero_iff not_sum_power2_lt_zero) lemma sum_power2_gt_zero_iff: "0 < x\<^sup>2 + y\<^sup>2 \ x \ 0 \ y \ 0" unfolding not_le [symmetric] by (simp add: sum_power2_le_zero_iff) lemma abs_le_square_iff: "\x\ \ \y\ \ x\<^sup>2 \ y\<^sup>2" (is "?lhs \ ?rhs") proof assume ?lhs then have "\x\\<^sup>2 \ \y\\<^sup>2" by (rule power_mono) simp then show ?rhs by simp next assume ?rhs then show ?lhs by (auto intro!: power2_le_imp_le [OF _ abs_ge_zero]) qed lemma power2_le_iff_abs_le: "y \ 0 \ x\<^sup>2 \ y\<^sup>2 \ \x\ \ y" by (metis abs_le_square_iff abs_of_nonneg) lemma abs_square_le_1:"x\<^sup>2 \ 1 \ \x\ \ 1" using abs_le_square_iff [of x 1] by simp lemma abs_square_eq_1: "x\<^sup>2 = 1 \ \x\ = 1" by (auto simp add: abs_if power2_eq_1_iff) lemma abs_square_less_1: "x\<^sup>2 < 1 \ \x\ < 1" using abs_square_eq_1 [of x] abs_square_le_1 [of x] by (auto simp add: le_less) lemma square_le_1: assumes "- 1 \ x" "x \ 1" shows "x\<^sup>2 \ 1" using assms by (metis add.inverse_inverse linear mult_le_one neg_equal_0_iff_equal neg_le_iff_le power2_eq_square power_minus_Bit0) end - subsection \Miscellaneous rules\ lemma (in linordered_semidom) self_le_power: "1 \ a \ 0 < n \ a \ a ^ n" using power_increasing [of 1 n a] power_one_right [of a] by auto +lemma power2_ge_1_iff: "x ^ 2 \ 1 \ x \ 1 \ x \ (-1 :: 'a :: linordered_idom)" + using abs_le_square_iff[of 1 x] by (auto simp: abs_if split: if_splits) + lemma (in power) power_eq_if: "p ^ m = (if m=0 then 1 else p * (p ^ (m - 1)))" unfolding One_nat_def by (cases m) simp_all lemma (in comm_semiring_1) power2_sum: "(x + y)\<^sup>2 = x\<^sup>2 + y\<^sup>2 + 2 * x * y" by (simp add: algebra_simps power2_eq_square mult_2_right) context comm_ring_1 begin lemma power2_diff: "(x - y)\<^sup>2 = x\<^sup>2 + y\<^sup>2 - 2 * x * y" by (simp add: algebra_simps power2_eq_square mult_2_right) lemma power2_commute: "(x - y)\<^sup>2 = (y - x)\<^sup>2" by (simp add: algebra_simps power2_eq_square) lemma minus_power_mult_self: "(- a) ^ n * (- a) ^ n = a ^ (2 * n)" by (simp add: power_mult_distrib [symmetric]) (simp add: power2_eq_square [symmetric] power_mult [symmetric]) lemma minus_one_mult_self [simp]: "(- 1) ^ n * (- 1) ^ n = 1" using minus_power_mult_self [of 1 n] by simp lemma left_minus_one_mult_self [simp]: "(- 1) ^ n * ((- 1) ^ n * a) = a" by (simp add: mult.assoc [symmetric]) end text \Simprules for comparisons where common factors can be cancelled.\ lemmas zero_compare_simps = add_strict_increasing add_strict_increasing2 add_increasing zero_le_mult_iff zero_le_divide_iff zero_less_mult_iff zero_less_divide_iff mult_le_0_iff divide_le_0_iff mult_less_0_iff divide_less_0_iff zero_le_power2 power2_less_0 subsection \Exponentiation for the Natural Numbers\ lemma nat_one_le_power [simp]: "Suc 0 \ i \ Suc 0 \ i ^ n" by (rule one_le_power [of i n, unfolded One_nat_def]) lemma nat_zero_less_power_iff [simp]: "x ^ n > 0 \ x > 0 \ n = 0" for x :: nat by (induct n) auto lemma nat_power_eq_Suc_0_iff [simp]: "x ^ m = Suc 0 \ m = 0 \ x = Suc 0" by (induct m) auto lemma power_Suc_0 [simp]: "Suc 0 ^ n = Suc 0" by simp text \ Valid for the naturals, but what if \0 < i < 1\? Premises cannot be weakened: consider the case where \i = 0\, \m = 1\ and \n = 0\. \ lemma nat_power_less_imp_less: fixes i :: nat assumes nonneg: "0 < i" assumes less: "i ^ m < i ^ n" shows "m < n" proof (cases "i = 1") case True with less power_one [where 'a = nat] show ?thesis by simp next case False with nonneg have "1 < i" by auto from power_strict_increasing_iff [OF this] less show ?thesis .. qed lemma power_gt_expt: "n > Suc 0 \ n^k > k" by (induction k) (auto simp: less_trans_Suc n_less_m_mult_n) lemma less_exp [simp]: \n < 2 ^ n\ by (simp add: power_gt_expt) lemma power_dvd_imp_le: fixes i :: nat assumes "i ^ m dvd i ^ n" "1 < i" shows "m \ n" using assms by (auto intro: power_le_imp_le_exp [OF \1 < i\ dvd_imp_le]) lemma dvd_power_iff_le: fixes k::nat shows "2 \ k \ ((k ^ m) dvd (k ^ n) \ m \ n)" using le_imp_power_dvd power_dvd_imp_le by force lemma power2_nat_le_eq_le: "m\<^sup>2 \ n\<^sup>2 \ m \ n" for m n :: nat by (auto intro: power2_le_imp_le power_mono) lemma power2_nat_le_imp_le: fixes m n :: nat assumes "m\<^sup>2 \ n" shows "m \ n" proof (cases m) case 0 then show ?thesis by simp next case (Suc k) show ?thesis proof (rule ccontr) assume "\ ?thesis" then have "n < m" by simp with assms Suc show False by (simp add: power2_eq_square) qed qed lemma ex_power_ivl1: fixes b k :: nat assumes "b \ 2" shows "k \ 1 \ \n. b^n \ k \ k < b^(n+1)" (is "_ \ \n. ?P k n") proof(induction k) case 0 thus ?case by simp next case (Suc k) show ?case proof cases assume "k=0" hence "?P (Suc k) 0" using assms by simp thus ?case .. next assume "k\0" with Suc obtain n where IH: "?P k n" by auto show ?case proof (cases "k = b^(n+1) - 1") case True hence "?P (Suc k) (n+1)" using assms by (simp add: power_less_power_Suc) thus ?thesis .. next case False hence "?P (Suc k) n" using IH by auto thus ?thesis .. qed qed qed lemma ex_power_ivl2: fixes b k :: nat assumes "b \ 2" "k \ 2" shows "\n. b^n < k \ k \ b^(n+1)" proof - have "1 \ k - 1" using assms(2) by arith from ex_power_ivl1[OF assms(1) this] obtain n where "b ^ n \ k - 1 \ k - 1 < b ^ (n + 1)" .. hence "b^n < k \ k \ b^(n+1)" using assms by auto thus ?thesis .. qed subsubsection \Cardinality of the Powerset\ lemma card_UNIV_bool [simp]: "card (UNIV :: bool set) = 2" unfolding UNIV_bool by simp lemma card_Pow: "finite A \ card (Pow A) = 2 ^ card A" proof (induct rule: finite_induct) case empty show ?case by simp next case (insert x A) from \x \ A\ have disjoint: "Pow A \ insert x ` Pow A = {}" by blast from \x \ A\ have inj_on: "inj_on (insert x) (Pow A)" unfolding inj_on_def by auto have "card (Pow (insert x A)) = card (Pow A \ insert x ` Pow A)" by (simp only: Pow_insert) also have "\ = card (Pow A) + card (insert x ` Pow A)" by (rule card_Un_disjoint) (use \finite A\ disjoint in simp_all) also from inj_on have "card (insert x ` Pow A) = card (Pow A)" by (rule card_image) also have "\ + \ = 2 * \" by (simp add: mult_2) also from insert(3) have "\ = 2 ^ Suc (card A)" by simp also from insert(1,2) have "Suc (card A) = card (insert x A)" by (rule card_insert_disjoint [symmetric]) finally show ?case . qed subsection \Code generator tweak\ code_identifier code_module Power \ (SML) Arith and (OCaml) Arith and (Haskell) Arith end diff --git a/src/HOL/Real_Vector_Spaces.thy b/src/HOL/Real_Vector_Spaces.thy --- a/src/HOL/Real_Vector_Spaces.thy +++ b/src/HOL/Real_Vector_Spaces.thy @@ -1,2482 +1,2487 @@ (* Title: HOL/Real_Vector_Spaces.thy Author: Brian Huffman Author: Johannes Hölzl *) section \Vector Spaces and Algebras over the Reals\ theory Real_Vector_Spaces imports Real Topological_Spaces Vector_Spaces begin subsection \Real vector spaces\ class scaleR = fixes scaleR :: "real \ 'a \ 'a" (infixr "*\<^sub>R" 75) begin abbreviation divideR :: "'a \ real \ 'a" (infixl "'/\<^sub>R" 70) where "x /\<^sub>R r \ inverse r *\<^sub>R x" end class real_vector = scaleR + ab_group_add + assumes scaleR_add_right: "a *\<^sub>R (x + y) = a *\<^sub>R x + a *\<^sub>R y" and scaleR_add_left: "(a + b) *\<^sub>R x = a *\<^sub>R x + b *\<^sub>R x" and scaleR_scaleR: "a *\<^sub>R b *\<^sub>R x = (a * b) *\<^sub>R x" and scaleR_one: "1 *\<^sub>R x = x" class real_algebra = real_vector + ring + assumes mult_scaleR_left [simp]: "a *\<^sub>R x * y = a *\<^sub>R (x * y)" and mult_scaleR_right [simp]: "x * a *\<^sub>R y = a *\<^sub>R (x * y)" class real_algebra_1 = real_algebra + ring_1 class real_div_algebra = real_algebra_1 + division_ring class real_field = real_div_algebra + field instantiation real :: real_field begin definition real_scaleR_def [simp]: "scaleR a x = a * x" instance by standard (simp_all add: algebra_simps) end locale linear = Vector_Spaces.linear "scaleR::_\_\'a::real_vector" "scaleR::_\_\'b::real_vector" begin lemmas scaleR = scale end global_interpretation real_vector?: vector_space "scaleR :: real \ 'a \ 'a :: real_vector" rewrites "Vector_Spaces.linear (*\<^sub>R) (*\<^sub>R) = linear" and "Vector_Spaces.linear (*) (*\<^sub>R) = linear" defines dependent_raw_def: dependent = real_vector.dependent and representation_raw_def: representation = real_vector.representation and subspace_raw_def: subspace = real_vector.subspace and span_raw_def: span = real_vector.span and extend_basis_raw_def: extend_basis = real_vector.extend_basis and dim_raw_def: dim = real_vector.dim proof unfold_locales show "Vector_Spaces.linear (*\<^sub>R) (*\<^sub>R) = linear" "Vector_Spaces.linear (*) (*\<^sub>R) = linear" by (force simp: linear_def real_scaleR_def[abs_def])+ qed (use scaleR_add_right scaleR_add_left scaleR_scaleR scaleR_one in auto) hide_const (open)\ \locale constants\ real_vector.dependent real_vector.independent real_vector.representation real_vector.subspace real_vector.span real_vector.extend_basis real_vector.dim abbreviation "independent x \ \ dependent x" global_interpretation real_vector?: vector_space_pair "scaleR::_\_\'a::real_vector" "scaleR::_\_\'b::real_vector" rewrites "Vector_Spaces.linear (*\<^sub>R) (*\<^sub>R) = linear" and "Vector_Spaces.linear (*) (*\<^sub>R) = linear" defines construct_raw_def: construct = real_vector.construct proof unfold_locales show "Vector_Spaces.linear (*) (*\<^sub>R) = linear" unfolding linear_def real_scaleR_def by auto qed (auto simp: linear_def) hide_const (open)\ \locale constants\ real_vector.construct lemma linear_compose: "linear f \ linear g \ linear (g \ f)" unfolding linear_def by (rule Vector_Spaces.linear_compose) text \Recover original theorem names\ lemmas scaleR_left_commute = real_vector.scale_left_commute lemmas scaleR_zero_left = real_vector.scale_zero_left lemmas scaleR_minus_left = real_vector.scale_minus_left lemmas scaleR_diff_left = real_vector.scale_left_diff_distrib lemmas scaleR_sum_left = real_vector.scale_sum_left lemmas scaleR_zero_right = real_vector.scale_zero_right lemmas scaleR_minus_right = real_vector.scale_minus_right lemmas scaleR_diff_right = real_vector.scale_right_diff_distrib lemmas scaleR_sum_right = real_vector.scale_sum_right lemmas scaleR_eq_0_iff = real_vector.scale_eq_0_iff lemmas scaleR_left_imp_eq = real_vector.scale_left_imp_eq lemmas scaleR_right_imp_eq = real_vector.scale_right_imp_eq lemmas scaleR_cancel_left = real_vector.scale_cancel_left lemmas scaleR_cancel_right = real_vector.scale_cancel_right lemma [field_simps]: "c \ 0 \ a = b /\<^sub>R c \ c *\<^sub>R a = b" "c \ 0 \ b /\<^sub>R c = a \ b = c *\<^sub>R a" "c \ 0 \ a + b /\<^sub>R c = (c *\<^sub>R a + b) /\<^sub>R c" "c \ 0 \ a /\<^sub>R c + b = (a + c *\<^sub>R b) /\<^sub>R c" "c \ 0 \ a - b /\<^sub>R c = (c *\<^sub>R a - b) /\<^sub>R c" "c \ 0 \ a /\<^sub>R c - b = (a - c *\<^sub>R b) /\<^sub>R c" "c \ 0 \ - (a /\<^sub>R c) + b = (- a + c *\<^sub>R b) /\<^sub>R c" "c \ 0 \ - (a /\<^sub>R c) - b = (- a - c *\<^sub>R b) /\<^sub>R c" for a b :: "'a :: real_vector" by (auto simp add: scaleR_add_right scaleR_add_left scaleR_diff_right scaleR_diff_left) text \Legacy names\ lemmas scaleR_left_distrib = scaleR_add_left lemmas scaleR_right_distrib = scaleR_add_right lemmas scaleR_left_diff_distrib = scaleR_diff_left lemmas scaleR_right_diff_distrib = scaleR_diff_right lemmas linear_injective_0 = linear_inj_iff_eq_0 and linear_injective_on_subspace_0 = linear_inj_on_iff_eq_0 and linear_cmul = linear_scale and linear_scaleR = linear_scale_self and subspace_mul = subspace_scale and span_linear_image = linear_span_image and span_0 = span_zero and span_mul = span_scale and injective_scaleR = injective_scale lemma scaleR_minus1_left [simp]: "scaleR (-1) x = - x" for x :: "'a::real_vector" using scaleR_minus_left [of 1 x] by simp lemma scaleR_2: fixes x :: "'a::real_vector" shows "scaleR 2 x = x + x" unfolding one_add_one [symmetric] scaleR_left_distrib by simp lemma scaleR_half_double [simp]: fixes a :: "'a::real_vector" shows "(1 / 2) *\<^sub>R (a + a) = a" proof - have "\r. r *\<^sub>R (a + a) = (r * 2) *\<^sub>R a" by (metis scaleR_2 scaleR_scaleR) then show ?thesis by simp qed +lemma shift_zero_ident [simp]: + fixes f :: "'a \ 'b::real_vector" + shows "(+)0 \ f = f" + by force + lemma linear_scale_real: fixes r::real shows "linear f \ f (r * b) = r * f b" using linear_scale by fastforce interpretation scaleR_left: additive "(\a. scaleR a x :: 'a::real_vector)" by standard (rule scaleR_left_distrib) interpretation scaleR_right: additive "(\x. scaleR a x :: 'a::real_vector)" by standard (rule scaleR_right_distrib) lemma nonzero_inverse_scaleR_distrib: "a \ 0 \ x \ 0 \ inverse (scaleR a x) = scaleR (inverse a) (inverse x)" for x :: "'a::real_div_algebra" by (rule inverse_unique) simp lemma inverse_scaleR_distrib: "inverse (scaleR a x) = scaleR (inverse a) (inverse x)" for x :: "'a::{real_div_algebra,division_ring}" by (metis inverse_zero nonzero_inverse_scaleR_distrib scale_eq_0_iff) lemmas sum_constant_scaleR = real_vector.sum_constant_scale\ \legacy name\ named_theorems vector_add_divide_simps "to simplify sums of scaled vectors" lemma [vector_add_divide_simps]: "v + (b / z) *\<^sub>R w = (if z = 0 then v else (z *\<^sub>R v + b *\<^sub>R w) /\<^sub>R z)" "a *\<^sub>R v + (b / z) *\<^sub>R w = (if z = 0 then a *\<^sub>R v else ((a * z) *\<^sub>R v + b *\<^sub>R w) /\<^sub>R z)" "(a / z) *\<^sub>R v + w = (if z = 0 then w else (a *\<^sub>R v + z *\<^sub>R w) /\<^sub>R z)" "(a / z) *\<^sub>R v + b *\<^sub>R w = (if z = 0 then b *\<^sub>R w else (a *\<^sub>R v + (b * z) *\<^sub>R w) /\<^sub>R z)" "v - (b / z) *\<^sub>R w = (if z = 0 then v else (z *\<^sub>R v - b *\<^sub>R w) /\<^sub>R z)" "a *\<^sub>R v - (b / z) *\<^sub>R w = (if z = 0 then a *\<^sub>R v else ((a * z) *\<^sub>R v - b *\<^sub>R w) /\<^sub>R z)" "(a / z) *\<^sub>R v - w = (if z = 0 then -w else (a *\<^sub>R v - z *\<^sub>R w) /\<^sub>R z)" "(a / z) *\<^sub>R v - b *\<^sub>R w = (if z = 0 then -b *\<^sub>R w else (a *\<^sub>R v - (b * z) *\<^sub>R w) /\<^sub>R z)" for v :: "'a :: real_vector" by (simp_all add: divide_inverse_commute scaleR_add_right scaleR_diff_right) lemma eq_vector_fraction_iff [vector_add_divide_simps]: fixes x :: "'a :: real_vector" shows "(x = (u / v) *\<^sub>R a) \ (if v=0 then x = 0 else v *\<^sub>R x = u *\<^sub>R a)" by auto (metis (no_types) divide_eq_1_iff divide_inverse_commute scaleR_one scaleR_scaleR) lemma vector_fraction_eq_iff [vector_add_divide_simps]: fixes x :: "'a :: real_vector" shows "((u / v) *\<^sub>R a = x) \ (if v=0 then x = 0 else u *\<^sub>R a = v *\<^sub>R x)" by (metis eq_vector_fraction_iff) lemma real_vector_affinity_eq: fixes x :: "'a :: real_vector" assumes m0: "m \ 0" shows "m *\<^sub>R x + c = y \ x = inverse m *\<^sub>R y - (inverse m *\<^sub>R c)" (is "?lhs \ ?rhs") proof assume ?lhs then have "m *\<^sub>R x = y - c" by (simp add: field_simps) then have "inverse m *\<^sub>R (m *\<^sub>R x) = inverse m *\<^sub>R (y - c)" by simp then show "x = inverse m *\<^sub>R y - (inverse m *\<^sub>R c)" using m0 by (simp add: scaleR_diff_right) next assume ?rhs with m0 show "m *\<^sub>R x + c = y" by (simp add: scaleR_diff_right) qed lemma real_vector_eq_affinity: "m \ 0 \ y = m *\<^sub>R x + c \ inverse m *\<^sub>R y - (inverse m *\<^sub>R c) = x" for x :: "'a::real_vector" using real_vector_affinity_eq[where m=m and x=x and y=y and c=c] by metis lemma scaleR_eq_iff [simp]: "b + u *\<^sub>R a = a + u *\<^sub>R b \ a = b \ u = 1" for a :: "'a::real_vector" proof (cases "u = 1") case True then show ?thesis by auto next case False have "a = b" if "b + u *\<^sub>R a = a + u *\<^sub>R b" proof - from that have "(u - 1) *\<^sub>R a = (u - 1) *\<^sub>R b" by (simp add: algebra_simps) with False show ?thesis by auto qed then show ?thesis by auto qed lemma scaleR_collapse [simp]: "(1 - u) *\<^sub>R a + u *\<^sub>R a = a" for a :: "'a::real_vector" by (simp add: algebra_simps) subsection \Embedding of the Reals into any \real_algebra_1\: \of_real\\ definition of_real :: "real \ 'a::real_algebra_1" where "of_real r = scaleR r 1" lemma scaleR_conv_of_real: "scaleR r x = of_real r * x" by (simp add: of_real_def) lemma of_real_0 [simp]: "of_real 0 = 0" by (simp add: of_real_def) lemma of_real_1 [simp]: "of_real 1 = 1" by (simp add: of_real_def) lemma of_real_add [simp]: "of_real (x + y) = of_real x + of_real y" by (simp add: of_real_def scaleR_left_distrib) lemma of_real_minus [simp]: "of_real (- x) = - of_real x" by (simp add: of_real_def) lemma of_real_diff [simp]: "of_real (x - y) = of_real x - of_real y" by (simp add: of_real_def scaleR_left_diff_distrib) lemma of_real_mult [simp]: "of_real (x * y) = of_real x * of_real y" by (simp add: of_real_def) lemma of_real_sum[simp]: "of_real (sum f s) = (\x\s. of_real (f x))" by (induct s rule: infinite_finite_induct) auto lemma of_real_prod[simp]: "of_real (prod f s) = (\x\s. of_real (f x))" by (induct s rule: infinite_finite_induct) auto lemma nonzero_of_real_inverse: "x \ 0 \ of_real (inverse x) = inverse (of_real x :: 'a::real_div_algebra)" by (simp add: of_real_def nonzero_inverse_scaleR_distrib) lemma of_real_inverse [simp]: "of_real (inverse x) = inverse (of_real x :: 'a::{real_div_algebra,division_ring})" by (simp add: of_real_def inverse_scaleR_distrib) lemma nonzero_of_real_divide: "y \ 0 \ of_real (x / y) = (of_real x / of_real y :: 'a::real_field)" by (simp add: divide_inverse nonzero_of_real_inverse) lemma of_real_divide [simp]: "of_real (x / y) = (of_real x / of_real y :: 'a::real_div_algebra)" by (simp add: divide_inverse) lemma of_real_power [simp]: "of_real (x ^ n) = (of_real x :: 'a::{real_algebra_1}) ^ n" by (induct n) simp_all lemma of_real_power_int [simp]: "of_real (power_int x n) = power_int (of_real x :: 'a :: {real_div_algebra,division_ring}) n" by (auto simp: power_int_def) lemma of_real_eq_iff [simp]: "of_real x = of_real y \ x = y" by (simp add: of_real_def) lemma inj_of_real: "inj of_real" by (auto intro: injI) lemmas of_real_eq_0_iff [simp] = of_real_eq_iff [of _ 0, simplified] lemmas of_real_eq_1_iff [simp] = of_real_eq_iff [of _ 1, simplified] lemma minus_of_real_eq_of_real_iff [simp]: "-of_real x = of_real y \ -x = y" using of_real_eq_iff[of "-x" y] by (simp only: of_real_minus) lemma of_real_eq_minus_of_real_iff [simp]: "of_real x = -of_real y \ x = -y" using of_real_eq_iff[of x "-y"] by (simp only: of_real_minus) lemma of_real_eq_id [simp]: "of_real = (id :: real \ real)" by (rule ext) (simp add: of_real_def) text \Collapse nested embeddings.\ lemma of_real_of_nat_eq [simp]: "of_real (of_nat n) = of_nat n" by (induct n) auto lemma of_real_of_int_eq [simp]: "of_real (of_int z) = of_int z" by (cases z rule: int_diff_cases) simp lemma of_real_numeral [simp]: "of_real (numeral w) = numeral w" using of_real_of_int_eq [of "numeral w"] by simp lemma of_real_neg_numeral [simp]: "of_real (- numeral w) = - numeral w" using of_real_of_int_eq [of "- numeral w"] by simp lemma numeral_power_int_eq_of_real_cancel_iff [simp]: "power_int (numeral x) n = (of_real y :: 'a :: {real_div_algebra, division_ring}) \ power_int (numeral x) n = y" proof - have "power_int (numeral x) n = (of_real (power_int (numeral x) n) :: 'a)" by simp also have "\ = of_real y \ power_int (numeral x) n = y" by (subst of_real_eq_iff) auto finally show ?thesis . qed lemma of_real_eq_numeral_power_int_cancel_iff [simp]: "(of_real y :: 'a :: {real_div_algebra, division_ring}) = power_int (numeral x) n \ y = power_int (numeral x) n" by (subst (1 2) eq_commute) simp lemma of_real_eq_of_real_power_int_cancel_iff [simp]: "power_int (of_real b :: 'a :: {real_div_algebra, division_ring}) w = of_real x \ power_int b w = x" by (metis of_real_power_int of_real_eq_iff) lemma of_real_in_Ints_iff [simp]: "of_real x \ \ \ x \ \" proof safe fix x assume "(of_real x :: 'a) \ \" then obtain n where "(of_real x :: 'a) = of_int n" by (auto simp: Ints_def) also have "of_int n = of_real (real_of_int n)" by simp finally have "x = real_of_int n" by (subst (asm) of_real_eq_iff) thus "x \ \" by auto qed (auto simp: Ints_def) lemma Ints_of_real [intro]: "x \ \ \ of_real x \ \" by simp text \Every real algebra has characteristic zero.\ instance real_algebra_1 < ring_char_0 proof from inj_of_real inj_of_nat have "inj (of_real \ of_nat)" by (rule inj_compose) then show "inj (of_nat :: nat \ 'a)" by (simp add: comp_def) qed lemma fraction_scaleR_times [simp]: fixes a :: "'a::real_algebra_1" shows "(numeral u / numeral v) *\<^sub>R (numeral w * a) = (numeral u * numeral w / numeral v) *\<^sub>R a" by (metis (no_types, lifting) of_real_numeral scaleR_conv_of_real scaleR_scaleR times_divide_eq_left) lemma inverse_scaleR_times [simp]: fixes a :: "'a::real_algebra_1" shows "(1 / numeral v) *\<^sub>R (numeral w * a) = (numeral w / numeral v) *\<^sub>R a" by (metis divide_inverse_commute inverse_eq_divide of_real_numeral scaleR_conv_of_real scaleR_scaleR) lemma scaleR_times [simp]: fixes a :: "'a::real_algebra_1" shows "(numeral u) *\<^sub>R (numeral w * a) = (numeral u * numeral w) *\<^sub>R a" by (simp add: scaleR_conv_of_real) instance real_field < field_char_0 .. subsection \The Set of Real Numbers\ definition Reals :: "'a::real_algebra_1 set" ("\") where "\ = range of_real" lemma Reals_of_real [simp]: "of_real r \ \" by (simp add: Reals_def) lemma Reals_of_int [simp]: "of_int z \ \" by (subst of_real_of_int_eq [symmetric], rule Reals_of_real) lemma Reals_of_nat [simp]: "of_nat n \ \" by (subst of_real_of_nat_eq [symmetric], rule Reals_of_real) lemma Reals_numeral [simp]: "numeral w \ \" by (subst of_real_numeral [symmetric], rule Reals_of_real) lemma Reals_0 [simp]: "0 \ \" and Reals_1 [simp]: "1 \ \" by (simp_all add: Reals_def) lemma Reals_add [simp]: "a \ \ \ b \ \ \ a + b \ \" by (metis (no_types, opaque_lifting) Reals_def Reals_of_real imageE of_real_add) lemma Reals_minus [simp]: "a \ \ \ - a \ \" by (auto simp: Reals_def) lemma Reals_minus_iff [simp]: "- a \ \ \ a \ \" using Reals_minus by fastforce lemma Reals_diff [simp]: "a \ \ \ b \ \ \ a - b \ \" by (metis Reals_add Reals_minus_iff add_uminus_conv_diff) lemma Reals_mult [simp]: "a \ \ \ b \ \ \ a * b \ \" by (metis (no_types, lifting) Reals_def Reals_of_real imageE of_real_mult) lemma nonzero_Reals_inverse: "a \ \ \ a \ 0 \ inverse a \ \" for a :: "'a::real_div_algebra" by (metis Reals_def Reals_of_real imageE of_real_inverse) lemma Reals_inverse: "a \ \ \ inverse a \ \" for a :: "'a::{real_div_algebra,division_ring}" using nonzero_Reals_inverse by fastforce lemma Reals_inverse_iff [simp]: "inverse x \ \ \ x \ \" for x :: "'a::{real_div_algebra,division_ring}" by (metis Reals_inverse inverse_inverse_eq) lemma nonzero_Reals_divide: "a \ \ \ b \ \ \ b \ 0 \ a / b \ \" for a b :: "'a::real_field" by (simp add: divide_inverse) lemma Reals_divide [simp]: "a \ \ \ b \ \ \ a / b \ \" for a b :: "'a::{real_field,field}" using nonzero_Reals_divide by fastforce lemma Reals_power [simp]: "a \ \ \ a ^ n \ \" for a :: "'a::real_algebra_1" by (metis Reals_def Reals_of_real imageE of_real_power) lemma Reals_cases [cases set: Reals]: assumes "q \ \" obtains (of_real) r where "q = of_real r" unfolding Reals_def proof - from \q \ \\ have "q \ range of_real" unfolding Reals_def . then obtain r where "q = of_real r" .. then show thesis .. qed lemma sum_in_Reals [intro,simp]: "(\i. i \ s \ f i \ \) \ sum f s \ \" proof (induct s rule: infinite_finite_induct) case infinite then show ?case by (metis Reals_0 sum.infinite) qed simp_all lemma prod_in_Reals [intro,simp]: "(\i. i \ s \ f i \ \) \ prod f s \ \" proof (induct s rule: infinite_finite_induct) case infinite then show ?case by (metis Reals_1 prod.infinite) qed simp_all lemma Reals_induct [case_names of_real, induct set: Reals]: "q \ \ \ (\r. P (of_real r)) \ P q" by (rule Reals_cases) auto subsection \Ordered real vector spaces\ class ordered_real_vector = real_vector + ordered_ab_group_add + assumes scaleR_left_mono: "x \ y \ 0 \ a \ a *\<^sub>R x \ a *\<^sub>R y" and scaleR_right_mono: "a \ b \ 0 \ x \ a *\<^sub>R x \ b *\<^sub>R x" begin lemma scaleR_mono: "a \ b \ x \ y \ 0 \ b \ 0 \ x \ a *\<^sub>R x \ b *\<^sub>R y" by (meson order_trans scaleR_left_mono scaleR_right_mono) lemma scaleR_mono': "a \ b \ c \ d \ 0 \ a \ 0 \ c \ a *\<^sub>R c \ b *\<^sub>R d" by (rule scaleR_mono) (auto intro: order.trans) lemma pos_le_divideR_eq [field_simps]: "a \ b /\<^sub>R c \ c *\<^sub>R a \ b" (is "?P \ ?Q") if "0 < c" proof assume ?P with scaleR_left_mono that have "c *\<^sub>R a \ c *\<^sub>R (b /\<^sub>R c)" by simp with that show ?Q by (simp add: scaleR_one scaleR_scaleR inverse_eq_divide) next assume ?Q with scaleR_left_mono that have "c *\<^sub>R a /\<^sub>R c \ b /\<^sub>R c" by simp with that show ?P by (simp add: scaleR_one scaleR_scaleR inverse_eq_divide) qed lemma pos_less_divideR_eq [field_simps]: "a < b /\<^sub>R c \ c *\<^sub>R a < b" if "c > 0" using that pos_le_divideR_eq [of c a b] by (auto simp add: le_less scaleR_scaleR scaleR_one) lemma pos_divideR_le_eq [field_simps]: "b /\<^sub>R c \ a \ b \ c *\<^sub>R a" if "c > 0" using that pos_le_divideR_eq [of "inverse c" b a] by simp lemma pos_divideR_less_eq [field_simps]: "b /\<^sub>R c < a \ b < c *\<^sub>R a" if "c > 0" using that pos_less_divideR_eq [of "inverse c" b a] by simp lemma pos_le_minus_divideR_eq [field_simps]: "a \ - (b /\<^sub>R c) \ c *\<^sub>R a \ - b" if "c > 0" using that by (metis add_minus_cancel diff_0 left_minus minus_minus neg_le_iff_le scaleR_add_right uminus_add_conv_diff pos_le_divideR_eq) lemma pos_less_minus_divideR_eq [field_simps]: "a < - (b /\<^sub>R c) \ c *\<^sub>R a < - b" if "c > 0" using that by (metis le_less less_le_not_le pos_divideR_le_eq pos_divideR_less_eq pos_le_minus_divideR_eq) lemma pos_minus_divideR_le_eq [field_simps]: "- (b /\<^sub>R c) \ a \ - b \ c *\<^sub>R a" if "c > 0" using that by (metis pos_divideR_le_eq pos_le_minus_divideR_eq that inverse_positive_iff_positive le_imp_neg_le minus_minus) lemma pos_minus_divideR_less_eq [field_simps]: "- (b /\<^sub>R c) < a \ - b < c *\<^sub>R a" if "c > 0" using that by (simp add: less_le_not_le pos_le_minus_divideR_eq pos_minus_divideR_le_eq) lemma scaleR_image_atLeastAtMost: "c > 0 \ scaleR c ` {x..y} = {c *\<^sub>R x..c *\<^sub>R y}" apply (auto intro!: scaleR_left_mono simp: image_iff Bex_def) using pos_divideR_le_eq [of c] pos_le_divideR_eq [of c] apply (meson local.order_eq_iff) done end lemma neg_le_divideR_eq [field_simps]: "a \ b /\<^sub>R c \ b \ c *\<^sub>R a" (is "?P \ ?Q") if "c < 0" for a b :: "'a :: ordered_real_vector" using that pos_le_divideR_eq [of "- c" a "- b"] by simp lemma neg_less_divideR_eq [field_simps]: "a < b /\<^sub>R c \ b < c *\<^sub>R a" if "c < 0" for a b :: "'a :: ordered_real_vector" using that neg_le_divideR_eq [of c a b] by (auto simp add: le_less) lemma neg_divideR_le_eq [field_simps]: "b /\<^sub>R c \ a \ c *\<^sub>R a \ b" if "c < 0" for a b :: "'a :: ordered_real_vector" using that pos_divideR_le_eq [of "- c" "- b" a] by simp lemma neg_divideR_less_eq [field_simps]: "b /\<^sub>R c < a \ c *\<^sub>R a < b" if "c < 0" for a b :: "'a :: ordered_real_vector" using that neg_divideR_le_eq [of c b a] by (auto simp add: le_less) lemma neg_le_minus_divideR_eq [field_simps]: "a \ - (b /\<^sub>R c) \ - b \ c *\<^sub>R a" if "c < 0" for a b :: "'a :: ordered_real_vector" using that pos_le_minus_divideR_eq [of "- c" a "- b"] by (simp add: minus_le_iff) lemma neg_less_minus_divideR_eq [field_simps]: "a < - (b /\<^sub>R c) \ - b < c *\<^sub>R a" if "c < 0" for a b :: "'a :: ordered_real_vector" proof - have *: "- b = c *\<^sub>R a \ b = - (c *\<^sub>R a)" by (metis add.inverse_inverse) from that neg_le_minus_divideR_eq [of c a b] show ?thesis by (auto simp add: le_less *) qed lemma neg_minus_divideR_le_eq [field_simps]: "- (b /\<^sub>R c) \ a \ c *\<^sub>R a \ - b" if "c < 0" for a b :: "'a :: ordered_real_vector" using that pos_minus_divideR_le_eq [of "- c" "- b" a] by (simp add: le_minus_iff) lemma neg_minus_divideR_less_eq [field_simps]: "- (b /\<^sub>R c) < a \ c *\<^sub>R a < - b" if "c < 0" for a b :: "'a :: ordered_real_vector" using that by (simp add: less_le_not_le neg_le_minus_divideR_eq neg_minus_divideR_le_eq) lemma [field_split_simps]: "a = b /\<^sub>R c \ (if c = 0 then a = 0 else c *\<^sub>R a = b)" "b /\<^sub>R c = a \ (if c = 0 then a = 0 else b = c *\<^sub>R a)" "a + b /\<^sub>R c = (if c = 0 then a else (c *\<^sub>R a + b) /\<^sub>R c)" "a /\<^sub>R c + b = (if c = 0 then b else (a + c *\<^sub>R b) /\<^sub>R c)" "a - b /\<^sub>R c = (if c = 0 then a else (c *\<^sub>R a - b) /\<^sub>R c)" "a /\<^sub>R c - b = (if c = 0 then - b else (a - c *\<^sub>R b) /\<^sub>R c)" "- (a /\<^sub>R c) + b = (if c = 0 then b else (- a + c *\<^sub>R b) /\<^sub>R c)" "- (a /\<^sub>R c) - b = (if c = 0 then - b else (- a - c *\<^sub>R b) /\<^sub>R c)" for a b :: "'a :: real_vector" by (auto simp add: field_simps) lemma [field_split_simps]: "0 < c \ a \ b /\<^sub>R c \ (if c > 0 then c *\<^sub>R a \ b else if c < 0 then b \ c *\<^sub>R a else a \ 0)" "0 < c \ a < b /\<^sub>R c \ (if c > 0 then c *\<^sub>R a < b else if c < 0 then b < c *\<^sub>R a else a < 0)" "0 < c \ b /\<^sub>R c \ a \ (if c > 0 then b \ c *\<^sub>R a else if c < 0 then c *\<^sub>R a \ b else a \ 0)" "0 < c \ b /\<^sub>R c < a \ (if c > 0 then b < c *\<^sub>R a else if c < 0 then c *\<^sub>R a < b else a > 0)" "0 < c \ a \ - (b /\<^sub>R c) \ (if c > 0 then c *\<^sub>R a \ - b else if c < 0 then - b \ c *\<^sub>R a else a \ 0)" "0 < c \ a < - (b /\<^sub>R c) \ (if c > 0 then c *\<^sub>R a < - b else if c < 0 then - b < c *\<^sub>R a else a < 0)" "0 < c \ - (b /\<^sub>R c) \ a \ (if c > 0 then - b \ c *\<^sub>R a else if c < 0 then c *\<^sub>R a \ - b else a \ 0)" "0 < c \ - (b /\<^sub>R c) < a \ (if c > 0 then - b < c *\<^sub>R a else if c < 0 then c *\<^sub>R a < - b else a > 0)" for a b :: "'a :: ordered_real_vector" by (clarsimp intro!: field_simps)+ lemma scaleR_nonneg_nonneg: "0 \ a \ 0 \ x \ 0 \ a *\<^sub>R x" for x :: "'a::ordered_real_vector" using scaleR_left_mono [of 0 x a] by simp lemma scaleR_nonneg_nonpos: "0 \ a \ x \ 0 \ a *\<^sub>R x \ 0" for x :: "'a::ordered_real_vector" using scaleR_left_mono [of x 0 a] by simp lemma scaleR_nonpos_nonneg: "a \ 0 \ 0 \ x \ a *\<^sub>R x \ 0" for x :: "'a::ordered_real_vector" using scaleR_right_mono [of a 0 x] by simp lemma split_scaleR_neg_le: "(0 \ a \ x \ 0) \ (a \ 0 \ 0 \ x) \ a *\<^sub>R x \ 0" for x :: "'a::ordered_real_vector" by (auto simp: scaleR_nonneg_nonpos scaleR_nonpos_nonneg) lemma le_add_iff1: "a *\<^sub>R e + c \ b *\<^sub>R e + d \ (a - b) *\<^sub>R e + c \ d" for c d e :: "'a::ordered_real_vector" by (simp add: algebra_simps) lemma le_add_iff2: "a *\<^sub>R e + c \ b *\<^sub>R e + d \ c \ (b - a) *\<^sub>R e + d" for c d e :: "'a::ordered_real_vector" by (simp add: algebra_simps) lemma scaleR_left_mono_neg: "b \ a \ c \ 0 \ c *\<^sub>R a \ c *\<^sub>R b" for a b :: "'a::ordered_real_vector" by (drule scaleR_left_mono [of _ _ "- c"], simp_all) lemma scaleR_right_mono_neg: "b \ a \ c \ 0 \ a *\<^sub>R c \ b *\<^sub>R c" for c :: "'a::ordered_real_vector" by (drule scaleR_right_mono [of _ _ "- c"], simp_all) lemma scaleR_nonpos_nonpos: "a \ 0 \ b \ 0 \ 0 \ a *\<^sub>R b" for b :: "'a::ordered_real_vector" using scaleR_right_mono_neg [of a 0 b] by simp lemma split_scaleR_pos_le: "(0 \ a \ 0 \ b) \ (a \ 0 \ b \ 0) \ 0 \ a *\<^sub>R b" for b :: "'a::ordered_real_vector" by (auto simp: scaleR_nonneg_nonneg scaleR_nonpos_nonpos) lemma zero_le_scaleR_iff: fixes b :: "'a::ordered_real_vector" shows "0 \ a *\<^sub>R b \ 0 < a \ 0 \ b \ a < 0 \ b \ 0 \ a = 0" (is "?lhs = ?rhs") proof (cases "a = 0") case True then show ?thesis by simp next case False show ?thesis proof assume ?lhs from \a \ 0\ consider "a > 0" | "a < 0" by arith then show ?rhs proof cases case 1 with \?lhs\ have "inverse a *\<^sub>R 0 \ inverse a *\<^sub>R (a *\<^sub>R b)" by (intro scaleR_mono) auto with 1 show ?thesis by simp next case 2 with \?lhs\ have "- inverse a *\<^sub>R 0 \ - inverse a *\<^sub>R (a *\<^sub>R b)" by (intro scaleR_mono) auto with 2 show ?thesis by simp qed next assume ?rhs then show ?lhs by (auto simp: not_le \a \ 0\ intro!: split_scaleR_pos_le) qed qed lemma scaleR_le_0_iff: "a *\<^sub>R b \ 0 \ 0 < a \ b \ 0 \ a < 0 \ 0 \ b \ a = 0" for b::"'a::ordered_real_vector" by (insert zero_le_scaleR_iff [of "-a" b]) force lemma scaleR_le_cancel_left: "c *\<^sub>R a \ c *\<^sub>R b \ (0 < c \ a \ b) \ (c < 0 \ b \ a)" for b :: "'a::ordered_real_vector" by (auto simp: neq_iff scaleR_left_mono scaleR_left_mono_neg dest: scaleR_left_mono[where a="inverse c"] scaleR_left_mono_neg[where c="inverse c"]) lemma scaleR_le_cancel_left_pos: "0 < c \ c *\<^sub>R a \ c *\<^sub>R b \ a \ b" for b :: "'a::ordered_real_vector" by (auto simp: scaleR_le_cancel_left) lemma scaleR_le_cancel_left_neg: "c < 0 \ c *\<^sub>R a \ c *\<^sub>R b \ b \ a" for b :: "'a::ordered_real_vector" by (auto simp: scaleR_le_cancel_left) lemma scaleR_left_le_one_le: "0 \ x \ a \ 1 \ a *\<^sub>R x \ x" for x :: "'a::ordered_real_vector" and a :: real using scaleR_right_mono[of a 1 x] by simp subsection \Real normed vector spaces\ class dist = fixes dist :: "'a \ 'a \ real" class norm = fixes norm :: "'a \ real" class sgn_div_norm = scaleR + norm + sgn + assumes sgn_div_norm: "sgn x = x /\<^sub>R norm x" class dist_norm = dist + norm + minus + assumes dist_norm: "dist x y = norm (x - y)" class uniformity_dist = dist + uniformity + assumes uniformity_dist: "uniformity = (INF e\{0 <..}. principal {(x, y). dist x y < e})" begin lemma eventually_uniformity_metric: "eventually P uniformity \ (\e>0. \x y. dist x y < e \ P (x, y))" unfolding uniformity_dist by (subst eventually_INF_base) (auto simp: eventually_principal subset_eq intro: bexI[of _ "min _ _"]) end class real_normed_vector = real_vector + sgn_div_norm + dist_norm + uniformity_dist + open_uniformity + assumes norm_eq_zero [simp]: "norm x = 0 \ x = 0" and norm_triangle_ineq: "norm (x + y) \ norm x + norm y" and norm_scaleR [simp]: "norm (scaleR a x) = \a\ * norm x" begin lemma norm_ge_zero [simp]: "0 \ norm x" proof - have "0 = norm (x + -1 *\<^sub>R x)" using scaleR_add_left[of 1 "-1" x] norm_scaleR[of 0 x] by (simp add: scaleR_one) also have "\ \ norm x + norm (-1 *\<^sub>R x)" by (rule norm_triangle_ineq) finally show ?thesis by simp qed lemma bdd_below_norm_image: "bdd_below (norm ` A)" by (meson bdd_belowI2 norm_ge_zero) end class real_normed_algebra = real_algebra + real_normed_vector + assumes norm_mult_ineq: "norm (x * y) \ norm x * norm y" class real_normed_algebra_1 = real_algebra_1 + real_normed_algebra + assumes norm_one [simp]: "norm 1 = 1" lemma (in real_normed_algebra_1) scaleR_power [simp]: "(scaleR x y) ^ n = scaleR (x^n) (y^n)" by (induct n) (simp_all add: scaleR_one scaleR_scaleR mult_ac) class real_normed_div_algebra = real_div_algebra + real_normed_vector + assumes norm_mult: "norm (x * y) = norm x * norm y" class real_normed_field = real_field + real_normed_div_algebra instance real_normed_div_algebra < real_normed_algebra_1 proof show "norm (x * y) \ norm x * norm y" for x y :: 'a by (simp add: norm_mult) next have "norm (1 * 1::'a) = norm (1::'a) * norm (1::'a)" by (rule norm_mult) then show "norm (1::'a) = 1" by simp qed context real_normed_vector begin lemma norm_zero [simp]: "norm (0::'a) = 0" by simp lemma zero_less_norm_iff [simp]: "norm x > 0 \ x \ 0" by (simp add: order_less_le) lemma norm_not_less_zero [simp]: "\ norm x < 0" by (simp add: linorder_not_less) lemma norm_le_zero_iff [simp]: "norm x \ 0 \ x = 0" by (simp add: order_le_less) lemma norm_minus_cancel [simp]: "norm (- x) = norm x" proof - have "- 1 *\<^sub>R x = - (1 *\<^sub>R x)" unfolding add_eq_0_iff2[symmetric] scaleR_add_left[symmetric] using norm_eq_zero by fastforce then have "norm (- x) = norm (scaleR (- 1) x)" by (simp only: scaleR_one) also have "\ = \- 1\ * norm x" by (rule norm_scaleR) finally show ?thesis by simp qed lemma norm_minus_commute: "norm (a - b) = norm (b - a)" proof - have "norm (- (b - a)) = norm (b - a)" by (rule norm_minus_cancel) then show ?thesis by simp qed lemma dist_add_cancel [simp]: "dist (a + b) (a + c) = dist b c" by (simp add: dist_norm) lemma dist_add_cancel2 [simp]: "dist (b + a) (c + a) = dist b c" by (simp add: dist_norm) lemma norm_uminus_minus: "norm (- x - y) = norm (x + y)" by (subst (2) norm_minus_cancel[symmetric], subst minus_add_distrib) simp lemma norm_triangle_ineq2: "norm a - norm b \ norm (a - b)" proof - have "norm (a - b + b) \ norm (a - b) + norm b" by (rule norm_triangle_ineq) then show ?thesis by simp qed lemma norm_triangle_ineq3: "\norm a - norm b\ \ norm (a - b)" proof - have "norm a - norm b \ norm (a - b)" by (simp add: norm_triangle_ineq2) moreover have "norm b - norm a \ norm (a - b)" by (metis norm_minus_commute norm_triangle_ineq2) ultimately show ?thesis by (simp add: abs_le_iff) qed lemma norm_triangle_ineq4: "norm (a - b) \ norm a + norm b" proof - have "norm (a + - b) \ norm a + norm (- b)" by (rule norm_triangle_ineq) then show ?thesis by simp qed lemma norm_triangle_le_diff: "norm x + norm y \ e \ norm (x - y) \ e" by (meson norm_triangle_ineq4 order_trans) lemma norm_diff_ineq: "norm a - norm b \ norm (a + b)" proof - have "norm a - norm (- b) \ norm (a - - b)" by (rule norm_triangle_ineq2) then show ?thesis by simp qed lemma norm_triangle_sub: "norm x \ norm y + norm (x - y)" using norm_triangle_ineq[of "y" "x - y"] by (simp add: field_simps) lemma norm_triangle_le: "norm x + norm y \ e \ norm (x + y) \ e" by (rule norm_triangle_ineq [THEN order_trans]) lemma norm_triangle_lt: "norm x + norm y < e \ norm (x + y) < e" by (rule norm_triangle_ineq [THEN le_less_trans]) lemma norm_add_leD: "norm (a + b) \ c \ norm b \ norm a + c" by (metis ab_semigroup_add_class.add.commute add_commute diff_le_eq norm_diff_ineq order_trans) lemma norm_diff_triangle_ineq: "norm ((a + b) - (c + d)) \ norm (a - c) + norm (b - d)" proof - have "norm ((a + b) - (c + d)) = norm ((a - c) + (b - d))" by (simp add: algebra_simps) also have "\ \ norm (a - c) + norm (b - d)" by (rule norm_triangle_ineq) finally show ?thesis . qed lemma norm_diff_triangle_le: "norm (x - z) \ e1 + e2" if "norm (x - y) \ e1" "norm (y - z) \ e2" proof - have "norm (x - (y + z - y)) \ norm (x - y) + norm (y - z)" using norm_diff_triangle_ineq that diff_diff_eq2 by presburger with that show ?thesis by simp qed lemma norm_diff_triangle_less: "norm (x - z) < e1 + e2" if "norm (x - y) < e1" "norm (y - z) < e2" proof - have "norm (x - z) \ norm (x - y) + norm (y - z)" by (metis norm_diff_triangle_ineq add_diff_cancel_left' diff_diff_eq2) with that show ?thesis by auto qed lemma norm_triangle_mono: "norm a \ r \ norm b \ s \ norm (a + b) \ r + s" by (metis (mono_tags) add_mono_thms_linordered_semiring(1) norm_triangle_ineq order.trans) lemma norm_sum: "norm (sum f A) \ (\i\A. norm (f i))" for f::"'b \ 'a" by (induct A rule: infinite_finite_induct) (auto intro: norm_triangle_mono) lemma sum_norm_le: "norm (sum f S) \ sum g S" if "\x. x \ S \ norm (f x) \ g x" for f::"'b \ 'a" by (rule order_trans [OF norm_sum sum_mono]) (simp add: that) lemma abs_norm_cancel [simp]: "\norm a\ = norm a" by (rule abs_of_nonneg [OF norm_ge_zero]) lemma sum_norm_bound: "norm (sum f S) \ of_nat (card S)*K" if "\x. x \ S \ norm (f x) \ K" for f :: "'b \ 'a" using sum_norm_le[OF that] sum_constant[symmetric] by simp lemma norm_add_less: "norm x < r \ norm y < s \ norm (x + y) < r + s" by (rule order_le_less_trans [OF norm_triangle_ineq add_strict_mono]) end lemma dist_scaleR [simp]: "dist (x *\<^sub>R a) (y *\<^sub>R a) = \x - y\ * norm a" for a :: "'a::real_normed_vector" by (metis dist_norm norm_scaleR scaleR_left.diff) lemma norm_mult_less: "norm x < r \ norm y < s \ norm (x * y) < r * s" for x y :: "'a::real_normed_algebra" by (rule order_le_less_trans [OF norm_mult_ineq]) (simp add: mult_strict_mono') lemma norm_of_real [simp]: "norm (of_real r :: 'a::real_normed_algebra_1) = \r\" by (simp add: of_real_def) lemma norm_numeral [simp]: "norm (numeral w::'a::real_normed_algebra_1) = numeral w" by (subst of_real_numeral [symmetric], subst norm_of_real, simp) lemma norm_neg_numeral [simp]: "norm (- numeral w::'a::real_normed_algebra_1) = numeral w" by (subst of_real_neg_numeral [symmetric], subst norm_of_real, simp) lemma norm_of_real_add1 [simp]: "norm (of_real x + 1 :: 'a :: real_normed_div_algebra) = \x + 1\" by (metis norm_of_real of_real_1 of_real_add) lemma norm_of_real_addn [simp]: "norm (of_real x + numeral b :: 'a :: real_normed_div_algebra) = \x + numeral b\" by (metis norm_of_real of_real_add of_real_numeral) lemma norm_of_int [simp]: "norm (of_int z::'a::real_normed_algebra_1) = \of_int z\" by (subst of_real_of_int_eq [symmetric], rule norm_of_real) lemma norm_of_nat [simp]: "norm (of_nat n::'a::real_normed_algebra_1) = of_nat n" by (metis abs_of_nat norm_of_real of_real_of_nat_eq) lemma nonzero_norm_inverse: "a \ 0 \ norm (inverse a) = inverse (norm a)" for a :: "'a::real_normed_div_algebra" by (metis inverse_unique norm_mult norm_one right_inverse) lemma norm_inverse: "norm (inverse a) = inverse (norm a)" for a :: "'a::{real_normed_div_algebra,division_ring}" by (metis inverse_zero nonzero_norm_inverse norm_zero) lemma nonzero_norm_divide: "b \ 0 \ norm (a / b) = norm a / norm b" for a b :: "'a::real_normed_field" by (simp add: divide_inverse norm_mult nonzero_norm_inverse) lemma norm_divide: "norm (a / b) = norm a / norm b" for a b :: "'a::{real_normed_field,field}" by (simp add: divide_inverse norm_mult norm_inverse) lemma norm_inverse_le_norm: fixes x :: "'a::real_normed_div_algebra" shows "r \ norm x \ 0 < r \ norm (inverse x) \ inverse r" by (simp add: le_imp_inverse_le norm_inverse) lemma norm_power_ineq: "norm (x ^ n) \ norm x ^ n" for x :: "'a::real_normed_algebra_1" proof (induct n) case 0 show "norm (x ^ 0) \ norm x ^ 0" by simp next case (Suc n) have "norm (x * x ^ n) \ norm x * norm (x ^ n)" by (rule norm_mult_ineq) also from Suc have "\ \ norm x * norm x ^ n" using norm_ge_zero by (rule mult_left_mono) finally show "norm (x ^ Suc n) \ norm x ^ Suc n" by simp qed lemma norm_power: "norm (x ^ n) = norm x ^ n" for x :: "'a::real_normed_div_algebra" by (induct n) (simp_all add: norm_mult) lemma norm_power_int: "norm (power_int x n) = power_int (norm x) n" for x :: "'a::real_normed_div_algebra" by (cases n rule: int_cases4) (auto simp: norm_power power_int_minus norm_inverse) lemma power_eq_imp_eq_norm: fixes w :: "'a::real_normed_div_algebra" assumes eq: "w ^ n = z ^ n" and "n > 0" shows "norm w = norm z" proof - have "norm w ^ n = norm z ^ n" by (metis (no_types) eq norm_power) then show ?thesis using assms by (force intro: power_eq_imp_eq_base) qed lemma power_eq_1_iff: fixes w :: "'a::real_normed_div_algebra" shows "w ^ n = 1 \ norm w = 1 \ n = 0" by (metis norm_one power_0_left power_eq_0_iff power_eq_imp_eq_norm power_one) lemma norm_mult_numeral1 [simp]: "norm (numeral w * a) = numeral w * norm a" for a b :: "'a::{real_normed_field,field}" by (simp add: norm_mult) lemma norm_mult_numeral2 [simp]: "norm (a * numeral w) = norm a * numeral w" for a b :: "'a::{real_normed_field,field}" by (simp add: norm_mult) lemma norm_divide_numeral [simp]: "norm (a / numeral w) = norm a / numeral w" for a b :: "'a::{real_normed_field,field}" by (simp add: norm_divide) lemma norm_of_real_diff [simp]: "norm (of_real b - of_real a :: 'a::real_normed_algebra_1) \ \b - a\" by (metis norm_of_real of_real_diff order_refl) text \Despite a superficial resemblance, \norm_eq_1\ is not relevant.\ lemma square_norm_one: fixes x :: "'a::real_normed_div_algebra" assumes "x\<^sup>2 = 1" shows "norm x = 1" by (metis assms norm_minus_cancel norm_one power2_eq_1_iff) lemma norm_less_p1: "norm x < norm (of_real (norm x) + 1 :: 'a)" for x :: "'a::real_normed_algebra_1" proof - have "norm x < norm (of_real (norm x + 1) :: 'a)" by (simp add: of_real_def) then show ?thesis by simp qed lemma prod_norm: "prod (\x. norm (f x)) A = norm (prod f A)" for f :: "'a \ 'b::{comm_semiring_1,real_normed_div_algebra}" by (induct A rule: infinite_finite_induct) (auto simp: norm_mult) lemma norm_prod_le: "norm (prod f A) \ (\a\A. norm (f a :: 'a :: {real_normed_algebra_1,comm_monoid_mult}))" proof (induct A rule: infinite_finite_induct) case empty then show ?case by simp next case (insert a A) then have "norm (prod f (insert a A)) \ norm (f a) * norm (prod f A)" by (simp add: norm_mult_ineq) also have "norm (prod f A) \ (\a\A. norm (f a))" by (rule insert) finally show ?case by (simp add: insert mult_left_mono) next case infinite then show ?case by simp qed lemma norm_prod_diff: fixes z w :: "'i \ 'a::{real_normed_algebra_1, comm_monoid_mult}" shows "(\i. i \ I \ norm (z i) \ 1) \ (\i. i \ I \ norm (w i) \ 1) \ norm ((\i\I. z i) - (\i\I. w i)) \ (\i\I. norm (z i - w i))" proof (induction I rule: infinite_finite_induct) case empty then show ?case by simp next case (insert i I) note insert.hyps[simp] have "norm ((\i\insert i I. z i) - (\i\insert i I. w i)) = norm ((\i\I. z i) * (z i - w i) + ((\i\I. z i) - (\i\I. w i)) * w i)" (is "_ = norm (?t1 + ?t2)") by (auto simp: field_simps) also have "\ \ norm ?t1 + norm ?t2" by (rule norm_triangle_ineq) also have "norm ?t1 \ norm (\i\I. z i) * norm (z i - w i)" by (rule norm_mult_ineq) also have "\ \ (\i\I. norm (z i)) * norm(z i - w i)" by (rule mult_right_mono) (auto intro: norm_prod_le) also have "(\i\I. norm (z i)) \ (\i\I. 1)" by (intro prod_mono) (auto intro!: insert) also have "norm ?t2 \ norm ((\i\I. z i) - (\i\I. w i)) * norm (w i)" by (rule norm_mult_ineq) also have "norm (w i) \ 1" by (auto intro: insert) also have "norm ((\i\I. z i) - (\i\I. w i)) \ (\i\I. norm (z i - w i))" using insert by auto finally show ?case by (auto simp: ac_simps mult_right_mono mult_left_mono) next case infinite then show ?case by simp qed lemma norm_power_diff: fixes z w :: "'a::{real_normed_algebra_1, comm_monoid_mult}" assumes "norm z \ 1" "norm w \ 1" shows "norm (z^m - w^m) \ m * norm (z - w)" proof - have "norm (z^m - w^m) = norm ((\ i < m. z) - (\ i < m. w))" by simp also have "\ \ (\i = m * norm (z - w)" by simp finally show ?thesis . qed subsection \Metric spaces\ class metric_space = uniformity_dist + open_uniformity + assumes dist_eq_0_iff [simp]: "dist x y = 0 \ x = y" and dist_triangle2: "dist x y \ dist x z + dist y z" begin lemma dist_self [simp]: "dist x x = 0" by simp lemma zero_le_dist [simp]: "0 \ dist x y" using dist_triangle2 [of x x y] by simp lemma zero_less_dist_iff: "0 < dist x y \ x \ y" by (simp add: less_le) lemma dist_not_less_zero [simp]: "\ dist x y < 0" by (simp add: not_less) lemma dist_le_zero_iff [simp]: "dist x y \ 0 \ x = y" by (simp add: le_less) lemma dist_commute: "dist x y = dist y x" proof (rule order_antisym) show "dist x y \ dist y x" using dist_triangle2 [of x y x] by simp show "dist y x \ dist x y" using dist_triangle2 [of y x y] by simp qed lemma dist_commute_lessI: "dist y x < e \ dist x y < e" by (simp add: dist_commute) lemma dist_triangle: "dist x z \ dist x y + dist y z" using dist_triangle2 [of x z y] by (simp add: dist_commute) lemma dist_triangle3: "dist x y \ dist a x + dist a y" using dist_triangle2 [of x y a] by (simp add: dist_commute) lemma abs_dist_diff_le: "\dist a b - dist b c\ \ dist a c" using dist_triangle3[of b c a] dist_triangle2[of a b c] by simp lemma dist_pos_lt: "x \ y \ 0 < dist x y" by (simp add: zero_less_dist_iff) lemma dist_nz: "x \ y \ 0 < dist x y" by (simp add: zero_less_dist_iff) declare dist_nz [symmetric, simp] lemma dist_triangle_le: "dist x z + dist y z \ e \ dist x y \ e" by (rule order_trans [OF dist_triangle2]) lemma dist_triangle_lt: "dist x z + dist y z < e \ dist x y < e" by (rule le_less_trans [OF dist_triangle2]) lemma dist_triangle_less_add: "dist x1 y < e1 \ dist x2 y < e2 \ dist x1 x2 < e1 + e2" by (rule dist_triangle_lt [where z=y]) simp lemma dist_triangle_half_l: "dist x1 y < e / 2 \ dist x2 y < e / 2 \ dist x1 x2 < e" by (rule dist_triangle_lt [where z=y]) simp lemma dist_triangle_half_r: "dist y x1 < e / 2 \ dist y x2 < e / 2 \ dist x1 x2 < e" by (rule dist_triangle_half_l) (simp_all add: dist_commute) lemma dist_triangle_third: assumes "dist x1 x2 < e/3" "dist x2 x3 < e/3" "dist x3 x4 < e/3" shows "dist x1 x4 < e" proof - have "dist x1 x3 < e/3 + e/3" by (metis assms(1) assms(2) dist_commute dist_triangle_less_add) then have "dist x1 x4 < (e/3 + e/3) + e/3" by (metis assms(3) dist_commute dist_triangle_less_add) then show ?thesis by simp qed subclass uniform_space proof fix E x assume "eventually E uniformity" then obtain e where E: "0 < e" "\x y. dist x y < e \ E (x, y)" by (auto simp: eventually_uniformity_metric) then show "E (x, x)" "\\<^sub>F (x, y) in uniformity. E (y, x)" by (auto simp: eventually_uniformity_metric dist_commute) show "\D. eventually D uniformity \ (\x y z. D (x, y) \ D (y, z) \ E (x, z))" using E dist_triangle_half_l[where e=e] unfolding eventually_uniformity_metric by (intro exI[of _ "\(x, y). dist x y < e / 2"] exI[of _ "e/2"] conjI) (auto simp: dist_commute) qed lemma open_dist: "open S \ (\x\S. \e>0. \y. dist y x < e \ y \ S)" by (simp add: dist_commute open_uniformity eventually_uniformity_metric) lemma open_ball: "open {y. dist x y < d}" unfolding open_dist proof (intro ballI) fix y assume *: "y \ {y. dist x y < d}" then show "\e>0. \z. dist z y < e \ z \ {y. dist x y < d}" by (auto intro!: exI[of _ "d - dist x y"] simp: field_simps dist_triangle_lt) qed subclass first_countable_topology proof fix x show "\A::nat \ 'a set. (\i. x \ A i \ open (A i)) \ (\S. open S \ x \ S \ (\i. A i \ S))" proof (safe intro!: exI[of _ "\n. {y. dist x y < inverse (Suc n)}"]) fix S assume "open S" "x \ S" then obtain e where e: "0 < e" and "{y. dist x y < e} \ S" by (auto simp: open_dist subset_eq dist_commute) moreover from e obtain i where "inverse (Suc i) < e" by (auto dest!: reals_Archimedean) then have "{y. dist x y < inverse (Suc i)} \ {y. dist x y < e}" by auto ultimately show "\i. {y. dist x y < inverse (Suc i)} \ S" by blast qed (auto intro: open_ball) qed end instance metric_space \ t2_space proof fix x y :: "'a::metric_space" assume xy: "x \ y" let ?U = "{y'. dist x y' < dist x y / 2}" let ?V = "{x'. dist y x' < dist x y / 2}" have *: "d x z \ d x y + d y z \ d y z = d z y \ \ (d x y * 2 < d x z \ d z y * 2 < d x z)" for d :: "'a \ 'a \ real" and x y z :: 'a by arith have "open ?U \ open ?V \ x \ ?U \ y \ ?V \ ?U \ ?V = {}" using dist_pos_lt[OF xy] *[of dist, OF dist_triangle dist_commute] using open_ball[of _ "dist x y / 2"] by auto then show "\U V. open U \ open V \ x \ U \ y \ V \ U \ V = {}" by blast qed text \Every normed vector space is a metric space.\ instance real_normed_vector < metric_space proof fix x y z :: 'a show "dist x y = 0 \ x = y" by (simp add: dist_norm) show "dist x y \ dist x z + dist y z" using norm_triangle_ineq4 [of "x - z" "y - z"] by (simp add: dist_norm) qed subsection \Class instances for real numbers\ instantiation real :: real_normed_field begin definition dist_real_def: "dist x y = \x - y\" definition uniformity_real_def [code del]: "(uniformity :: (real \ real) filter) = (INF e\{0 <..}. principal {(x, y). dist x y < e})" definition open_real_def [code del]: "open (U :: real set) \ (\x\U. eventually (\(x', y). x' = x \ y \ U) uniformity)" definition real_norm_def [simp]: "norm r = \r\" instance by intro_classes (auto simp: abs_mult open_real_def dist_real_def sgn_real_def uniformity_real_def) end declare uniformity_Abort[where 'a=real, code] lemma dist_of_real [simp]: "dist (of_real x :: 'a) (of_real y) = dist x y" for a :: "'a::real_normed_div_algebra" by (metis dist_norm norm_of_real of_real_diff real_norm_def) declare [[code abort: "open :: real set \ bool"]] instance real :: linorder_topology proof show "(open :: real set \ bool) = generate_topology (range lessThan \ range greaterThan)" proof (rule ext, safe) fix S :: "real set" assume "open S" then obtain f where "\x\S. 0 < f x \ (\y. dist y x < f x \ y \ S)" unfolding open_dist bchoice_iff .. then have *: "(\x\S. {x - f x <..} \ {..< x + f x}) = S" (is "?S = S") by (fastforce simp: dist_real_def) moreover have "generate_topology (range lessThan \ range greaterThan) ?S" by (force intro: generate_topology.Basis generate_topology_Union generate_topology.Int) ultimately show "generate_topology (range lessThan \ range greaterThan) S" by simp next fix S :: "real set" assume "generate_topology (range lessThan \ range greaterThan) S" moreover have "\a::real. open {.. (\y. \y - x\ < a - x \ y \ {..e>0. \y. \y - x\ < e \ y \ {..a::real. open {a <..}" unfolding open_dist dist_real_def proof clarify fix x a :: real assume "a < x" then have "0 < x - a \ (\y. \y - x\ < x - a \ y \ {a<..})" by auto then show "\e>0. \y. \y - x\ < e \ y \ {a<..}" .. qed ultimately show "open S" by induct auto qed qed instance real :: linear_continuum_topology .. lemmas open_real_greaterThan = open_greaterThan[where 'a=real] lemmas open_real_lessThan = open_lessThan[where 'a=real] lemmas open_real_greaterThanLessThan = open_greaterThanLessThan[where 'a=real] lemmas closed_real_atMost = closed_atMost[where 'a=real] lemmas closed_real_atLeast = closed_atLeast[where 'a=real] lemmas closed_real_atLeastAtMost = closed_atLeastAtMost[where 'a=real] instance real :: ordered_real_vector by standard (auto intro: mult_left_mono mult_right_mono) subsection \Extra type constraints\ text \Only allow \<^term>\open\ in class \topological_space\.\ setup \Sign.add_const_constraint (\<^const_name>\open\, SOME \<^typ>\'a::topological_space set \ bool\)\ text \Only allow \<^term>\uniformity\ in class \uniform_space\.\ setup \Sign.add_const_constraint (\<^const_name>\uniformity\, SOME \<^typ>\('a::uniformity \ 'a) filter\)\ text \Only allow \<^term>\dist\ in class \metric_space\.\ setup \Sign.add_const_constraint (\<^const_name>\dist\, SOME \<^typ>\'a::metric_space \ 'a \ real\)\ text \Only allow \<^term>\norm\ in class \real_normed_vector\.\ setup \Sign.add_const_constraint (\<^const_name>\norm\, SOME \<^typ>\'a::real_normed_vector \ real\)\ subsection \Sign function\ lemma norm_sgn: "norm (sgn x) = (if x = 0 then 0 else 1)" for x :: "'a::real_normed_vector" by (simp add: sgn_div_norm) lemma sgn_zero [simp]: "sgn (0::'a::real_normed_vector) = 0" by (simp add: sgn_div_norm) lemma sgn_zero_iff: "sgn x = 0 \ x = 0" for x :: "'a::real_normed_vector" by (simp add: sgn_div_norm) lemma sgn_minus: "sgn (- x) = - sgn x" for x :: "'a::real_normed_vector" by (simp add: sgn_div_norm) lemma sgn_scaleR: "sgn (scaleR r x) = scaleR (sgn r) (sgn x)" for x :: "'a::real_normed_vector" by (simp add: sgn_div_norm ac_simps) lemma sgn_one [simp]: "sgn (1::'a::real_normed_algebra_1) = 1" by (simp add: sgn_div_norm) lemma sgn_of_real: "sgn (of_real r :: 'a::real_normed_algebra_1) = of_real (sgn r)" unfolding of_real_def by (simp only: sgn_scaleR sgn_one) lemma sgn_mult: "sgn (x * y) = sgn x * sgn y" for x y :: "'a::real_normed_div_algebra" by (simp add: sgn_div_norm norm_mult) hide_fact (open) sgn_mult lemma real_sgn_eq: "sgn x = x / \x\" for x :: real by (simp add: sgn_div_norm divide_inverse) lemma zero_le_sgn_iff [simp]: "0 \ sgn x \ 0 \ x" for x :: real by (cases "0::real" x rule: linorder_cases) simp_all lemma sgn_le_0_iff [simp]: "sgn x \ 0 \ x \ 0" for x :: real by (cases "0::real" x rule: linorder_cases) simp_all lemma norm_conv_dist: "norm x = dist x 0" unfolding dist_norm by simp declare norm_conv_dist [symmetric, simp] lemma dist_0_norm [simp]: "dist 0 x = norm x" for x :: "'a::real_normed_vector" by (simp add: dist_norm) lemma dist_diff [simp]: "dist a (a - b) = norm b" "dist (a - b) a = norm b" by (simp_all add: dist_norm) lemma dist_of_int: "dist (of_int m) (of_int n :: 'a :: real_normed_algebra_1) = of_int \m - n\" proof - have "dist (of_int m) (of_int n :: 'a) = dist (of_int m :: 'a) (of_int m - (of_int (m - n)))" by simp also have "\ = of_int \m - n\" by (subst dist_diff, subst norm_of_int) simp finally show ?thesis . qed lemma dist_of_nat: "dist (of_nat m) (of_nat n :: 'a :: real_normed_algebra_1) = of_int \int m - int n\" by (subst (1 2) of_int_of_nat_eq [symmetric]) (rule dist_of_int) subsection \Bounded Linear and Bilinear Operators\ lemma linearI: "linear f" if "\b1 b2. f (b1 + b2) = f b1 + f b2" "\r b. f (r *\<^sub>R b) = r *\<^sub>R f b" using that by unfold_locales (auto simp: algebra_simps) lemma linear_iff: "linear f \ (\x y. f (x + y) = f x + f y) \ (\c x. f (c *\<^sub>R x) = c *\<^sub>R f x)" (is "linear f \ ?rhs") proof assume "linear f" then interpret f: linear f . show "?rhs" by (simp add: f.add f.scale) next assume "?rhs" then show "linear f" by (intro linearI) auto qed lemmas linear_scaleR_left = linear_scale_left lemmas linear_imp_scaleR = linear_imp_scale corollary real_linearD: fixes f :: "real \ real" assumes "linear f" obtains c where "f = (*) c" by (rule linear_imp_scaleR [OF assms]) (force simp: scaleR_conv_of_real) lemma linear_times_of_real: "linear (\x. a * of_real x)" by (auto intro!: linearI simp: distrib_left) (metis mult_scaleR_right scaleR_conv_of_real) locale bounded_linear = linear f for f :: "'a::real_normed_vector \ 'b::real_normed_vector" + assumes bounded: "\K. \x. norm (f x) \ norm x * K" begin lemma pos_bounded: "\K>0. \x. norm (f x) \ norm x * K" proof - obtain K where K: "\x. norm (f x) \ norm x * K" using bounded by blast show ?thesis proof (intro exI impI conjI allI) show "0 < max 1 K" by (rule order_less_le_trans [OF zero_less_one max.cobounded1]) next fix x have "norm (f x) \ norm x * K" using K . also have "\ \ norm x * max 1 K" by (rule mult_left_mono [OF max.cobounded2 norm_ge_zero]) finally show "norm (f x) \ norm x * max 1 K" . qed qed lemma nonneg_bounded: "\K\0. \x. norm (f x) \ norm x * K" using pos_bounded by (auto intro: order_less_imp_le) lemma linear: "linear f" by (fact local.linear_axioms) end lemma bounded_linear_intro: assumes "\x y. f (x + y) = f x + f y" and "\r x. f (scaleR r x) = scaleR r (f x)" and "\x. norm (f x) \ norm x * K" shows "bounded_linear f" by standard (blast intro: assms)+ locale bounded_bilinear = fixes prod :: "'a::real_normed_vector \ 'b::real_normed_vector \ 'c::real_normed_vector" (infixl "**" 70) assumes add_left: "prod (a + a') b = prod a b + prod a' b" and add_right: "prod a (b + b') = prod a b + prod a b'" and scaleR_left: "prod (scaleR r a) b = scaleR r (prod a b)" and scaleR_right: "prod a (scaleR r b) = scaleR r (prod a b)" and bounded: "\K. \a b. norm (prod a b) \ norm a * norm b * K" begin lemma pos_bounded: "\K>0. \a b. norm (a ** b) \ norm a * norm b * K" proof - obtain K where "\a b. norm (a ** b) \ norm a * norm b * K" using bounded by blast then have "norm (a ** b) \ norm a * norm b * (max 1 K)" for a b by (rule order.trans) (simp add: mult_left_mono) then show ?thesis by force qed lemma nonneg_bounded: "\K\0. \a b. norm (a ** b) \ norm a * norm b * K" using pos_bounded by (auto intro: order_less_imp_le) lemma additive_right: "additive (\b. prod a b)" by (rule additive.intro, rule add_right) lemma additive_left: "additive (\a. prod a b)" by (rule additive.intro, rule add_left) lemma zero_left: "prod 0 b = 0" by (rule additive.zero [OF additive_left]) lemma zero_right: "prod a 0 = 0" by (rule additive.zero [OF additive_right]) lemma minus_left: "prod (- a) b = - prod a b" by (rule additive.minus [OF additive_left]) lemma minus_right: "prod a (- b) = - prod a b" by (rule additive.minus [OF additive_right]) lemma diff_left: "prod (a - a') b = prod a b - prod a' b" by (rule additive.diff [OF additive_left]) lemma diff_right: "prod a (b - b') = prod a b - prod a b'" by (rule additive.diff [OF additive_right]) lemma sum_left: "prod (sum g S) x = sum ((\i. prod (g i) x)) S" by (rule additive.sum [OF additive_left]) lemma sum_right: "prod x (sum g S) = sum ((\i. (prod x (g i)))) S" by (rule additive.sum [OF additive_right]) lemma bounded_linear_left: "bounded_linear (\a. a ** b)" proof - obtain K where "\a b. norm (a ** b) \ norm a * norm b * K" using pos_bounded by blast then show ?thesis by (rule_tac K="norm b * K" in bounded_linear_intro) (auto simp: algebra_simps scaleR_left add_left) qed lemma bounded_linear_right: "bounded_linear (\b. a ** b)" proof - obtain K where "\a b. norm (a ** b) \ norm a * norm b * K" using pos_bounded by blast then show ?thesis by (rule_tac K="norm a * K" in bounded_linear_intro) (auto simp: algebra_simps scaleR_right add_right) qed lemma prod_diff_prod: "(x ** y - a ** b) = (x - a) ** (y - b) + (x - a) ** b + a ** (y - b)" by (simp add: diff_left diff_right) lemma flip: "bounded_bilinear (\x y. y ** x)" proof show "\K. \a b. norm (b ** a) \ norm a * norm b * K" by (metis bounded mult.commute) qed (simp_all add: add_right add_left scaleR_right scaleR_left) lemma comp1: assumes "bounded_linear g" shows "bounded_bilinear (\x. (**) (g x))" proof unfold_locales interpret g: bounded_linear g by fact show "\a a' b. g (a + a') ** b = g a ** b + g a' ** b" "\a b b'. g a ** (b + b') = g a ** b + g a ** b'" "\r a b. g (r *\<^sub>R a) ** b = r *\<^sub>R (g a ** b)" "\a r b. g a ** (r *\<^sub>R b) = r *\<^sub>R (g a ** b)" by (auto simp: g.add add_left add_right g.scaleR scaleR_left scaleR_right) from g.nonneg_bounded nonneg_bounded obtain K L where nn: "0 \ K" "0 \ L" and K: "\x. norm (g x) \ norm x * K" and L: "\a b. norm (a ** b) \ norm a * norm b * L" by auto have "norm (g a ** b) \ norm a * K * norm b * L" for a b by (auto intro!: order_trans[OF K] order_trans[OF L] mult_mono simp: nn) then show "\K. \a b. norm (g a ** b) \ norm a * norm b * K" by (auto intro!: exI[where x="K * L"] simp: ac_simps) qed lemma comp: "bounded_linear f \ bounded_linear g \ bounded_bilinear (\x y. f x ** g y)" by (rule bounded_bilinear.flip[OF bounded_bilinear.comp1[OF bounded_bilinear.flip[OF comp1]]]) end lemma bounded_linear_ident[simp]: "bounded_linear (\x. x)" by standard (auto intro!: exI[of _ 1]) lemma bounded_linear_zero[simp]: "bounded_linear (\x. 0)" by standard (auto intro!: exI[of _ 1]) lemma bounded_linear_add: assumes "bounded_linear f" and "bounded_linear g" shows "bounded_linear (\x. f x + g x)" proof - interpret f: bounded_linear f by fact interpret g: bounded_linear g by fact show ?thesis proof from f.bounded obtain Kf where Kf: "norm (f x) \ norm x * Kf" for x by blast from g.bounded obtain Kg where Kg: "norm (g x) \ norm x * Kg" for x by blast show "\K. \x. norm (f x + g x) \ norm x * K" using add_mono[OF Kf Kg] by (intro exI[of _ "Kf + Kg"]) (auto simp: field_simps intro: norm_triangle_ineq order_trans) qed (simp_all add: f.add g.add f.scaleR g.scaleR scaleR_right_distrib) qed lemma bounded_linear_minus: assumes "bounded_linear f" shows "bounded_linear (\x. - f x)" proof - interpret f: bounded_linear f by fact show ?thesis by unfold_locales (simp_all add: f.add f.scaleR f.bounded) qed lemma bounded_linear_sub: "bounded_linear f \ bounded_linear g \ bounded_linear (\x. f x - g x)" using bounded_linear_add[of f "\x. - g x"] bounded_linear_minus[of g] by (auto simp: algebra_simps) lemma bounded_linear_sum: fixes f :: "'i \ 'a::real_normed_vector \ 'b::real_normed_vector" shows "(\i. i \ I \ bounded_linear (f i)) \ bounded_linear (\x. \i\I. f i x)" by (induct I rule: infinite_finite_induct) (auto intro!: bounded_linear_add) lemma bounded_linear_compose: assumes "bounded_linear f" and "bounded_linear g" shows "bounded_linear (\x. f (g x))" proof - interpret f: bounded_linear f by fact interpret g: bounded_linear g by fact show ?thesis proof unfold_locales show "f (g (x + y)) = f (g x) + f (g y)" for x y by (simp only: f.add g.add) show "f (g (scaleR r x)) = scaleR r (f (g x))" for r x by (simp only: f.scaleR g.scaleR) from f.pos_bounded obtain Kf where f: "\x. norm (f x) \ norm x * Kf" and Kf: "0 < Kf" by blast from g.pos_bounded obtain Kg where g: "\x. norm (g x) \ norm x * Kg" by blast show "\K. \x. norm (f (g x)) \ norm x * K" proof (intro exI allI) fix x have "norm (f (g x)) \ norm (g x) * Kf" using f . also have "\ \ (norm x * Kg) * Kf" using g Kf [THEN order_less_imp_le] by (rule mult_right_mono) also have "(norm x * Kg) * Kf = norm x * (Kg * Kf)" by (rule mult.assoc) finally show "norm (f (g x)) \ norm x * (Kg * Kf)" . qed qed qed lemma bounded_bilinear_mult: "bounded_bilinear ((*) :: 'a \ 'a \ 'a::real_normed_algebra)" proof (rule bounded_bilinear.intro) show "\K. \a b::'a. norm (a * b) \ norm a * norm b * K" by (rule_tac x=1 in exI) (simp add: norm_mult_ineq) qed (auto simp: algebra_simps) lemma bounded_linear_mult_left: "bounded_linear (\x::'a::real_normed_algebra. x * y)" using bounded_bilinear_mult by (rule bounded_bilinear.bounded_linear_left) lemma bounded_linear_mult_right: "bounded_linear (\y::'a::real_normed_algebra. x * y)" using bounded_bilinear_mult by (rule bounded_bilinear.bounded_linear_right) lemmas bounded_linear_mult_const = bounded_linear_mult_left [THEN bounded_linear_compose] lemmas bounded_linear_const_mult = bounded_linear_mult_right [THEN bounded_linear_compose] lemma bounded_linear_divide: "bounded_linear (\x. x / y)" for y :: "'a::real_normed_field" unfolding divide_inverse by (rule bounded_linear_mult_left) lemma bounded_bilinear_scaleR: "bounded_bilinear scaleR" proof (rule bounded_bilinear.intro) show "\K. \a b. norm (a *\<^sub>R b) \ norm a * norm b * K" using less_eq_real_def by auto qed (auto simp: algebra_simps) lemma bounded_linear_scaleR_left: "bounded_linear (\r. scaleR r x)" using bounded_bilinear_scaleR by (rule bounded_bilinear.bounded_linear_left) lemma bounded_linear_scaleR_right: "bounded_linear (\x. scaleR r x)" using bounded_bilinear_scaleR by (rule bounded_bilinear.bounded_linear_right) lemmas bounded_linear_scaleR_const = bounded_linear_scaleR_left[THEN bounded_linear_compose] lemmas bounded_linear_const_scaleR = bounded_linear_scaleR_right[THEN bounded_linear_compose] lemma bounded_linear_of_real: "bounded_linear (\r. of_real r)" unfolding of_real_def by (rule bounded_linear_scaleR_left) lemma real_bounded_linear: "bounded_linear f \ (\c::real. f = (\x. x * c))" for f :: "real \ real" proof - { fix x assume "bounded_linear f" then interpret bounded_linear f . from scaleR[of x 1] have "f x = x * f 1" by simp } then show ?thesis by (auto intro: exI[of _ "f 1"] bounded_linear_mult_left) qed instance real_normed_algebra_1 \ perfect_space proof fix x::'a have "\e. 0 < e \ \y. norm (y - x) < e \ y \ x" by (rule_tac x = "x + of_real (e/2)" in exI) auto then show "\ open {x}" by (clarsimp simp: open_dist dist_norm) qed subsection \Filters and Limits on Metric Space\ lemma (in metric_space) nhds_metric: "nhds x = (INF e\{0 <..}. principal {y. dist y x < e})" unfolding nhds_def proof (safe intro!: INF_eq) fix S assume "open S" "x \ S" then obtain e where "{y. dist y x < e} \ S" "0 < e" by (auto simp: open_dist subset_eq) then show "\e\{0<..}. principal {y. dist y x < e} \ principal S" by auto qed (auto intro!: exI[of _ "{y. dist x y < e}" for e] open_ball simp: dist_commute) (* Contributed by Dominique Unruh *) lemma tendsto_iff_uniformity: \ \More general analogus of \tendsto_iff\ below. Applies to all uniform spaces, not just metric ones.\ fixes l :: \'b :: uniform_space\ shows \(f \ l) F \ (\E. eventually E uniformity \ (\\<^sub>F x in F. E (f x, l)))\ proof (intro iffI allI impI) fix E :: \('b \ 'b) \ bool\ assume \(f \ l) F\ and \eventually E uniformity\ from \eventually E uniformity\ have \eventually (\(x, y). E (y, x)) uniformity\ by (simp add: uniformity_sym) then have \\\<^sub>F (y, x) in uniformity. y = l \ E (x, y)\ using eventually_mono by fastforce with \(f \ l) F\ have \eventually (\x. E (x ,l)) (filtermap f F)\ by (simp add: filterlim_def le_filter_def eventually_nhds_uniformity) then show \\\<^sub>F x in F. E (f x, l)\ by (simp add: eventually_filtermap) next assume assm: \\E. eventually E uniformity \ (\\<^sub>F x in F. E (f x, l))\ have \eventually P (filtermap f F)\ if \\\<^sub>F (x, y) in uniformity. x = l \ P y\ for P proof - from that have \\\<^sub>F (y, x) in uniformity. x = l \ P y\ using uniformity_sym[where E=\\(x,y). x=l \ P y\] by auto with assm have \\\<^sub>F x in F. P (f x)\ by auto then show ?thesis by (auto simp: eventually_filtermap) qed then show \(f \ l) F\ by (simp add: filterlim_def le_filter_def eventually_nhds_uniformity) qed lemma (in metric_space) tendsto_iff: "(f \ l) F \ (\e>0. eventually (\x. dist (f x) l < e) F)" unfolding nhds_metric filterlim_INF filterlim_principal by auto lemma tendsto_dist_iff: "((f \ l) F) \ (((\x. dist (f x) l) \ 0) F)" unfolding tendsto_iff by simp lemma (in metric_space) tendstoI [intro?]: "(\e. 0 < e \ eventually (\x. dist (f x) l < e) F) \ (f \ l) F" by (auto simp: tendsto_iff) lemma (in metric_space) tendstoD: "(f \ l) F \ 0 < e \ eventually (\x. dist (f x) l < e) F" by (auto simp: tendsto_iff) lemma (in metric_space) eventually_nhds_metric: "eventually P (nhds a) \ (\d>0. \x. dist x a < d \ P x)" unfolding nhds_metric by (subst eventually_INF_base) (auto simp: eventually_principal Bex_def subset_eq intro: exI[of _ "min a b" for a b]) lemma eventually_at: "eventually P (at a within S) \ (\d>0. \x\S. x \ a \ dist x a < d \ P x)" for a :: "'a :: metric_space" by (auto simp: eventually_at_filter eventually_nhds_metric) lemma frequently_at: "frequently P (at a within S) \ (\d>0. \x\S. x \ a \ dist x a < d \ P x)" for a :: "'a :: metric_space" unfolding frequently_def eventually_at by auto lemma eventually_at_le: "eventually P (at a within S) \ (\d>0. \x\S. x \ a \ dist x a \ d \ P x)" for a :: "'a::metric_space" unfolding eventually_at_filter eventually_nhds_metric apply safe apply (rule_tac x="d / 2" in exI, auto) done lemma eventually_at_left_real: "a > (b :: real) \ eventually (\x. x \ {b<.. eventually (\x. x \ {a<.. a) F" and le: "eventually (\x. dist (g x) b \ dist (f x) a) F" shows "(g \ b) F" proof (rule tendstoI) fix e :: real assume "0 < e" with f have "eventually (\x. dist (f x) a < e) F" by (rule tendstoD) with le show "eventually (\x. dist (g x) b < e) F" using le_less_trans by (rule eventually_elim2) qed lemma filterlim_real_sequentially: "LIM x sequentially. real x :> at_top" proof (clarsimp simp: filterlim_at_top) fix Z show "\\<^sub>F x in sequentially. Z \ real x" by (meson eventually_sequentiallyI nat_ceiling_le_eq) qed lemma filterlim_nat_sequentially: "filterlim nat sequentially at_top" proof - have "\\<^sub>F x in at_top. Z \ nat x" for Z by (auto intro!: eventually_at_top_linorderI[where c="int Z"]) then show ?thesis unfolding filterlim_at_top .. qed lemma filterlim_floor_sequentially: "filterlim floor at_top at_top" proof - have "\\<^sub>F x in at_top. Z \ \x\" for Z by (auto simp: le_floor_iff intro!: eventually_at_top_linorderI[where c="of_int Z"]) then show ?thesis unfolding filterlim_at_top .. qed lemma filterlim_sequentially_iff_filterlim_real: "filterlim f sequentially F \ filterlim (\x. real (f x)) at_top F" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs using filterlim_compose filterlim_real_sequentially by blast next assume R: ?rhs show ?lhs proof - have "filterlim (\x. nat (floor (real (f x)))) sequentially F" by (intro filterlim_compose[OF filterlim_nat_sequentially] filterlim_compose[OF filterlim_floor_sequentially] R) then show ?thesis by simp qed qed subsubsection \Limits of Sequences\ lemma lim_sequentially: "X \ L \ (\r>0. \no. \n\no. dist (X n) L < r)" for L :: "'a::metric_space" unfolding tendsto_iff eventually_sequentially .. lemmas LIMSEQ_def = lim_sequentially (*legacy binding*) lemma LIMSEQ_iff_nz: "X \ L \ (\r>0. \no>0. \n\no. dist (X n) L < r)" for L :: "'a::metric_space" unfolding lim_sequentially by (metis Suc_leD zero_less_Suc) lemma metric_LIMSEQ_I: "(\r. 0 < r \ \no. \n\no. dist (X n) L < r) \ X \ L" for L :: "'a::metric_space" by (simp add: lim_sequentially) lemma metric_LIMSEQ_D: "X \ L \ 0 < r \ \no. \n\no. dist (X n) L < r" for L :: "'a::metric_space" by (simp add: lim_sequentially) lemma LIMSEQ_norm_0: assumes "\n::nat. norm (f n) < 1 / real (Suc n)" shows "f \ 0" proof (rule metric_LIMSEQ_I) fix \ :: "real" assume "\ > 0" then obtain N::nat where "\ > inverse N" "N > 0" by (metis neq0_conv real_arch_inverse) then have "norm (f n) < \" if "n \ N" for n proof - have "1 / (Suc n) \ 1 / N" using \0 < N\ inverse_of_nat_le le_SucI that by blast also have "\ < \" by (metis (no_types) \inverse (real N) < \\ inverse_eq_divide) finally show ?thesis by (meson assms less_eq_real_def not_le order_trans) qed then show "\no. \n\no. dist (f n) 0 < \" by auto qed subsubsection \Limits of Functions\ lemma LIM_def: "f \a\ L \ (\r > 0. \s > 0. \x. x \ a \ dist x a < s \ dist (f x) L < r)" for a :: "'a::metric_space" and L :: "'b::metric_space" unfolding tendsto_iff eventually_at by simp lemma metric_LIM_I: "(\r. 0 < r \ \s>0. \x. x \ a \ dist x a < s \ dist (f x) L < r) \ f \a\ L" for a :: "'a::metric_space" and L :: "'b::metric_space" by (simp add: LIM_def) lemma metric_LIM_D: "f \a\ L \ 0 < r \ \s>0. \x. x \ a \ dist x a < s \ dist (f x) L < r" for a :: "'a::metric_space" and L :: "'b::metric_space" by (simp add: LIM_def) lemma metric_LIM_imp_LIM: fixes l :: "'a::metric_space" and m :: "'b::metric_space" assumes f: "f \a\ l" and le: "\x. x \ a \ dist (g x) m \ dist (f x) l" shows "g \a\ m" by (rule metric_tendsto_imp_tendsto [OF f]) (auto simp: eventually_at_topological le) lemma metric_LIM_equal2: fixes a :: "'a::metric_space" assumes "g \a\ l" "0 < R" and "\x. x \ a \ dist x a < R \ f x = g x" shows "f \a\ l" proof - have "\S. \open S; l \ S; \\<^sub>F x in at a. g x \ S\ \ \\<^sub>F x in at a. f x \ S" apply (simp add: eventually_at) by (metis assms(2) assms(3) dual_order.strict_trans linorder_neqE_linordered_idom) then show ?thesis using assms by (simp add: tendsto_def) qed lemma metric_LIM_compose2: fixes a :: "'a::metric_space" assumes f: "f \a\ b" and g: "g \b\ c" and inj: "\d>0. \x. x \ a \ dist x a < d \ f x \ b" shows "(\x. g (f x)) \a\ c" using inj by (intro tendsto_compose_eventually[OF g f]) (auto simp: eventually_at) lemma metric_isCont_LIM_compose2: fixes f :: "'a :: metric_space \ _" assumes f [unfolded isCont_def]: "isCont f a" and g: "g \f a\ l" and inj: "\d>0. \x. x \ a \ dist x a < d \ f x \ f a" shows "(\x. g (f x)) \a\ l" by (rule metric_LIM_compose2 [OF f g inj]) subsection \Complete metric spaces\ subsection \Cauchy sequences\ lemma (in metric_space) Cauchy_def: "Cauchy X = (\e>0. \M. \m\M. \n\M. dist (X m) (X n) < e)" proof - have *: "eventually P (INF M. principal {(X m, X n) | n m. m \ M \ n \ M}) \ (\M. \m\M. \n\M. P (X m, X n))" for P apply (subst eventually_INF_base) subgoal by simp subgoal for a b by (intro bexI[of _ "max a b"]) (auto simp: eventually_principal subset_eq) subgoal by (auto simp: eventually_principal, blast) done have "Cauchy X \ (INF M. principal {(X m, X n) | n m. m \ M \ n \ M}) \ uniformity" unfolding Cauchy_uniform_iff le_filter_def * .. also have "\ = (\e>0. \M. \m\M. \n\M. dist (X m) (X n) < e)" unfolding uniformity_dist le_INF_iff by (auto simp: * le_principal) finally show ?thesis . qed lemma (in metric_space) Cauchy_altdef: "Cauchy f \ (\e>0. \M. \m\M. \n>m. dist (f m) (f n) < e)" (is "?lhs \ ?rhs") proof assume ?rhs show ?lhs unfolding Cauchy_def proof (intro allI impI) fix e :: real assume e: "e > 0" with \?rhs\ obtain M where M: "m \ M \ n > m \ dist (f m) (f n) < e" for m n by blast have "dist (f m) (f n) < e" if "m \ M" "n \ M" for m n using M[of m n] M[of n m] e that by (cases m n rule: linorder_cases) (auto simp: dist_commute) then show "\M. \m\M. \n\M. dist (f m) (f n) < e" by blast qed next assume ?lhs show ?rhs proof (intro allI impI) fix e :: real assume e: "e > 0" with \Cauchy f\ obtain M where "\m n. m \ M \ n \ M \ dist (f m) (f n) < e" unfolding Cauchy_def by blast then show "\M. \m\M. \n>m. dist (f m) (f n) < e" by (intro exI[of _ M]) force qed qed lemma (in metric_space) Cauchy_altdef2: "Cauchy s \ (\e>0. \N::nat. \n\N. dist(s n)(s N) < e)" (is "?lhs = ?rhs") proof assume "Cauchy s" then show ?rhs by (force simp: Cauchy_def) next assume ?rhs { fix e::real assume "e>0" with \?rhs\ obtain N where N: "\n\N. dist (s n) (s N) < e/2" by (erule_tac x="e/2" in allE) auto { fix n m assume nm: "N \ m \ N \ n" then have "dist (s m) (s n) < e" using N using dist_triangle_half_l[of "s m" "s N" "e" "s n"] by blast } then have "\N. \m n. N \ m \ N \ n \ dist (s m) (s n) < e" by blast } then have ?lhs unfolding Cauchy_def by blast then show ?lhs by blast qed lemma (in metric_space) metric_CauchyI: "(\e. 0 < e \ \M. \m\M. \n\M. dist (X m) (X n) < e) \ Cauchy X" by (simp add: Cauchy_def) lemma (in metric_space) CauchyI': "(\e. 0 < e \ \M. \m\M. \n>m. dist (X m) (X n) < e) \ Cauchy X" unfolding Cauchy_altdef by blast lemma (in metric_space) metric_CauchyD: "Cauchy X \ 0 < e \ \M. \m\M. \n\M. dist (X m) (X n) < e" by (simp add: Cauchy_def) lemma (in metric_space) metric_Cauchy_iff2: "Cauchy X = (\j. (\M. \m \ M. \n \ M. dist (X m) (X n) < inverse(real (Suc j))))" apply (auto simp add: Cauchy_def) by (metis less_trans of_nat_Suc reals_Archimedean) lemma Cauchy_iff2: "Cauchy X \ (\j. (\M. \m \ M. \n \ M. \X m - X n\ < inverse (real (Suc j))))" by (simp only: metric_Cauchy_iff2 dist_real_def) lemma lim_1_over_n [tendsto_intros]: "((\n. 1 / of_nat n) \ (0::'a::real_normed_field)) sequentially" proof (subst lim_sequentially, intro allI impI exI) fix e::real and n assume e: "e > 0" have "inverse e < of_nat (nat \inverse e + 1\)" by linarith also assume "n \ nat \inverse e + 1\" finally show "dist (1 / of_nat n :: 'a) 0 < e" using e by (simp add: field_split_simps norm_divide) qed lemma (in metric_space) complete_def: shows "complete S = (\f. (\n. f n \ S) \ Cauchy f \ (\l\S. f \ l))" unfolding complete_uniform proof safe fix f :: "nat \ 'a" assume f: "\n. f n \ S" "Cauchy f" and *: "\F\principal S. F \ bot \ cauchy_filter F \ (\x\S. F \ nhds x)" then show "\l\S. f \ l" unfolding filterlim_def using f by (intro *[rule_format]) (auto simp: filtermap_sequentually_ne_bot le_principal eventually_filtermap Cauchy_uniform) next fix F :: "'a filter" assume "F \ principal S" "F \ bot" "cauchy_filter F" assume seq: "\f. (\n. f n \ S) \ Cauchy f \ (\l\S. f \ l)" from \F \ principal S\ \cauchy_filter F\ have FF_le: "F \\<^sub>F F \ uniformity_on S" by (simp add: cauchy_filter_def principal_prod_principal[symmetric] prod_filter_mono) let ?P = "\P e. eventually P F \ (\x. P x \ x \ S) \ (\x y. P x \ P y \ dist x y < e)" have P: "\P. ?P P \" if "0 < \" for \ :: real proof - from that have "eventually (\(x, y). x \ S \ y \ S \ dist x y < \) (uniformity_on S)" by (auto simp: eventually_inf_principal eventually_uniformity_metric) from filter_leD[OF FF_le this] show ?thesis by (auto simp: eventually_prod_same) qed have "\P. \n. ?P (P n) (1 / Suc n) \ P (Suc n) \ P n" proof (rule dependent_nat_choice) show "\P. ?P P (1 / Suc 0)" using P[of 1] by auto next fix P n assume "?P P (1/Suc n)" moreover obtain Q where "?P Q (1 / Suc (Suc n))" using P[of "1/Suc (Suc n)"] by auto ultimately show "\Q. ?P Q (1 / Suc (Suc n)) \ Q \ P" by (intro exI[of _ "\x. P x \ Q x"]) (auto simp: eventually_conj_iff) qed then obtain P where P: "eventually (P n) F" "P n x \ x \ S" "P n x \ P n y \ dist x y < 1 / Suc n" "P (Suc n) \ P n" for n x y by metis have "antimono P" using P(4) by (rule decseq_SucI) obtain X where X: "P n (X n)" for n using P(1)[THEN eventually_happens'[OF \F \ bot\]] by metis have "Cauchy X" unfolding metric_Cauchy_iff2 inverse_eq_divide proof (intro exI allI impI) fix j m n :: nat assume "j \ m" "j \ n" with \antimono P\ X have "P j (X m)" "P j (X n)" by (auto simp: antimono_def) then show "dist (X m) (X n) < 1 / Suc j" by (rule P) qed moreover have "\n. X n \ S" using P(2) X by auto ultimately obtain x where "X \ x" "x \ S" using seq by blast show "\x\S. F \ nhds x" proof (rule bexI) have "eventually (\y. dist y x < e) F" if "0 < e" for e :: real proof - from that have "(\n. 1 / Suc n :: real) \ 0 \ 0 < e / 2" by (subst filterlim_sequentially_Suc) (auto intro!: lim_1_over_n) then have "\\<^sub>F n in sequentially. dist (X n) x < e / 2 \ 1 / Suc n < e / 2" using \X \ x\ unfolding tendsto_iff order_tendsto_iff[where 'a=real] eventually_conj_iff by blast then obtain n where "dist x (X n) < e / 2" "1 / Suc n < e / 2" by (auto simp: eventually_sequentially dist_commute) show ?thesis using \eventually (P n) F\ proof eventually_elim case (elim y) then have "dist y (X n) < 1 / Suc n" by (intro X P) also have "\ < e / 2" by fact finally show "dist y x < e" by (rule dist_triangle_half_l) fact qed qed then show "F \ nhds x" unfolding nhds_metric le_INF_iff le_principal by auto qed fact qed text\apparently unused\ lemma (in metric_space) totally_bounded_metric: "totally_bounded S \ (\e>0. \k. finite k \ S \ (\x\k. {y. dist x y < e}))" unfolding totally_bounded_def eventually_uniformity_metric imp_ex apply (subst all_comm) apply (intro arg_cong[where f=All] ext, safe) subgoal for e apply (erule allE[of _ "\(x, y). dist x y < e"]) apply auto done subgoal for e P k apply (intro exI[of _ k]) apply (force simp: subset_eq) done done setup \Sign.add_const_constraint (\<^const_name>\dist\, SOME \<^typ>\'a::dist \ 'a \ real\)\ (* Contributed by Dominique Unruh *) lemma cauchy_filter_metric: fixes F :: "'a::{uniformity_dist,uniform_space} filter" shows "cauchy_filter F \ (\e. e>0 \ (\P. eventually P F \ (\x y. P x \ P y \ dist x y < e)))" proof (unfold cauchy_filter_def le_filter_def, auto) assume assm: \\e>0. \P. eventually P F \ (\x y. P x \ P y \ dist x y < e)\ then show \eventually P uniformity \ eventually P (F \\<^sub>F F)\ for P apply (auto simp: eventually_uniformity_metric) using eventually_prod_same by blast next fix e :: real assume \e > 0\ assume asm: \\P. eventually P uniformity \ eventually P (F \\<^sub>F F)\ define P where \P \ \(x,y :: 'a). dist x y < e\ with asm \e > 0\ have \eventually P (F \\<^sub>F F)\ by (metis case_prod_conv eventually_uniformity_metric) then show \\P. eventually P F \ (\x y. P x \ P y \ dist x y < e)\ by (auto simp add: eventually_prod_same P_def) qed (* Contributed by Dominique Unruh *) lemma cauchy_filter_metric_filtermap: fixes f :: "'a \ 'b::{uniformity_dist,uniform_space}" shows "cauchy_filter (filtermap f F) \ (\e. e>0 \ (\P. eventually P F \ (\x y. P x \ P y \ dist (f x) (f y) < e)))" proof (subst cauchy_filter_metric, intro iffI allI impI) assume \\e>0. \P. eventually P (filtermap f F) \ (\x y. P x \ P y \ dist x y < e)\ then show \e>0 \ \P. eventually P F \ (\x y. P x \ P y \ dist (f x) (f y) < e)\ for e unfolding eventually_filtermap by blast next assume asm: \\e>0. \P. eventually P F \ (\x y. P x \ P y \ dist (f x) (f y) < e)\ fix e::real assume \e > 0\ then obtain P where \eventually P F\ and PPe: \P x \ P y \ dist (f x) (f y) < e\ for x y using asm by blast show \\P. eventually P (filtermap f F) \ (\x y. P x \ P y \ dist x y < e)\ apply (rule exI[of _ \\x. \y. P y \ x = f y\]) using PPe \eventually P F\ apply (auto simp: eventually_filtermap) by (smt (verit, ccfv_SIG) eventually_elim2) qed setup \Sign.add_const_constraint (\<^const_name>\dist\, SOME \<^typ>\'a::metric_space \ 'a \ real\)\ subsubsection \Cauchy Sequences are Convergent\ (* TODO: update to uniform_space *) class complete_space = metric_space + assumes Cauchy_convergent: "Cauchy X \ convergent X" lemma Cauchy_convergent_iff: "Cauchy X \ convergent X" for X :: "nat \ 'a::complete_space" by (blast intro: Cauchy_convergent convergent_Cauchy) text \To prove that a Cauchy sequence converges, it suffices to show that a subsequence converges.\ lemma Cauchy_converges_subseq: fixes u::"nat \ 'a::metric_space" assumes "Cauchy u" "strict_mono r" "(u \ r) \ l" shows "u \ l" proof - have *: "eventually (\n. dist (u n) l < e) sequentially" if "e > 0" for e proof - have "e/2 > 0" using that by auto then obtain N1 where N1: "\m n. m \ N1 \ n \ N1 \ dist (u m) (u n) < e/2" using \Cauchy u\ unfolding Cauchy_def by blast obtain N2 where N2: "\n. n \ N2 \ dist ((u \ r) n) l < e / 2" using order_tendstoD(2)[OF iffD1[OF tendsto_dist_iff \(u \ r) \ l\] \e/2 > 0\] unfolding eventually_sequentially by auto have "dist (u n) l < e" if "n \ max N1 N2" for n proof - have "dist (u n) l \ dist (u n) ((u \ r) n) + dist ((u \ r) n) l" by (rule dist_triangle) also have "\ < e/2 + e/2" proof (intro add_strict_mono) show "dist (u n) ((u \ r) n) < e / 2" using N1[of n "r n"] N2[of n] that unfolding comp_def by (meson assms(2) le_trans max.bounded_iff strict_mono_imp_increasing) show "dist ((u \ r) n) l < e / 2" using N2 that by auto qed finally show ?thesis by simp qed then show ?thesis unfolding eventually_sequentially by blast qed have "(\n. dist (u n) l) \ 0" by (simp add: less_le_trans * order_tendstoI) then show ?thesis using tendsto_dist_iff by auto qed subsection \The set of real numbers is a complete metric space\ text \ Proof that Cauchy sequences converge based on the one from \<^url>\http://pirate.shu.edu/~wachsmut/ira/numseq/proofs/cauconv.html\ \ text \ If sequence \<^term>\X\ is Cauchy, then its limit is the lub of \<^term>\{r::real. \N. \n\N. r < X n}\ \ lemma increasing_LIMSEQ: fixes f :: "nat \ real" assumes inc: "\n. f n \ f (Suc n)" and bdd: "\n. f n \ l" and en: "\e. 0 < e \ \n. l \ f n + e" shows "f \ l" proof (rule increasing_tendsto) fix x assume "x < l" with dense[of 0 "l - x"] obtain e where "0 < e" "e < l - x" by auto from en[OF \0 < e\] obtain n where "l - e \ f n" by (auto simp: field_simps) with \e < l - x\ \0 < e\ have "x < f n" by simp with incseq_SucI[of f, OF inc] show "eventually (\n. x < f n) sequentially" by (auto simp: eventually_sequentially incseq_def intro: less_le_trans) qed (use bdd in auto) lemma real_Cauchy_convergent: fixes X :: "nat \ real" assumes X: "Cauchy X" shows "convergent X" proof - define S :: "real set" where "S = {x. \N. \n\N. x < X n}" then have mem_S: "\N x. \n\N. x < X n \ x \ S" by auto have bound_isUb: "y \ x" if N: "\n\N. X n < x" and "y \ S" for N and x y :: real proof - from that have "\M. \n\M. y < X n" by (simp add: S_def) then obtain M where "\n\M. y < X n" .. then have "y < X (max M N)" by simp also have "\ < x" using N by simp finally show ?thesis by (rule order_less_imp_le) qed obtain N where "\m\N. \n\N. dist (X m) (X n) < 1" using X[THEN metric_CauchyD, OF zero_less_one] by auto then have N: "\n\N. dist (X n) (X N) < 1" by simp have [simp]: "S \ {}" proof (intro exI ex_in_conv[THEN iffD1]) from N have "\n\N. X N - 1 < X n" by (simp add: abs_diff_less_iff dist_real_def) then show "X N - 1 \ S" by (rule mem_S) qed have [simp]: "bdd_above S" proof from N have "\n\N. X n < X N + 1" by (simp add: abs_diff_less_iff dist_real_def) then show "\s. s \ S \ s \ X N + 1" by (rule bound_isUb) qed have "X \ Sup S" proof (rule metric_LIMSEQ_I) fix r :: real assume "0 < r" then have r: "0 < r/2" by simp obtain N where "\n\N. \m\N. dist (X n) (X m) < r/2" using metric_CauchyD [OF X r] by auto then have "\n\N. dist (X n) (X N) < r/2" by simp then have N: "\n\N. X N - r/2 < X n \ X n < X N + r/2" by (simp only: dist_real_def abs_diff_less_iff) from N have "\n\N. X N - r/2 < X n" by blast then have "X N - r/2 \ S" by (rule mem_S) then have 1: "X N - r/2 \ Sup S" by (simp add: cSup_upper) from N have "\n\N. X n < X N + r/2" by blast from bound_isUb[OF this] have 2: "Sup S \ X N + r/2" by (intro cSup_least) simp_all show "\N. \n\N. dist (X n) (Sup S) < r" proof (intro exI allI impI) fix n assume n: "N \ n" from N n have "X n < X N + r/2" and "X N - r/2 < X n" by simp_all then show "dist (X n) (Sup S) < r" using 1 2 by (simp add: abs_diff_less_iff dist_real_def) qed qed then show ?thesis by (auto simp: convergent_def) qed instance real :: complete_space by intro_classes (rule real_Cauchy_convergent) class banach = real_normed_vector + complete_space instance real :: banach .. lemma tendsto_at_topI_sequentially: fixes f :: "real \ 'b::first_countable_topology" assumes *: "\X. filterlim X at_top sequentially \ (\n. f (X n)) \ y" shows "(f \ y) at_top" proof - obtain A where A: "decseq A" "open (A n)" "y \ A n" "nhds y = (INF n. principal (A n))" for n by (rule nhds_countable[of y]) (rule that) have "\m. \k. \x\k. f x \ A m" proof (rule ccontr) assume "\ (\m. \k. \x\k. f x \ A m)" then obtain m where "\k. \x\k. f x \ A m" by auto then have "\X. \n. (f (X n) \ A m) \ max n (X n) + 1 \ X (Suc n)" by (intro dependent_nat_choice) (auto simp del: max.bounded_iff) then obtain X where X: "\n. f (X n) \ A m" "\n. max n (X n) + 1 \ X (Suc n)" by auto have "1 \ n \ real n \ X n" for n using X[of "n - 1"] by auto then have "filterlim X at_top sequentially" by (force intro!: filterlim_at_top_mono[OF filterlim_real_sequentially] simp: eventually_sequentially) from topological_tendstoD[OF *[OF this] A(2, 3), of m] X(1) show False by auto qed then obtain k where "k m \ x \ f x \ A m" for m x by metis then show ?thesis unfolding at_top_def A by (intro filterlim_base[where i=k]) auto qed lemma tendsto_at_topI_sequentially_real: fixes f :: "real \ real" assumes mono: "mono f" and limseq: "(\n. f (real n)) \ y" shows "(f \ y) at_top" proof (rule tendstoI) fix e :: real assume "0 < e" with limseq obtain N :: nat where N: "N \ n \ \f (real n) - y\ < e" for n by (auto simp: lim_sequentially dist_real_def) have le: "f x \ y" for x :: real proof - obtain n where "x \ real_of_nat n" using real_arch_simple[of x] .. note monoD[OF mono this] also have "f (real_of_nat n) \ y" by (rule LIMSEQ_le_const[OF limseq]) (auto intro!: exI[of _ n] monoD[OF mono]) finally show ?thesis . qed have "eventually (\x. real N \ x) at_top" by (rule eventually_ge_at_top) then show "eventually (\x. dist (f x) y < e) at_top" proof eventually_elim case (elim x) with N[of N] le have "y - f (real N) < e" by auto moreover note monoD[OF mono elim] ultimately show "dist (f x) y < e" using le[of x] by (auto simp: dist_real_def field_simps) qed qed end diff --git a/src/HOL/Topological_Spaces.thy b/src/HOL/Topological_Spaces.thy --- a/src/HOL/Topological_Spaces.thy +++ b/src/HOL/Topological_Spaces.thy @@ -1,3819 +1,3833 @@ (* Title: HOL/Topological_Spaces.thy Author: Brian Huffman Author: Johannes Hölzl *) section \Topological Spaces\ theory Topological_Spaces imports Main begin named_theorems continuous_intros "structural introduction rules for continuity" subsection \Topological space\ class "open" = fixes "open" :: "'a set \ bool" class topological_space = "open" + assumes open_UNIV [simp, intro]: "open UNIV" assumes open_Int [intro]: "open S \ open T \ open (S \ T)" assumes open_Union [intro]: "\S\K. open S \ open (\K)" begin definition closed :: "'a set \ bool" where "closed S \ open (- S)" lemma open_empty [continuous_intros, intro, simp]: "open {}" using open_Union [of "{}"] by simp lemma open_Un [continuous_intros, intro]: "open S \ open T \ open (S \ T)" using open_Union [of "{S, T}"] by simp lemma open_UN [continuous_intros, intro]: "\x\A. open (B x) \ open (\x\A. B x)" using open_Union [of "B ` A"] by simp lemma open_Inter [continuous_intros, intro]: "finite S \ \T\S. open T \ open (\S)" by (induction set: finite) auto lemma open_INT [continuous_intros, intro]: "finite A \ \x\A. open (B x) \ open (\x\A. B x)" using open_Inter [of "B ` A"] by simp lemma openI: assumes "\x. x \ S \ \T. open T \ x \ T \ T \ S" shows "open S" proof - have "open (\{T. open T \ T \ S})" by auto moreover have "\{T. open T \ T \ S} = S" by (auto dest!: assms) ultimately show "open S" by simp qed lemma open_subopen: "open S \ (\x\S. \T. open T \ x \ T \ T \ S)" by (auto intro: openI) lemma closed_empty [continuous_intros, intro, simp]: "closed {}" unfolding closed_def by simp lemma closed_Un [continuous_intros, intro]: "closed S \ closed T \ closed (S \ T)" unfolding closed_def by auto lemma closed_UNIV [continuous_intros, intro, simp]: "closed UNIV" unfolding closed_def by simp lemma closed_Int [continuous_intros, intro]: "closed S \ closed T \ closed (S \ T)" unfolding closed_def by auto lemma closed_INT [continuous_intros, intro]: "\x\A. closed (B x) \ closed (\x\A. B x)" unfolding closed_def by auto lemma closed_Inter [continuous_intros, intro]: "\S\K. closed S \ closed (\K)" unfolding closed_def uminus_Inf by auto lemma closed_Union [continuous_intros, intro]: "finite S \ \T\S. closed T \ closed (\S)" by (induct set: finite) auto lemma closed_UN [continuous_intros, intro]: "finite A \ \x\A. closed (B x) \ closed (\x\A. B x)" using closed_Union [of "B ` A"] by simp lemma open_closed: "open S \ closed (- S)" by (simp add: closed_def) lemma closed_open: "closed S \ open (- S)" by (rule closed_def) lemma open_Diff [continuous_intros, intro]: "open S \ closed T \ open (S - T)" by (simp add: closed_open Diff_eq open_Int) lemma closed_Diff [continuous_intros, intro]: "closed S \ open T \ closed (S - T)" by (simp add: open_closed Diff_eq closed_Int) lemma open_Compl [continuous_intros, intro]: "closed S \ open (- S)" by (simp add: closed_open) lemma closed_Compl [continuous_intros, intro]: "open S \ closed (- S)" by (simp add: open_closed) lemma open_Collect_neg: "closed {x. P x} \ open {x. \ P x}" unfolding Collect_neg_eq by (rule open_Compl) lemma open_Collect_conj: assumes "open {x. P x}" "open {x. Q x}" shows "open {x. P x \ Q x}" using open_Int[OF assms] by (simp add: Int_def) lemma open_Collect_disj: assumes "open {x. P x}" "open {x. Q x}" shows "open {x. P x \ Q x}" using open_Un[OF assms] by (simp add: Un_def) lemma open_Collect_ex: "(\i. open {x. P i x}) \ open {x. \i. P i x}" using open_UN[of UNIV "\i. {x. P i x}"] unfolding Collect_ex_eq by simp lemma open_Collect_imp: "closed {x. P x} \ open {x. Q x} \ open {x. P x \ Q x}" unfolding imp_conv_disj by (intro open_Collect_disj open_Collect_neg) lemma open_Collect_const: "open {x. P}" by (cases P) auto lemma closed_Collect_neg: "open {x. P x} \ closed {x. \ P x}" unfolding Collect_neg_eq by (rule closed_Compl) lemma closed_Collect_conj: assumes "closed {x. P x}" "closed {x. Q x}" shows "closed {x. P x \ Q x}" using closed_Int[OF assms] by (simp add: Int_def) lemma closed_Collect_disj: assumes "closed {x. P x}" "closed {x. Q x}" shows "closed {x. P x \ Q x}" using closed_Un[OF assms] by (simp add: Un_def) lemma closed_Collect_all: "(\i. closed {x. P i x}) \ closed {x. \i. P i x}" using closed_INT[of UNIV "\i. {x. P i x}"] by (simp add: Collect_all_eq) lemma closed_Collect_imp: "open {x. P x} \ closed {x. Q x} \ closed {x. P x \ Q x}" unfolding imp_conv_disj by (intro closed_Collect_disj closed_Collect_neg) lemma closed_Collect_const: "closed {x. P}" by (cases P) auto end subsection \Hausdorff and other separation properties\ class t0_space = topological_space + assumes t0_space: "x \ y \ \U. open U \ \ (x \ U \ y \ U)" class t1_space = topological_space + assumes t1_space: "x \ y \ \U. open U \ x \ U \ y \ U" instance t1_space \ t0_space by standard (fast dest: t1_space) context t1_space begin lemma separation_t1: "x \ y \ (\U. open U \ x \ U \ y \ U)" using t1_space[of x y] by blast lemma closed_singleton [iff]: "closed {a}" proof - let ?T = "\{S. open S \ a \ S}" have "open ?T" by (simp add: open_Union) also have "?T = - {a}" by (auto simp add: set_eq_iff separation_t1) finally show "closed {a}" by (simp only: closed_def) qed lemma closed_insert [continuous_intros, simp]: assumes "closed S" shows "closed (insert a S)" proof - from closed_singleton assms have "closed ({a} \ S)" by (rule closed_Un) then show "closed (insert a S)" by simp qed lemma finite_imp_closed: "finite S \ closed S" by (induct pred: finite) simp_all end text \T2 spaces are also known as Hausdorff spaces.\ class t2_space = topological_space + assumes hausdorff: "x \ y \ \U V. open U \ open V \ x \ U \ y \ V \ U \ V = {}" instance t2_space \ t1_space by standard (fast dest: hausdorff) lemma (in t2_space) separation_t2: "x \ y \ (\U V. open U \ open V \ x \ U \ y \ V \ U \ V = {})" using hausdorff [of x y] by blast lemma (in t0_space) separation_t0: "x \ y \ (\U. open U \ \ (x \ U \ y \ U))" using t0_space [of x y] by blast text \A classical separation axiom for topological space, the T3 axiom -- also called regularity: if a point is not in a closed set, then there are open sets separating them.\ class t3_space = t2_space + assumes t3_space: "closed S \ y \ S \ \U V. open U \ open V \ y \ U \ S \ V \ U \ V = {}" text \A classical separation axiom for topological space, the T4 axiom -- also called normality: if two closed sets are disjoint, then there are open sets separating them.\ class t4_space = t2_space + assumes t4_space: "closed S \ closed T \ S \ T = {} \ \U V. open U \ open V \ S \ U \ T \ V \ U \ V = {}" text \T4 is stronger than T3, and weaker than metric.\ instance t4_space \ t3_space proof fix S and y::'a assume "closed S" "y \ S" then show "\U V. open U \ open V \ y \ U \ S \ V \ U \ V = {}" using t4_space[of "{y}" S] by auto qed text \A perfect space is a topological space with no isolated points.\ class perfect_space = topological_space + assumes not_open_singleton: "\ open {x}" lemma (in perfect_space) UNIV_not_singleton: "UNIV \ {x}" for x::'a by (metis (no_types) open_UNIV not_open_singleton) subsection \Generators for toplogies\ inductive generate_topology :: "'a set set \ 'a set \ bool" for S :: "'a set set" where UNIV: "generate_topology S UNIV" | Int: "generate_topology S (a \ b)" if "generate_topology S a" and "generate_topology S b" | UN: "generate_topology S (\K)" if "(\k. k \ K \ generate_topology S k)" | Basis: "generate_topology S s" if "s \ S" hide_fact (open) UNIV Int UN Basis lemma generate_topology_Union: "(\k. k \ I \ generate_topology S (K k)) \ generate_topology S (\k\I. K k)" using generate_topology.UN [of "K ` I"] by auto lemma topological_space_generate_topology: "class.topological_space (generate_topology S)" by standard (auto intro: generate_topology.intros) subsection \Order topologies\ class order_topology = order + "open" + assumes open_generated_order: "open = generate_topology (range (\a. {..< a}) \ range (\a. {a <..}))" begin subclass topological_space unfolding open_generated_order by (rule topological_space_generate_topology) lemma open_greaterThan [continuous_intros, simp]: "open {a <..}" unfolding open_generated_order by (auto intro: generate_topology.Basis) lemma open_lessThan [continuous_intros, simp]: "open {..< a}" unfolding open_generated_order by (auto intro: generate_topology.Basis) lemma open_greaterThanLessThan [continuous_intros, simp]: "open {a <..< b}" unfolding greaterThanLessThan_eq by (simp add: open_Int) end class linorder_topology = linorder + order_topology lemma closed_atMost [continuous_intros, simp]: "closed {..a}" for a :: "'a::linorder_topology" by (simp add: closed_open) lemma closed_atLeast [continuous_intros, simp]: "closed {a..}" for a :: "'a::linorder_topology" by (simp add: closed_open) lemma closed_atLeastAtMost [continuous_intros, simp]: "closed {a..b}" for a b :: "'a::linorder_topology" proof - have "{a .. b} = {a ..} \ {.. b}" by auto then show ?thesis by (simp add: closed_Int) qed lemma (in order) less_separate: assumes "x < y" shows "\a b. x \ {..< a} \ y \ {b <..} \ {..< a} \ {b <..} = {}" proof (cases "\z. x < z \ z < y") case True then obtain z where "x < z \ z < y" .. then have "x \ {..< z} \ y \ {z <..} \ {z <..} \ {..< z} = {}" by auto then show ?thesis by blast next case False with \x < y\ have "x \ {..< y}" "y \ {x <..}" "{x <..} \ {..< y} = {}" by auto then show ?thesis by blast qed instance linorder_topology \ t2_space proof fix x y :: 'a show "x \ y \ \U V. open U \ open V \ x \ U \ y \ V \ U \ V = {}" using less_separate [of x y] less_separate [of y x] by (elim neqE; metis open_lessThan open_greaterThan Int_commute) qed lemma (in linorder_topology) open_right: assumes "open S" "x \ S" and gt_ex: "x < y" shows "\b>x. {x ..< b} \ S" using assms unfolding open_generated_order proof induct case UNIV then show ?case by blast next case (Int A B) then obtain a b where "a > x" "{x ..< a} \ A" "b > x" "{x ..< b} \ B" by auto then show ?case by (auto intro!: exI[of _ "min a b"]) next case UN then show ?case by blast next case Basis then show ?case by (fastforce intro: exI[of _ y] gt_ex) qed lemma (in linorder_topology) open_left: assumes "open S" "x \ S" and lt_ex: "y < x" shows "\b S" using assms unfolding open_generated_order proof induction case UNIV then show ?case by blast next case (Int A B) then obtain a b where "a < x" "{a <.. x} \ A" "b < x" "{b <.. x} \ B" by auto then show ?case by (auto intro!: exI[of _ "max a b"]) next case UN then show ?case by blast next case Basis then show ?case by (fastforce intro: exI[of _ y] lt_ex) qed subsection \Setup some topologies\ subsubsection \Boolean is an order topology\ class discrete_topology = topological_space + assumes open_discrete: "\A. open A" instance discrete_topology < t2_space proof fix x y :: 'a assume "x \ y" then show "\U V. open U \ open V \ x \ U \ y \ V \ U \ V = {}" by (intro exI[of _ "{_}"]) (auto intro!: open_discrete) qed instantiation bool :: linorder_topology begin definition open_bool :: "bool set \ bool" where "open_bool = generate_topology (range (\a. {..< a}) \ range (\a. {a <..}))" instance by standard (rule open_bool_def) end instance bool :: discrete_topology proof fix A :: "bool set" have *: "{False <..} = {True}" "{..< True} = {False}" by auto have "A = UNIV \ A = {} \ A = {False <..} \ A = {..< True}" using subset_UNIV[of A] unfolding UNIV_bool * by blast then show "open A" by auto qed instantiation nat :: linorder_topology begin definition open_nat :: "nat set \ bool" where "open_nat = generate_topology (range (\a. {..< a}) \ range (\a. {a <..}))" instance by standard (rule open_nat_def) end instance nat :: discrete_topology proof fix A :: "nat set" have "open {n}" for n :: nat proof (cases n) case 0 moreover have "{0} = {..<1::nat}" by auto ultimately show ?thesis by auto next case (Suc n') then have "{n} = {.. {n' <..}" by auto with Suc show ?thesis by (auto intro: open_lessThan open_greaterThan) qed then have "open (\a\A. {a})" by (intro open_UN) auto then show "open A" by simp qed instantiation int :: linorder_topology begin definition open_int :: "int set \ bool" where "open_int = generate_topology (range (\a. {..< a}) \ range (\a. {a <..}))" instance by standard (rule open_int_def) end instance int :: discrete_topology proof fix A :: "int set" have "{.. {i-1 <..} = {i}" for i :: int by auto then have "open {i}" for i :: int using open_Int[OF open_lessThan[of "i + 1"] open_greaterThan[of "i - 1"]] by auto then have "open (\a\A. {a})" by (intro open_UN) auto then show "open A" by simp qed subsubsection \Topological filters\ definition (in topological_space) nhds :: "'a \ 'a filter" where "nhds a = (INF S\{S. open S \ a \ S}. principal S)" definition (in topological_space) at_within :: "'a \ 'a set \ 'a filter" ("at (_)/ within (_)" [1000, 60] 60) where "at a within s = inf (nhds a) (principal (s - {a}))" abbreviation (in topological_space) at :: "'a \ 'a filter" ("at") where "at x \ at x within (CONST UNIV)" abbreviation (in order_topology) at_right :: "'a \ 'a filter" where "at_right x \ at x within {x <..}" abbreviation (in order_topology) at_left :: "'a \ 'a filter" where "at_left x \ at x within {..< x}" lemma (in topological_space) nhds_generated_topology: "open = generate_topology T \ nhds x = (INF S\{S\T. x \ S}. principal S)" unfolding nhds_def proof (safe intro!: antisym INF_greatest) fix S assume "generate_topology T S" "x \ S" then show "(INF S\{S \ T. x \ S}. principal S) \ principal S" by induct (auto intro: INF_lower order_trans simp: inf_principal[symmetric] simp del: inf_principal) qed (auto intro!: INF_lower intro: generate_topology.intros) lemma (in topological_space) eventually_nhds: "eventually P (nhds a) \ (\S. open S \ a \ S \ (\x\S. P x))" unfolding nhds_def by (subst eventually_INF_base) (auto simp: eventually_principal) lemma eventually_eventually: "eventually (\y. eventually P (nhds y)) (nhds x) = eventually P (nhds x)" by (auto simp: eventually_nhds) lemma (in topological_space) eventually_nhds_in_open: "open s \ x \ s \ eventually (\y. y \ s) (nhds x)" by (subst eventually_nhds) blast lemma (in topological_space) eventually_nhds_x_imp_x: "eventually P (nhds x) \ P x" by (subst (asm) eventually_nhds) blast lemma (in topological_space) nhds_neq_bot [simp]: "nhds a \ bot" by (simp add: trivial_limit_def eventually_nhds) lemma (in t1_space) t1_space_nhds: "x \ y \ (\\<^sub>F x in nhds x. x \ y)" by (drule t1_space) (auto simp: eventually_nhds) lemma (in topological_space) nhds_discrete_open: "open {x} \ nhds x = principal {x}" by (auto simp: nhds_def intro!: antisym INF_greatest INF_lower2[of "{x}"]) lemma (in discrete_topology) nhds_discrete: "nhds x = principal {x}" by (simp add: nhds_discrete_open open_discrete) lemma (in discrete_topology) at_discrete: "at x within S = bot" unfolding at_within_def nhds_discrete by simp lemma (in discrete_topology) tendsto_discrete: "filterlim (f :: 'b \ 'a) (nhds y) F \ eventually (\x. f x = y) F" by (auto simp: nhds_discrete filterlim_principal) lemma (in topological_space) at_within_eq: "at x within s = (INF S\{S. open S \ x \ S}. principal (S \ s - {x}))" unfolding nhds_def at_within_def by (subst INF_inf_const2[symmetric]) (auto simp: Diff_Int_distrib) lemma (in topological_space) eventually_at_filter: "eventually P (at a within s) \ eventually (\x. x \ a \ x \ s \ P x) (nhds a)" by (simp add: at_within_def eventually_inf_principal imp_conjL[symmetric] conj_commute) lemma (in topological_space) at_le: "s \ t \ at x within s \ at x within t" unfolding at_within_def by (intro inf_mono) auto lemma (in topological_space) eventually_at_topological: "eventually P (at a within s) \ (\S. open S \ a \ S \ (\x\S. x \ a \ x \ s \ P x))" by (simp add: eventually_nhds eventually_at_filter) +lemma eventually_at_in_open: + assumes "open A" "x \ A" + shows "eventually (\y. y \ A - {x}) (at x)" + using assms eventually_at_topological by blast + +lemma eventually_at_in_open': + assumes "open A" "x \ A" + shows "eventually (\y. y \ A) (at x)" + using assms eventually_at_topological by blast + lemma (in topological_space) at_within_open: "a \ S \ open S \ at a within S = at a" unfolding filter_eq_iff eventually_at_topological by (metis open_Int Int_iff UNIV_I) lemma (in topological_space) at_within_open_NO_MATCH: "a \ s \ open s \ NO_MATCH UNIV s \ at a within s = at a" by (simp only: at_within_open) lemma (in topological_space) at_within_open_subset: "a \ S \ open S \ S \ T \ at a within T = at a" by (metis at_le at_within_open dual_order.antisym subset_UNIV) lemma (in topological_space) at_within_nhd: assumes "x \ S" "open S" "T \ S - {x} = U \ S - {x}" shows "at x within T = at x within U" unfolding filter_eq_iff eventually_at_filter proof (intro allI eventually_subst) have "eventually (\x. x \ S) (nhds x)" using \x \ S\ \open S\ by (auto simp: eventually_nhds) then show "\\<^sub>F n in nhds x. (n \ x \ n \ T \ P n) = (n \ x \ n \ U \ P n)" for P by eventually_elim (insert \T \ S - {x} = U \ S - {x}\, blast) qed lemma (in topological_space) at_within_empty [simp]: "at a within {} = bot" unfolding at_within_def by simp lemma (in topological_space) at_within_union: "at x within (S \ T) = sup (at x within S) (at x within T)" unfolding filter_eq_iff eventually_sup eventually_at_filter by (auto elim!: eventually_rev_mp) lemma (in topological_space) at_eq_bot_iff: "at a = bot \ open {a}" unfolding trivial_limit_def eventually_at_topological by (metis UNIV_I empty_iff is_singletonE is_singletonI' singleton_iff) +lemma (in t1_space) eventually_neq_at_within: + "eventually (\w. w \ x) (at z within A)" + by (smt (verit, ccfv_threshold) eventually_True eventually_at_topological separation_t1) + lemma (in perfect_space) at_neq_bot [simp]: "at a \ bot" by (simp add: at_eq_bot_iff not_open_singleton) lemma (in order_topology) nhds_order: "nhds x = inf (INF a\{x <..}. principal {..< a}) (INF a\{..< x}. principal {a <..})" proof - have 1: "{S \ range lessThan \ range greaterThan. x \ S} = (\a. {..< a}) ` {x <..} \ (\a. {a <..}) ` {..< x}" by auto show ?thesis by (simp only: nhds_generated_topology[OF open_generated_order] INF_union 1 INF_image comp_def) qed lemma (in topological_space) filterlim_at_within_If: assumes "filterlim f G (at x within (A \ {x. P x}))" and "filterlim g G (at x within (A \ {x. \P x}))" shows "filterlim (\x. if P x then f x else g x) G (at x within A)" proof (rule filterlim_If) note assms(1) also have "at x within (A \ {x. P x}) = inf (nhds x) (principal (A \ Collect P - {x}))" by (simp add: at_within_def) also have "A \ Collect P - {x} = (A - {x}) \ Collect P" by blast also have "inf (nhds x) (principal \) = inf (at x within A) (principal (Collect P))" by (simp add: at_within_def inf_assoc) finally show "filterlim f G (inf (at x within A) (principal (Collect P)))" . next note assms(2) also have "at x within (A \ {x. \ P x}) = inf (nhds x) (principal (A \ {x. \ P x} - {x}))" by (simp add: at_within_def) also have "A \ {x. \ P x} - {x} = (A - {x}) \ {x. \ P x}" by blast also have "inf (nhds x) (principal \) = inf (at x within A) (principal {x. \ P x})" by (simp add: at_within_def inf_assoc) finally show "filterlim g G (inf (at x within A) (principal {x. \ P x}))" . qed lemma (in topological_space) filterlim_at_If: assumes "filterlim f G (at x within {x. P x})" and "filterlim g G (at x within {x. \P x})" shows "filterlim (\x. if P x then f x else g x) G (at x)" using assms by (intro filterlim_at_within_If) simp_all lemma (in linorder_topology) at_within_order: assumes "UNIV \ {x}" shows "at x within s = inf (INF a\{x <..}. principal ({..< a} \ s - {x})) (INF a\{..< x}. principal ({a <..} \ s - {x}))" proof (cases "{x <..} = {}" "{..< x} = {}" rule: case_split [case_product case_split]) case True_True have "UNIV = {..< x} \ {x} \ {x <..}" by auto with assms True_True show ?thesis by auto qed (auto simp del: inf_principal simp: at_within_def nhds_order Int_Diff inf_principal[symmetric] INF_inf_const2 inf_sup_aci[where 'a="'a filter"]) lemma (in linorder_topology) at_left_eq: "y < x \ at_left x = (INF a\{..< x}. principal {a <..< x})" by (subst at_within_order) (auto simp: greaterThan_Int_greaterThan greaterThanLessThan_eq[symmetric] min.absorb2 INF_constant intro!: INF_lower2 inf_absorb2) lemma (in linorder_topology) eventually_at_left: "y < x \ eventually P (at_left x) \ (\by>b. y < x \ P y)" unfolding at_left_eq by (subst eventually_INF_base) (auto simp: eventually_principal Ball_def) lemma (in linorder_topology) at_right_eq: "x < y \ at_right x = (INF a\{x <..}. principal {x <..< a})" by (subst at_within_order) (auto simp: lessThan_Int_lessThan greaterThanLessThan_eq[symmetric] max.absorb2 INF_constant Int_commute intro!: INF_lower2 inf_absorb1) lemma (in linorder_topology) eventually_at_right: "x < y \ eventually P (at_right x) \ (\b>x. \y>x. y < b \ P y)" unfolding at_right_eq by (subst eventually_INF_base) (auto simp: eventually_principal Ball_def) lemma eventually_at_right_less: "\\<^sub>F y in at_right (x::'a::{linorder_topology, no_top}). x < y" using gt_ex[of x] eventually_at_right[of x] by auto lemma trivial_limit_at_right_top: "at_right (top::_::{order_top,linorder_topology}) = bot" by (auto simp: filter_eq_iff eventually_at_topological) lemma trivial_limit_at_left_bot: "at_left (bot::_::{order_bot,linorder_topology}) = bot" by (auto simp: filter_eq_iff eventually_at_topological) lemma trivial_limit_at_left_real [simp]: "\ trivial_limit (at_left x)" for x :: "'a::{no_bot,dense_order,linorder_topology}" using lt_ex [of x] by safe (auto simp add: trivial_limit_def eventually_at_left dest: dense) lemma trivial_limit_at_right_real [simp]: "\ trivial_limit (at_right x)" for x :: "'a::{no_top,dense_order,linorder_topology}" using gt_ex[of x] by safe (auto simp add: trivial_limit_def eventually_at_right dest: dense) lemma (in linorder_topology) at_eq_sup_left_right: "at x = sup (at_left x) (at_right x)" by (auto simp: eventually_at_filter filter_eq_iff eventually_sup elim: eventually_elim2 eventually_mono) lemma (in linorder_topology) eventually_at_split: "eventually P (at x) \ eventually P (at_left x) \ eventually P (at_right x)" by (subst at_eq_sup_left_right) (simp add: eventually_sup) lemma (in order_topology) eventually_at_leftI: assumes "\x. x \ {a<.. P x" "a < b" shows "eventually P (at_left b)" using assms unfolding eventually_at_topological by (intro exI[of _ "{a<..}"]) auto lemma (in order_topology) eventually_at_rightI: assumes "\x. x \ {a<.. P x" "a < b" shows "eventually P (at_right a)" using assms unfolding eventually_at_topological by (intro exI[of _ "{.. (\S. open S \ x \ S \ (\x. f x \ S \ P x))" unfolding eventually_filtercomap eventually_nhds by auto lemma eventually_filtercomap_at_topological: "eventually P (filtercomap f (at A within B)) \ (\S. open S \ A \ S \ (\x. f x \ S \ B - {A} \ P x))" (is "?lhs = ?rhs") unfolding at_within_def filtercomap_inf eventually_inf_principal filtercomap_principal eventually_filtercomap_nhds eventually_principal by blast lemma eventually_at_right_field: "eventually P (at_right x) \ (\b>x. \y>x. y < b \ P y)" for x :: "'a::{linordered_field, linorder_topology}" using linordered_field_no_ub[rule_format, of x] by (auto simp: eventually_at_right) lemma eventually_at_left_field: "eventually P (at_left x) \ (\by>b. y < x \ P y)" for x :: "'a::{linordered_field, linorder_topology}" using linordered_field_no_lb[rule_format, of x] by (auto simp: eventually_at_left) subsubsection \Tendsto\ abbreviation (in topological_space) tendsto :: "('b \ 'a) \ 'a \ 'b filter \ bool" (infixr "\" 55) where "(f \ l) F \ filterlim f (nhds l) F" definition (in t2_space) Lim :: "'f filter \ ('f \ 'a) \ 'a" where "Lim A f = (THE l. (f \ l) A)" lemma (in topological_space) tendsto_eq_rhs: "(f \ x) F \ x = y \ (f \ y) F" by simp named_theorems tendsto_intros "introduction rules for tendsto" setup \ Global_Theory.add_thms_dynamic (\<^binding>\tendsto_eq_intros\, fn context => Named_Theorems.get (Context.proof_of context) \<^named_theorems>\tendsto_intros\ |> map_filter (try (fn thm => @{thm tendsto_eq_rhs} OF [thm]))) \ context topological_space begin lemma tendsto_def: "(f \ l) F \ (\S. open S \ l \ S \ eventually (\x. f x \ S) F)" unfolding nhds_def filterlim_INF filterlim_principal by auto lemma tendsto_cong: "(f \ c) F \ (g \ c) F" if "eventually (\x. f x = g x) F" by (rule filterlim_cong [OF refl refl that]) lemma tendsto_mono: "F \ F' \ (f \ l) F' \ (f \ l) F" unfolding tendsto_def le_filter_def by fast lemma tendsto_ident_at [tendsto_intros, simp, intro]: "((\x. x) \ a) (at a within s)" by (auto simp: tendsto_def eventually_at_topological) lemma tendsto_const [tendsto_intros, simp, intro]: "((\x. k) \ k) F" by (simp add: tendsto_def) lemma filterlim_at: "(LIM x F. f x :> at b within s) \ eventually (\x. f x \ s \ f x \ b) F \ (f \ b) F" by (simp add: at_within_def filterlim_inf filterlim_principal conj_commute) lemma (in -) assumes "filterlim f (nhds L) F" shows tendsto_imp_filterlim_at_right: "eventually (\x. f x > L) F \ filterlim f (at_right L) F" and tendsto_imp_filterlim_at_left: "eventually (\x. f x < L) F \ filterlim f (at_left L) F" using assms by (auto simp: filterlim_at elim: eventually_mono) lemma filterlim_at_withinI: assumes "filterlim f (nhds c) F" assumes "eventually (\x. f x \ A - {c}) F" shows "filterlim f (at c within A) F" using assms by (simp add: filterlim_at) lemma filterlim_atI: assumes "filterlim f (nhds c) F" assumes "eventually (\x. f x \ c) F" shows "filterlim f (at c) F" using assms by (intro filterlim_at_withinI) simp_all lemma topological_tendstoI: "(\S. open S \ l \ S \ eventually (\x. f x \ S) F) \ (f \ l) F" by (auto simp: tendsto_def) lemma topological_tendstoD: "(f \ l) F \ open S \ l \ S \ eventually (\x. f x \ S) F" by (auto simp: tendsto_def) lemma tendsto_bot [simp]: "(f \ a) bot" by (simp add: tendsto_def) lemma tendsto_eventually: "eventually (\x. f x = l) net \ ((\x. f x) \ l) net" by (rule topological_tendstoI) (auto elim: eventually_mono) (* Contributed by Dominique Unruh *) lemma tendsto_principal_singleton[simp]: shows "(f \ f x) (principal {x})" unfolding tendsto_def eventually_principal by simp end lemma (in topological_space) filterlim_within_subset: "filterlim f l (at x within S) \ T \ S \ filterlim f l (at x within T)" by (blast intro: filterlim_mono at_le) lemmas tendsto_within_subset = filterlim_within_subset lemma (in order_topology) order_tendsto_iff: "(f \ x) F \ (\lx. l < f x) F) \ (\u>x. eventually (\x. f x < u) F)" by (auto simp: nhds_order filterlim_inf filterlim_INF filterlim_principal) lemma (in order_topology) order_tendstoI: "(\a. a < y \ eventually (\x. a < f x) F) \ (\a. y < a \ eventually (\x. f x < a) F) \ (f \ y) F" by (auto simp: order_tendsto_iff) lemma (in order_topology) order_tendstoD: assumes "(f \ y) F" shows "a < y \ eventually (\x. a < f x) F" and "y < a \ eventually (\x. f x < a) F" using assms by (auto simp: order_tendsto_iff) lemma (in linorder_topology) tendsto_max[tendsto_intros]: assumes X: "(X \ x) net" and Y: "(Y \ y) net" shows "((\x. max (X x) (Y x)) \ max x y) net" proof (rule order_tendstoI) fix a assume "a < max x y" then show "eventually (\x. a < max (X x) (Y x)) net" using order_tendstoD(1)[OF X, of a] order_tendstoD(1)[OF Y, of a] by (auto simp: less_max_iff_disj elim: eventually_mono) next fix a assume "max x y < a" then show "eventually (\x. max (X x) (Y x) < a) net" using order_tendstoD(2)[OF X, of a] order_tendstoD(2)[OF Y, of a] by (auto simp: eventually_conj_iff) qed lemma (in linorder_topology) tendsto_min[tendsto_intros]: assumes X: "(X \ x) net" and Y: "(Y \ y) net" shows "((\x. min (X x) (Y x)) \ min x y) net" proof (rule order_tendstoI) fix a assume "a < min x y" then show "eventually (\x. a < min (X x) (Y x)) net" using order_tendstoD(1)[OF X, of a] order_tendstoD(1)[OF Y, of a] by (auto simp: eventually_conj_iff) next fix a assume "min x y < a" then show "eventually (\x. min (X x) (Y x) < a) net" using order_tendstoD(2)[OF X, of a] order_tendstoD(2)[OF Y, of a] by (auto simp: min_less_iff_disj elim: eventually_mono) qed lemma (in order_topology) assumes "a < b" shows at_within_Icc_at_right: "at a within {a..b} = at_right a" and at_within_Icc_at_left: "at b within {a..b} = at_left b" using order_tendstoD(2)[OF tendsto_ident_at assms, of "{a<..}"] using order_tendstoD(1)[OF tendsto_ident_at assms, of "{.. x < b \ at x within {a..b} = at x" by (rule at_within_open_subset[where S="{a<.. bot" and "(f \ a) F" and "(f \ b) F" shows "a = b" proof (rule ccontr) assume "a \ b" obtain U V where "open U" "open V" "a \ U" "b \ V" "U \ V = {}" using hausdorff [OF \a \ b\] by fast have "eventually (\x. f x \ U) F" using \(f \ a) F\ \open U\ \a \ U\ by (rule topological_tendstoD) moreover have "eventually (\x. f x \ V) F" using \(f \ b) F\ \open V\ \b \ V\ by (rule topological_tendstoD) ultimately have "eventually (\x. False) F" proof eventually_elim case (elim x) then have "f x \ U \ V" by simp with \U \ V = {}\ show ?case by simp qed with \\ trivial_limit F\ show "False" by (simp add: trivial_limit_def) qed lemma (in t2_space) tendsto_const_iff: fixes a b :: 'a assumes "\ trivial_limit F" shows "((\x. a) \ b) F \ a = b" by (auto intro!: tendsto_unique [OF assms tendsto_const]) lemma (in t2_space) tendsto_unique': assumes "F \ bot" shows "\\<^sub>\\<^sub>1l. (f \ l) F" using Uniq_def assms local.tendsto_unique by fastforce lemma Lim_in_closed_set: assumes "closed S" "eventually (\x. f(x) \ S) F" "F \ bot" "(f \ l) F" shows "l \ S" proof (rule ccontr) assume "l \ S" with \closed S\ have "open (- S)" "l \ - S" by (simp_all add: open_Compl) with assms(4) have "eventually (\x. f x \ - S) F" by (rule topological_tendstoD) with assms(2) have "eventually (\x. False) F" by (rule eventually_elim2) simp with assms(3) show "False" by (simp add: eventually_False) qed lemma (in t3_space) nhds_closed: assumes "x \ A" and "open A" shows "\A'. x \ A' \ closed A' \ A' \ A \ eventually (\y. y \ A') (nhds x)" proof - from assms have "\U V. open U \ open V \ x \ U \ - A \ V \ U \ V = {}" by (intro t3_space) auto then obtain U V where UV: "open U" "open V" "x \ U" "-A \ V" "U \ V = {}" by auto have "eventually (\y. y \ U) (nhds x)" using \open U\ and \x \ U\ by (intro eventually_nhds_in_open) hence "eventually (\y. y \ -V) (nhds x)" by eventually_elim (use UV in auto) with UV show ?thesis by (intro exI[of _ "-V"]) auto qed lemma (in order_topology) increasing_tendsto: assumes bdd: "eventually (\n. f n \ l) F" and en: "\x. x < l \ eventually (\n. x < f n) F" shows "(f \ l) F" using assms by (intro order_tendstoI) (auto elim!: eventually_mono) lemma (in order_topology) decreasing_tendsto: assumes bdd: "eventually (\n. l \ f n) F" and en: "\x. l < x \ eventually (\n. f n < x) F" shows "(f \ l) F" using assms by (intro order_tendstoI) (auto elim!: eventually_mono) lemma (in order_topology) tendsto_sandwich: assumes ev: "eventually (\n. f n \ g n) net" "eventually (\n. g n \ h n) net" assumes lim: "(f \ c) net" "(h \ c) net" shows "(g \ c) net" proof (rule order_tendstoI) fix a show "a < c \ eventually (\x. a < g x) net" using order_tendstoD[OF lim(1), of a] ev by (auto elim: eventually_elim2) next fix a show "c < a \ eventually (\x. g x < a) net" using order_tendstoD[OF lim(2), of a] ev by (auto elim: eventually_elim2) qed lemma (in t1_space) limit_frequently_eq: assumes "F \ bot" and "frequently (\x. f x = c) F" and "(f \ d) F" shows "d = c" proof (rule ccontr) assume "d \ c" from t1_space[OF this] obtain U where "open U" "d \ U" "c \ U" by blast with assms have "eventually (\x. f x \ U) F" unfolding tendsto_def by blast then have "eventually (\x. f x \ c) F" by eventually_elim (insert \c \ U\, blast) with assms(2) show False unfolding frequently_def by contradiction qed lemma (in t1_space) tendsto_imp_eventually_ne: assumes "(f \ c) F" "c \ c'" shows "eventually (\z. f z \ c') F" proof (cases "F=bot") case True thus ?thesis by auto next case False show ?thesis proof (rule ccontr) assume "\ eventually (\z. f z \ c') F" then have "frequently (\z. f z = c') F" by (simp add: frequently_def) from limit_frequently_eq[OF False this \(f \ c) F\] and \c \ c'\ show False by contradiction qed qed lemma (in linorder_topology) tendsto_le: assumes F: "\ trivial_limit F" and x: "(f \ x) F" and y: "(g \ y) F" and ev: "eventually (\x. g x \ f x) F" shows "y \ x" proof (rule ccontr) assume "\ y \ x" with less_separate[of x y] obtain a b where xy: "x < a" "b < y" "{.. {b<..} = {}" by (auto simp: not_le) then have "eventually (\x. f x < a) F" "eventually (\x. b < g x) F" using x y by (auto intro: order_tendstoD) with ev have "eventually (\x. False) F" by eventually_elim (insert xy, fastforce) with F show False by (simp add: eventually_False) qed lemma (in linorder_topology) tendsto_lowerbound: assumes x: "(f \ x) F" and ev: "eventually (\i. a \ f i) F" and F: "\ trivial_limit F" shows "a \ x" using F x tendsto_const ev by (rule tendsto_le) lemma (in linorder_topology) tendsto_upperbound: assumes x: "(f \ x) F" and ev: "eventually (\i. a \ f i) F" and F: "\ trivial_limit F" shows "a \ x" by (rule tendsto_le [OF F tendsto_const x ev]) lemma filterlim_at_within_not_equal: fixes f::"'a \ 'b::t2_space" assumes "filterlim f (at a within s) F" shows "eventually (\w. f w\s \ f w \b) F" proof (cases "a=b") case True then show ?thesis using assms by (simp add: filterlim_at) next case False from hausdorff[OF this] obtain U V where UV:"open U" "open V" "a \ U" "b \ V" "U \ V = {}" by auto have "(f \ a) F" using assms filterlim_at by auto then have "\\<^sub>F x in F. f x \ U" using UV unfolding tendsto_def by auto moreover have "\\<^sub>F x in F. f x \ s \ f x\a" using assms filterlim_at by auto ultimately show ?thesis apply eventually_elim using UV by auto qed subsubsection \Rules about \<^const>\Lim\\ lemma tendsto_Lim: "\ trivial_limit net \ (f \ l) net \ Lim net f = l" unfolding Lim_def using tendsto_unique [of net f] by auto lemma Lim_ident_at: "\ trivial_limit (at x within s) \ Lim (at x within s) (\x. x) = x" by (rule tendsto_Lim[OF _ tendsto_ident_at]) auto lemma eventually_Lim_ident_at: "(\\<^sub>F y in at x within X. P (Lim (at x within X) (\x. x)) y) \ (\\<^sub>F y in at x within X. P x y)" for x::"'a::t2_space" by (cases "at x within X = bot") (auto simp: Lim_ident_at) lemma filterlim_at_bot_at_right: fixes f :: "'a::linorder_topology \ 'b::linorder" assumes mono: "\x y. Q x \ Q y \ x \ y \ f x \ f y" and bij: "\x. P x \ f (g x) = x" "\x. P x \ Q (g x)" and Q: "eventually Q (at_right a)" and bound: "\b. Q b \ a < b" and P: "eventually P at_bot" shows "filterlim f at_bot (at_right a)" proof - from P obtain x where x: "\y. y \ x \ P y" unfolding eventually_at_bot_linorder by auto show ?thesis proof (intro filterlim_at_bot_le[THEN iffD2] allI impI) fix z assume "z \ x" with x have "P z" by auto have "eventually (\x. x \ g z) (at_right a)" using bound[OF bij(2)[OF \P z\]] unfolding eventually_at_right[OF bound[OF bij(2)[OF \P z\]]] by (auto intro!: exI[of _ "g z"]) with Q show "eventually (\x. f x \ z) (at_right a)" by eventually_elim (metis bij \P z\ mono) qed qed lemma filterlim_at_top_at_left: fixes f :: "'a::linorder_topology \ 'b::linorder" assumes mono: "\x y. Q x \ Q y \ x \ y \ f x \ f y" and bij: "\x. P x \ f (g x) = x" "\x. P x \ Q (g x)" and Q: "eventually Q (at_left a)" and bound: "\b. Q b \ b < a" and P: "eventually P at_top" shows "filterlim f at_top (at_left a)" proof - from P obtain x where x: "\y. x \ y \ P y" unfolding eventually_at_top_linorder by auto show ?thesis proof (intro filterlim_at_top_ge[THEN iffD2] allI impI) fix z assume "x \ z" with x have "P z" by auto have "eventually (\x. g z \ x) (at_left a)" using bound[OF bij(2)[OF \P z\]] unfolding eventually_at_left[OF bound[OF bij(2)[OF \P z\]]] by (auto intro!: exI[of _ "g z"]) with Q show "eventually (\x. z \ f x) (at_left a)" by eventually_elim (metis bij \P z\ mono) qed qed lemma filterlim_split_at: "filterlim f F (at_left x) \ filterlim f F (at_right x) \ filterlim f F (at x)" for x :: "'a::linorder_topology" by (subst at_eq_sup_left_right) (rule filterlim_sup) lemma filterlim_at_split: "filterlim f F (at x) \ filterlim f F (at_left x) \ filterlim f F (at_right x)" for x :: "'a::linorder_topology" by (subst at_eq_sup_left_right) (simp add: filterlim_def filtermap_sup) lemma eventually_nhds_top: fixes P :: "'a :: {order_top,linorder_topology} \ bool" and b :: 'a assumes "b < top" shows "eventually P (nhds top) \ (\bz. b < z \ P z))" unfolding eventually_nhds proof safe fix S :: "'a set" assume "open S" "top \ S" note open_left[OF this \b < top\] moreover assume "\s\S. P s" ultimately show "\bz>b. P z" by (auto simp: subset_eq Ball_def) next fix b assume "b < top" "\z>b. P z" then show "\S. open S \ top \ S \ (\xa\S. P xa)" by (intro exI[of _ "{b <..}"]) auto qed lemma tendsto_at_within_iff_tendsto_nhds: "(g \ g l) (at l within S) \ (g \ g l) (inf (nhds l) (principal S))" unfolding tendsto_def eventually_at_filter eventually_inf_principal by (intro ext all_cong imp_cong) (auto elim!: eventually_mono) subsection \Limits on sequences\ abbreviation (in topological_space) LIMSEQ :: "[nat \ 'a, 'a] \ bool" ("((_)/ \ (_))" [60, 60] 60) where "X \ L \ (X \ L) sequentially" abbreviation (in t2_space) lim :: "(nat \ 'a) \ 'a" where "lim X \ Lim sequentially X" definition (in topological_space) convergent :: "(nat \ 'a) \ bool" where "convergent X = (\L. X \ L)" lemma lim_def: "lim X = (THE L. X \ L)" unfolding Lim_def .. lemma lim_explicit: "f \ f0 \ (\S. open S \ f0 \ S \ (\N. \n\N. f n \ S))" unfolding tendsto_def eventually_sequentially by auto subsection \Monotone sequences and subsequences\ text \ Definition of monotonicity. The use of disjunction here complicates proofs considerably. One alternative is to add a Boolean argument to indicate the direction. Another is to develop the notions of increasing and decreasing first. \ definition monoseq :: "(nat \ 'a::order) \ bool" where "monoseq X \ (\m. \n\m. X m \ X n) \ (\m. \n\m. X n \ X m)" abbreviation incseq :: "(nat \ 'a::order) \ bool" where "incseq X \ mono X" lemma incseq_def: "incseq X \ (\m. \n\m. X n \ X m)" unfolding mono_def .. abbreviation decseq :: "(nat \ 'a::order) \ bool" where "decseq X \ antimono X" lemma decseq_def: "decseq X \ (\m. \n\m. X n \ X m)" unfolding antimono_def .. subsubsection \Definition of subsequence.\ (* For compatibility with the old "subseq" *) lemma strict_mono_leD: "strict_mono r \ m \ n \ r m \ r n" by (erule (1) monoD [OF strict_mono_mono]) lemma strict_mono_id: "strict_mono id" by (simp add: strict_mono_def) lemma incseq_SucI: "(\n. X n \ X (Suc n)) \ incseq X" using lift_Suc_mono_le[of X] by (auto simp: incseq_def) lemma incseqD: "incseq f \ i \ j \ f i \ f j" by (auto simp: incseq_def) lemma incseq_SucD: "incseq A \ A i \ A (Suc i)" using incseqD[of A i "Suc i"] by auto lemma incseq_Suc_iff: "incseq f \ (\n. f n \ f (Suc n))" by (auto intro: incseq_SucI dest: incseq_SucD) lemma incseq_const[simp, intro]: "incseq (\x. k)" unfolding incseq_def by auto lemma decseq_SucI: "(\n. X (Suc n) \ X n) \ decseq X" using order.lift_Suc_mono_le[OF dual_order, of X] by (auto simp: decseq_def) lemma decseqD: "decseq f \ i \ j \ f j \ f i" by (auto simp: decseq_def) lemma decseq_SucD: "decseq A \ A (Suc i) \ A i" using decseqD[of A i "Suc i"] by auto lemma decseq_Suc_iff: "decseq f \ (\n. f (Suc n) \ f n)" by (auto intro: decseq_SucI dest: decseq_SucD) lemma decseq_const[simp, intro]: "decseq (\x. k)" unfolding decseq_def by auto lemma monoseq_iff: "monoseq X \ incseq X \ decseq X" unfolding monoseq_def incseq_def decseq_def .. lemma monoseq_Suc: "monoseq X \ (\n. X n \ X (Suc n)) \ (\n. X (Suc n) \ X n)" unfolding monoseq_iff incseq_Suc_iff decseq_Suc_iff .. lemma monoI1: "\m. \n \ m. X m \ X n \ monoseq X" by (simp add: monoseq_def) lemma monoI2: "\m. \n \ m. X n \ X m \ monoseq X" by (simp add: monoseq_def) lemma mono_SucI1: "\n. X n \ X (Suc n) \ monoseq X" by (simp add: monoseq_Suc) lemma mono_SucI2: "\n. X (Suc n) \ X n \ monoseq X" by (simp add: monoseq_Suc) lemma monoseq_minus: fixes a :: "nat \ 'a::ordered_ab_group_add" assumes "monoseq a" shows "monoseq (\ n. - a n)" proof (cases "\m. \n \ m. a m \ a n") case True then have "\m. \n \ m. - a n \ - a m" by auto then show ?thesis by (rule monoI2) next case False then have "\m. \n \ m. - a m \ - a n" using \monoseq a\[unfolded monoseq_def] by auto then show ?thesis by (rule monoI1) qed subsubsection \Subsequence (alternative definition, (e.g. Hoskins)\ lemma strict_mono_Suc_iff: "strict_mono f \ (\n. f n < f (Suc n))" proof (intro iffI strict_monoI) assume *: "\n. f n < f (Suc n)" fix m n :: nat assume "m < n" thus "f m < f n" by (induction rule: less_Suc_induct) (use * in auto) qed (auto simp: strict_mono_def) lemma strict_mono_add: "strict_mono (\n::'a::linordered_semidom. n + k)" by (auto simp: strict_mono_def) text \For any sequence, there is a monotonic subsequence.\ lemma seq_monosub: fixes s :: "nat \ 'a::linorder" shows "\f. strict_mono f \ monoseq (\n. (s (f n)))" proof (cases "\n. \p>n. \m\p. s m \ s p") case True then have "\f. \n. (\m\f n. s m \ s (f n)) \ f n < f (Suc n)" by (intro dependent_nat_choice) (auto simp: conj_commute) then obtain f :: "nat \ nat" where f: "strict_mono f" and mono: "\n m. f n \ m \ s m \ s (f n)" by (auto simp: strict_mono_Suc_iff) then have "incseq f" unfolding strict_mono_Suc_iff incseq_Suc_iff by (auto intro: less_imp_le) then have "monoseq (\n. s (f n))" by (auto simp add: incseq_def intro!: mono monoI2) with f show ?thesis by auto next case False then obtain N where N: "p > N \ \m>p. s p < s m" for p by (force simp: not_le le_less) have "\f. \n. N < f n \ f n < f (Suc n) \ s (f n) \ s (f (Suc n))" proof (intro dependent_nat_choice) fix x assume "N < x" with N[of x] show "\y>N. x < y \ s x \ s y" by (auto intro: less_trans) qed auto then show ?thesis by (auto simp: monoseq_iff incseq_Suc_iff strict_mono_Suc_iff) qed lemma seq_suble: assumes sf: "strict_mono (f :: nat \ nat)" shows "n \ f n" proof (induct n) case 0 show ?case by simp next case (Suc n) with sf [unfolded strict_mono_Suc_iff, rule_format, of n] have "n < f (Suc n)" by arith then show ?case by arith qed lemma eventually_subseq: "strict_mono r \ eventually P sequentially \ eventually (\n. P (r n)) sequentially" unfolding eventually_sequentially by (metis seq_suble le_trans) lemma not_eventually_sequentiallyD: assumes "\ eventually P sequentially" shows "\r::nat\nat. strict_mono r \ (\n. \ P (r n))" proof - from assms have "\n. \m\n. \ P m" unfolding eventually_sequentially by (simp add: not_less) then obtain r where "\n. r n \ n" "\n. \ P (r n)" by (auto simp: choice_iff) then show ?thesis by (auto intro!: exI[of _ "\n. r (((Suc \ r) ^^ Suc n) 0)"] simp: less_eq_Suc_le strict_mono_Suc_iff) qed lemma sequentially_offset: assumes "eventually (\i. P i) sequentially" shows "eventually (\i. P (i + k)) sequentially" using assms by (rule eventually_sequentially_seg [THEN iffD2]) lemma seq_offset_neg: "(f \ l) sequentially \ ((\i. f(i - k)) \ l) sequentially" apply (erule filterlim_compose) apply (simp add: filterlim_def le_sequentially eventually_filtermap eventually_sequentially, arith) done lemma filterlim_subseq: "strict_mono f \ filterlim f sequentially sequentially" unfolding filterlim_iff by (metis eventually_subseq) lemma strict_mono_o: "strict_mono r \ strict_mono s \ strict_mono (r \ s)" unfolding strict_mono_def by simp lemma strict_mono_compose: "strict_mono r \ strict_mono s \ strict_mono (\x. r (s x))" using strict_mono_o[of r s] by (simp add: o_def) lemma incseq_imp_monoseq: "incseq X \ monoseq X" by (simp add: incseq_def monoseq_def) lemma decseq_imp_monoseq: "decseq X \ monoseq X" by (simp add: decseq_def monoseq_def) lemma decseq_eq_incseq: "decseq X = incseq (\n. - X n)" for X :: "nat \ 'a::ordered_ab_group_add" by (simp add: decseq_def incseq_def) lemma INT_decseq_offset: assumes "decseq F" shows "(\i. F i) = (\i\{n..}. F i)" proof safe fix x i assume x: "x \ (\i\{n..}. F i)" show "x \ F i" proof cases from x have "x \ F n" by auto also assume "i \ n" with \decseq F\ have "F n \ F i" unfolding decseq_def by simp finally show ?thesis . qed (insert x, simp) qed auto lemma LIMSEQ_const_iff: "(\n. k) \ l \ k = l" for k l :: "'a::t2_space" using trivial_limit_sequentially by (rule tendsto_const_iff) lemma LIMSEQ_SUP: "incseq X \ X \ (SUP i. X i :: 'a::{complete_linorder,linorder_topology})" by (intro increasing_tendsto) (auto simp: SUP_upper less_SUP_iff incseq_def eventually_sequentially intro: less_le_trans) lemma LIMSEQ_INF: "decseq X \ X \ (INF i. X i :: 'a::{complete_linorder,linorder_topology})" by (intro decreasing_tendsto) (auto simp: INF_lower INF_less_iff decseq_def eventually_sequentially intro: le_less_trans) lemma LIMSEQ_ignore_initial_segment: "f \ a \ (\n. f (n + k)) \ a" unfolding tendsto_def by (subst eventually_sequentially_seg[where k=k]) lemma LIMSEQ_offset: "(\n. f (n + k)) \ a \ f \ a" unfolding tendsto_def by (subst (asm) eventually_sequentially_seg[where k=k]) lemma LIMSEQ_Suc: "f \ l \ (\n. f (Suc n)) \ l" by (drule LIMSEQ_ignore_initial_segment [where k="Suc 0"]) simp lemma LIMSEQ_imp_Suc: "(\n. f (Suc n)) \ l \ f \ l" by (rule LIMSEQ_offset [where k="Suc 0"]) simp lemma LIMSEQ_lessThan_iff_atMost: shows "(\n. f {.. x \ (\n. f {..n}) \ x" apply (subst filterlim_sequentially_Suc [symmetric]) apply (simp only: lessThan_Suc_atMost) done lemma (in t2_space) LIMSEQ_Uniq: "\\<^sub>\\<^sub>1l. X \ l" by (simp add: tendsto_unique') lemma (in t2_space) LIMSEQ_unique: "X \ a \ X \ b \ a = b" using trivial_limit_sequentially by (rule tendsto_unique) lemma LIMSEQ_le_const: "X \ x \ \N. \n\N. a \ X n \ a \ x" for a x :: "'a::linorder_topology" by (simp add: eventually_at_top_linorder tendsto_lowerbound) lemma LIMSEQ_le: "X \ x \ Y \ y \ \N. \n\N. X n \ Y n \ x \ y" for x y :: "'a::linorder_topology" using tendsto_le[of sequentially Y y X x] by (simp add: eventually_sequentially) lemma LIMSEQ_le_const2: "X \ x \ \N. \n\N. X n \ a \ x \ a" for a x :: "'a::linorder_topology" by (rule LIMSEQ_le[of X x "\n. a"]) auto lemma Lim_bounded: "f \ l \ \n\M. f n \ C \ l \ C" for l :: "'a::linorder_topology" by (intro LIMSEQ_le_const2) auto lemma Lim_bounded2: fixes f :: "nat \ 'a::linorder_topology" assumes lim:"f \ l" and ge: "\n\N. f n \ C" shows "l \ C" using ge by (intro tendsto_le[OF trivial_limit_sequentially lim tendsto_const]) (auto simp: eventually_sequentially) lemma lim_mono: fixes X Y :: "nat \ 'a::linorder_topology" assumes "\n. N \ n \ X n \ Y n" and "X \ x" and "Y \ y" shows "x \ y" using assms(1) by (intro LIMSEQ_le[OF assms(2,3)]) auto lemma Sup_lim: fixes a :: "'a::{complete_linorder,linorder_topology}" assumes "\n. b n \ s" and "b \ a" shows "a \ Sup s" by (metis Lim_bounded assms complete_lattice_class.Sup_upper) lemma Inf_lim: fixes a :: "'a::{complete_linorder,linorder_topology}" assumes "\n. b n \ s" and "b \ a" shows "Inf s \ a" by (metis Lim_bounded2 assms complete_lattice_class.Inf_lower) lemma SUP_Lim: fixes X :: "nat \ 'a::{complete_linorder,linorder_topology}" assumes inc: "incseq X" and l: "X \ l" shows "(SUP n. X n) = l" using LIMSEQ_SUP[OF inc] tendsto_unique[OF trivial_limit_sequentially l] by simp lemma INF_Lim: fixes X :: "nat \ 'a::{complete_linorder,linorder_topology}" assumes dec: "decseq X" and l: "X \ l" shows "(INF n. X n) = l" using LIMSEQ_INF[OF dec] tendsto_unique[OF trivial_limit_sequentially l] by simp lemma convergentD: "convergent X \ \L. X \ L" by (simp add: convergent_def) lemma convergentI: "X \ L \ convergent X" by (auto simp add: convergent_def) lemma convergent_LIMSEQ_iff: "convergent X \ X \ lim X" by (auto intro: theI LIMSEQ_unique simp add: convergent_def lim_def) lemma convergent_const: "convergent (\n. c)" by (rule convergentI) (rule tendsto_const) lemma monoseq_le: "monoseq a \ a \ x \ (\n. a n \ x) \ (\m. \n\m. a m \ a n) \ (\n. x \ a n) \ (\m. \n\m. a n \ a m)" for x :: "'a::linorder_topology" by (metis LIMSEQ_le_const LIMSEQ_le_const2 decseq_def incseq_def monoseq_iff) lemma LIMSEQ_subseq_LIMSEQ: "X \ L \ strict_mono f \ (X \ f) \ L" unfolding comp_def by (rule filterlim_compose [of X, OF _ filterlim_subseq]) lemma convergent_subseq_convergent: "convergent X \ strict_mono f \ convergent (X \ f)" by (auto simp: convergent_def intro: LIMSEQ_subseq_LIMSEQ) lemma limI: "X \ L \ lim X = L" by (rule tendsto_Lim) (rule trivial_limit_sequentially) lemma lim_le: "convergent f \ (\n. f n \ x) \ lim f \ x" for x :: "'a::linorder_topology" using LIMSEQ_le_const2[of f "lim f" x] by (simp add: convergent_LIMSEQ_iff) lemma lim_const [simp]: "lim (\m. a) = a" by (simp add: limI) subsubsection \Increasing and Decreasing Series\ lemma incseq_le: "incseq X \ X \ L \ X n \ L" for L :: "'a::linorder_topology" by (metis incseq_def LIMSEQ_le_const) lemma decseq_ge: "decseq X \ X \ L \ L \ X n" for L :: "'a::linorder_topology" by (metis decseq_def LIMSEQ_le_const2) subsection \First countable topologies\ class first_countable_topology = topological_space + assumes first_countable_basis: "\A::nat \ 'a set. (\i. x \ A i \ open (A i)) \ (\S. open S \ x \ S \ (\i. A i \ S))" lemma (in first_countable_topology) countable_basis_at_decseq: obtains A :: "nat \ 'a set" where "\i. open (A i)" "\i. x \ (A i)" "\S. open S \ x \ S \ eventually (\i. A i \ S) sequentially" proof atomize_elim from first_countable_basis[of x] obtain A :: "nat \ 'a set" where nhds: "\i. open (A i)" "\i. x \ A i" and incl: "\S. open S \ x \ S \ \i. A i \ S" by auto define F where "F n = (\i\n. A i)" for n show "\A. (\i. open (A i)) \ (\i. x \ A i) \ (\S. open S \ x \ S \ eventually (\i. A i \ S) sequentially)" proof (safe intro!: exI[of _ F]) fix i show "open (F i)" using nhds(1) by (auto simp: F_def) show "x \ F i" using nhds(2) by (auto simp: F_def) next fix S assume "open S" "x \ S" from incl[OF this] obtain i where "F i \ S" unfolding F_def by auto moreover have "\j. i \ j \ F j \ F i" by (simp add: Inf_superset_mono F_def image_mono) ultimately show "eventually (\i. F i \ S) sequentially" by (auto simp: eventually_sequentially) qed qed lemma (in first_countable_topology) nhds_countable: obtains X :: "nat \ 'a set" where "decseq X" "\n. open (X n)" "\n. x \ X n" "nhds x = (INF n. principal (X n))" proof - from first_countable_basis obtain A :: "nat \ 'a set" where *: "\n. x \ A n" "\n. open (A n)" "\S. open S \ x \ S \ \i. A i \ S" by metis show thesis proof show "decseq (\n. \i\n. A i)" by (simp add: antimono_iff_le_Suc atMost_Suc) show "x \ (\i\n. A i)" "\n. open (\i\n. A i)" for n using * by auto with * show "nhds x = (INF n. principal (\i\n. A i))" unfolding nhds_def apply (intro INF_eq) apply fastforce apply blast done qed qed lemma (in first_countable_topology) countable_basis: obtains A :: "nat \ 'a set" where "\i. open (A i)" "\i. x \ A i" "\F. (\n. F n \ A n) \ F \ x" proof atomize_elim obtain A :: "nat \ 'a set" where *: "\i. open (A i)" "\i. x \ A i" "\S. open S \ x \ S \ eventually (\i. A i \ S) sequentially" by (rule countable_basis_at_decseq) blast have "eventually (\n. F n \ S) sequentially" if "\n. F n \ A n" "open S" "x \ S" for F S using *(3)[of S] that by (auto elim: eventually_mono simp: subset_eq) with * show "\A. (\i. open (A i)) \ (\i. x \ A i) \ (\F. (\n. F n \ A n) \ F \ x)" by (intro exI[of _ A]) (auto simp: tendsto_def) qed lemma (in first_countable_topology) sequentially_imp_eventually_nhds_within: assumes "\f. (\n. f n \ s) \ f \ a \ eventually (\n. P (f n)) sequentially" shows "eventually P (inf (nhds a) (principal s))" proof (rule ccontr) obtain A :: "nat \ 'a set" where *: "\i. open (A i)" "\i. a \ A i" "\F. \n. F n \ A n \ F \ a" by (rule countable_basis) blast assume "\ ?thesis" with * have "\F. \n. F n \ s \ F n \ A n \ \ P (F n)" unfolding eventually_inf_principal eventually_nhds by (intro choice) fastforce then obtain F where F: "\n. F n \ s" and "\n. F n \ A n" and F': "\n. \ P (F n)" by blast with * have "F \ a" by auto then have "eventually (\n. P (F n)) sequentially" using assms F by simp then show False by (simp add: F') qed lemma (in first_countable_topology) eventually_nhds_within_iff_sequentially: "eventually P (inf (nhds a) (principal s)) \ (\f. (\n. f n \ s) \ f \ a \ eventually (\n. P (f n)) sequentially)" proof (safe intro!: sequentially_imp_eventually_nhds_within) assume "eventually P (inf (nhds a) (principal s))" then obtain S where "open S" "a \ S" "\x\S. x \ s \ P x" by (auto simp: eventually_inf_principal eventually_nhds) moreover fix f assume "\n. f n \ s" "f \ a" ultimately show "eventually (\n. P (f n)) sequentially" by (auto dest!: topological_tendstoD elim: eventually_mono) qed lemma (in first_countable_topology) eventually_nhds_iff_sequentially: "eventually P (nhds a) \ (\f. f \ a \ eventually (\n. P (f n)) sequentially)" using eventually_nhds_within_iff_sequentially[of P a UNIV] by simp (*Thanks to Sébastien Gouëzel*) lemma Inf_as_limit: fixes A::"'a::{linorder_topology, first_countable_topology, complete_linorder} set" assumes "A \ {}" shows "\u. (\n. u n \ A) \ u \ Inf A" proof (cases "Inf A \ A") case True show ?thesis by (rule exI[of _ "\n. Inf A"], auto simp add: True) next case False obtain y where "y \ A" using assms by auto then have "Inf A < y" using False Inf_lower less_le by auto obtain F :: "nat \ 'a set" where F: "\i. open (F i)" "\i. Inf A \ F i" "\u. (\n. u n \ F n) \ u \ Inf A" by (metis first_countable_topology_class.countable_basis) define u where "u = (\n. SOME z. z \ F n \ z \ A)" have "\z. z \ U \ z \ A" if "Inf A \ U" "open U" for U proof - obtain b where "b > Inf A" "{Inf A .. U" using open_right[OF \open U\ \Inf A \ U\ \Inf A < y\] by auto obtain z where "z < b" "z \ A" using \Inf A < b\ Inf_less_iff by auto then have "z \ {Inf A ..z \ A\ \{Inf A .. U\ by auto qed then have *: "u n \ F n \ u n \ A" for n using \Inf A \ F n\ \open (F n)\ unfolding u_def by (metis (no_types, lifting) someI_ex) then have "u \ Inf A" using F(3) by simp then show ?thesis using * by auto qed lemma tendsto_at_iff_sequentially: "(f \ a) (at x within s) \ (\X. (\i. X i \ s - {x}) \ X \ x \ ((f \ X) \ a))" for f :: "'a::first_countable_topology \ _" unfolding filterlim_def[of _ "nhds a"] le_filter_def eventually_filtermap at_within_def eventually_nhds_within_iff_sequentially comp_def by metis lemma approx_from_above_dense_linorder: fixes x::"'a::{dense_linorder, linorder_topology, first_countable_topology}" assumes "x < y" shows "\u. (\n. u n > x) \ (u \ x)" proof - obtain A :: "nat \ 'a set" where A: "\i. open (A i)" "\i. x \ A i" "\F. (\n. F n \ A n) \ F \ x" by (metis first_countable_topology_class.countable_basis) define u where "u = (\n. SOME z. z \ A n \ z > x)" have "\z. z \ U \ x < z" if "x \ U" "open U" for U using open_right[OF \open U\ \x \ U\ \x < y\] by (meson atLeastLessThan_iff dense less_imp_le subset_eq) then have *: "u n \ A n \ x < u n" for n using \x \ A n\ \open (A n)\ unfolding u_def by (metis (no_types, lifting) someI_ex) then have "u \ x" using A(3) by simp then show ?thesis using * by auto qed lemma approx_from_below_dense_linorder: fixes x::"'a::{dense_linorder, linorder_topology, first_countable_topology}" assumes "x > y" shows "\u. (\n. u n < x) \ (u \ x)" proof - obtain A :: "nat \ 'a set" where A: "\i. open (A i)" "\i. x \ A i" "\F. (\n. F n \ A n) \ F \ x" by (metis first_countable_topology_class.countable_basis) define u where "u = (\n. SOME z. z \ A n \ z < x)" have "\z. z \ U \ z < x" if "x \ U" "open U" for U using open_left[OF \open U\ \x \ U\ \x > y\] by (meson dense greaterThanAtMost_iff less_imp_le subset_eq) then have *: "u n \ A n \ u n < x" for n using \x \ A n\ \open (A n)\ unfolding u_def by (metis (no_types, lifting) someI_ex) then have "u \ x" using A(3) by simp then show ?thesis using * by auto qed subsection \Function limit at a point\ abbreviation LIM :: "('a::topological_space \ 'b::topological_space) \ 'a \ 'b \ bool" ("((_)/ \(_)/\ (_))" [60, 0, 60] 60) where "f \a\ L \ (f \ L) (at a)" lemma tendsto_within_open: "a \ S \ open S \ (f \ l) (at a within S) \ (f \a\ l)" by (simp add: tendsto_def at_within_open[where S = S]) lemma tendsto_within_open_NO_MATCH: "a \ S \ NO_MATCH UNIV S \ open S \ (f \ l)(at a within S) \ (f \ l)(at a)" for f :: "'a::topological_space \ 'b::topological_space" using tendsto_within_open by blast lemma LIM_const_not_eq[tendsto_intros]: "k \ L \ \ (\x. k) \a\ L" for a :: "'a::perfect_space" and k L :: "'b::t2_space" by (simp add: tendsto_const_iff) lemmas LIM_not_zero = LIM_const_not_eq [where L = 0] lemma LIM_const_eq: "(\x. k) \a\ L \ k = L" for a :: "'a::perfect_space" and k L :: "'b::t2_space" by (simp add: tendsto_const_iff) lemma LIM_unique: "f \a\ L \ f \a\ M \ L = M" for a :: "'a::perfect_space" and L M :: "'b::t2_space" using at_neq_bot by (rule tendsto_unique) lemma LIM_Uniq: "\\<^sub>\\<^sub>1L::'b::t2_space. f \a\ L" for a :: "'a::perfect_space" by (auto simp add: Uniq_def LIM_unique) text \Limits are equal for functions equal except at limit point.\ lemma LIM_equal: "\x. x \ a \ f x = g x \ (f \a\ l) \ (g \a\ l)" by (simp add: tendsto_def eventually_at_topological) lemma LIM_cong: "a = b \ (\x. x \ b \ f x = g x) \ l = m \ (f \a\ l) \ (g \b\ m)" by (simp add: LIM_equal) lemma tendsto_cong_limit: "(f \ l) F \ k = l \ (f \ k) F" by simp lemma tendsto_at_iff_tendsto_nhds: "g \l\ g l \ (g \ g l) (nhds l)" unfolding tendsto_def eventually_at_filter by (intro ext all_cong imp_cong) (auto elim!: eventually_mono) lemma tendsto_compose: "g \l\ g l \ (f \ l) F \ ((\x. g (f x)) \ g l) F" unfolding tendsto_at_iff_tendsto_nhds by (rule filterlim_compose[of g]) lemma tendsto_compose_eventually: "g \l\ m \ (f \ l) F \ eventually (\x. f x \ l) F \ ((\x. g (f x)) \ m) F" by (rule filterlim_compose[of g _ "at l"]) (auto simp add: filterlim_at) lemma LIM_compose_eventually: assumes "f \a\ b" and "g \b\ c" and "eventually (\x. f x \ b) (at a)" shows "(\x. g (f x)) \a\ c" using assms(2,1,3) by (rule tendsto_compose_eventually) lemma tendsto_compose_filtermap: "((g \ f) \ T) F \ (g \ T) (filtermap f F)" by (simp add: filterlim_def filtermap_filtermap comp_def) lemma tendsto_compose_at: assumes f: "(f \ y) F" and g: "(g \ z) (at y)" and fg: "eventually (\w. f w = y \ g y = z) F" shows "((g \ f) \ z) F" proof - have "(\\<^sub>F a in F. f a \ y) \ g y = z" using fg by force moreover have "(g \ z) (filtermap f F) \ \ (\\<^sub>F a in F. f a \ y)" by (metis (no_types) filterlim_atI filterlim_def tendsto_mono f g) ultimately show ?thesis by (metis (no_types) f filterlim_compose filterlim_filtermap g tendsto_at_iff_tendsto_nhds tendsto_compose_filtermap) qed subsubsection \Relation of \LIM\ and \LIMSEQ\\ lemma (in first_countable_topology) sequentially_imp_eventually_within: "(\f. (\n. f n \ s \ f n \ a) \ f \ a \ eventually (\n. P (f n)) sequentially) \ eventually P (at a within s)" unfolding at_within_def by (intro sequentially_imp_eventually_nhds_within) auto lemma (in first_countable_topology) sequentially_imp_eventually_at: "(\f. (\n. f n \ a) \ f \ a \ eventually (\n. P (f n)) sequentially) \ eventually P (at a)" using sequentially_imp_eventually_within [where s=UNIV] by simp lemma LIMSEQ_SEQ_conv: "(\S. (\n. S n \ a) \ S \ a \ (\n. X (S n)) \ L) \ X \a\ L" (is "?lhs=?rhs") for a :: "'a::first_countable_topology" and L :: "'b::topological_space" proof assume ?lhs then show ?rhs by (simp add: sequentially_imp_eventually_within tendsto_def) next assume ?rhs then show ?lhs using tendsto_compose_eventually eventuallyI by blast qed lemma sequentially_imp_eventually_at_left: fixes a :: "'a::{linorder_topology,first_countable_topology}" assumes b[simp]: "b < a" and *: "\f. (\n. b < f n) \ (\n. f n < a) \ incseq f \ f \ a \ eventually (\n. P (f n)) sequentially" shows "eventually P (at_left a)" proof (safe intro!: sequentially_imp_eventually_within) fix X assume X: "\n. X n \ {..< a} \ X n \ a" "X \ a" show "eventually (\n. P (X n)) sequentially" proof (rule ccontr) assume neg: "\ ?thesis" have "\s. \n. (\ P (X (s n)) \ b < X (s n)) \ (X (s n) \ X (s (Suc n)) \ Suc (s n) \ s (Suc n))" (is "\s. ?P s") proof (rule dependent_nat_choice) have "\ eventually (\n. b < X n \ P (X n)) sequentially" by (intro not_eventually_impI neg order_tendstoD(1) [OF X(2) b]) then show "\x. \ P (X x) \ b < X x" by (auto dest!: not_eventuallyD) next fix x n have "\ eventually (\n. Suc x \ n \ b < X n \ X x < X n \ P (X n)) sequentially" using X by (intro not_eventually_impI order_tendstoD(1)[OF X(2)] eventually_ge_at_top neg) auto then show "\n. (\ P (X n) \ b < X n) \ (X x \ X n \ Suc x \ n)" by (auto dest!: not_eventuallyD) qed then obtain s where "?P s" .. with X have "b < X (s n)" and "X (s n) < a" and "incseq (\n. X (s n))" and "(\n. X (s n)) \ a" and "\ P (X (s n))" for n by (auto simp: strict_mono_Suc_iff Suc_le_eq incseq_Suc_iff intro!: LIMSEQ_subseq_LIMSEQ[OF \X \ a\, unfolded comp_def]) from *[OF this(1,2,3,4)] this(5) show False by auto qed qed lemma tendsto_at_left_sequentially: fixes a b :: "'b::{linorder_topology,first_countable_topology}" assumes "b < a" assumes *: "\S. (\n. S n < a) \ (\n. b < S n) \ incseq S \ S \ a \ (\n. X (S n)) \ L" shows "(X \ L) (at_left a)" using assms by (simp add: tendsto_def [where l=L] sequentially_imp_eventually_at_left) lemma sequentially_imp_eventually_at_right: fixes a b :: "'a::{linorder_topology,first_countable_topology}" assumes b[simp]: "a < b" assumes *: "\f. (\n. a < f n) \ (\n. f n < b) \ decseq f \ f \ a \ eventually (\n. P (f n)) sequentially" shows "eventually P (at_right a)" proof (safe intro!: sequentially_imp_eventually_within) fix X assume X: "\n. X n \ {a <..} \ X n \ a" "X \ a" show "eventually (\n. P (X n)) sequentially" proof (rule ccontr) assume neg: "\ ?thesis" have "\s. \n. (\ P (X (s n)) \ X (s n) < b) \ (X (s (Suc n)) \ X (s n) \ Suc (s n) \ s (Suc n))" (is "\s. ?P s") proof (rule dependent_nat_choice) have "\ eventually (\n. X n < b \ P (X n)) sequentially" by (intro not_eventually_impI neg order_tendstoD(2) [OF X(2) b]) then show "\x. \ P (X x) \ X x < b" by (auto dest!: not_eventuallyD) next fix x n have "\ eventually (\n. Suc x \ n \ X n < b \ X n < X x \ P (X n)) sequentially" using X by (intro not_eventually_impI order_tendstoD(2)[OF X(2)] eventually_ge_at_top neg) auto then show "\n. (\ P (X n) \ X n < b) \ (X n \ X x \ Suc x \ n)" by (auto dest!: not_eventuallyD) qed then obtain s where "?P s" .. with X have "a < X (s n)" and "X (s n) < b" and "decseq (\n. X (s n))" and "(\n. X (s n)) \ a" and "\ P (X (s n))" for n by (auto simp: strict_mono_Suc_iff Suc_le_eq decseq_Suc_iff intro!: LIMSEQ_subseq_LIMSEQ[OF \X \ a\, unfolded comp_def]) from *[OF this(1,2,3,4)] this(5) show False by auto qed qed lemma tendsto_at_right_sequentially: fixes a :: "_ :: {linorder_topology, first_countable_topology}" assumes "a < b" and *: "\S. (\n. a < S n) \ (\n. S n < b) \ decseq S \ S \ a \ (\n. X (S n)) \ L" shows "(X \ L) (at_right a)" using assms by (simp add: tendsto_def [where l=L] sequentially_imp_eventually_at_right) subsection \Continuity\ subsubsection \Continuity on a set\ definition continuous_on :: "'a set \ ('a::topological_space \ 'b::topological_space) \ bool" where "continuous_on s f \ (\x\s. (f \ f x) (at x within s))" lemma continuous_on_cong [cong]: "s = t \ (\x. x \ t \ f x = g x) \ continuous_on s f \ continuous_on t g" unfolding continuous_on_def by (intro ball_cong filterlim_cong) (auto simp: eventually_at_filter) lemma continuous_on_cong_simp: "s = t \ (\x. x \ t =simp=> f x = g x) \ continuous_on s f \ continuous_on t g" unfolding simp_implies_def by (rule continuous_on_cong) lemma continuous_on_topological: "continuous_on s f \ (\x\s. \B. open B \ f x \ B \ (\A. open A \ x \ A \ (\y\s. y \ A \ f y \ B)))" unfolding continuous_on_def tendsto_def eventually_at_topological by metis lemma continuous_on_open_invariant: "continuous_on s f \ (\B. open B \ (\A. open A \ A \ s = f -` B \ s))" proof safe fix B :: "'b set" assume "continuous_on s f" "open B" then have "\x\f -` B \ s. (\A. open A \ x \ A \ s \ A \ f -` B)" by (auto simp: continuous_on_topological subset_eq Ball_def imp_conjL) then obtain A where "\x\f -` B \ s. open (A x) \ x \ A x \ s \ A x \ f -` B" unfolding bchoice_iff .. then show "\A. open A \ A \ s = f -` B \ s" by (intro exI[of _ "\x\f -` B \ s. A x"]) auto next assume B: "\B. open B \ (\A. open A \ A \ s = f -` B \ s)" show "continuous_on s f" unfolding continuous_on_topological proof safe fix x B assume "x \ s" "open B" "f x \ B" with B obtain A where A: "open A" "A \ s = f -` B \ s" by auto with \x \ s\ \f x \ B\ show "\A. open A \ x \ A \ (\y\s. y \ A \ f y \ B)" by (intro exI[of _ A]) auto qed qed lemma continuous_on_open_vimage: "open s \ continuous_on s f \ (\B. open B \ open (f -` B \ s))" unfolding continuous_on_open_invariant by (metis open_Int Int_absorb Int_commute[of s] Int_assoc[of _ _ s]) corollary continuous_imp_open_vimage: assumes "continuous_on s f" "open s" "open B" "f -` B \ s" shows "open (f -` B)" by (metis assms continuous_on_open_vimage le_iff_inf) corollary open_vimage[continuous_intros]: assumes "open s" and "continuous_on UNIV f" shows "open (f -` s)" using assms by (simp add: continuous_on_open_vimage [OF open_UNIV]) lemma continuous_on_closed_invariant: "continuous_on s f \ (\B. closed B \ (\A. closed A \ A \ s = f -` B \ s))" proof - have *: "(\A. P A \ Q (- A)) \ (\A. P A) \ (\A. Q A)" for P Q :: "'b set \ bool" by (metis double_compl) show ?thesis unfolding continuous_on_open_invariant by (intro *) (auto simp: open_closed[symmetric]) qed lemma continuous_on_closed_vimage: "closed s \ continuous_on s f \ (\B. closed B \ closed (f -` B \ s))" unfolding continuous_on_closed_invariant by (metis closed_Int Int_absorb Int_commute[of s] Int_assoc[of _ _ s]) corollary closed_vimage_Int[continuous_intros]: assumes "closed s" and "continuous_on t f" and t: "closed t" shows "closed (f -` s \ t)" using assms by (simp add: continuous_on_closed_vimage [OF t]) corollary closed_vimage[continuous_intros]: assumes "closed s" and "continuous_on UNIV f" shows "closed (f -` s)" using closed_vimage_Int [OF assms] by simp lemma continuous_on_empty [simp]: "continuous_on {} f" by (simp add: continuous_on_def) lemma continuous_on_sing [simp]: "continuous_on {x} f" by (simp add: continuous_on_def at_within_def) lemma continuous_on_open_Union: "(\s. s \ S \ open s) \ (\s. s \ S \ continuous_on s f) \ continuous_on (\S) f" unfolding continuous_on_def by safe (metis open_Union at_within_open UnionI) lemma continuous_on_open_UN: "(\s. s \ S \ open (A s)) \ (\s. s \ S \ continuous_on (A s) f) \ continuous_on (\s\S. A s) f" by (rule continuous_on_open_Union) auto lemma continuous_on_open_Un: "open s \ open t \ continuous_on s f \ continuous_on t f \ continuous_on (s \ t) f" using continuous_on_open_Union [of "{s,t}"] by auto lemma continuous_on_closed_Un: "closed s \ closed t \ continuous_on s f \ continuous_on t f \ continuous_on (s \ t) f" by (auto simp add: continuous_on_closed_vimage closed_Un Int_Un_distrib) lemma continuous_on_closed_Union: assumes "finite I" "\i. i \ I \ closed (U i)" "\i. i \ I \ continuous_on (U i) f" shows "continuous_on (\ i \ I. U i) f" using assms by (induction I) (auto intro!: continuous_on_closed_Un) lemma continuous_on_If: assumes closed: "closed s" "closed t" and cont: "continuous_on s f" "continuous_on t g" and P: "\x. x \ s \ \ P x \ f x = g x" "\x. x \ t \ P x \ f x = g x" shows "continuous_on (s \ t) (\x. if P x then f x else g x)" (is "continuous_on _ ?h") proof- from P have "\x\s. f x = ?h x" "\x\t. g x = ?h x" by auto with cont have "continuous_on s ?h" "continuous_on t ?h" by simp_all with closed show ?thesis by (rule continuous_on_closed_Un) qed lemma continuous_on_cases: "closed s \ closed t \ continuous_on s f \ continuous_on t g \ \x. (x\s \ \ P x) \ (x \ t \ P x) \ f x = g x \ continuous_on (s \ t) (\x. if P x then f x else g x)" by (rule continuous_on_If) auto lemma continuous_on_id[continuous_intros,simp]: "continuous_on s (\x. x)" unfolding continuous_on_def by fast lemma continuous_on_id'[continuous_intros,simp]: "continuous_on s id" unfolding continuous_on_def id_def by fast lemma continuous_on_const[continuous_intros,simp]: "continuous_on s (\x. c)" unfolding continuous_on_def by auto lemma continuous_on_subset: "continuous_on s f \ t \ s \ continuous_on t f" unfolding continuous_on_def by (metis subset_eq tendsto_within_subset) lemma continuous_on_compose[continuous_intros]: "continuous_on s f \ continuous_on (f ` s) g \ continuous_on s (g \ f)" unfolding continuous_on_topological by simp metis lemma continuous_on_compose2: "continuous_on t g \ continuous_on s f \ f ` s \ t \ continuous_on s (\x. g (f x))" using continuous_on_compose[of s f g] continuous_on_subset by (force simp add: comp_def) lemma continuous_on_generate_topology: assumes *: "open = generate_topology X" and **: "\B. B \ X \ \C. open C \ C \ A = f -` B \ A" shows "continuous_on A f" unfolding continuous_on_open_invariant proof safe fix B :: "'a set" assume "open B" then show "\C. open C \ C \ A = f -` B \ A" unfolding * proof induct case (UN K) then obtain C where "\k. k \ K \ open (C k)" "\k. k \ K \ C k \ A = f -` k \ A" by metis then show ?case by (intro exI[of _ "\k\K. C k"]) blast qed (auto intro: **) qed lemma continuous_onI_mono: fixes f :: "'a::linorder_topology \ 'b::{dense_order,linorder_topology}" assumes "open (f`A)" and mono: "\x y. x \ A \ y \ A \ x \ y \ f x \ f y" shows "continuous_on A f" proof (rule continuous_on_generate_topology[OF open_generated_order], safe) have monoD: "\x y. x \ A \ y \ A \ f x < f y \ x < y" by (auto simp: not_le[symmetric] mono) have "\x. x \ A \ f x < b \ a < x" if a: "a \ A" and fa: "f a < b" for a b proof - obtain y where "f a < y" "{f a ..< y} \ f`A" using open_right[OF \open (f`A)\, of "f a" b] a fa by auto obtain z where z: "f a < z" "z < min b y" using dense[of "f a" "min b y"] \f a < y\ \f a < b\ by auto then obtain c where "z = f c" "c \ A" using \{f a ..< y} \ f`A\[THEN subsetD, of z] by (auto simp: less_imp_le) with a z show ?thesis by (auto intro!: exI[of _ c] simp: monoD) qed then show "\C. open C \ C \ A = f -` {.. A" for b by (intro exI[of _ "(\x\{x\A. f x < b}. {..< x})"]) (auto intro: le_less_trans[OF mono] less_imp_le) have "\x. x \ A \ b < f x \ x < a" if a: "a \ A" and fa: "b < f a" for a b proof - note a fa moreover obtain y where "y < f a" "{y <.. f a} \ f`A" using open_left[OF \open (f`A)\, of "f a" b] a fa by auto then obtain z where z: "max b y < z" "z < f a" using dense[of "max b y" "f a"] \y < f a\ \b < f a\ by auto then obtain c where "z = f c" "c \ A" using \{y <.. f a} \ f`A\[THEN subsetD, of z] by (auto simp: less_imp_le) with a z show ?thesis by (auto intro!: exI[of _ c] simp: monoD) qed then show "\C. open C \ C \ A = f -` {b <..} \ A" for b by (intro exI[of _ "(\x\{x\A. b < f x}. {x <..})"]) (auto intro: less_le_trans[OF _ mono] less_imp_le) qed lemma continuous_on_IccI: "\(f \ f a) (at_right a); (f \ f b) (at_left b); (\x. a < x \ x < b \ f \x\ f x); a < b\ \ continuous_on {a .. b} f" for a::"'a::linorder_topology" using at_within_open[of _ "{a<.. f a) (at_right a)" and continuous_on_Icc_at_leftD: "(f \ f b) (at_left b)" using assms by (auto simp: at_within_Icc_at_right at_within_Icc_at_left continuous_on_def dest: bspec[where x=a] bspec[where x=b]) lemma continuous_on_discrete [simp]: "continuous_on A (f :: 'a :: discrete_topology \ _)" by (auto simp: continuous_on_def at_discrete) subsubsection \Continuity at a point\ definition continuous :: "'a::t2_space filter \ ('a \ 'b::topological_space) \ bool" where "continuous F f \ (f \ f (Lim F (\x. x))) F" lemma continuous_bot[continuous_intros, simp]: "continuous bot f" unfolding continuous_def by auto lemma continuous_trivial_limit: "trivial_limit net \ continuous net f" by simp lemma continuous_within: "continuous (at x within s) f \ (f \ f x) (at x within s)" by (cases "trivial_limit (at x within s)") (auto simp add: Lim_ident_at continuous_def) lemma continuous_within_topological: "continuous (at x within s) f \ (\B. open B \ f x \ B \ (\A. open A \ x \ A \ (\y\s. y \ A \ f y \ B)))" unfolding continuous_within tendsto_def eventually_at_topological by metis lemma continuous_within_compose[continuous_intros]: "continuous (at x within s) f \ continuous (at (f x) within f ` s) g \ continuous (at x within s) (g \ f)" by (simp add: continuous_within_topological) metis lemma continuous_within_compose2: "continuous (at x within s) f \ continuous (at (f x) within f ` s) g \ continuous (at x within s) (\x. g (f x))" using continuous_within_compose[of x s f g] by (simp add: comp_def) lemma continuous_at: "continuous (at x) f \ f \x\ f x" using continuous_within[of x UNIV f] by simp lemma continuous_ident[continuous_intros, simp]: "continuous (at x within S) (\x. x)" unfolding continuous_within by (rule tendsto_ident_at) lemma continuous_id[continuous_intros, simp]: "continuous (at x within S) id" by (simp add: id_def) lemma continuous_const[continuous_intros, simp]: "continuous F (\x. c)" unfolding continuous_def by (rule tendsto_const) lemma continuous_on_eq_continuous_within: "continuous_on s f \ (\x\s. continuous (at x within s) f)" unfolding continuous_on_def continuous_within .. lemma continuous_discrete [simp]: "continuous (at x within A) (f :: 'a :: discrete_topology \ _)" by (auto simp: continuous_def at_discrete) abbreviation isCont :: "('a::t2_space \ 'b::topological_space) \ 'a \ bool" where "isCont f a \ continuous (at a) f" lemma isCont_def: "isCont f a \ f \a\ f a" by (rule continuous_at) lemma isContD: "isCont f x \ f \x\ f x" by (simp add: isCont_def) lemma isCont_cong: assumes "eventually (\x. f x = g x) (nhds x)" shows "isCont f x \ isCont g x" proof - from assms have [simp]: "f x = g x" by (rule eventually_nhds_x_imp_x) from assms have "eventually (\x. f x = g x) (at x)" by (auto simp: eventually_at_filter elim!: eventually_mono) with assms have "isCont f x \ isCont g x" unfolding isCont_def by (intro filterlim_cong) (auto elim!: eventually_mono) with assms show ?thesis by simp qed lemma continuous_at_imp_continuous_at_within: "isCont f x \ continuous (at x within s) f" by (auto intro: tendsto_mono at_le simp: continuous_at continuous_within) lemma continuous_on_eq_continuous_at: "open s \ continuous_on s f \ (\x\s. isCont f x)" by (simp add: continuous_on_def continuous_at at_within_open[of _ s]) lemma continuous_within_open: "a \ A \ open A \ continuous (at a within A) f \ isCont f a" by (simp add: at_within_open_NO_MATCH) lemma continuous_at_imp_continuous_on: "\x\s. isCont f x \ continuous_on s f" by (auto intro: continuous_at_imp_continuous_at_within simp: continuous_on_eq_continuous_within) lemma isCont_o2: "isCont f a \ isCont g (f a) \ isCont (\x. g (f x)) a" unfolding isCont_def by (rule tendsto_compose) lemma continuous_at_compose[continuous_intros]: "isCont f a \ isCont g (f a) \ isCont (g \ f) a" unfolding o_def by (rule isCont_o2) lemma isCont_tendsto_compose: "isCont g l \ (f \ l) F \ ((\x. g (f x)) \ g l) F" unfolding isCont_def by (rule tendsto_compose) lemma continuous_on_tendsto_compose: assumes f_cont: "continuous_on s f" and g: "(g \ l) F" and l: "l \ s" and ev: "\\<^sub>Fx in F. g x \ s" shows "((\x. f (g x)) \ f l) F" proof - from f_cont l have f: "(f \ f l) (at l within s)" by (simp add: continuous_on_def) have i: "((\x. if g x = l then f l else f (g x)) \ f l) F" by (rule filterlim_If) (auto intro!: filterlim_compose[OF f] eventually_conj tendsto_mono[OF _ g] simp: filterlim_at eventually_inf_principal eventually_mono[OF ev]) show ?thesis by (rule filterlim_cong[THEN iffD1[OF _ i]]) auto qed lemma continuous_within_compose3: "isCont g (f x) \ continuous (at x within s) f \ continuous (at x within s) (\x. g (f x))" using continuous_at_imp_continuous_at_within continuous_within_compose2 by blast lemma filtermap_nhds_open_map: assumes cont: "isCont f a" and open_map: "\S. open S \ open (f`S)" shows "filtermap f (nhds a) = nhds (f a)" unfolding filter_eq_iff proof safe fix P assume "eventually P (filtermap f (nhds a))" then obtain S where "open S" "a \ S" "\x\S. P (f x)" by (auto simp: eventually_filtermap eventually_nhds) then show "eventually P (nhds (f a))" unfolding eventually_nhds by (intro exI[of _ "f`S"]) (auto intro!: open_map) qed (metis filterlim_iff tendsto_at_iff_tendsto_nhds isCont_def eventually_filtermap cont) lemma continuous_at_split: "continuous (at x) f \ continuous (at_left x) f \ continuous (at_right x) f" for x :: "'a::linorder_topology" by (simp add: continuous_within filterlim_at_split) lemma continuous_on_max [continuous_intros]: fixes f g :: "'a::topological_space \ 'b::linorder_topology" shows "continuous_on A f \ continuous_on A g \ continuous_on A (\x. max (f x) (g x))" by (auto simp: continuous_on_def intro!: tendsto_max) lemma continuous_on_min [continuous_intros]: fixes f g :: "'a::topological_space \ 'b::linorder_topology" shows "continuous_on A f \ continuous_on A g \ continuous_on A (\x. min (f x) (g x))" by (auto simp: continuous_on_def intro!: tendsto_min) lemma continuous_max [continuous_intros]: fixes f :: "'a::t2_space \ 'b::linorder_topology" shows "\continuous F f; continuous F g\ \ continuous F (\x. (max (f x) (g x)))" by (simp add: tendsto_max continuous_def) lemma continuous_min [continuous_intros]: fixes f :: "'a::t2_space \ 'b::linorder_topology" shows "\continuous F f; continuous F g\ \ continuous F (\x. (min (f x) (g x)))" by (simp add: tendsto_min continuous_def) text \ The following open/closed Collect lemmas are ported from Sébastien Gouëzel's \Ergodic_Theory\. \ lemma open_Collect_neq: fixes f g :: "'a::topological_space \ 'b::t2_space" assumes f: "continuous_on UNIV f" and g: "continuous_on UNIV g" shows "open {x. f x \ g x}" proof (rule openI) fix t assume "t \ {x. f x \ g x}" then obtain U V where *: "open U" "open V" "f t \ U" "g t \ V" "U \ V = {}" by (auto simp add: separation_t2) with open_vimage[OF \open U\ f] open_vimage[OF \open V\ g] show "\T. open T \ t \ T \ T \ {x. f x \ g x}" by (intro exI[of _ "f -` U \ g -` V"]) auto qed lemma closed_Collect_eq: fixes f g :: "'a::topological_space \ 'b::t2_space" assumes f: "continuous_on UNIV f" and g: "continuous_on UNIV g" shows "closed {x. f x = g x}" using open_Collect_neq[OF f g] by (simp add: closed_def Collect_neg_eq) lemma open_Collect_less: fixes f g :: "'a::topological_space \ 'b::linorder_topology" assumes f: "continuous_on UNIV f" and g: "continuous_on UNIV g" shows "open {x. f x < g x}" proof (rule openI) fix t assume t: "t \ {x. f x < g x}" show "\T. open T \ t \ T \ T \ {x. f x < g x}" proof (cases "\z. f t < z \ z < g t") case True then obtain z where "f t < z \ z < g t" by blast then show ?thesis using open_vimage[OF _ f, of "{..< z}"] open_vimage[OF _ g, of "{z <..}"] by (intro exI[of _ "f -` {.. g -` {z<..}"]) auto next case False then have *: "{g t ..} = {f t <..}" "{..< g t} = {.. f t}" using t by (auto intro: leI) show ?thesis using open_vimage[OF _ f, of "{..< g t}"] open_vimage[OF _ g, of "{f t <..}"] t apply (intro exI[of _ "f -` {..< g t} \ g -` {f t<..}"]) apply (simp add: open_Int) apply (auto simp add: *) done qed qed lemma closed_Collect_le: fixes f g :: "'a :: topological_space \ 'b::linorder_topology" assumes f: "continuous_on UNIV f" and g: "continuous_on UNIV g" shows "closed {x. f x \ g x}" using open_Collect_less [OF g f] by (simp add: closed_def Collect_neg_eq[symmetric] not_le) subsubsection \Open-cover compactness\ context topological_space begin definition compact :: "'a set \ bool" where compact_eq_Heine_Borel: (* This name is used for backwards compatibility *) "compact S \ (\C. (\c\C. open c) \ S \ \C \ (\D\C. finite D \ S \ \D))" lemma compactI: assumes "\C. \t\C. open t \ s \ \C \ \C'. C' \ C \ finite C' \ s \ \C'" shows "compact s" unfolding compact_eq_Heine_Borel using assms by metis lemma compact_empty[simp]: "compact {}" by (auto intro!: compactI) lemma compactE: (*related to COMPACT_IMP_HEINE_BOREL in HOL Light*) assumes "compact S" "S \ \\" "\B. B \ \ \ open B" obtains \' where "\' \ \" "finite \'" "S \ \\'" by (meson assms compact_eq_Heine_Borel) lemma compactE_image: assumes "compact S" and opn: "\T. T \ C \ open (f T)" and S: "S \ (\c\C. f c)" obtains C' where "C' \ C" and "finite C'" and "S \ (\c\C'. f c)" apply (rule compactE[OF \compact S\ S]) using opn apply force by (metis finite_subset_image) lemma compact_Int_closed [intro]: assumes "compact S" and "closed T" shows "compact (S \ T)" proof (rule compactI) fix C assume C: "\c\C. open c" assume cover: "S \ T \ \C" from C \closed T\ have "\c\C \ {- T}. open c" by auto moreover from cover have "S \ \(C \ {- T})" by auto ultimately have "\D\C \ {- T}. finite D \ S \ \D" using \compact S\ unfolding compact_eq_Heine_Borel by auto then obtain D where "D \ C \ {- T} \ finite D \ S \ \D" .. then show "\D\C. finite D \ S \ T \ \D" by (intro exI[of _ "D - {-T}"]) auto qed lemma compact_diff: "\compact S; open T\ \ compact(S - T)" by (simp add: Diff_eq compact_Int_closed open_closed) lemma inj_setminus: "inj_on uminus (A::'a set set)" by (auto simp: inj_on_def) subsection \Finite intersection property\ lemma compact_fip: "compact U \ (\A. (\a\A. closed a) \ (\B \ A. finite B \ U \ \B \ {}) \ U \ \A \ {})" (is "_ \ ?R") proof (safe intro!: compact_eq_Heine_Borel[THEN iffD2]) fix A assume "compact U" assume A: "\a\A. closed a" "U \ \A = {}" assume fin: "\B \ A. finite B \ U \ \B \ {}" from A have "(\a\uminus`A. open a) \ U \ \(uminus`A)" by auto with \compact U\ obtain B where "B \ A" "finite (uminus`B)" "U \ \(uminus`B)" unfolding compact_eq_Heine_Borel by (metis subset_image_iff) with fin[THEN spec, of B] show False by (auto dest: finite_imageD intro: inj_setminus) next fix A assume ?R assume "\a\A. open a" "U \ \A" then have "U \ \(uminus`A) = {}" "\a\uminus`A. closed a" by auto with \?R\ obtain B where "B \ A" "finite (uminus`B)" "U \ \(uminus`B) = {}" by (metis subset_image_iff) then show "\T\A. finite T \ U \ \T" by (auto intro!: exI[of _ B] inj_setminus dest: finite_imageD) qed lemma compact_imp_fip: assumes "compact S" and "\T. T \ F \ closed T" and "\F'. finite F' \ F' \ F \ S \ (\F') \ {}" shows "S \ (\F) \ {}" using assms unfolding compact_fip by auto lemma compact_imp_fip_image: assumes "compact s" and P: "\i. i \ I \ closed (f i)" and Q: "\I'. finite I' \ I' \ I \ (s \ (\i\I'. f i) \ {})" shows "s \ (\i\I. f i) \ {}" proof - from P have "\i \ f ` I. closed i" by blast moreover have "\A. finite A \ A \ f ` I \ (s \ (\A) \ {})" by (metis Q finite_subset_image) ultimately show "s \ (\(f ` I)) \ {}" by (metis \compact s\ compact_imp_fip) qed end lemma (in t2_space) compact_imp_closed: assumes "compact s" shows "closed s" unfolding closed_def proof (rule openI) fix y assume "y \ - s" let ?C = "\x\s. {u. open u \ x \ u \ eventually (\y. y \ u) (nhds y)}" have "s \ \?C" proof fix x assume "x \ s" with \y \ - s\ have "x \ y" by clarsimp then have "\u v. open u \ open v \ x \ u \ y \ v \ u \ v = {}" by (rule hausdorff) with \x \ s\ show "x \ \?C" unfolding eventually_nhds by auto qed then obtain D where "D \ ?C" and "finite D" and "s \ \D" by (rule compactE [OF \compact s\]) auto from \D \ ?C\ have "\x\D. eventually (\y. y \ x) (nhds y)" by auto with \finite D\ have "eventually (\y. y \ \D) (nhds y)" by (simp add: eventually_ball_finite) with \s \ \D\ have "eventually (\y. y \ s) (nhds y)" by (auto elim!: eventually_mono) then show "\t. open t \ y \ t \ t \ - s" by (simp add: eventually_nhds subset_eq) qed lemma compact_continuous_image: assumes f: "continuous_on s f" and s: "compact s" shows "compact (f ` s)" proof (rule compactI) fix C assume "\c\C. open c" and cover: "f`s \ \C" with f have "\c\C. \A. open A \ A \ s = f -` c \ s" unfolding continuous_on_open_invariant by blast then obtain A where A: "\c\C. open (A c) \ A c \ s = f -` c \ s" unfolding bchoice_iff .. with cover have "\c. c \ C \ open (A c)" "s \ (\c\C. A c)" by (fastforce simp add: subset_eq set_eq_iff)+ from compactE_image[OF s this] obtain D where "D \ C" "finite D" "s \ (\c\D. A c)" . with A show "\D \ C. finite D \ f`s \ \D" by (intro exI[of _ D]) (fastforce simp add: subset_eq set_eq_iff)+ qed lemma continuous_on_inv: fixes f :: "'a::topological_space \ 'b::t2_space" assumes "continuous_on s f" and "compact s" and "\x\s. g (f x) = x" shows "continuous_on (f ` s) g" unfolding continuous_on_topological proof (clarsimp simp add: assms(3)) fix x :: 'a and B :: "'a set" assume "x \ s" and "open B" and "x \ B" have 1: "\x\s. f x \ f ` (s - B) \ x \ s - B" using assms(3) by (auto, metis) have "continuous_on (s - B) f" using \continuous_on s f\ Diff_subset by (rule continuous_on_subset) moreover have "compact (s - B)" using \open B\ and \compact s\ unfolding Diff_eq by (intro compact_Int_closed closed_Compl) ultimately have "compact (f ` (s - B))" by (rule compact_continuous_image) then have "closed (f ` (s - B))" by (rule compact_imp_closed) then have "open (- f ` (s - B))" by (rule open_Compl) moreover have "f x \ - f ` (s - B)" using \x \ s\ and \x \ B\ by (simp add: 1) moreover have "\y\s. f y \ - f ` (s - B) \ y \ B" by (simp add: 1) ultimately show "\A. open A \ f x \ A \ (\y\s. f y \ A \ y \ B)" by fast qed lemma continuous_on_inv_into: fixes f :: "'a::topological_space \ 'b::t2_space" assumes s: "continuous_on s f" "compact s" and f: "inj_on f s" shows "continuous_on (f ` s) (the_inv_into s f)" by (rule continuous_on_inv[OF s]) (auto simp: the_inv_into_f_f[OF f]) lemma (in linorder_topology) compact_attains_sup: assumes "compact S" "S \ {}" shows "\s\S. \t\S. t \ s" proof (rule classical) assume "\ (\s\S. \t\S. t \ s)" then obtain t where t: "\s\S. t s \ S" and "\s\S. s < t s" by (metis not_le) then have "\s. s\S \ open {..< t s}" "S \ (\s\S. {..< t s})" by auto with \compact S\ obtain C where "C \ S" "finite C" and C: "S \ (\s\C. {..< t s})" by (metis compactE_image) with \S \ {}\ have Max: "Max (t`C) \ t`C" and "\s\t`C. s \ Max (t`C)" by (auto intro!: Max_in) with C have "S \ {..< Max (t`C)}" by (auto intro: less_le_trans simp: subset_eq) with t Max \C \ S\ show ?thesis by fastforce qed lemma (in linorder_topology) compact_attains_inf: assumes "compact S" "S \ {}" shows "\s\S. \t\S. s \ t" proof (rule classical) assume "\ (\s\S. \t\S. s \ t)" then obtain t where t: "\s\S. t s \ S" and "\s\S. t s < s" by (metis not_le) then have "\s. s\S \ open {t s <..}" "S \ (\s\S. {t s <..})" by auto with \compact S\ obtain C where "C \ S" "finite C" and C: "S \ (\s\C. {t s <..})" by (metis compactE_image) with \S \ {}\ have Min: "Min (t`C) \ t`C" and "\s\t`C. Min (t`C) \ s" by (auto intro!: Min_in) with C have "S \ {Min (t`C) <..}" by (auto intro: le_less_trans simp: subset_eq) with t Min \C \ S\ show ?thesis by fastforce qed lemma continuous_attains_sup: fixes f :: "'a::topological_space \ 'b::linorder_topology" shows "compact s \ s \ {} \ continuous_on s f \ (\x\s. \y\s. f y \ f x)" using compact_attains_sup[of "f ` s"] compact_continuous_image[of s f] by auto lemma continuous_attains_inf: fixes f :: "'a::topological_space \ 'b::linorder_topology" shows "compact s \ s \ {} \ continuous_on s f \ (\x\s. \y\s. f x \ f y)" using compact_attains_inf[of "f ` s"] compact_continuous_image[of s f] by auto subsection \Connectedness\ context topological_space begin definition "connected S \ \ (\A B. open A \ open B \ S \ A \ B \ A \ B \ S = {} \ A \ S \ {} \ B \ S \ {})" lemma connectedI: "(\A B. open A \ open B \ A \ U \ {} \ B \ U \ {} \ A \ B \ U = {} \ U \ A \ B \ False) \ connected U" by (auto simp: connected_def) lemma connected_empty [simp]: "connected {}" by (auto intro!: connectedI) lemma connected_sing [simp]: "connected {x}" by (auto intro!: connectedI) lemma connectedD: "connected A \ open U \ open V \ U \ V \ A = {} \ A \ U \ V \ U \ A = {} \ V \ A = {}" by (auto simp: connected_def) end lemma connected_closed: "connected s \ \ (\A B. closed A \ closed B \ s \ A \ B \ A \ B \ s = {} \ A \ s \ {} \ B \ s \ {})" apply (simp add: connected_def del: ex_simps, safe) apply (drule_tac x="-A" in spec) apply (drule_tac x="-B" in spec) apply (fastforce simp add: closed_def [symmetric]) apply (drule_tac x="-A" in spec) apply (drule_tac x="-B" in spec) apply (fastforce simp add: open_closed [symmetric]) done lemma connected_closedD: "\connected s; A \ B \ s = {}; s \ A \ B; closed A; closed B\ \ A \ s = {} \ B \ s = {}" by (simp add: connected_closed) lemma connected_Union: assumes cs: "\s. s \ S \ connected s" and ne: "\S \ {}" shows "connected(\S)" proof (rule connectedI) fix A B assume A: "open A" and B: "open B" and Alap: "A \ \S \ {}" and Blap: "B \ \S \ {}" and disj: "A \ B \ \S = {}" and cover: "\S \ A \ B" have disjs:"\s. s \ S \ A \ B \ s = {}" using disj by auto obtain sa where sa: "sa \ S" "A \ sa \ {}" using Alap by auto obtain sb where sb: "sb \ S" "B \ sb \ {}" using Blap by auto obtain x where x: "\s. s \ S \ x \ s" using ne by auto then have "x \ \S" using \sa \ S\ by blast then have "x \ A \ x \ B" using cover by auto then show False using cs [unfolded connected_def] by (metis A B IntI Sup_upper sa sb disjs x cover empty_iff subset_trans) qed lemma connected_Un: "connected s \ connected t \ s \ t \ {} \ connected (s \ t)" using connected_Union [of "{s,t}"] by auto lemma connected_diff_open_from_closed: assumes st: "s \ t" and tu: "t \ u" and s: "open s" and t: "closed t" and u: "connected u" and ts: "connected (t - s)" shows "connected(u - s)" proof (rule connectedI) fix A B assume AB: "open A" "open B" "A \ (u - s) \ {}" "B \ (u - s) \ {}" and disj: "A \ B \ (u - s) = {}" and cover: "u - s \ A \ B" then consider "A \ (t - s) = {}" | "B \ (t - s) = {}" using st ts tu connectedD [of "t-s" "A" "B"] by auto then show False proof cases case 1 then have "(A - t) \ (B \ s) \ u = {}" using disj st by auto moreover have "u \ (A - t) \ (B \ s)" using 1 cover by auto ultimately show False using connectedD [of u "A - t" "B \ s"] AB s t 1 u by auto next case 2 then have "(A \ s) \ (B - t) \ u = {}" using disj st by auto moreover have "u \ (A \ s) \ (B - t)" using 2 cover by auto ultimately show False using connectedD [of u "A \ s" "B - t"] AB s t 2 u by auto qed qed lemma connected_iff_const: fixes S :: "'a::topological_space set" shows "connected S \ (\P::'a \ bool. continuous_on S P \ (\c. \s\S. P s = c))" proof safe fix P :: "'a \ bool" assume "connected S" "continuous_on S P" then have "\b. \A. open A \ A \ S = P -` {b} \ S" unfolding continuous_on_open_invariant by (simp add: open_discrete) from this[of True] this[of False] obtain t f where "open t" "open f" and *: "f \ S = P -` {False} \ S" "t \ S = P -` {True} \ S" by meson then have "t \ S = {} \ f \ S = {}" by (intro connectedD[OF \connected S\]) auto then show "\c. \s\S. P s = c" proof (rule disjE) assume "t \ S = {}" then show ?thesis unfolding * by (intro exI[of _ False]) auto next assume "f \ S = {}" then show ?thesis unfolding * by (intro exI[of _ True]) auto qed next assume P: "\P::'a \ bool. continuous_on S P \ (\c. \s\S. P s = c)" show "connected S" proof (rule connectedI) fix A B assume *: "open A" "open B" "A \ S \ {}" "B \ S \ {}" "A \ B \ S = {}" "S \ A \ B" have "continuous_on S (\x. x \ A)" unfolding continuous_on_open_invariant proof safe fix C :: "bool set" have "C = UNIV \ C = {True} \ C = {False} \ C = {}" using subset_UNIV[of C] unfolding UNIV_bool by auto with * show "\T. open T \ T \ S = (\x. x \ A) -` C \ S" by (intro exI[of _ "(if True \ C then A else {}) \ (if False \ C then B else {})"]) auto qed from P[rule_format, OF this] obtain c where "\s. s \ S \ (s \ A) = c" by blast with * show False by (cases c) auto qed qed lemma connectedD_const: "connected S \ continuous_on S P \ \c. \s\S. P s = c" for P :: "'a::topological_space \ bool" by (auto simp: connected_iff_const) lemma connectedI_const: "(\P::'a::topological_space \ bool. continuous_on S P \ \c. \s\S. P s = c) \ connected S" by (auto simp: connected_iff_const) lemma connected_local_const: assumes "connected A" "a \ A" "b \ A" and *: "\a\A. eventually (\b. f a = f b) (at a within A)" shows "f a = f b" proof - obtain S where S: "\a. a \ A \ a \ S a" "\a. a \ A \ open (S a)" "\a x. a \ A \ x \ S a \ x \ A \ f a = f x" using * unfolding eventually_at_topological by metis let ?P = "\b\{b\A. f a = f b}. S b" and ?N = "\b\{b\A. f a \ f b}. S b" have "?P \ A = {} \ ?N \ A = {}" using \connected A\ S \a\A\ by (intro connectedD) (auto, metis) then show "f a = f b" proof assume "?N \ A = {}" then have "\x\A. f a = f x" using S(1) by auto with \b\A\ show ?thesis by auto next assume "?P \ A = {}" then show ?thesis using \a \ A\ S(1)[of a] by auto qed qed lemma (in linorder_topology) connectedD_interval: assumes "connected U" and xy: "x \ U" "y \ U" and "x \ z" "z \ y" shows "z \ U" proof - have eq: "{.. {z<..} = - {z}" by auto have "\ connected U" if "z \ U" "x < z" "z < y" using xy that apply (simp only: connected_def simp_thms) apply (rule_tac exI[of _ "{..< z}"]) apply (rule_tac exI[of _ "{z <..}"]) apply (auto simp add: eq) done with assms show "z \ U" by (metis less_le) qed lemma (in linorder_topology) not_in_connected_cases: assumes conn: "connected S" assumes nbdd: "x \ S" assumes ne: "S \ {}" obtains "bdd_above S" "\y. y \ S \ x \ y" | "bdd_below S" "\y. y \ S \ x \ y" proof - obtain s where "s \ S" using ne by blast { assume "s \ x" have "False" if "x \ y" "y \ S" for y using connectedD_interval[OF conn \s \ S\ \y \ S\ \s \ x\ \x \ y\] \x \ S\ by simp then have wit: "y \ S \ x \ y" for y using le_cases by blast then have "bdd_above S" by (rule local.bdd_aboveI) note this wit } moreover { assume "x \ s" have "False" if "x \ y" "y \ S" for y using connectedD_interval[OF conn \y \ S\ \s \ S\ \x \ y\ \s \ x\ ] \x \ S\ by simp then have wit: "y \ S \ x \ y" for y using le_cases by blast then have "bdd_below S" by (rule bdd_belowI) note this wit } ultimately show ?thesis by (meson le_cases that) qed lemma connected_continuous_image: assumes *: "continuous_on s f" and "connected s" shows "connected (f ` s)" proof (rule connectedI_const) fix P :: "'b \ bool" assume "continuous_on (f ` s) P" then have "continuous_on s (P \ f)" by (rule continuous_on_compose[OF *]) from connectedD_const[OF \connected s\ this] show "\c. \s\f ` s. P s = c" by auto qed section \Linear Continuum Topologies\ class linear_continuum_topology = linorder_topology + linear_continuum begin lemma Inf_notin_open: assumes A: "open A" and bnd: "\a\A. x < a" shows "Inf A \ A" proof assume "Inf A \ A" then obtain b where "b < Inf A" "{b <.. Inf A} \ A" using open_left[of A "Inf A" x] assms by auto with dense[of b "Inf A"] obtain c where "c < Inf A" "c \ A" by (auto simp: subset_eq) then show False using cInf_lower[OF \c \ A\] bnd by (metis not_le less_imp_le bdd_belowI) qed lemma Sup_notin_open: assumes A: "open A" and bnd: "\a\A. a < x" shows "Sup A \ A" proof assume "Sup A \ A" with assms obtain b where "Sup A < b" "{Sup A ..< b} \ A" using open_right[of A "Sup A" x] by auto with dense[of "Sup A" b] obtain c where "Sup A < c" "c \ A" by (auto simp: subset_eq) then show False using cSup_upper[OF \c \ A\] bnd by (metis less_imp_le not_le bdd_aboveI) qed end instance linear_continuum_topology \ perfect_space proof fix x :: 'a obtain y where "x < y \ y < x" using ex_gt_or_lt [of x] .. with Inf_notin_open[of "{x}" y] Sup_notin_open[of "{x}" y] show "\ open {x}" by auto qed lemma connectedI_interval: fixes U :: "'a :: linear_continuum_topology set" assumes *: "\x y z. x \ U \ y \ U \ x \ z \ z \ y \ z \ U" shows "connected U" proof (rule connectedI) { fix A B assume "open A" "open B" "A \ B \ U = {}" "U \ A \ B" fix x y assume "x < y" "x \ A" "y \ B" "x \ U" "y \ U" let ?z = "Inf (B \ {x <..})" have "x \ ?z" "?z \ y" using \y \ B\ \x < y\ by (auto intro: cInf_lower cInf_greatest) with \x \ U\ \y \ U\ have "?z \ U" by (rule *) moreover have "?z \ B \ {x <..}" using \open B\ by (intro Inf_notin_open) auto ultimately have "?z \ A" using \x \ ?z\ \A \ B \ U = {}\ \x \ A\ \U \ A \ B\ by auto have "\b\B. b \ A \ b \ U" if "?z < y" proof - obtain a where "?z < a" "{?z ..< a} \ A" using open_right[OF \open A\ \?z \ A\ \?z < y\] by auto moreover obtain b where "b \ B" "x < b" "b < min a y" using cInf_less_iff[of "B \ {x <..}" "min a y"] \?z < a\ \?z < y\ \x < y\ \y \ B\ by auto moreover have "?z \ b" using \b \ B\ \x < b\ by (intro cInf_lower) auto moreover have "b \ U" using \x \ ?z\ \?z \ b\ \b < min a y\ by (intro *[OF \x \ U\ \y \ U\]) (auto simp: less_imp_le) ultimately show ?thesis by (intro bexI[of _ b]) auto qed then have False using \?z \ y\ \?z \ A\ \y \ B\ \y \ U\ \A \ B \ U = {}\ unfolding le_less by blast } note not_disjoint = this fix A B assume AB: "open A" "open B" "U \ A \ B" "A \ B \ U = {}" moreover assume "A \ U \ {}" then obtain x where x: "x \ U" "x \ A" by auto moreover assume "B \ U \ {}" then obtain y where y: "y \ U" "y \ B" by auto moreover note not_disjoint[of B A y x] not_disjoint[of A B x y] ultimately show False by (cases x y rule: linorder_cases) auto qed lemma connected_iff_interval: "connected U \ (\x\U. \y\U. \z. x \ z \ z \ y \ z \ U)" for U :: "'a::linear_continuum_topology set" by (auto intro: connectedI_interval dest: connectedD_interval) lemma connected_UNIV[simp]: "connected (UNIV::'a::linear_continuum_topology set)" by (simp add: connected_iff_interval) lemma connected_Ioi[simp]: "connected {a<..}" for a :: "'a::linear_continuum_topology" by (auto simp: connected_iff_interval) lemma connected_Ici[simp]: "connected {a..}" for a :: "'a::linear_continuum_topology" by (auto simp: connected_iff_interval) lemma connected_Iio[simp]: "connected {.. A" "b \ A" shows "{a <..< b} \ A" using connectedD_interval[OF assms] by (simp add: subset_eq Ball_def less_imp_le) lemma connected_contains_Icc: fixes A :: "'a::linorder_topology set" assumes "connected A" "a \ A" "b \ A" shows "{a..b} \ A" proof fix x assume "x \ {a..b}" then have "x = a \ x = b \ x \ {a<.. A" using assms connected_contains_Ioo[of A a b] by auto qed subsection \Intermediate Value Theorem\ lemma IVT': fixes f :: "'a::linear_continuum_topology \ 'b::linorder_topology" assumes y: "f a \ y" "y \ f b" "a \ b" and *: "continuous_on {a .. b} f" shows "\x. a \ x \ x \ b \ f x = y" proof - have "connected {a..b}" unfolding connected_iff_interval by auto from connected_continuous_image[OF * this, THEN connectedD_interval, of "f a" "f b" y] y show ?thesis by (auto simp add: atLeastAtMost_def atLeast_def atMost_def) qed lemma IVT2': fixes f :: "'a :: linear_continuum_topology \ 'b :: linorder_topology" assumes y: "f b \ y" "y \ f a" "a \ b" and *: "continuous_on {a .. b} f" shows "\x. a \ x \ x \ b \ f x = y" proof - have "connected {a..b}" unfolding connected_iff_interval by auto from connected_continuous_image[OF * this, THEN connectedD_interval, of "f b" "f a" y] y show ?thesis by (auto simp add: atLeastAtMost_def atLeast_def atMost_def) qed lemma IVT: fixes f :: "'a::linear_continuum_topology \ 'b::linorder_topology" shows "f a \ y \ y \ f b \ a \ b \ (\x. a \ x \ x \ b \ isCont f x) \ \x. a \ x \ x \ b \ f x = y" by (rule IVT') (auto intro: continuous_at_imp_continuous_on) lemma IVT2: fixes f :: "'a::linear_continuum_topology \ 'b::linorder_topology" shows "f b \ y \ y \ f a \ a \ b \ (\x. a \ x \ x \ b \ isCont f x) \ \x. a \ x \ x \ b \ f x = y" by (rule IVT2') (auto intro: continuous_at_imp_continuous_on) lemma continuous_inj_imp_mono: fixes f :: "'a::linear_continuum_topology \ 'b::linorder_topology" assumes x: "a < x" "x < b" and cont: "continuous_on {a..b} f" and inj: "inj_on f {a..b}" shows "(f a < f x \ f x < f b) \ (f b < f x \ f x < f a)" proof - note I = inj_on_eq_iff[OF inj] { assume "f x < f a" "f x < f b" then obtain s t where "x \ s" "s \ b" "a \ t" "t \ x" "f s = f t" "f x < f s" using IVT'[of f x "min (f a) (f b)" b] IVT2'[of f x "min (f a) (f b)" a] x by (auto simp: continuous_on_subset[OF cont] less_imp_le) with x I have False by auto } moreover { assume "f a < f x" "f b < f x" then obtain s t where "x \ s" "s \ b" "a \ t" "t \ x" "f s = f t" "f s < f x" using IVT'[of f a "max (f a) (f b)" x] IVT2'[of f b "max (f a) (f b)" x] x by (auto simp: continuous_on_subset[OF cont] less_imp_le) with x I have False by auto } ultimately show ?thesis using I[of a x] I[of x b] x less_trans[OF x] by (auto simp add: le_less less_imp_neq neq_iff) qed lemma continuous_at_Sup_mono: fixes f :: "'a::{linorder_topology,conditionally_complete_linorder} \ 'b::{linorder_topology,conditionally_complete_linorder}" assumes "mono f" and cont: "continuous (at_left (Sup S)) f" and S: "S \ {}" "bdd_above S" shows "f (Sup S) = (SUP s\S. f s)" proof (rule antisym) have f: "(f \ f (Sup S)) (at_left (Sup S))" using cont unfolding continuous_within . show "f (Sup S) \ (SUP s\S. f s)" proof cases assume "Sup S \ S" then show ?thesis by (rule cSUP_upper) (auto intro: bdd_above_image_mono S \mono f\) next assume "Sup S \ S" from \S \ {}\ obtain s where "s \ S" by auto with \Sup S \ S\ S have "s < Sup S" unfolding less_le by (blast intro: cSup_upper) show ?thesis proof (rule ccontr) assume "\ ?thesis" with order_tendstoD(1)[OF f, of "SUP s\S. f s"] obtain b where "b < Sup S" and *: "\y. b < y \ y < Sup S \ (SUP s\S. f s) < f y" by (auto simp: not_le eventually_at_left[OF \s < Sup S\]) with \S \ {}\ obtain c where "c \ S" "b < c" using less_cSupD[of S b] by auto with \Sup S \ S\ S have "c < Sup S" unfolding less_le by (blast intro: cSup_upper) from *[OF \b < c\ \c < Sup S\] cSUP_upper[OF \c \ S\ bdd_above_image_mono[of f]] show False by (auto simp: assms) qed qed qed (intro cSUP_least \mono f\[THEN monoD] cSup_upper S) lemma continuous_at_Sup_antimono: fixes f :: "'a::{linorder_topology,conditionally_complete_linorder} \ 'b::{linorder_topology,conditionally_complete_linorder}" assumes "antimono f" and cont: "continuous (at_left (Sup S)) f" and S: "S \ {}" "bdd_above S" shows "f (Sup S) = (INF s\S. f s)" proof (rule antisym) have f: "(f \ f (Sup S)) (at_left (Sup S))" using cont unfolding continuous_within . show "(INF s\S. f s) \ f (Sup S)" proof cases assume "Sup S \ S" then show ?thesis by (intro cINF_lower) (auto intro: bdd_below_image_antimono S \antimono f\) next assume "Sup S \ S" from \S \ {}\ obtain s where "s \ S" by auto with \Sup S \ S\ S have "s < Sup S" unfolding less_le by (blast intro: cSup_upper) show ?thesis proof (rule ccontr) assume "\ ?thesis" with order_tendstoD(2)[OF f, of "INF s\S. f s"] obtain b where "b < Sup S" and *: "\y. b < y \ y < Sup S \ f y < (INF s\S. f s)" by (auto simp: not_le eventually_at_left[OF \s < Sup S\]) with \S \ {}\ obtain c where "c \ S" "b < c" using less_cSupD[of S b] by auto with \Sup S \ S\ S have "c < Sup S" unfolding less_le by (blast intro: cSup_upper) from *[OF \b < c\ \c < Sup S\] cINF_lower[OF bdd_below_image_antimono, of f S c] \c \ S\ show False by (auto simp: assms) qed qed qed (intro cINF_greatest \antimono f\[THEN antimonoD] cSup_upper S) lemma continuous_at_Inf_mono: fixes f :: "'a::{linorder_topology,conditionally_complete_linorder} \ 'b::{linorder_topology,conditionally_complete_linorder}" assumes "mono f" and cont: "continuous (at_right (Inf S)) f" and S: "S \ {}" "bdd_below S" shows "f (Inf S) = (INF s\S. f s)" proof (rule antisym) have f: "(f \ f (Inf S)) (at_right (Inf S))" using cont unfolding continuous_within . show "(INF s\S. f s) \ f (Inf S)" proof cases assume "Inf S \ S" then show ?thesis by (rule cINF_lower[rotated]) (auto intro: bdd_below_image_mono S \mono f\) next assume "Inf S \ S" from \S \ {}\ obtain s where "s \ S" by auto with \Inf S \ S\ S have "Inf S < s" unfolding less_le by (blast intro: cInf_lower) show ?thesis proof (rule ccontr) assume "\ ?thesis" with order_tendstoD(2)[OF f, of "INF s\S. f s"] obtain b where "Inf S < b" and *: "\y. Inf S < y \ y < b \ f y < (INF s\S. f s)" by (auto simp: not_le eventually_at_right[OF \Inf S < s\]) with \S \ {}\ obtain c where "c \ S" "c < b" using cInf_lessD[of S b] by auto with \Inf S \ S\ S have "Inf S < c" unfolding less_le by (blast intro: cInf_lower) from *[OF \Inf S < c\ \c < b\] cINF_lower[OF bdd_below_image_mono[of f] \c \ S\] show False by (auto simp: assms) qed qed qed (intro cINF_greatest \mono f\[THEN monoD] cInf_lower \bdd_below S\ \S \ {}\) lemma continuous_at_Inf_antimono: fixes f :: "'a::{linorder_topology,conditionally_complete_linorder} \ 'b::{linorder_topology,conditionally_complete_linorder}" assumes "antimono f" and cont: "continuous (at_right (Inf S)) f" and S: "S \ {}" "bdd_below S" shows "f (Inf S) = (SUP s\S. f s)" proof (rule antisym) have f: "(f \ f (Inf S)) (at_right (Inf S))" using cont unfolding continuous_within . show "f (Inf S) \ (SUP s\S. f s)" proof cases assume "Inf S \ S" then show ?thesis by (rule cSUP_upper) (auto intro: bdd_above_image_antimono S \antimono f\) next assume "Inf S \ S" from \S \ {}\ obtain s where "s \ S" by auto with \Inf S \ S\ S have "Inf S < s" unfolding less_le by (blast intro: cInf_lower) show ?thesis proof (rule ccontr) assume "\ ?thesis" with order_tendstoD(1)[OF f, of "SUP s\S. f s"] obtain b where "Inf S < b" and *: "\y. Inf S < y \ y < b \ (SUP s\S. f s) < f y" by (auto simp: not_le eventually_at_right[OF \Inf S < s\]) with \S \ {}\ obtain c where "c \ S" "c < b" using cInf_lessD[of S b] by auto with \Inf S \ S\ S have "Inf S < c" unfolding less_le by (blast intro: cInf_lower) from *[OF \Inf S < c\ \c < b\] cSUP_upper[OF \c \ S\ bdd_above_image_antimono[of f]] show False by (auto simp: assms) qed qed qed (intro cSUP_least \antimono f\[THEN antimonoD] cInf_lower S) subsection \Uniform spaces\ class uniformity = fixes uniformity :: "('a \ 'a) filter" begin abbreviation uniformity_on :: "'a set \ ('a \ 'a) filter" where "uniformity_on s \ inf uniformity (principal (s\s))" end lemma uniformity_Abort: "uniformity = Filter.abstract_filter (\u. Code.abort (STR ''uniformity is not executable'') (\u. uniformity))" by simp class open_uniformity = "open" + uniformity + assumes open_uniformity: "\U. open U \ (\x\U. eventually (\(x', y). x' = x \ y \ U) uniformity)" begin subclass topological_space by standard (force elim: eventually_mono eventually_elim2 simp: split_beta' open_uniformity)+ end class uniform_space = open_uniformity + assumes uniformity_refl: "eventually E uniformity \ E (x, x)" and uniformity_sym: "eventually E uniformity \ eventually (\(x, y). E (y, x)) uniformity" and uniformity_trans: "eventually E uniformity \ \D. eventually D uniformity \ (\x y z. D (x, y) \ D (y, z) \ E (x, z))" begin lemma uniformity_bot: "uniformity \ bot" using uniformity_refl by auto lemma uniformity_trans': "eventually E uniformity \ eventually (\((x, y), (y', z)). y = y' \ E (x, z)) (uniformity \\<^sub>F uniformity)" by (drule uniformity_trans) (auto simp add: eventually_prod_same) lemma uniformity_transE: assumes "eventually E uniformity" obtains D where "eventually D uniformity" "\x y z. D (x, y) \ D (y, z) \ E (x, z)" using uniformity_trans [OF assms] by auto lemma eventually_nhds_uniformity: "eventually P (nhds x) \ eventually (\(x', y). x' = x \ P y) uniformity" (is "_ \ ?N P x") unfolding eventually_nhds proof safe assume *: "?N P x" have "?N (?N P) x" if "?N P x" for x proof - from that obtain D where ev: "eventually D uniformity" and D: "D (a, b) \ D (b, c) \ case (a, c) of (x', y) \ x' = x \ P y" for a b c by (rule uniformity_transE) simp from ev show ?thesis by eventually_elim (insert ev D, force elim: eventually_mono split: prod.split) qed then have "open {x. ?N P x}" by (simp add: open_uniformity) then show "\S. open S \ x \ S \ (\x\S. P x)" by (intro exI[of _ "{x. ?N P x}"]) (auto dest: uniformity_refl simp: *) qed (force simp add: open_uniformity elim: eventually_mono) subsubsection \Totally bounded sets\ definition totally_bounded :: "'a set \ bool" where "totally_bounded S \ (\E. eventually E uniformity \ (\X. finite X \ (\s\S. \x\X. E (x, s))))" lemma totally_bounded_empty[iff]: "totally_bounded {}" by (auto simp add: totally_bounded_def) lemma totally_bounded_subset: "totally_bounded S \ T \ S \ totally_bounded T" by (fastforce simp add: totally_bounded_def) lemma totally_bounded_Union[intro]: assumes M: "finite M" "\S. S \ M \ totally_bounded S" shows "totally_bounded (\M)" unfolding totally_bounded_def proof safe fix E assume "eventually E uniformity" with M obtain X where "\S\M. finite (X S) \ (\s\S. \x\X S. E (x, s))" by (metis totally_bounded_def) with \finite M\ show "\X. finite X \ (\s\\M. \x\X. E (x, s))" by (intro exI[of _ "\S\M. X S"]) force qed subsubsection \Cauchy filter\ definition cauchy_filter :: "'a filter \ bool" where "cauchy_filter F \ F \\<^sub>F F \ uniformity" definition Cauchy :: "(nat \ 'a) \ bool" where Cauchy_uniform: "Cauchy X = cauchy_filter (filtermap X sequentially)" lemma Cauchy_uniform_iff: "Cauchy X \ (\P. eventually P uniformity \ (\N. \n\N. \m\N. P (X n, X m)))" unfolding Cauchy_uniform cauchy_filter_def le_filter_def eventually_prod_same eventually_filtermap eventually_sequentially proof safe let ?U = "\P. eventually P uniformity" { fix P assume "?U P" "\P. ?U P \ (\Q. (\N. \n\N. Q (X n)) \ (\x y. Q x \ Q y \ P (x, y)))" then obtain Q N where "\n. n \ N \ Q (X n)" "\x y. Q x \ Q y \ P (x, y)" by metis then show "\N. \n\N. \m\N. P (X n, X m)" by blast next fix P assume "?U P" and P: "\P. ?U P \ (\N. \n\N. \m\N. P (X n, X m))" then obtain Q where "?U Q" and Q: "\x y z. Q (x, y) \ Q (y, z) \ P (x, z)" by (auto elim: uniformity_transE) then have "?U (\x. Q x \ (\(x, y). Q (y, x)) x)" unfolding eventually_conj_iff by (simp add: uniformity_sym) from P[rule_format, OF this] obtain N where N: "\n m. n \ N \ m \ N \ Q (X n, X m) \ Q (X m, X n)" by auto show "\Q. (\N. \n\N. Q (X n)) \ (\x y. Q x \ Q y \ P (x, y))" proof (safe intro!: exI[of _ "\x. \n\N. Q (x, X n) \ Q (X n, x)"] exI[of _ N] N) fix x y assume "\n\N. Q (x, X n) \ Q (X n, x)" "\n\N. Q (y, X n) \ Q (X n, y)" then have "Q (x, X N)" "Q (X N, y)" by auto then show "P (x, y)" by (rule Q) qed } qed lemma nhds_imp_cauchy_filter: assumes *: "F \ nhds x" shows "cauchy_filter F" proof - have "F \\<^sub>F F \ nhds x \\<^sub>F nhds x" by (intro prod_filter_mono *) also have "\ \ uniformity" unfolding le_filter_def eventually_nhds_uniformity eventually_prod_same proof safe fix P assume "eventually P uniformity" then obtain Ql where ev: "eventually Ql uniformity" and "Ql (x, y) \ Ql (y, z) \ P (x, z)" for x y z by (rule uniformity_transE) simp with ev[THEN uniformity_sym] show "\Q. eventually (\(x', y). x' = x \ Q y) uniformity \ (\x y. Q x \ Q y \ P (x, y))" by (rule_tac exI[of _ "\y. Ql (y, x) \ Ql (x, y)"]) (fastforce elim: eventually_elim2) qed finally show ?thesis by (simp add: cauchy_filter_def) qed lemma LIMSEQ_imp_Cauchy: "X \ x \ Cauchy X" unfolding Cauchy_uniform filterlim_def by (intro nhds_imp_cauchy_filter) lemma Cauchy_subseq_Cauchy: assumes "Cauchy X" "strict_mono f" shows "Cauchy (X \ f)" unfolding Cauchy_uniform comp_def filtermap_filtermap[symmetric] cauchy_filter_def by (rule order_trans[OF _ \Cauchy X\[unfolded Cauchy_uniform cauchy_filter_def]]) (intro prod_filter_mono filtermap_mono filterlim_subseq[OF \strict_mono f\, unfolded filterlim_def]) lemma convergent_Cauchy: "convergent X \ Cauchy X" unfolding convergent_def by (erule exE, erule LIMSEQ_imp_Cauchy) definition complete :: "'a set \ bool" where complete_uniform: "complete S \ (\F \ principal S. F \ bot \ cauchy_filter F \ (\x\S. F \ nhds x))" lemma (in uniform_space) cauchy_filter_complete_converges: assumes "cauchy_filter F" "complete A" "F \ principal A" "F \ bot" shows "\c. F \ nhds c" using assms unfolding complete_uniform by blast end subsubsection \Uniformly continuous functions\ definition uniformly_continuous_on :: "'a set \ ('a::uniform_space \ 'b::uniform_space) \ bool" where uniformly_continuous_on_uniformity: "uniformly_continuous_on s f \ (LIM (x, y) (uniformity_on s). (f x, f y) :> uniformity)" lemma uniformly_continuous_onD: "uniformly_continuous_on s f \ eventually E uniformity \ eventually (\(x, y). x \ s \ y \ s \ E (f x, f y)) uniformity" by (simp add: uniformly_continuous_on_uniformity filterlim_iff eventually_inf_principal split_beta' mem_Times_iff imp_conjL) lemma uniformly_continuous_on_const[continuous_intros]: "uniformly_continuous_on s (\x. c)" by (auto simp: uniformly_continuous_on_uniformity filterlim_iff uniformity_refl) lemma uniformly_continuous_on_id[continuous_intros]: "uniformly_continuous_on s (\x. x)" by (auto simp: uniformly_continuous_on_uniformity filterlim_def) lemma uniformly_continuous_on_compose: "uniformly_continuous_on s g \ uniformly_continuous_on (g`s) f \ uniformly_continuous_on s (\x. f (g x))" using filterlim_compose[of "\(x, y). (f x, f y)" uniformity "uniformity_on (g`s)" "\(x, y). (g x, g y)" "uniformity_on s"] by (simp add: split_beta' uniformly_continuous_on_uniformity filterlim_inf filterlim_principal eventually_inf_principal mem_Times_iff) lemma uniformly_continuous_imp_continuous: assumes f: "uniformly_continuous_on s f" shows "continuous_on s f" by (auto simp: filterlim_iff eventually_at_filter eventually_nhds_uniformity continuous_on_def elim: eventually_mono dest!: uniformly_continuous_onD[OF f]) section \Product Topology\ subsection \Product is a topological space\ instantiation prod :: (topological_space, topological_space) topological_space begin definition open_prod_def[code del]: "open (S :: ('a \ 'b) set) \ (\x\S. \A B. open A \ open B \ x \ A \ B \ A \ B \ S)" lemma open_prod_elim: assumes "open S" and "x \ S" obtains A B where "open A" and "open B" and "x \ A \ B" and "A \ B \ S" using assms unfolding open_prod_def by fast lemma open_prod_intro: assumes "\x. x \ S \ \A B. open A \ open B \ x \ A \ B \ A \ B \ S" shows "open S" using assms unfolding open_prod_def by fast instance proof show "open (UNIV :: ('a \ 'b) set)" unfolding open_prod_def by auto next fix S T :: "('a \ 'b) set" assume "open S" "open T" show "open (S \ T)" proof (rule open_prod_intro) fix x assume x: "x \ S \ T" from x have "x \ S" by simp obtain Sa Sb where A: "open Sa" "open Sb" "x \ Sa \ Sb" "Sa \ Sb \ S" using \open S\ and \x \ S\ by (rule open_prod_elim) from x have "x \ T" by simp obtain Ta Tb where B: "open Ta" "open Tb" "x \ Ta \ Tb" "Ta \ Tb \ T" using \open T\ and \x \ T\ by (rule open_prod_elim) let ?A = "Sa \ Ta" and ?B = "Sb \ Tb" have "open ?A \ open ?B \ x \ ?A \ ?B \ ?A \ ?B \ S \ T" using A B by (auto simp add: open_Int) then show "\A B. open A \ open B \ x \ A \ B \ A \ B \ S \ T" by fast qed next fix K :: "('a \ 'b) set set" assume "\S\K. open S" then show "open (\K)" unfolding open_prod_def by fast qed end declare [[code abort: "open :: ('a::topological_space \ 'b::topological_space) set \ bool"]] lemma open_Times: "open S \ open T \ open (S \ T)" unfolding open_prod_def by auto lemma fst_vimage_eq_Times: "fst -` S = S \ UNIV" by auto lemma snd_vimage_eq_Times: "snd -` S = UNIV \ S" by auto lemma open_vimage_fst: "open S \ open (fst -` S)" by (simp add: fst_vimage_eq_Times open_Times) lemma open_vimage_snd: "open S \ open (snd -` S)" by (simp add: snd_vimage_eq_Times open_Times) lemma closed_vimage_fst: "closed S \ closed (fst -` S)" unfolding closed_open vimage_Compl [symmetric] by (rule open_vimage_fst) lemma closed_vimage_snd: "closed S \ closed (snd -` S)" unfolding closed_open vimage_Compl [symmetric] by (rule open_vimage_snd) lemma closed_Times: "closed S \ closed T \ closed (S \ T)" proof - have "S \ T = (fst -` S) \ (snd -` T)" by auto then show "closed S \ closed T \ closed (S \ T)" by (simp add: closed_vimage_fst closed_vimage_snd closed_Int) qed lemma subset_fst_imageI: "A \ B \ S \ y \ B \ A \ fst ` S" unfolding image_def subset_eq by force lemma subset_snd_imageI: "A \ B \ S \ x \ A \ B \ snd ` S" unfolding image_def subset_eq by force lemma open_image_fst: assumes "open S" shows "open (fst ` S)" proof (rule openI) fix x assume "x \ fst ` S" then obtain y where "(x, y) \ S" by auto then obtain A B where "open A" "open B" "x \ A" "y \ B" "A \ B \ S" using \open S\ unfolding open_prod_def by auto from \A \ B \ S\ \y \ B\ have "A \ fst ` S" by (rule subset_fst_imageI) with \open A\ \x \ A\ have "open A \ x \ A \ A \ fst ` S" by simp then show "\T. open T \ x \ T \ T \ fst ` S" .. qed lemma open_image_snd: assumes "open S" shows "open (snd ` S)" proof (rule openI) fix y assume "y \ snd ` S" then obtain x where "(x, y) \ S" by auto then obtain A B where "open A" "open B" "x \ A" "y \ B" "A \ B \ S" using \open S\ unfolding open_prod_def by auto from \A \ B \ S\ \x \ A\ have "B \ snd ` S" by (rule subset_snd_imageI) with \open B\ \y \ B\ have "open B \ y \ B \ B \ snd ` S" by simp then show "\T. open T \ y \ T \ T \ snd ` S" .. qed lemma nhds_prod: "nhds (a, b) = nhds a \\<^sub>F nhds b" unfolding nhds_def proof (subst prod_filter_INF, auto intro!: antisym INF_greatest simp: principal_prod_principal) fix S T assume "open S" "a \ S" "open T" "b \ T" then show "(INF x \ {S. open S \ (a, b) \ S}. principal x) \ principal (S \ T)" by (intro INF_lower) (auto intro!: open_Times) next fix S' assume "open S'" "(a, b) \ S'" then obtain S T where "open S" "a \ S" "open T" "b \ T" "S \ T \ S'" by (auto elim: open_prod_elim) then show "(INF x \ {S. open S \ a \ S}. INF y \ {S. open S \ b \ S}. principal (x \ y)) \ principal S'" by (auto intro!: INF_lower2) qed subsubsection \Continuity of operations\ lemma tendsto_fst [tendsto_intros]: assumes "(f \ a) F" shows "((\x. fst (f x)) \ fst a) F" proof (rule topological_tendstoI) fix S assume "open S" and "fst a \ S" then have "open (fst -` S)" and "a \ fst -` S" by (simp_all add: open_vimage_fst) with assms have "eventually (\x. f x \ fst -` S) F" by (rule topological_tendstoD) then show "eventually (\x. fst (f x) \ S) F" by simp qed lemma tendsto_snd [tendsto_intros]: assumes "(f \ a) F" shows "((\x. snd (f x)) \ snd a) F" proof (rule topological_tendstoI) fix S assume "open S" and "snd a \ S" then have "open (snd -` S)" and "a \ snd -` S" by (simp_all add: open_vimage_snd) with assms have "eventually (\x. f x \ snd -` S) F" by (rule topological_tendstoD) then show "eventually (\x. snd (f x) \ S) F" by simp qed lemma tendsto_Pair [tendsto_intros]: assumes "(f \ a) F" and "(g \ b) F" shows "((\x. (f x, g x)) \ (a, b)) F" unfolding nhds_prod using assms by (rule filterlim_Pair) lemma continuous_fst[continuous_intros]: "continuous F f \ continuous F (\x. fst (f x))" unfolding continuous_def by (rule tendsto_fst) lemma continuous_snd[continuous_intros]: "continuous F f \ continuous F (\x. snd (f x))" unfolding continuous_def by (rule tendsto_snd) lemma continuous_Pair[continuous_intros]: "continuous F f \ continuous F g \ continuous F (\x. (f x, g x))" unfolding continuous_def by (rule tendsto_Pair) lemma continuous_on_fst[continuous_intros]: "continuous_on s f \ continuous_on s (\x. fst (f x))" unfolding continuous_on_def by (auto intro: tendsto_fst) lemma continuous_on_snd[continuous_intros]: "continuous_on s f \ continuous_on s (\x. snd (f x))" unfolding continuous_on_def by (auto intro: tendsto_snd) lemma continuous_on_Pair[continuous_intros]: "continuous_on s f \ continuous_on s g \ continuous_on s (\x. (f x, g x))" unfolding continuous_on_def by (auto intro: tendsto_Pair) lemma continuous_on_swap[continuous_intros]: "continuous_on A prod.swap" by (simp add: prod.swap_def continuous_on_fst continuous_on_snd continuous_on_Pair continuous_on_id) lemma continuous_on_swap_args: assumes "continuous_on (A\B) (\(x,y). d x y)" shows "continuous_on (B\A) (\(x,y). d y x)" proof - have "(\(x,y). d y x) = (\(x,y). d x y) \ prod.swap" by force then show ?thesis by (metis assms continuous_on_compose continuous_on_swap product_swap) qed lemma isCont_fst [simp]: "isCont f a \ isCont (\x. fst (f x)) a" by (fact continuous_fst) lemma isCont_snd [simp]: "isCont f a \ isCont (\x. snd (f x)) a" by (fact continuous_snd) lemma isCont_Pair [simp]: "\isCont f a; isCont g a\ \ isCont (\x. (f x, g x)) a" by (fact continuous_Pair) lemma continuous_on_compose_Pair: assumes f: "continuous_on (Sigma A B) (\(a, b). f a b)" assumes g: "continuous_on C g" assumes h: "continuous_on C h" assumes subset: "\c. c \ C \ g c \ A" "\c. c \ C \ h c \ B (g c)" shows "continuous_on C (\c. f (g c) (h c))" using continuous_on_compose2[OF f continuous_on_Pair[OF g h]] subset by auto subsubsection \Connectedness of products\ proposition connected_Times: assumes S: "connected S" and T: "connected T" shows "connected (S \ T)" proof (rule connectedI_const) fix P::"'a \ 'b \ bool" assume P[THEN continuous_on_compose2, continuous_intros]: "continuous_on (S \ T) P" have "continuous_on S (\s. P (s, t))" if "t \ T" for t by (auto intro!: continuous_intros that) from connectedD_const[OF S this] obtain c1 where c1: "\s t. t \ T \ s \ S \ P (s, t) = c1 t" by metis moreover have "continuous_on T (\t. P (s, t))" if "s \ S" for s by (auto intro!: continuous_intros that) from connectedD_const[OF T this] obtain c2 where "\s t. t \ T \ s \ S \ P (s, t) = c2 s" by metis ultimately show "\c. \s\S \ T. P s = c" by auto qed corollary connected_Times_eq [simp]: "connected (S \ T) \ S = {} \ T = {} \ connected S \ connected T" (is "?lhs = ?rhs") proof assume L: ?lhs show ?rhs proof cases assume "S \ {} \ T \ {}" moreover have "connected (fst ` (S \ T))" "connected (snd ` (S \ T))" using continuous_on_fst continuous_on_snd continuous_on_id by (blast intro: connected_continuous_image [OF _ L])+ ultimately show ?thesis by auto qed auto qed (auto simp: connected_Times) subsubsection \Separation axioms\ instance prod :: (t0_space, t0_space) t0_space proof fix x y :: "'a \ 'b" assume "x \ y" then have "fst x \ fst y \ snd x \ snd y" by (simp add: prod_eq_iff) then show "\U. open U \ (x \ U) \ (y \ U)" by (fast dest: t0_space elim: open_vimage_fst open_vimage_snd) qed instance prod :: (t1_space, t1_space) t1_space proof fix x y :: "'a \ 'b" assume "x \ y" then have "fst x \ fst y \ snd x \ snd y" by (simp add: prod_eq_iff) then show "\U. open U \ x \ U \ y \ U" by (fast dest: t1_space elim: open_vimage_fst open_vimage_snd) qed instance prod :: (t2_space, t2_space) t2_space proof fix x y :: "'a \ 'b" assume "x \ y" then have "fst x \ fst y \ snd x \ snd y" by (simp add: prod_eq_iff) then show "\U V. open U \ open V \ x \ U \ y \ V \ U \ V = {}" by (fast dest: hausdorff elim: open_vimage_fst open_vimage_snd) qed lemma isCont_swap[continuous_intros]: "isCont prod.swap a" using continuous_on_eq_continuous_within continuous_on_swap by blast lemma open_diagonal_complement: "open {(x,y) |x y. x \ (y::('a::t2_space))}" proof - have "open {(x, y). x \ (y::'a)}" unfolding split_def by (intro open_Collect_neq continuous_intros) also have "{(x, y). x \ (y::'a)} = {(x, y) |x y. x \ (y::'a)}" by auto finally show ?thesis . qed lemma closed_diagonal: "closed {y. \ x::('a::t2_space). y = (x,x)}" proof - have "{y. \ x::'a. y = (x,x)} = UNIV - {(x,y) | x y. x \ y}" by auto then show ?thesis using open_diagonal_complement closed_Diff by auto qed lemma open_superdiagonal: "open {(x,y) | x y. x > (y::'a::{linorder_topology})}" proof - have "open {(x, y). x > (y::'a)}" unfolding split_def by (intro open_Collect_less continuous_intros) also have "{(x, y). x > (y::'a)} = {(x, y) |x y. x > (y::'a)}" by auto finally show ?thesis . qed lemma closed_subdiagonal: "closed {(x,y) | x y. x \ (y::'a::{linorder_topology})}" proof - have "{(x,y) | x y. x \ (y::'a)} = UNIV - {(x,y) | x y. x > (y::'a)}" by auto then show ?thesis using open_superdiagonal closed_Diff by auto qed lemma open_subdiagonal: "open {(x,y) | x y. x < (y::'a::{linorder_topology})}" proof - have "open {(x, y). x < (y::'a)}" unfolding split_def by (intro open_Collect_less continuous_intros) also have "{(x, y). x < (y::'a)} = {(x, y) |x y. x < (y::'a)}" by auto finally show ?thesis . qed lemma closed_superdiagonal: "closed {(x,y) | x y. x \ (y::('a::{linorder_topology}))}" proof - have "{(x,y) | x y. x \ (y::'a)} = UNIV - {(x,y) | x y. x < y}" by auto then show ?thesis using open_subdiagonal closed_Diff by auto qed end diff --git a/src/HOL/Transcendental.thy b/src/HOL/Transcendental.thy --- a/src/HOL/Transcendental.thy +++ b/src/HOL/Transcendental.thy @@ -1,7428 +1,7463 @@ (* Title: HOL/Transcendental.thy Author: Jacques D. Fleuriot, University of Cambridge, University of Edinburgh Author: Lawrence C Paulson Author: Jeremy Avigad *) section \Power Series, Transcendental Functions etc.\ theory Transcendental imports Series Deriv NthRoot begin text \A theorem about the factcorial function on the reals.\ lemma square_fact_le_2_fact: "fact n * fact n \ (fact (2 * n) :: real)" proof (induct n) case 0 then show ?case by simp next case (Suc n) have "(fact (Suc n)) * (fact (Suc n)) = of_nat (Suc n) * of_nat (Suc n) * (fact n * fact n :: real)" by (simp add: field_simps) also have "\ \ of_nat (Suc n) * of_nat (Suc n) * fact (2 * n)" by (rule mult_left_mono [OF Suc]) simp also have "\ \ of_nat (Suc (Suc (2 * n))) * of_nat (Suc (2 * n)) * fact (2 * n)" by (rule mult_right_mono)+ (auto simp: field_simps) also have "\ = fact (2 * Suc n)" by (simp add: field_simps) finally show ?case . qed lemma fact_in_Reals: "fact n \ \" by (induction n) auto lemma of_real_fact [simp]: "of_real (fact n) = fact n" by (metis of_nat_fact of_real_of_nat_eq) lemma pochhammer_of_real: "pochhammer (of_real x) n = of_real (pochhammer x n)" by (simp add: pochhammer_prod) lemma norm_fact [simp]: "norm (fact n :: 'a::real_normed_algebra_1) = fact n" proof - have "(fact n :: 'a) = of_real (fact n)" by simp also have "norm \ = fact n" by (subst norm_of_real) simp finally show ?thesis . qed lemma root_test_convergence: fixes f :: "nat \ 'a::banach" assumes f: "(\n. root n (norm (f n))) \ x" \ \could be weakened to lim sup\ and "x < 1" shows "summable f" proof - have "0 \ x" by (rule LIMSEQ_le[OF tendsto_const f]) (auto intro!: exI[of _ 1]) from \x < 1\ obtain z where z: "x < z" "z < 1" by (metis dense) from f \x < z\ have "eventually (\n. root n (norm (f n)) < z) sequentially" by (rule order_tendstoD) then have "eventually (\n. norm (f n) \ z^n) sequentially" using eventually_ge_at_top proof eventually_elim fix n assume less: "root n (norm (f n)) < z" and n: "1 \ n" from power_strict_mono[OF less, of n] n show "norm (f n) \ z ^ n" by simp qed then show "summable f" unfolding eventually_sequentially using z \0 \ x\ by (auto intro!: summable_comparison_test[OF _ summable_geometric]) qed subsection \More facts about binomial coefficients\ text \ These facts could have been proven before, but having real numbers makes the proofs a lot easier. \ lemma central_binomial_odd: "odd n \ n choose (Suc (n div 2)) = n choose (n div 2)" proof - assume "odd n" hence "Suc (n div 2) \ n" by presburger hence "n choose (Suc (n div 2)) = n choose (n - Suc (n div 2))" by (rule binomial_symmetric) also from \odd n\ have "n - Suc (n div 2) = n div 2" by presburger finally show ?thesis . qed lemma binomial_less_binomial_Suc: assumes k: "k < n div 2" shows "n choose k < n choose (Suc k)" proof - from k have k': "k \ n" "Suc k \ n" by simp_all from k' have "real (n choose k) = fact n / (fact k * fact (n - k))" by (simp add: binomial_fact) also from k' have "n - k = Suc (n - Suc k)" by simp also from k' have "fact \ = (real n - real k) * fact (n - Suc k)" by (subst fact_Suc) (simp_all add: of_nat_diff) also from k have "fact k = fact (Suc k) / (real k + 1)" by (simp add: field_simps) also have "fact n / (fact (Suc k) / (real k + 1) * ((real n - real k) * fact (n - Suc k))) = (n choose (Suc k)) * ((real k + 1) / (real n - real k))" using k by (simp add: field_split_simps binomial_fact) also from assms have "(real k + 1) / (real n - real k) < 1" by simp finally show ?thesis using k by (simp add: mult_less_cancel_left) qed lemma binomial_strict_mono: assumes "k < k'" "2*k' \ n" shows "n choose k < n choose k'" proof - from assms have "k \ k' - 1" by simp thus ?thesis proof (induction rule: inc_induct) case base with assms binomial_less_binomial_Suc[of "k' - 1" n] show ?case by simp next case (step k) from step.prems step.hyps assms have "n choose k < n choose (Suc k)" by (intro binomial_less_binomial_Suc) simp_all also have "\ < n choose k'" by (rule step.IH) finally show ?case . qed qed lemma binomial_mono: assumes "k \ k'" "2*k' \ n" shows "n choose k \ n choose k'" using assms binomial_strict_mono[of k k' n] by (cases "k = k'") simp_all lemma binomial_strict_antimono: assumes "k < k'" "2 * k \ n" "k' \ n" shows "n choose k > n choose k'" proof - from assms have "n choose (n - k) > n choose (n - k')" by (intro binomial_strict_mono) (simp_all add: algebra_simps) with assms show ?thesis by (simp add: binomial_symmetric [symmetric]) qed lemma binomial_antimono: assumes "k \ k'" "k \ n div 2" "k' \ n" shows "n choose k \ n choose k'" proof (cases "k = k'") case False note not_eq = False show ?thesis proof (cases "k = n div 2 \ odd n") case False with assms(2) have "2*k \ n" by presburger with not_eq assms binomial_strict_antimono[of k k' n] show ?thesis by simp next case True have "n choose k' \ n choose (Suc (n div 2))" proof (cases "k' = Suc (n div 2)") case False with assms True not_eq have "Suc (n div 2) < k'" by simp with assms binomial_strict_antimono[of "Suc (n div 2)" k' n] True show ?thesis by auto qed simp_all also from True have "\ = n choose k" by (simp add: central_binomial_odd) finally show ?thesis . qed qed simp_all lemma binomial_maximum: "n choose k \ n choose (n div 2)" proof - have "k \ n div 2 \ 2*k \ n" by linarith consider "2*k \ n" | "2*k \ n" "k \ n" | "k > n" by linarith thus ?thesis proof cases case 1 thus ?thesis by (intro binomial_mono) linarith+ next case 2 thus ?thesis by (intro binomial_antimono) simp_all qed (simp_all add: binomial_eq_0) qed lemma binomial_maximum': "(2*n) choose k \ (2*n) choose n" using binomial_maximum[of "2*n"] by simp lemma central_binomial_lower_bound: assumes "n > 0" shows "4^n / (2*real n) \ real ((2*n) choose n)" proof - from binomial[of 1 1 "2*n"] have "4 ^ n = (\k\2*n. (2*n) choose k)" by (simp add: power_mult power2_eq_square One_nat_def [symmetric] del: One_nat_def) also have "{..2*n} = {0<..<2*n} \ {0,2*n}" by auto also have "(\k\\. (2*n) choose k) = (\k\{0<..<2*n}. (2*n) choose k) + (\k\{0,2*n}. (2*n) choose k)" by (subst sum.union_disjoint) auto also have "(\k\{0,2*n}. (2*n) choose k) \ (\k\1. (n choose k)\<^sup>2)" by (cases n) simp_all also from assms have "\ \ (\k\n. (n choose k)\<^sup>2)" by (intro sum_mono2) auto also have "\ = (2*n) choose n" by (rule choose_square_sum) also have "(\k\{0<..<2*n}. (2*n) choose k) \ (\k\{0<..<2*n}. (2*n) choose n)" by (intro sum_mono binomial_maximum') also have "\ = card {0<..<2*n} * ((2*n) choose n)" by simp also have "card {0<..<2*n} \ 2*n - 1" by (cases n) simp_all also have "(2 * n - 1) * (2 * n choose n) + (2 * n choose n) = ((2*n) choose n) * (2*n)" using assms by (simp add: algebra_simps) finally have "4 ^ n \ (2 * n choose n) * (2 * n)" by simp_all hence "real (4 ^ n) \ real ((2 * n choose n) * (2 * n))" by (subst of_nat_le_iff) with assms show ?thesis by (simp add: field_simps) qed subsection \Properties of Power Series\ lemma powser_zero [simp]: "(\n. f n * 0 ^ n) = f 0" for f :: "nat \ 'a::real_normed_algebra_1" proof - have "(\n<1. f n * 0 ^ n) = (\n. f n * 0 ^ n)" by (subst suminf_finite[where N="{0}"]) (auto simp: power_0_left) then show ?thesis by simp qed lemma powser_sums_zero: "(\n. a n * 0^n) sums a 0" for a :: "nat \ 'a::real_normed_div_algebra" using sums_finite [of "{0}" "\n. a n * 0 ^ n"] by simp lemma powser_sums_zero_iff [simp]: "(\n. a n * 0^n) sums x \ a 0 = x" for a :: "nat \ 'a::real_normed_div_algebra" using powser_sums_zero sums_unique2 by blast text \ Power series has a circle or radius of convergence: if it sums for \x\, then it sums absolutely for \z\ with \<^term>\\z\ < \x\\.\ lemma powser_insidea: fixes x z :: "'a::real_normed_div_algebra" assumes 1: "summable (\n. f n * x^n)" and 2: "norm z < norm x" shows "summable (\n. norm (f n * z ^ n))" proof - from 2 have x_neq_0: "x \ 0" by clarsimp from 1 have "(\n. f n * x^n) \ 0" by (rule summable_LIMSEQ_zero) then have "convergent (\n. f n * x^n)" by (rule convergentI) then have "Cauchy (\n. f n * x^n)" by (rule convergent_Cauchy) then have "Bseq (\n. f n * x^n)" by (rule Cauchy_Bseq) then obtain K where 3: "0 < K" and 4: "\n. norm (f n * x^n) \ K" by (auto simp: Bseq_def) have "\N. \n\N. norm (norm (f n * z ^ n)) \ K * norm (z ^ n) * inverse (norm (x^n))" proof (intro exI allI impI) fix n :: nat assume "0 \ n" have "norm (norm (f n * z ^ n)) * norm (x^n) = norm (f n * x^n) * norm (z ^ n)" by (simp add: norm_mult abs_mult) also have "\ \ K * norm (z ^ n)" by (simp only: mult_right_mono 4 norm_ge_zero) also have "\ = K * norm (z ^ n) * (inverse (norm (x^n)) * norm (x^n))" by (simp add: x_neq_0) also have "\ = K * norm (z ^ n) * inverse (norm (x^n)) * norm (x^n)" by (simp only: mult.assoc) finally show "norm (norm (f n * z ^ n)) \ K * norm (z ^ n) * inverse (norm (x^n))" by (simp add: mult_le_cancel_right x_neq_0) qed moreover have "summable (\n. K * norm (z ^ n) * inverse (norm (x^n)))" proof - from 2 have "norm (norm (z * inverse x)) < 1" using x_neq_0 by (simp add: norm_mult nonzero_norm_inverse divide_inverse [where 'a=real, symmetric]) then have "summable (\n. norm (z * inverse x) ^ n)" by (rule summable_geometric) then have "summable (\n. K * norm (z * inverse x) ^ n)" by (rule summable_mult) then show "summable (\n. K * norm (z ^ n) * inverse (norm (x^n)))" using x_neq_0 by (simp add: norm_mult nonzero_norm_inverse power_mult_distrib power_inverse norm_power mult.assoc) qed ultimately show "summable (\n. norm (f n * z ^ n))" by (rule summable_comparison_test) qed lemma powser_inside: fixes f :: "nat \ 'a::{real_normed_div_algebra,banach}" shows "summable (\n. f n * (x^n)) \ norm z < norm x \ summable (\n. f n * (z ^ n))" by (rule powser_insidea [THEN summable_norm_cancel]) lemma powser_times_n_limit_0: fixes x :: "'a::{real_normed_div_algebra,banach}" assumes "norm x < 1" shows "(\n. of_nat n * x ^ n) \ 0" proof - have "norm x / (1 - norm x) \ 0" using assms by (auto simp: field_split_simps) moreover obtain N where N: "norm x / (1 - norm x) < of_int N" using ex_le_of_int by (meson ex_less_of_int) ultimately have N0: "N>0" by auto then have *: "real_of_int (N + 1) * norm x / real_of_int N < 1" using N assms by (auto simp: field_simps) have **: "real_of_int N * (norm x * (real_of_nat (Suc n) * norm (x ^ n))) \ real_of_nat n * (norm x * ((1 + N) * norm (x ^ n)))" if "N \ int n" for n :: nat proof - from that have "real_of_int N * real_of_nat (Suc n) \ real_of_nat n * real_of_int (1 + N)" by (simp add: algebra_simps) then have "(real_of_int N * real_of_nat (Suc n)) * (norm x * norm (x ^ n)) \ (real_of_nat n * (1 + N)) * (norm x * norm (x ^ n))" using N0 mult_mono by fastforce then show ?thesis by (simp add: algebra_simps) qed show ?thesis using * by (rule summable_LIMSEQ_zero [OF summable_ratio_test, where N1="nat N"]) (simp add: N0 norm_mult field_simps ** del: of_nat_Suc of_int_add) qed corollary lim_n_over_pown: fixes x :: "'a::{real_normed_field,banach}" shows "1 < norm x \ ((\n. of_nat n / x^n) \ 0) sequentially" using powser_times_n_limit_0 [of "inverse x"] by (simp add: norm_divide field_split_simps) lemma sum_split_even_odd: fixes f :: "nat \ real" shows "(\i<2 * n. if even i then f i else g i) = (\iii<2 * Suc n. if even i then f i else g i) = (\ii = (\ii real" assumes "g sums x" shows "(\ n. if even n then 0 else g ((n - 1) div 2)) sums x" unfolding sums_def proof (rule LIMSEQ_I) fix r :: real assume "0 < r" from \g sums x\[unfolded sums_def, THEN LIMSEQ_D, OF this] obtain no where no_eq: "\n. n \ no \ (norm (sum g {.. 2 * no" for m proof - from that have "m div 2 \ no" by auto have sum_eq: "?SUM (2 * (m div 2)) = sum g {..< m div 2}" using sum_split_even_odd by auto then have "(norm (?SUM (2 * (m div 2)) - x) < r)" using no_eq unfolding sum_eq using \m div 2 \ no\ by auto moreover have "?SUM (2 * (m div 2)) = ?SUM m" proof (cases "even m") case True then show ?thesis by (auto simp: even_two_times_div_two) next case False then have eq: "Suc (2 * (m div 2)) = m" by simp then have "even (2 * (m div 2))" using \odd m\ by auto have "?SUM m = ?SUM (Suc (2 * (m div 2)))" unfolding eq .. also have "\ = ?SUM (2 * (m div 2))" using \even (2 * (m div 2))\ by auto finally show ?thesis by auto qed ultimately show ?thesis by auto qed then show "\no. \ m \ no. norm (?SUM m - x) < r" by blast qed lemma sums_if: fixes g :: "nat \ real" assumes "g sums x" and "f sums y" shows "(\ n. if even n then f (n div 2) else g ((n - 1) div 2)) sums (x + y)" proof - let ?s = "\ n. if even n then 0 else f ((n - 1) div 2)" have if_sum: "(if B then (0 :: real) else E) + (if B then T else 0) = (if B then T else E)" for B T E by (cases B) auto have g_sums: "(\ n. if even n then 0 else g ((n - 1) div 2)) sums x" using sums_if'[OF \g sums x\] . have if_eq: "\B T E. (if \ B then T else E) = (if B then E else T)" by auto have "?s sums y" using sums_if'[OF \f sums y\] . from this[unfolded sums_def, THEN LIMSEQ_Suc] have "(\n. if even n then f (n div 2) else 0) sums y" by (simp add: lessThan_Suc_eq_insert_0 sum.atLeast1_atMost_eq image_Suc_lessThan if_eq sums_def cong del: if_weak_cong) from sums_add[OF g_sums this] show ?thesis by (simp only: if_sum) qed subsection \Alternating series test / Leibniz formula\ (* FIXME: generalise these results from the reals via type classes? *) lemma sums_alternating_upper_lower: fixes a :: "nat \ real" assumes mono: "\n. a (Suc n) \ a n" and a_pos: "\n. 0 \ a n" and "a \ 0" shows "\l. ((\n. (\i<2*n. (- 1)^i*a i) \ l) \ (\ n. \i<2*n. (- 1)^i*a i) \ l) \ ((\n. l \ (\i<2*n + 1. (- 1)^i*a i)) \ (\ n. \i<2*n + 1. (- 1)^i*a i) \ l)" (is "\l. ((\n. ?f n \ l) \ _) \ ((\n. l \ ?g n) \ _)") proof (rule nested_sequence_unique) have fg_diff: "\n. ?f n - ?g n = - a (2 * n)" by auto show "\n. ?f n \ ?f (Suc n)" proof show "?f n \ ?f (Suc n)" for n using mono[of "2*n"] by auto qed show "\n. ?g (Suc n) \ ?g n" proof show "?g (Suc n) \ ?g n" for n using mono[of "Suc (2*n)"] by auto qed show "\n. ?f n \ ?g n" proof show "?f n \ ?g n" for n using fg_diff a_pos by auto qed show "(\n. ?f n - ?g n) \ 0" unfolding fg_diff proof (rule LIMSEQ_I) fix r :: real assume "0 < r" with \a \ 0\[THEN LIMSEQ_D] obtain N where "\ n. n \ N \ norm (a n - 0) < r" by auto then have "\n \ N. norm (- a (2 * n) - 0) < r" by auto then show "\N. \n \ N. norm (- a (2 * n) - 0) < r" by auto qed qed lemma summable_Leibniz': fixes a :: "nat \ real" assumes a_zero: "a \ 0" and a_pos: "\n. 0 \ a n" and a_monotone: "\n. a (Suc n) \ a n" shows summable: "summable (\ n. (-1)^n * a n)" and "\n. (\i<2*n. (-1)^i*a i) \ (\i. (-1)^i*a i)" and "(\n. \i<2*n. (-1)^i*a i) \ (\i. (-1)^i*a i)" and "\n. (\i. (-1)^i*a i) \ (\i<2*n+1. (-1)^i*a i)" and "(\n. \i<2*n+1. (-1)^i*a i) \ (\i. (-1)^i*a i)" proof - let ?S = "\n. (-1)^n * a n" let ?P = "\n. \i n. ?f n \ l" and "?f \ l" and above_l: "\ n. l \ ?g n" and "?g \ l" using sums_alternating_upper_lower[OF a_monotone a_pos a_zero] by blast let ?Sa = "\m. \n l" proof (rule LIMSEQ_I) fix r :: real assume "0 < r" with \?f \ l\[THEN LIMSEQ_D] obtain f_no where f: "\n. n \ f_no \ norm (?f n - l) < r" by auto from \0 < r\ \?g \ l\[THEN LIMSEQ_D] obtain g_no where g: "\n. n \ g_no \ norm (?g n - l) < r" by auto have "norm (?Sa n - l) < r" if "n \ (max (2 * f_no) (2 * g_no))" for n proof - from that have "n \ 2 * f_no" and "n \ 2 * g_no" by auto show ?thesis proof (cases "even n") case True then have n_eq: "2 * (n div 2) = n" by (simp add: even_two_times_div_two) with \n \ 2 * f_no\ have "n div 2 \ f_no" by auto from f[OF this] show ?thesis unfolding n_eq atLeastLessThanSuc_atLeastAtMost . next case False then have "even (n - 1)" by simp then have n_eq: "2 * ((n - 1) div 2) = n - 1" by (simp add: even_two_times_div_two) then have range_eq: "n - 1 + 1 = n" using odd_pos[OF False] by auto from n_eq \n \ 2 * g_no\ have "(n - 1) div 2 \ g_no" by auto from g[OF this] show ?thesis by (simp only: n_eq range_eq) qed qed then show "\no. \n \ no. norm (?Sa n - l) < r" by blast qed then have sums_l: "(\i. (-1)^i * a i) sums l" by (simp only: sums_def) then show "summable ?S" by (auto simp: summable_def) have "l = suminf ?S" by (rule sums_unique[OF sums_l]) fix n show "suminf ?S \ ?g n" unfolding sums_unique[OF sums_l, symmetric] using above_l by auto show "?f n \ suminf ?S" unfolding sums_unique[OF sums_l, symmetric] using below_l by auto show "?g \ suminf ?S" using \?g \ l\ \l = suminf ?S\ by auto show "?f \ suminf ?S" using \?f \ l\ \l = suminf ?S\ by auto qed theorem summable_Leibniz: fixes a :: "nat \ real" assumes a_zero: "a \ 0" and "monoseq a" shows "summable (\ n. (-1)^n * a n)" (is "?summable") and "0 < a 0 \ (\n. (\i. (- 1)^i*a i) \ { \i<2*n. (- 1)^i * a i .. \i<2*n+1. (- 1)^i * a i})" (is "?pos") and "a 0 < 0 \ (\n. (\i. (- 1)^i*a i) \ { \i<2*n+1. (- 1)^i * a i .. \i<2*n. (- 1)^i * a i})" (is "?neg") and "(\n. \i<2*n. (- 1)^i*a i) \ (\i. (- 1)^i*a i)" (is "?f") and "(\n. \i<2*n+1. (- 1)^i*a i) \ (\i. (- 1)^i*a i)" (is "?g") proof - have "?summable \ ?pos \ ?neg \ ?f \ ?g" proof (cases "(\n. 0 \ a n) \ (\m. \n\m. a n \ a m)") case True then have ord: "\n m. m \ n \ a n \ a m" and ge0: "\n. 0 \ a n" by auto have mono: "a (Suc n) \ a n" for n using ord[where n="Suc n" and m=n] by auto note leibniz = summable_Leibniz'[OF \a \ 0\ ge0] from leibniz[OF mono] show ?thesis using \0 \ a 0\ by auto next let ?a = "\n. - a n" case False with monoseq_le[OF \monoseq a\ \a \ 0\] have "(\ n. a n \ 0) \ (\m. \n\m. a m \ a n)" by auto then have ord: "\n m. m \ n \ ?a n \ ?a m" and ge0: "\ n. 0 \ ?a n" by auto have monotone: "?a (Suc n) \ ?a n" for n using ord[where n="Suc n" and m=n] by auto note leibniz = summable_Leibniz'[OF _ ge0, of "\x. x", OF tendsto_minus[OF \a \ 0\, unfolded minus_zero] monotone] have "summable (\ n. (-1)^n * ?a n)" using leibniz(1) by auto then obtain l where "(\ n. (-1)^n * ?a n) sums l" unfolding summable_def by auto from this[THEN sums_minus] have "(\ n. (-1)^n * a n) sums -l" by auto then have ?summable by (auto simp: summable_def) moreover have "\- a - - b\ = \a - b\" for a b :: real unfolding minus_diff_minus by auto from suminf_minus[OF leibniz(1), unfolded mult_minus_right minus_minus] have move_minus: "(\n. - ((- 1) ^ n * a n)) = - (\n. (- 1) ^ n * a n)" by auto have ?pos using \0 \ ?a 0\ by auto moreover have ?neg using leibniz(2,4) unfolding mult_minus_right sum_negf move_minus neg_le_iff_le by auto moreover have ?f and ?g using leibniz(3,5)[unfolded mult_minus_right sum_negf move_minus, THEN tendsto_minus_cancel] by auto ultimately show ?thesis by auto qed then show ?summable and ?pos and ?neg and ?f and ?g by safe qed subsection \Term-by-Term Differentiability of Power Series\ definition diffs :: "(nat \ 'a::ring_1) \ nat \ 'a" where "diffs c = (\n. of_nat (Suc n) * c (Suc n))" text \Lemma about distributing negation over it.\ lemma diffs_minus: "diffs (\n. - c n) = (\n. - diffs c n)" by (simp add: diffs_def) lemma diffs_equiv: fixes x :: "'a::{real_normed_vector,ring_1}" shows "summable (\n. diffs c n * x^n) \ (\n. of_nat n * c n * x^(n - Suc 0)) sums (\n. diffs c n * x^n)" unfolding diffs_def by (simp add: summable_sums sums_Suc_imp) lemma lemma_termdiff1: fixes z :: "'a :: {monoid_mult,comm_ring}" shows "(\ppi 0" shows "((z + h) ^ n - z ^ n) / h - of_nat n * z ^ (n - Suc 0) = h * (\p< n - Suc 0. \q< n - Suc 0 - p. (z + h) ^ q * z ^ (n - 2 - q))" (is "?lhs = ?rhs") proof (cases n) case (Suc m) have 0: "\x k. (\njiijppip::nat. p < n \ f p \ K" and K: "0 \ K" shows "sum f {.. of_nat n * K" proof - have "sum f {.. (\i of_nat n * K" by (auto simp: mult_right_mono K) finally show ?thesis . qed lemma lemma_termdiff3: fixes h z :: "'a::real_normed_field" assumes 1: "h \ 0" and 2: "norm z \ K" and 3: "norm (z + h) \ K" shows "norm (((z + h) ^ n - z ^ n) / h - of_nat n * z ^ (n - Suc 0)) \ of_nat n * of_nat (n - Suc 0) * K ^ (n - 2) * norm h" proof - have "norm (((z + h) ^ n - z ^ n) / h - of_nat n * z ^ (n - Suc 0)) = norm (\pq \ of_nat n * (of_nat (n - Suc 0) * K ^ (n - 2)) * norm h" proof (rule mult_right_mono [OF _ norm_ge_zero]) from norm_ge_zero 2 have K: "0 \ K" by (rule order_trans) have le_Kn: "norm ((z + h) ^ i * z ^ j) \ K ^ n" if "i + j = n" for i j n proof - have "norm (z + h) ^ i * norm z ^ j \ K ^ i * K ^ j" by (intro mult_mono power_mono 2 3 norm_ge_zero zero_le_power K) also have "... = K^n" by (metis power_add that) finally show ?thesis by (simp add: norm_mult norm_power) qed then have "\p q. \p < n; q < n - Suc 0\ \ norm ((z + h) ^ q * z ^ (n - 2 - q)) \ K ^ (n - 2)" by (simp del: subst_all) then show "norm (\pq of_nat n * (of_nat (n - Suc 0) * K ^ (n - 2))" by (intro order_trans [OF norm_sum] real_sum_nat_ivl_bounded2 mult_nonneg_nonneg of_nat_0_le_iff zero_le_power K) qed also have "\ = of_nat n * of_nat (n - Suc 0) * K ^ (n - 2) * norm h" by (simp only: mult.assoc) finally show ?thesis . qed lemma lemma_termdiff4: fixes f :: "'a::real_normed_vector \ 'b::real_normed_vector" and k :: real assumes k: "0 < k" and le: "\h. h \ 0 \ norm h < k \ norm (f h) \ K * norm h" shows "f \0\ 0" proof (rule tendsto_norm_zero_cancel) show "(\h. norm (f h)) \0\ 0" proof (rule real_tendsto_sandwich) show "eventually (\h. 0 \ norm (f h)) (at 0)" by simp show "eventually (\h. norm (f h) \ K * norm h) (at 0)" using k by (auto simp: eventually_at dist_norm le) show "(\h. 0) \(0::'a)\ (0::real)" by (rule tendsto_const) have "(\h. K * norm h) \(0::'a)\ K * norm (0::'a)" by (intro tendsto_intros) then show "(\h. K * norm h) \(0::'a)\ 0" by simp qed qed lemma lemma_termdiff5: fixes g :: "'a::real_normed_vector \ nat \ 'b::banach" and k :: real assumes k: "0 < k" and f: "summable f" and le: "\h n. h \ 0 \ norm h < k \ norm (g h n) \ f n * norm h" shows "(\h. suminf (g h)) \0\ 0" proof (rule lemma_termdiff4 [OF k]) fix h :: 'a assume "h \ 0" and "norm h < k" then have 1: "\n. norm (g h n) \ f n * norm h" by (simp add: le) then have "\N. \n\N. norm (norm (g h n)) \ f n * norm h" by simp moreover from f have 2: "summable (\n. f n * norm h)" by (rule summable_mult2) ultimately have 3: "summable (\n. norm (g h n))" by (rule summable_comparison_test) then have "norm (suminf (g h)) \ (\n. norm (g h n))" by (rule summable_norm) also from 1 3 2 have "(\n. norm (g h n)) \ (\n. f n * norm h)" by (simp add: suminf_le) also from f have "(\n. f n * norm h) = suminf f * norm h" by (rule suminf_mult2 [symmetric]) finally show "norm (suminf (g h)) \ suminf f * norm h" . qed (* FIXME: Long proofs *) lemma termdiffs_aux: fixes x :: "'a::{real_normed_field,banach}" assumes 1: "summable (\n. diffs (diffs c) n * K ^ n)" and 2: "norm x < norm K" shows "(\h. \n. c n * (((x + h) ^ n - x^n) / h - of_nat n * x ^ (n - Suc 0))) \0\ 0" proof - from dense [OF 2] obtain r where r1: "norm x < r" and r2: "r < norm K" by fast from norm_ge_zero r1 have r: "0 < r" by (rule order_le_less_trans) then have r_neq_0: "r \ 0" by simp show ?thesis proof (rule lemma_termdiff5) show "0 < r - norm x" using r1 by simp from r r2 have "norm (of_real r::'a) < norm K" by simp with 1 have "summable (\n. norm (diffs (diffs c) n * (of_real r ^ n)))" by (rule powser_insidea) then have "summable (\n. diffs (diffs (\n. norm (c n))) n * r ^ n)" using r by (simp add: diffs_def norm_mult norm_power del: of_nat_Suc) then have "summable (\n. of_nat n * diffs (\n. norm (c n)) n * r ^ (n - Suc 0))" by (rule diffs_equiv [THEN sums_summable]) also have "(\n. of_nat n * diffs (\n. norm (c n)) n * r ^ (n - Suc 0)) = (\n. diffs (\m. of_nat (m - Suc 0) * norm (c m) * inverse r) n * (r ^ n))" by (simp add: diffs_def r_neq_0 fun_eq_iff split: nat_diff_split) finally have "summable (\n. of_nat n * (of_nat (n - Suc 0) * norm (c n) * inverse r) * r ^ (n - Suc 0))" by (rule diffs_equiv [THEN sums_summable]) also have "(\n. of_nat n * (of_nat (n - Suc 0) * norm (c n) * inverse r) * r ^ (n - Suc 0)) = (\n. norm (c n) * of_nat n * of_nat (n - Suc 0) * r ^ (n - 2))" by (rule ext) (simp add: r_neq_0 split: nat_diff_split) finally show "summable (\n. norm (c n) * of_nat n * of_nat (n - Suc 0) * r ^ (n - 2))" . next fix h :: 'a and n assume h: "h \ 0" assume "norm h < r - norm x" then have "norm x + norm h < r" by simp with norm_triangle_ineq have xh: "norm (x + h) < r" by (rule order_le_less_trans) have "norm (((x + h) ^ n - x ^ n) / h - of_nat n * x ^ (n - Suc 0)) \ real n * (real (n - Suc 0) * (r ^ (n - 2) * norm h))" by (metis (mono_tags, lifting) h mult.assoc lemma_termdiff3 less_eq_real_def r1 xh) then show "norm (c n * (((x + h) ^ n - x^n) / h - of_nat n * x ^ (n - Suc 0))) \ norm (c n) * of_nat n * of_nat (n - Suc 0) * r ^ (n - 2) * norm h" by (simp only: norm_mult mult.assoc mult_left_mono [OF _ norm_ge_zero]) qed qed lemma termdiffs: fixes K x :: "'a::{real_normed_field,banach}" assumes 1: "summable (\n. c n * K ^ n)" and 2: "summable (\n. (diffs c) n * K ^ n)" and 3: "summable (\n. (diffs (diffs c)) n * K ^ n)" and 4: "norm x < norm K" shows "DERIV (\x. \n. c n * x^n) x :> (\n. (diffs c) n * x^n)" unfolding DERIV_def proof (rule LIM_zero_cancel) show "(\h. (suminf (\n. c n * (x + h) ^ n) - suminf (\n. c n * x^n)) / h - suminf (\n. diffs c n * x^n)) \0\ 0" proof (rule LIM_equal2) show "0 < norm K - norm x" using 4 by (simp add: less_diff_eq) next fix h :: 'a assume "norm (h - 0) < norm K - norm x" then have "norm x + norm h < norm K" by simp then have 5: "norm (x + h) < norm K" by (rule norm_triangle_ineq [THEN order_le_less_trans]) have "summable (\n. c n * x^n)" and "summable (\n. c n * (x + h) ^ n)" and "summable (\n. diffs c n * x^n)" using 1 2 4 5 by (auto elim: powser_inside) then have "((\n. c n * (x + h) ^ n) - (\n. c n * x^n)) / h - (\n. diffs c n * x^n) = (\n. (c n * (x + h) ^ n - c n * x^n) / h - of_nat n * c n * x ^ (n - Suc 0))" by (intro sums_unique sums_diff sums_divide diffs_equiv summable_sums) then show "((\n. c n * (x + h) ^ n) - (\n. c n * x^n)) / h - (\n. diffs c n * x^n) = (\n. c n * (((x + h) ^ n - x^n) / h - of_nat n * x ^ (n - Suc 0)))" by (simp add: algebra_simps) next show "(\h. \n. c n * (((x + h) ^ n - x^n) / h - of_nat n * x ^ (n - Suc 0))) \0\ 0" by (rule termdiffs_aux [OF 3 4]) qed qed subsection \The Derivative of a Power Series Has the Same Radius of Convergence\ lemma termdiff_converges: fixes x :: "'a::{real_normed_field,banach}" assumes K: "norm x < K" and sm: "\x. norm x < K \ summable(\n. c n * x ^ n)" shows "summable (\n. diffs c n * x ^ n)" proof (cases "x = 0") case True then show ?thesis using powser_sums_zero sums_summable by auto next case False then have "K > 0" using K less_trans zero_less_norm_iff by blast then obtain r :: real where r: "norm x < norm r" "norm r < K" "r > 0" using K False by (auto simp: field_simps abs_less_iff add_pos_pos intro: that [of "(norm x + K) / 2"]) have to0: "(\n. of_nat n * (x / of_real r) ^ n) \ 0" using r by (simp add: norm_divide powser_times_n_limit_0 [of "x / of_real r"]) obtain N where N: "\n. n\N \ real_of_nat n * norm x ^ n < r ^ n" using r LIMSEQ_D [OF to0, of 1] by (auto simp: norm_divide norm_mult norm_power field_simps) have "summable (\n. (of_nat n * c n) * x ^ n)" proof (rule summable_comparison_test') show "summable (\n. norm (c n * of_real r ^ n))" apply (rule powser_insidea [OF sm [of "of_real ((r+K)/2)"]]) using N r norm_of_real [of "r + K", where 'a = 'a] by auto show "\n. N \ n \ norm (of_nat n * c n * x ^ n) \ norm (c n * of_real r ^ n)" using N r by (fastforce simp add: norm_mult norm_power less_eq_real_def) qed then have "summable (\n. (of_nat (Suc n) * c(Suc n)) * x ^ Suc n)" using summable_iff_shift [of "\n. of_nat n * c n * x ^ n" 1] by simp then have "summable (\n. (of_nat (Suc n) * c(Suc n)) * x ^ n)" using False summable_mult2 [of "\n. (of_nat (Suc n) * c(Suc n) * x ^ n) * x" "inverse x"] by (simp add: mult.assoc) (auto simp: ac_simps) then show ?thesis by (simp add: diffs_def) qed lemma termdiff_converges_all: fixes x :: "'a::{real_normed_field,banach}" assumes "\x. summable (\n. c n * x^n)" shows "summable (\n. diffs c n * x^n)" by (rule termdiff_converges [where K = "1 + norm x"]) (use assms in auto) lemma termdiffs_strong: fixes K x :: "'a::{real_normed_field,banach}" assumes sm: "summable (\n. c n * K ^ n)" and K: "norm x < norm K" shows "DERIV (\x. \n. c n * x^n) x :> (\n. diffs c n * x^n)" proof - have "norm K + norm x < norm K + norm K" using K by force then have K2: "norm ((of_real (norm K) + of_real (norm x)) / 2 :: 'a) < norm K" by (auto simp: norm_triangle_lt norm_divide field_simps) then have [simp]: "norm ((of_real (norm K) + of_real (norm x)) :: 'a) < norm K * 2" by simp have "summable (\n. c n * (of_real (norm x + norm K) / 2) ^ n)" by (metis K2 summable_norm_cancel [OF powser_insidea [OF sm]] add.commute of_real_add) moreover have "\x. norm x < norm K \ summable (\n. diffs c n * x ^ n)" by (blast intro: sm termdiff_converges powser_inside) moreover have "\x. norm x < norm K \ summable (\n. diffs(diffs c) n * x ^ n)" by (blast intro: sm termdiff_converges powser_inside) ultimately show ?thesis by (rule termdiffs [where K = "of_real (norm x + norm K) / 2"]) (use K in \auto simp: field_simps simp flip: of_real_add\) qed lemma termdiffs_strong_converges_everywhere: fixes K x :: "'a::{real_normed_field,banach}" assumes "\y. summable (\n. c n * y ^ n)" shows "((\x. \n. c n * x^n) has_field_derivative (\n. diffs c n * x^n)) (at x)" using termdiffs_strong[OF assms[of "of_real (norm x + 1)"], of x] by (force simp del: of_real_add) lemma termdiffs_strong': fixes z :: "'a :: {real_normed_field,banach}" assumes "\z. norm z < K \ summable (\n. c n * z ^ n)" assumes "norm z < K" shows "((\z. \n. c n * z^n) has_field_derivative (\n. diffs c n * z^n)) (at z)" proof (rule termdiffs_strong) define L :: real where "L = (norm z + K) / 2" have "0 \ norm z" by simp also note \norm z < K\ finally have K: "K \ 0" by simp from assms K have L: "L \ 0" "norm z < L" "L < K" by (simp_all add: L_def) from L show "norm z < norm (of_real L :: 'a)" by simp from L show "summable (\n. c n * of_real L ^ n)" by (intro assms(1)) simp_all qed lemma termdiffs_sums_strong: fixes z :: "'a :: {banach,real_normed_field}" assumes sums: "\z. norm z < K \ (\n. c n * z ^ n) sums f z" assumes deriv: "(f has_field_derivative f') (at z)" assumes norm: "norm z < K" shows "(\n. diffs c n * z ^ n) sums f'" proof - have summable: "summable (\n. diffs c n * z^n)" by (intro termdiff_converges[OF norm] sums_summable[OF sums]) from norm have "eventually (\z. z \ norm -` {..z. (\n. c n * z^n) = f z) (nhds z)" by eventually_elim (insert sums, simp add: sums_iff) have "((\z. \n. c n * z^n) has_field_derivative (\n. diffs c n * z^n)) (at z)" by (intro termdiffs_strong'[OF _ norm] sums_summable[OF sums]) hence "(f has_field_derivative (\n. diffs c n * z^n)) (at z)" by (subst (asm) DERIV_cong_ev[OF refl eq refl]) from this and deriv have "(\n. diffs c n * z^n) = f'" by (rule DERIV_unique) with summable show ?thesis by (simp add: sums_iff) qed lemma isCont_powser: fixes K x :: "'a::{real_normed_field,banach}" assumes "summable (\n. c n * K ^ n)" assumes "norm x < norm K" shows "isCont (\x. \n. c n * x^n) x" using termdiffs_strong[OF assms] by (blast intro!: DERIV_isCont) lemmas isCont_powser' = isCont_o2[OF _ isCont_powser] lemma isCont_powser_converges_everywhere: fixes K x :: "'a::{real_normed_field,banach}" assumes "\y. summable (\n. c n * y ^ n)" shows "isCont (\x. \n. c n * x^n) x" using termdiffs_strong[OF assms[of "of_real (norm x + 1)"], of x] by (force intro!: DERIV_isCont simp del: of_real_add) lemma powser_limit_0: fixes a :: "nat \ 'a::{real_normed_field,banach}" assumes s: "0 < s" and sm: "\x. norm x < s \ (\n. a n * x ^ n) sums (f x)" shows "(f \ a 0) (at 0)" proof - have "norm (of_real s / 2 :: 'a) < s" using s by (auto simp: norm_divide) then have "summable (\n. a n * (of_real s / 2) ^ n)" by (rule sums_summable [OF sm]) then have "((\x. \n. a n * x ^ n) has_field_derivative (\n. diffs a n * 0 ^ n)) (at 0)" by (rule termdiffs_strong) (use s in \auto simp: norm_divide\) then have "isCont (\x. \n. a n * x ^ n) 0" by (blast intro: DERIV_continuous) then have "((\x. \n. a n * x ^ n) \ a 0) (at 0)" by (simp add: continuous_within) moreover have "(\x. f x - (\n. a n * x ^ n)) \0\ 0" apply (clarsimp simp: LIM_eq) apply (rule_tac x=s in exI) using s sm sums_unique by fastforce ultimately show ?thesis by (rule Lim_transform) qed lemma powser_limit_0_strong: fixes a :: "nat \ 'a::{real_normed_field,banach}" assumes s: "0 < s" and sm: "\x. x \ 0 \ norm x < s \ (\n. a n * x ^ n) sums (f x)" shows "(f \ a 0) (at 0)" proof - have *: "((\x. if x = 0 then a 0 else f x) \ a 0) (at 0)" by (rule powser_limit_0 [OF s]) (auto simp: powser_sums_zero sm) show ?thesis using "*" by (auto cong: Lim_cong_within) qed subsection \Derivability of power series\ lemma DERIV_series': fixes f :: "real \ nat \ real" assumes DERIV_f: "\ n. DERIV (\ x. f x n) x0 :> (f' x0 n)" and allf_summable: "\ x. x \ {a <..< b} \ summable (f x)" and x0_in_I: "x0 \ {a <..< b}" and "summable (f' x0)" and "summable L" and L_def: "\n x y. x \ {a <..< b} \ y \ {a <..< b} \ \f x n - f y n\ \ L n * \x - y\" shows "DERIV (\ x. suminf (f x)) x0 :> (suminf (f' x0))" unfolding DERIV_def proof (rule LIM_I) fix r :: real assume "0 < r" then have "0 < r/3" by auto obtain N_L where N_L: "\ n. N_L \ n \ \ \ i. L (i + n) \ < r/3" using suminf_exist_split[OF \0 < r/3\ \summable L\] by auto obtain N_f' where N_f': "\ n. N_f' \ n \ \ \ i. f' x0 (i + n) \ < r/3" using suminf_exist_split[OF \0 < r/3\ \summable (f' x0)\] by auto let ?N = "Suc (max N_L N_f')" have "\ \ i. f' x0 (i + ?N) \ < r/3" (is "?f'_part < r/3") and L_estimate: "\ \ i. L (i + ?N) \ < r/3" using N_L[of "?N"] and N_f' [of "?N"] by auto let ?diff = "\i x. (f (x0 + x) i - f x0 i) / x" let ?r = "r / (3 * real ?N)" from \0 < r\ have "0 < ?r" by simp let ?s = "\n. SOME s. 0 < s \ (\ x. x \ 0 \ \ x \ < s \ \ ?diff n x - f' x0 n \ < ?r)" define S' where "S' = Min (?s ` {..< ?N })" have "0 < S'" unfolding S'_def proof (rule iffD2[OF Min_gr_iff]) show "\x \ (?s ` {..< ?N }). 0 < x" proof fix x assume "x \ ?s ` {.. {..0 < ?r\, unfolded real_norm_def] obtain s where s_bound: "0 < s \ (\x. x \ 0 \ \x\ < s \ \?diff n x - f' x0 n\ < ?r)" by auto have "0 < ?s n" by (rule someI2[where a=s]) (auto simp: s_bound simp del: of_nat_Suc) then show "0 < x" by (simp only: \x = ?s n\) qed qed auto define S where "S = min (min (x0 - a) (b - x0)) S'" then have "0 < S" and S_a: "S \ x0 - a" and S_b: "S \ b - x0" and "S \ S'" using x0_in_I and \0 < S'\ by auto have "\(suminf (f (x0 + x)) - suminf (f x0)) / x - suminf (f' x0)\ < r" if "x \ 0" and "\x\ < S" for x proof - from that have x_in_I: "x0 + x \ {a <..< b}" using S_a S_b by auto note diff_smbl = summable_diff[OF allf_summable[OF x_in_I] allf_summable[OF x0_in_I]] note div_smbl = summable_divide[OF diff_smbl] note all_smbl = summable_diff[OF div_smbl \summable (f' x0)\] note ign = summable_ignore_initial_segment[where k="?N"] note diff_shft_smbl = summable_diff[OF ign[OF allf_summable[OF x_in_I]] ign[OF allf_summable[OF x0_in_I]]] note div_shft_smbl = summable_divide[OF diff_shft_smbl] note all_shft_smbl = summable_diff[OF div_smbl ign[OF \summable (f' x0)\]] have 1: "\(\?diff (n + ?N) x\)\ \ L (n + ?N)" for n proof - have "\?diff (n + ?N) x\ \ L (n + ?N) * \(x0 + x) - x0\ / \x\" using divide_right_mono[OF L_def[OF x_in_I x0_in_I] abs_ge_zero] by (simp only: abs_divide) with \x \ 0\ show ?thesis by auto qed note 2 = summable_rabs_comparison_test[OF _ ign[OF \summable L\]] from 1 have "\ \ i. ?diff (i + ?N) x \ \ (\ i. L (i + ?N))" by (metis (lifting) abs_idempotent order_trans[OF summable_rabs[OF 2] suminf_le[OF _ 2 ign[OF \summable L\]]]) then have "\\i. ?diff (i + ?N) x\ \ r / 3" (is "?L_part \ r/3") using L_estimate by auto have "\\n \ (\n?diff n x - f' x0 n\)" .. also have "\ < (\n {..< ?N}" have "\x\ < S" using \\x\ < S\ . also have "S \ S'" using \S \ S'\ . also have "S' \ ?s n" unfolding S'_def proof (rule Min_le_iff[THEN iffD2]) have "?s n \ (?s ` {.. ?s n \ ?s n" using \n \ {..< ?N}\ by auto then show "\ a \ (?s ` {.. ?s n" by blast qed auto finally have "\x\ < ?s n" . from DERIV_D[OF DERIV_f[where n=n], THEN LIM_D, OF \0 < ?r\, unfolded real_norm_def diff_0_right, unfolded some_eq_ex[symmetric], THEN conjunct2] have "\x. x \ 0 \ \x\ < ?s n \ \?diff n x - f' x0 n\ < ?r" . with \x \ 0\ and \\x\ < ?s n\ show "\?diff n x - f' x0 n\ < ?r" by blast qed auto also have "\ = of_nat (card {.. = real ?N * ?r" by simp also have "\ = r/3" by (auto simp del: of_nat_Suc) finally have "\\n < r / 3" (is "?diff_part < r / 3") . from suminf_diff[OF allf_summable[OF x_in_I] allf_summable[OF x0_in_I]] have "\(suminf (f (x0 + x)) - (suminf (f x0))) / x - suminf (f' x0)\ = \\n. ?diff n x - f' x0 n\" unfolding suminf_diff[OF div_smbl \summable (f' x0)\, symmetric] using suminf_divide[OF diff_smbl, symmetric] by auto also have "\ \ ?diff_part + \(\n. ?diff (n + ?N) x) - (\ n. f' x0 (n + ?N))\" unfolding suminf_split_initial_segment[OF all_smbl, where k="?N"] unfolding suminf_diff[OF div_shft_smbl ign[OF \summable (f' x0)\]] apply (simp only: add.commute) using abs_triangle_ineq by blast also have "\ \ ?diff_part + ?L_part + ?f'_part" using abs_triangle_ineq4 by auto also have "\ < r /3 + r/3 + r/3" using \?diff_part < r/3\ \?L_part \ r/3\ and \?f'_part < r/3\ by (rule add_strict_mono [OF add_less_le_mono]) finally show ?thesis by auto qed then show "\s > 0. \ x. x \ 0 \ norm (x - 0) < s \ norm (((\n. f (x0 + x) n) - (\n. f x0 n)) / x - (\n. f' x0 n)) < r" using \0 < S\ by auto qed lemma DERIV_power_series': fixes f :: "nat \ real" assumes converges: "\x. x \ {-R <..< R} \ summable (\n. f n * real (Suc n) * x^n)" and x0_in_I: "x0 \ {-R <..< R}" and "0 < R" shows "DERIV (\x. (\n. f n * x^(Suc n))) x0 :> (\n. f n * real (Suc n) * x0^n)" (is "DERIV (\x. suminf (?f x)) x0 :> suminf (?f' x0)") proof - have for_subinterval: "DERIV (\x. suminf (?f x)) x0 :> suminf (?f' x0)" if "0 < R'" and "R' < R" and "-R' < x0" and "x0 < R'" for R' proof - from that have "x0 \ {-R' <..< R'}" and "R' \ {-R <..< R}" and "x0 \ {-R <..< R}" by auto show ?thesis proof (rule DERIV_series') show "summable (\ n. \f n * real (Suc n) * R'^n\)" proof - have "(R' + R) / 2 < R" and "0 < (R' + R) / 2" using \0 < R'\ \0 < R\ \R' < R\ by (auto simp: field_simps) then have in_Rball: "(R' + R) / 2 \ {-R <..< R}" using \R' < R\ by auto have "norm R' < norm ((R' + R) / 2)" using \0 < R'\ \0 < R\ \R' < R\ by (auto simp: field_simps) from powser_insidea[OF converges[OF in_Rball] this] show ?thesis by auto qed next fix n x y assume "x \ {-R' <..< R'}" and "y \ {-R' <..< R'}" show "\?f x n - ?f y n\ \ \f n * real (Suc n) * R'^n\ * \x-y\" proof - have "\f n * x ^ (Suc n) - f n * y ^ (Suc n)\ = (\f n\ * \x-y\) * \\p" unfolding right_diff_distrib[symmetric] diff_power_eq_sum abs_mult by auto also have "\ \ (\f n\ * \x-y\) * (\real (Suc n)\ * \R' ^ n\)" proof (rule mult_left_mono) have "\\p \ (\px ^ p * y ^ (n - p)\)" by (rule sum_abs) also have "\ \ (\p {.. n" by auto have "\x^n\ \ R'^n" if "x \ {-R'<..x\ \ R'" by auto then show ?thesis unfolding power_abs by (rule power_mono) auto qed from mult_mono[OF this[OF \x \ {-R'<.., of p] this[OF \y \ {-R'<.., of "n-p"]] and \0 < R'\ have "\x^p * y^(n - p)\ \ R'^p * R'^(n - p)" unfolding abs_mult by auto then show "\x^p * y^(n - p)\ \ R'^n" unfolding power_add[symmetric] using \p \ n\ by auto qed also have "\ = real (Suc n) * R' ^ n" unfolding sum_constant card_atLeastLessThan by auto finally show "\\p \ \real (Suc n)\ * \R' ^ n\" unfolding abs_of_nonneg[OF zero_le_power[OF less_imp_le[OF \0 < R'\]]] by linarith show "0 \ \f n\ * \x - y\" unfolding abs_mult[symmetric] by auto qed also have "\ = \f n * real (Suc n) * R' ^ n\ * \x - y\" unfolding abs_mult mult.assoc[symmetric] by algebra finally show ?thesis . qed next show "DERIV (\x. ?f x n) x0 :> ?f' x0 n" for n by (auto intro!: derivative_eq_intros simp del: power_Suc) next fix x assume "x \ {-R' <..< R'}" then have "R' \ {-R <..< R}" and "norm x < norm R'" using assms \R' < R\ by auto have "summable (\n. f n * x^n)" proof (rule summable_comparison_test, intro exI allI impI) fix n have le: "\f n\ * 1 \ \f n\ * real (Suc n)" by (rule mult_left_mono) auto show "norm (f n * x^n) \ norm (f n * real (Suc n) * x^n)" unfolding real_norm_def abs_mult using le mult_right_mono by fastforce qed (rule powser_insidea[OF converges[OF \R' \ {-R <..< R}\] \norm x < norm R'\]) from this[THEN summable_mult2[where c=x], simplified mult.assoc, simplified mult.commute] show "summable (?f x)" by auto next show "summable (?f' x0)" using converges[OF \x0 \ {-R <..< R}\] . show "x0 \ {-R' <..< R'}" using \x0 \ {-R' <..< R'}\ . qed qed let ?R = "(R + \x0\) / 2" have "\x0\ < ?R" using assms by (auto simp: field_simps) then have "- ?R < x0" proof (cases "x0 < 0") case True then have "- x0 < ?R" using \\x0\ < ?R\ by auto then show ?thesis unfolding neg_less_iff_less[symmetric, of "- x0"] by auto next case False have "- ?R < 0" using assms by auto also have "\ \ x0" using False by auto finally show ?thesis . qed then have "0 < ?R" "?R < R" "- ?R < x0" and "x0 < ?R" using assms by (auto simp: field_simps) from for_subinterval[OF this] show ?thesis . qed lemma geometric_deriv_sums: fixes z :: "'a :: {real_normed_field,banach}" assumes "norm z < 1" shows "(\n. of_nat (Suc n) * z ^ n) sums (1 / (1 - z)^2)" proof - have "(\n. diffs (\n. 1) n * z^n) sums (1 / (1 - z)^2)" proof (rule termdiffs_sums_strong) fix z :: 'a assume "norm z < 1" thus "(\n. 1 * z^n) sums (1 / (1 - z))" by (simp add: geometric_sums) qed (insert assms, auto intro!: derivative_eq_intros simp: power2_eq_square) thus ?thesis unfolding diffs_def by simp qed lemma isCont_pochhammer [continuous_intros]: "isCont (\z. pochhammer z n) z" for z :: "'a::real_normed_field" by (induct n) (auto simp: pochhammer_rec') lemma continuous_on_pochhammer [continuous_intros]: "continuous_on A (\z. pochhammer z n)" for A :: "'a::real_normed_field set" by (intro continuous_at_imp_continuous_on ballI isCont_pochhammer) lemmas continuous_on_pochhammer' [continuous_intros] = continuous_on_compose2[OF continuous_on_pochhammer _ subset_UNIV] subsection \Exponential Function\ definition exp :: "'a \ 'a::{real_normed_algebra_1,banach}" where "exp = (\x. \n. x^n /\<^sub>R fact n)" lemma summable_exp_generic: fixes x :: "'a::{real_normed_algebra_1,banach}" defines S_def: "S \ \n. x^n /\<^sub>R fact n" shows "summable S" proof - have S_Suc: "\n. S (Suc n) = (x * S n) /\<^sub>R (Suc n)" unfolding S_def by (simp del: mult_Suc) obtain r :: real where r0: "0 < r" and r1: "r < 1" using dense [OF zero_less_one] by fast obtain N :: nat where N: "norm x < real N * r" using ex_less_of_nat_mult r0 by auto from r1 show ?thesis proof (rule summable_ratio_test [rule_format]) fix n :: nat assume n: "N \ n" have "norm x \ real N * r" using N by (rule order_less_imp_le) also have "real N * r \ real (Suc n) * r" using r0 n by (simp add: mult_right_mono) finally have "norm x * norm (S n) \ real (Suc n) * r * norm (S n)" using norm_ge_zero by (rule mult_right_mono) then have "norm (x * S n) \ real (Suc n) * r * norm (S n)" by (rule order_trans [OF norm_mult_ineq]) then have "norm (x * S n) / real (Suc n) \ r * norm (S n)" by (simp add: pos_divide_le_eq ac_simps) then show "norm (S (Suc n)) \ r * norm (S n)" by (simp add: S_Suc inverse_eq_divide) qed qed lemma summable_norm_exp: "summable (\n. norm (x^n /\<^sub>R fact n))" for x :: "'a::{real_normed_algebra_1,banach}" proof (rule summable_norm_comparison_test [OF exI, rule_format]) show "summable (\n. norm x^n /\<^sub>R fact n)" by (rule summable_exp_generic) show "norm (x^n /\<^sub>R fact n) \ norm x^n /\<^sub>R fact n" for n by (simp add: norm_power_ineq) qed lemma summable_exp: "summable (\n. inverse (fact n) * x^n)" for x :: "'a::{real_normed_field,banach}" using summable_exp_generic [where x=x] by (simp add: scaleR_conv_of_real nonzero_of_real_inverse) lemma exp_converges: "(\n. x^n /\<^sub>R fact n) sums exp x" unfolding exp_def by (rule summable_exp_generic [THEN summable_sums]) lemma exp_fdiffs: "diffs (\n. inverse (fact n)) = (\n. inverse (fact n :: 'a::{real_normed_field,banach}))" by (simp add: diffs_def mult_ac nonzero_inverse_mult_distrib nonzero_of_real_inverse del: mult_Suc of_nat_Suc) lemma diffs_of_real: "diffs (\n. of_real (f n)) = (\n. of_real (diffs f n))" by (simp add: diffs_def) lemma DERIV_exp [simp]: "DERIV exp x :> exp x" unfolding exp_def scaleR_conv_of_real proof (rule DERIV_cong) have sinv: "summable (\n. of_real (inverse (fact n)) * x ^ n)" for x::'a by (rule exp_converges [THEN sums_summable, unfolded scaleR_conv_of_real]) note xx = exp_converges [THEN sums_summable, unfolded scaleR_conv_of_real] show "((\x. \n. of_real (inverse (fact n)) * x ^ n) has_field_derivative (\n. diffs (\n. of_real (inverse (fact n))) n * x ^ n)) (at x)" by (rule termdiffs [where K="of_real (1 + norm x)"]) (simp_all only: diffs_of_real exp_fdiffs sinv norm_of_real) show "(\n. diffs (\n. of_real (inverse (fact n))) n * x ^ n) = (\n. of_real (inverse (fact n)) * x ^ n)" by (simp add: diffs_of_real exp_fdiffs) qed declare DERIV_exp[THEN DERIV_chain2, derivative_intros] and DERIV_exp[THEN DERIV_chain2, unfolded has_field_derivative_def, derivative_intros] lemmas has_derivative_exp[derivative_intros] = DERIV_exp[THEN DERIV_compose_FDERIV] lemma norm_exp: "norm (exp x) \ exp (norm x)" proof - from summable_norm[OF summable_norm_exp, of x] have "norm (exp x) \ (\n. inverse (fact n) * norm (x^n))" by (simp add: exp_def) also have "\ \ exp (norm x)" using summable_exp_generic[of "norm x"] summable_norm_exp[of x] by (auto simp: exp_def intro!: suminf_le norm_power_ineq) finally show ?thesis . qed lemma isCont_exp: "isCont exp x" for x :: "'a::{real_normed_field,banach}" by (rule DERIV_exp [THEN DERIV_isCont]) lemma isCont_exp' [simp]: "isCont f a \ isCont (\x. exp (f x)) a" for f :: "_ \'a::{real_normed_field,banach}" by (rule isCont_o2 [OF _ isCont_exp]) lemma tendsto_exp [tendsto_intros]: "(f \ a) F \ ((\x. exp (f x)) \ exp a) F" for f:: "_ \'a::{real_normed_field,banach}" by (rule isCont_tendsto_compose [OF isCont_exp]) lemma continuous_exp [continuous_intros]: "continuous F f \ continuous F (\x. exp (f x))" for f :: "_ \'a::{real_normed_field,banach}" unfolding continuous_def by (rule tendsto_exp) lemma continuous_on_exp [continuous_intros]: "continuous_on s f \ continuous_on s (\x. exp (f x))" for f :: "_ \'a::{real_normed_field,banach}" unfolding continuous_on_def by (auto intro: tendsto_exp) subsubsection \Properties of the Exponential Function\ lemma exp_zero [simp]: "exp 0 = 1" unfolding exp_def by (simp add: scaleR_conv_of_real) lemma exp_series_add_commuting: fixes x y :: "'a::{real_normed_algebra_1,banach}" defines S_def: "S \ \x n. x^n /\<^sub>R fact n" assumes comm: "x * y = y * x" shows "S (x + y) n = (\i\n. S x i * S y (n - i))" proof (induct n) case 0 show ?case unfolding S_def by simp next case (Suc n) have S_Suc: "\x n. S x (Suc n) = (x * S x n) /\<^sub>R real (Suc n)" unfolding S_def by (simp del: mult_Suc) then have times_S: "\x n. x * S x n = real (Suc n) *\<^sub>R S x (Suc n)" by simp have S_comm: "\n. S x n * y = y * S x n" by (simp add: power_commuting_commutes comm S_def) have "real (Suc n) *\<^sub>R S (x + y) (Suc n) = (x + y) * (\i\n. S x i * S y (n - i))" by (metis Suc.hyps times_S) also have "\ = x * (\i\n. S x i * S y (n - i)) + y * (\i\n. S x i * S y (n - i))" by (rule distrib_right) also have "\ = (\i\n. x * S x i * S y (n - i)) + (\i\n. S x i * y * S y (n - i))" by (simp add: sum_distrib_left ac_simps S_comm) also have "\ = (\i\n. x * S x i * S y (n - i)) + (\i\n. S x i * (y * S y (n - i)))" by (simp add: ac_simps) also have "\ = (\i\n. real (Suc i) *\<^sub>R (S x (Suc i) * S y (n - i))) + (\i\n. real (Suc n - i) *\<^sub>R (S x i * S y (Suc n - i)))" by (simp add: times_S Suc_diff_le) also have "(\i\n. real (Suc i) *\<^sub>R (S x (Suc i) * S y (n - i))) = (\i\Suc n. real i *\<^sub>R (S x i * S y (Suc n - i)))" by (subst sum.atMost_Suc_shift) simp also have "(\i\n. real (Suc n - i) *\<^sub>R (S x i * S y (Suc n - i))) = (\i\Suc n. real (Suc n - i) *\<^sub>R (S x i * S y (Suc n - i)))" by simp also have "(\i\Suc n. real i *\<^sub>R (S x i * S y (Suc n - i))) + (\i\Suc n. real (Suc n - i) *\<^sub>R (S x i * S y (Suc n - i))) = (\i\Suc n. real (Suc n) *\<^sub>R (S x i * S y (Suc n - i)))" by (simp flip: sum.distrib scaleR_add_left of_nat_add) also have "\ = real (Suc n) *\<^sub>R (\i\Suc n. S x i * S y (Suc n - i))" by (simp only: scaleR_right.sum) finally show "S (x + y) (Suc n) = (\i\Suc n. S x i * S y (Suc n - i))" by (simp del: sum.cl_ivl_Suc) qed lemma exp_add_commuting: "x * y = y * x \ exp (x + y) = exp x * exp y" by (simp only: exp_def Cauchy_product summable_norm_exp exp_series_add_commuting) lemma exp_times_arg_commute: "exp A * A = A * exp A" by (simp add: exp_def suminf_mult[symmetric] summable_exp_generic power_commutes suminf_mult2) lemma exp_add: "exp (x + y) = exp x * exp y" for x y :: "'a::{real_normed_field,banach}" by (rule exp_add_commuting) (simp add: ac_simps) lemma exp_double: "exp(2 * z) = exp z ^ 2" by (simp add: exp_add_commuting mult_2 power2_eq_square) lemmas mult_exp_exp = exp_add [symmetric] lemma exp_of_real: "exp (of_real x) = of_real (exp x)" unfolding exp_def apply (subst suminf_of_real [OF summable_exp_generic]) apply (simp add: scaleR_conv_of_real) done lemmas of_real_exp = exp_of_real[symmetric] corollary exp_in_Reals [simp]: "z \ \ \ exp z \ \" by (metis Reals_cases Reals_of_real exp_of_real) lemma exp_not_eq_zero [simp]: "exp x \ 0" proof have "exp x * exp (- x) = 1" by (simp add: exp_add_commuting[symmetric]) also assume "exp x = 0" finally show False by simp qed lemma exp_minus_inverse: "exp x * exp (- x) = 1" by (simp add: exp_add_commuting[symmetric]) lemma exp_minus: "exp (- x) = inverse (exp x)" for x :: "'a::{real_normed_field,banach}" by (intro inverse_unique [symmetric] exp_minus_inverse) lemma exp_diff: "exp (x - y) = exp x / exp y" for x :: "'a::{real_normed_field,banach}" using exp_add [of x "- y"] by (simp add: exp_minus divide_inverse) lemma exp_of_nat_mult: "exp (of_nat n * x) = exp x ^ n" for x :: "'a::{real_normed_field,banach}" by (induct n) (auto simp: distrib_left exp_add mult.commute) corollary exp_of_nat2_mult: "exp (x * of_nat n) = exp x ^ n" for x :: "'a::{real_normed_field,banach}" by (metis exp_of_nat_mult mult_of_nat_commute) lemma exp_sum: "finite I \ exp (sum f I) = prod (\x. exp (f x)) I" by (induct I rule: finite_induct) (auto simp: exp_add_commuting mult.commute) lemma exp_divide_power_eq: fixes x :: "'a::{real_normed_field,banach}" assumes "n > 0" shows "exp (x / of_nat n) ^ n = exp x" using assms proof (induction n arbitrary: x) case (Suc n) show ?case proof (cases "n = 0") case True then show ?thesis by simp next case False have [simp]: "1 + (of_nat n * of_nat n + of_nat n * 2) \ (0::'a)" using of_nat_eq_iff [of "1 + n * n + n * 2" "0"] by simp from False have [simp]: "x * of_nat n / (1 + of_nat n) / of_nat n = x / (1 + of_nat n)" by simp have [simp]: "x / (1 + of_nat n) + x * of_nat n / (1 + of_nat n) = x" using of_nat_neq_0 by (auto simp add: field_split_simps) show ?thesis using Suc.IH [of "x * of_nat n / (1 + of_nat n)"] False by (simp add: exp_add [symmetric]) qed qed simp subsubsection \Properties of the Exponential Function on Reals\ text \Comparisons of \<^term>\exp x\ with zero.\ text \Proof: because every exponential can be seen as a square.\ lemma exp_ge_zero [simp]: "0 \ exp x" for x :: real proof - have "0 \ exp (x/2) * exp (x/2)" by simp then show ?thesis by (simp add: exp_add [symmetric]) qed lemma exp_gt_zero [simp]: "0 < exp x" for x :: real by (simp add: order_less_le) lemma not_exp_less_zero [simp]: "\ exp x < 0" for x :: real by (simp add: not_less) lemma not_exp_le_zero [simp]: "\ exp x \ 0" for x :: real by (simp add: not_le) lemma abs_exp_cancel [simp]: "\exp x\ = exp x" for x :: real by simp text \Strict monotonicity of exponential.\ lemma exp_ge_add_one_self_aux: fixes x :: real assumes "0 \ x" shows "1 + x \ exp x" using order_le_imp_less_or_eq [OF assms] proof assume "0 < x" have "1 + x \ (\n<2. inverse (fact n) * x^n)" by (auto simp: numeral_2_eq_2) also have "\ \ (\n. inverse (fact n) * x^n)" using \0 < x\ by (auto simp add: zero_le_mult_iff intro: sum_le_suminf [OF summable_exp]) finally show "1 + x \ exp x" by (simp add: exp_def) qed auto lemma exp_gt_one: "0 < x \ 1 < exp x" for x :: real proof - assume x: "0 < x" then have "1 < 1 + x" by simp also from x have "1 + x \ exp x" by (simp add: exp_ge_add_one_self_aux) finally show ?thesis . qed lemma exp_less_mono: fixes x y :: real assumes "x < y" shows "exp x < exp y" proof - from \x < y\ have "0 < y - x" by simp then have "1 < exp (y - x)" by (rule exp_gt_one) then have "1 < exp y / exp x" by (simp only: exp_diff) then show "exp x < exp y" by simp qed lemma exp_less_cancel: "exp x < exp y \ x < y" for x y :: real unfolding linorder_not_le [symmetric] by (auto simp: order_le_less exp_less_mono) lemma exp_less_cancel_iff [iff]: "exp x < exp y \ x < y" for x y :: real by (auto intro: exp_less_mono exp_less_cancel) lemma exp_le_cancel_iff [iff]: "exp x \ exp y \ x \ y" for x y :: real by (auto simp: linorder_not_less [symmetric]) lemma exp_inj_iff [iff]: "exp x = exp y \ x = y" for x y :: real by (simp add: order_eq_iff) text \Comparisons of \<^term>\exp x\ with one.\ lemma one_less_exp_iff [simp]: "1 < exp x \ 0 < x" for x :: real using exp_less_cancel_iff [where x = 0 and y = x] by simp lemma exp_less_one_iff [simp]: "exp x < 1 \ x < 0" for x :: real using exp_less_cancel_iff [where x = x and y = 0] by simp lemma one_le_exp_iff [simp]: "1 \ exp x \ 0 \ x" for x :: real using exp_le_cancel_iff [where x = 0 and y = x] by simp lemma exp_le_one_iff [simp]: "exp x \ 1 \ x \ 0" for x :: real using exp_le_cancel_iff [where x = x and y = 0] by simp lemma exp_eq_one_iff [simp]: "exp x = 1 \ x = 0" for x :: real using exp_inj_iff [where x = x and y = 0] by simp lemma lemma_exp_total: "1 \ y \ \x. 0 \ x \ x \ y - 1 \ exp x = y" for y :: real proof (rule IVT) assume "1 \ y" then have "0 \ y - 1" by simp then have "1 + (y - 1) \ exp (y - 1)" by (rule exp_ge_add_one_self_aux) then show "y \ exp (y - 1)" by simp qed (simp_all add: le_diff_eq) lemma exp_total: "0 < y \ \x. exp x = y" for y :: real proof (rule linorder_le_cases [of 1 y]) assume "1 \ y" then show "\x. exp x = y" by (fast dest: lemma_exp_total) next assume "0 < y" and "y \ 1" then have "1 \ inverse y" by (simp add: one_le_inverse_iff) then obtain x where "exp x = inverse y" by (fast dest: lemma_exp_total) then have "exp (- x) = y" by (simp add: exp_minus) then show "\x. exp x = y" .. qed subsection \Natural Logarithm\ class ln = real_normed_algebra_1 + banach + fixes ln :: "'a \ 'a" assumes ln_one [simp]: "ln 1 = 0" definition powr :: "'a \ 'a \ 'a::ln" (infixr "powr" 80) \ \exponentation via ln and exp\ where "x powr a \ if x = 0 then 0 else exp (a * ln x)" lemma powr_0 [simp]: "0 powr z = 0" by (simp add: powr_def) instantiation real :: ln begin definition ln_real :: "real \ real" where "ln_real x = (THE u. exp u = x)" instance by intro_classes (simp add: ln_real_def) end lemma powr_eq_0_iff [simp]: "w powr z = 0 \ w = 0" by (simp add: powr_def) lemma ln_exp [simp]: "ln (exp x) = x" for x :: real by (simp add: ln_real_def) lemma exp_ln [simp]: "0 < x \ exp (ln x) = x" for x :: real by (auto dest: exp_total) lemma exp_ln_iff [simp]: "exp (ln x) = x \ 0 < x" for x :: real by (metis exp_gt_zero exp_ln) lemma ln_unique: "exp y = x \ ln x = y" for x :: real by (erule subst) (rule ln_exp) lemma ln_mult: "0 < x \ 0 < y \ ln (x * y) = ln x + ln y" for x :: real by (rule ln_unique) (simp add: exp_add) lemma ln_prod: "finite I \ (\i. i \ I \ f i > 0) \ ln (prod f I) = sum (\x. ln(f x)) I" for f :: "'a \ real" by (induct I rule: finite_induct) (auto simp: ln_mult prod_pos) lemma ln_inverse: "0 < x \ ln (inverse x) = - ln x" for x :: real by (rule ln_unique) (simp add: exp_minus) lemma ln_div: "0 < x \ 0 < y \ ln (x / y) = ln x - ln y" for x :: real by (rule ln_unique) (simp add: exp_diff) lemma ln_realpow: "0 < x \ ln (x^n) = real n * ln x" by (rule ln_unique) (simp add: exp_of_nat_mult) lemma ln_less_cancel_iff [simp]: "0 < x \ 0 < y \ ln x < ln y \ x < y" for x :: real by (subst exp_less_cancel_iff [symmetric]) simp lemma ln_le_cancel_iff [simp]: "0 < x \ 0 < y \ ln x \ ln y \ x \ y" for x :: real by (simp add: linorder_not_less [symmetric]) lemma ln_inj_iff [simp]: "0 < x \ 0 < y \ ln x = ln y \ x = y" for x :: real by (simp add: order_eq_iff) lemma ln_add_one_self_le_self: "0 \ x \ ln (1 + x) \ x" for x :: real by (rule exp_le_cancel_iff [THEN iffD1]) (simp add: exp_ge_add_one_self_aux) lemma ln_less_self [simp]: "0 < x \ ln x < x" for x :: real by (rule order_less_le_trans [where y = "ln (1 + x)"]) (simp_all add: ln_add_one_self_le_self) lemma ln_ge_iff: "\x::real. 0 < x \ y \ ln x \ exp y \ x" using exp_le_cancel_iff exp_total by force lemma ln_ge_zero [simp]: "1 \ x \ 0 \ ln x" for x :: real using ln_le_cancel_iff [of 1 x] by simp lemma ln_ge_zero_imp_ge_one: "0 \ ln x \ 0 < x \ 1 \ x" for x :: real using ln_le_cancel_iff [of 1 x] by simp lemma ln_ge_zero_iff [simp]: "0 < x \ 0 \ ln x \ 1 \ x" for x :: real using ln_le_cancel_iff [of 1 x] by simp lemma ln_less_zero_iff [simp]: "0 < x \ ln x < 0 \ x < 1" for x :: real using ln_less_cancel_iff [of x 1] by simp lemma ln_le_zero_iff [simp]: "0 < x \ ln x \ 0 \ x \ 1" for x :: real by (metis less_numeral_extra(1) ln_le_cancel_iff ln_one) lemma ln_gt_zero: "1 < x \ 0 < ln x" for x :: real using ln_less_cancel_iff [of 1 x] by simp lemma ln_gt_zero_imp_gt_one: "0 < ln x \ 0 < x \ 1 < x" for x :: real using ln_less_cancel_iff [of 1 x] by simp lemma ln_gt_zero_iff [simp]: "0 < x \ 0 < ln x \ 1 < x" for x :: real using ln_less_cancel_iff [of 1 x] by simp lemma ln_eq_zero_iff [simp]: "0 < x \ ln x = 0 \ x = 1" for x :: real using ln_inj_iff [of x 1] by simp lemma ln_less_zero: "0 < x \ x < 1 \ ln x < 0" for x :: real by simp lemma ln_neg_is_const: "x \ 0 \ ln x = (THE x. False)" for x :: real by (auto simp: ln_real_def intro!: arg_cong[where f = The]) lemma powr_eq_one_iff [simp]: "a powr x = 1 \ x = 0" if "a > 1" for a x :: real using that by (auto simp: powr_def split: if_splits) lemma isCont_ln: fixes x :: real assumes "x \ 0" shows "isCont ln x" proof (cases "0 < x") case True then have "isCont ln (exp (ln x))" by (intro isCont_inverse_function[where d = "\x\" and f = exp]) auto with True show ?thesis by simp next case False with \x \ 0\ show "isCont ln x" unfolding isCont_def by (subst filterlim_cong[OF _ refl, of _ "nhds (ln 0)" _ "\_. ln 0"]) (auto simp: ln_neg_is_const not_less eventually_at dist_real_def intro!: exI[of _ "\x\"]) qed lemma tendsto_ln [tendsto_intros]: "(f \ a) F \ a \ 0 \ ((\x. ln (f x)) \ ln a) F" for a :: real by (rule isCont_tendsto_compose [OF isCont_ln]) lemma continuous_ln: "continuous F f \ f (Lim F (\x. x)) \ 0 \ continuous F (\x. ln (f x :: real))" unfolding continuous_def by (rule tendsto_ln) lemma isCont_ln' [continuous_intros]: "continuous (at x) f \ f x \ 0 \ continuous (at x) (\x. ln (f x :: real))" unfolding continuous_at by (rule tendsto_ln) lemma continuous_within_ln [continuous_intros]: "continuous (at x within s) f \ f x \ 0 \ continuous (at x within s) (\x. ln (f x :: real))" unfolding continuous_within by (rule tendsto_ln) lemma continuous_on_ln [continuous_intros]: "continuous_on s f \ (\x\s. f x \ 0) \ continuous_on s (\x. ln (f x :: real))" unfolding continuous_on_def by (auto intro: tendsto_ln) lemma DERIV_ln: "0 < x \ DERIV ln x :> inverse x" for x :: real by (rule DERIV_inverse_function [where f=exp and a=0 and b="x+1"]) (auto intro: DERIV_cong [OF DERIV_exp exp_ln] isCont_ln) lemma DERIV_ln_divide: "0 < x \ DERIV ln x :> 1 / x" for x :: real by (rule DERIV_ln[THEN DERIV_cong]) (simp_all add: divide_inverse) declare DERIV_ln_divide[THEN DERIV_chain2, derivative_intros] and DERIV_ln_divide[THEN DERIV_chain2, unfolded has_field_derivative_def, derivative_intros] lemmas has_derivative_ln[derivative_intros] = DERIV_ln[THEN DERIV_compose_FDERIV] lemma ln_series: assumes "0 < x" and "x < 2" shows "ln x = (\ n. (-1)^n * (1 / real (n + 1)) * (x - 1)^(Suc n))" (is "ln x = suminf (?f (x - 1))") proof - let ?f' = "\x n. (-1)^n * (x - 1)^n" have "ln x - suminf (?f (x - 1)) = ln 1 - suminf (?f (1 - 1))" proof (rule DERIV_isconst3 [where x = x]) fix x :: real assume "x \ {0 <..< 2}" then have "0 < x" and "x < 2" by auto have "norm (1 - x) < 1" using \0 < x\ and \x < 2\ by auto have "1 / x = 1 / (1 - (1 - x))" by auto also have "\ = (\ n. (1 - x)^n)" using geometric_sums[OF \norm (1 - x) < 1\] by (rule sums_unique) also have "\ = suminf (?f' x)" unfolding power_mult_distrib[symmetric] by (rule arg_cong[where f=suminf], rule arg_cong[where f="(^)"], auto) finally have "DERIV ln x :> suminf (?f' x)" using DERIV_ln[OF \0 < x\] unfolding divide_inverse by auto moreover have repos: "\ h x :: real. h - 1 + x = h + x - 1" by auto have "DERIV (\x. suminf (?f x)) (x - 1) :> (\n. (-1)^n * (1 / real (n + 1)) * real (Suc n) * (x - 1) ^ n)" proof (rule DERIV_power_series') show "x - 1 \ {- 1<..<1}" and "(0 :: real) < 1" using \0 < x\ \x < 2\ by auto next fix x :: real assume "x \ {- 1<..<1}" then show "summable (\n. (- 1) ^ n * (1 / real (n + 1)) * real (Suc n) * x^n)" by (simp add: abs_if flip: power_mult_distrib) qed then have "DERIV (\x. suminf (?f x)) (x - 1) :> suminf (?f' x)" unfolding One_nat_def by auto then have "DERIV (\x. suminf (?f (x - 1))) x :> suminf (?f' x)" unfolding DERIV_def repos . ultimately have "DERIV (\x. ln x - suminf (?f (x - 1))) x :> suminf (?f' x) - suminf (?f' x)" by (rule DERIV_diff) then show "DERIV (\x. ln x - suminf (?f (x - 1))) x :> 0" by auto qed (auto simp: assms) then show ?thesis by auto qed lemma exp_first_terms: fixes x :: "'a::{real_normed_algebra_1,banach}" shows "exp x = (\nR (x ^ n)) + (\n. inverse(fact (n + k)) *\<^sub>R (x ^ (n + k)))" proof - have "exp x = suminf (\n. inverse(fact n) *\<^sub>R (x^n))" by (simp add: exp_def) also from summable_exp_generic have "\ = (\ n. inverse(fact(n+k)) *\<^sub>R (x ^ (n + k))) + (\ n::natR (x^n))" (is "_ = _ + ?a") by (rule suminf_split_initial_segment) finally show ?thesis by simp qed lemma exp_first_term: "exp x = 1 + (\n. inverse (fact (Suc n)) *\<^sub>R (x ^ Suc n))" for x :: "'a::{real_normed_algebra_1,banach}" using exp_first_terms[of x 1] by simp lemma exp_first_two_terms: "exp x = 1 + x + (\n. inverse (fact (n + 2)) *\<^sub>R (x ^ (n + 2)))" for x :: "'a::{real_normed_algebra_1,banach}" using exp_first_terms[of x 2] by (simp add: eval_nat_numeral) lemma exp_bound: fixes x :: real assumes a: "0 \ x" and b: "x \ 1" shows "exp x \ 1 + x + x\<^sup>2" proof - have "suminf (\n. inverse(fact (n+2)) * (x ^ (n + 2))) \ x\<^sup>2" proof - - have "(\n. x\<^sup>2 / 2 * (1 / 2) ^ n) sums (x\<^sup>2 / 2 * (1 / (1 - 1 / 2)))" + have "(\n. x\<^sup>2 / 2 * (1/2) ^ n) sums (x\<^sup>2 / 2 * (1 / (1 - 1/2)))" by (intro sums_mult geometric_sums) simp - then have sumsx: "(\n. x\<^sup>2 / 2 * (1 / 2) ^ n) sums x\<^sup>2" + then have sumsx: "(\n. x\<^sup>2 / 2 * (1/2) ^ n) sums x\<^sup>2" by simp have "suminf (\n. inverse(fact (n+2)) * (x ^ (n + 2))) \ suminf (\n. (x\<^sup>2/2) * ((1/2)^n))" proof (intro suminf_le allI) show "inverse (fact (n + 2)) * x ^ (n + 2) \ (x\<^sup>2/2) * ((1/2)^n)" for n :: nat proof - have "(2::nat) * 2 ^ n \ fact (n + 2)" by (induct n) simp_all then have "real ((2::nat) * 2 ^ n) \ real_of_nat (fact (n + 2))" by (simp only: of_nat_le_iff) then have "((2::real) * 2 ^ n) \ fact (n + 2)" unfolding of_nat_fact by simp then have "inverse (fact (n + 2)) \ inverse ((2::real) * 2 ^ n)" by (rule le_imp_inverse_le) simp then have "inverse (fact (n + 2)) \ 1/(2::real) * (1/2)^n" by (simp add: power_inverse [symmetric]) then have "inverse (fact (n + 2)) * (x^n * x\<^sup>2) \ 1/2 * (1/2)^n * (1 * x\<^sup>2)" by (rule mult_mono) (rule mult_mono, simp_all add: power_le_one a b) then show ?thesis unfolding power_add by (simp add: ac_simps del: fact_Suc) qed show "summable (\n. inverse (fact (n + 2)) * x ^ (n + 2))" by (rule summable_exp [THEN summable_ignore_initial_segment]) - show "summable (\n. x\<^sup>2 / 2 * (1 / 2) ^ n)" + show "summable (\n. x\<^sup>2 / 2 * (1/2) ^ n)" by (rule sums_summable [OF sumsx]) qed also have "\ = x\<^sup>2" by (rule sums_unique [THEN sym]) (rule sumsx) finally show ?thesis . qed then show ?thesis unfolding exp_first_two_terms by auto qed corollary exp_half_le2: "exp(1/2) \ (2::real)" using exp_bound [of "1/2"] by (simp add: field_simps) corollary exp_le: "exp 1 \ (3::real)" using exp_bound [of 1] by (simp add: field_simps) lemma exp_bound_half: "norm z \ 1/2 \ norm (exp z) \ 2" by (blast intro: order_trans intro!: exp_half_le2 norm_exp) lemma exp_bound_lemma: assumes "norm z \ 1/2" shows "norm (exp z) \ 1 + 2 * norm z" proof - have *: "(norm z)\<^sup>2 \ norm z * 1" unfolding power2_eq_square by (rule mult_left_mono) (use assms in auto) have "norm (exp z) \ exp (norm z)" by (rule norm_exp) also have "\ \ 1 + (norm z) + (norm z)\<^sup>2" using assms exp_bound by auto also have "\ \ 1 + 2 * norm z" using * by auto finally show ?thesis . qed lemma real_exp_bound_lemma: "0 \ x \ x \ 1/2 \ exp x \ 1 + 2 * x" for x :: real using exp_bound_lemma [of x] by simp lemma ln_one_minus_pos_upper_bound: fixes x :: real assumes a: "0 \ x" and b: "x < 1" shows "ln (1 - x) \ - x" proof - have "(1 - x) * (1 + x + x\<^sup>2) = 1 - x^3" by (simp add: algebra_simps power2_eq_square power3_eq_cube) also have "\ \ 1" by (auto simp: a) finally have "(1 - x) * (1 + x + x\<^sup>2) \ 1" . moreover have c: "0 < 1 + x + x\<^sup>2" by (simp add: add_pos_nonneg a) ultimately have "1 - x \ 1 / (1 + x + x\<^sup>2)" by (elim mult_imp_le_div_pos) also have "\ \ 1 / exp x" by (metis a abs_one b exp_bound exp_gt_zero frac_le less_eq_real_def real_sqrt_abs real_sqrt_pow2_iff real_sqrt_power) also have "\ = exp (- x)" by (auto simp: exp_minus divide_inverse) finally have "1 - x \ exp (- x)" . also have "1 - x = exp (ln (1 - x))" by (metis b diff_0 exp_ln_iff less_iff_diff_less_0 minus_diff_eq) finally have "exp (ln (1 - x)) \ exp (- x)" . then show ?thesis by (auto simp only: exp_le_cancel_iff) qed lemma exp_ge_add_one_self [simp]: "1 + x \ exp x" for x :: real proof (cases "0 \ x \ x \ -1") case True then show ?thesis by (meson exp_ge_add_one_self_aux exp_ge_zero order.trans real_add_le_0_iff) next case False then have ln1: "ln (1 + x) \ x" using ln_one_minus_pos_upper_bound [of "-x"] by simp have "1 + x = exp (ln (1 + x))" using False by auto also have "\ \ exp x" by (simp add: ln1) finally show ?thesis . qed lemma ln_one_plus_pos_lower_bound: fixes x :: real assumes a: "0 \ x" and b: "x \ 1" shows "x - x\<^sup>2 \ ln (1 + x)" proof - have "exp (x - x\<^sup>2) = exp x / exp (x\<^sup>2)" by (rule exp_diff) also have "\ \ (1 + x + x\<^sup>2) / exp (x \<^sup>2)" by (metis a b divide_right_mono exp_bound exp_ge_zero) also have "\ \ (1 + x + x\<^sup>2) / (1 + x\<^sup>2)" by (simp add: a divide_left_mono add_pos_nonneg) also from a have "\ \ 1 + x" by (simp add: field_simps add_strict_increasing zero_le_mult_iff) finally have "exp (x - x\<^sup>2) \ 1 + x" . also have "\ = exp (ln (1 + x))" proof - from a have "0 < 1 + x" by auto then show ?thesis by (auto simp only: exp_ln_iff [THEN sym]) qed finally have "exp (x - x\<^sup>2) \ exp (ln (1 + x))" . then show ?thesis by (metis exp_le_cancel_iff) qed lemma ln_one_minus_pos_lower_bound: fixes x :: real - assumes a: "0 \ x" and b: "x \ 1 / 2" + assumes a: "0 \ x" and b: "x \ 1/2" shows "- x - 2 * x\<^sup>2 \ ln (1 - x)" proof - from b have c: "x < 1" by auto then have "ln (1 - x) = - ln (1 + x / (1 - x))" by (auto simp: ln_inverse [symmetric] field_simps intro: arg_cong [where f=ln]) also have "- (x / (1 - x)) \ \" proof - have "ln (1 + x / (1 - x)) \ x / (1 - x)" using a c by (intro ln_add_one_self_le_self) auto then show ?thesis by auto qed also have "- (x / (1 - x)) = - x / (1 - x)" by auto finally have d: "- x / (1 - x) \ ln (1 - x)" . have "0 < 1 - x" using a b by simp then have e: "- x - 2 * x\<^sup>2 \ - x / (1 - x)" using mult_right_le_one_le[of "x * x" "2 * x"] a b by (simp add: field_simps power2_eq_square) from e d show "- x - 2 * x\<^sup>2 \ ln (1 - x)" by (rule order_trans) qed lemma ln_add_one_self_le_self2: fixes x :: real shows "-1 < x \ ln (1 + x) \ x" by (metis diff_gt_0_iff_gt diff_minus_eq_add exp_ge_add_one_self exp_le_cancel_iff exp_ln minus_less_iff) lemma abs_ln_one_plus_x_minus_x_bound_nonneg: fixes x :: real assumes x: "0 \ x" and x1: "x \ 1" shows "\ln (1 + x) - x\ \ x\<^sup>2" proof - from x have "ln (1 + x) \ x" by (rule ln_add_one_self_le_self) then have "ln (1 + x) - x \ 0" by simp then have "\ln(1 + x) - x\ = - (ln(1 + x) - x)" by (rule abs_of_nonpos) also have "\ = x - ln (1 + x)" by simp also have "\ \ x\<^sup>2" proof - from x x1 have "x - x\<^sup>2 \ ln (1 + x)" by (intro ln_one_plus_pos_lower_bound) then show ?thesis by simp qed finally show ?thesis . qed lemma abs_ln_one_plus_x_minus_x_bound_nonpos: fixes x :: real - assumes a: "-(1 / 2) \ x" and b: "x \ 0" + assumes a: "-(1/2) \ x" and b: "x \ 0" shows "\ln (1 + x) - x\ \ 2 * x\<^sup>2" proof - have *: "- (-x) - 2 * (-x)\<^sup>2 \ ln (1 - (- x))" by (metis a b diff_zero ln_one_minus_pos_lower_bound minus_diff_eq neg_le_iff_le) have "\ln (1 + x) - x\ = x - ln (1 - (- x))" using a ln_add_one_self_le_self2 [of x] by (simp add: abs_if) also have "\ \ 2 * x\<^sup>2" using * by (simp add: algebra_simps) finally show ?thesis . qed lemma abs_ln_one_plus_x_minus_x_bound: fixes x :: real - assumes "\x\ \ 1 / 2" + assumes "\x\ \ 1/2" shows "\ln (1 + x) - x\ \ 2 * x\<^sup>2" proof (cases "0 \ x") case True then show ?thesis using abs_ln_one_plus_x_minus_x_bound_nonneg assms by fastforce next case False then show ?thesis using abs_ln_one_plus_x_minus_x_bound_nonpos assms by auto qed lemma ln_x_over_x_mono: fixes x :: real assumes x: "exp 1 \ x" "x \ y" shows "ln y / y \ ln x / x" proof - note x moreover have "0 < exp (1::real)" by simp ultimately have a: "0 < x" and b: "0 < y" by (fast intro: less_le_trans order_trans)+ have "x * ln y - x * ln x = x * (ln y - ln x)" by (simp add: algebra_simps) also have "\ = x * ln (y / x)" by (simp only: ln_div a b) also have "y / x = (x + (y - x)) / x" by simp also have "\ = 1 + (y - x) / x" using x a by (simp add: field_simps) also have "x * ln (1 + (y - x) / x) \ x * ((y - x) / x)" using x a by (intro mult_left_mono ln_add_one_self_le_self) simp_all also have "\ = y - x" using a by simp also have "\ = (y - x) * ln (exp 1)" by simp also have "\ \ (y - x) * ln x" using a x exp_total of_nat_1 x(1) by (fastforce intro: mult_left_mono) also have "\ = y * ln x - x * ln x" by (rule left_diff_distrib) finally have "x * ln y \ y * ln x" by arith then have "ln y \ (y * ln x) / x" using a by (simp add: field_simps) also have "\ = y * (ln x / x)" by simp finally show ?thesis using b by (simp add: field_simps) qed lemma ln_le_minus_one: "0 < x \ ln x \ x - 1" for x :: real using exp_ge_add_one_self[of "ln x"] by simp corollary ln_diff_le: "0 < x \ 0 < y \ ln x - ln y \ (x - y) / y" for x :: real by (simp add: ln_div [symmetric] diff_divide_distrib ln_le_minus_one) lemma ln_eq_minus_one: fixes x :: real assumes "0 < x" "ln x = x - 1" shows "x = 1" proof - let ?l = "\y. ln y - y + 1" have D: "\x::real. 0 < x \ DERIV ?l x :> (1 / x - 1)" by (auto intro!: derivative_eq_intros) show ?thesis proof (cases rule: linorder_cases) assume "x < 1" from dense[OF \x < 1\] obtain a where "x < a" "a < 1" by blast from \x < a\ have "?l x < ?l a" proof (rule DERIV_pos_imp_increasing) fix y assume "x \ y" "y \ a" with \0 < x\ \a < 1\ have "0 < 1 / y - 1" "0 < y" by (auto simp: field_simps) with D show "\z. DERIV ?l y :> z \ 0 < z" by blast qed also have "\ \ 0" using ln_le_minus_one \0 < x\ \x < a\ by (auto simp: field_simps) finally show "x = 1" using assms by auto next assume "1 < x" from dense[OF this] obtain a where "1 < a" "a < x" by blast from \a < x\ have "?l x < ?l a" proof (rule DERIV_neg_imp_decreasing) fix y assume "a \ y" "y \ x" with \1 < a\ have "1 / y - 1 < 0" "0 < y" by (auto simp: field_simps) with D show "\z. DERIV ?l y :> z \ z < 0" by blast qed also have "\ \ 0" using ln_le_minus_one \1 < a\ by (auto simp: field_simps) finally show "x = 1" using assms by auto next assume "x = 1" then show ?thesis by simp qed qed lemma ln_x_over_x_tendsto_0: "((\x::real. ln x / x) \ 0) at_top" proof (rule lhospital_at_top_at_top[where f' = inverse and g' = "\_. 1"]) from eventually_gt_at_top[of "0::real"] show "\\<^sub>F x in at_top. (ln has_real_derivative inverse x) (at x)" by eventually_elim (auto intro!: derivative_eq_intros simp: field_simps) qed (use tendsto_inverse_0 in \auto simp: filterlim_ident dest!: tendsto_mono[OF at_top_le_at_infinity]\) lemma exp_ge_one_plus_x_over_n_power_n: assumes "x \ - real n" "n > 0" shows "(1 + x / of_nat n) ^ n \ exp x" proof (cases "x = - of_nat n") case False from assms False have "(1 + x / of_nat n) ^ n = exp (of_nat n * ln (1 + x / of_nat n))" by (subst exp_of_nat_mult, subst exp_ln) (simp_all add: field_simps) also from assms False have "ln (1 + x / real n) \ x / real n" by (intro ln_add_one_self_le_self2) (simp_all add: field_simps) with assms have "exp (of_nat n * ln (1 + x / of_nat n)) \ exp x" by (simp add: field_simps) finally show ?thesis . next case True then show ?thesis by (simp add: zero_power) qed lemma exp_ge_one_minus_x_over_n_power_n: assumes "x \ real n" "n > 0" shows "(1 - x / of_nat n) ^ n \ exp (-x)" using exp_ge_one_plus_x_over_n_power_n[of n "-x"] assms by simp lemma exp_at_bot: "(exp \ (0::real)) at_bot" unfolding tendsto_Zfun_iff proof (rule ZfunI, simp add: eventually_at_bot_dense) fix r :: real assume "0 < r" have "exp x < r" if "x < ln r" for x by (metis \0 < r\ exp_less_mono exp_ln that) then show "\k. \n at_top" by (rule filterlim_at_top_at_top[where Q="\x. True" and P="\x. 0 < x" and g=ln]) (auto intro: eventually_gt_at_top) lemma lim_exp_minus_1: "((\z::'a. (exp(z) - 1) / z) \ 1) (at 0)" for x :: "'a::{real_normed_field,banach}" proof - have "((\z::'a. exp(z) - 1) has_field_derivative 1) (at 0)" by (intro derivative_eq_intros | simp)+ then show ?thesis by (simp add: Deriv.has_field_derivative_iff) qed lemma ln_at_0: "LIM x at_right 0. ln (x::real) :> at_bot" by (rule filterlim_at_bot_at_right[where Q="\x. 0 < x" and P="\x. True" and g=exp]) (auto simp: eventually_at_filter) lemma ln_at_top: "LIM x at_top. ln (x::real) :> at_top" by (rule filterlim_at_top_at_top[where Q="\x. 0 < x" and P="\x. True" and g=exp]) (auto intro: eventually_gt_at_top) lemma filtermap_ln_at_top: "filtermap (ln::real \ real) at_top = at_top" by (intro filtermap_fun_inverse[of exp] exp_at_top ln_at_top) auto lemma filtermap_exp_at_top: "filtermap (exp::real \ real) at_top = at_top" by (intro filtermap_fun_inverse[of ln] exp_at_top ln_at_top) (auto simp: eventually_at_top_dense) lemma filtermap_ln_at_right: "filtermap ln (at_right (0::real)) = at_bot" by (auto intro!: filtermap_fun_inverse[where g="\x. exp x"] ln_at_0 simp: filterlim_at exp_at_bot) lemma tendsto_power_div_exp_0: "((\x. x ^ k / exp x) \ (0::real)) at_top" proof (induct k) case 0 show "((\x. x ^ 0 / exp x) \ (0::real)) at_top" by (simp add: inverse_eq_divide[symmetric]) (metis filterlim_compose[OF tendsto_inverse_0] exp_at_top filterlim_mono at_top_le_at_infinity order_refl) next case (Suc k) show ?case proof (rule lhospital_at_top_at_top) show "eventually (\x. DERIV (\x. x ^ Suc k) x :> (real (Suc k) * x^k)) at_top" by eventually_elim (intro derivative_eq_intros, auto) show "eventually (\x. DERIV exp x :> exp x) at_top" by eventually_elim auto show "eventually (\x. exp x \ 0) at_top" by auto from tendsto_mult[OF tendsto_const Suc, of "real (Suc k)"] show "((\x. real (Suc k) * x ^ k / exp x) \ 0) at_top" by simp qed (rule exp_at_top) qed subsubsection\ A couple of simple bounds\ lemma exp_plus_inverse_exp: fixes x::real shows "2 \ exp x + inverse (exp x)" proof - have "2 \ exp x + exp (-x)" using exp_ge_add_one_self [of x] exp_ge_add_one_self [of "-x"] by linarith then show ?thesis by (simp add: exp_minus) qed lemma real_le_x_sinh: fixes x::real assumes "0 \ x" shows "x \ (exp x - inverse(exp x)) / 2" proof - have *: "exp a - inverse(exp a) - 2*a \ exp b - inverse(exp b) - 2*b" if "a \ b" for a b::real using exp_plus_inverse_exp by (fastforce intro: derivative_eq_intros DERIV_nonneg_imp_nondecreasing [OF that]) show ?thesis using*[OF assms] by simp qed lemma real_le_abs_sinh: fixes x::real shows "abs x \ abs((exp x - inverse(exp x)) / 2)" proof (cases "0 \ x") case True show ?thesis using real_le_x_sinh [OF True] True by (simp add: abs_if) next case False have "-x \ (exp(-x) - inverse(exp(-x))) / 2" by (meson False linear neg_le_0_iff_le real_le_x_sinh) also have "\ \ \(exp x - inverse (exp x)) / 2\" by (metis (no_types, opaque_lifting) abs_divide abs_le_iff abs_minus_cancel add.inverse_inverse exp_minus minus_diff_eq order_refl) finally show ?thesis using False by linarith qed subsection\The general logarithm\ definition log :: "real \ real \ real" \ \logarithm of \<^term>\x\ to base \<^term>\a\\ where "log a x = ln x / ln a" lemma tendsto_log [tendsto_intros]: "(f \ a) F \ (g \ b) F \ 0 < a \ a \ 1 \ 0 < b \ ((\x. log (f x) (g x)) \ log a b) F" unfolding log_def by (intro tendsto_intros) auto lemma continuous_log: assumes "continuous F f" and "continuous F g" and "0 < f (Lim F (\x. x))" and "f (Lim F (\x. x)) \ 1" and "0 < g (Lim F (\x. x))" shows "continuous F (\x. log (f x) (g x))" using assms unfolding continuous_def by (rule tendsto_log) lemma continuous_at_within_log[continuous_intros]: assumes "continuous (at a within s) f" and "continuous (at a within s) g" and "0 < f a" and "f a \ 1" and "0 < g a" shows "continuous (at a within s) (\x. log (f x) (g x))" using assms unfolding continuous_within by (rule tendsto_log) lemma isCont_log[continuous_intros, simp]: assumes "isCont f a" "isCont g a" "0 < f a" "f a \ 1" "0 < g a" shows "isCont (\x. log (f x) (g x)) a" using assms unfolding continuous_at by (rule tendsto_log) lemma continuous_on_log[continuous_intros]: assumes "continuous_on s f" "continuous_on s g" and "\x\s. 0 < f x" "\x\s. f x \ 1" "\x\s. 0 < g x" shows "continuous_on s (\x. log (f x) (g x))" using assms unfolding continuous_on_def by (fast intro: tendsto_log) lemma powr_one_eq_one [simp]: "1 powr a = 1" by (simp add: powr_def) lemma powr_zero_eq_one [simp]: "x powr 0 = (if x = 0 then 0 else 1)" by (simp add: powr_def) lemma powr_one_gt_zero_iff [simp]: "x powr 1 = x \ 0 \ x" for x :: real by (auto simp: powr_def) declare powr_one_gt_zero_iff [THEN iffD2, simp] lemma powr_diff: fixes w:: "'a::{ln,real_normed_field}" shows "w powr (z1 - z2) = w powr z1 / w powr z2" by (simp add: powr_def algebra_simps exp_diff) lemma powr_mult: "0 \ x \ 0 \ y \ (x * y) powr a = (x powr a) * (y powr a)" for a x y :: real by (simp add: powr_def exp_add [symmetric] ln_mult distrib_left) lemma powr_ge_pzero [simp]: "0 \ x powr y" for x y :: real by (simp add: powr_def) lemma powr_non_neg[simp]: "\a powr x < 0" for a x::real using powr_ge_pzero[of a x] by arith lemma inverse_powr: "\y::real. 0 \ y \ inverse y powr a = inverse (y powr a)" by (simp add: exp_minus ln_inverse powr_def) lemma powr_divide: "\0 \ x; 0 \ y\ \ (x / y) powr a = (x powr a) / (y powr a)" for a b x :: real by (simp add: divide_inverse powr_mult inverse_powr) lemma powr_add: "x powr (a + b) = (x powr a) * (x powr b)" for a b x :: "'a::{ln,real_normed_field}" by (simp add: powr_def exp_add [symmetric] distrib_right) lemma powr_mult_base: "0 \ x \x * x powr y = x powr (1 + y)" for x :: real by (auto simp: powr_add) lemma powr_powr: "(x powr a) powr b = x powr (a * b)" for a b x :: real by (simp add: powr_def) lemma powr_powr_swap: "(x powr a) powr b = (x powr b) powr a" for a b x :: real by (simp add: powr_powr mult.commute) lemma powr_minus: "x powr (- a) = inverse (x powr a)" for a x :: "'a::{ln,real_normed_field}" by (simp add: powr_def exp_minus [symmetric]) lemma powr_minus_divide: "x powr (- a) = 1/(x powr a)" for a x :: "'a::{ln,real_normed_field}" by (simp add: divide_inverse powr_minus) lemma divide_powr_uminus: "a / b powr c = a * b powr (- c)" for a b c :: real by (simp add: powr_minus_divide) lemma powr_less_mono: "a < b \ 1 < x \ x powr a < x powr b" for a b x :: real by (simp add: powr_def) lemma powr_less_cancel: "x powr a < x powr b \ 1 < x \ a < b" for a b x :: real by (simp add: powr_def) lemma powr_less_cancel_iff [simp]: "1 < x \ x powr a < x powr b \ a < b" for a b x :: real by (blast intro: powr_less_cancel powr_less_mono) lemma powr_le_cancel_iff [simp]: "1 < x \ x powr a \ x powr b \ a \ b" for a b x :: real by (simp add: linorder_not_less [symmetric]) lemma powr_realpow: "0 < x \ x powr (real n) = x^n" by (induction n) (simp_all add: ac_simps powr_add) lemma powr_real_of_int': assumes "x \ 0" "x \ 0 \ n > 0" shows "x powr real_of_int n = power_int x n" proof (cases "x = 0") case False with assms have "x > 0" by simp show ?thesis proof (cases n rule: int_cases4) case (nonneg n) thus ?thesis using \x > 0\ by (auto simp add: powr_realpow) next case (neg n) thus ?thesis using \x > 0\ by (auto simp add: powr_realpow powr_minus power_int_minus) qed qed (use assms in auto) lemma log_ln: "ln x = log (exp(1)) x" by (simp add: log_def) lemma DERIV_log: assumes "x > 0" shows "DERIV (\y. log b y) x :> 1 / (ln b * x)" proof - define lb where "lb = 1 / ln b" moreover have "DERIV (\y. lb * ln y) x :> lb / x" using \x > 0\ by (auto intro!: derivative_eq_intros) ultimately show ?thesis by (simp add: log_def) qed lemmas DERIV_log[THEN DERIV_chain2, derivative_intros] and DERIV_log[THEN DERIV_chain2, unfolded has_field_derivative_def, derivative_intros] lemma powr_log_cancel [simp]: "0 < a \ a \ 1 \ 0 < x \ a powr (log a x) = x" by (simp add: powr_def log_def) lemma log_powr_cancel [simp]: "0 < a \ a \ 1 \ log a (a powr y) = y" by (simp add: log_def powr_def) lemma log_mult: "0 < a \ a \ 1 \ 0 < x \ 0 < y \ log a (x * y) = log a x + log a y" by (simp add: log_def ln_mult divide_inverse distrib_right) lemma log_eq_div_ln_mult_log: "0 < a \ a \ 1 \ 0 < b \ b \ 1 \ 0 < x \ log a x = (ln b/ln a) * log b x" by (simp add: log_def divide_inverse) text\Base 10 logarithms\ lemma log_base_10_eq1: "0 < x \ log 10 x = (ln (exp 1) / ln 10) * ln x" by (simp add: log_def) lemma log_base_10_eq2: "0 < x \ log 10 x = (log 10 (exp 1)) * ln x" by (simp add: log_def) lemma log_one [simp]: "log a 1 = 0" by (simp add: log_def) lemma log_eq_one [simp]: "0 < a \ a \ 1 \ log a a = 1" by (simp add: log_def) lemma log_inverse: "0 < a \ a \ 1 \ 0 < x \ log a (inverse x) = - log a x" using ln_inverse log_def by auto lemma log_divide: "0 < a \ a \ 1 \ 0 < x \ 0 < y \ log a (x/y) = log a x - log a y" by (simp add: log_mult divide_inverse log_inverse) lemma powr_gt_zero [simp]: "0 < x powr a \ x \ 0" for a x :: real by (simp add: powr_def) lemma powr_nonneg_iff[simp]: "a powr x \ 0 \ a = 0" for a x::real by (meson not_less powr_gt_zero) lemma log_add_eq_powr: "0 < b \ b \ 1 \ 0 < x \ log b x + y = log b (x * b powr y)" and add_log_eq_powr: "0 < b \ b \ 1 \ 0 < x \ y + log b x = log b (b powr y * x)" and log_minus_eq_powr: "0 < b \ b \ 1 \ 0 < x \ log b x - y = log b (x * b powr -y)" and minus_log_eq_powr: "0 < b \ b \ 1 \ 0 < x \ y - log b x = log b (b powr y / x)" by (simp_all add: log_mult log_divide) lemma log_less_cancel_iff [simp]: "1 < a \ 0 < x \ 0 < y \ log a x < log a y \ x < y" using powr_less_cancel_iff [of a] powr_log_cancel [of a x] powr_log_cancel [of a y] by (metis less_eq_real_def less_trans not_le zero_less_one) lemma log_inj: assumes "1 < b" shows "inj_on (log b) {0 <..}" proof (rule inj_onI, simp) fix x y assume pos: "0 < x" "0 < y" and *: "log b x = log b y" show "x = y" proof (cases rule: linorder_cases) assume "x = y" then show ?thesis by simp next assume "x < y" then have "log b x < log b y" using log_less_cancel_iff[OF \1 < b\] pos by simp then show ?thesis using * by simp next assume "y < x" then have "log b y < log b x" using log_less_cancel_iff[OF \1 < b\] pos by simp then show ?thesis using * by simp qed qed lemma log_le_cancel_iff [simp]: "1 < a \ 0 < x \ 0 < y \ log a x \ log a y \ x \ y" by (simp add: linorder_not_less [symmetric]) lemma zero_less_log_cancel_iff[simp]: "1 < a \ 0 < x \ 0 < log a x \ 1 < x" using log_less_cancel_iff[of a 1 x] by simp lemma zero_le_log_cancel_iff[simp]: "1 < a \ 0 < x \ 0 \ log a x \ 1 \ x" using log_le_cancel_iff[of a 1 x] by simp lemma log_less_zero_cancel_iff[simp]: "1 < a \ 0 < x \ log a x < 0 \ x < 1" using log_less_cancel_iff[of a x 1] by simp lemma log_le_zero_cancel_iff[simp]: "1 < a \ 0 < x \ log a x \ 0 \ x \ 1" using log_le_cancel_iff[of a x 1] by simp lemma one_less_log_cancel_iff[simp]: "1 < a \ 0 < x \ 1 < log a x \ a < x" using log_less_cancel_iff[of a a x] by simp lemma one_le_log_cancel_iff[simp]: "1 < a \ 0 < x \ 1 \ log a x \ a \ x" using log_le_cancel_iff[of a a x] by simp lemma log_less_one_cancel_iff[simp]: "1 < a \ 0 < x \ log a x < 1 \ x < a" using log_less_cancel_iff[of a x a] by simp lemma log_le_one_cancel_iff[simp]: "1 < a \ 0 < x \ log a x \ 1 \ x \ a" using log_le_cancel_iff[of a x a] by simp lemma le_log_iff: fixes b x y :: real assumes "1 < b" "x > 0" shows "y \ log b x \ b powr y \ x" using assms by (metis less_irrefl less_trans powr_le_cancel_iff powr_log_cancel zero_less_one) lemma less_log_iff: assumes "1 < b" "x > 0" shows "y < log b x \ b powr y < x" by (metis assms dual_order.strict_trans less_irrefl powr_less_cancel_iff powr_log_cancel zero_less_one) lemma assumes "1 < b" "x > 0" shows log_less_iff: "log b x < y \ x < b powr y" and log_le_iff: "log b x \ y \ x \ b powr y" using le_log_iff[OF assms, of y] less_log_iff[OF assms, of y] by auto lemmas powr_le_iff = le_log_iff[symmetric] and powr_less_iff = less_log_iff[symmetric] and less_powr_iff = log_less_iff[symmetric] and le_powr_iff = log_le_iff[symmetric] lemma le_log_of_power: assumes "b ^ n \ m" "1 < b" shows "n \ log b m" proof - from assms have "0 < m" by (metis less_trans zero_less_power less_le_trans zero_less_one) thus ?thesis using assms by (simp add: le_log_iff powr_realpow) qed lemma le_log2_of_power: "2 ^ n \ m \ n \ log 2 m" for m n :: nat using le_log_of_power[of 2] by simp lemma log_of_power_le: "\ m \ b ^ n; b > 1; m > 0 \ \ log b (real m) \ n" by (simp add: log_le_iff powr_realpow) lemma log2_of_power_le: "\ m \ 2 ^ n; m > 0 \ \ log 2 m \ n" for m n :: nat using log_of_power_le[of _ 2] by simp lemma log_of_power_less: "\ m < b ^ n; b > 1; m > 0 \ \ log b (real m) < n" by (simp add: log_less_iff powr_realpow) lemma log2_of_power_less: "\ m < 2 ^ n; m > 0 \ \ log 2 m < n" for m n :: nat using log_of_power_less[of _ 2] by simp lemma less_log_of_power: assumes "b ^ n < m" "1 < b" shows "n < log b m" proof - have "0 < m" by (metis assms less_trans zero_less_power zero_less_one) thus ?thesis using assms by (simp add: less_log_iff powr_realpow) qed lemma less_log2_of_power: "2 ^ n < m \ n < log 2 m" for m n :: nat using less_log_of_power[of 2] by simp lemma gr_one_powr[simp]: fixes x y :: real shows "\ x > 1; y > 0 \ \ 1 < x powr y" by(simp add: less_powr_iff) lemma log_pow_cancel [simp]: "a > 0 \ a \ 1 \ log a (a ^ b) = b" by (simp add: ln_realpow log_def) lemma floor_log_eq_powr_iff: "x > 0 \ b > 1 \ \log b x\ = k \ b powr k \ x \ x < b powr (k + 1)" by (auto simp: floor_eq_iff powr_le_iff less_powr_iff) lemma floor_log_nat_eq_powr_iff: fixes b n k :: nat shows "\ b \ 2; k > 0 \ \ floor (log b (real k)) = n \ b^n \ k \ k < b^(n+1)" by (auto simp: floor_log_eq_powr_iff powr_add powr_realpow of_nat_power[symmetric] of_nat_mult[symmetric] ac_simps simp del: of_nat_power of_nat_mult) lemma floor_log_nat_eq_if: fixes b n k :: nat assumes "b^n \ k" "k < b^(n+1)" "b \ 2" shows "floor (log b (real k)) = n" proof - have "k \ 1" using assms(1,3) one_le_power[of b n] by linarith with assms show ?thesis by(simp add: floor_log_nat_eq_powr_iff) qed lemma ceiling_log_eq_powr_iff: "\ x > 0; b > 1 \ \ \log b x\ = int k + 1 \ b powr k < x \ x \ b powr (k + 1)" by (auto simp: ceiling_eq_iff powr_less_iff le_powr_iff) lemma ceiling_log_nat_eq_powr_iff: fixes b n k :: nat shows "\ b \ 2; k > 0 \ \ ceiling (log b (real k)) = int n + 1 \ (b^n < k \ k \ b^(n+1))" using ceiling_log_eq_powr_iff by (auto simp: powr_add powr_realpow of_nat_power[symmetric] of_nat_mult[symmetric] ac_simps simp del: of_nat_power of_nat_mult) lemma ceiling_log_nat_eq_if: fixes b n k :: nat assumes "b^n < k" "k \ b^(n+1)" "b \ 2" shows "ceiling (log b (real k)) = int n + 1" proof - have "k \ 1" using assms(1,3) one_le_power[of b n] by linarith with assms show ?thesis by(simp add: ceiling_log_nat_eq_powr_iff) qed lemma floor_log2_div2: fixes n :: nat assumes "n \ 2" shows "floor(log 2 n) = floor(log 2 (n div 2)) + 1" proof cases assume "n=2" thus ?thesis by simp next let ?m = "n div 2" assume "n\2" hence "1 \ ?m" using assms by arith then obtain i where i: "2 ^ i \ ?m" "?m < 2 ^ (i + 1)" using ex_power_ivl1[of 2 ?m] by auto have "2^(i+1) \ 2*?m" using i(1) by simp also have "2*?m \ n" by arith finally have *: "2^(i+1) \ \" . have "n < 2^(i+1+1)" using i(2) by simp from floor_log_nat_eq_if[OF * this] floor_log_nat_eq_if[OF i] show ?thesis by simp qed lemma ceiling_log2_div2: assumes "n \ 2" shows "ceiling(log 2 (real n)) = ceiling(log 2 ((n-1) div 2 + 1)) + 1" proof cases assume "n=2" thus ?thesis by simp next let ?m = "(n-1) div 2 + 1" assume "n\2" hence "2 \ ?m" using assms by arith then obtain i where i: "2 ^ i < ?m" "?m \ 2 ^ (i + 1)" using ex_power_ivl2[of 2 ?m] by auto have "n \ 2*?m" by arith also have "2*?m \ 2 ^ ((i+1)+1)" using i(2) by simp finally have *: "n \ \" . have "2^(i+1) < n" using i(1) by (auto simp: less_Suc_eq_0_disj) from ceiling_log_nat_eq_if[OF this *] ceiling_log_nat_eq_if[OF i] show ?thesis by simp qed lemma powr_real_of_int: "x > 0 \ x powr real_of_int n = (if n \ 0 then x ^ nat n else inverse (x ^ nat (- n)))" using powr_realpow[of x "nat n"] powr_realpow[of x "nat (-n)"] by (auto simp: field_simps powr_minus) lemma powr_numeral [simp]: "0 \ x \ x powr (numeral n :: real) = x ^ (numeral n)" by (metis less_le power_zero_numeral powr_0 of_nat_numeral powr_realpow) lemma powr_int: assumes "x > 0" shows "x powr i = (if i \ 0 then x ^ nat i else 1 / x ^ nat (-i))" proof (cases "i < 0") case True have r: "x powr i = 1 / x powr (- i)" by (simp add: powr_minus field_simps) show ?thesis using \i < 0\ \x > 0\ by (simp add: r field_simps powr_realpow[symmetric]) next case False then show ?thesis by (simp add: assms powr_realpow[symmetric]) qed definition powr_real :: "real \ real \ real" where [code_abbrev, simp]: "powr_real = Transcendental.powr" lemma compute_powr_real [code]: "powr_real b i = (if b \ 0 then Code.abort (STR ''powr_real with nonpositive base'') (\_. powr_real b i) else if \i\ = i then (if 0 \ i then b ^ nat \i\ else 1 / b ^ nat \- i\) else Code.abort (STR ''powr_real with non-integer exponent'') (\_. powr_real b i))" for b i :: real by (auto simp: powr_int) lemma powr_one: "0 \ x \ x powr 1 = x" for x :: real using powr_realpow [of x 1] by simp lemma powr_neg_one: "0 < x \ x powr - 1 = 1 / x" for x :: real using powr_int [of x "- 1"] by simp lemma powr_neg_numeral: "0 < x \ x powr - numeral n = 1 / x ^ numeral n" for x :: real using powr_int [of x "- numeral n"] by simp lemma root_powr_inverse: "0 < n \ 0 < x \ root n x = x powr (1/n)" by (rule real_root_pos_unique) (auto simp: powr_realpow[symmetric] powr_powr) lemma ln_powr: "x \ 0 \ ln (x powr y) = y * ln x" for x :: real by (simp add: powr_def) lemma ln_root: "n > 0 \ b > 0 \ ln (root n b) = ln b / n" by (simp add: root_powr_inverse ln_powr) lemma ln_sqrt: "0 < x \ ln (sqrt x) = ln x / 2" by (simp add: ln_powr ln_powr[symmetric] mult.commute) lemma log_root: "n > 0 \ a > 0 \ log b (root n a) = log b a / n" by (simp add: log_def ln_root) lemma log_powr: "x \ 0 \ log b (x powr y) = y * log b x" by (simp add: log_def ln_powr) (* [simp] is not worth it, interferes with some proofs *) lemma log_nat_power: "0 < x \ log b (x^n) = real n * log b x" by (simp add: log_powr powr_realpow [symmetric]) lemma log_of_power_eq: assumes "m = b ^ n" "b > 1" shows "n = log b (real m)" proof - have "n = log b (b ^ n)" using assms(2) by (simp add: log_nat_power) also have "\ = log b m" using assms by simp finally show ?thesis . qed lemma log2_of_power_eq: "m = 2 ^ n \ n = log 2 m" for m n :: nat using log_of_power_eq[of _ 2] by simp lemma log_base_change: "0 < a \ a \ 1 \ log b x = log a x / log a b" by (simp add: log_def) lemma log_base_pow: "0 < a \ log (a ^ n) x = log a x / n" by (simp add: log_def ln_realpow) lemma log_base_powr: "a \ 0 \ log (a powr b) x = log a x / b" by (simp add: log_def ln_powr) lemma log_base_root: "n > 0 \ b > 0 \ log (root n b) x = n * (log b x)" by (simp add: log_def ln_root) lemma ln_bound: "0 < x \ ln x \ x" for x :: real using ln_le_minus_one by force lemma powr_mono: fixes x :: real assumes "a \ b" and "1 \ x" shows "x powr a \ x powr b" using assms less_eq_real_def by auto lemma ge_one_powr_ge_zero: "1 \ x \ 0 \ a \ 1 \ x powr a" for x :: real using powr_mono by fastforce lemma powr_less_mono2: "0 < a \ 0 \ x \ x < y \ x powr a < y powr a" for x :: real by (simp add: powr_def) lemma powr_less_mono2_neg: "a < 0 \ 0 < x \ x < y \ y powr a < x powr a" for x :: real by (simp add: powr_def) lemma powr_mono2: "x powr a \ y powr a" if "0 \ a" "0 \ x" "x \ y" for x :: real using less_eq_real_def powr_less_mono2 that by auto lemma powr_le1: "0 \ a \ 0 \ x \ x \ 1 \ x powr a \ 1" for x :: real using powr_mono2 by fastforce lemma powr_mono2': fixes a x y :: real assumes "a \ 0" "x > 0" "x \ y" shows "x powr a \ y powr a" proof - from assms have "x powr - a \ y powr - a" by (intro powr_mono2) simp_all with assms show ?thesis by (auto simp: powr_minus field_simps) qed lemma powr_mono_both: fixes x :: real assumes "0 \ a" "a \ b" "1 \ x" "x \ y" shows "x powr a \ y powr b" by (meson assms order.trans powr_mono powr_mono2 zero_le_one) lemma powr_inj: "0 < a \ a \ 1 \ a powr x = a powr y \ x = y" for x :: real unfolding powr_def exp_inj_iff by simp lemma powr_half_sqrt: "0 \ x \ x powr (1/2) = sqrt x" by (simp add: powr_def root_powr_inverse sqrt_def) lemma square_powr_half [simp]: fixes x::real shows "x\<^sup>2 powr (1/2) = \x\" by (simp add: powr_half_sqrt) lemma ln_powr_bound: "1 \ x \ 0 < a \ ln x \ (x powr a) / a" for x :: real by (metis exp_gt_zero linear ln_eq_zero_iff ln_exp ln_less_self ln_powr mult.commute mult_imp_le_div_pos not_less powr_gt_zero) lemma ln_powr_bound2: fixes x :: real assumes "1 < x" and "0 < a" shows "(ln x) powr a \ (a powr a) * x" proof - from assms have "ln x \ (x powr (1 / a)) / (1 / a)" by (metis less_eq_real_def ln_powr_bound zero_less_divide_1_iff) also have "\ = a * (x powr (1 / a))" by simp finally have "(ln x) powr a \ (a * (x powr (1 / a))) powr a" by (metis assms less_imp_le ln_gt_zero powr_mono2) also have "\ = (a powr a) * ((x powr (1 / a)) powr a)" using assms powr_mult by auto also have "(x powr (1 / a)) powr a = x powr ((1 / a) * a)" by (rule powr_powr) also have "\ = x" using assms by auto finally show ?thesis . qed lemma tendsto_powr: fixes a b :: real assumes f: "(f \ a) F" and g: "(g \ b) F" and a: "a \ 0" shows "((\x. f x powr g x) \ a powr b) F" unfolding powr_def proof (rule filterlim_If) from f show "((\x. 0) \ (if a = 0 then 0 else exp (b * ln a))) (inf F (principal {x. f x = 0}))" by simp (auto simp: filterlim_iff eventually_inf_principal elim: eventually_mono dest: t1_space_nhds) from f g a show "((\x. exp (g x * ln (f x))) \ (if a = 0 then 0 else exp (b * ln a))) (inf F (principal {x. f x \ 0}))" by (auto intro!: tendsto_intros intro: tendsto_mono inf_le1) qed lemma tendsto_powr'[tendsto_intros]: fixes a :: real assumes f: "(f \ a) F" and g: "(g \ b) F" and a: "a \ 0 \ (b > 0 \ eventually (\x. f x \ 0) F)" shows "((\x. f x powr g x) \ a powr b) F" proof - from a consider "a \ 0" | "a = 0" "b > 0" "eventually (\x. f x \ 0) F" by auto then show ?thesis proof cases case 1 with f g show ?thesis by (rule tendsto_powr) next case 2 have "((\x. if f x = 0 then 0 else exp (g x * ln (f x))) \ 0) F" proof (intro filterlim_If) have "filterlim f (principal {0<..}) (inf F (principal {z. f z \ 0}))" using \eventually (\x. f x \ 0) F\ by (auto simp: filterlim_iff eventually_inf_principal eventually_principal elim: eventually_mono) moreover have "filterlim f (nhds a) (inf F (principal {z. f z \ 0}))" by (rule tendsto_mono[OF _ f]) simp_all ultimately have f: "filterlim f (at_right 0) (inf F (principal {x. f x \ 0}))" by (simp add: at_within_def filterlim_inf \a = 0\) have g: "(g \ b) (inf F (principal {z. f z \ 0}))" by (rule tendsto_mono[OF _ g]) simp_all show "((\x. exp (g x * ln (f x))) \ 0) (inf F (principal {x. f x \ 0}))" by (rule filterlim_compose[OF exp_at_bot] filterlim_tendsto_pos_mult_at_bot filterlim_compose[OF ln_at_0] f g \b > 0\)+ qed simp_all with \a = 0\ show ?thesis by (simp add: powr_def) qed qed lemma continuous_powr: assumes "continuous F f" and "continuous F g" and "f (Lim F (\x. x)) \ 0" shows "continuous F (\x. (f x) powr (g x :: real))" using assms unfolding continuous_def by (rule tendsto_powr) lemma continuous_at_within_powr[continuous_intros]: fixes f g :: "_ \ real" assumes "continuous (at a within s) f" and "continuous (at a within s) g" and "f a \ 0" shows "continuous (at a within s) (\x. (f x) powr (g x))" using assms unfolding continuous_within by (rule tendsto_powr) lemma isCont_powr[continuous_intros, simp]: fixes f g :: "_ \ real" assumes "isCont f a" "isCont g a" "f a \ 0" shows "isCont (\x. (f x) powr g x) a" using assms unfolding continuous_at by (rule tendsto_powr) lemma continuous_on_powr[continuous_intros]: fixes f g :: "_ \ real" assumes "continuous_on s f" "continuous_on s g" and "\x\s. f x \ 0" shows "continuous_on s (\x. (f x) powr (g x))" using assms unfolding continuous_on_def by (fast intro: tendsto_powr) lemma tendsto_powr2: fixes a :: real assumes f: "(f \ a) F" and g: "(g \ b) F" and "\\<^sub>F x in F. 0 \ f x" and b: "0 < b" shows "((\x. f x powr g x) \ a powr b) F" using tendsto_powr'[of f a F g b] assms by auto lemma has_derivative_powr[derivative_intros]: assumes g[derivative_intros]: "(g has_derivative g') (at x within X)" and f[derivative_intros]:"(f has_derivative f') (at x within X)" assumes pos: "0 < g x" and "x \ X" shows "((\x. g x powr f x::real) has_derivative (\h. (g x powr f x) * (f' h * ln (g x) + g' h * f x / g x))) (at x within X)" proof - have "\\<^sub>F x in at x within X. g x > 0" by (rule order_tendstoD[OF _ pos]) (rule has_derivative_continuous[OF g, unfolded continuous_within]) then obtain d where "d > 0" and pos': "\x'. x' \ X \ dist x' x < d \ 0 < g x'" using pos unfolding eventually_at by force have "((\x. exp (f x * ln (g x))) has_derivative (\h. (g x powr f x) * (f' h * ln (g x) + g' h * f x / g x))) (at x within X)" using pos by (auto intro!: derivative_eq_intros simp: field_split_simps powr_def) then show ?thesis by (rule has_derivative_transform_within[OF _ \d > 0\ \x \ X\]) (auto simp: powr_def dest: pos') qed lemma DERIV_powr: fixes r :: real assumes g: "DERIV g x :> m" and pos: "g x > 0" and f: "DERIV f x :> r" shows "DERIV (\x. g x powr f x) x :> (g x powr f x) * (r * ln (g x) + m * f x / g x)" using assms by (auto intro!: derivative_eq_intros ext simp: has_field_derivative_def algebra_simps) lemma DERIV_fun_powr: fixes r :: real assumes g: "DERIV g x :> m" and pos: "g x > 0" shows "DERIV (\x. (g x) powr r) x :> r * (g x) powr (r - of_nat 1) * m" using DERIV_powr[OF g pos DERIV_const, of r] pos by (simp add: powr_diff field_simps) lemma has_real_derivative_powr: assumes "z > 0" shows "((\z. z powr r) has_real_derivative r * z powr (r - 1)) (at z)" proof (subst DERIV_cong_ev[OF refl _ refl]) from assms have "eventually (\z. z \ 0) (nhds z)" by (intro t1_space_nhds) auto then show "eventually (\z. z powr r = exp (r * ln z)) (nhds z)" unfolding powr_def by eventually_elim simp from assms show "((\z. exp (r * ln z)) has_real_derivative r * z powr (r - 1)) (at z)" by (auto intro!: derivative_eq_intros simp: powr_def field_simps exp_diff) qed declare has_real_derivative_powr[THEN DERIV_chain2, derivative_intros] lemma tendsto_zero_powrI: assumes "(f \ (0::real)) F" "(g \ b) F" "\\<^sub>F x in F. 0 \ f x" "0 < b" shows "((\x. f x powr g x) \ 0) F" using tendsto_powr2[OF assms] by simp lemma continuous_on_powr': fixes f g :: "_ \ real" assumes "continuous_on s f" "continuous_on s g" and "\x\s. f x \ 0 \ (f x = 0 \ g x > 0)" shows "continuous_on s (\x. (f x) powr (g x))" unfolding continuous_on_def proof fix x assume x: "x \ s" from assms x show "((\x. f x powr g x) \ f x powr g x) (at x within s)" proof (cases "f x = 0") case True from assms(3) have "eventually (\x. f x \ 0) (at x within s)" by (auto simp: at_within_def eventually_inf_principal) with True x assms show ?thesis by (auto intro!: tendsto_zero_powrI[of f _ g "g x"] simp: continuous_on_def) next case False with assms x show ?thesis by (auto intro!: tendsto_powr' simp: continuous_on_def) qed qed lemma tendsto_neg_powr: assumes "s < 0" and f: "LIM x F. f x :> at_top" shows "((\x. f x powr s) \ (0::real)) F" proof - have "((\x. exp (s * ln (f x))) \ (0::real)) F" (is "?X") by (auto intro!: filterlim_compose[OF exp_at_bot] filterlim_compose[OF ln_at_top] filterlim_tendsto_neg_mult_at_bot assms) also have "?X \ ((\x. f x powr s) \ (0::real)) F" using f filterlim_at_top_dense[of f F] by (intro filterlim_cong[OF refl refl]) (auto simp: neq_iff powr_def elim: eventually_mono) finally show ?thesis . qed lemma tendsto_exp_limit_at_right: "((\y. (1 + x * y) powr (1 / y)) \ exp x) (at_right 0)" for x :: real proof (cases "x = 0") case True then show ?thesis by simp next case False have "((\y. ln (1 + x * y)::real) has_real_derivative 1 * x) (at 0)" by (auto intro!: derivative_eq_intros) then have "((\y. ln (1 + x * y) / y) \ x) (at 0)" by (auto simp: has_field_derivative_def field_has_derivative_at) then have *: "((\y. exp (ln (1 + x * y) / y)) \ exp x) (at 0)" by (rule tendsto_intros) then show ?thesis proof (rule filterlim_mono_eventually) show "eventually (\xa. exp (ln (1 + x * xa) / xa) = (1 + x * xa) powr (1 / xa)) (at_right 0)" unfolding eventually_at_right[OF zero_less_one] using False by (intro exI[of _ "1 / \x\"]) (auto simp: field_simps powr_def abs_if add_nonneg_eq_0_iff) qed (simp_all add: at_eq_sup_left_right) qed lemma tendsto_exp_limit_at_top: "((\y. (1 + x / y) powr y) \ exp x) at_top" for x :: real by (simp add: filterlim_at_top_to_right inverse_eq_divide tendsto_exp_limit_at_right) lemma tendsto_exp_limit_sequentially: "(\n. (1 + x / n) ^ n) \ exp x" for x :: real proof (rule filterlim_mono_eventually) from reals_Archimedean2 [of "\x\"] obtain n :: nat where *: "real n > \x\" .. then have "eventually (\n :: nat. 0 < 1 + x / real n) at_top" by (intro eventually_sequentiallyI [of n]) (auto simp: field_split_simps) then show "eventually (\n. (1 + x / n) powr n = (1 + x / n) ^ n) at_top" by (rule eventually_mono) (erule powr_realpow) show "(\n. (1 + x / real n) powr real n) \ exp x" by (rule filterlim_compose [OF tendsto_exp_limit_at_top filterlim_real_sequentially]) qed auto subsection \Sine and Cosine\ definition sin_coeff :: "nat \ real" where "sin_coeff = (\n. if even n then 0 else (- 1) ^ ((n - Suc 0) div 2) / (fact n))" definition cos_coeff :: "nat \ real" where "cos_coeff = (\n. if even n then ((- 1) ^ (n div 2)) / (fact n) else 0)" definition sin :: "'a \ 'a::{real_normed_algebra_1,banach}" where "sin = (\x. \n. sin_coeff n *\<^sub>R x^n)" definition cos :: "'a \ 'a::{real_normed_algebra_1,banach}" where "cos = (\x. \n. cos_coeff n *\<^sub>R x^n)" lemma sin_coeff_0 [simp]: "sin_coeff 0 = 0" unfolding sin_coeff_def by simp lemma cos_coeff_0 [simp]: "cos_coeff 0 = 1" unfolding cos_coeff_def by simp lemma sin_coeff_Suc: "sin_coeff (Suc n) = cos_coeff n / real (Suc n)" unfolding cos_coeff_def sin_coeff_def by (simp del: mult_Suc) lemma cos_coeff_Suc: "cos_coeff (Suc n) = - sin_coeff n / real (Suc n)" unfolding cos_coeff_def sin_coeff_def by (simp del: mult_Suc) (auto elim: oddE) lemma summable_norm_sin: "summable (\n. norm (sin_coeff n *\<^sub>R x^n))" for x :: "'a::{real_normed_algebra_1,banach}" proof (rule summable_comparison_test [OF _ summable_norm_exp]) show "\N. \n\N. norm (norm (sin_coeff n *\<^sub>R x ^ n)) \ norm (x ^ n /\<^sub>R fact n)" unfolding sin_coeff_def by (auto simp: divide_inverse abs_mult power_abs [symmetric] zero_le_mult_iff) qed lemma summable_norm_cos: "summable (\n. norm (cos_coeff n *\<^sub>R x^n))" for x :: "'a::{real_normed_algebra_1,banach}" proof (rule summable_comparison_test [OF _ summable_norm_exp]) show "\N. \n\N. norm (norm (cos_coeff n *\<^sub>R x ^ n)) \ norm (x ^ n /\<^sub>R fact n)" unfolding cos_coeff_def by (auto simp: divide_inverse abs_mult power_abs [symmetric] zero_le_mult_iff) qed lemma sin_converges: "(\n. sin_coeff n *\<^sub>R x^n) sums sin x" unfolding sin_def by (metis (full_types) summable_norm_cancel summable_norm_sin summable_sums) lemma cos_converges: "(\n. cos_coeff n *\<^sub>R x^n) sums cos x" unfolding cos_def by (metis (full_types) summable_norm_cancel summable_norm_cos summable_sums) lemma sin_of_real: "sin (of_real x) = of_real (sin x)" for x :: real proof - have "(\n. of_real (sin_coeff n *\<^sub>R x^n)) = (\n. sin_coeff n *\<^sub>R (of_real x)^n)" proof show "of_real (sin_coeff n *\<^sub>R x^n) = sin_coeff n *\<^sub>R of_real x^n" for n by (simp add: scaleR_conv_of_real) qed also have "\ sums (sin (of_real x))" by (rule sin_converges) finally have "(\n. of_real (sin_coeff n *\<^sub>R x^n)) sums (sin (of_real x))" . then show ?thesis using sums_unique2 sums_of_real [OF sin_converges] by blast qed corollary sin_in_Reals [simp]: "z \ \ \ sin z \ \" by (metis Reals_cases Reals_of_real sin_of_real) lemma cos_of_real: "cos (of_real x) = of_real (cos x)" for x :: real proof - have "(\n. of_real (cos_coeff n *\<^sub>R x^n)) = (\n. cos_coeff n *\<^sub>R (of_real x)^n)" proof show "of_real (cos_coeff n *\<^sub>R x^n) = cos_coeff n *\<^sub>R of_real x^n" for n by (simp add: scaleR_conv_of_real) qed also have "\ sums (cos (of_real x))" by (rule cos_converges) finally have "(\n. of_real (cos_coeff n *\<^sub>R x^n)) sums (cos (of_real x))" . then show ?thesis using sums_unique2 sums_of_real [OF cos_converges] by blast qed corollary cos_in_Reals [simp]: "z \ \ \ cos z \ \" by (metis Reals_cases Reals_of_real cos_of_real) lemma diffs_sin_coeff: "diffs sin_coeff = cos_coeff" by (simp add: diffs_def sin_coeff_Suc del: of_nat_Suc) lemma diffs_cos_coeff: "diffs cos_coeff = (\n. - sin_coeff n)" by (simp add: diffs_def cos_coeff_Suc del: of_nat_Suc) lemma sin_int_times_real: "sin (of_int m * of_real x) = of_real (sin (of_int m * x))" by (metis sin_of_real of_real_mult of_real_of_int_eq) lemma cos_int_times_real: "cos (of_int m * of_real x) = of_real (cos (of_int m * x))" by (metis cos_of_real of_real_mult of_real_of_int_eq) text \Now at last we can get the derivatives of exp, sin and cos.\ lemma DERIV_sin [simp]: "DERIV sin x :> cos x" for x :: "'a::{real_normed_field,banach}" unfolding sin_def cos_def scaleR_conv_of_real apply (rule DERIV_cong) apply (rule termdiffs [where K="of_real (norm x) + 1 :: 'a"]) apply (simp_all add: norm_less_p1 diffs_of_real diffs_sin_coeff diffs_cos_coeff summable_minus_iff scaleR_conv_of_real [symmetric] summable_norm_sin [THEN summable_norm_cancel] summable_norm_cos [THEN summable_norm_cancel]) done declare DERIV_sin[THEN DERIV_chain2, derivative_intros] and DERIV_sin[THEN DERIV_chain2, unfolded has_field_derivative_def, derivative_intros] lemmas has_derivative_sin[derivative_intros] = DERIV_sin[THEN DERIV_compose_FDERIV] lemma DERIV_cos [simp]: "DERIV cos x :> - sin x" for x :: "'a::{real_normed_field,banach}" unfolding sin_def cos_def scaleR_conv_of_real apply (rule DERIV_cong) apply (rule termdiffs [where K="of_real (norm x) + 1 :: 'a"]) apply (simp_all add: norm_less_p1 diffs_of_real diffs_minus suminf_minus diffs_sin_coeff diffs_cos_coeff summable_minus_iff scaleR_conv_of_real [symmetric] summable_norm_sin [THEN summable_norm_cancel] summable_norm_cos [THEN summable_norm_cancel]) done declare DERIV_cos[THEN DERIV_chain2, derivative_intros] and DERIV_cos[THEN DERIV_chain2, unfolded has_field_derivative_def, derivative_intros] lemmas has_derivative_cos[derivative_intros] = DERIV_cos[THEN DERIV_compose_FDERIV] lemma isCont_sin: "isCont sin x" for x :: "'a::{real_normed_field,banach}" by (rule DERIV_sin [THEN DERIV_isCont]) lemma continuous_on_sin_real: "continuous_on {a..b} sin" for a::real using continuous_at_imp_continuous_on isCont_sin by blast lemma isCont_cos: "isCont cos x" for x :: "'a::{real_normed_field,banach}" by (rule DERIV_cos [THEN DERIV_isCont]) lemma continuous_on_cos_real: "continuous_on {a..b} cos" for a::real using continuous_at_imp_continuous_on isCont_cos by blast context fixes f :: "'a::t2_space \ 'b::{real_normed_field,banach}" begin lemma isCont_sin' [simp]: "isCont f a \ isCont (\x. sin (f x)) a" by (rule isCont_o2 [OF _ isCont_sin]) lemma isCont_cos' [simp]: "isCont f a \ isCont (\x. cos (f x)) a" by (rule isCont_o2 [OF _ isCont_cos]) lemma tendsto_sin [tendsto_intros]: "(f \ a) F \ ((\x. sin (f x)) \ sin a) F" by (rule isCont_tendsto_compose [OF isCont_sin]) lemma tendsto_cos [tendsto_intros]: "(f \ a) F \ ((\x. cos (f x)) \ cos a) F" by (rule isCont_tendsto_compose [OF isCont_cos]) lemma continuous_sin [continuous_intros]: "continuous F f \ continuous F (\x. sin (f x))" unfolding continuous_def by (rule tendsto_sin) lemma continuous_on_sin [continuous_intros]: "continuous_on s f \ continuous_on s (\x. sin (f x))" unfolding continuous_on_def by (auto intro: tendsto_sin) lemma continuous_cos [continuous_intros]: "continuous F f \ continuous F (\x. cos (f x))" unfolding continuous_def by (rule tendsto_cos) lemma continuous_on_cos [continuous_intros]: "continuous_on s f \ continuous_on s (\x. cos (f x))" unfolding continuous_on_def by (auto intro: tendsto_cos) end lemma continuous_within_sin: "continuous (at z within s) sin" for z :: "'a::{real_normed_field,banach}" by (simp add: continuous_within tendsto_sin) lemma continuous_within_cos: "continuous (at z within s) cos" for z :: "'a::{real_normed_field,banach}" by (simp add: continuous_within tendsto_cos) subsection \Properties of Sine and Cosine\ lemma sin_zero [simp]: "sin 0 = 0" by (simp add: sin_def sin_coeff_def scaleR_conv_of_real) lemma cos_zero [simp]: "cos 0 = 1" by (simp add: cos_def cos_coeff_def scaleR_conv_of_real) lemma DERIV_fun_sin: "DERIV g x :> m \ DERIV (\x. sin (g x)) x :> cos (g x) * m" by (fact derivative_intros) lemma DERIV_fun_cos: "DERIV g x :> m \ DERIV (\x. cos(g x)) x :> - sin (g x) * m" by (fact derivative_intros) subsection \Deriving the Addition Formulas\ text \The product of two cosine series.\ lemma cos_x_cos_y: fixes x :: "'a::{real_normed_field,banach}" shows "(\p. \n\p. if even p \ even n then ((-1) ^ (p div 2) * (p choose n) / (fact p)) *\<^sub>R (x^n) * y^(p-n) else 0) sums (cos x * cos y)" proof - have "(cos_coeff n * cos_coeff (p - n)) *\<^sub>R (x^n * y^(p - n)) = (if even p \ even n then ((-1) ^ (p div 2) * (p choose n) / (fact p)) *\<^sub>R (x^n) * y^(p - n) else 0)" if "n \ p" for n p :: nat proof - from that have *: "even n \ even p \ (-1) ^ (n div 2) * (-1) ^ ((p - n) div 2) = (-1 :: real) ^ (p div 2)" by (metis div_add power_add le_add_diff_inverse odd_add) with that show ?thesis by (auto simp: algebra_simps cos_coeff_def binomial_fact) qed then have "(\p. \n\p. if even p \ even n then ((-1) ^ (p div 2) * (p choose n) / (fact p)) *\<^sub>R (x^n) * y^(p-n) else 0) = (\p. \n\p. (cos_coeff n * cos_coeff (p - n)) *\<^sub>R (x^n * y^(p-n)))" by simp also have "\ = (\p. \n\p. (cos_coeff n *\<^sub>R x^n) * (cos_coeff (p - n) *\<^sub>R y^(p-n)))" by (simp add: algebra_simps) also have "\ sums (cos x * cos y)" using summable_norm_cos by (auto simp: cos_def scaleR_conv_of_real intro!: Cauchy_product_sums) finally show ?thesis . qed text \The product of two sine series.\ lemma sin_x_sin_y: fixes x :: "'a::{real_normed_field,banach}" shows "(\p. \n\p. if even p \ odd n then - ((-1) ^ (p div 2) * (p choose n) / (fact p)) *\<^sub>R (x^n) * y^(p-n) else 0) sums (sin x * sin y)" proof - have "(sin_coeff n * sin_coeff (p - n)) *\<^sub>R (x^n * y^(p-n)) = (if even p \ odd n then -((-1) ^ (p div 2) * (p choose n) / (fact p)) *\<^sub>R (x^n) * y^(p-n) else 0)" if "n \ p" for n p :: nat proof - have "(-1) ^ ((n - Suc 0) div 2) * (-1) ^ ((p - Suc n) div 2) = - ((-1 :: real) ^ (p div 2))" if np: "odd n" "even p" proof - have "p > 0" using \n \ p\ neq0_conv that(1) by blast then have \
: "(- 1::real) ^ (p div 2 - Suc 0) = - ((- 1) ^ (p div 2))" using \even p\ by (auto simp add: dvd_def power_eq_if) from \n \ p\ np have *: "n - Suc 0 + (p - Suc n) = p - Suc (Suc 0)" "Suc (Suc 0) \ p" by arith+ have "(p - Suc (Suc 0)) div 2 = p div 2 - Suc 0" by simp with \n \ p\ np \
* show ?thesis by (simp add: flip: div_add power_add) qed then show ?thesis using \n\p\ by (auto simp: algebra_simps sin_coeff_def binomial_fact) qed then have "(\p. \n\p. if even p \ odd n then - ((-1) ^ (p div 2) * (p choose n) / (fact p)) *\<^sub>R (x^n) * y^(p-n) else 0) = (\p. \n\p. (sin_coeff n * sin_coeff (p - n)) *\<^sub>R (x^n * y^(p-n)))" by simp also have "\ = (\p. \n\p. (sin_coeff n *\<^sub>R x^n) * (sin_coeff (p - n) *\<^sub>R y^(p-n)))" by (simp add: algebra_simps) also have "\ sums (sin x * sin y)" using summable_norm_sin by (auto simp: sin_def scaleR_conv_of_real intro!: Cauchy_product_sums) finally show ?thesis . qed lemma sums_cos_x_plus_y: fixes x :: "'a::{real_normed_field,banach}" shows "(\p. \n\p. if even p then ((-1) ^ (p div 2) * (p choose n) / (fact p)) *\<^sub>R (x^n) * y^(p-n) else 0) sums cos (x + y)" proof - have "(\n\p. if even p then ((-1) ^ (p div 2) * (p choose n) / (fact p)) *\<^sub>R (x^n) * y^(p-n) else 0) = cos_coeff p *\<^sub>R ((x + y) ^ p)" for p :: nat proof - have "(\n\p. if even p then ((-1) ^ (p div 2) * (p choose n) / (fact p)) *\<^sub>R (x^n) * y^(p-n) else 0) = (if even p then \n\p. ((-1) ^ (p div 2) * (p choose n) / (fact p)) *\<^sub>R (x^n) * y^(p-n) else 0)" by simp also have "\ = (if even p then of_real ((-1) ^ (p div 2) / (fact p)) * (\n\p. (p choose n) *\<^sub>R (x^n) * y^(p-n)) else 0)" by (auto simp: sum_distrib_left field_simps scaleR_conv_of_real nonzero_of_real_divide) also have "\ = cos_coeff p *\<^sub>R ((x + y) ^ p)" by (simp add: cos_coeff_def binomial_ring [of x y] scaleR_conv_of_real atLeast0AtMost) finally show ?thesis . qed then have "(\p. \n\p. if even p then ((-1) ^ (p div 2) * (p choose n) / (fact p)) *\<^sub>R (x^n) * y^(p-n) else 0) = (\p. cos_coeff p *\<^sub>R ((x+y)^p))" by simp also have "\ sums cos (x + y)" by (rule cos_converges) finally show ?thesis . qed theorem cos_add: fixes x :: "'a::{real_normed_field,banach}" shows "cos (x + y) = cos x * cos y - sin x * sin y" proof - have "(if even p \ even n then ((- 1) ^ (p div 2) * int (p choose n) / (fact p)) *\<^sub>R (x^n) * y^(p-n) else 0) - (if even p \ odd n then - ((- 1) ^ (p div 2) * int (p choose n) / (fact p)) *\<^sub>R (x^n) * y^(p-n) else 0) = (if even p then ((-1) ^ (p div 2) * (p choose n) / (fact p)) *\<^sub>R (x^n) * y^(p-n) else 0)" if "n \ p" for n p :: nat by simp then have "(\p. \n\p. (if even p then ((-1) ^ (p div 2) * (p choose n) / (fact p)) *\<^sub>R (x^n) * y^(p-n) else 0)) sums (cos x * cos y - sin x * sin y)" using sums_diff [OF cos_x_cos_y [of x y] sin_x_sin_y [of x y]] by (simp add: sum_subtractf [symmetric]) then show ?thesis by (blast intro: sums_cos_x_plus_y sums_unique2) qed lemma sin_minus_converges: "(\n. - (sin_coeff n *\<^sub>R (-x)^n)) sums sin x" proof - have [simp]: "\n. - (sin_coeff n *\<^sub>R (-x)^n) = (sin_coeff n *\<^sub>R x^n)" by (auto simp: sin_coeff_def elim!: oddE) show ?thesis by (simp add: sin_def summable_norm_sin [THEN summable_norm_cancel, THEN summable_sums]) qed lemma sin_minus [simp]: "sin (- x) = - sin x" for x :: "'a::{real_normed_algebra_1,banach}" using sin_minus_converges [of x] by (auto simp: sin_def summable_norm_sin [THEN summable_norm_cancel] suminf_minus sums_iff equation_minus_iff) lemma cos_minus_converges: "(\n. (cos_coeff n *\<^sub>R (-x)^n)) sums cos x" proof - have [simp]: "\n. (cos_coeff n *\<^sub>R (-x)^n) = (cos_coeff n *\<^sub>R x^n)" by (auto simp: Transcendental.cos_coeff_def elim!: evenE) show ?thesis by (simp add: cos_def summable_norm_cos [THEN summable_norm_cancel, THEN summable_sums]) qed lemma cos_minus [simp]: "cos (-x) = cos x" for x :: "'a::{real_normed_algebra_1,banach}" - using cos_minus_converges [of x] - by (simp add: cos_def summable_norm_cos [THEN summable_norm_cancel] - suminf_minus sums_iff equation_minus_iff) + using cos_minus_converges [of x] by (metis cos_def sums_unique) + +lemma cos_abs_real [simp]: "cos \x :: real\ = cos x" + by (simp add: abs_if) lemma sin_cos_squared_add [simp]: "(sin x)\<^sup>2 + (cos x)\<^sup>2 = 1" for x :: "'a::{real_normed_field,banach}" using cos_add [of x "-x"] by (simp add: power2_eq_square algebra_simps) lemma sin_cos_squared_add2 [simp]: "(cos x)\<^sup>2 + (sin x)\<^sup>2 = 1" for x :: "'a::{real_normed_field,banach}" by (subst add.commute, rule sin_cos_squared_add) lemma sin_cos_squared_add3 [simp]: "cos x * cos x + sin x * sin x = 1" for x :: "'a::{real_normed_field,banach}" using sin_cos_squared_add2 [unfolded power2_eq_square] . lemma sin_squared_eq: "(sin x)\<^sup>2 = 1 - (cos x)\<^sup>2" for x :: "'a::{real_normed_field,banach}" unfolding eq_diff_eq by (rule sin_cos_squared_add) lemma cos_squared_eq: "(cos x)\<^sup>2 = 1 - (sin x)\<^sup>2" for x :: "'a::{real_normed_field,banach}" unfolding eq_diff_eq by (rule sin_cos_squared_add2) lemma abs_sin_le_one [simp]: "\sin x\ \ 1" for x :: real by (rule power2_le_imp_le) (simp_all add: sin_squared_eq) lemma sin_ge_minus_one [simp]: "- 1 \ sin x" for x :: real using abs_sin_le_one [of x] by (simp add: abs_le_iff) lemma sin_le_one [simp]: "sin x \ 1" for x :: real using abs_sin_le_one [of x] by (simp add: abs_le_iff) lemma abs_cos_le_one [simp]: "\cos x\ \ 1" for x :: real by (rule power2_le_imp_le) (simp_all add: cos_squared_eq) lemma cos_ge_minus_one [simp]: "- 1 \ cos x" for x :: real using abs_cos_le_one [of x] by (simp add: abs_le_iff) lemma cos_le_one [simp]: "cos x \ 1" for x :: real using abs_cos_le_one [of x] by (simp add: abs_le_iff) lemma cos_diff: "cos (x - y) = cos x * cos y + sin x * sin y" for x :: "'a::{real_normed_field,banach}" using cos_add [of x "- y"] by simp lemma cos_double: "cos(2*x) = (cos x)\<^sup>2 - (sin x)\<^sup>2" for x :: "'a::{real_normed_field,banach}" using cos_add [where x=x and y=x] by (simp add: power2_eq_square) lemma sin_cos_le1: "\sin x * sin y + cos x * cos y\ \ 1" for x :: real using cos_diff [of x y] by (metis abs_cos_le_one add.commute) lemma DERIV_fun_pow: "DERIV g x :> m \ DERIV (\x. (g x) ^ n) x :> real n * (g x) ^ (n - 1) * m" by (auto intro!: derivative_eq_intros simp:) lemma DERIV_fun_exp: "DERIV g x :> m \ DERIV (\x. exp (g x)) x :> exp (g x) * m" by (auto intro!: derivative_intros) subsection \The Constant Pi\ definition pi :: real where "pi = 2 * (THE x. 0 \ x \ x \ 2 \ cos x = 0)" text \Show that there's a least positive \<^term>\x\ with \<^term>\cos x = 0\; hence define pi.\ lemma sin_paired: "(\n. (- 1) ^ n / (fact (2 * n + 1)) * x ^ (2 * n + 1)) sums sin x" for x :: real proof - have "(\n. \k = n*2..n. 0 < ?f n" proof fix n :: nat let ?k2 = "real (Suc (Suc (4 * n)))" let ?k3 = "real (Suc (Suc (Suc (4 * n))))" have "x * x < ?k2 * ?k3" using assms by (intro mult_strict_mono', simp_all) then have "x * x * x * x ^ (n * 4) < ?k2 * ?k3 * x * x ^ (n * 4)" by (intro mult_strict_right_mono zero_less_power \0 < x\) then show "0 < ?f n" by (simp add: ac_simps divide_less_eq) qed have sums: "?f sums sin x" by (rule sin_paired [THEN sums_group]) simp show "0 < sin x" unfolding sums_unique [OF sums] using sums_summable [OF sums] pos by (simp add: suminf_pos) qed lemma cos_double_less_one: "0 < x \ x < 2 \ cos (2 * x) < 1" for x :: real using sin_gt_zero_02 [where x = x] by (auto simp: cos_squared_eq cos_double) lemma cos_paired: "(\n. (- 1) ^ n / (fact (2 * n)) * x ^ (2 * n)) sums cos x" for x :: real proof - have "(\n. \k = n * 2.. real" assumes f: "summable f" and fplus: "\d. 0 < f (k + (Suc(Suc 0) * d)) + f (k + ((Suc (Suc 0) * d) + 1))" shows "sum f {..n. \n = n * Suc (Suc 0)..n. f (n + k))" proof (rule sums_group) show "(\n. f (n + k)) sums (\n. f (n + k))" by (simp add: f summable_iff_shift summable_sums) qed auto with fplus have "0 < (\n. f (n + k))" apply (simp add: add.commute) apply (metis (no_types, lifting) suminf_pos summable_def sums_unique) done then show ?thesis by (simp add: f suminf_minus_initial_segment) qed lemma cos_two_less_zero [simp]: "cos 2 < (0::real)" proof - note fact_Suc [simp del] from sums_minus [OF cos_paired] have *: "(\n. - ((- 1) ^ n * 2 ^ (2 * n) / fact (2 * n))) sums - cos (2::real)" by simp then have sm: "summable (\n. - ((- 1::real) ^ n * 2 ^ (2 * n) / (fact (2 * n))))" by (rule sums_summable) have "0 < (\nnn. - ((- 1) ^ n * 2 ^ (2 * n) / (fact (2 * n))))" proof - { fix d let ?six4d = "Suc (Suc (Suc (Suc (Suc (Suc (4 * d))))))" have "(4::real) * (fact (?six4d)) < (Suc (Suc (?six4d)) * fact (Suc (?six4d)))" unfolding of_nat_mult by (rule mult_strict_mono) (simp_all add: fact_less_mono) then have "(4::real) * (fact (?six4d)) < (fact (Suc (Suc (?six4d))))" by (simp only: fact_Suc [of "Suc (?six4d)"] of_nat_mult of_nat_fact) then have "(4::real) * inverse (fact (Suc (Suc (?six4d)))) < inverse (fact (?six4d))" by (simp add: inverse_eq_divide less_divide_eq) } then show ?thesis by (force intro!: sum_pos_lt_pair [OF sm] simp add: divide_inverse algebra_simps) qed ultimately have "0 < (\n. - ((- 1::real) ^ n * 2 ^ (2 * n) / (fact (2 * n))))" by (rule order_less_trans) moreover from * have "- cos 2 = (\n. - ((- 1::real) ^ n * 2 ^ (2 * n) / (fact (2 * n))))" by (rule sums_unique) ultimately have "(0::real) < - cos 2" by simp then show ?thesis by simp qed lemmas cos_two_neq_zero [simp] = cos_two_less_zero [THEN less_imp_neq] lemmas cos_two_le_zero [simp] = cos_two_less_zero [THEN order_less_imp_le] lemma cos_is_zero: "\!x::real. 0 \ x \ x \ 2 \ cos x = 0" proof (rule ex_ex1I) show "\x::real. 0 \ x \ x \ 2 \ cos x = 0" by (rule IVT2) simp_all next fix a b :: real assume ab: "0 \ a \ a \ 2 \ cos a = 0" "0 \ b \ b \ 2 \ cos b = 0" have cosd: "\x::real. cos differentiable (at x)" unfolding real_differentiable_def by (auto intro: DERIV_cos) show "a = b" proof (cases a b rule: linorder_cases) case less then obtain z where "a < z" "z < b" "(cos has_real_derivative 0) (at z)" using Rolle by (metis cosd continuous_on_cos_real ab) then have "sin z = 0" using DERIV_cos DERIV_unique neg_equal_0_iff_equal by blast then show ?thesis by (metis \a < z\ \z < b\ ab order_less_le_trans less_le sin_gt_zero_02) next case greater then obtain z where "b < z" "z < a" "(cos has_real_derivative 0) (at z)" using Rolle by (metis cosd continuous_on_cos_real ab) then have "sin z = 0" using DERIV_cos DERIV_unique neg_equal_0_iff_equal by blast then show ?thesis by (metis \b < z\ \z < a\ ab order_less_le_trans less_le sin_gt_zero_02) qed auto qed lemma pi_half: "pi/2 = (THE x. 0 \ x \ x \ 2 \ cos x = 0)" by (simp add: pi_def) lemma cos_pi_half [simp]: "cos (pi/2) = 0" by (simp add: pi_half cos_is_zero [THEN theI']) lemma cos_of_real_pi_half [simp]: "cos ((of_real pi/2) :: 'a) = 0" if "SORT_CONSTRAINT('a::{real_field,banach,real_normed_algebra_1})" by (metis cos_pi_half cos_of_real eq_numeral_simps(4) nonzero_of_real_divide of_real_0 of_real_numeral) lemma pi_half_gt_zero [simp]: "0 < pi/2" proof - have "0 \ pi/2" by (simp add: pi_half cos_is_zero [THEN theI']) then show ?thesis by (metis cos_pi_half cos_zero less_eq_real_def one_neq_zero) qed lemmas pi_half_neq_zero [simp] = pi_half_gt_zero [THEN less_imp_neq, symmetric] lemmas pi_half_ge_zero [simp] = pi_half_gt_zero [THEN order_less_imp_le] lemma pi_half_less_two [simp]: "pi/2 < 2" proof - have "pi/2 \ 2" by (simp add: pi_half cos_is_zero [THEN theI']) then show ?thesis by (metis cos_pi_half cos_two_neq_zero le_less) qed lemmas pi_half_neq_two [simp] = pi_half_less_two [THEN less_imp_neq] lemmas pi_half_le_two [simp] = pi_half_less_two [THEN order_less_imp_le] lemma pi_gt_zero [simp]: "0 < pi" using pi_half_gt_zero by simp lemma pi_ge_zero [simp]: "0 \ pi" by (rule pi_gt_zero [THEN order_less_imp_le]) lemma pi_neq_zero [simp]: "pi \ 0" by (rule pi_gt_zero [THEN less_imp_neq, symmetric]) lemma pi_not_less_zero [simp]: "\ pi < 0" by (simp add: linorder_not_less) lemma minus_pi_half_less_zero: "-(pi/2) < 0" by simp lemma m2pi_less_pi: "- (2*pi) < pi" by simp lemma sin_pi_half [simp]: "sin(pi/2) = 1" using sin_cos_squared_add2 [where x = "pi/2"] using sin_gt_zero_02 [OF pi_half_gt_zero pi_half_less_two] by (simp add: power2_eq_1_iff) lemma sin_of_real_pi_half [simp]: "sin ((of_real pi/2) :: 'a) = 1" if "SORT_CONSTRAINT('a::{real_field,banach,real_normed_algebra_1})" using sin_pi_half by (metis sin_pi_half eq_numeral_simps(4) nonzero_of_real_divide of_real_1 of_real_numeral sin_of_real) lemma sin_cos_eq: "sin x = cos (of_real pi/2 - x)" for x :: "'a::{real_normed_field,banach}" by (simp add: cos_diff) lemma minus_sin_cos_eq: "- sin x = cos (x + of_real pi/2)" for x :: "'a::{real_normed_field,banach}" by (simp add: cos_add nonzero_of_real_divide) lemma cos_sin_eq: "cos x = sin (of_real pi/2 - x)" for x :: "'a::{real_normed_field,banach}" using sin_cos_eq [of "of_real pi/2 - x"] by simp lemma sin_add: "sin (x + y) = sin x * cos y + cos x * sin y" for x :: "'a::{real_normed_field,banach}" using cos_add [of "of_real pi/2 - x" "-y"] by (simp add: cos_sin_eq) (simp add: sin_cos_eq) lemma sin_diff: "sin (x - y) = sin x * cos y - cos x * sin y" for x :: "'a::{real_normed_field,banach}" using sin_add [of x "- y"] by simp lemma sin_double: "sin(2 * x) = 2 * sin x * cos x" for x :: "'a::{real_normed_field,banach}" using sin_add [where x=x and y=x] by simp lemma cos_of_real_pi [simp]: "cos (of_real pi) = -1" using cos_add [where x = "pi/2" and y = "pi/2"] by (simp add: cos_of_real) lemma sin_of_real_pi [simp]: "sin (of_real pi) = 0" using sin_add [where x = "pi/2" and y = "pi/2"] by (simp add: sin_of_real) lemma cos_pi [simp]: "cos pi = -1" using cos_add [where x = "pi/2" and y = "pi/2"] by simp lemma sin_pi [simp]: "sin pi = 0" using sin_add [where x = "pi/2" and y = "pi/2"] by simp lemma sin_periodic_pi [simp]: "sin (x + pi) = - sin x" by (simp add: sin_add) lemma sin_periodic_pi2 [simp]: "sin (pi + x) = - sin x" by (simp add: sin_add) lemma cos_periodic_pi [simp]: "cos (x + pi) = - cos x" by (simp add: cos_add) lemma cos_periodic_pi2 [simp]: "cos (pi + x) = - cos x" by (simp add: cos_add) lemma sin_periodic [simp]: "sin (x + 2 * pi) = sin x" by (simp add: sin_add sin_double cos_double) lemma cos_periodic [simp]: "cos (x + 2 * pi) = cos x" by (simp add: cos_add sin_double cos_double) lemma cos_npi [simp]: "cos (real n * pi) = (- 1) ^ n" by (induct n) (auto simp: distrib_right) lemma cos_npi2 [simp]: "cos (pi * real n) = (- 1) ^ n" by (metis cos_npi mult.commute) lemma sin_npi [simp]: "sin (real n * pi) = 0" for n :: nat by (induct n) (auto simp: distrib_right) lemma sin_npi2 [simp]: "sin (pi * real n) = 0" for n :: nat by (simp add: mult.commute [of pi]) lemma cos_two_pi [simp]: "cos (2 * pi) = 1" by (simp add: cos_double) lemma sin_two_pi [simp]: "sin (2 * pi) = 0" by (simp add: sin_double) context fixes w :: "'a::{real_normed_field,banach}" begin lemma sin_times_sin: "sin w * sin z = (cos (w - z) - cos (w + z)) / 2" by (simp add: cos_diff cos_add) lemma sin_times_cos: "sin w * cos z = (sin (w + z) + sin (w - z)) / 2" by (simp add: sin_diff sin_add) lemma cos_times_sin: "cos w * sin z = (sin (w + z) - sin (w - z)) / 2" by (simp add: sin_diff sin_add) lemma cos_times_cos: "cos w * cos z = (cos (w - z) + cos (w + z)) / 2" by (simp add: cos_diff cos_add) lemma cos_double_cos: "cos (2 * w) = 2 * cos w ^ 2 - 1" by (simp add: cos_double sin_squared_eq) lemma cos_double_sin: "cos (2 * w) = 1 - 2 * sin w ^ 2" by (simp add: cos_double sin_squared_eq) end lemma sin_plus_sin: "sin w + sin z = 2 * sin ((w + z) / 2) * cos ((w - z) / 2)" for w :: "'a::{real_normed_field,banach}" apply (simp add: mult.assoc sin_times_cos) apply (simp add: field_simps) done lemma sin_diff_sin: "sin w - sin z = 2 * sin ((w - z) / 2) * cos ((w + z) / 2)" for w :: "'a::{real_normed_field,banach}" apply (simp add: mult.assoc sin_times_cos) apply (simp add: field_simps) done lemma cos_plus_cos: "cos w + cos z = 2 * cos ((w + z) / 2) * cos ((w - z) / 2)" for w :: "'a::{real_normed_field,banach,field}" apply (simp add: mult.assoc cos_times_cos) apply (simp add: field_simps) done lemma cos_diff_cos: "cos w - cos z = 2 * sin ((w + z) / 2) * sin ((z - w) / 2)" for w :: "'a::{real_normed_field,banach,field}" apply (simp add: mult.assoc sin_times_sin) apply (simp add: field_simps) done lemma sin_pi_minus [simp]: "sin (pi - x) = sin x" by (metis sin_minus sin_periodic_pi minus_minus uminus_add_conv_diff) lemma cos_pi_minus [simp]: "cos (pi - x) = - (cos x)" by (metis cos_minus cos_periodic_pi uminus_add_conv_diff) lemma sin_minus_pi [simp]: "sin (x - pi) = - (sin x)" by (simp add: sin_diff) lemma cos_minus_pi [simp]: "cos (x - pi) = - (cos x)" by (simp add: cos_diff) lemma sin_2pi_minus [simp]: "sin (2 * pi - x) = - (sin x)" by (metis sin_periodic_pi2 add_diff_eq mult_2 sin_pi_minus) lemma cos_2pi_minus [simp]: "cos (2 * pi - x) = cos x" by (metis (no_types, opaque_lifting) cos_add cos_minus cos_two_pi sin_minus sin_two_pi diff_0_right minus_diff_eq mult_1 mult_zero_left uminus_add_conv_diff) lemma sin_gt_zero2: "0 < x \ x < pi/2 \ 0 < sin x" by (metis sin_gt_zero_02 order_less_trans pi_half_less_two) lemma sin_less_zero: assumes "- pi/2 < x" and "x < 0" shows "sin x < 0" proof - have "0 < sin (- x)" using assms by (simp only: sin_gt_zero2) then show ?thesis by simp qed lemma pi_less_4: "pi < 4" using pi_half_less_two by auto lemma cos_gt_zero: "0 < x \ x < pi/2 \ 0 < cos x" by (simp add: cos_sin_eq sin_gt_zero2) lemma cos_gt_zero_pi: "-(pi/2) < x \ x < pi/2 \ 0 < cos x" using cos_gt_zero [of x] cos_gt_zero [of "-x"] by (cases rule: linorder_cases [of x 0]) auto lemma cos_ge_zero: "-(pi/2) \ x \ x \ pi/2 \ 0 \ cos x" by (auto simp: order_le_less cos_gt_zero_pi) (metis cos_pi_half eq_divide_eq eq_numeral_simps(4)) lemma sin_gt_zero: "0 < x \ x < pi \ 0 < sin x" by (simp add: sin_cos_eq cos_gt_zero_pi) lemma sin_lt_zero: "pi < x \ x < 2 * pi \ sin x < 0" using sin_gt_zero [of "x - pi"] by (simp add: sin_diff) lemma pi_ge_two: "2 \ pi" proof (rule ccontr) assume "\ ?thesis" then have "pi < 2" by auto have "\y > pi. y < 2 \ y < 2 * pi" proof (cases "2 < 2 * pi") case True with dense[OF \pi < 2\] show ?thesis by auto next case False have "pi < 2 * pi" by auto from dense[OF this] and False show ?thesis by auto qed then obtain y where "pi < y" and "y < 2" and "y < 2 * pi" by blast then have "0 < sin y" using sin_gt_zero_02 by auto moreover have "sin y < 0" using sin_gt_zero[of "y - pi"] \pi < y\ and \y < 2 * pi\ sin_periodic_pi[of "y - pi"] by auto ultimately show False by auto qed lemma sin_ge_zero: "0 \ x \ x \ pi \ 0 \ sin x" by (auto simp: order_le_less sin_gt_zero) lemma sin_le_zero: "pi \ x \ x < 2 * pi \ sin x \ 0" using sin_ge_zero [of "x - pi"] by (simp add: sin_diff) lemma sin_pi_divide_n_ge_0 [simp]: assumes "n \ 0" - shows "0 \ sin (pi / real n)" + shows "0 \ sin (pi/real n)" by (rule sin_ge_zero) (use assms in \simp_all add: field_split_simps\) lemma sin_pi_divide_n_gt_0: assumes "2 \ n" - shows "0 < sin (pi / real n)" + shows "0 < sin (pi/real n)" by (rule sin_gt_zero) (use assms in \simp_all add: field_split_simps\) text\Proof resembles that of \cos_is_zero\ but with \<^term>\pi\ for the upper bound\ lemma cos_total: assumes y: "-1 \ y" "y \ 1" shows "\!x. 0 \ x \ x \ pi \ cos x = y" proof (rule ex_ex1I) show "\x::real. 0 \ x \ x \ pi \ cos x = y" by (rule IVT2) (simp_all add: y) next fix a b :: real assume ab: "0 \ a \ a \ pi \ cos a = y" "0 \ b \ b \ pi \ cos b = y" have cosd: "\x::real. cos differentiable (at x)" unfolding real_differentiable_def by (auto intro: DERIV_cos) show "a = b" proof (cases a b rule: linorder_cases) case less then obtain z where "a < z" "z < b" "(cos has_real_derivative 0) (at z)" using Rolle by (metis cosd continuous_on_cos_real ab) then have "sin z = 0" using DERIV_cos DERIV_unique neg_equal_0_iff_equal by blast then show ?thesis by (metis \a < z\ \z < b\ ab order_less_le_trans less_le sin_gt_zero) next case greater then obtain z where "b < z" "z < a" "(cos has_real_derivative 0) (at z)" using Rolle by (metis cosd continuous_on_cos_real ab) then have "sin z = 0" using DERIV_cos DERIV_unique neg_equal_0_iff_equal by blast then show ?thesis by (metis \b < z\ \z < a\ ab order_less_le_trans less_le sin_gt_zero) qed auto qed lemma sin_total: assumes y: "-1 \ y" "y \ 1" shows "\!x. - (pi/2) \ x \ x \ pi/2 \ sin x = y" proof - from cos_total [OF y] obtain x where x: "0 \ x" "x \ pi" "cos x = y" and uniq: "\x'. 0 \ x' \ x' \ pi \ cos x' = y \ x' = x " by blast show ?thesis unfolding sin_cos_eq proof (rule ex1I [where a="pi/2 - x"]) show "- (pi/2) \ z \ z \ pi/2 \ cos (of_real pi/2 - z) = y \ z = pi/2 - x" for z using uniq [of "pi/2 -z"] by auto qed (use x in auto) qed lemma cos_zero_lemma: assumes "0 \ x" "cos x = 0" shows "\n. odd n \ x = of_nat n * (pi/2)" proof - have xle: "x < (1 + real_of_int \x/pi\) * pi" using floor_correct [of "x/pi"] by (simp add: add.commute divide_less_eq) obtain n where "real n * pi \ x" "x < real (Suc n) * pi" proof show "real (nat \x / pi\) * pi \ x" using assms floor_divide_lower [of pi x] by auto show "x < real (Suc (nat \x / pi\)) * pi" using assms floor_divide_upper [of pi x] by (simp add: xle) qed then have x: "0 \ x - n * pi" "(x - n * pi) \ pi" "cos (x - n * pi) = 0" by (auto simp: algebra_simps cos_diff assms) then have "\!x. 0 \ x \ x \ pi \ cos x = 0" by (auto simp: intro!: cos_total) then obtain \ where \: "0 \ \" "\ \ pi" "cos \ = 0" and uniq: "\\. 0 \ \ \ \ \ pi \ cos \ = 0 \ \ = \" by blast then have "x - real n * pi = \" using x by blast moreover have "pi/2 = \" using pi_half_ge_zero uniq by fastforce ultimately show ?thesis by (rule_tac x = "Suc (2 * n)" in exI) (simp add: algebra_simps) qed lemma sin_zero_lemma: assumes "0 \ x" "sin x = 0" shows "\n::nat. even n \ x = real n * (pi/2)" proof - obtain n where "odd n" and n: "x + pi/2 = of_nat n * (pi/2)" "n > 0" using cos_zero_lemma [of "x + pi/2"] assms by (auto simp add: cos_add) - then have "x = real (n - 1) * (pi / 2)" + then have "x = real (n - 1) * (pi/2)" by (simp add: algebra_simps of_nat_diff) then show ?thesis by (simp add: \odd n\) qed lemma cos_zero_iff: "cos x = 0 \ ((\n. odd n \ x = real n * (pi/2)) \ (\n. odd n \ x = - (real n * (pi/2))))" (is "?lhs = ?rhs") proof - have *: "cos (real n * pi/2) = 0" if "odd n" for n :: nat proof - from that obtain m where "n = 2 * m + 1" .. then show ?thesis by (simp add: field_simps) (simp add: cos_add add_divide_distrib) qed show ?thesis proof show ?rhs if ?lhs using that cos_zero_lemma [of x] cos_zero_lemma [of "-x"] by force show ?lhs if ?rhs using that by (auto dest: * simp del: eq_divide_eq_numeral1) qed qed lemma sin_zero_iff: "sin x = 0 \ ((\n. even n \ x = real n * (pi/2)) \ (\n. even n \ x = - (real n * (pi/2))))" (is "?lhs = ?rhs") proof show ?rhs if ?lhs using that sin_zero_lemma [of x] sin_zero_lemma [of "-x"] by force show ?lhs if ?rhs using that by (auto elim: evenE) qed lemma sin_zero_pi_iff: fixes x::real assumes "\x\ < pi" shows "sin x = 0 \ x = 0" proof show "x = 0" if "sin x = 0" using that assms by (auto simp: sin_zero_iff) qed auto lemma cos_zero_iff_int: "cos x = 0 \ (\i. odd i \ x = of_int i * (pi/2))" proof - have 1: "\n. odd n \ \i. odd i \ real n = real_of_int i" by (metis even_of_nat_iff of_int_of_nat_eq) have 2: "\n. odd n \ \i. odd i \ - (real n * pi) = real_of_int i * pi" by (metis even_minus even_of_nat_iff mult.commute mult_minus_right of_int_minus of_int_of_nat_eq) have 3: "\odd i; \n. even n \ real_of_int i \ - (real n)\ \ \n. odd n \ real_of_int i = real n" for i by (cases i rule: int_cases2) auto show ?thesis by (force simp: cos_zero_iff intro!: 1 2 3) qed lemma sin_zero_iff_int: "sin x = 0 \ (\i. even i \ x = of_int i * (pi/2))" (is "?lhs = ?rhs") proof safe assume ?lhs then consider (plus) n where "even n" "x = real n * (pi/2)" | (minus) n where "even n" "x = - (real n * (pi/2))" using sin_zero_iff by auto then show "\n. even n \ x = of_int n * (pi/2)" proof cases case plus then show ?rhs by (metis even_of_nat_iff of_int_of_nat_eq) next case minus then show ?thesis by (rule_tac x="- (int n)" in exI) simp qed next fix i :: int assume "even i" then show "sin (of_int i * (pi/2)) = 0" by (cases i rule: int_cases2, simp_all add: sin_zero_iff) qed lemma sin_zero_iff_int2: "sin x = 0 \ (\i::int. x = of_int i * pi)" proof - - have "sin x = 0 \ (\i. even i \ x = real_of_int i * (pi / 2))" + have "sin x = 0 \ (\i. even i \ x = real_of_int i * (pi/2))" by (auto simp: sin_zero_iff_int) - also have "... = (\j. x = real_of_int (2*j) * (pi / 2))" + also have "... = (\j. x = real_of_int (2*j) * (pi/2))" using dvd_triv_left by blast also have "... = (\i::int. x = of_int i * pi)" by auto finally show ?thesis . qed lemma sin_npi_int [simp]: "sin (pi * of_int n) = 0" by (simp add: sin_zero_iff_int2) lemma cos_monotone_0_pi: assumes "0 \ y" and "y < x" and "x \ pi" shows "cos x < cos y" proof - have "- (x - y) < 0" using assms by auto from MVT2[OF \y < x\ DERIV_cos] obtain z where "y < z" and "z < x" and cos_diff: "cos x - cos y = (x - y) * - sin z" by auto then have "0 < z" and "z < pi" using assms by auto then have "0 < sin z" using sin_gt_zero by auto then have "cos x - cos y < 0" unfolding cos_diff minus_mult_commute[symmetric] using \- (x - y) < 0\ by (rule mult_pos_neg2) then show ?thesis by auto qed lemma cos_monotone_0_pi_le: assumes "0 \ y" and "y \ x" and "x \ pi" shows "cos x \ cos y" proof (cases "y < x") case True show ?thesis using cos_monotone_0_pi[OF \0 \ y\ True \x \ pi\] by auto next case False then have "y = x" using \y \ x\ by auto then show ?thesis by auto qed lemma cos_monotone_minus_pi_0: assumes "- pi \ y" and "y < x" and "x \ 0" shows "cos y < cos x" proof - have "0 \ - x" and "- x < - y" and "- y \ pi" using assms by auto from cos_monotone_0_pi[OF this] show ?thesis unfolding cos_minus . qed lemma cos_monotone_minus_pi_0': assumes "- pi \ y" and "y \ x" and "x \ 0" shows "cos y \ cos x" proof (cases "y < x") case True show ?thesis using cos_monotone_minus_pi_0[OF \-pi \ y\ True \x \ 0\] by auto next case False then have "y = x" using \y \ x\ by auto then show ?thesis by auto qed lemma sin_monotone_2pi: assumes "- (pi/2) \ y" and "y < x" and "x \ pi/2" shows "sin y < sin x" unfolding sin_cos_eq using assms by (auto intro: cos_monotone_0_pi) lemma sin_monotone_2pi_le: assumes "- (pi/2) \ y" and "y \ x" and "x \ pi/2" shows "sin y \ sin x" by (metis assms le_less sin_monotone_2pi) lemma sin_x_le_x: fixes x :: real assumes "x \ 0" shows "sin x \ x" proof - let ?f = "\x. x - sin x" have "\u. \0 \ u; u \ x\ \ \y. (?f has_real_derivative 1 - cos u) (at u)" by (auto intro!: derivative_eq_intros simp: field_simps) then have "?f x \ ?f 0" by (metis cos_le_one diff_ge_0_iff_ge DERIV_nonneg_imp_nondecreasing [OF assms]) then show "sin x \ x" by simp qed lemma sin_x_ge_neg_x: fixes x :: real assumes x: "x \ 0" shows "sin x \ - x" proof - let ?f = "\x. x + sin x" have \
: "\u. \0 \ u; u \ x\ \ \y. (?f has_real_derivative 1 + cos u) (at u)" by (auto intro!: derivative_eq_intros simp: field_simps) have "?f x \ ?f 0" by (rule DERIV_nonneg_imp_nondecreasing [OF assms]) (use \
real_0_le_add_iff in force) then show "sin x \ -x" by simp qed lemma abs_sin_x_le_abs_x: "\sin x\ \ \x\" for x :: real using sin_x_ge_neg_x [of x] sin_x_le_x [of x] sin_x_ge_neg_x [of "-x"] sin_x_le_x [of "-x"] by (auto simp: abs_real_def) subsection \More Corollaries about Sine and Cosine\ lemma sin_cos_npi [simp]: "sin (real (Suc (2 * n)) * pi/2) = (-1) ^ n" proof - have "sin ((real n + 1/2) * pi) = cos (real n * pi)" by (auto simp: algebra_simps sin_add) then show ?thesis by (simp add: distrib_right add_divide_distrib add.commute mult.commute [of pi]) qed lemma cos_2npi [simp]: "cos (2 * real n * pi) = 1" for n :: nat by (cases "even n") (simp_all add: cos_double mult.assoc) lemma cos_3over2_pi [simp]: "cos (3/2*pi) = 0" proof - have "cos (3/2*pi) = cos (pi + pi/2)" by simp also have "... = 0" by (subst cos_add, simp) finally show ?thesis . qed lemma sin_2npi [simp]: "sin (2 * real n * pi) = 0" for n :: nat by (auto simp: mult.assoc sin_double) lemma sin_3over2_pi [simp]: "sin (3/2*pi) = - 1" proof - have "sin (3/2*pi) = sin (pi + pi/2)" by simp also have "... = -1" by (subst sin_add, simp) finally show ?thesis . qed lemma cos_pi_eq_zero [simp]: "cos (pi * real (Suc (2 * m)) / 2) = 0" by (simp only: cos_add sin_add of_nat_Suc distrib_right distrib_left add_divide_distrib, auto) lemma DERIV_cos_add [simp]: "DERIV (\x. cos (x + k)) xa :> - sin (xa + k)" by (auto intro!: derivative_eq_intros) lemma sin_zero_norm_cos_one: fixes x :: "'a::{real_normed_field,banach}" assumes "sin x = 0" shows "norm (cos x) = 1" using sin_cos_squared_add [of x, unfolded assms] by (simp add: square_norm_one) lemma sin_zero_abs_cos_one: "sin x = 0 \ \cos x\ = (1::real)" using sin_zero_norm_cos_one by fastforce lemma cos_one_sin_zero: fixes x :: "'a::{real_normed_field,banach}" assumes "cos x = 1" shows "sin x = 0" using sin_cos_squared_add [of x, unfolded assms] by simp lemma sin_times_pi_eq_0: "sin (x * pi) = 0 \ x \ \" by (simp add: sin_zero_iff_int2) (metis Ints_cases Ints_of_int) lemma cos_one_2pi: "cos x = 1 \ (\n::nat. x = n * 2 * pi) \ (\n::nat. x = - (n * 2 * pi))" (is "?lhs = ?rhs") proof assume ?lhs then have "sin x = 0" by (simp add: cos_one_sin_zero) then show ?rhs proof (simp only: sin_zero_iff, elim exE disjE conjE) fix n :: nat assume n: "even n" "x = real n * (pi/2)" then obtain m where m: "n = 2 * m" using dvdE by blast then have me: "even m" using \?lhs\ n by (auto simp: field_simps) (metis one_neq_neg_one power_minus_odd power_one) show ?rhs using m me n by (auto simp: field_simps elim!: evenE) next fix n :: nat assume n: "even n" "x = - (real n * (pi/2))" then obtain m where m: "n = 2 * m" using dvdE by blast then have me: "even m" using \?lhs\ n by (auto simp: field_simps) (metis one_neq_neg_one power_minus_odd power_one) show ?rhs using m me n by (auto simp: field_simps elim!: evenE) qed next assume ?rhs then show "cos x = 1" by (metis cos_2npi cos_minus mult.assoc mult.left_commute) qed lemma cos_one_2pi_int: "cos x = 1 \ (\n::int. x = n * 2 * pi)" (is "?lhs = ?rhs") proof assume "cos x = 1" then show ?rhs by (metis cos_one_2pi mult.commute mult_minus_right of_int_minus of_int_of_nat_eq) next assume ?rhs then show "cos x = 1" by (clarsimp simp add: cos_one_2pi) (metis mult_minus_right of_int_of_nat) qed lemma cos_npi_int [simp]: fixes n::int shows "cos (pi * of_int n) = (if even n then 1 else -1)" by (auto simp: algebra_simps cos_one_2pi_int elim!: oddE evenE) lemma sin_cos_sqrt: "0 \ sin x \ sin x = sqrt (1 - (cos(x) ^ 2))" using sin_squared_eq real_sqrt_unique by fastforce lemma sin_eq_0_pi: "- pi < x \ x < pi \ sin x = 0 \ x = 0" by (metis sin_gt_zero sin_minus minus_less_iff neg_0_less_iff_less not_less_iff_gr_or_eq) lemma cos_treble_cos: "cos (3 * x) = 4 * cos x ^ 3 - 3 * cos x" for x :: "'a::{real_normed_field,banach}" proof - have *: "(sin x * (sin x * 3)) = 3 - (cos x * (cos x * 3))" by (simp add: mult.assoc [symmetric] sin_squared_eq [unfolded power2_eq_square]) have "cos(3 * x) = cos(2*x + x)" by simp also have "\ = 4 * cos x ^ 3 - 3 * cos x" unfolding cos_add cos_double sin_double by (simp add: * field_simps power2_eq_square power3_eq_cube) finally show ?thesis . qed -lemma cos_45: "cos (pi / 4) = sqrt 2 / 2" -proof - - let ?c = "cos (pi / 4)" - let ?s = "sin (pi / 4)" +lemma cos_45: "cos (pi/4) = sqrt 2 / 2" +proof - + let ?c = "cos (pi/4)" + let ?s = "sin (pi/4)" have nonneg: "0 \ ?c" by (simp add: cos_ge_zero) - have "0 = cos (pi / 4 + pi / 4)" + have "0 = cos (pi/4 + pi/4)" by simp - also have "cos (pi / 4 + pi / 4) = ?c\<^sup>2 - ?s\<^sup>2" + also have "cos (pi/4 + pi/4) = ?c\<^sup>2 - ?s\<^sup>2" by (simp only: cos_add power2_eq_square) also have "\ = 2 * ?c\<^sup>2 - 1" by (simp add: sin_squared_eq) finally have "?c\<^sup>2 = (sqrt 2 / 2)\<^sup>2" by (simp add: power_divide) then show ?thesis using nonneg by (rule power2_eq_imp_eq) simp qed -lemma cos_30: "cos (pi / 6) = sqrt 3/2" -proof - - let ?c = "cos (pi / 6)" - let ?s = "sin (pi / 6)" +lemma cos_30: "cos (pi/6) = sqrt 3/2" +proof - + let ?c = "cos (pi/6)" + let ?s = "sin (pi/6)" have pos_c: "0 < ?c" by (rule cos_gt_zero) simp_all - have "0 = cos (pi / 6 + pi / 6 + pi / 6)" + have "0 = cos (pi/6 + pi/6 + pi/6)" by simp also have "\ = (?c * ?c - ?s * ?s) * ?c - (?s * ?c + ?c * ?s) * ?s" by (simp only: cos_add sin_add) also have "\ = ?c * (?c\<^sup>2 - 3 * ?s\<^sup>2)" by (simp add: algebra_simps power2_eq_square) finally have "?c\<^sup>2 = (sqrt 3/2)\<^sup>2" using pos_c by (simp add: sin_squared_eq power_divide) then show ?thesis using pos_c [THEN order_less_imp_le] by (rule power2_eq_imp_eq) simp qed -lemma sin_45: "sin (pi / 4) = sqrt 2 / 2" +lemma sin_45: "sin (pi/4) = sqrt 2 / 2" by (simp add: sin_cos_eq cos_45) -lemma sin_60: "sin (pi / 3) = sqrt 3/2" +lemma sin_60: "sin (pi/3) = sqrt 3/2" by (simp add: sin_cos_eq cos_30) -lemma cos_60: "cos (pi / 3) = 1 / 2" -proof - - have "0 \ cos (pi / 3)" +lemma cos_60: "cos (pi/3) = 1/2" +proof - + have "0 \ cos (pi/3)" by (rule cos_ge_zero) (use pi_half_ge_zero in \linarith+\) then show ?thesis by (simp add: cos_squared_eq sin_60 power_divide power2_eq_imp_eq) qed -lemma sin_30: "sin (pi / 6) = 1 / 2" +lemma sin_30: "sin (pi/6) = 1/2" by (simp add: sin_cos_eq cos_60) +lemma cos_120: "cos (2 * pi/3) = -1/2" + and sin_120: "sin (2 * pi/3) = sqrt 3 / 2" + using sin_double[of "pi/3"] cos_double[of "pi/3"] + by (simp_all add: power2_eq_square sin_60 cos_60) + +lemma cos_120': "cos (pi * 2 / 3) = -1/2" + using cos_120 by (subst mult.commute) + +lemma sin_120': "sin (pi * 2 / 3) = sqrt 3 / 2" + using sin_120 by (subst mult.commute) + lemma cos_integer_2pi: "n \ \ \ cos(2 * pi * n) = 1" by (metis Ints_cases cos_one_2pi_int mult.assoc mult.commute) lemma sin_integer_2pi: "n \ \ \ sin(2 * pi * n) = 0" by (metis sin_two_pi Ints_mult mult.assoc mult.commute sin_times_pi_eq_0) lemma cos_int_2pin [simp]: "cos ((2 * pi) * of_int n) = 1" by (simp add: cos_one_2pi_int) lemma sin_int_2pin [simp]: "sin ((2 * pi) * of_int n) = 0" by (metis Ints_of_int sin_integer_2pi) lemma sincos_principal_value: "\y. (- pi < y \ y \ pi) \ (sin y = sin x \ cos y = cos x)" proof - define y where "y \ pi - (2 * pi) * frac ((pi - x) / (2 * pi))" have "-pi < y"" y \ pi" by (auto simp: field_simps frac_lt_1 y_def) moreover have "sin y = sin x" "cos y = cos x" unfolding y_def apply (simp_all add: frac_def divide_simps sin_add cos_add) by (metis sin_int_2pin cos_int_2pin diff_zero add.right_neutral mult.commute mult.left_neutral mult_zero_left)+ ultimately show ?thesis by metis qed subsection \Tangent\ definition tan :: "'a \ 'a::{real_normed_field,banach}" where "tan = (\x. sin x / cos x)" lemma tan_of_real: "of_real (tan x) = (tan (of_real x) :: 'a::{real_normed_field,banach})" by (simp add: tan_def sin_of_real cos_of_real) lemma tan_in_Reals [simp]: "z \ \ \ tan z \ \" for z :: "'a::{real_normed_field,banach}" by (simp add: tan_def) lemma tan_zero [simp]: "tan 0 = 0" by (simp add: tan_def) lemma tan_pi [simp]: "tan pi = 0" by (simp add: tan_def) lemma tan_npi [simp]: "tan (real n * pi) = 0" for n :: nat by (simp add: tan_def) lemma tan_minus [simp]: "tan (- x) = - tan x" by (simp add: tan_def) lemma tan_periodic [simp]: "tan (x + 2 * pi) = tan x" by (simp add: tan_def) lemma lemma_tan_add1: "cos x \ 0 \ cos y \ 0 \ 1 - tan x * tan y = cos (x + y)/(cos x * cos y)" by (simp add: tan_def cos_add field_simps) lemma add_tan_eq: "cos x \ 0 \ cos y \ 0 \ tan x + tan y = sin(x + y)/(cos x * cos y)" for x :: "'a::{real_normed_field,banach}" by (simp add: tan_def sin_add field_simps) lemma tan_add: "cos x \ 0 \ cos y \ 0 \ cos (x + y) \ 0 \ tan (x + y) = (tan x + tan y)/(1 - tan x * tan y)" for x :: "'a::{real_normed_field,banach}" by (simp add: add_tan_eq lemma_tan_add1 field_simps) (simp add: tan_def) lemma tan_double: "cos x \ 0 \ cos (2 * x) \ 0 \ tan (2 * x) = (2 * tan x) / (1 - (tan x)\<^sup>2)" for x :: "'a::{real_normed_field,banach}" using tan_add [of x x] by (simp add: power2_eq_square) lemma tan_gt_zero: "0 < x \ x < pi/2 \ 0 < tan x" by (simp add: tan_def zero_less_divide_iff sin_gt_zero2 cos_gt_zero_pi) lemma tan_less_zero: assumes "- pi/2 < x" and "x < 0" shows "tan x < 0" proof - have "0 < tan (- x)" using assms by (simp only: tan_gt_zero) then show ?thesis by simp qed lemma tan_half: "tan x = sin (2 * x) / (cos (2 * x) + 1)" for x :: "'a::{real_normed_field,banach,field}" unfolding tan_def sin_double cos_double sin_squared_eq by (simp add: power2_eq_square) -lemma tan_30: "tan (pi / 6) = 1 / sqrt 3" +lemma tan_30: "tan (pi/6) = 1 / sqrt 3" unfolding tan_def by (simp add: sin_30 cos_30) -lemma tan_45: "tan (pi / 4) = 1" +lemma tan_45: "tan (pi/4) = 1" unfolding tan_def by (simp add: sin_45 cos_45) -lemma tan_60: "tan (pi / 3) = sqrt 3" +lemma tan_60: "tan (pi/3) = sqrt 3" unfolding tan_def by (simp add: sin_60 cos_60) lemma DERIV_tan [simp]: "cos x \ 0 \ DERIV tan x :> inverse ((cos x)\<^sup>2)" for x :: "'a::{real_normed_field,banach}" unfolding tan_def by (auto intro!: derivative_eq_intros, simp add: divide_inverse power2_eq_square) declare DERIV_tan[THEN DERIV_chain2, derivative_intros] and DERIV_tan[THEN DERIV_chain2, unfolded has_field_derivative_def, derivative_intros] lemmas has_derivative_tan[derivative_intros] = DERIV_tan[THEN DERIV_compose_FDERIV] lemma isCont_tan: "cos x \ 0 \ isCont tan x" for x :: "'a::{real_normed_field,banach}" by (rule DERIV_tan [THEN DERIV_isCont]) lemma isCont_tan' [simp,continuous_intros]: fixes a :: "'a::{real_normed_field,banach}" and f :: "'a \ 'a" shows "isCont f a \ cos (f a) \ 0 \ isCont (\x. tan (f x)) a" by (rule isCont_o2 [OF _ isCont_tan]) lemma tendsto_tan [tendsto_intros]: fixes f :: "'a \ 'a::{real_normed_field,banach}" shows "(f \ a) F \ cos a \ 0 \ ((\x. tan (f x)) \ tan a) F" by (rule isCont_tendsto_compose [OF isCont_tan]) lemma continuous_tan: fixes f :: "'a \ 'a::{real_normed_field,banach}" shows "continuous F f \ cos (f (Lim F (\x. x))) \ 0 \ continuous F (\x. tan (f x))" unfolding continuous_def by (rule tendsto_tan) lemma continuous_on_tan [continuous_intros]: fixes f :: "'a \ 'a::{real_normed_field,banach}" shows "continuous_on s f \ (\x\s. cos (f x) \ 0) \ continuous_on s (\x. tan (f x))" unfolding continuous_on_def by (auto intro: tendsto_tan) lemma continuous_within_tan [continuous_intros]: fixes f :: "'a \ 'a::{real_normed_field,banach}" shows "continuous (at x within s) f \ cos (f x) \ 0 \ continuous (at x within s) (\x. tan (f x))" unfolding continuous_within by (rule tendsto_tan) lemma LIM_cos_div_sin: "(\x. cos(x)/sin(x)) \pi/2\ 0" by (rule tendsto_cong_limit, (rule tendsto_intros)+, simp_all) lemma lemma_tan_total: assumes "0 < y" shows "\x. 0 < x \ x < pi/2 \ y < tan x" proof - obtain s where "0 < s" and s: "\x. \x \ pi/2; norm (x - pi/2) < s\ \ norm (cos x / sin x - 0) < inverse y" using LIM_D [OF LIM_cos_div_sin, of "inverse y"] that assms by force obtain e where e: "0 < e" "e < s" "e < pi/2" using \0 < s\ field_lbound_gt_zero pi_half_gt_zero by blast show ?thesis proof (intro exI conjI) have "0 < sin e" "0 < cos e" using e by (auto intro: cos_gt_zero sin_gt_zero2 simp: mult.commute) then show "y < tan (pi/2 - e)" using s [of "pi/2 - e"] e assms by (simp add: tan_def sin_diff cos_diff) (simp add: field_simps split: if_split_asm) qed (use e in auto) qed lemma tan_total_pos: assumes "0 \ y" shows "\x. 0 \ x \ x < pi/2 \ tan x = y" proof (cases "y = 0") case True then show ?thesis using pi_half_gt_zero tan_zero by blast next case False with assms have "y > 0" by linarith obtain x where x: "0 < x" "x < pi/2" "y < tan x" using lemma_tan_total \0 < y\ by blast have "\u\0. u \ x \ tan u = y" proof (intro IVT allI impI) show "isCont tan u" if "0 \ u \ u \ x" for u proof - have "cos u \ 0" using antisym_conv2 cos_gt_zero that x(2) by fastforce with assms show ?thesis by (auto intro!: DERIV_tan [THEN DERIV_isCont]) qed qed (use assms x in auto) then show ?thesis using x(2) by auto qed lemma lemma_tan_total1: "\x. -(pi/2) < x \ x < (pi/2) \ tan x = y" proof (cases "0::real" y rule: le_cases) case le then show ?thesis by (meson less_le_trans minus_pi_half_less_zero tan_total_pos) next case ge - with tan_total_pos [of "-y"] obtain x where "0 \ x" "x < pi / 2" "tan x = - y" + with tan_total_pos [of "-y"] obtain x where "0 \ x" "x < pi/2" "tan x = - y" by force then show ?thesis by (rule_tac x="-x" in exI) auto qed proposition tan_total: "\! x. -(pi/2) < x \ x < (pi/2) \ tan x = y" proof - - have "u = v" if u: "- (pi / 2) < u" "u < pi / 2" and v: "- (pi / 2) < v" "v < pi / 2" + have "u = v" if u: "- (pi/2) < u" "u < pi/2" and v: "- (pi/2) < v" "v < pi/2" and eq: "tan u = tan v" for u v proof (cases u v rule: linorder_cases) case less have "\x. u \ x \ x \ v \ isCont tan x" by (metis cos_gt_zero_pi isCont_tan le_less_trans less_irrefl less_le_trans u(1) v(2)) then have "continuous_on {u..v} tan" by (simp add: continuous_at_imp_continuous_on) moreover have "\x. u < x \ x < v \ tan differentiable (at x)" by (metis DERIV_tan cos_gt_zero_pi real_differentiable_def less_numeral_extra(3) order.strict_trans u(1) v(2)) ultimately obtain z where "u < z" "z < v" "DERIV tan z :> 0" by (metis less Rolle eq) moreover have "cos z \ 0" by (metis (no_types) \u < z\ \z < v\ cos_gt_zero_pi less_le_trans linorder_not_less not_less_iff_gr_or_eq u(1) v(2)) ultimately show ?thesis using DERIV_unique [OF _ DERIV_tan] by fastforce next case greater have "\x. v \ x \ x \ u \ isCont tan x" by (metis cos_gt_zero_pi isCont_tan le_less_trans less_irrefl less_le_trans u(2) v(1)) then have "continuous_on {v..u} tan" by (simp add: continuous_at_imp_continuous_on) moreover have "\x. v < x \ x < u \ tan differentiable (at x)" by (metis DERIV_tan cos_gt_zero_pi real_differentiable_def less_numeral_extra(3) order.strict_trans u(2) v(1)) ultimately obtain z where "v < z" "z < u" "DERIV tan z :> 0" by (metis greater Rolle eq) moreover have "cos z \ 0" by (metis \v < z\ \z < u\ cos_gt_zero_pi less_eq_real_def less_le_trans order_less_irrefl u(2) v(1)) ultimately show ?thesis using DERIV_unique [OF _ DERIV_tan] by fastforce qed auto - then have "\!x. - (pi / 2) < x \ x < pi / 2 \ tan x = y" - if x: "- (pi / 2) < x" "x < pi / 2" "tan x = y" for x + then have "\!x. - (pi/2) < x \ x < pi/2 \ tan x = y" + if x: "- (pi/2) < x" "x < pi/2" "tan x = y" for x using that by auto then show ?thesis using lemma_tan_total1 [where y = y] by auto qed lemma tan_monotone: assumes "- (pi/2) < y" and "y < x" and "x < pi/2" shows "tan y < tan x" proof - have "DERIV tan x' :> inverse ((cos x')\<^sup>2)" if "y \ x'" "x' \ x" for x' proof - have "-(pi/2) < x'" and "x' < pi/2" using that assms by auto with cos_gt_zero_pi have "cos x' \ 0" by force then show "DERIV tan x' :> inverse ((cos x')\<^sup>2)" by (rule DERIV_tan) qed from MVT2[OF \y < x\ this] obtain z where "y < z" and "z < x" and tan_diff: "tan x - tan y = (x - y) * inverse ((cos z)\<^sup>2)" by auto then have "- (pi/2) < z" and "z < pi/2" using assms by auto then have "0 < cos z" using cos_gt_zero_pi by auto then have inv_pos: "0 < inverse ((cos z)\<^sup>2)" by auto have "0 < x - y" using \y < x\ by auto with inv_pos have "0 < tan x - tan y" unfolding tan_diff by auto then show ?thesis by auto qed lemma tan_monotone': assumes "- (pi/2) < y" and "y < pi/2" and "- (pi/2) < x" and "x < pi/2" shows "y < x \ tan y < tan x" proof assume "y < x" then show "tan y < tan x" using tan_monotone and \- (pi/2) < y\ and \x < pi/2\ by auto next assume "tan y < tan x" show "y < x" proof (rule ccontr) assume "\ ?thesis" then have "x \ y" by auto then have "tan x \ tan y" proof (cases "x = y") case True then show ?thesis by auto next case False then have "x < y" using \x \ y\ by auto from tan_monotone[OF \- (pi/2) < x\ this \y < pi/2\] show ?thesis by auto qed then show False using \tan y < tan x\ by auto qed qed lemma tan_inverse: "1 / (tan y) = tan (pi/2 - y)" unfolding tan_def sin_cos_eq[of y] cos_sin_eq[of y] by auto lemma tan_periodic_pi[simp]: "tan (x + pi) = tan x" by (simp add: tan_def) lemma tan_periodic_nat[simp]: "tan (x + real n * pi) = tan x" for n :: nat proof (induct n arbitrary: x) case 0 then show ?case by simp next case (Suc n) have split_pi_off: "x + real (Suc n) * pi = (x + real n * pi) + pi" unfolding Suc_eq_plus1 of_nat_add distrib_right by auto show ?case unfolding split_pi_off using Suc by auto qed lemma tan_periodic_int[simp]: "tan (x + of_int i * pi) = tan x" proof (cases "0 \ i") case True then have i_nat: "of_int i = of_int (nat i)" by auto show ?thesis unfolding i_nat by (metis of_int_of_nat_eq tan_periodic_nat) next case False then have i_nat: "of_int i = - of_int (nat (- i))" by auto have "tan x = tan (x + of_int i * pi - of_int i * pi)" by auto also have "\ = tan (x + of_int i * pi)" unfolding i_nat mult_minus_left diff_minus_eq_add by (metis of_int_of_nat_eq tan_periodic_nat) finally show ?thesis by auto qed lemma tan_periodic_n[simp]: "tan (x + numeral n * pi) = tan x" using tan_periodic_int[of _ "numeral n" ] by simp lemma tan_minus_45: "tan (-(pi/4)) = -1" unfolding tan_def by (simp add: sin_45 cos_45) lemma tan_diff: "cos x \ 0 \ cos y \ 0 \ cos (x - y) \ 0 \ tan (x - y) = (tan x - tan y)/(1 + tan x * tan y)" for x :: "'a::{real_normed_field,banach}" using tan_add [of x "-y"] by simp lemma tan_pos_pi2_le: "0 \ x \ x < pi/2 \ 0 \ tan x" using less_eq_real_def tan_gt_zero by auto lemma cos_tan: "\x\ < pi/2 \ cos x = 1 / sqrt (1 + tan x ^ 2)" using cos_gt_zero_pi [of x] by (simp add: field_split_simps tan_def real_sqrt_divide abs_if split: if_split_asm) lemma cos_tan_half: "cos x \0 \ cos (2*x) = (1 - (tan x)^2) / (1 + (tan x)^2)" unfolding cos_double tan_def by (auto simp add:field_simps ) lemma sin_tan: "\x\ < pi/2 \ sin x = tan x / sqrt (1 + tan x ^ 2)" using cos_gt_zero [of "x"] cos_gt_zero [of "-x"] by (force simp: field_split_simps tan_def real_sqrt_divide abs_if split: if_split_asm) lemma sin_tan_half: "sin (2*x) = 2 * tan x / (1 + (tan x)^2)" unfolding sin_double tan_def by (cases "cos x=0") (auto simp add:field_simps power2_eq_square) lemma tan_mono_le: "-(pi/2) < x \ x \ y \ y < pi/2 \ tan x \ tan y" using less_eq_real_def tan_monotone by auto lemma tan_mono_lt_eq: "-(pi/2) < x \ x < pi/2 \ -(pi/2) < y \ y < pi/2 \ tan x < tan y \ x < y" using tan_monotone' by blast lemma tan_mono_le_eq: "-(pi/2) < x \ x < pi/2 \ -(pi/2) < y \ y < pi/2 \ tan x \ tan y \ x \ y" by (meson tan_mono_le not_le tan_monotone) lemma tan_bound_pi2: "\x\ < pi/4 \ \tan x\ < 1" using tan_45 tan_monotone [of x "pi/4"] tan_monotone [of "-x" "pi/4"] by (auto simp: abs_if split: if_split_asm) lemma tan_cot: "tan(pi/2 - x) = inverse(tan x)" by (simp add: tan_def sin_diff cos_diff) subsection \Cotangent\ definition cot :: "'a \ 'a::{real_normed_field,banach}" where "cot = (\x. cos x / sin x)" lemma cot_of_real: "of_real (cot x) = (cot (of_real x) :: 'a::{real_normed_field,banach})" by (simp add: cot_def sin_of_real cos_of_real) lemma cot_in_Reals [simp]: "z \ \ \ cot z \ \" for z :: "'a::{real_normed_field,banach}" by (simp add: cot_def) lemma cot_zero [simp]: "cot 0 = 0" by (simp add: cot_def) lemma cot_pi [simp]: "cot pi = 0" by (simp add: cot_def) lemma cot_npi [simp]: "cot (real n * pi) = 0" for n :: nat by (simp add: cot_def) lemma cot_minus [simp]: "cot (- x) = - cot x" by (simp add: cot_def) lemma cot_periodic [simp]: "cot (x + 2 * pi) = cot x" by (simp add: cot_def) lemma cot_altdef: "cot x = inverse (tan x)" by (simp add: cot_def tan_def) lemma tan_altdef: "tan x = inverse (cot x)" by (simp add: cot_def tan_def) lemma tan_cot': "tan (pi/2 - x) = cot x" by (simp add: tan_cot cot_altdef) lemma cot_gt_zero: "0 < x \ x < pi/2 \ 0 < cot x" by (simp add: cot_def zero_less_divide_iff sin_gt_zero2 cos_gt_zero_pi) lemma cot_less_zero: assumes lb: "- pi/2 < x" and "x < 0" shows "cot x < 0" proof - have "0 < cot (- x)" using assms by (simp only: cot_gt_zero) then show ?thesis by simp qed lemma DERIV_cot [simp]: "sin x \ 0 \ DERIV cot x :> -inverse ((sin x)\<^sup>2)" for x :: "'a::{real_normed_field,banach}" unfolding cot_def using cos_squared_eq[of x] by (auto intro!: derivative_eq_intros) (simp add: divide_inverse power2_eq_square) lemma isCont_cot: "sin x \ 0 \ isCont cot x" for x :: "'a::{real_normed_field,banach}" by (rule DERIV_cot [THEN DERIV_isCont]) lemma isCont_cot' [simp,continuous_intros]: "isCont f a \ sin (f a) \ 0 \ isCont (\x. cot (f x)) a" for a :: "'a::{real_normed_field,banach}" and f :: "'a \ 'a" by (rule isCont_o2 [OF _ isCont_cot]) lemma tendsto_cot [tendsto_intros]: "(f \ a) F \ sin a \ 0 \ ((\x. cot (f x)) \ cot a) F" for f :: "'a \ 'a::{real_normed_field,banach}" by (rule isCont_tendsto_compose [OF isCont_cot]) lemma continuous_cot: "continuous F f \ sin (f (Lim F (\x. x))) \ 0 \ continuous F (\x. cot (f x))" for f :: "'a \ 'a::{real_normed_field,banach}" unfolding continuous_def by (rule tendsto_cot) lemma continuous_on_cot [continuous_intros]: fixes f :: "'a \ 'a::{real_normed_field,banach}" shows "continuous_on s f \ (\x\s. sin (f x) \ 0) \ continuous_on s (\x. cot (f x))" unfolding continuous_on_def by (auto intro: tendsto_cot) lemma continuous_within_cot [continuous_intros]: fixes f :: "'a \ 'a::{real_normed_field,banach}" shows "continuous (at x within s) f \ sin (f x) \ 0 \ continuous (at x within s) (\x. cot (f x))" unfolding continuous_within by (rule tendsto_cot) subsection \Inverse Trigonometric Functions\ definition arcsin :: "real \ real" where "arcsin y = (THE x. -(pi/2) \ x \ x \ pi/2 \ sin x = y)" definition arccos :: "real \ real" where "arccos y = (THE x. 0 \ x \ x \ pi \ cos x = y)" definition arctan :: "real \ real" where "arctan y = (THE x. -(pi/2) < x \ x < pi/2 \ tan x = y)" lemma arcsin: "- 1 \ y \ y \ 1 \ - (pi/2) \ arcsin y \ arcsin y \ pi/2 \ sin (arcsin y) = y" unfolding arcsin_def by (rule theI' [OF sin_total]) lemma arcsin_pi: "- 1 \ y \ y \ 1 \ - (pi/2) \ arcsin y \ arcsin y \ pi \ sin (arcsin y) = y" by (drule (1) arcsin) (force intro: order_trans) lemma sin_arcsin [simp]: "- 1 \ y \ y \ 1 \ sin (arcsin y) = y" by (blast dest: arcsin) lemma arcsin_bounded: "- 1 \ y \ y \ 1 \ - (pi/2) \ arcsin y \ arcsin y \ pi/2" by (blast dest: arcsin) lemma arcsin_lbound: "- 1 \ y \ y \ 1 \ - (pi/2) \ arcsin y" by (blast dest: arcsin) lemma arcsin_ubound: "- 1 \ y \ y \ 1 \ arcsin y \ pi/2" by (blast dest: arcsin) lemma arcsin_lt_bounded: assumes "- 1 < y" "y < 1" shows "- (pi/2) < arcsin y \ arcsin y < pi/2" proof - have "arcsin y \ pi/2" by (metis arcsin assms not_less not_less_iff_gr_or_eq sin_pi_half) moreover have "arcsin y \ - pi/2" by (metis arcsin assms minus_divide_left not_less not_less_iff_gr_or_eq sin_minus sin_pi_half) ultimately show ?thesis using arcsin_bounded [of y] assms by auto qed lemma arcsin_sin: "- (pi/2) \ x \ x \ pi/2 \ arcsin (sin x) = x" unfolding arcsin_def using the1_equality [OF sin_total] by simp +lemma arcsin_unique: + assumes "-pi/2 \ x" and "x \ pi/2" and "sin x = y" shows "arcsin y = x" + using arcsin_sin[of x] assms by force + lemma arcsin_0 [simp]: "arcsin 0 = 0" using arcsin_sin [of 0] by simp lemma arcsin_1 [simp]: "arcsin 1 = pi/2" using arcsin_sin [of "pi/2"] by simp lemma arcsin_minus_1 [simp]: "arcsin (- 1) = - (pi/2)" using arcsin_sin [of "- pi/2"] by simp lemma arcsin_minus: "- 1 \ x \ x \ 1 \ arcsin (- x) = - arcsin x" by (metis (no_types, opaque_lifting) arcsin arcsin_sin minus_minus neg_le_iff_le sin_minus) +lemma arcsin_one_half [simp]: "arcsin (1/2) = pi / 6" + and arcsin_minus_one_half [simp]: "arcsin (-(1/2)) = -pi / 6" + by (intro arcsin_unique; simp add: sin_30 field_simps)+ + +lemma arcsin_one_over_sqrt_2: "arcsin (1 / sqrt 2) = pi / 4" + by (rule arcsin_unique) (auto simp: sin_45 field_simps) + lemma arcsin_eq_iff: "\x\ \ 1 \ \y\ \ 1 \ arcsin x = arcsin y \ x = y" by (metis abs_le_iff arcsin minus_le_iff) lemma cos_arcsin_nonzero: "- 1 < x \ x < 1 \ cos (arcsin x) \ 0" using arcsin_lt_bounded cos_gt_zero_pi by force lemma arccos: "- 1 \ y \ y \ 1 \ 0 \ arccos y \ arccos y \ pi \ cos (arccos y) = y" unfolding arccos_def by (rule theI' [OF cos_total]) lemma cos_arccos [simp]: "- 1 \ y \ y \ 1 \ cos (arccos y) = y" by (blast dest: arccos) lemma arccos_bounded: "- 1 \ y \ y \ 1 \ 0 \ arccos y \ arccos y \ pi" by (blast dest: arccos) lemma arccos_lbound: "- 1 \ y \ y \ 1 \ 0 \ arccos y" by (blast dest: arccos) lemma arccos_ubound: "- 1 \ y \ y \ 1 \ arccos y \ pi" by (blast dest: arccos) lemma arccos_lt_bounded: assumes "- 1 < y" "y < 1" shows "0 < arccos y \ arccos y < pi" proof - have "arccos y \ 0" by (metis (no_types) arccos assms(1) assms(2) cos_zero less_eq_real_def less_irrefl) moreover have "arccos y \ -pi" by (metis arccos assms(1) assms(2) cos_minus cos_pi not_less not_less_iff_gr_or_eq) ultimately show ?thesis using arccos_bounded [of y] assms by (metis arccos cos_pi not_less not_less_iff_gr_or_eq) qed lemma arccos_cos: "0 \ x \ x \ pi \ arccos (cos x) = x" by (auto simp: arccos_def intro!: the1_equality cos_total) lemma arccos_cos2: "x \ 0 \ - pi \ x \ arccos (cos x) = -x" by (auto simp: arccos_def intro!: the1_equality cos_total) +lemma arccos_unique: + assumes "0 \ x" and "x \ pi" and "cos x = y" shows "arccos y = x" + using arccos_cos assms by blast + lemma cos_arcsin: assumes "- 1 \ x" "x \ 1" shows "cos (arcsin x) = sqrt (1 - x\<^sup>2)" proof (rule power2_eq_imp_eq) show "(cos (arcsin x))\<^sup>2 = (sqrt (1 - x\<^sup>2))\<^sup>2" by (simp add: square_le_1 assms cos_squared_eq) show "0 \ cos (arcsin x)" using arcsin assms cos_ge_zero by blast show "0 \ sqrt (1 - x\<^sup>2)" by (simp add: square_le_1 assms) qed lemma sin_arccos: assumes "- 1 \ x" "x \ 1" shows "sin (arccos x) = sqrt (1 - x\<^sup>2)" proof (rule power2_eq_imp_eq) show "(sin (arccos x))\<^sup>2 = (sqrt (1 - x\<^sup>2))\<^sup>2" by (simp add: square_le_1 assms sin_squared_eq) show "0 \ sin (arccos x)" by (simp add: arccos_bounded assms sin_ge_zero) show "0 \ sqrt (1 - x\<^sup>2)" by (simp add: square_le_1 assms) qed lemma arccos_0 [simp]: "arccos 0 = pi/2" - by (metis arccos_cos cos_gt_zero cos_pi cos_pi_half pi_gt_zero - pi_half_ge_zero not_le not_zero_less_neg_numeral numeral_One) + using arccos_cos pi_half_ge_zero by fastforce lemma arccos_1 [simp]: "arccos 1 = 0" using arccos_cos by force lemma arccos_minus_1 [simp]: "arccos (- 1) = pi" by (metis arccos_cos cos_pi order_refl pi_ge_zero) lemma arccos_minus: "-1 \ x \ x \ 1 \ arccos (- x) = pi - arccos x" - by (metis arccos_cos arccos_cos2 cos_minus_pi cos_total diff_le_0_iff_le le_add_same_cancel1 - minus_diff_eq uminus_add_conv_diff) + by (smt (verit, ccfv_threshold) arccos arccos_cos cos_minus cos_minus_pi) + +lemma arccos_one_half [simp]: "arccos (1/2) = pi / 3" + and arccos_minus_one_half [simp]: "arccos (-(1/2)) = 2 * pi / 3" + by (intro arccos_unique; simp add: cos_60 cos_120)+ + +lemma arccos_one_over_sqrt_2: "arccos (1 / sqrt 2) = pi / 4" + by (rule arccos_unique) (auto simp: cos_45 field_simps) corollary arccos_minus_abs: assumes "\x\ \ 1" shows "arccos (- x) = pi - arccos x" using assms by (simp add: arccos_minus) lemma sin_arccos_nonzero: "- 1 < x \ x < 1 \ sin (arccos x) \ 0" using arccos_lt_bounded sin_gt_zero by force lemma arctan: "- (pi/2) < arctan y \ arctan y < pi/2 \ tan (arctan y) = y" unfolding arctan_def by (rule theI' [OF tan_total]) lemma tan_arctan: "tan (arctan y) = y" by (simp add: arctan) lemma arctan_bounded: "- (pi/2) < arctan y \ arctan y < pi/2" by (auto simp only: arctan) lemma arctan_lbound: "- (pi/2) < arctan y" by (simp add: arctan) lemma arctan_ubound: "arctan y < pi/2" by (auto simp only: arctan) lemma arctan_unique: assumes "-(pi/2) < x" and "x < pi/2" and "tan x = y" shows "arctan y = x" using assms arctan [of y] tan_total [of y] by (fast elim: ex1E) lemma arctan_tan: "-(pi/2) < x \ x < pi/2 \ arctan (tan x) = x" by (rule arctan_unique) simp_all lemma arctan_zero_zero [simp]: "arctan 0 = 0" by (rule arctan_unique) simp_all lemma arctan_minus: "arctan (- x) = - arctan x" using arctan [of "x"] by (auto simp: arctan_unique) lemma cos_arctan_not_zero [simp]: "cos (arctan x) \ 0" by (intro less_imp_neq [symmetric] cos_gt_zero_pi arctan_lbound arctan_ubound) lemma cos_arctan: "cos (arctan x) = 1 / sqrt (1 + x\<^sup>2)" proof (rule power2_eq_imp_eq) have "0 < 1 + x\<^sup>2" by (simp add: add_pos_nonneg) show "0 \ 1 / sqrt (1 + x\<^sup>2)" by simp show "0 \ cos (arctan x)" by (intro less_imp_le cos_gt_zero_pi arctan_lbound arctan_ubound) have "(cos (arctan x))\<^sup>2 * (1 + (tan (arctan x))\<^sup>2) = 1" unfolding tan_def by (simp add: distrib_left power_divide) then show "(cos (arctan x))\<^sup>2 = (1 / sqrt (1 + x\<^sup>2))\<^sup>2" using \0 < 1 + x\<^sup>2\ by (simp add: arctan power_divide eq_divide_eq) qed lemma sin_arctan: "sin (arctan x) = x / sqrt (1 + x\<^sup>2)" using add_pos_nonneg [OF zero_less_one zero_le_power2 [of x]] using tan_arctan [of x] unfolding tan_def cos_arctan by (simp add: eq_divide_eq) lemma tan_sec: "cos x \ 0 \ 1 + (tan x)\<^sup>2 = (inverse (cos x))\<^sup>2" for x :: "'a::{real_normed_field,banach,field}" by (simp add: add_divide_eq_iff inverse_eq_divide power2_eq_square tan_def) lemma arctan_less_iff: "arctan x < arctan y \ x < y" by (metis tan_monotone' arctan_lbound arctan_ubound tan_arctan) lemma arctan_le_iff: "arctan x \ arctan y \ x \ y" by (simp only: not_less [symmetric] arctan_less_iff) lemma arctan_eq_iff: "arctan x = arctan y \ x = y" by (simp only: eq_iff [where 'a=real] arctan_le_iff) lemma zero_less_arctan_iff [simp]: "0 < arctan x \ 0 < x" using arctan_less_iff [of 0 x] by simp lemma arctan_less_zero_iff [simp]: "arctan x < 0 \ x < 0" using arctan_less_iff [of x 0] by simp lemma zero_le_arctan_iff [simp]: "0 \ arctan x \ 0 \ x" using arctan_le_iff [of 0 x] by simp lemma arctan_le_zero_iff [simp]: "arctan x \ 0 \ x \ 0" using arctan_le_iff [of x 0] by simp lemma arctan_eq_zero_iff [simp]: "arctan x = 0 \ x = 0" using arctan_eq_iff [of x 0] by simp lemma continuous_on_arcsin': "continuous_on {-1 .. 1} arcsin" proof - have "continuous_on (sin ` {- pi/2 .. pi/2}) arcsin" by (rule continuous_on_inv) (auto intro: continuous_intros simp: arcsin_sin) also have "sin ` {- pi/2 .. pi/2} = {-1 .. 1}" proof safe fix x :: real assume "x \ {-1..1}" then show "x \ sin ` {- pi/2..pi/2}" using arcsin_lbound arcsin_ubound by (intro image_eqI[where x="arcsin x"]) auto qed simp finally show ?thesis . qed lemma continuous_on_arcsin [continuous_intros]: "continuous_on s f \ (\x\s. -1 \ f x \ f x \ 1) \ continuous_on s (\x. arcsin (f x))" using continuous_on_compose[of s f, OF _ continuous_on_subset[OF continuous_on_arcsin']] by (auto simp: comp_def subset_eq) lemma isCont_arcsin: "-1 < x \ x < 1 \ isCont arcsin x" using continuous_on_arcsin'[THEN continuous_on_subset, of "{ -1 <..< 1 }"] by (auto simp: continuous_on_eq_continuous_at subset_eq) lemma continuous_on_arccos': "continuous_on {-1 .. 1} arccos" proof - have "continuous_on (cos ` {0 .. pi}) arccos" by (rule continuous_on_inv) (auto intro: continuous_intros simp: arccos_cos) also have "cos ` {0 .. pi} = {-1 .. 1}" proof safe fix x :: real assume "x \ {-1..1}" then show "x \ cos ` {0..pi}" using arccos_lbound arccos_ubound by (intro image_eqI[where x="arccos x"]) auto qed simp finally show ?thesis . qed lemma continuous_on_arccos [continuous_intros]: "continuous_on s f \ (\x\s. -1 \ f x \ f x \ 1) \ continuous_on s (\x. arccos (f x))" using continuous_on_compose[of s f, OF _ continuous_on_subset[OF continuous_on_arccos']] by (auto simp: comp_def subset_eq) lemma isCont_arccos: "-1 < x \ x < 1 \ isCont arccos x" using continuous_on_arccos'[THEN continuous_on_subset, of "{ -1 <..< 1 }"] by (auto simp: continuous_on_eq_continuous_at subset_eq) lemma isCont_arctan: "isCont arctan x" proof - - obtain u where u: "- (pi / 2) < u" "u < arctan x" + obtain u where u: "- (pi/2) < u" "u < arctan x" by (meson arctan arctan_less_iff linordered_field_no_lb) - obtain v where v: "arctan x < v" "v < pi / 2" + obtain v where v: "arctan x < v" "v < pi/2" by (meson arctan_less_iff arctan_ubound linordered_field_no_ub) have "isCont arctan (tan (arctan x))" proof (rule isCont_inverse_function2 [of u "arctan x" v]) show "\z. \u \ z; z \ v\ \ arctan (tan z) = z" using arctan_unique u(1) v(2) by auto then show "\z. \u \ z; z \ v\ \ isCont tan z" by (metis arctan cos_gt_zero_pi isCont_tan less_irrefl) qed (use u v in auto) then show ?thesis by (simp add: arctan) qed lemma tendsto_arctan [tendsto_intros]: "(f \ x) F \ ((\x. arctan (f x)) \ arctan x) F" by (rule isCont_tendsto_compose [OF isCont_arctan]) lemma continuous_arctan [continuous_intros]: "continuous F f \ continuous F (\x. arctan (f x))" unfolding continuous_def by (rule tendsto_arctan) lemma continuous_on_arctan [continuous_intros]: "continuous_on s f \ continuous_on s (\x. arctan (f x))" unfolding continuous_on_def by (auto intro: tendsto_arctan) lemma DERIV_arcsin: assumes "- 1 < x" "x < 1" shows "DERIV arcsin x :> inverse (sqrt (1 - x\<^sup>2))" proof (rule DERIV_inverse_function) show "(sin has_real_derivative sqrt (1 - x\<^sup>2)) (at (arcsin x))" by (rule derivative_eq_intros | use assms cos_arcsin in force)+ show "sqrt (1 - x\<^sup>2) \ 0" using abs_square_eq_1 assms by force qed (use assms isCont_arcsin in auto) lemma DERIV_arccos: assumes "- 1 < x" "x < 1" shows "DERIV arccos x :> inverse (- sqrt (1 - x\<^sup>2))" proof (rule DERIV_inverse_function) show "(cos has_real_derivative - sqrt (1 - x\<^sup>2)) (at (arccos x))" by (rule derivative_eq_intros | use assms sin_arccos in force)+ show "- sqrt (1 - x\<^sup>2) \ 0" using abs_square_eq_1 assms by force qed (use assms isCont_arccos in auto) lemma DERIV_arctan: "DERIV arctan x :> inverse (1 + x\<^sup>2)" proof (rule DERIV_inverse_function) have "inverse ((cos (arctan x))\<^sup>2) = 1 + x\<^sup>2" by (metis arctan cos_arctan_not_zero power_inverse tan_sec) then show "(tan has_real_derivative 1 + x\<^sup>2) (at (arctan x))" by (auto intro!: derivative_eq_intros) show "\y. \x - 1 < y; y < x + 1\ \ tan (arctan y) = y" using tan_arctan by blast show "1 + x\<^sup>2 \ 0" by (metis power_one sum_power2_eq_zero_iff zero_neq_one) qed (use isCont_arctan in auto) declare DERIV_arcsin[THEN DERIV_chain2, derivative_intros] DERIV_arcsin[THEN DERIV_chain2, unfolded has_field_derivative_def, derivative_intros] DERIV_arccos[THEN DERIV_chain2, derivative_intros] DERIV_arccos[THEN DERIV_chain2, unfolded has_field_derivative_def, derivative_intros] DERIV_arctan[THEN DERIV_chain2, derivative_intros] DERIV_arctan[THEN DERIV_chain2, unfolded has_field_derivative_def, derivative_intros] lemmas has_derivative_arctan[derivative_intros] = DERIV_arctan[THEN DERIV_compose_FDERIV] and has_derivative_arccos[derivative_intros] = DERIV_arccos[THEN DERIV_compose_FDERIV] and has_derivative_arcsin[derivative_intros] = DERIV_arcsin[THEN DERIV_compose_FDERIV] lemma filterlim_tan_at_right: "filterlim tan at_bot (at_right (- (pi/2)))" by (rule filterlim_at_bot_at_right[where Q="\x. - pi/2 < x \ x < pi/2" and P="\x. True" and g=arctan]) (auto simp: arctan le_less eventually_at dist_real_def simp del: less_divide_eq_numeral1 intro!: tan_monotone exI[of _ "pi/2"]) lemma filterlim_tan_at_left: "filterlim tan at_top (at_left (pi/2))" by (rule filterlim_at_top_at_left[where Q="\x. - pi/2 < x \ x < pi/2" and P="\x. True" and g=arctan]) (auto simp: arctan le_less eventually_at dist_real_def simp del: less_divide_eq_numeral1 intro!: tan_monotone exI[of _ "pi/2"]) lemma tendsto_arctan_at_top: "(arctan \ (pi/2)) at_top" proof (rule tendstoI) fix e :: real assume "0 < e" define y where "y = pi/2 - min (pi/2) e" then have y: "0 \ y" "y < pi/2" "pi/2 \ e + y" using \0 < e\ by auto show "eventually (\x. dist (arctan x) (pi/2) < e) at_top" proof (intro eventually_at_top_dense[THEN iffD2] exI allI impI) fix x assume "tan y < x" then have "arctan (tan y) < arctan x" by (simp add: arctan_less_iff) with y have "y < arctan x" by (subst (asm) arctan_tan) simp_all with arctan_ubound[of x, arith] y \0 < e\ show "dist (arctan x) (pi/2) < e" by (simp add: dist_real_def) qed qed lemma tendsto_arctan_at_bot: "(arctan \ - (pi/2)) at_bot" unfolding filterlim_at_bot_mirror arctan_minus by (intro tendsto_minus tendsto_arctan_at_top) subsection \Prove Totality of the Trigonometric Functions\ lemma cos_arccos_abs: "\y\ \ 1 \ cos (arccos y) = y" by (simp add: abs_le_iff) lemma sin_arccos_abs: "\y\ \ 1 \ sin (arccos y) = sqrt (1 - y\<^sup>2)" by (simp add: sin_arccos abs_le_iff) lemma sin_mono_less_eq: "- (pi/2) \ x \ x \ pi/2 \ - (pi/2) \ y \ y \ pi/2 \ sin x < sin y \ x < y" by (metis not_less_iff_gr_or_eq sin_monotone_2pi) lemma sin_mono_le_eq: "- (pi/2) \ x \ x \ pi/2 \ - (pi/2) \ y \ y \ pi/2 \ sin x \ sin y \ x \ y" by (meson leD le_less_linear sin_monotone_2pi sin_monotone_2pi_le) lemma sin_inj_pi: "- (pi/2) \ x \ x \ pi/2 \ - (pi/2) \ y \ y \ pi/2 \ sin x = sin y \ x = y" by (metis arcsin_sin) lemma arcsin_le_iff: assumes "x \ -1" "x \ 1" "y \ -pi/2" "y \ pi/2" shows "arcsin x \ y \ x \ sin y" proof - have "arcsin x \ y \ sin (arcsin x) \ sin y" using arcsin_bounded[of x] assms by (subst sin_mono_le_eq) auto also from assms have "sin (arcsin x) = x" by simp finally show ?thesis . qed lemma le_arcsin_iff: assumes "x \ -1" "x \ 1" "y \ -pi/2" "y \ pi/2" shows "arcsin x \ y \ x \ sin y" proof - have "arcsin x \ y \ sin (arcsin x) \ sin y" using arcsin_bounded[of x] assms by (subst sin_mono_le_eq) auto also from assms have "sin (arcsin x) = x" by simp finally show ?thesis . qed lemma cos_mono_less_eq: "0 \ x \ x \ pi \ 0 \ y \ y \ pi \ cos x < cos y \ y < x" by (meson cos_monotone_0_pi cos_monotone_0_pi_le leD le_less_linear) lemma cos_mono_le_eq: "0 \ x \ x \ pi \ 0 \ y \ y \ pi \ cos x \ cos y \ y \ x" by (metis arccos_cos cos_monotone_0_pi_le eq_iff linear) lemma cos_inj_pi: "0 \ x \ x \ pi \ 0 \ y \ y \ pi \ cos x = cos y \ x = y" by (metis arccos_cos) lemma arccos_le_pi2: "\0 \ y; y \ 1\ \ arccos y \ pi/2" by (metis (mono_tags) arccos_0 arccos cos_le_one cos_monotone_0_pi_le cos_pi cos_pi_half pi_half_ge_zero antisym_conv less_eq_neg_nonpos linear minus_minus order.trans order_refl) lemma sincos_total_pi_half: assumes "0 \ x" "0 \ y" "x\<^sup>2 + y\<^sup>2 = 1" shows "\t. 0 \ t \ t \ pi/2 \ x = cos t \ y = sin t" proof - have x1: "x \ 1" using assms by (metis le_add_same_cancel1 power2_le_imp_le power_one zero_le_power2) with assms have *: "0 \ arccos x" "cos (arccos x) = x" by (auto simp: arccos) from assms have "y = sqrt (1 - x\<^sup>2)" by (metis abs_of_nonneg add.commute add_diff_cancel real_sqrt_abs) with x1 * assms arccos_le_pi2 [of x] show ?thesis by (rule_tac x="arccos x" in exI) (auto simp: sin_arccos) qed lemma sincos_total_pi: assumes "0 \ y" "x\<^sup>2 + y\<^sup>2 = 1" shows "\t. 0 \ t \ t \ pi \ x = cos t \ y = sin t" proof (cases rule: le_cases [of 0 x]) case le from sincos_total_pi_half [OF le] show ?thesis by (metis pi_ge_two pi_half_le_two add.commute add_le_cancel_left add_mono assms) next case ge then have "0 \ -x" by simp then obtain t where t: "t\0" "t \ pi/2" "-x = cos t" "y = sin t" using sincos_total_pi_half assms by auto (metis \0 \ - x\ power2_minus) show ?thesis by (rule exI [where x = "pi -t"]) (use t in auto) qed lemma sincos_total_2pi_le: assumes "x\<^sup>2 + y\<^sup>2 = 1" shows "\t. 0 \ t \ t \ 2 * pi \ x = cos t \ y = sin t" proof (cases rule: le_cases [of 0 y]) case le from sincos_total_pi [OF le] show ?thesis by (metis assms le_add_same_cancel1 mult.commute mult_2_right order.trans) next case ge then have "0 \ -y" by simp then obtain t where t: "t\0" "t \ pi" "x = cos t" "-y = sin t" using sincos_total_pi assms by auto (metis \0 \ - y\ power2_minus) show ?thesis by (rule exI [where x = "2 * pi - t"]) (use t in auto) qed lemma sincos_total_2pi: assumes "x\<^sup>2 + y\<^sup>2 = 1" obtains t where "0 \ t" "t < 2*pi" "x = cos t" "y = sin t" proof - from sincos_total_2pi_le [OF assms] obtain t where t: "0 \ t" "t \ 2*pi" "x = cos t" "y = sin t" by blast show ?thesis by (cases "t = 2 * pi") (use t that in \force+\) qed lemma arcsin_less_mono: "\x\ \ 1 \ \y\ \ 1 \ arcsin x < arcsin y \ x < y" by (rule trans [OF sin_mono_less_eq [symmetric]]) (use arcsin_ubound arcsin_lbound in auto) lemma arcsin_le_mono: "\x\ \ 1 \ \y\ \ 1 \ arcsin x \ arcsin y \ x \ y" using arcsin_less_mono not_le by blast lemma arcsin_less_arcsin: "- 1 \ x \ x < y \ y \ 1 \ arcsin x < arcsin y" using arcsin_less_mono by auto lemma arcsin_le_arcsin: "- 1 \ x \ x \ y \ y \ 1 \ arcsin x \ arcsin y" using arcsin_le_mono by auto +lemma arcsin_nonneg: "x \ {0..1} \ arcsin x \ 0" + using arcsin_le_arcsin[of 0 x] by simp + lemma arccos_less_mono: "\x\ \ 1 \ \y\ \ 1 \ arccos x < arccos y \ y < x" by (rule trans [OF cos_mono_less_eq [symmetric]]) (use arccos_ubound arccos_lbound in auto) lemma arccos_le_mono: "\x\ \ 1 \ \y\ \ 1 \ arccos x \ arccos y \ y \ x" using arccos_less_mono [of y x] by (simp add: not_le [symmetric]) lemma arccos_less_arccos: "- 1 \ x \ x < y \ y \ 1 \ arccos y < arccos x" using arccos_less_mono by auto lemma arccos_le_arccos: "- 1 \ x \ x \ y \ y \ 1 \ arccos y \ arccos x" using arccos_le_mono by auto lemma arccos_eq_iff: "\x\ \ 1 \ \y\ \ 1 \ arccos x = arccos y \ x = y" using cos_arccos_abs by fastforce lemma arccos_cos_eq_abs: assumes "\\\ \ pi" shows "arccos (cos \) = \\\" unfolding arccos_def proof (intro the_equality conjI; clarify?) show "cos \\\ = cos \" by (simp add: abs_real_def) show "x = \\\" if "cos x = cos \" "0 \ x" "x \ pi" for x by (simp add: \cos \\\ = cos \\ assms cos_inj_pi that) qed (use assms in auto) lemma arccos_cos_eq_abs_2pi: obtains k where "arccos (cos \) = \\ - of_int k * (2 * pi)\" proof - define k where "k \ \(\ + pi) / (2 * pi)\" have lepi: "\\ - of_int k * (2 * pi)\ \ pi" using floor_divide_lower [of "2*pi" "\ + pi"] floor_divide_upper [of "2*pi" "\ + pi"] by (auto simp: k_def abs_if algebra_simps) have "arccos (cos \) = arccos (cos (\ - of_int k * (2 * pi)))" using cos_int_2pin sin_int_2pin by (simp add: cos_diff mult.commute) also have "\ = \\ - of_int k * (2 * pi)\" using arccos_cos_eq_abs lepi by blast finally show ?thesis using that by metis qed 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)" + 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" + 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" + 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 arcsin_arccos_sqrt_pos: "0 \ x \ x \ 1 \ arcsin x = arccos(sqrt(1 - x\<^sup>2))" by (smt (verit, del_insts) arccos_cos arcsin_0 arcsin_le_arcsin arcsin_pi cos_arcsin) 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))" by (smt (verit, del_insts) arccos_lbound arccos_le_pi2 arcsin_sin sin_arccos) 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) lemma cos_limit_1: assumes "(\j. cos (\ j)) \ 1" shows "\k. (\j. \ j - of_int (k j) * (2 * pi)) \ 0" proof - have "\\<^sub>F j in sequentially. cos (\ j) \ {- 1..1}" by auto then have "(\j. arccos (cos (\ j))) \ arccos 1" using continuous_on_tendsto_compose [OF continuous_on_arccos' assms] by auto moreover have "\j. \k. arccos (cos (\ j)) = \\ j - of_int k * (2 * pi)\" using arccos_cos_eq_abs_2pi by metis then have "\k. \j. arccos (cos (\ j)) = \\ j - of_int (k j) * (2 * pi)\" by metis ultimately have "\k. (\j. \\ j - of_int (k j) * (2 * pi)\) \ 0" by auto then show ?thesis by (simp add: tendsto_rabs_zero_iff) qed lemma cos_diff_limit_1: assumes "(\j. cos (\ j - \)) \ 1" obtains k where "(\j. \ j - of_int (k j) * (2 * pi)) \ \" proof - obtain k where "(\j. (\ j - \) - of_int (k j) * (2 * pi)) \ 0" using cos_limit_1 [OF assms] by auto then have "(\j. \ + ((\ j - \) - of_int (k j) * (2 * pi))) \ \ + 0" by (rule tendsto_add [OF tendsto_const]) with that show ?thesis by auto qed subsection \Machin's formula\ -lemma arctan_one: "arctan 1 = pi / 4" +lemma arctan_one: "arctan 1 = pi/4" by (rule arctan_unique) (simp_all add: tan_45 m2pi_less_pi) lemma tan_total_pi4: assumes "\x\ < 1" - shows "\z. - (pi / 4) < z \ z < pi / 4 \ tan z = x" + shows "\z. - (pi/4) < z \ z < pi/4 \ tan z = x" proof - show "- (pi / 4) < arctan x \ arctan x < pi / 4 \ tan (arctan x) = x" + show "- (pi/4) < arctan x \ arctan x < pi/4 \ tan (arctan x) = x" unfolding arctan_one [symmetric] arctan_minus [symmetric] unfolding arctan_less_iff using assms by (auto simp: arctan) qed lemma arctan_add: assumes "\x\ \ 1" "\y\ < 1" shows "arctan x + arctan y = arctan ((x + y) / (1 - x * y))" proof (rule arctan_unique [symmetric]) - have "- (pi / 4) \ arctan x" "- (pi / 4) < arctan y" + have "- (pi/4) \ arctan x" "- (pi/4) < arctan y" unfolding arctan_one [symmetric] arctan_minus [symmetric] unfolding arctan_le_iff arctan_less_iff using assms by auto from add_le_less_mono [OF this] show 1: "- (pi/2) < arctan x + arctan y" by simp - have "arctan x \ pi / 4" "arctan y < pi / 4" + have "arctan x \ pi/4" "arctan y < pi/4" unfolding arctan_one [symmetric] unfolding arctan_le_iff arctan_less_iff using assms by auto from add_le_less_mono [OF this] show 2: "arctan x + arctan y < pi/2" by simp show "tan (arctan x + arctan y) = (x + y) / (1 - x * y)" using cos_gt_zero_pi [OF 1 2] by (simp add: arctan tan_add) qed lemma arctan_double: "\x\ < 1 \ 2 * arctan x = arctan ((2 * x) / (1 - x\<^sup>2))" by (metis arctan_add linear mult_2 not_less power2_eq_square) -theorem machin: "pi / 4 = 4 * arctan (1 / 5) - arctan (1 / 239)" +theorem machin: "pi/4 = 4 * arctan (1 / 5) - arctan (1/239)" proof - have "\1 / 5\ < (1 :: real)" by auto from arctan_add[OF less_imp_le[OF this] this] have "2 * arctan (1 / 5) = arctan (5 / 12)" by auto moreover have "\5 / 12\ < (1 :: real)" by auto from arctan_add[OF less_imp_le[OF this] this] have "2 * arctan (5 / 12) = arctan (120 / 119)" by auto moreover - have "\1\ \ (1::real)" and "\1 / 239\ < (1::real)" + have "\1\ \ (1::real)" and "\1/239\ < (1::real)" by auto - from arctan_add[OF this] have "arctan 1 + arctan (1 / 239) = arctan (120 / 119)" + from arctan_add[OF this] have "arctan 1 + arctan (1/239) = arctan (120 / 119)" by auto - ultimately have "arctan 1 + arctan (1 / 239) = 4 * arctan (1 / 5)" + ultimately have "arctan 1 + arctan (1/239) = 4 * arctan (1 / 5)" by auto then show ?thesis unfolding arctan_one by algebra qed -lemma machin_Euler: "5 * arctan (1 / 7) + 2 * arctan (3 / 79) = pi / 4" +lemma machin_Euler: "5 * arctan (1 / 7) + 2 * arctan (3 / 79) = pi/4" proof - have 17: "\1 / 7\ < (1 :: real)" by auto with arctan_double have "2 * arctan (1 / 7) = arctan (7 / 24)" by simp (simp add: field_simps) moreover have "\7 / 24\ < (1 :: real)" by auto with arctan_double have "2 * arctan (7 / 24) = arctan (336 / 527)" by simp (simp add: field_simps) moreover have "\336 / 527\ < (1 :: real)" by auto from arctan_add[OF less_imp_le[OF 17] this] have "arctan(1/7) + arctan (336 / 527) = arctan (2879 / 3353)" by auto ultimately have I: "5 * arctan (1 / 7) = arctan (2879 / 3353)" by auto have 379: "\3 / 79\ < (1 :: real)" by auto with arctan_double have II: "2 * arctan (3 / 79) = arctan (237 / 3116)" by simp (simp add: field_simps) have *: "\2879 / 3353\ < (1 :: real)" by auto have "\237 / 3116\ < (1 :: real)" by auto from arctan_add[OF less_imp_le[OF *] this] have "arctan (2879/3353) + arctan (237/3116) = pi/4" by (simp add: arctan_one) with I II show ?thesis by auto qed (*But could also prove MACHIN_GAUSS: 12 * arctan(1/18) + 8 * arctan(1/57) - 5 * arctan(1/239) = pi/4*) subsection \Introducing the inverse tangent power series\ lemma monoseq_arctan_series: fixes x :: real assumes "\x\ \ 1" shows "monoseq (\n. 1 / real (n * 2 + 1) * x^(n * 2 + 1))" (is "monoseq ?a") proof (cases "x = 0") case True then show ?thesis by (auto simp: monoseq_def) next case False have "norm x \ 1" and "x \ 1" and "-1 \ x" using assms by auto show "monoseq ?a" proof - have mono: "1 / real (Suc (Suc n * 2)) * x ^ Suc (Suc n * 2) \ 1 / real (Suc (n * 2)) * x ^ Suc (n * 2)" if "0 \ x" and "x \ 1" for n and x :: real proof (rule mult_mono) show "1 / real (Suc (Suc n * 2)) \ 1 / real (Suc (n * 2))" by (rule frac_le) simp_all show "0 \ 1 / real (Suc (n * 2))" by auto show "x ^ Suc (Suc n * 2) \ x ^ Suc (n * 2)" by (rule power_decreasing) (simp_all add: \0 \ x\ \x \ 1\) show "0 \ x ^ Suc (Suc n * 2)" by (rule zero_le_power) (simp add: \0 \ x\) qed show ?thesis proof (cases "0 \ x") case True from mono[OF this \x \ 1\, THEN allI] show ?thesis unfolding Suc_eq_plus1[symmetric] by (rule mono_SucI2) next case False then have "0 \ - x" and "- x \ 1" using \-1 \ x\ by auto from mono[OF this] have "1 / real (Suc (Suc n * 2)) * x ^ Suc (Suc n * 2) \ 1 / real (Suc (n * 2)) * x ^ Suc (n * 2)" for n using \0 \ -x\ by auto then show ?thesis unfolding Suc_eq_plus1[symmetric] by (rule mono_SucI1[OF allI]) qed qed qed lemma zeroseq_arctan_series: fixes x :: real assumes "\x\ \ 1" shows "(\n. 1 / real (n * 2 + 1) * x^(n * 2 + 1)) \ 0" (is "?a \ 0") proof (cases "x = 0") case True then show ?thesis by simp next case False have "norm x \ 1" and "x \ 1" and "-1 \ x" using assms by auto show "?a \ 0" proof (cases "\x\ < 1") case True then have "norm x < 1" by auto from tendsto_mult[OF LIMSEQ_inverse_real_of_nat LIMSEQ_power_zero[OF \norm x < 1\, THEN LIMSEQ_Suc]] have "(\n. 1 / real (n + 1) * x ^ (n + 1)) \ 0" unfolding inverse_eq_divide Suc_eq_plus1 by simp then show ?thesis using pos2 by (rule LIMSEQ_linear) next case False then have "x = -1 \ x = 1" using \\x\ \ 1\ by auto then have n_eq: "\ n. x ^ (n * 2 + 1) = x" unfolding One_nat_def by auto from tendsto_mult[OF LIMSEQ_inverse_real_of_nat[THEN LIMSEQ_linear, OF pos2, unfolded inverse_eq_divide] tendsto_const[of x]] show ?thesis unfolding n_eq Suc_eq_plus1 by auto qed qed lemma summable_arctan_series: fixes n :: nat assumes "\x\ \ 1" shows "summable (\ k. (-1)^k * (1 / real (k*2+1) * x ^ (k*2+1)))" (is "summable (?c x)") by (rule summable_Leibniz(1), rule zeroseq_arctan_series[OF assms], rule monoseq_arctan_series[OF assms]) lemma DERIV_arctan_series: assumes "\x\ < 1" shows "DERIV (\x'. \k. (-1)^k * (1 / real (k * 2 + 1) * x' ^ (k * 2 + 1))) x :> (\k. (-1)^k * x^(k * 2))" (is "DERIV ?arctan _ :> ?Int") proof - let ?f = "\n. if even n then (-1)^(n div 2) * 1 / real (Suc n) else 0" have n_even: "even n \ 2 * (n div 2) = n" for n :: nat by presburger then have if_eq: "?f n * real (Suc n) * x'^n = (if even n then (-1)^(n div 2) * x'^(2 * (n div 2)) else 0)" for n x' by auto have summable_Integral: "summable (\ n. (- 1) ^ n * x^(2 * n))" if "\x\ < 1" for x :: real proof - from that have "x\<^sup>2 < 1" by (simp add: abs_square_less_1) have "summable (\ n. (- 1) ^ n * (x\<^sup>2) ^n)" by (rule summable_Leibniz(1)) (auto intro!: LIMSEQ_realpow_zero monoseq_realpow \x\<^sup>2 < 1\ order_less_imp_le[OF \x\<^sup>2 < 1\]) then show ?thesis by (simp only: power_mult) qed have sums_even: "(sums) f = (sums) (\ n. if even n then f (n div 2) else 0)" for f :: "nat \ real" proof - have "f sums x = (\ n. if even n then f (n div 2) else 0) sums x" for x :: real proof assume "f sums x" from sums_if[OF sums_zero this] show "(\n. if even n then f (n div 2) else 0) sums x" by auto next assume "(\ n. if even n then f (n div 2) else 0) sums x" from LIMSEQ_linear[OF this[simplified sums_def] pos2, simplified sum_split_even_odd[simplified mult.commute]] show "f sums x" unfolding sums_def by auto qed then show ?thesis .. qed have Int_eq: "(\n. ?f n * real (Suc n) * x^n) = ?Int" unfolding if_eq mult.commute[of _ 2] suminf_def sums_even[of "\ n. (- 1) ^ n * x ^ (2 * n)", symmetric] by auto have arctan_eq: "(\n. ?f n * x^(Suc n)) = ?arctan x" for x proof - have if_eq': "\n. (if even n then (- 1) ^ (n div 2) * 1 / real (Suc n) else 0) * x ^ Suc n = (if even n then (- 1) ^ (n div 2) * (1 / real (Suc (2 * (n div 2))) * x ^ Suc (2 * (n div 2))) else 0)" using n_even by auto have idx_eq: "\n. n * 2 + 1 = Suc (2 * n)" by auto then show ?thesis unfolding if_eq' idx_eq suminf_def sums_even[of "\ n. (- 1) ^ n * (1 / real (Suc (2 * n)) * x ^ Suc (2 * n))", symmetric] by auto qed have "DERIV (\ x. \ n. ?f n * x^(Suc n)) x :> (\n. ?f n * real (Suc n) * x^n)" proof (rule DERIV_power_series') show "x \ {- 1 <..< 1}" using \\ x \ < 1\ by auto show "summable (\ n. ?f n * real (Suc n) * x'^n)" if x'_bounds: "x' \ {- 1 <..< 1}" for x' :: real proof - from that have "\x'\ < 1" by auto then show ?thesis using that sums_summable sums_if [OF sums_0 [of "\x. 0"] summable_sums [OF summable_Integral]] by (auto simp add: if_distrib [of "\x. x * y" for y] cong: if_cong) qed qed auto then show ?thesis by (simp only: Int_eq arctan_eq) qed lemma arctan_series: assumes "\x\ \ 1" shows "arctan x = (\k. (-1)^k * (1 / real (k * 2 + 1) * x ^ (k * 2 + 1)))" (is "_ = suminf (\ n. ?c x n)") proof - let ?c' = "\x n. (-1)^n * x^(n*2)" have DERIV_arctan_suminf: "DERIV (\ x. suminf (?c x)) x :> (suminf (?c' x))" if "0 < r" and "r < 1" and "\x\ < r" for r x :: real proof (rule DERIV_arctan_series) from that show "\x\ < 1" using \r < 1\ and \\x\ < r\ by auto qed { fix x :: real assume "\x\ \ 1" note summable_Leibniz[OF zeroseq_arctan_series[OF this] monoseq_arctan_series[OF this]] } note arctan_series_borders = this have when_less_one: "arctan x = (\k. ?c x k)" if "\x\ < 1" for x :: real proof - obtain r where "\x\ < r" and "r < 1" using dense[OF \\x\ < 1\] by blast then have "0 < r" and "- r < x" and "x < r" by auto have suminf_eq_arctan_bounded: "suminf (?c x) - arctan x = suminf (?c a) - arctan a" if "-r < a" and "b < r" and "a < b" and "a \ x" and "x \ b" for x a b proof - from that have "\x\ < r" by auto show "suminf (?c x) - arctan x = suminf (?c a) - arctan a" proof (rule DERIV_isconst2[of "a" "b"]) show "a < b" and "a \ x" and "x \ b" using \a < b\ \a \ x\ \x \ b\ by auto have "\x. - r < x \ x < r \ DERIV (\ x. suminf (?c x) - arctan x) x :> 0" proof (rule allI, rule impI) fix x assume "-r < x \ x < r" then have "\x\ < r" by auto with \r < 1\ have "\x\ < 1" by auto have "\- (x\<^sup>2)\ < 1" using abs_square_less_1 \\x\ < 1\ by auto then have "(\n. (- (x\<^sup>2)) ^ n) sums (1 / (1 - (- (x\<^sup>2))))" unfolding real_norm_def[symmetric] by (rule geometric_sums) then have "(?c' x) sums (1 / (1 - (- (x\<^sup>2))))" unfolding power_mult_distrib[symmetric] power_mult mult.commute[of _ 2] by auto then have suminf_c'_eq_geom: "inverse (1 + x\<^sup>2) = suminf (?c' x)" using sums_unique unfolding inverse_eq_divide by auto have "DERIV (\ x. suminf (?c x)) x :> (inverse (1 + x\<^sup>2))" unfolding suminf_c'_eq_geom by (rule DERIV_arctan_suminf[OF \0 < r\ \r < 1\ \\x\ < r\]) from DERIV_diff [OF this DERIV_arctan] show "DERIV (\x. suminf (?c x) - arctan x) x :> 0" by auto qed then have DERIV_in_rball: "\y. a \ y \ y \ b \ DERIV (\x. suminf (?c x) - arctan x) y :> 0" using \-r < a\ \b < r\ by auto then show "\y. \a < y; y < b\ \ DERIV (\x. suminf (?c x) - arctan x) y :> 0" using \\x\ < r\ by auto show "continuous_on {a..b} (\x. suminf (?c x) - arctan x)" using DERIV_in_rball DERIV_atLeastAtMost_imp_continuous_on by blast qed qed have suminf_arctan_zero: "suminf (?c 0) - arctan 0 = 0" unfolding Suc_eq_plus1[symmetric] power_Suc2 mult_zero_right arctan_zero_zero suminf_zero by auto have "suminf (?c x) - arctan x = 0" proof (cases "x = 0") case True then show ?thesis using suminf_arctan_zero by auto next case False then have "0 < \x\" and "- \x\ < \x\" by auto have "suminf (?c (- \x\)) - arctan (- \x\) = suminf (?c 0) - arctan 0" by (rule suminf_eq_arctan_bounded[where x1=0 and a1="-\x\" and b1="\x\", symmetric]) (simp_all only: \\x\ < r\ \-\x\ < \x\\ neg_less_iff_less) moreover have "suminf (?c x) - arctan x = suminf (?c (- \x\)) - arctan (- \x\)" by (rule suminf_eq_arctan_bounded[where x1=x and a1="- \x\" and b1="\x\"]) (simp_all only: \\x\ < r\ \- \x\ < \x\\ neg_less_iff_less) ultimately show ?thesis using suminf_arctan_zero by auto qed then show ?thesis by auto qed show "arctan x = suminf (\n. ?c x n)" proof (cases "\x\ < 1") case True then show ?thesis by (rule when_less_one) next case False then have "\x\ = 1" using \\x\ \ 1\ by auto let ?a = "\x n. \1 / real (n * 2 + 1) * x^(n * 2 + 1)\" let ?diff = "\x n. \arctan x - (\i" have "?diff 1 n \ ?a 1 n" for n :: nat proof - have "0 < (1 :: real)" by auto moreover have "?diff x n \ ?a x n" if "0 < x" and "x < 1" for x :: real proof - from that have "\x\ \ 1" and "\x\ < 1" by auto from \0 < x\ have "0 < 1 / real (0 * 2 + (1::nat)) * x ^ (0 * 2 + 1)" by auto note bounds = mp[OF arctan_series_borders(2)[OF \\x\ \ 1\] this, unfolded when_less_one[OF \\x\ < 1\, symmetric], THEN spec] have "0 < 1 / real (n*2+1) * x^(n*2+1)" by (rule mult_pos_pos) (simp_all only: zero_less_power[OF \0 < x\], auto) then have a_pos: "?a x n = 1 / real (n*2+1) * x^(n*2+1)" by (rule abs_of_pos) show ?thesis proof (cases "even n") case True then have sgn_pos: "(-1)^n = (1::real)" by auto from \even n\ obtain m where "n = 2 * m" .. then have "2 * m = n" .. from bounds[of m, unfolded this atLeastAtMost_iff] have "\arctan x - (\i \ (\ii = ?c x n" by auto also have "\ = ?a x n" unfolding sgn_pos a_pos by auto finally show ?thesis . next case False then have sgn_neg: "(-1)^n = (-1::real)" by auto from \odd n\ obtain m where "n = 2 * m + 1" .. then have m_def: "2 * m + 1 = n" .. then have m_plus: "2 * (m + 1) = n + 1" by auto from bounds[of "m + 1", unfolded this atLeastAtMost_iff, THEN conjunct1] bounds[of m, unfolded m_def atLeastAtMost_iff, THEN conjunct2] have "\arctan x - (\i \ (\ii = - ?c x n" by auto also have "\ = ?a x n" unfolding sgn_neg a_pos by auto finally show ?thesis . qed qed hence "\x \ { 0 <..< 1 }. 0 \ ?a x n - ?diff x n" by auto moreover have "isCont (\ x. ?a x n - ?diff x n) x" for x unfolding diff_conv_add_uminus divide_inverse by (auto intro!: isCont_add isCont_rabs continuous_ident isCont_minus isCont_arctan continuous_at_within_inverse isCont_mult isCont_power continuous_const isCont_sum simp del: add_uminus_conv_diff) ultimately have "0 \ ?a 1 n - ?diff 1 n" by (rule LIM_less_bound) then show ?thesis by auto qed have "?a 1 \ 0" unfolding tendsto_rabs_zero_iff power_one divide_inverse One_nat_def by (auto intro!: tendsto_mult LIMSEQ_linear LIMSEQ_inverse_real_of_nat simp del: of_nat_Suc) have "?diff 1 \ 0" proof (rule LIMSEQ_I) fix r :: real assume "0 < r" obtain N :: nat where N_I: "N \ n \ ?a 1 n < r" for n using LIMSEQ_D[OF \?a 1 \ 0\ \0 < r\] by auto have "norm (?diff 1 n - 0) < r" if "N \ n" for n using \?diff 1 n \ ?a 1 n\ N_I[OF that] by auto then show "\N. \ n \ N. norm (?diff 1 n - 0) < r" by blast qed from this [unfolded tendsto_rabs_zero_iff, THEN tendsto_add [OF _ tendsto_const], of "- arctan 1", THEN tendsto_minus] have "(?c 1) sums (arctan 1)" unfolding sums_def by auto then have "arctan 1 = (\i. ?c 1 i)" by (rule sums_unique) show ?thesis proof (cases "x = 1") case True then show ?thesis by (simp add: \arctan 1 = (\ i. ?c 1 i)\) next case False then have "x = -1" using \\x\ = 1\ by auto have "- (pi/2) < 0" using pi_gt_zero by auto have "- (2 * pi) < 0" using pi_gt_zero by auto have c_minus_minus: "?c (- 1) i = - ?c 1 i" for i by auto - have "arctan (- 1) = arctan (tan (-(pi / 4)))" + have "arctan (- 1) = arctan (tan (-(pi/4)))" unfolding tan_45 tan_minus .. - also have "\ = - (pi / 4)" + also have "\ = - (pi/4)" by (rule arctan_tan) (auto simp: order_less_trans[OF \- (pi/2) < 0\ pi_gt_zero]) - also have "\ = - (arctan (tan (pi / 4)))" + also have "\ = - (arctan (tan (pi/4)))" unfolding neg_equal_iff_equal by (rule arctan_tan[symmetric]) (auto simp: order_less_trans[OF \- (2 * pi) < 0\ pi_gt_zero]) also have "\ = - (arctan 1)" unfolding tan_45 .. also have "\ = - (\ i. ?c 1 i)" using \arctan 1 = (\ i. ?c 1 i)\ by auto also have "\ = (\ i. ?c (- 1) i)" using suminf_minus[OF sums_summable[OF \(?c 1) sums (arctan 1)\]] unfolding c_minus_minus by auto finally show ?thesis using \x = -1\ by auto qed qed qed lemma arctan_half: "arctan x = 2 * arctan (x / (1 + sqrt(1 + x\<^sup>2)))" for x :: real proof - obtain y where low: "- (pi/2) < y" and high: "y < pi/2" and y_eq: "tan y = x" using tan_total by blast then have low2: "- (pi/2) < y / 2" and high2: "y / 2 < pi/2" by auto have "0 < cos y" by (rule cos_gt_zero_pi[OF low high]) then have "cos y \ 0" and cos_sqrt: "sqrt ((cos y)\<^sup>2) = cos y" by auto have "1 + (tan y)\<^sup>2 = 1 + (sin y)\<^sup>2 / (cos y)\<^sup>2" unfolding tan_def power_divide .. also have "\ = (cos y)\<^sup>2 / (cos y)\<^sup>2 + (sin y)\<^sup>2 / (cos y)\<^sup>2" using \cos y \ 0\ by auto also have "\ = 1 / (cos y)\<^sup>2" unfolding add_divide_distrib[symmetric] sin_cos_squared_add2 .. finally have "1 + (tan y)\<^sup>2 = 1 / (cos y)\<^sup>2" . have "sin y / (cos y + 1) = tan y / ((cos y + 1) / cos y)" unfolding tan_def using \cos y \ 0\ by (simp add: field_simps) also have "\ = tan y / (1 + 1 / cos y)" using \cos y \ 0\ unfolding add_divide_distrib by auto also have "\ = tan y / (1 + 1 / sqrt ((cos y)\<^sup>2))" unfolding cos_sqrt .. also have "\ = tan y / (1 + sqrt (1 / (cos y)\<^sup>2))" unfolding real_sqrt_divide by auto finally have eq: "sin y / (cos y + 1) = tan y / (1 + sqrt(1 + (tan y)\<^sup>2))" unfolding \1 + (tan y)\<^sup>2 = 1 / (cos y)\<^sup>2\ . have "arctan x = y" using arctan_tan low high y_eq by auto also have "\ = 2 * (arctan (tan (y/2)))" using arctan_tan[OF low2 high2] by auto also have "\ = 2 * (arctan (sin y / (cos y + 1)))" unfolding tan_half by auto finally show ?thesis unfolding eq \tan y = x\ . qed lemma arctan_monotone: "x < y \ arctan x < arctan y" by (simp only: arctan_less_iff) lemma arctan_monotone': "x \ y \ arctan x \ arctan y" by (simp only: arctan_le_iff) lemma arctan_inverse: assumes "x \ 0" shows "arctan (1 / x) = sgn x * pi/2 - arctan x" proof (rule arctan_unique) have \
: "x > 0 \ arctan x < pi" using arctan_bounded [of x] by linarith show "- (pi/2) < sgn x * pi/2 - arctan x" using assms by (auto simp: sgn_real_def arctan algebra_simps \
) show "sgn x * pi/2 - arctan x < pi/2" using arctan_bounded [of "- x"] assms by (auto simp: algebra_simps sgn_real_def arctan_minus) show "tan (sgn x * pi/2 - arctan x) = 1 / x" unfolding tan_inverse [of "arctan x", unfolded tan_arctan] sgn_real_def by (simp add: tan_def cos_arctan sin_arctan sin_diff cos_diff) qed -theorem pi_series: "pi / 4 = (\k. (-1)^k * 1 / real (k * 2 + 1))" +theorem pi_series: "pi/4 = (\k. (-1)^k * 1 / real (k * 2 + 1))" (is "_ = ?SUM") proof - - have "pi / 4 = arctan 1" + have "pi/4 = arctan 1" using arctan_one by auto also have "\ = ?SUM" using arctan_series[of 1] by auto finally show ?thesis by auto qed subsection \Existence of Polar Coordinates\ lemma cos_x_y_le_one: "\x / sqrt (x\<^sup>2 + y\<^sup>2)\ \ 1" by (rule power2_le_imp_le [OF _ zero_le_one]) (simp add: power_divide divide_le_eq not_sum_power2_lt_zero) lemma polar_Ex: "\r::real. \a. x = r * cos a \ y = r * sin a" proof - have polar_ex1: "\r a. x = r * cos a \ y = r * sin a" if "0 < y" for y proof - have "x = sqrt (x\<^sup>2 + y\<^sup>2) * cos (arccos (x / sqrt (x\<^sup>2 + y\<^sup>2)))" by (simp add: cos_arccos_abs [OF cos_x_y_le_one]) moreover have "y = sqrt (x\<^sup>2 + y\<^sup>2) * sin (arccos (x / sqrt (x\<^sup>2 + y\<^sup>2)))" using that by (simp add: sin_arccos_abs [OF cos_x_y_le_one] power_divide right_diff_distrib flip: real_sqrt_mult) ultimately show ?thesis by blast qed show ?thesis proof (cases "0::real" y rule: linorder_cases) case less then show ?thesis by (rule polar_ex1) next case equal then show ?thesis by (force simp: intro!: cos_zero sin_zero) next case greater with polar_ex1 [where y="-y"] show ?thesis by auto (metis cos_minus minus_minus minus_mult_right sin_minus) qed qed subsection \Basics about polynomial functions: products, extremal behaviour and root counts\ lemma pairs_le_eq_Sigma: "{(i, j). i + j \ m} = Sigma (atMost m) (\r. atMost (m - r))" for m :: nat by auto lemma sum_up_index_split: "(\k\m + n. f k) = (\k\m. f k) + (\k = Suc m..m + n. f k)" by (metis atLeast0AtMost Suc_eq_plus1 le0 sum.ub_add_nat) lemma Sigma_interval_disjoint: "(SIGMA i:A. {..v i}) \ (SIGMA i:A.{v i<..w}) = {}" for w :: "'a::order" by auto lemma product_atMost_eq_Un: "A \ {..m} = (SIGMA i:A.{..m - i}) \ (SIGMA i:A.{m - i<..m})" for m :: nat by auto lemma polynomial_product: (*with thanks to Chaitanya Mangla*) fixes x :: "'a::idom" assumes m: "\i. i > m \ a i = 0" and n: "\j. j > n \ b j = 0" shows "(\i\m. (a i) * x ^ i) * (\j\n. (b j) * x ^ j) = (\r\m + n. (\k\r. (a k) * (b (r - k))) * x ^ r)" proof - have "\i j. \m + n - i < j; a i \ 0\ \ b j = 0" by (meson le_add_diff leI le_less_trans m n) then have \
: "(\(i,j)\(SIGMA i:{..m+n}. {m+n - i<..m+n}). a i * x ^ i * (b j * x ^ j)) = 0" by (clarsimp simp add: sum_Un Sigma_interval_disjoint intro!: sum.neutral) have "(\i\m. (a i) * x ^ i) * (\j\n. (b j) * x ^ j) = (\i\m. \j\n. (a i * x ^ i) * (b j * x ^ j))" by (rule sum_product) also have "\ = (\i\m + n. \j\n + m. a i * x ^ i * (b j * x ^ j))" using assms by (auto simp: sum_up_index_split) also have "\ = (\r\m + n. \j\m + n - r. a r * x ^ r * (b j * x ^ j))" by (simp add: add_ac sum.Sigma product_atMost_eq_Un sum_Un Sigma_interval_disjoint \
) also have "\ = (\(i,j)\{(i,j). i+j \ m+n}. (a i * x ^ i) * (b j * x ^ j))" by (auto simp: pairs_le_eq_Sigma sum.Sigma) also have "... = (\k\m + n. \i\k. a i * x ^ i * (b (k - i) * x ^ (k - i)))" by (rule sum.triangle_reindex_eq) also have "\ = (\r\m + n. (\k\r. (a k) * (b (r - k))) * x ^ r)" by (auto simp: algebra_simps sum_distrib_left simp flip: power_add intro!: sum.cong) finally show ?thesis . qed lemma polynomial_product_nat: fixes x :: nat assumes m: "\i. i > m \ a i = 0" and n: "\j. j > n \ b j = 0" shows "(\i\m. (a i) * x ^ i) * (\j\n. (b j) * x ^ j) = (\r\m + n. (\k\r. (a k) * (b (r - k))) * x ^ r)" using polynomial_product [of m a n b x] assms by (simp only: of_nat_mult [symmetric] of_nat_power [symmetric] of_nat_eq_iff Int.int_sum [symmetric]) lemma polyfun_diff: (*COMPLEX_SUB_POLYFUN in HOL Light*) fixes x :: "'a::idom" assumes "1 \ n" shows "(\i\n. a i * x^i) - (\i\n. a i * y^i) = (x - y) * (\ji=Suc j..n. a i * y^(i - j - 1)) * x^j)" proof - have h: "bij_betw (\(i,j). (j,i)) ((SIGMA i : atMost n. lessThan i)) (SIGMA j : lessThan n. {Suc j..n})" by (auto simp: bij_betw_def inj_on_def) have "(\i\n. a i * x^i) - (\i\n. a i * y^i) = (\i\n. a i * (x^i - y^i))" by (simp add: right_diff_distrib sum_subtractf) also have "\ = (\i\n. a i * (x - y) * (\j = (\i\n. \j = (\(i,j) \ (SIGMA i : atMost n. lessThan i). a i * (x - y) * (y^(i - Suc j) * x^j))" by (simp add: sum.Sigma) also have "\ = (\(j,i) \ (SIGMA j : lessThan n. {Suc j..n}). a i * (x - y) * (y^(i - Suc j) * x^j))" by (auto simp: sum.reindex_bij_betw [OF h, symmetric] intro: sum.cong_simp) also have "\ = (\ji=Suc j..n. a i * (x - y) * (y^(i - Suc j) * x^j))" by (simp add: sum.Sigma) also have "\ = (x - y) * (\ji=Suc j..n. a i * y^(i - j - 1)) * x^j)" by (simp add: sum_distrib_left mult_ac) finally show ?thesis . qed lemma polyfun_diff_alt: (*COMPLEX_SUB_POLYFUN_ALT in HOL Light*) fixes x :: "'a::idom" assumes "1 \ n" shows "(\i\n. a i * x^i) - (\i\n. a i * y^i) = (x - y) * ((\jki=Suc j..n. a i * y^(i - j - 1)) = (\kk. k < n - j \ k \ (\i. i - Suc j) ` {Suc j..n}" by (rule_tac x="k + Suc j" in image_eqI, auto) then have h: "bij_betw (\i. i - (j + 1)) {Suc j..n} (lessThan (n-j))" by (auto simp: bij_betw_def inj_on_def) then show ?thesis by (auto simp: sum.reindex_bij_betw [OF h, symmetric] intro: sum.cong_simp) qed then show ?thesis by (simp add: polyfun_diff [OF assms] sum_distrib_right) qed lemma polyfun_linear_factor: (*COMPLEX_POLYFUN_LINEAR_FACTOR in HOL Light*) fixes a :: "'a::idom" shows "\b. \z. (\i\n. c(i) * z^i) = (z - a) * (\ii\n. c(i) * a^i)" proof (cases "n = 0") case True then show ?thesis by simp next case False have "(\b. \z. (\i\n. c i * z^i) = (z - a) * (\ii\n. c i * a^i)) \ (\b. \z. (\i\n. c i * z^i) - (\i\n. c i * a^i) = (z - a) * (\i \ (\b. \z. (z - a) * (\ji = Suc j..n. c i * a^(i - Suc j)) * z^j) = (z - a) * (\i = True" by auto finally show ?thesis by simp qed lemma polyfun_linear_factor_root: (*COMPLEX_POLYFUN_LINEAR_FACTOR_ROOT in HOL Light*) fixes a :: "'a::idom" assumes "(\i\n. c(i) * a^i) = 0" obtains b where "\z. (\i\n. c i * z^i) = (z - a) * (\iw. \i\n. c i * w^i) a" for c :: "nat \ 'a::real_normed_div_algebra" by simp lemma zero_polynom_imp_zero_coeffs: fixes c :: "nat \ 'a::{ab_semigroup_mult,real_normed_div_algebra}" assumes "\w. (\i\n. c i * w^i) = 0" "k \ n" shows "c k = 0" using assms proof (induction n arbitrary: c k) case 0 then show ?case by simp next case (Suc n c k) have [simp]: "c 0 = 0" using Suc.prems(1) [of 0] by simp have "(\i\Suc n. c i * w^i) = w * (\i\n. c (Suc i) * w^i)" for w proof - have "(\i\Suc n. c i * w^i) = (\i\n. c (Suc i) * w ^ Suc i)" unfolding Set_Interval.sum.atMost_Suc_shift by simp also have "\ = w * (\i\n. c (Suc i) * w^i)" by (simp add: sum_distrib_left ac_simps) finally show ?thesis . qed then have w: "\w. w \ 0 \ (\i\n. c (Suc i) * w^i) = 0" using Suc by auto then have "(\h. \i\n. c (Suc i) * h^i) \0\ 0" by (simp cong: LIM_cong) \ \the case \w = 0\ by continuity\ then have "(\i\n. c (Suc i) * 0^i) = 0" using isCont_polynom [of 0 "\i. c (Suc i)" n] LIM_unique by (force simp: Limits.isCont_iff) then have "\w. (\i\n. c (Suc i) * w^i) = 0" using w by metis then have "\i. i \ n \ c (Suc i) = 0" using Suc.IH [of "\i. c (Suc i)"] by blast then show ?case using \k \ Suc n\ by (cases k) auto qed lemma polyfun_rootbound: (*COMPLEX_POLYFUN_ROOTBOUND in HOL Light*) fixes c :: "nat \ 'a::{idom,real_normed_div_algebra}" assumes "c k \ 0" "k\n" shows "finite {z. (\i\n. c(i) * z^i) = 0} \ card {z. (\i\n. c(i) * z^i) = 0} \ n" using assms proof (induction n arbitrary: c k) case 0 then show ?case by simp next case (Suc m c k) let ?succase = ?case show ?case proof (cases "{z. (\i\Suc m. c(i) * z^i) = 0} = {}") case True then show ?succase by simp next case False then obtain z0 where z0: "(\i\Suc m. c(i) * z0^i) = 0" by blast then obtain b where b: "\w. (\i\Suc m. c i * w^i) = (w - z0) * (\i\m. b i * w^i)" using polyfun_linear_factor_root [OF z0, unfolded lessThan_Suc_atMost] by blast then have eq: "{z. (\i\Suc m. c i * z^i) = 0} = insert z0 {z. (\i\m. b i * z^i) = 0}" by auto have "\ (\k\m. b k = 0)" proof assume [simp]: "\k\m. b k = 0" then have "\w. (\i\m. b i * w^i) = 0" by simp then have "\w. (\i\Suc m. c i * w^i) = 0" using b by simp then have "\k. k \ Suc m \ c k = 0" using zero_polynom_imp_zero_coeffs by blast then show False using Suc.prems by blast qed then obtain k' where bk': "b k' \ 0" "k' \ m" by blast show ?succase using Suc.IH [of b k'] bk' by (simp add: eq card_insert_if del: sum.atMost_Suc) qed qed lemma fixes c :: "nat \ 'a::{idom,real_normed_div_algebra}" assumes "c k \ 0" "k\n" shows polyfun_roots_finite: "finite {z. (\i\n. c(i) * z^i) = 0}" and polyfun_roots_card: "card {z. (\i\n. c(i) * z^i) = 0} \ n" using polyfun_rootbound assms by auto lemma polyfun_finite_roots: (*COMPLEX_POLYFUN_FINITE_ROOTS in HOL Light*) fixes c :: "nat \ 'a::{idom,real_normed_div_algebra}" shows "finite {x. (\i\n. c i * x^i) = 0} \ (\i\n. c i \ 0)" (is "?lhs = ?rhs") proof assume ?lhs moreover have "\ finite {x. (\i\n. c i * x^i) = 0}" if "\i\n. c i = 0" proof - from that have "\x. (\i\n. c i * x^i) = 0" by simp then show ?thesis using ex_new_if_finite [OF infinite_UNIV_char_0 [where 'a='a]] by auto qed ultimately show ?rhs by metis next assume ?rhs with polyfun_rootbound show ?lhs by blast qed lemma polyfun_eq_0: "(\x. (\i\n. c i * x^i) = 0) \ (\i\n. c i = 0)" for c :: "nat \ 'a::{idom,real_normed_div_algebra}" (*COMPLEX_POLYFUN_EQ_0 in HOL Light*) using zero_polynom_imp_zero_coeffs by auto lemma polyfun_eq_coeffs: "(\x. (\i\n. c i * x^i) = (\i\n. d i * x^i)) \ (\i\n. c i = d i)" for c :: "nat \ 'a::{idom,real_normed_div_algebra}" proof - have "(\x. (\i\n. c i * x^i) = (\i\n. d i * x^i)) \ (\x. (\i\n. (c i - d i) * x^i) = 0)" by (simp add: left_diff_distrib Groups_Big.sum_subtractf) also have "\ \ (\i\n. c i - d i = 0)" by (rule polyfun_eq_0) finally show ?thesis by simp qed lemma polyfun_eq_const: (*COMPLEX_POLYFUN_EQ_CONST in HOL Light*) fixes c :: "nat \ 'a::{idom,real_normed_div_algebra}" shows "(\x. (\i\n. c i * x^i) = k) \ c 0 = k \ (\i \ {1..n}. c i = 0)" (is "?lhs = ?rhs") proof - have *: "\x. (\i\n. (if i=0 then k else 0) * x^i) = k" by (induct n) auto show ?thesis proof assume ?lhs with * have "(\i\n. c i = (if i=0 then k else 0))" by (simp add: polyfun_eq_coeffs [symmetric]) then show ?rhs by simp next assume ?rhs then show ?lhs by (induct n) auto qed qed lemma root_polyfun: fixes z :: "'a::idom" assumes "1 \ n" shows "z^n = a \ (\i\n. (if i = 0 then -a else if i=n then 1 else 0) * z^i) = 0" using assms by (cases n) (simp_all add: sum.atLeast_Suc_atMost atLeast0AtMost [symmetric]) lemma assumes "SORT_CONSTRAINT('a::{idom,real_normed_div_algebra})" and "1 \ n" shows finite_roots_unity: "finite {z::'a. z^n = 1}" and card_roots_unity: "card {z::'a. z^n = 1} \ n" using polyfun_rootbound [of "\i. if i = 0 then -1 else if i=n then 1 else 0" n n] assms(2) by (auto simp: root_polyfun [OF assms(2)]) subsection \Hyperbolic functions\ definition sinh :: "'a :: {banach, real_normed_algebra_1} \ 'a" where "sinh x = (exp x - exp (-x)) /\<^sub>R 2" definition cosh :: "'a :: {banach, real_normed_algebra_1} \ 'a" where "cosh x = (exp x + exp (-x)) /\<^sub>R 2" definition tanh :: "'a :: {banach, real_normed_field} \ 'a" where "tanh x = sinh x / cosh x" definition arsinh :: "'a :: {banach, real_normed_algebra_1, ln} \ 'a" where "arsinh x = ln (x + (x^2 + 1) powr of_real (1/2))" definition arcosh :: "'a :: {banach, real_normed_algebra_1, ln} \ 'a" where "arcosh x = ln (x + (x^2 - 1) powr of_real (1/2))" definition artanh :: "'a :: {banach, real_normed_field, ln} \ 'a" where "artanh x = ln ((1 + x) / (1 - x)) / 2" lemma arsinh_0 [simp]: "arsinh 0 = 0" by (simp add: arsinh_def) lemma arcosh_1 [simp]: "arcosh 1 = 0" by (simp add: arcosh_def) lemma artanh_0 [simp]: "artanh 0 = 0" by (simp add: artanh_def) lemma tanh_altdef: "tanh x = (exp x - exp (-x)) / (exp x + exp (-x))" proof - have "tanh x = (2 *\<^sub>R sinh x) / (2 *\<^sub>R cosh x)" by (simp add: tanh_def scaleR_conv_of_real) also have "2 *\<^sub>R sinh x = exp x - exp (-x)" by (simp add: sinh_def) also have "2 *\<^sub>R cosh x = exp x + exp (-x)" by (simp add: cosh_def) finally show ?thesis . qed lemma tanh_real_altdef: "tanh (x::real) = (1 - exp (- 2 * x)) / (1 + exp (- 2 * x))" proof - have [simp]: "exp (2 * x) = exp x * exp x" "exp (x * 2) = exp x * exp x" by (subst exp_add [symmetric]; simp)+ have "tanh x = (2 * exp (-x) * sinh x) / (2 * exp (-x) * cosh x)" by (simp add: tanh_def) also have "2 * exp (-x) * sinh x = 1 - exp (-2*x)" by (simp add: exp_minus field_simps sinh_def) also have "2 * exp (-x) * cosh x = 1 + exp (-2*x)" by (simp add: exp_minus field_simps cosh_def) finally show ?thesis . qed lemma sinh_converges: "(\n. if even n then 0 else x ^ n /\<^sub>R fact n) sums sinh x" proof - have "(\n. (x ^ n /\<^sub>R fact n - (-x) ^ n /\<^sub>R fact n) /\<^sub>R 2) sums sinh x" unfolding sinh_def by (intro sums_scaleR_right sums_diff exp_converges) also have "(\n. (x ^ n /\<^sub>R fact n - (-x) ^ n /\<^sub>R fact n) /\<^sub>R 2) = (\n. if even n then 0 else x ^ n /\<^sub>R fact n)" by auto finally show ?thesis . qed lemma cosh_converges: "(\n. if even n then x ^ n /\<^sub>R fact n else 0) sums cosh x" proof - have "(\n. (x ^ n /\<^sub>R fact n + (-x) ^ n /\<^sub>R fact n) /\<^sub>R 2) sums cosh x" unfolding cosh_def by (intro sums_scaleR_right sums_add exp_converges) also have "(\n. (x ^ n /\<^sub>R fact n + (-x) ^ n /\<^sub>R fact n) /\<^sub>R 2) = (\n. if even n then x ^ n /\<^sub>R fact n else 0)" by auto finally show ?thesis . qed lemma sinh_0 [simp]: "sinh 0 = 0" by (simp add: sinh_def) lemma cosh_0 [simp]: "cosh 0 = 1" proof - have "cosh 0 = (1/2) *\<^sub>R (1 + 1)" by (simp add: cosh_def) also have "\ = 1" by (rule scaleR_half_double) finally show ?thesis . qed lemma tanh_0 [simp]: "tanh 0 = 0" by (simp add: tanh_def) lemma sinh_minus [simp]: "sinh (- x) = -sinh x" by (simp add: sinh_def algebra_simps) lemma cosh_minus [simp]: "cosh (- x) = cosh x" by (simp add: cosh_def algebra_simps) lemma tanh_minus [simp]: "tanh (-x) = -tanh x" by (simp add: tanh_def) lemma sinh_ln_real: "x > 0 \ sinh (ln x :: real) = (x - inverse x) / 2" by (simp add: sinh_def exp_minus) lemma cosh_ln_real: "x > 0 \ cosh (ln x :: real) = (x + inverse x) / 2" by (simp add: cosh_def exp_minus) lemma tanh_ln_real: "tanh (ln x :: real) = (x ^ 2 - 1) / (x ^ 2 + 1)" if "x > 0" proof - from that have "(x * 2 - inverse x * 2) * (x\<^sup>2 + 1) = (x\<^sup>2 - 1) * (2 * x + 2 * inverse x)" by (simp add: field_simps power2_eq_square) moreover have "x\<^sup>2 + 1 > 0" using that by (simp add: ac_simps add_pos_nonneg) moreover have "2 * x + 2 * inverse x > 0" using that by (simp add: add_pos_pos) ultimately have "(x * 2 - inverse x * 2) / (2 * x + 2 * inverse x) = (x\<^sup>2 - 1) / (x\<^sup>2 + 1)" by (simp add: frac_eq_eq) with that show ?thesis by (simp add: tanh_def sinh_ln_real cosh_ln_real) qed lemma has_field_derivative_scaleR_right [derivative_intros]: "(f has_field_derivative D) F \ ((\x. c *\<^sub>R f x) has_field_derivative (c *\<^sub>R D)) F" unfolding has_field_derivative_def using has_derivative_scaleR_right[of f "\x. D * x" F c] by (simp add: mult_scaleR_left [symmetric] del: mult_scaleR_left) lemma has_field_derivative_sinh [THEN DERIV_chain2, derivative_intros]: "(sinh has_field_derivative cosh x) (at (x :: 'a :: {banach, real_normed_field}))" unfolding sinh_def cosh_def by (auto intro!: derivative_eq_intros) lemma has_field_derivative_cosh [THEN DERIV_chain2, derivative_intros]: "(cosh has_field_derivative sinh x) (at (x :: 'a :: {banach, real_normed_field}))" unfolding sinh_def cosh_def by (auto intro!: derivative_eq_intros) lemma has_field_derivative_tanh [THEN DERIV_chain2, derivative_intros]: "cosh x \ 0 \ (tanh has_field_derivative 1 - tanh x ^ 2) (at (x :: 'a :: {banach, real_normed_field}))" unfolding tanh_def by (auto intro!: derivative_eq_intros simp: power2_eq_square field_split_simps) lemma has_derivative_sinh [derivative_intros]: fixes g :: "'a \ ('a :: {banach, real_normed_field})" assumes "(g has_derivative (\x. Db * x)) (at x within s)" shows "((\x. sinh (g x)) has_derivative (\y. (cosh (g x) * Db) * y)) (at x within s)" proof - have "((\x. - g x) has_derivative (\y. -(Db * y))) (at x within s)" using assms by (intro derivative_intros) also have "(\y. -(Db * y)) = (\x. (-Db) * x)" by (simp add: fun_eq_iff) finally have "((\x. sinh (g x)) has_derivative (\y. (exp (g x) * Db * y - exp (-g x) * (-Db) * y) /\<^sub>R 2)) (at x within s)" unfolding sinh_def by (intro derivative_intros assms) also have "(\y. (exp (g x) * Db * y - exp (-g x) * (-Db) * y) /\<^sub>R 2) = (\y. (cosh (g x) * Db) * y)" by (simp add: fun_eq_iff cosh_def algebra_simps) finally show ?thesis . qed lemma has_derivative_cosh [derivative_intros]: fixes g :: "'a \ ('a :: {banach, real_normed_field})" assumes "(g has_derivative (\y. Db * y)) (at x within s)" shows "((\x. cosh (g x)) has_derivative (\y. (sinh (g x) * Db) * y)) (at x within s)" proof - have "((\x. - g x) has_derivative (\y. -(Db * y))) (at x within s)" using assms by (intro derivative_intros) also have "(\y. -(Db * y)) = (\y. (-Db) * y)" by (simp add: fun_eq_iff) finally have "((\x. cosh (g x)) has_derivative (\y. (exp (g x) * Db * y + exp (-g x) * (-Db) * y) /\<^sub>R 2)) (at x within s)" unfolding cosh_def by (intro derivative_intros assms) also have "(\y. (exp (g x) * Db * y + exp (-g x) * (-Db) * y) /\<^sub>R 2) = (\y. (sinh (g x) * Db) * y)" by (simp add: fun_eq_iff sinh_def algebra_simps) finally show ?thesis . qed lemma sinh_plus_cosh: "sinh x + cosh x = exp x" proof - - have "sinh x + cosh x = (1 / 2) *\<^sub>R (exp x + exp x)" + have "sinh x + cosh x = (1/2) *\<^sub>R (exp x + exp x)" by (simp add: sinh_def cosh_def algebra_simps) also have "\ = exp x" by (rule scaleR_half_double) finally show ?thesis . qed lemma cosh_plus_sinh: "cosh x + sinh x = exp x" by (subst add.commute) (rule sinh_plus_cosh) lemma cosh_minus_sinh: "cosh x - sinh x = exp (-x)" proof - - have "cosh x - sinh x = (1 / 2) *\<^sub>R (exp (-x) + exp (-x))" + have "cosh x - sinh x = (1/2) *\<^sub>R (exp (-x) + exp (-x))" by (simp add: sinh_def cosh_def algebra_simps) also have "\ = exp (-x)" by (rule scaleR_half_double) finally show ?thesis . qed lemma sinh_minus_cosh: "sinh x - cosh x = -exp (-x)" using cosh_minus_sinh[of x] by (simp add: algebra_simps) context fixes x :: "'a :: {real_normed_field, banach}" begin lemma sinh_zero_iff: "sinh x = 0 \ exp x \ {1, -1}" by (auto simp: sinh_def field_simps exp_minus power2_eq_square square_eq_1_iff) lemma cosh_zero_iff: "cosh x = 0 \ exp x ^ 2 = -1" by (auto simp: cosh_def exp_minus field_simps power2_eq_square eq_neg_iff_add_eq_0) lemma cosh_square_eq: "cosh x ^ 2 = sinh x ^ 2 + 1" by (simp add: cosh_def sinh_def algebra_simps power2_eq_square exp_add [symmetric] scaleR_conv_of_real) lemma sinh_square_eq: "sinh x ^ 2 = cosh x ^ 2 - 1" by (simp add: cosh_square_eq) lemma hyperbolic_pythagoras: "cosh x ^ 2 - sinh x ^ 2 = 1" by (simp add: cosh_square_eq) lemma sinh_add: "sinh (x + y) = sinh x * cosh y + cosh x * sinh y" by (simp add: sinh_def cosh_def algebra_simps scaleR_conv_of_real exp_add [symmetric]) lemma sinh_diff: "sinh (x - y) = sinh x * cosh y - cosh x * sinh y" by (simp add: sinh_def cosh_def algebra_simps scaleR_conv_of_real exp_add [symmetric]) lemma cosh_add: "cosh (x + y) = cosh x * cosh y + sinh x * sinh y" by (simp add: sinh_def cosh_def algebra_simps scaleR_conv_of_real exp_add [symmetric]) lemma cosh_diff: "cosh (x - y) = cosh x * cosh y - sinh x * sinh y" by (simp add: sinh_def cosh_def algebra_simps scaleR_conv_of_real exp_add [symmetric]) lemma tanh_add: "tanh (x + y) = (tanh x + tanh y) / (1 + tanh x * tanh y)" if "cosh x \ 0" "cosh y \ 0" proof - have "(sinh x * cosh y + cosh x * sinh y) * (1 + sinh x * sinh y / (cosh x * cosh y)) = (cosh x * cosh y + sinh x * sinh y) * ((sinh x * cosh y + sinh y * cosh x) / (cosh y * cosh x))" using that by (simp add: field_split_simps) also have "(sinh x * cosh y + sinh y * cosh x) / (cosh y * cosh x) = sinh x / cosh x + sinh y / cosh y" using that by (simp add: field_split_simps) finally have "(sinh x * cosh y + cosh x * sinh y) * (1 + sinh x * sinh y / (cosh x * cosh y)) = (sinh x / cosh x + sinh y / cosh y) * (cosh x * cosh y + sinh x * sinh y)" by simp then show ?thesis using that by (auto simp add: tanh_def sinh_add cosh_add eq_divide_eq) (simp_all add: field_split_simps) qed lemma sinh_double: "sinh (2 * x) = 2 * sinh x * cosh x" using sinh_add[of x] by simp lemma cosh_double: "cosh (2 * x) = cosh x ^ 2 + sinh x ^ 2" using cosh_add[of x] by (simp add: power2_eq_square) end lemma sinh_field_def: "sinh z = (exp z - exp (-z)) / (2 :: 'a :: {banach, real_normed_field})" by (simp add: sinh_def scaleR_conv_of_real) lemma cosh_field_def: "cosh z = (exp z + exp (-z)) / (2 :: 'a :: {banach, real_normed_field})" by (simp add: cosh_def scaleR_conv_of_real) subsubsection \More specific properties of the real functions\ lemma sinh_real_zero_iff [simp]: "sinh (x::real) = 0 \ x = 0" proof - have "(-1 :: real) < 0" by simp also have "0 < exp x" by simp finally have "exp x \ -1" by (intro notI) simp thus ?thesis by (subst sinh_zero_iff) simp qed lemma plus_inverse_ge_2: fixes x :: real assumes "x > 0" shows "x + inverse x \ 2" proof - have "0 \ (x - 1) ^ 2" by simp also have "\ = x^2 - 2*x + 1" by (simp add: power2_eq_square algebra_simps) finally show ?thesis using assms by (simp add: field_simps power2_eq_square) qed lemma sinh_real_nonneg_iff [simp]: "sinh (x :: real) \ 0 \ x \ 0" by (simp add: sinh_def) lemma sinh_real_pos_iff [simp]: "sinh (x :: real) > 0 \ x > 0" by (simp add: sinh_def) lemma sinh_real_nonpos_iff [simp]: "sinh (x :: real) \ 0 \ x \ 0" by (simp add: sinh_def) lemma sinh_real_neg_iff [simp]: "sinh (x :: real) < 0 \ x < 0" by (simp add: sinh_def) lemma cosh_real_ge_1: "cosh (x :: real) \ 1" using plus_inverse_ge_2[of "exp x"] by (simp add: cosh_def exp_minus) lemma cosh_real_pos [simp]: "cosh (x :: real) > 0" using cosh_real_ge_1[of x] by simp lemma cosh_real_nonneg[simp]: "cosh (x :: real) \ 0" using cosh_real_ge_1[of x] by simp lemma cosh_real_nonzero [simp]: "cosh (x :: real) \ 0" using cosh_real_ge_1[of x] by simp lemma tanh_real_nonneg_iff [simp]: "tanh (x :: real) \ 0 \ x \ 0" by (simp add: tanh_def field_simps) lemma tanh_real_pos_iff [simp]: "tanh (x :: real) > 0 \ x > 0" by (simp add: tanh_def field_simps) lemma tanh_real_nonpos_iff [simp]: "tanh (x :: real) \ 0 \ x \ 0" by (simp add: tanh_def field_simps) lemma tanh_real_neg_iff [simp]: "tanh (x :: real) < 0 \ x < 0" by (simp add: tanh_def field_simps) lemma tanh_real_zero_iff [simp]: "tanh (x :: real) = 0 \ x = 0" by (simp add: tanh_def field_simps) lemma arsinh_real_def: "arsinh (x::real) = ln (x + sqrt (x^2 + 1))" by (simp add: arsinh_def powr_half_sqrt) lemma arcosh_real_def: "x \ 1 \ arcosh (x::real) = ln (x + sqrt (x^2 - 1))" by (simp add: arcosh_def powr_half_sqrt) lemma arsinh_real_aux: "0 < x + sqrt (x ^ 2 + 1 :: real)" proof (cases "x < 0") case True have "(-x) ^ 2 = x ^ 2" by simp also have "x ^ 2 < x ^ 2 + 1" by simp finally have "sqrt ((-x) ^ 2) < sqrt (x ^ 2 + 1)" by (rule real_sqrt_less_mono) thus ?thesis using True by simp qed (auto simp: add_nonneg_pos) lemma arsinh_minus_real [simp]: "arsinh (-x::real) = -arsinh x" proof - have "arsinh (-x) = ln (sqrt (x\<^sup>2 + 1) - x)" by (simp add: arsinh_real_def) also have "sqrt (x^2 + 1) - x = inverse (sqrt (x^2 + 1) + x)" using arsinh_real_aux[of x] by (simp add: field_split_simps algebra_simps power2_eq_square) also have "ln \ = -arsinh x" using arsinh_real_aux[of x] by (simp add: arsinh_real_def ln_inverse) finally show ?thesis . qed lemma artanh_minus_real [simp]: assumes "abs x < 1" shows "artanh (-x::real) = -artanh x" using assms by (simp add: artanh_def ln_div field_simps) lemma sinh_less_cosh_real: "sinh (x :: real) < cosh x" by (simp add: sinh_def cosh_def) lemma sinh_le_cosh_real: "sinh (x :: real) \ cosh x" by (simp add: sinh_def cosh_def) lemma tanh_real_lt_1: "tanh (x :: real) < 1" by (simp add: tanh_def sinh_less_cosh_real) lemma tanh_real_gt_neg1: "tanh (x :: real) > -1" proof - have "- cosh x < sinh x" by (simp add: sinh_def cosh_def field_split_simps) thus ?thesis by (simp add: tanh_def field_simps) qed lemma tanh_real_bounds: "tanh (x :: real) \ {-1<..<1}" using tanh_real_lt_1 tanh_real_gt_neg1 by simp context fixes x :: real begin lemma arsinh_sinh_real: "arsinh (sinh x) = x" by (simp add: arsinh_real_def powr_def sinh_square_eq sinh_plus_cosh) lemma arcosh_cosh_real: "x \ 0 \ arcosh (cosh x) = x" by (simp add: arcosh_real_def powr_def cosh_square_eq cosh_real_ge_1 cosh_plus_sinh) lemma artanh_tanh_real: "artanh (tanh x) = x" proof - have "artanh (tanh x) = ln (cosh x * (cosh x + sinh x) / (cosh x * (cosh x - sinh x))) / 2" by (simp add: artanh_def tanh_def field_split_simps) also have "cosh x * (cosh x + sinh x) / (cosh x * (cosh x - sinh x)) = (cosh x + sinh x) / (cosh x - sinh x)" by simp also have "\ = (exp x)^2" by (simp add: cosh_plus_sinh cosh_minus_sinh exp_minus field_simps power2_eq_square) also have "ln ((exp x)^2) / 2 = x" by (simp add: ln_realpow) finally show ?thesis . qed end lemma sinh_real_strict_mono: "strict_mono (sinh :: real \ real)" by (rule pos_deriv_imp_strict_mono derivative_intros)+ auto lemma cosh_real_strict_mono: assumes "0 \ x" and "x < (y::real)" shows "cosh x < cosh y" proof - from assms have "\z>x. z < y \ cosh y - cosh x = (y - x) * sinh z" by (intro MVT2) (auto dest: connectedD_interval intro!: derivative_eq_intros) then obtain z where z: "z > x" "z < y" "cosh y - cosh x = (y - x) * sinh z" by blast note \cosh y - cosh x = (y - x) * sinh z\ also from \z > x\ and assms have "(y - x) * sinh z > 0" by (intro mult_pos_pos) auto finally show "cosh x < cosh y" by simp qed lemma tanh_real_strict_mono: "strict_mono (tanh :: real \ real)" proof - have *: "tanh x ^ 2 < 1" for x :: real using tanh_real_bounds[of x] by (simp add: abs_square_less_1 abs_if) show ?thesis by (rule pos_deriv_imp_strict_mono) (insert *, auto intro!: derivative_intros) qed lemma sinh_real_abs [simp]: "sinh (abs x :: real) = abs (sinh x)" by (simp add: abs_if) lemma cosh_real_abs [simp]: "cosh (abs x :: real) = cosh x" by (simp add: abs_if) lemma tanh_real_abs [simp]: "tanh (abs x :: real) = abs (tanh x)" by (auto simp: abs_if) lemma sinh_real_eq_iff [simp]: "sinh x = sinh y \ x = (y :: real)" using sinh_real_strict_mono by (simp add: strict_mono_eq) lemma tanh_real_eq_iff [simp]: "tanh x = tanh y \ x = (y :: real)" using tanh_real_strict_mono by (simp add: strict_mono_eq) lemma cosh_real_eq_iff [simp]: "cosh x = cosh y \ abs x = abs (y :: real)" proof - have "cosh x = cosh y \ x = y" if "x \ 0" "y \ 0" for x y :: real using cosh_real_strict_mono[of x y] cosh_real_strict_mono[of y x] that by (cases x y rule: linorder_cases) auto from this[of "abs x" "abs y"] show ?thesis by simp qed lemma sinh_real_le_iff [simp]: "sinh x \ sinh y \ x \ (y::real)" using sinh_real_strict_mono by (simp add: strict_mono_less_eq) lemma cosh_real_nonneg_le_iff: "x \ 0 \ y \ 0 \ cosh x \ cosh y \ x \ (y::real)" using cosh_real_strict_mono[of x y] cosh_real_strict_mono[of y x] by (cases x y rule: linorder_cases) auto lemma cosh_real_nonpos_le_iff: "x \ 0 \ y \ 0 \ cosh x \ cosh y \ x \ (y::real)" using cosh_real_nonneg_le_iff[of "-x" "-y"] by simp lemma tanh_real_le_iff [simp]: "tanh x \ tanh y \ x \ (y::real)" using tanh_real_strict_mono by (simp add: strict_mono_less_eq) lemma sinh_real_less_iff [simp]: "sinh x < sinh y \ x < (y::real)" using sinh_real_strict_mono by (simp add: strict_mono_less) lemma cosh_real_nonneg_less_iff: "x \ 0 \ y \ 0 \ cosh x < cosh y \ x < (y::real)" using cosh_real_strict_mono[of x y] cosh_real_strict_mono[of y x] by (cases x y rule: linorder_cases) auto lemma cosh_real_nonpos_less_iff: "x \ 0 \ y \ 0 \ cosh x < cosh y \ x > (y::real)" using cosh_real_nonneg_less_iff[of "-x" "-y"] by simp lemma tanh_real_less_iff [simp]: "tanh x < tanh y \ x < (y::real)" using tanh_real_strict_mono by (simp add: strict_mono_less) subsubsection \Limits\ lemma sinh_real_at_top: "filterlim (sinh :: real \ real) at_top at_top" proof - have *: "((\x. - exp (- x)) \ (-0::real)) at_top" by (intro tendsto_minus filterlim_compose[OF exp_at_bot] filterlim_uminus_at_bot_at_top) - have "filterlim (\x. (1 / 2) * (-exp (-x) + exp x) :: real) at_top at_top" + have "filterlim (\x. (1/2) * (-exp (-x) + exp x) :: real) at_top at_top" by (rule filterlim_tendsto_pos_mult_at_top[OF _ _ filterlim_tendsto_add_at_top[OF *]] tendsto_const)+ (auto simp: exp_at_top) - also have "(\x. (1 / 2) * (-exp (-x) + exp x) :: real) = sinh" + also have "(\x. (1/2) * (-exp (-x) + exp x) :: real) = sinh" by (simp add: fun_eq_iff sinh_def) finally show ?thesis . qed lemma sinh_real_at_bot: "filterlim (sinh :: real \ real) at_bot at_bot" proof - have "filterlim (\x. -sinh x :: real) at_bot at_top" by (simp add: filterlim_uminus_at_top [symmetric] sinh_real_at_top) also have "(\x. -sinh x :: real) = (\x. sinh (-x))" by simp finally show ?thesis by (subst filterlim_at_bot_mirror) qed lemma cosh_real_at_top: "filterlim (cosh :: real \ real) at_top at_top" proof - have *: "((\x. exp (- x)) \ (0::real)) at_top" by (intro filterlim_compose[OF exp_at_bot] filterlim_uminus_at_bot_at_top) - have "filterlim (\x. (1 / 2) * (exp (-x) + exp x) :: real) at_top at_top" + have "filterlim (\x. (1/2) * (exp (-x) + exp x) :: real) at_top at_top" by (rule filterlim_tendsto_pos_mult_at_top[OF _ _ filterlim_tendsto_add_at_top[OF *]] tendsto_const)+ (auto simp: exp_at_top) - also have "(\x. (1 / 2) * (exp (-x) + exp x) :: real) = cosh" + also have "(\x. (1/2) * (exp (-x) + exp x) :: real) = cosh" by (simp add: fun_eq_iff cosh_def) finally show ?thesis . qed lemma cosh_real_at_bot: "filterlim (cosh :: real \ real) at_top at_bot" proof - have "filterlim (\x. cosh (-x) :: real) at_top at_top" by (simp add: cosh_real_at_top) thus ?thesis by (subst filterlim_at_bot_mirror) qed lemma tanh_real_at_top: "(tanh \ (1::real)) at_top" proof - have "((\x::real. (1 - exp (- 2 * x)) / (1 + exp (- 2 * x))) \ (1 - 0) / (1 + 0)) at_top" by (intro tendsto_intros filterlim_compose[OF exp_at_bot] filterlim_tendsto_neg_mult_at_bot[OF tendsto_const] filterlim_ident) auto also have "(\x::real. (1 - exp (- 2 * x)) / (1 + exp (- 2 * x))) = tanh" by (rule ext) (simp add: tanh_real_altdef) finally show ?thesis by simp qed lemma tanh_real_at_bot: "(tanh \ (-1::real)) at_bot" proof - have "((\x::real. -tanh x) \ -1) at_top" by (intro tendsto_minus tanh_real_at_top) also have "(\x. -tanh x :: real) = (\x. tanh (-x))" by simp finally show ?thesis by (subst filterlim_at_bot_mirror) qed subsubsection \Properties of the inverse hyperbolic functions\ lemma isCont_sinh: "isCont sinh (x :: 'a :: {real_normed_field, banach})" unfolding sinh_def [abs_def] by (auto intro!: continuous_intros) lemma isCont_cosh: "isCont cosh (x :: 'a :: {real_normed_field, banach})" unfolding cosh_def [abs_def] by (auto intro!: continuous_intros) lemma isCont_tanh: "cosh x \ 0 \ isCont tanh (x :: 'a :: {real_normed_field, banach})" unfolding tanh_def [abs_def] by (auto intro!: continuous_intros isCont_divide isCont_sinh isCont_cosh) lemma continuous_on_sinh [continuous_intros]: fixes f :: "_ \'a::{real_normed_field,banach}" assumes "continuous_on A f" shows "continuous_on A (\x. sinh (f x))" unfolding sinh_def using assms by (intro continuous_intros) lemma continuous_on_cosh [continuous_intros]: fixes f :: "_ \'a::{real_normed_field,banach}" assumes "continuous_on A f" shows "continuous_on A (\x. cosh (f x))" unfolding cosh_def using assms by (intro continuous_intros) lemma continuous_sinh [continuous_intros]: fixes f :: "_ \'a::{real_normed_field,banach}" assumes "continuous F f" shows "continuous F (\x. sinh (f x))" unfolding sinh_def using assms by (intro continuous_intros) lemma continuous_cosh [continuous_intros]: fixes f :: "_ \'a::{real_normed_field,banach}" assumes "continuous F f" shows "continuous F (\x. cosh (f x))" unfolding cosh_def using assms by (intro continuous_intros) lemma continuous_on_tanh [continuous_intros]: fixes f :: "_ \'a::{real_normed_field,banach}" assumes "continuous_on A f" "\x. x \ A \ cosh (f x) \ 0" shows "continuous_on A (\x. tanh (f x))" unfolding tanh_def using assms by (intro continuous_intros) auto lemma continuous_at_within_tanh [continuous_intros]: fixes f :: "_ \'a::{real_normed_field,banach}" assumes "continuous (at x within A) f" "cosh (f x) \ 0" shows "continuous (at x within A) (\x. tanh (f x))" unfolding tanh_def using assms by (intro continuous_intros continuous_divide) auto lemma continuous_tanh [continuous_intros]: fixes f :: "_ \'a::{real_normed_field,banach}" assumes "continuous F f" "cosh (f (Lim F (\x. x))) \ 0" shows "continuous F (\x. tanh (f x))" unfolding tanh_def using assms by (intro continuous_intros continuous_divide) auto lemma tendsto_sinh [tendsto_intros]: fixes f :: "_ \'a::{real_normed_field,banach}" shows "(f \ a) F \ ((\x. sinh (f x)) \ sinh a) F" by (rule isCont_tendsto_compose [OF isCont_sinh]) lemma tendsto_cosh [tendsto_intros]: fixes f :: "_ \'a::{real_normed_field,banach}" shows "(f \ a) F \ ((\x. cosh (f x)) \ cosh a) F" by (rule isCont_tendsto_compose [OF isCont_cosh]) lemma tendsto_tanh [tendsto_intros]: fixes f :: "_ \'a::{real_normed_field,banach}" shows "(f \ a) F \ cosh a \ 0 \ ((\x. tanh (f x)) \ tanh a) F" by (rule isCont_tendsto_compose [OF isCont_tanh]) lemma arsinh_real_has_field_derivative [derivative_intros]: fixes x :: real shows "(arsinh has_field_derivative (1 / (sqrt (x ^ 2 + 1)))) (at x within A)" proof - have pos: "1 + x ^ 2 > 0" by (intro add_pos_nonneg) auto from pos arsinh_real_aux[of x] show ?thesis unfolding arsinh_def [abs_def] by (auto intro!: derivative_eq_intros simp: powr_minus powr_half_sqrt field_split_simps) qed lemma arcosh_real_has_field_derivative [derivative_intros]: fixes x :: real assumes "x > 1" shows "(arcosh has_field_derivative (1 / (sqrt (x ^ 2 - 1)))) (at x within A)" proof - from assms have "x + sqrt (x\<^sup>2 - 1) > 0" by (simp add: add_pos_pos) thus ?thesis using assms unfolding arcosh_def [abs_def] by (auto intro!: derivative_eq_intros simp: powr_minus powr_half_sqrt field_split_simps power2_eq_1_iff) qed lemma artanh_real_has_field_derivative [derivative_intros]: "(artanh has_field_derivative (1 / (1 - x ^ 2))) (at x within A)" if "\x\ < 1" for x :: real proof - from that have "- 1 < x" "x < 1" by linarith+ hence "(artanh has_field_derivative (4 - 4 * x) / ((1 + x) * (1 - x) * (1 - x) * 4)) (at x within A)" unfolding artanh_def [abs_def] by (auto intro!: derivative_eq_intros simp: powr_minus powr_half_sqrt) also have "(4 - 4 * x) / ((1 + x) * (1 - x) * (1 - x) * 4) = 1 / ((1 + x) * (1 - x))" using \-1 < x\ \x < 1\ by (simp add: frac_eq_eq) also have "(1 + x) * (1 - x) = 1 - x ^ 2" by (simp add: algebra_simps power2_eq_square) finally show ?thesis . qed lemma continuous_on_arsinh [continuous_intros]: "continuous_on A (arsinh :: real \ real)" by (rule DERIV_continuous_on derivative_intros)+ lemma continuous_on_arcosh [continuous_intros]: assumes "A \ {1..}" shows "continuous_on A (arcosh :: real \ real)" proof - have pos: "x + sqrt (x ^ 2 - 1) > 0" if "x \ 1" for x using that by (intro add_pos_nonneg) auto show ?thesis unfolding arcosh_def [abs_def] by (intro continuous_on_subset [OF _ assms] continuous_on_ln continuous_on_add continuous_on_id continuous_on_powr') (auto dest: pos simp: powr_half_sqrt intro!: continuous_intros) qed lemma continuous_on_artanh [continuous_intros]: assumes "A \ {-1<..<1}" shows "continuous_on A (artanh :: real \ real)" unfolding artanh_def [abs_def] by (intro continuous_on_subset [OF _ assms]) (auto intro!: continuous_intros) lemma continuous_on_arsinh' [continuous_intros]: fixes f :: "real \ real" assumes "continuous_on A f" shows "continuous_on A (\x. arsinh (f x))" by (rule continuous_on_compose2[OF continuous_on_arsinh assms]) auto lemma continuous_on_arcosh' [continuous_intros]: fixes f :: "real \ real" assumes "continuous_on A f" "\x. x \ A \ f x \ 1" shows "continuous_on A (\x. arcosh (f x))" by (rule continuous_on_compose2[OF continuous_on_arcosh assms(1) order.refl]) (use assms(2) in auto) lemma continuous_on_artanh' [continuous_intros]: fixes f :: "real \ real" assumes "continuous_on A f" "\x. x \ A \ f x \ {-1<..<1}" shows "continuous_on A (\x. artanh (f x))" by (rule continuous_on_compose2[OF continuous_on_artanh assms(1) order.refl]) (use assms(2) in auto) lemma isCont_arsinh [continuous_intros]: "isCont arsinh (x :: real)" using continuous_on_arsinh[of UNIV] by (auto simp: continuous_on_eq_continuous_at) lemma isCont_arcosh [continuous_intros]: assumes "x > 1" shows "isCont arcosh (x :: real)" proof - have "continuous_on {1::real<..} arcosh" by (rule continuous_on_arcosh) auto with assms show ?thesis by (auto simp: continuous_on_eq_continuous_at) qed lemma isCont_artanh [continuous_intros]: assumes "x > -1" "x < 1" shows "isCont artanh (x :: real)" proof - have "continuous_on {-1<..<(1::real)} artanh" by (rule continuous_on_artanh) auto with assms show ?thesis by (auto simp: continuous_on_eq_continuous_at) qed lemma tendsto_arsinh [tendsto_intros]: "(f \ a) F \ ((\x. arsinh (f x)) \ arsinh a) F" for f :: "_ \ real" by (rule isCont_tendsto_compose [OF isCont_arsinh]) lemma tendsto_arcosh_strong [tendsto_intros]: fixes f :: "_ \ real" assumes "(f \ a) F" "a \ 1" "eventually (\x. f x \ 1) F" shows "((\x. arcosh (f x)) \ arcosh a) F" by (rule continuous_on_tendsto_compose[OF continuous_on_arcosh[OF order.refl]]) (use assms in auto) lemma tendsto_arcosh: fixes f :: "_ \ real" assumes "(f \ a) F" "a > 1" shows "((\x. arcosh (f x)) \ arcosh a) F" by (rule isCont_tendsto_compose [OF isCont_arcosh]) (use assms in auto) lemma tendsto_arcosh_at_left_1: "(arcosh \ 0) (at_right (1::real))" proof - have "(arcosh \ arcosh 1) (at_right (1::real))" by (rule tendsto_arcosh_strong) (auto simp: eventually_at intro!: exI[of _ 1]) thus ?thesis by simp qed lemma tendsto_artanh [tendsto_intros]: fixes f :: "'a \ real" assumes "(f \ a) F" "a > -1" "a < 1" shows "((\x. artanh (f x)) \ artanh a) F" by (rule isCont_tendsto_compose [OF isCont_artanh]) (use assms in auto) lemma continuous_arsinh [continuous_intros]: "continuous F f \ continuous F (\x. arsinh (f x :: real))" unfolding continuous_def by (rule tendsto_arsinh) (* TODO: This rule does not work for one-sided continuity at 1 *) lemma continuous_arcosh_strong [continuous_intros]: assumes "continuous F f" "eventually (\x. f x \ 1) F" shows "continuous F (\x. arcosh (f x :: real))" proof (cases "F = bot") case False show ?thesis unfolding continuous_def proof (intro tendsto_arcosh_strong) show "1 \ f (Lim F (\x. x))" using assms False unfolding continuous_def by (rule tendsto_lowerbound) qed (insert assms, auto simp: continuous_def) qed auto lemma continuous_arcosh: "continuous F f \ f (Lim F (\x. x)) > 1 \ continuous F (\x. arcosh (f x :: real))" unfolding continuous_def by (rule tendsto_arcosh) auto lemma continuous_artanh [continuous_intros]: "continuous F f \ f (Lim F (\x. x)) \ {-1<..<1} \ continuous F (\x. artanh (f x :: real))" unfolding continuous_def by (rule tendsto_artanh) auto lemma arsinh_real_at_top: "filterlim (arsinh :: real \ real) at_top at_top" proof (subst filterlim_cong[OF refl refl]) show "filterlim (\x. ln (x + sqrt (1 + x\<^sup>2))) at_top at_top" by (intro filterlim_compose[OF ln_at_top filterlim_at_top_add_at_top] filterlim_ident filterlim_compose[OF sqrt_at_top] filterlim_tendsto_add_at_top[OF tendsto_const] filterlim_pow_at_top) auto qed (auto intro!: eventually_mono[OF eventually_ge_at_top[of 1]] simp: arsinh_real_def add_ac) lemma arsinh_real_at_bot: "filterlim (arsinh :: real \ real) at_bot at_bot" proof - have "filterlim (\x::real. -arsinh x) at_bot at_top" by (subst filterlim_uminus_at_top [symmetric]) (rule arsinh_real_at_top) also have "(\x::real. -arsinh x) = (\x. arsinh (-x))" by simp finally show ?thesis by (subst filterlim_at_bot_mirror) qed lemma arcosh_real_at_top: "filterlim (arcosh :: real \ real) at_top at_top" proof (subst filterlim_cong[OF refl refl]) show "filterlim (\x. ln (x + sqrt (-1 + x\<^sup>2))) at_top at_top" by (intro filterlim_compose[OF ln_at_top filterlim_at_top_add_at_top] filterlim_ident filterlim_compose[OF sqrt_at_top] filterlim_tendsto_add_at_top[OF tendsto_const] filterlim_pow_at_top) auto qed (auto intro!: eventually_mono[OF eventually_ge_at_top[of 1]] simp: arcosh_real_def) lemma artanh_real_at_left_1: "filterlim (artanh :: real \ real) at_top (at_left 1)" proof - have *: "filterlim (\x::real. (1 + x) / (1 - x)) at_top (at_left 1)" by (rule LIM_at_top_divide) (auto intro!: tendsto_eq_intros eventually_mono[OF eventually_at_left_real[of 0]]) have "filterlim (\x::real. (1/2) * ln ((1 + x) / (1 - x))) at_top (at_left 1)" by (intro filterlim_tendsto_pos_mult_at_top[OF tendsto_const] * filterlim_compose[OF ln_at_top]) auto also have "(\x::real. (1/2) * ln ((1 + x) / (1 - x))) = artanh" by (simp add: artanh_def [abs_def]) finally show ?thesis . qed lemma artanh_real_at_right_1: "filterlim (artanh :: real \ real) at_bot (at_right (-1))" proof - have "?thesis \ filterlim (\x::real. -artanh x) at_top (at_right (-1))" by (simp add: filterlim_uminus_at_bot) also have "\ \ filterlim (\x::real. artanh (-x)) at_top (at_right (-1))" by (intro filterlim_cong refl eventually_mono[OF eventually_at_right_real[of "-1" "1"]]) auto also have "\ \ filterlim (artanh :: real \ real) at_top (at_left 1)" by (simp add: filterlim_at_left_to_right) also have \ by (rule artanh_real_at_left_1) finally show ?thesis . qed subsection \Simprocs for root and power literals\ lemma numeral_powr_numeral_real [simp]: "numeral m powr numeral n = (numeral m ^ numeral n :: real)" by (simp add: powr_numeral) context begin private lemma sqrt_numeral_simproc_aux: assumes "m * m \ n" shows "sqrt (numeral n :: real) \ numeral m" proof - have "numeral n \ numeral m * (numeral m :: real)" by (simp add: assms [symmetric]) moreover have "sqrt \ \ numeral m" by (subst real_sqrt_abs2) simp ultimately show "sqrt (numeral n :: real) \ numeral m" by simp qed private lemma root_numeral_simproc_aux: assumes "Num.pow m n \ x" shows "root (numeral n) (numeral x :: real) \ numeral m" by (subst assms [symmetric], subst numeral_pow, subst real_root_pos2) simp_all private lemma powr_numeral_simproc_aux: assumes "Num.pow y n = x" shows "numeral x powr (m / numeral n :: real) \ numeral y powr m" by (subst assms [symmetric], subst numeral_pow, subst powr_numeral [symmetric]) (simp, subst powr_powr, simp_all) private lemma numeral_powr_inverse_eq: "numeral x powr (inverse (numeral n)) = numeral x powr (1 / numeral n :: real)" by simp ML \ signature ROOT_NUMERAL_SIMPROC = sig val sqrt : int option -> int -> int option val sqrt' : int option -> int -> int option val nth_root : int option -> int -> int -> int option val nth_root' : int option -> int -> int -> int option val sqrt_simproc : Proof.context -> cterm -> thm option val root_simproc : int * int -> Proof.context -> cterm -> thm option val powr_simproc : int * int -> Proof.context -> cterm -> thm option end structure Root_Numeral_Simproc : ROOT_NUMERAL_SIMPROC = struct fun iterate NONE p f x = let fun go x = if p x then x else go (f x) in SOME (go x) end | iterate (SOME threshold) p f x = let fun go (threshold, x) = if p x then SOME x else if threshold = 0 then NONE else go (threshold - 1, f x) in go (threshold, x) end fun nth_root _ 1 x = SOME x | nth_root _ _ 0 = SOME 0 | nth_root _ _ 1 = SOME 1 | nth_root threshold n x = let fun newton_step y = ((n - 1) * y + x div Integer.pow (n - 1) y) div n fun is_root y = Integer.pow n y <= x andalso x < Integer.pow n (y + 1) in if x < n then SOME 1 else if x < Integer.pow n 2 then SOME 1 else let val y = Real.floor (Math.pow (Real.fromInt x, Real.fromInt 1 / Real.fromInt n)) in if is_root y then SOME y else iterate threshold is_root newton_step ((x + n - 1) div n) end end fun nth_root' _ 1 x = SOME x | nth_root' _ _ 0 = SOME 0 | nth_root' _ _ 1 = SOME 1 | nth_root' threshold n x = if x < n then NONE else if x < Integer.pow n 2 then NONE else case nth_root threshold n x of NONE => NONE | SOME y => if Integer.pow n y = x then SOME y else NONE fun sqrt _ 0 = SOME 0 | sqrt _ 1 = SOME 1 | sqrt threshold n = let fun aux (a, b) = if n >= b * b then aux (b, b * b) else (a, b) val (lower_root, lower_n) = aux (1, 2) fun newton_step x = (x + n div x) div 2 fun is_sqrt r = r*r <= n andalso n < (r+1)*(r+1) val y = Real.floor (Math.sqrt (Real.fromInt n)) in if is_sqrt y then SOME y else Option.mapPartial (iterate threshold is_sqrt newton_step o (fn x => x * lower_root)) (sqrt threshold (n div lower_n)) end fun sqrt' threshold x = case sqrt threshold x of NONE => NONE | SOME y => if y * y = x then SOME y else NONE fun sqrt_simproc ctxt ct = let val n = ct |> Thm.term_of |> dest_comb |> snd |> dest_comb |> snd |> HOLogic.dest_numeral in case sqrt' (SOME 10000) n of NONE => NONE | SOME m => SOME (Thm.instantiate' [] (map (SOME o Thm.cterm_of ctxt o HOLogic.mk_numeral) [m, n]) @{thm sqrt_numeral_simproc_aux}) end handle TERM _ => NONE fun root_simproc (threshold1, threshold2) ctxt ct = let val [n, x] = ct |> Thm.term_of |> strip_comb |> snd |> map (dest_comb #> snd #> HOLogic.dest_numeral) in if n > threshold1 orelse x > threshold2 then NONE else case nth_root' (SOME 100) n x of NONE => NONE | SOME m => SOME (Thm.instantiate' [] (map (SOME o Thm.cterm_of ctxt o HOLogic.mk_numeral) [m, n, x]) @{thm root_numeral_simproc_aux}) end handle TERM _ => NONE | Match => NONE fun powr_simproc (threshold1, threshold2) ctxt ct = let val eq_thm = Conv.try_conv (Conv.rewr_conv @{thm numeral_powr_inverse_eq}) ct val ct = Thm.dest_equals_rhs (Thm.cprop_of eq_thm) val (_, [x, t]) = strip_comb (Thm.term_of ct) val (_, [m, n]) = strip_comb t val [x, n] = map (dest_comb #> snd #> HOLogic.dest_numeral) [x, n] in if n > threshold1 orelse x > threshold2 then NONE else case nth_root' (SOME 100) n x of NONE => NONE | SOME y => let val [y, n, x] = map HOLogic.mk_numeral [y, n, x] val thm = Thm.instantiate' [] (map (SOME o Thm.cterm_of ctxt) [y, n, x, m]) @{thm powr_numeral_simproc_aux} in SOME (@{thm transitive} OF [eq_thm, thm]) end end handle TERM _ => NONE | Match => NONE end \ end simproc_setup sqrt_numeral ("sqrt (numeral n)") = \K Root_Numeral_Simproc.sqrt_simproc\ simproc_setup root_numeral ("root (numeral n) (numeral x)") = \K (Root_Numeral_Simproc.root_simproc (200, Integer.pow 200 2))\ simproc_setup powr_divide_numeral ("numeral x powr (m / numeral n :: real)" | "numeral x powr (inverse (numeral n) :: real)") = \K (Root_Numeral_Simproc.powr_simproc (200, Integer.pow 200 2))\ lemma "root 100 1267650600228229401496703205376 = 2" by simp lemma "sqrt 196 = 14" by simp lemma "256 powr (7 / 4 :: real) = 16384" by simp lemma "27 powr (inverse 3) = (3::real)" by simp end