# HG changeset patch # User hoelzl # Date 1307613016 -7200 # Node ID a150d16bf77c0a7b00a9e36989c22b1c22017526 # Parent 57a1c19f8e3bde67dde50dcd10f83d52b1568ed2 lemmas about right derivative and limits diff -r 57a1c19f8e3b -r a150d16bf77c src/HOL/Multivariate_Analysis/Derivative.thy --- a/src/HOL/Multivariate_Analysis/Derivative.thy Thu Jun 09 11:50:16 2011 +0200 +++ b/src/HOL/Multivariate_Analysis/Derivative.thy Thu Jun 09 11:50:16 2011 +0200 @@ -105,6 +105,22 @@ "a \ s \ open s \ ((f has_derivative f') (at a within s) \ (f has_derivative f') (at a))" by (simp only: at_within_interior interior_open) +lemma has_derivative_right: + fixes f :: "real \ real" and y :: "real" + shows "(f has_derivative (op * y)) (at x within ({x <..} \ I)) \ + ((\t. (f x - f t) / (x - t)) ---> y) (at x within ({x <..} \ I))" +proof - + have "((\t. (f t - (f x + y * (t - x))) / \t - x\) ---> 0) (at x within ({x<..} \ I)) \ + ((\t. (f t - f x) / (t - x) - y) ---> 0) (at x within ({x<..} \ I))" + by (intro Lim_cong_within) (auto simp add: divide.diff divide.add) + also have "\ \ ((\t. (f t - f x) / (t - x)) ---> y) (at x within ({x<..} \ I))" + by (simp add: Lim_null[symmetric]) + also have "\ \ ((\t. (f x - f t) / (x - t)) ---> y) (at x within ({x<..} \ I))" + by (intro Lim_cong_within) (simp_all add: times_divide_eq field_simps) + finally show ?thesis + by (simp add: mult.bounded_linear_right has_derivative_within) +qed + lemma bounded_linear_imp_linear: "bounded_linear f \ linear f" (* TODO: move elsewhere *) proof - assume "bounded_linear f" diff -r 57a1c19f8e3b -r a150d16bf77c src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy --- a/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Thu Jun 09 11:50:16 2011 +0200 +++ b/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Thu Jun 09 11:50:16 2011 +0200 @@ -1122,6 +1122,100 @@ thus ?lhs by (rule Lim_at_within) qed +lemma Lim_within_LIMSEQ: + fixes a :: real and L :: "'a::metric_space" + assumes "\S. (\n. S n \ a \ S n \ T) \ S ----> a \ (\n. X (S n)) ----> L" + shows "(X ---> L) (at a within T)" +proof (rule ccontr) + assume "\ (X ---> L) (at a within T)" + hence "\r>0. \s>0. \x\T. x \ a \ \x - a\ < s \ r \ dist (X x) L" + unfolding tendsto_iff eventually_within dist_norm by (simp add: not_less[symmetric]) + then obtain r where r: "r > 0" "\s. s > 0 \ \x\T-{a}. \x - a\ < s \ dist (X x) L \ r" by blast + + let ?F = "\n::nat. SOME x. x \ T \ x \ a \ \x - a\ < inverse (real (Suc n)) \ dist (X x) L \ r" + have "\n. \x. x \ T \ x \ a \ \x - a\ < inverse (real (Suc n)) \ dist (X x) L \ r" + using r by (simp add: Bex_def) + hence F: "\n. ?F n \ T \ ?F n \ a \ \?F n - a\ < inverse (real (Suc n)) \ dist (X (?F n)) L \ r" + by (rule someI_ex) + hence F1: "\n. ?F n \ T \ ?F n \ a" + and F2: "\n. \?F n - a\ < inverse (real (Suc n))" + and F3: "\n. dist (X (?F n)) L \ r" + by fast+ + + have "?F ----> a" + proof (rule LIMSEQ_I, unfold real_norm_def) + fix e::real + assume "0 < e" + (* choose no such that inverse (real (Suc n)) < e *) + then have "\no. inverse (real (Suc no)) < e" by (rule reals_Archimedean) + then obtain m where nodef: "inverse (real (Suc m)) < e" by auto + show "\no. \n. no \ n --> \?F n - a\ < e" + proof (intro exI allI impI) + fix n + assume mlen: "m \ n" + have "\?F n - a\ < inverse (real (Suc n))" + by (rule F2) + also have "inverse (real (Suc n)) \ inverse (real (Suc m))" + using mlen by auto + also from nodef have + "inverse (real (Suc m)) < e" . + finally show "\?F n - a\ < e" . + qed + qed + moreover note `\S. (\n. S n \ a \ S n \ T) \ S ----> a \ (\n. X (S n)) ----> L` + ultimately have "(\n. X (?F n)) ----> L" using F1 by simp + + moreover have "\ ((\n. X (?F n)) ----> L)" + proof - + { + fix no::nat + obtain n where "n = no + 1" by simp + then have nolen: "no \ n" by simp + (* We prove this by showing that for any m there is an n\m such that |X (?F n) - L| \ r *) + have "dist (X (?F n)) L \ r" + by (rule F3) + with nolen have "\n. no \ n \ dist (X (?F n)) L \ r" by fast + } + then have "(\no. \n. no \ n \ dist (X (?F n)) L \ r)" by simp + with r have "\e>0. (\no. \n. no \ n \ dist (X (?F n)) L \ e)" by auto + thus ?thesis by (unfold LIMSEQ_def, auto simp add: not_less) + qed + ultimately show False by simp +qed + +lemma Lim_right_bound: + fixes f :: "real \ real" + assumes mono: "\a b. a \ I \ b \ I \ x < a \ a \ b \ f a \ f b" + assumes bnd: "\a. a \ I \ x < a \ K \ f a" + shows "(f ---> Inf (f ` ({x<..} \ I))) (at x within ({x<..} \ I))" +proof cases + assume "{x<..} \ I = {}" then show ?thesis by (simp add: Lim_within_empty) +next + assume [simp]: "{x<..} \ I \ {}" + show ?thesis + proof (rule Lim_within_LIMSEQ, safe) + fix S assume S: "\n. S n \ x \ S n \ {x <..} \ I" "S ----> x" + + show "(\n. f (S n)) ----> Inf (f ` ({x<..} \ I))" + proof (rule LIMSEQ_I, rule ccontr) + fix r :: real assume "0 < r" + with Inf_close[of "f ` ({x<..} \ I)" r] + obtain y where y: "x < y" "y \ I" "f y < Inf (f ` ({x <..} \ I)) + r" by auto + from `x < y` have "0 < y - x" by auto + from S(2)[THEN LIMSEQ_D, OF this] + obtain N where N: "\n. N \ n \ \S n - x\ < y - x" by auto + + assume "\ (\N. \n\N. norm (f (S n) - Inf (f ` ({x<..} \ I))) < r)" + moreover have "\n. Inf (f ` ({x<..} \ I)) \ f (S n)" + using S bnd by (intro Inf_lower[where z=K]) auto + ultimately obtain n where n: "N \ n" "r + Inf (f ` ({x<..} \ I)) \ f (S n)" + by (auto simp: not_less field_simps) + with N[OF n(1)] mono[OF _ `y \ I`, of "S n"] S(1)[THEN spec, of n] y + show False by auto + qed + qed +qed + text{* Another limit point characterization. *} lemma islimpt_sequential: @@ -1505,14 +1599,16 @@ (* FIXME: Only one congruence rule for tendsto can be used at a time! *) lemma Lim_cong_within(*[cong add]*): - assumes "\x. x \ a \ f x = g x" - shows "((\x. f x) ---> l) (at a within S) \ ((g ---> l) (at a within S))" + assumes "a = b" "x = y" "S = T" + assumes "\x. x \ b \ x \ T \ f x = g x" + shows "(f ---> x) (at a within S) \ (g ---> y) (at b within T)" unfolding tendsto_def Limits.eventually_within eventually_at_topological using assms by simp lemma Lim_cong_at(*[cong add]*): + assumes "a = b" "x = y" assumes "\x. x \ a \ f x = g x" - shows "((\x. f x) ---> l) (at a) \ ((g ---> l) (at a))" + shows "((\x. f x) ---> x) (at a) \ ((g ---> y) (at a))" unfolding tendsto_def eventually_at_topological using assms by simp