# HG changeset patch # User huffman # Date 1313435626 25200 # Node ID bd7c586b902e905cdbc58133af302d2d34742a19 # Parent eba74571833bb648126f5985486c04241da588b7 remove duplicate lemmas eventually_conjI, eventually_and, eventually_false diff -r eba74571833b -r bd7c586b902e src/HOL/Multivariate_Analysis/Cartesian_Euclidean_Space.thy --- a/src/HOL/Multivariate_Analysis/Cartesian_Euclidean_Space.thy Mon Aug 15 10:49:48 2011 -0700 +++ b/src/HOL/Multivariate_Analysis/Cartesian_Euclidean_Space.thy Mon Aug 15 12:13:46 2011 -0700 @@ -1370,7 +1370,7 @@ lemma Lim_component_eq_cart: fixes f :: "'a \ real^'n" assumes net:"(f ---> l) net" "~(trivial_limit net)" and ev:"eventually (\x. f(x)$i = b) net" shows "l$i = b" - using ev[unfolded order_eq_iff eventually_and] using Lim_component_ge_cart[OF net, of b i] and + using ev[unfolded order_eq_iff eventually_conj_iff] using Lim_component_ge_cart[OF net, of b i] and Lim_component_le_cart[OF net, of i b] by auto lemma connected_ivt_component_cart: fixes x::"real^'n" shows diff -r eba74571833b -r bd7c586b902e src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy --- a/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Mon Aug 15 10:49:48 2011 -0700 +++ b/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Mon Aug 15 12:13:46 2011 -0700 @@ -976,21 +976,10 @@ text{* Combining theorems for "eventually" *} -lemma eventually_conjI: - "\eventually (\x. P x) net; eventually (\x. Q x) net\ - \ eventually (\x. P x \ Q x) net" -by (rule eventually_conj) (* FIXME: delete *) - lemma eventually_rev_mono: "eventually P net \ (\x. P x \ Q x) \ eventually Q net" using eventually_mono [of P Q] by fast -lemma eventually_and: " eventually (\x. P x \ Q x) net \ eventually P net \ eventually Q net" - by (rule eventually_conj_iff) (* FIXME: delete *) - -lemma eventually_false: "eventually (\x. False) net \ trivial_limit net" - by (rule eventually_False) (* FIXME: delete *) - lemma not_eventually: "(\x. \ P x ) \ ~(trivial_limit net) ==> ~(eventually (\x. P x) net)" by (simp add: eventually_False) @@ -1255,7 +1244,7 @@ hence "dist (f x) 0 < e" by (simp add: dist_norm) } thus "eventually (\x. dist (f x) 0 < e) net" - using eventually_and[of "\x. norm(f x) <= g x" "\x. dist (g x) 0 < e" net] + using eventually_conj_iff[of "\x. norm(f x) <= g x" "\x. dist (g x) 0 < e" net] using eventually_mono[of "(\x. norm (f x) \ g x \ dist (g x) 0 < e)" "(\x. dist (f x) 0 < e)" net] using assms `e>0` unfolding tendsto_iff by auto qed @@ -1271,7 +1260,7 @@ assume "norm (f x) \ norm (g x)" "dist (g x) 0 < e" hence "dist (f x) 0 < e" by (simp add: dist_norm)} thus "eventually (\x. dist (f x) 0 < e) net" - using eventually_and[of "\x. norm (f x) \ norm (g x)" "\x. dist (g x) 0 < e" net] + using eventually_conj_iff[of "\x. norm (f x) \ norm (g x)" "\x. dist (g x) 0 < e" net] using eventually_mono[of "\x. norm (f x) \ norm (g x) \ dist (g x) 0 < e" "\x. dist (f x) 0 < e" net] using assms `e>0` unfolding tendsto_iff by blast qed @@ -1304,7 +1293,7 @@ with assms(2) have "eventually (\x. dist (f x) l < dist a l - e) net" by (rule tendstoD) with assms(3) have "eventually (\x. dist a (f x) \ e \ dist (f x) l < dist a l - e) net" - by (rule eventually_conjI) + by (rule eventually_conj) then obtain w where "dist a (f w) \ e" "dist (f w) l < dist a l - e" using assms(1) eventually_happens by auto hence "dist a (f w) + dist (f w) l < e + (dist a l - e)" @@ -1326,7 +1315,7 @@ with assms(2) have "eventually (\x. dist (f x) l < norm l - e) net" by (rule tendstoD) with assms(3) have "eventually (\x. norm (f x) \ e \ dist (f x) l < norm l - e) net" - by (rule eventually_conjI) + by (rule eventually_conj) then obtain w where "norm (f w) \ e" "dist (f w) l < norm l - e" using assms(1) eventually_happens by auto hence "norm (f w - l) < norm l - e" "norm (f w) \ e" by (simp_all add: dist_norm) @@ -1345,7 +1334,7 @@ with assms(2) have "eventually (\x. dist (f x) l < e - norm l) net" by (rule tendstoD) with assms(3) have "eventually (\x. e \ norm (f x) \ dist (f x) l < e - norm l) net" - by (rule eventually_conjI) + by (rule eventually_conj) then obtain w where "e \ norm (f w)" "dist (f w) l < e - norm l" using assms(1) eventually_happens by auto hence "norm (f w - l) + norm l < e" "e \ norm (f w)" by (simp_all add: dist_norm) @@ -4236,7 +4225,7 @@ { fix x and e::real assume "x\s" "e>0" have "eventually (\n. \x\s. norm (f n x - g x) < e / 3) net" using `e>0` assms(3)[THEN spec[where x="e/3"]] by auto then obtain n where n:"\xa\s. norm (f n xa - g xa) < e / 3" "continuous_on s (f n)" - using eventually_and[of "(\n. \x\s. norm (f n x - g x) < e / 3)" "(\n. continuous_on s (f n))" net] assms(1,2) eventually_happens by blast + using eventually_conj_iff[of "(\n. \x\s. norm (f n x - g x) < e / 3)" "(\n. continuous_on s (f n))" net] assms(1,2) eventually_happens by blast have "e / 3 > 0" using `e>0` by auto then obtain d where "d>0" and d:"\x'\s. dist x' x < d \ dist (f n x') (f n x) < e / 3" using n(2)[unfolded continuous_on_iff, THEN bspec[where x=x], OF `x\s`, THEN spec[where x="e/3"]] by blast @@ -5305,7 +5294,7 @@ lemma Lim_component_eq: fixes f :: "'a \ 'b::euclidean_space" assumes net:"(f ---> l) net" "~(trivial_limit net)" and ev:"eventually (\x. f(x)$$i = b) net" shows "l$$i = b" - using ev[unfolded order_eq_iff eventually_and] using Lim_component_ge[OF net, of b i] and Lim_component_le[OF net, of i b] by auto + using ev[unfolded order_eq_iff eventually_conj_iff] using Lim_component_ge[OF net, of b i] and Lim_component_le[OF net, of i b] by auto text{* Limits relative to a union. *} lemma eventually_within_Un: