# HG changeset patch # User huffman # Date 1243993592 25200 # Node ID 8cbcab09ce2acb6a1577839a049f4fcd2b85eab9 # Parent 8d8417abb14f6f8bc0764cb597051d0dcc022b31 generalize Lim_transform lemmas diff -r 8d8417abb14f -r 8cbcab09ce2a src/HOL/Library/Topology_Euclidean_Space.thy --- a/src/HOL/Library/Topology_Euclidean_Space.thy Tue Jun 02 18:31:11 2009 -0700 +++ b/src/HOL/Library/Topology_Euclidean_Space.thy Tue Jun 02 18:46:32 2009 -0700 @@ -1607,25 +1607,29 @@ qed lemma Lim_transform_eventually: - fixes f g :: "'a::type \ 'b::real_normed_vector" - (* FIXME: generalize to metric_space *) + fixes f g :: "'a::type \ 'b::metric_space" shows "eventually (\x. f x = g x) net \ (f ---> l) net ==> (g ---> l) net" - using Lim_eventually[of "\x. f x - g x" 0 net] Lim_transform[of f g net l] by auto + unfolding tendsto_def + apply (clarify, drule spec, drule (1) mp) + apply (erule (1) eventually_elim2, simp) + done lemma Lim_transform_within: - fixes f g :: "'a::perfect_space \ 'b::real_normed_vector" - (* FIXME: generalize to metric_space *) + fixes f g :: "'a::metric_space \ 'b::metric_space" assumes "0 < d" "(\x'\S. 0 < dist x' x \ dist x' x < d \ f x' = g x')" "(f ---> l) (at x within S)" shows "(g ---> l) (at x within S)" -proof- - have "((\x. f x - g x) ---> 0) (at x within S)" unfolding Lim_within[of "\x. f x - g x" 0 x S] using assms(1,2) by auto - thus ?thesis using Lim_transform[of f g "at x within S" l] using assms(3) by auto -qed + using assms(1,3) unfolding Lim_within + apply - + apply (clarify, rename_tac e) + apply (drule_tac x=e in spec, clarsimp, rename_tac r) + apply (rule_tac x="min d r" in exI, clarsimp, rename_tac y) + apply (drule_tac x=y in bspec, assumption, clarsimp) + apply (simp add: assms(2)) + done lemma Lim_transform_at: - fixes f g :: "'a::perfect_space \ 'b::real_normed_vector" - (* FIXME: generalize to metric_space *) + fixes f g :: "'a::metric_space \ 'b::metric_space" shows "0 < d \ (\x'. 0 < dist x' x \ dist x' x < d \ f x' = g x') \ (f ---> l) (at x) ==> (g ---> l) (at x)" apply (subst within_UNIV[symmetric]) @@ -1635,8 +1639,7 @@ text{* Common case assuming being away from some crucial point like 0. *} lemma Lim_transform_away_within: - fixes f:: "'a::perfect_space \ 'b::real_normed_vector" - (* FIXME: generalize to metric_space *) + fixes f:: "'a::metric_space \ 'b::metric_space" assumes "a\b" "\x\ S. x \ a \ x \ b \ f x = g x" and "(f ---> l) (at a within S)" shows "(g ---> l) (at a within S)" @@ -1647,8 +1650,7 @@ qed lemma Lim_transform_away_at: - fixes f:: "'a::perfect_space \ 'b::real_normed_vector" - (* FIXME: generalize to metric_space *) + fixes f:: "'a::metric_space \ 'b::metric_space" assumes ab: "a\b" and fg: "\x. x \ a \ x \ b \ f x = g x" and fl: "(f ---> l) (at a)" shows "(g ---> l) (at a)" @@ -1658,7 +1660,7 @@ text{* Alternatively, within an open set. *} lemma Lim_transform_within_open: - fixes f:: "'a::perfect_space \ 'b::real_normed_vector" + fixes f:: "'a::metric_space \ 'b::metric_space" (* FIXME: generalize to metric_space *) assumes "open S" "a \ S" "\x\S. x \ a \ f x = g x" "(f ---> l) (at a)" shows "(g ---> l) (at a)" @@ -1671,6 +1673,8 @@ text{* A congruence rule allowing us to transform limits assuming not at point. *} +(* FIXME: Only one congruence rule for tendsto can be used at a time! *) + lemma Lim_cong_within[cong add]: "(\x. x \ a \ f x = g x) ==> ((\x. f x) ---> l) (at a within S) \ ((g ---> l) (at a within S))" by (simp add: Lim_within dist_nz[symmetric])