# HG changeset patch # User himmelma # Date 1243520374 -7200 # Node ID b4d4dbc5b04f6d09c6a6251ae6f0305bb9242631 # Parent 8ef7ba78bf260e5cf4d44cedb276412a456e24a8 Corrected definition of is_interval diff -r 8ef7ba78bf26 -r b4d4dbc5b04f src/HOL/Library/Convex_Euclidean_Space.thy --- a/src/HOL/Library/Convex_Euclidean_Space.thy Thu May 28 15:54:20 2009 +0200 +++ b/src/HOL/Library/Convex_Euclidean_Space.thy Thu May 28 16:19:34 2009 +0200 @@ -35,15 +35,6 @@ "a \ b \ dest_vec1 a \ dest_vec1 b" "dest_vec1 (1::real^1) = 1" by(auto simp add:vector_component_simps all_1 Cart_eq) -definition - "is_interval_new (s::((real^'n::finite) set)) \ - (\a\s. \b\s. \x. (\i. ((a$i \ x$i \ x$i \ b$i) \ (b$i \ x$i \ x$i \ a$i))) \ x \ s)" - -lemma is_interval_interval_new: "is_interval_new {a .. b}" (is ?th1) "is_interval_new {a<..x y z::real. x < y \ y < z \ x < z" by auto - show ?th1 ?th2 unfolding is_interval_new_def mem_interval Ball_def atLeastAtMost_iff - by(meson real_le_trans le_less_trans less_le_trans *)+ qed - lemma nequals0I:"x\A \ A \ {}" by auto lemma norm_not_0:"(x::real^'n::finite)\0 \ norm x \ 0" by auto @@ -2175,7 +2166,7 @@ subsection {* Convexity of general and special intervals. *} -lemma is_interval_convex: assumes "is_interval_new s" shows "convex s" +lemma is_interval_convex: assumes "is_interval s" shows "convex s" unfolding convex_def apply(rule,rule,rule,rule,rule,rule,rule) proof- fix x y u v assume as:"x \ s" "y \ s" "0 \ u" "0 \ v" "u + v = (1::real)" hence *:"u = 1 - v" "1 - v \ 0" and **:"v = 1 - u" "1 - u \ 0" by auto @@ -2188,22 +2179,22 @@ hence "v * b > (1 - u) * a" unfolding not_le using as(4) by(auto simp add: field_simps) hence "a < b" unfolding * using as(4) apply(rule_tac mult_left_less_imp_less) by(auto simp add: ring_simps) hence "u * a + v * b \ b" unfolding ** using **(2) as(3) by(auto simp add: field_simps intro!:mult_right_mono) } - ultimately show "u *s x + v *s y \ s" apply- apply(rule assms[unfolded is_interval_new_def, rule_format, OF as(1,2)]) + ultimately show "u *s x + v *s y \ s" apply- apply(rule assms[unfolded is_interval_def, rule_format, OF as(1,2)]) using as(3-) dimindex_ge_1 apply- by(auto simp add: vector_component) qed -lemma is_interval_connected: "is_interval_new s \ connected s" +lemma is_interval_connected: "is_interval s \ connected s" using is_interval_convex convex_connected by auto lemma convex_interval: "convex {a .. b}" "convex {a<.. (\a\s. \b\s. \ x. dest_vec1 a \ dest_vec1 x \ dest_vec1 x \ dest_vec1 b \ x \ s)" - unfolding is_interval_new_def dest_vec1_def forall_1 by auto - -lemma is_interval_connected_1: "is_interval_new s \ connected (s::(real^1) set)" + "is_interval s \ (\a\s. \b\s. \ x. dest_vec1 a \ dest_vec1 x \ dest_vec1 x \ dest_vec1 b \ x \ s)" + unfolding is_interval_def dest_vec1_def forall_1 by auto + +lemma is_interval_connected_1: "is_interval s \ connected (s::(real^1) set)" apply(rule, rule is_interval_connected, assumption) unfolding is_interval_1 apply(rule,rule,rule,rule,erule conjE,rule ccontr) proof- fix a b x assume as:"connected s" "a \ s" "b \ s" "dest_vec1 a \ dest_vec1 x" "dest_vec1 x \ dest_vec1 b" "x\s" @@ -2219,7 +2210,7 @@ by(auto simp add: basis_component field_simps) qed lemma is_interval_convex_1: - "is_interval_new s \ convex (s::(real^1) set)" + "is_interval s \ convex (s::(real^1) set)" using is_interval_convex convex_connected is_interval_connected_1 by auto lemma convex_connected_1: diff -r 8ef7ba78bf26 -r b4d4dbc5b04f src/HOL/Library/Topology_Euclidean_Space.thy --- a/src/HOL/Library/Topology_Euclidean_Space.thy Thu May 28 15:54:20 2009 +0200 +++ b/src/HOL/Library/Topology_Euclidean_Space.thy Thu May 28 16:19:34 2009 +0200 @@ -4751,16 +4751,12 @@ subsection{* Intervals in general, including infinite and mixtures of open and closed. *} -definition "is_interval s \ (\a\s. \b\s. \x. a \ x \ x \ b \ x \ s)" - -lemma is_interval_interval: fixes a::"real^'n::finite" shows - "is_interval {a<.. (\a\s. \b\s. \x. (\i. ((a$i \ x$i \ x$i \ b$i) \ (b$i \ x$i \ x$i \ a$i))) \ x \ s)" + +lemma is_interval_interval: "is_interval {a .. b::real^'n::finite}" (is ?th1) "is_interval {a<..x y z::real. x < y \ y < z \ x < z" by auto + show ?th1 ?th2 unfolding is_interval_def mem_interval Ball_def atLeastAtMost_iff + by(meson real_le_trans le_less_trans less_le_trans *)+ qed lemma is_interval_empty: "is_interval {}"