# HG changeset patch # User himmelma # Date 1243522994 -7200 # Node ID b98cbfabe8249c54dda7198d52726b3720664051 # Parent b4d4dbc5b04f6d09c6a6251ae6f0305bb9242631 Moved some lemmas about intervals to Topology diff -r b4d4dbc5b04f -r b98cbfabe824 src/HOL/Library/Topology_Euclidean_Space.thy --- a/src/HOL/Library/Topology_Euclidean_Space.thy Thu May 28 16:19:34 2009 +0200 +++ b/src/HOL/Library/Topology_Euclidean_Space.thy Thu May 28 17:03:14 2009 +0200 @@ -3791,11 +3791,11 @@ unfolding continuous_on_def apply (simp del: dist_sym) unfolding dist_vec1 unfolding dist_def .. lemma continuous_at_vec1_norm: - "\x. continuous (at x) (vec1 o norm)" + "continuous (at x) (vec1 o norm)" unfolding continuous_at_vec1_range using real_abs_sub_norm order_le_less_trans by blast lemma continuous_on_vec1_norm: - "\s. continuous_on s (vec1 o norm)" + "continuous_on s (vec1 o norm)" unfolding continuous_on_vec1_range norm_vec1[THEN sym] by (metis norm_vec1 order_le_less_trans real_abs_sub_norm) lemma continuous_at_vec1_component: @@ -3860,7 +3860,7 @@ lemma continuous_attains_inf: "compact s \ s \ {} \ continuous_on s (vec1 o f) - ==> (\x \ s. \y \ s. f x \ f y)" + \ (\x \ s. \y \ s. f x \ f y)" using compact_attains_inf[of "f ` s"] using compact_continuous_image[of s "vec1 \ f"] unfolding image_compose by auto @@ -4299,8 +4299,12 @@ lemma mem_interval: fixes a :: "'a::ord^'n::finite" shows "x \ {a<.. (\i. a$i < x$i \ x$i < b$i)" "x \ {a .. b} \ (\i. a$i \ x$i \ x$i \ b$i)" - using interval[of a b] - by(auto simp add: expand_set_eq vector_less_def vector_less_eq_def) + using interval[of a b] by(auto simp add: expand_set_eq vector_less_def vector_less_eq_def) + +lemma mem_interval_1: fixes x :: "real^1" shows + "(x \ {a .. b} \ dest_vec1 a \ dest_vec1 x \ dest_vec1 x \ dest_vec1 b)" + "(x \ {a<.. dest_vec1 a < dest_vec1 x \ dest_vec1 x < dest_vec1 b)" +by(simp_all add: Cart_eq vector_less_def vector_less_eq_def dest_vec1_def forall_1) lemma interval_eq_empty: fixes a :: "real^'n::finite" shows "({a <..< b} = {} \ (\i. b$i \ a$i))" (is ?th1) and @@ -4357,7 +4361,6 @@ apply (auto simp add: not_less less_imp_le) done - lemma interval_open_subset_closed: fixes a :: "'a::preorder^'n::finite" shows "{a<.. {a .. b}" proof(simp add: subset_eq, rule) @@ -5427,7 +5430,11 @@ ultimately show ?thesis using False by auto qed -subsection{* Banach fixed point theorem (not really topological...) *} +lemma image_smult_interval:"(\x. m *s (x::real^'n::finite)) ` {a..b} = + (if {a..b} = {} then {} else if 0 \ m then {m *s a..m *s b} else {m *s b..m *s a})" + using image_affinity_interval[of m 0 a b] by auto + +subsection{* Banach fixed point theorem (not really topological...) *} lemma banach_fix: assumes s:"complete s" "s \ {}" and c:"0 \ c" "c < 1" and f:"(f ` s) \ s" and