--- a/src/HOL/Multivariate_Analysis/Path_Connected.thy Sat Sep 14 22:50:15 2013 +0200
+++ b/src/HOL/Multivariate_Analysis/Path_Connected.thy Sat Sep 14 23:52:36 2013 +0200
@@ -11,7 +11,7 @@
subsection {* Paths. *}
definition path :: "(real \<Rightarrow> 'a::topological_space) \<Rightarrow> bool"
- where "path g \<longleftrightarrow> continuous_on {0 .. 1} g"
+ where "path g \<longleftrightarrow> continuous_on {0..1} g"
definition pathstart :: "(real \<Rightarrow> 'a::topological_space) \<Rightarrow> 'a"
where "pathstart g = g 0"
@@ -22,10 +22,10 @@
definition path_image :: "(real \<Rightarrow> 'a::topological_space) \<Rightarrow> 'a set"
where "path_image g = g ` {0 .. 1}"
-definition reversepath :: "(real \<Rightarrow> 'a::topological_space) \<Rightarrow> (real \<Rightarrow> 'a)"
+definition reversepath :: "(real \<Rightarrow> 'a::topological_space) \<Rightarrow> real \<Rightarrow> 'a"
where "reversepath g = (\<lambda>x. g(1 - x))"
-definition joinpaths :: "(real \<Rightarrow> 'a::topological_space) \<Rightarrow> (real \<Rightarrow> 'a) \<Rightarrow> (real \<Rightarrow> 'a)"
+definition joinpaths :: "(real \<Rightarrow> 'a::topological_space) \<Rightarrow> (real \<Rightarrow> 'a) \<Rightarrow> real \<Rightarrow> 'a"
(infixr "+++" 75)
where "g1 +++ g2 = (\<lambda>x. if x \<le> 1/2 then g1 (2 * x) else g2 (2 * x - 1))"
@@ -40,62 +40,77 @@
subsection {* Some lemmas about these concepts. *}
lemma injective_imp_simple_path: "injective_path g \<Longrightarrow> simple_path g"
- unfolding injective_path_def simple_path_def by auto
+ unfolding injective_path_def simple_path_def
+ by auto
lemma path_image_nonempty: "path_image g \<noteq> {}"
- unfolding path_image_def image_is_empty interval_eq_empty by auto
-
-lemma pathstart_in_path_image[intro]: "(pathstart g) \<in> path_image g"
- unfolding pathstart_def path_image_def by auto
+ unfolding path_image_def image_is_empty interval_eq_empty
+ by auto
-lemma pathfinish_in_path_image[intro]: "(pathfinish g) \<in> path_image g"
- unfolding pathfinish_def path_image_def by auto
+lemma pathstart_in_path_image[intro]: "pathstart g \<in> path_image g"
+ unfolding pathstart_def path_image_def
+ by auto
-lemma connected_path_image[intro]: "path g \<Longrightarrow> connected(path_image g)"
+lemma pathfinish_in_path_image[intro]: "pathfinish g \<in> path_image g"
+ unfolding pathfinish_def path_image_def
+ by auto
+
+lemma connected_path_image[intro]: "path g \<Longrightarrow> connected (path_image g)"
unfolding path_def path_image_def
apply (erule connected_continuous_image)
apply (rule convex_connected, rule convex_real_interval)
done
-lemma compact_path_image[intro]: "path g \<Longrightarrow> compact(path_image g)"
+lemma compact_path_image[intro]: "path g \<Longrightarrow> compact (path_image g)"
unfolding path_def path_image_def
- by (erule compact_continuous_image, rule compact_interval)
+ apply (erule compact_continuous_image)
+ apply (rule compact_interval)
+ done
-lemma reversepath_reversepath[simp]: "reversepath(reversepath g) = g"
- unfolding reversepath_def by auto
+lemma reversepath_reversepath[simp]: "reversepath (reversepath g) = g"
+ unfolding reversepath_def
+ by auto
-lemma pathstart_reversepath[simp]: "pathstart(reversepath g) = pathfinish g"
- unfolding pathstart_def reversepath_def pathfinish_def by auto
+lemma pathstart_reversepath[simp]: "pathstart (reversepath g) = pathfinish g"
+ unfolding pathstart_def reversepath_def pathfinish_def
+ by auto
-lemma pathfinish_reversepath[simp]: "pathfinish(reversepath g) = pathstart g"
- unfolding pathstart_def reversepath_def pathfinish_def by auto
+lemma pathfinish_reversepath[simp]: "pathfinish (reversepath g) = pathstart g"
+ unfolding pathstart_def reversepath_def pathfinish_def
+ by auto
lemma pathstart_join[simp]: "pathstart (g1 +++ g2) = pathstart g1"
- unfolding pathstart_def joinpaths_def pathfinish_def by auto
+ unfolding pathstart_def joinpaths_def pathfinish_def
+ by auto
lemma pathfinish_join[simp]: "pathfinish (g1 +++ g2) = pathfinish g2"
- unfolding pathstart_def joinpaths_def pathfinish_def by auto
+ unfolding pathstart_def joinpaths_def pathfinish_def
+ by auto
-lemma path_image_reversepath[simp]: "path_image(reversepath g) = path_image g"
+lemma path_image_reversepath[simp]: "path_image (reversepath g) = path_image g"
proof -
- have *: "\<And>g. path_image(reversepath g) \<subseteq> path_image g"
+ have *: "\<And>g. path_image (reversepath g) \<subseteq> path_image g"
unfolding path_image_def subset_eq reversepath_def Ball_def image_iff
- apply(rule,rule,erule bexE)
- apply(rule_tac x="1 - xa" in bexI)
+ apply rule
+ apply rule
+ apply (erule bexE)
+ apply (rule_tac x="1 - xa" in bexI)
apply auto
done
show ?thesis
using *[of g] *[of "reversepath g"]
- unfolding reversepath_reversepath by auto
+ unfolding reversepath_reversepath
+ by auto
qed
-lemma path_reversepath[simp]: "path (reversepath g) \<longleftrightarrow> path g"
+lemma path_reversepath [simp]: "path (reversepath g) \<longleftrightarrow> path g"
proof -
have *: "\<And>g. path g \<Longrightarrow> path (reversepath g)"
unfolding path_def reversepath_def
apply (rule continuous_on_compose[unfolded o_def, of _ "\<lambda>x. 1 - x"])
apply (intro continuous_on_intros)
- apply (rule continuous_on_subset[of "{0..1}"], assumption)
+ apply (rule continuous_on_subset[of "{0..1}"])
+ apply assumption
apply auto
done
show ?thesis
@@ -116,43 +131,59 @@
have g1: "continuous_on {0..1} g1 \<longleftrightarrow> continuous_on {0..1} ((g1 +++ g2) \<circ> (\<lambda>x. x / 2))"
by (intro continuous_on_cong refl) (auto simp: joinpaths_def)
have g2: "continuous_on {0..1} g2 \<longleftrightarrow> continuous_on {0..1} ((g1 +++ g2) \<circ> (\<lambda>x. x / 2 + 1/2))"
- using assms by (intro continuous_on_cong refl) (auto simp: joinpaths_def pathfinish_def pathstart_def)
- show "continuous_on {0..1} g1" "continuous_on {0..1} g2"
+ using assms
+ by (intro continuous_on_cong refl) (auto simp: joinpaths_def pathfinish_def pathstart_def)
+ show "continuous_on {0..1} g1" and "continuous_on {0..1} g2"
unfolding g1 g2
by (auto intro!: continuous_on_intros continuous_on_subset[OF cont] simp del: o_apply)
next
assume g1g2: "continuous_on {0..1} g1" "continuous_on {0..1} g2"
have 01: "{0 .. 1} = {0..1/2} \<union> {1/2 .. 1::real}"
by auto
- { fix x :: real assume "0 \<le> x" "x \<le> 1" then have "x \<in> (\<lambda>x. x * 2) ` {0..1 / 2}"
- by (intro image_eqI[where x="x/2"]) auto }
+ {
+ fix x :: real
+ assume "0 \<le> x" and "x \<le> 1"
+ then have "x \<in> (\<lambda>x. x * 2) ` {0..1 / 2}"
+ by (intro image_eqI[where x="x/2"]) auto
+ }
note 1 = this
- { fix x :: real assume "0 \<le> x" "x \<le> 1" then have "x \<in> (\<lambda>x. x * 2 - 1) ` {1 / 2..1}"
- by (intro image_eqI[where x="x/2 + 1/2"]) auto }
+ {
+ fix x :: real
+ assume "0 \<le> x" and "x \<le> 1"
+ then have "x \<in> (\<lambda>x. x * 2 - 1) ` {1 / 2..1}"
+ by (intro image_eqI[where x="x/2 + 1/2"]) auto
+ }
note 2 = this
show "continuous_on {0..1} (g1 +++ g2)"
- using assms unfolding joinpaths_def 01
- by (intro continuous_on_cases closed_atLeastAtMost g1g2[THEN continuous_on_compose2] continuous_on_intros)
- (auto simp: field_simps pathfinish_def pathstart_def intro!: 1 2)
+ using assms
+ unfolding joinpaths_def 01
+ apply (intro continuous_on_cases closed_atLeastAtMost g1g2[THEN continuous_on_compose2] continuous_on_intros)
+ apply (auto simp: field_simps pathfinish_def pathstart_def intro!: 1 2)
+ done
qed
-lemma path_image_join_subset: "path_image(g1 +++ g2) \<subseteq> (path_image g1 \<union> path_image g2)"
- unfolding path_image_def joinpaths_def by auto
+lemma path_image_join_subset: "path_image (g1 +++ g2) \<subseteq> path_image g1 \<union> path_image g2"
+ unfolding path_image_def joinpaths_def
+ by auto
lemma subset_path_image_join:
- assumes "path_image g1 \<subseteq> s" "path_image g2 \<subseteq> s"
- shows "path_image(g1 +++ g2) \<subseteq> s"
- using path_image_join_subset[of g1 g2] and assms by auto
+ assumes "path_image g1 \<subseteq> s"
+ and "path_image g2 \<subseteq> s"
+ shows "path_image (g1 +++ g2) \<subseteq> s"
+ using path_image_join_subset[of g1 g2] and assms
+ by auto
lemma path_image_join:
assumes "pathfinish g1 = pathstart g2"
- shows "path_image(g1 +++ g2) = (path_image g1) \<union> (path_image g2)"
- apply (rule, rule path_image_join_subset, rule)
+ shows "path_image (g1 +++ g2) = path_image g1 \<union> path_image g2"
+ apply rule
+ apply (rule path_image_join_subset)
+ apply rule
unfolding Un_iff
proof (erule disjE)
fix x
assume "x \<in> path_image g1"
- then obtain y where y: "y\<in>{0..1}" "x = g1 y"
+ then obtain y where y: "y \<in> {0..1}" "x = g1 y"
unfolding path_image_def image_iff by auto
then show "x \<in> path_image (g1 +++ g2)"
unfolding joinpaths_def path_image_def image_iff
@@ -162,20 +193,22 @@
next
fix x
assume "x \<in> path_image g2"
- then obtain y where y: "y\<in>{0..1}" "x = g2 y"
+ then obtain y where y: "y \<in> {0..1}" "x = g2 y"
unfolding path_image_def image_iff by auto
then show "x \<in> path_image (g1 +++ g2)"
unfolding joinpaths_def path_image_def image_iff
- apply(rule_tac x="(1/2) *\<^sub>R (y + 1)" in bexI)
+ apply (rule_tac x="(1/2) *\<^sub>R (y + 1)" in bexI)
using assms(1)[unfolded pathfinish_def pathstart_def]
- apply (auto simp add: add_divide_distrib)
+ apply (auto simp add: add_divide_distrib)
done
qed
lemma not_in_path_image_join:
- assumes "x \<notin> path_image g1" "x \<notin> path_image g2"
- shows "x \<notin> path_image(g1 +++ g2)"
- using assms and path_image_join_subset[of g1 g2] by auto
+ assumes "x \<notin> path_image g1"
+ and "x \<notin> path_image g2"
+ shows "x \<notin> path_image (g1 +++ g2)"
+ using assms and path_image_join_subset[of g1 g2]
+ by auto
lemma simple_path_reversepath:
assumes "simple_path g"
@@ -184,82 +217,111 @@
unfolding simple_path_def reversepath_def
apply -
apply (rule ballI)+
- apply (erule_tac x="1-x" in ballE, erule_tac x="1-y" in ballE)
+ apply (erule_tac x="1-x" in ballE)
+ apply (erule_tac x="1-y" in ballE)
apply auto
done
lemma simple_path_join_loop:
- assumes "injective_path g1" "injective_path g2" "pathfinish g2 = pathstart g1"
- "(path_image g1 \<inter> path_image g2) \<subseteq> {pathstart g1,pathstart g2}"
- shows "simple_path(g1 +++ g2)"
+ assumes "injective_path g1"
+ and "injective_path g2"
+ and "pathfinish g2 = pathstart g1"
+ and "path_image g1 \<inter> path_image g2 \<subseteq> {pathstart g1, pathstart g2}"
+ shows "simple_path (g1 +++ g2)"
unfolding simple_path_def
-proof ((rule ballI)+, rule impI)
+proof (intro ballI impI)
let ?g = "g1 +++ g2"
note inj = assms(1,2)[unfolded injective_path_def, rule_format]
fix x y :: real
assume xy: "x \<in> {0..1}" "y \<in> {0..1}" "?g x = ?g y"
show "x = y \<or> x = 0 \<and> y = 1 \<or> x = 1 \<and> y = 0"
- proof (case_tac "x \<le> 1/2", case_tac[!] "y \<le> 1/2", unfold not_le)
+ proof (cases "x \<le> 1/2", case_tac[!] "y \<le> 1/2", unfold not_le)
assume as: "x \<le> 1 / 2" "y \<le> 1 / 2"
then have "g1 (2 *\<^sub>R x) = g1 (2 *\<^sub>R y)"
- using xy(3) unfolding joinpaths_def by auto
- moreover
- have "2 *\<^sub>R x \<in> {0..1}" "2 *\<^sub>R y \<in> {0..1}" using xy(1,2) as
+ using xy(3)
+ unfolding joinpaths_def
+ by auto
+ moreover have "2 *\<^sub>R x \<in> {0..1}" "2 *\<^sub>R y \<in> {0..1}"
+ using xy(1,2) as
by auto
- ultimately
- show ?thesis using inj(1)[of "2*\<^sub>R x" "2*\<^sub>R y"] by auto
+ ultimately show ?thesis
+ using inj(1)[of "2*\<^sub>R x" "2*\<^sub>R y"]
+ by auto
next
- assume as:"x > 1 / 2" "y > 1 / 2"
+ assume as: "x > 1 / 2" "y > 1 / 2"
then have "g2 (2 *\<^sub>R x - 1) = g2 (2 *\<^sub>R y - 1)"
- using xy(3) unfolding joinpaths_def by auto
- moreover
- have "2 *\<^sub>R x - 1 \<in> {0..1}" "2 *\<^sub>R y - 1 \<in> {0..1}"
- using xy(1,2) as by auto
- ultimately
- show ?thesis using inj(2)[of "2*\<^sub>R x - 1" "2*\<^sub>R y - 1"] by auto
+ using xy(3)
+ unfolding joinpaths_def
+ by auto
+ moreover have "2 *\<^sub>R x - 1 \<in> {0..1}" "2 *\<^sub>R y - 1 \<in> {0..1}"
+ using xy(1,2) as
+ by auto
+ ultimately show ?thesis
+ using inj(2)[of "2*\<^sub>R x - 1" "2*\<^sub>R y - 1"] by auto
next
- assume as:"x \<le> 1 / 2" "y > 1 / 2"
+ assume as: "x \<le> 1 / 2" "y > 1 / 2"
then have "?g x \<in> path_image g1" "?g y \<in> path_image g2"
unfolding path_image_def joinpaths_def
using xy(1,2) by auto
- moreover
- have "?g y \<noteq> pathstart g2" using as(2) unfolding pathstart_def joinpaths_def
+ moreover have "?g y \<noteq> pathstart g2"
+ using as(2)
+ unfolding pathstart_def joinpaths_def
using inj(2)[of "2 *\<^sub>R y - 1" 0] and xy(2)
by (auto simp add: field_simps)
- ultimately
- have *:"?g x = pathstart g1" using assms(4) unfolding xy(3) by auto
- then have "x = 0" unfolding pathstart_def joinpaths_def using as(1) and xy(1)
- using inj(1)[of "2 *\<^sub>R x" 0] by auto
- moreover
- have "y = 1" using * unfolding xy(3) assms(3)[THEN sym]
- unfolding joinpaths_def pathfinish_def using as(2) and xy(2)
- using inj(2)[of "2 *\<^sub>R y - 1" 1] by auto
- ultimately show ?thesis by auto
+ ultimately have *: "?g x = pathstart g1"
+ using assms(4)
+ unfolding xy(3)
+ by auto
+ then have "x = 0"
+ unfolding pathstart_def joinpaths_def
+ using as(1) and xy(1)
+ using inj(1)[of "2 *\<^sub>R x" 0]
+ by auto
+ moreover have "y = 1"
+ using *
+ unfolding xy(3) assms(3)[symmetric]
+ unfolding joinpaths_def pathfinish_def
+ using as(2) and xy(2)
+ using inj(2)[of "2 *\<^sub>R y - 1" 1]
+ by auto
+ ultimately show ?thesis
+ by auto
next
assume as: "x > 1 / 2" "y \<le> 1 / 2"
- then have "?g x \<in> path_image g2" "?g y \<in> path_image g1"
+ then have "?g x \<in> path_image g2" and "?g y \<in> path_image g1"
unfolding path_image_def joinpaths_def
using xy(1,2) by auto
- moreover
- have "?g x \<noteq> pathstart g2" using as(1) unfolding pathstart_def joinpaths_def
+ moreover have "?g x \<noteq> pathstart g2"
+ using as(1)
+ unfolding pathstart_def joinpaths_def
using inj(2)[of "2 *\<^sub>R x - 1" 0] and xy(1)
by (auto simp add: field_simps)
- ultimately
- have *: "?g y = pathstart g1" using assms(4) unfolding xy(3) by auto
- then have "y = 0" unfolding pathstart_def joinpaths_def using as(2) and xy(2)
- using inj(1)[of "2 *\<^sub>R y" 0] by auto
- moreover
- have "x = 1" using * unfolding xy(3)[THEN sym] assms(3)[THEN sym]
+ ultimately have *: "?g y = pathstart g1"
+ using assms(4)
+ unfolding xy(3)
+ by auto
+ then have "y = 0"
+ unfolding pathstart_def joinpaths_def
+ using as(2) and xy(2)
+ using inj(1)[of "2 *\<^sub>R y" 0]
+ by auto
+ moreover have "x = 1"
+ using *
+ unfolding xy(3)[symmetric] assms(3)[symmetric]
unfolding joinpaths_def pathfinish_def using as(1) and xy(1)
- using inj(2)[of "2 *\<^sub>R x - 1" 1] by auto
- ultimately show ?thesis by auto
+ using inj(2)[of "2 *\<^sub>R x - 1" 1]
+ by auto
+ ultimately show ?thesis
+ by auto
qed
qed
lemma injective_path_join:
- assumes "injective_path g1" "injective_path g2" "pathfinish g1 = pathstart g2"
- "(path_image g1 \<inter> path_image g2) \<subseteq> {pathstart g2}"
- shows "injective_path(g1 +++ g2)"
+ assumes "injective_path g1"
+ and "injective_path g2"
+ and "pathfinish g1 = pathstart g2"
+ and "path_image g1 \<inter> path_image g2 \<subseteq> {pathstart g2}"
+ shows "injective_path (g1 +++ g2)"
unfolding injective_path_def
proof (rule, rule, rule)
let ?g = "g1 +++ g2"
@@ -268,31 +330,39 @@
assume xy: "x \<in> {0..1}" "y \<in> {0..1}" "(g1 +++ g2) x = (g1 +++ g2) y"
show "x = y"
proof (cases "x \<le> 1/2", case_tac[!] "y \<le> 1/2", unfold not_le)
- assume "x \<le> 1 / 2" "y \<le> 1 / 2"
- then show ?thesis using inj(1)[of "2*\<^sub>R x" "2*\<^sub>R y"] and xy
+ assume "x \<le> 1 / 2" and "y \<le> 1 / 2"
+ then show ?thesis
+ using inj(1)[of "2*\<^sub>R x" "2*\<^sub>R y"] and xy
unfolding joinpaths_def by auto
next
- assume "x > 1 / 2" "y > 1 / 2"
- then show ?thesis using inj(2)[of "2*\<^sub>R x - 1" "2*\<^sub>R y - 1"] and xy
+ assume "x > 1 / 2" and "y > 1 / 2"
+ then show ?thesis
+ using inj(2)[of "2*\<^sub>R x - 1" "2*\<^sub>R y - 1"] and xy
unfolding joinpaths_def by auto
next
assume as: "x \<le> 1 / 2" "y > 1 / 2"
- then have "?g x \<in> path_image g1" "?g y \<in> path_image g2"
+ then have "?g x \<in> path_image g1" and "?g y \<in> path_image g2"
unfolding path_image_def joinpaths_def
- using xy(1,2) by auto
- then have "?g x = pathfinish g1" "?g y = pathstart g2"
- using assms(4) unfolding assms(3) xy(3) by auto
+ using xy(1,2)
+ by auto
+ then have "?g x = pathfinish g1" and "?g y = pathstart g2"
+ using assms(4)
+ unfolding assms(3) xy(3)
+ by auto
then show ?thesis
using as and inj(1)[of "2 *\<^sub>R x" 1] inj(2)[of "2 *\<^sub>R y - 1" 0] and xy(1,2)
unfolding pathstart_def pathfinish_def joinpaths_def
by auto
next
- assume as:"x > 1 / 2" "y \<le> 1 / 2"
- then have "?g x \<in> path_image g2" "?g y \<in> path_image g1"
+ assume as:"x > 1 / 2" "y \<le> 1 / 2"
+ then have "?g x \<in> path_image g2" and "?g y \<in> path_image g1"
unfolding path_image_def joinpaths_def
- using xy(1,2) by auto
- then have "?g x = pathstart g2" "?g y = pathfinish g1"
- using assms(4) unfolding assms(3) xy(3) by auto
+ using xy(1,2)
+ by auto
+ then have "?g x = pathstart g2" and "?g y = pathfinish g1"
+ using assms(4)
+ unfolding assms(3) xy(3)
+ by auto
then show ?thesis using as and inj(2)[of "2 *\<^sub>R x - 1" 0] inj(1)[of "2 *\<^sub>R y" 1] and xy(1,2)
unfolding pathstart_def pathfinish_def joinpaths_def
by auto
@@ -300,64 +370,85 @@
qed
lemmas join_paths_simps = path_join path_image_join pathstart_join pathfinish_join
-
+
-subsection {* Reparametrizing a closed curve to start at some chosen point. *}
+subsection {* Reparametrizing a closed curve to start at some chosen point *}
-definition "shiftpath a (f::real \<Rightarrow> 'a::topological_space) =
- (\<lambda>x. if (a + x) \<le> 1 then f(a + x) else f(a + x - 1))"
+definition shiftpath :: "real \<Rightarrow> (real \<Rightarrow> 'a::topological_space) \<Rightarrow> real \<Rightarrow> 'a"
+ where "shiftpath a f = (\<lambda>x. if (a + x) \<le> 1 then f (a + x) else f (a + x - 1))"
-lemma pathstart_shiftpath: "a \<le> 1 \<Longrightarrow> pathstart(shiftpath a g) = g a"
+lemma pathstart_shiftpath: "a \<le> 1 \<Longrightarrow> pathstart (shiftpath a g) = g a"
unfolding pathstart_def shiftpath_def by auto
lemma pathfinish_shiftpath:
- assumes "0 \<le> a" "pathfinish g = pathstart g"
- shows "pathfinish(shiftpath a g) = g a"
- using assms unfolding pathstart_def pathfinish_def shiftpath_def
+ assumes "0 \<le> a"
+ and "pathfinish g = pathstart g"
+ shows "pathfinish (shiftpath a g) = g a"
+ using assms
+ unfolding pathstart_def pathfinish_def shiftpath_def
by auto
lemma endpoints_shiftpath:
- assumes "pathfinish g = pathstart g" "a \<in> {0 .. 1}"
- shows "pathfinish(shiftpath a g) = g a" "pathstart(shiftpath a g) = g a"
- using assms by(auto intro!:pathfinish_shiftpath pathstart_shiftpath)
+ assumes "pathfinish g = pathstart g"
+ and "a \<in> {0 .. 1}"
+ shows "pathfinish (shiftpath a g) = g a"
+ and "pathstart (shiftpath a g) = g a"
+ using assms
+ by (auto intro!: pathfinish_shiftpath pathstart_shiftpath)
lemma closed_shiftpath:
- assumes "pathfinish g = pathstart g" "a \<in> {0..1}"
- shows "pathfinish(shiftpath a g) = pathstart(shiftpath a g)"
- using endpoints_shiftpath[OF assms] by auto
+ assumes "pathfinish g = pathstart g"
+ and "a \<in> {0..1}"
+ shows "pathfinish (shiftpath a g) = pathstart (shiftpath a g)"
+ using endpoints_shiftpath[OF assms]
+ by auto
lemma path_shiftpath:
- assumes "path g" "pathfinish g = pathstart g" "a \<in> {0..1}"
- shows "path(shiftpath a g)"
+ assumes "path g"
+ and "pathfinish g = pathstart g"
+ and "a \<in> {0..1}"
+ shows "path (shiftpath a g)"
proof -
- have *: "{0 .. 1} = {0 .. 1-a} \<union> {1-a .. 1}" using assms(3) by auto
+ have *: "{0 .. 1} = {0 .. 1-a} \<union> {1-a .. 1}"
+ using assms(3) by auto
have **: "\<And>x. x + a = 1 \<Longrightarrow> g (x + a - 1) = g (x + a)"
- using assms(2)[unfolded pathfinish_def pathstart_def] by auto
+ using assms(2)[unfolded pathfinish_def pathstart_def]
+ by auto
show ?thesis
unfolding path_def shiftpath_def *
apply (rule continuous_on_union)
apply (rule closed_real_atLeastAtMost)+
- apply (rule continuous_on_eq[of _ "g \<circ> (\<lambda>x. a + x)"]) prefer 3
- apply (rule continuous_on_eq[of _ "g \<circ> (\<lambda>x. a - 1 + x)"]) defer prefer 3
- apply (rule continuous_on_intros)+ prefer 2
+ apply (rule continuous_on_eq[of _ "g \<circ> (\<lambda>x. a + x)"])
+ prefer 3
+ apply (rule continuous_on_eq[of _ "g \<circ> (\<lambda>x. a - 1 + x)"])
+ defer
+ prefer 3
+ apply (rule continuous_on_intros)+
+ prefer 2
apply (rule continuous_on_intros)+
apply (rule_tac[1-2] continuous_on_subset[OF assms(1)[unfolded path_def]])
using assms(3) and **
- apply (auto, auto simp add: field_simps)
+ apply auto
+ apply (auto simp add: field_simps)
done
qed
lemma shiftpath_shiftpath:
- assumes "pathfinish g = pathstart g" "a \<in> {0..1}" "x \<in> {0..1}"
+ assumes "pathfinish g = pathstart g"
+ and "a \<in> {0..1}"
+ and "x \<in> {0..1}"
shows "shiftpath (1 - a) (shiftpath a g) x = g x"
- using assms unfolding pathfinish_def pathstart_def shiftpath_def by auto
+ using assms
+ unfolding pathfinish_def pathstart_def shiftpath_def
+ by auto
lemma path_image_shiftpath:
- assumes "a \<in> {0..1}" "pathfinish g = pathstart g"
- shows "path_image(shiftpath a g) = path_image g"
+ assumes "a \<in> {0..1}"
+ and "pathfinish g = pathstart g"
+ shows "path_image (shiftpath a g) = path_image g"
proof -
{ fix x
- assume as:"g 1 = g 0" "x \<in> {0..1::real}" " \<forall>y\<in>{0..1} \<inter> {x. \<not> a + x \<le> 1}. g x \<noteq> g (a + y - 1)"
+ assume as: "g 1 = g 0" "x \<in> {0..1::real}" " \<forall>y\<in>{0..1} \<inter> {x. \<not> a + x \<le> 1}. g x \<noteq> g (a + y - 1)"
then have "\<exists>y\<in>{0..1} \<inter> {x. a + x \<le> 1}. g x = g (a + y)"
proof (cases "a \<le> x")
case False
@@ -368,46 +459,57 @@
done
next
case True
- then show ?thesis using as(1-2) and assms(1) apply(rule_tac x="x - a" in bexI)
- by(auto simp add: field_simps)
+ then show ?thesis
+ using as(1-2) and assms(1)
+ apply (rule_tac x="x - a" in bexI)
+ apply (auto simp add: field_simps)
+ done
qed
}
then show ?thesis
- using assms unfolding shiftpath_def path_image_def pathfinish_def pathstart_def
- by(auto simp add: image_iff)
+ using assms
+ unfolding shiftpath_def path_image_def pathfinish_def pathstart_def
+ by (auto simp add: image_iff)
qed
-subsection {* Special case of straight-line paths. *}
+subsection {* Special case of straight-line paths *}
definition linepath :: "'a::real_normed_vector \<Rightarrow> 'a \<Rightarrow> real \<Rightarrow> 'a"
where "linepath a b = (\<lambda>x. (1 - x) *\<^sub>R a + x *\<^sub>R b)"
-lemma pathstart_linepath[simp]: "pathstart(linepath a b) = a"
- unfolding pathstart_def linepath_def by auto
+lemma pathstart_linepath[simp]: "pathstart (linepath a b) = a"
+ unfolding pathstart_def linepath_def
+ by auto
-lemma pathfinish_linepath[simp]: "pathfinish(linepath a b) = b"
- unfolding pathfinish_def linepath_def by auto
+lemma pathfinish_linepath[simp]: "pathfinish (linepath a b) = b"
+ unfolding pathfinish_def linepath_def
+ by auto
lemma continuous_linepath_at[intro]: "continuous (at x) (linepath a b)"
- unfolding linepath_def by (intro continuous_intros)
+ unfolding linepath_def
+ by (intro continuous_intros)
lemma continuous_on_linepath[intro]: "continuous_on s (linepath a b)"
- using continuous_linepath_at by(auto intro!: continuous_at_imp_continuous_on)
+ using continuous_linepath_at
+ by (auto intro!: continuous_at_imp_continuous_on)
-lemma path_linepath[intro]: "path(linepath a b)"
- unfolding path_def by(rule continuous_on_linepath)
+lemma path_linepath[intro]: "path (linepath a b)"
+ unfolding path_def
+ by (rule continuous_on_linepath)
-lemma path_image_linepath[simp]: "path_image(linepath a b) = (closed_segment a b)"
+lemma path_image_linepath[simp]: "path_image (linepath a b) = closed_segment a b"
unfolding path_image_def segment linepath_def
- apply (rule set_eqI, rule) defer
+ apply (rule set_eqI)
+ apply rule
+ defer
unfolding mem_Collect_eq image_iff
- apply(erule exE)
- apply(rule_tac x="u *\<^sub>R 1" in bexI)
+ apply (erule exE)
+ apply (rule_tac x="u *\<^sub>R 1" in bexI)
apply auto
done
-lemma reversepath_linepath[simp]: "reversepath(linepath a b) = linepath b a"
+lemma reversepath_linepath[simp]: "reversepath (linepath a b) = linepath b a"
unfolding reversepath_def linepath_def
by auto
@@ -415,25 +517,30 @@
assumes "a \<noteq> b"
shows "injective_path (linepath a b)"
proof -
- { fix x y :: "real"
+ {
+ fix x y :: "real"
assume "x *\<^sub>R b + y *\<^sub>R a = x *\<^sub>R a + y *\<^sub>R b"
- then have "(x - y) *\<^sub>R a = (x - y) *\<^sub>R b" by (simp add: algebra_simps)
- with assms have "x = y" by simp }
+ then have "(x - y) *\<^sub>R a = (x - y) *\<^sub>R b"
+ by (simp add: algebra_simps)
+ with assms have "x = y"
+ by simp
+ }
then show ?thesis
unfolding injective_path_def linepath_def
by (auto simp add: algebra_simps)
qed
-lemma simple_path_linepath[intro]: "a \<noteq> b \<Longrightarrow> simple_path(linepath a b)"
- by(auto intro!: injective_imp_simple_path injective_path_linepath)
+lemma simple_path_linepath[intro]: "a \<noteq> b \<Longrightarrow> simple_path (linepath a b)"
+ by (auto intro!: injective_imp_simple_path injective_path_linepath)
-subsection {* Bounding a point away from a path. *}
+subsection {* Bounding a point away from a path *}
lemma not_on_path_ball:
fixes g :: "real \<Rightarrow> 'a::heine_borel"
- assumes "path g" "z \<notin> path_image g"
- shows "\<exists>e > 0. ball z e \<inter> (path_image g) = {}"
+ assumes "path g"
+ and "z \<notin> path_image g"
+ shows "\<exists>e > 0. ball z e \<inter> path_image g = {}"
proof -
obtain a where "a \<in> path_image g" "\<forall>y \<in> path_image g. dist z a \<le> dist z y"
using distance_attains_inf[OF _ path_image_nonempty, of g z]
@@ -447,34 +554,43 @@
lemma not_on_path_cball:
fixes g :: "real \<Rightarrow> 'a::heine_borel"
- assumes "path g" "z \<notin> path_image g"
+ assumes "path g"
+ and "z \<notin> path_image g"
shows "\<exists>e>0. cball z e \<inter> (path_image g) = {}"
proof -
- obtain e where "ball z e \<inter> path_image g = {}" "e>0"
+ obtain e where "ball z e \<inter> path_image g = {}" "e > 0"
using not_on_path_ball[OF assms] by auto
- moreover have "cball z (e/2) \<subseteq> ball z e" using `e>0` by auto
- ultimately show ?thesis apply(rule_tac x="e/2" in exI) by auto
+ moreover have "cball z (e/2) \<subseteq> ball z e"
+ using `e > 0` by auto
+ ultimately show ?thesis
+ apply (rule_tac x="e/2" in exI)
+ apply auto
+ done
qed
-subsection {* Path component, considered as a "joinability" relation (from Tom Hales). *}
+subsection {* Path component, considered as a "joinability" relation (from Tom Hales) *}
definition "path_component s x y \<longleftrightarrow>
(\<exists>g. path g \<and> path_image g \<subseteq> s \<and> pathstart g = x \<and> pathfinish g = y)"
-lemmas path_defs = path_def pathstart_def pathfinish_def path_image_def path_component_def
+lemmas path_defs = path_def pathstart_def pathfinish_def path_image_def path_component_def
lemma path_component_mem:
assumes "path_component s x y"
- shows "x \<in> s" "y \<in> s"
- using assms unfolding path_defs by auto
+ shows "x \<in> s" and "y \<in> s"
+ using assms
+ unfolding path_defs
+ by auto
lemma path_component_refl:
assumes "x \<in> s"
shows "path_component s x x"
unfolding path_defs
apply (rule_tac x="\<lambda>u. x" in exI)
- using assms apply (auto intro!:continuous_on_intros) done
+ using assms
+ apply (auto intro!: continuous_on_intros)
+ done
lemma path_component_refl_eq: "path_component s x x \<longleftrightarrow> x \<in> s"
by (auto intro!: path_component_mem path_component_refl)
@@ -488,21 +604,21 @@
done
lemma path_component_trans:
- assumes "path_component s x y" "path_component s y z"
+ assumes "path_component s x y"
+ and "path_component s y z"
shows "path_component s x z"
using assms
unfolding path_component_def
- apply -
- apply (erule exE)+
+ apply (elim exE)
apply (rule_tac x="g +++ ga" in exI)
apply (auto simp add: path_image_join)
done
-lemma path_component_of_subset: "s \<subseteq> t \<Longrightarrow> path_component s x y \<Longrightarrow> path_component t x y"
+lemma path_component_of_subset: "s \<subseteq> t \<Longrightarrow> path_component s x y \<Longrightarrow> path_component t x y"
unfolding path_component_def by auto
-subsection {* Can also consider it as a set, as the name suggests. *}
+text {* Can also consider it as a set, as the name suggests. *}
lemma path_component_set:
"{y. path_component s x y} =
@@ -514,13 +630,15 @@
done
lemma path_component_subset: "{y. path_component s x y} \<subseteq> s"
- apply (rule, rule path_component_mem(2))
+ apply rule
+ apply (rule path_component_mem(2))
apply auto
done
lemma path_component_eq_empty: "{y. path_component s x y} = {} \<longleftrightarrow> x \<notin> s"
apply rule
- apply (drule equals0D[of _ x]) defer
+ apply (drule equals0D[of _ x])
+ defer
apply (rule equals0I)
unfolding mem_Collect_eq
apply (drule path_component_mem(1))
@@ -529,29 +647,35 @@
done
-subsection {* Path connectedness of a space. *}
+subsection {* Path connectedness of a space *}
definition "path_connected s \<longleftrightarrow>
- (\<forall>x\<in>s. \<forall>y\<in>s. \<exists>g. path g \<and> (path_image g) \<subseteq> s \<and> pathstart g = x \<and> pathfinish g = y)"
+ (\<forall>x\<in>s. \<forall>y\<in>s. \<exists>g. path g \<and> path_image g \<subseteq> s \<and> pathstart g = x \<and> pathfinish g = y)"
lemma path_connected_component: "path_connected s \<longleftrightarrow> (\<forall>x\<in>s. \<forall>y\<in>s. path_component s x y)"
unfolding path_connected_def path_component_def by auto
-lemma path_connected_component_set: "path_connected s \<longleftrightarrow> (\<forall>x\<in>s. {y. path_component s x y} = s)"
+lemma path_connected_component_set: "path_connected s \<longleftrightarrow> (\<forall>x\<in>s. {y. path_component s x y} = s)"
unfolding path_connected_component
- apply (rule, rule, rule, rule path_component_subset)
+ apply rule
+ apply rule
+ apply rule
+ apply (rule path_component_subset)
unfolding subset_eq mem_Collect_eq Ball_def
apply auto
done
-subsection {* Some useful lemmas about path-connectedness. *}
+subsection {* Some useful lemmas about path-connectedness *}
lemma convex_imp_path_connected:
fixes s :: "'a::real_normed_vector set"
- assumes "convex s" shows "path_connected s"
+ assumes "convex s"
+ shows "path_connected s"
unfolding path_connected_def
- apply (rule, rule, rule_tac x = "linepath x y" in exI)
+ apply rule
+ apply rule
+ apply (rule_tac x = "linepath x y" in exI)
unfolding path_image_linepath
using assms [unfolded convex_contains_segment]
apply auto
@@ -561,26 +685,33 @@
assumes "path_connected s"
shows "connected s"
unfolding connected_def not_ex
- apply (rule, rule, rule ccontr)
+ apply rule
+ apply rule
+ apply (rule ccontr)
unfolding not_not
- apply (erule conjE)+
+ apply (elim conjE)
proof -
fix e1 e2
assume as: "open e1" "open e2" "s \<subseteq> e1 \<union> e2" "e1 \<inter> e2 \<inter> s = {}" "e1 \<inter> s \<noteq> {}" "e2 \<inter> s \<noteq> {}"
- then obtain x1 x2 where obt:"x1\<in>e1\<inter>s" "x2\<in>e2\<inter>s" by auto
- then obtain g where g:"path g" "path_image g \<subseteq> s" "pathstart g = x1" "pathfinish g = x2"
+ then obtain x1 x2 where obt:"x1 \<in> e1 \<inter> s" "x2 \<in> e2 \<inter> s"
+ by auto
+ then obtain g where g: "path g" "path_image g \<subseteq> s" "pathstart g = x1" "pathfinish g = x2"
using assms[unfolded path_connected_def,rule_format,of x1 x2] by auto
have *: "connected {0..1::real}"
by (auto intro!: convex_connected convex_real_interval)
have "{0..1} \<subseteq> {x \<in> {0..1}. g x \<in> e1} \<union> {x \<in> {0..1}. g x \<in> e2}"
using as(3) g(2)[unfolded path_defs] by blast
moreover have "{x \<in> {0..1}. g x \<in> e1} \<inter> {x \<in> {0..1}. g x \<in> e2} = {}"
- using as(4) g(2)[unfolded path_defs] unfolding subset_eq by auto
+ using as(4) g(2)[unfolded path_defs]
+ unfolding subset_eq
+ by auto
moreover have "{x \<in> {0..1}. g x \<in> e1} \<noteq> {} \<and> {x \<in> {0..1}. g x \<in> e2} \<noteq> {}"
- using g(3,4)[unfolded path_defs] using obt
+ using g(3,4)[unfolded path_defs]
+ using obt
by (simp add: ex_in_conv [symmetric], metis zero_le_one order_refl)
ultimately show False
- using *[unfolded connected_local not_ex, rule_format, of "{x\<in>{0..1}. g x \<in> e1}" "{x\<in>{0..1}. g x \<in> e2}"]
+ using *[unfolded connected_local not_ex, rule_format,
+ of "{x\<in>{0..1}. g x \<in> e1}" "{x\<in>{0..1}. g x \<in> e2}"]
using continuous_open_in_preimage[OF g(1)[unfolded path_def] as(1)]
using continuous_open_in_preimage[OF g(1)[unfolded path_def] as(2)]
by auto
@@ -600,20 +731,23 @@
unfolding mem_Collect_eq
apply auto
done
- then obtain e where e:"e>0" "ball y e \<subseteq> s"
- using assms[unfolded open_contains_ball] by auto
+ then obtain e where e: "e > 0" "ball y e \<subseteq> s"
+ using assms[unfolded open_contains_ball]
+ by auto
show "\<exists>e > 0. ball y e \<subseteq> {y. path_component s x y}"
apply (rule_tac x=e in exI)
- apply (rule,rule `e>0`, rule)
+ apply (rule,rule `e>0`)
+ apply rule
unfolding mem_ball mem_Collect_eq
proof -
fix z
assume "dist y z < e"
then show "path_component s x z"
- apply (rule_tac path_component_trans[of _ _ y]) defer
+ apply (rule_tac path_component_trans[of _ _ y])
+ defer
apply (rule path_component_of_subset[OF e(2)])
apply (rule convex_imp_path_connected[OF convex_ball, unfolded path_connected_component, rule_format])
- using `e>0` as
+ using `e > 0` as
apply auto
done
qed
@@ -622,16 +756,21 @@
lemma open_non_path_component:
fixes s :: "'a::real_normed_vector set"
assumes "open s"
- shows "open(s - {y. path_component s x y})"
+ shows "open (s - {y. path_component s x y})"
unfolding open_contains_ball
proof
fix y
- assume as: "y\<in>s - {y. path_component s x y}"
- then obtain e where e:"e>0" "ball y e \<subseteq> s"
- using assms [unfolded open_contains_ball] by auto
+ assume as: "y \<in> s - {y. path_component s x y}"
+ then obtain e where e: "e > 0" "ball y e \<subseteq> s"
+ using assms [unfolded open_contains_ball]
+ by auto
show "\<exists>e>0. ball y e \<subseteq> s - {y. path_component s x y}"
apply (rule_tac x=e in exI)
- apply (rule, rule `e>0`, rule, rule) defer
+ apply rule
+ apply (rule `e>0`)
+ apply rule
+ apply rule
+ defer
proof (rule ccontr)
fix z
assume "z \<in> ball y e" "\<not> z \<notin> {y. path_component s x y}"
@@ -643,43 +782,49 @@
apply (rule convex_imp_path_connected[OF convex_ball, unfolded path_connected_component, rule_format])
apply auto
done
- then show False using as by auto
+ then show False
+ using as by auto
qed (insert e(2), auto)
qed
lemma connected_open_path_connected:
fixes s :: "'a::real_normed_vector set"
- assumes "open s" "connected s"
+ assumes "open s"
+ and "connected s"
shows "path_connected s"
unfolding path_connected_component_set
proof (rule, rule, rule path_component_subset, rule)
fix x y
- assume "x \<in> s" "y \<in> s"
+ assume "x \<in> s" and "y \<in> s"
show "y \<in> {y. path_component s x y}"
proof (rule ccontr)
- assume "y \<notin> {y. path_component s x y}"
- moreover
- have "{y. path_component s x y} \<inter> s \<noteq> {}"
- using `x\<in>s` path_component_eq_empty path_component_subset[of s x] by auto
+ assume "\<not> ?thesis"
+ moreover have "{y. path_component s x y} \<inter> s \<noteq> {}"
+ using `x \<in> s` path_component_eq_empty path_component_subset[of s x]
+ by auto
ultimately
show False
- using `y\<in>s` open_non_path_component[OF assms(1)] open_path_component[OF assms(1)]
- using assms(2)[unfolded connected_def not_ex, rule_format, of"{y. path_component s x y}" "s - {y. path_component s x y}"]
+ using `y \<in> s` open_non_path_component[OF assms(1)] open_path_component[OF assms(1)]
+ using assms(2)[unfolded connected_def not_ex, rule_format,
+ of"{y. path_component s x y}" "s - {y. path_component s x y}"]
by auto
qed
qed
lemma path_connected_continuous_image:
- assumes "continuous_on s f" "path_connected s"
+ assumes "continuous_on s f"
+ and "path_connected s"
shows "path_connected (f ` s)"
unfolding path_connected_def
proof (rule, rule)
fix x' y'
assume "x' \<in> f ` s" "y' \<in> f ` s"
- then obtain x y where xy: "x\<in>s" "y\<in>s" "x' = f x" "y' = f y" by auto
- guess g using assms(2)[unfolded path_connected_def, rule_format, OF xy(1,2)] ..
+ then obtain x y where x: "x \<in> s" and y: "y \<in> s" and x': "x' = f x" and y': "y' = f y"
+ by auto
+ from x y obtain g where "path g \<and> path_image g \<subseteq> s \<and> pathstart g = x \<and> pathfinish g = y"
+ using assms(2)[unfolded path_connected_def] by fast
then show "\<exists>g. path g \<and> path_image g \<subseteq> f ` s \<and> pathstart g = x' \<and> pathfinish g = y'"
- unfolding xy
+ unfolding x' y'
apply (rule_tac x="f \<circ> g" in exI)
unfolding path_defs
apply (intro conjI continuous_on_compose continuous_on_subset[OF assms(1)])
@@ -688,11 +833,12 @@
qed
lemma homeomorphic_path_connectedness:
- "s homeomorphic t \<Longrightarrow> (path_connected s \<longleftrightarrow> path_connected t)"
+ "s homeomorphic t \<Longrightarrow> path_connected s \<longleftrightarrow> path_connected t"
unfolding homeomorphic_def homeomorphism_def
- apply (erule exE|erule conjE)+
+ apply (erule exE|erule conjE)+
apply rule
- apply (drule_tac f=f in path_connected_continuous_image) prefer 3
+ apply (drule_tac f=f in path_connected_continuous_image)
+ prefer 3
apply (drule_tac f=g in path_connected_continuous_image)
apply auto
done
@@ -702,21 +848,26 @@
lemma path_connected_singleton: "path_connected {a}"
unfolding path_connected_def pathstart_def pathfinish_def path_image_def
- apply (clarify, rule_tac x="\<lambda>x. a" in exI, simp add: image_constant_conv)
+ apply clarify
+ apply (rule_tac x="\<lambda>x. a" in exI)
+ apply (simp add: image_constant_conv)
apply (simp add: path_def continuous_on_const)
done
lemma path_connected_Un:
- assumes "path_connected s" "path_connected t" "s \<inter> t \<noteq> {}"
+ assumes "path_connected s"
+ and "path_connected t"
+ and "s \<inter> t \<noteq> {}"
shows "path_connected (s \<union> t)"
unfolding path_connected_component
proof (rule, rule)
fix x y
assume as: "x \<in> s \<union> t" "y \<in> s \<union> t"
- from assms(3) obtain z where "z \<in> s \<inter> t" by auto
+ from assms(3) obtain z where "z \<in> s \<inter> t"
+ by auto
then show "path_component (s \<union> t) x y"
using as and assms(1-2)[unfolded path_connected_component]
- apply -
+ apply -
apply (erule_tac[!] UnE)+
apply (rule_tac[2-3] path_component_trans[of _ _ z])
apply (auto simp add:path_component_of_subset [OF Un_upper1] path_component_of_subset[OF Un_upper2])
@@ -740,11 +891,11 @@
qed
-subsection {* sphere is path-connected. *}
+subsection {* Sphere is path-connected *}
lemma path_connected_punctured_universe:
assumes "2 \<le> DIM('a::euclidean_space)"
- shows "path_connected((UNIV::'a::euclidean_space set) - {a})"
+ shows "path_connected ((UNIV::'a set) - {a})"
proof -
let ?A = "{x::'a. \<exists>i\<in>Basis. x \<bullet> i < a \<bullet> i}"
let ?B = "{x::'a. \<exists>i\<in>Basis. a \<bullet> i < x \<bullet> i}"
@@ -754,70 +905,87 @@
proof (rule path_connected_UNION)
fix i :: 'a
assume "i \<in> Basis"
- then show "(\<Sum>i\<in>Basis. (a \<bullet> i - 1)*\<^sub>R i) \<in> {x::'a. x \<bullet> i < a \<bullet> i}" by simp
+ then show "(\<Sum>i\<in>Basis. (a \<bullet> i - 1)*\<^sub>R i) \<in> {x::'a. x \<bullet> i < a \<bullet> i}"
+ by simp
show "path_connected {x. x \<bullet> i < a \<bullet> i}"
using convex_imp_path_connected [OF convex_halfspace_lt, of i "a \<bullet> i"]
by (simp add: inner_commute)
qed
- have B: "path_connected ?B" unfolding Collect_bex_eq
+ have B: "path_connected ?B"
+ unfolding Collect_bex_eq
proof (rule path_connected_UNION)
fix i :: 'a
assume "i \<in> Basis"
- then show "(\<Sum>i\<in>Basis. (a \<bullet> i + 1) *\<^sub>R i) \<in> {x::'a. a \<bullet> i < x \<bullet> i}" by simp
+ then show "(\<Sum>i\<in>Basis. (a \<bullet> i + 1) *\<^sub>R i) \<in> {x::'a. a \<bullet> i < x \<bullet> i}"
+ by simp
show "path_connected {x. a \<bullet> i < x \<bullet> i}"
using convex_imp_path_connected [OF convex_halfspace_gt, of "a \<bullet> i" i]
by (simp add: inner_commute)
qed
- obtain S :: "'a set" where "S \<subseteq> Basis" "card S = Suc (Suc 0)"
- using ex_card[OF assms] by auto
- then obtain b0 b1 :: 'a where "b0 \<in> Basis" "b1 \<in> Basis" "b0 \<noteq> b1"
+ obtain S :: "'a set" where "S \<subseteq> Basis" and "card S = Suc (Suc 0)"
+ using ex_card[OF assms]
+ by auto
+ then obtain b0 b1 :: 'a where "b0 \<in> Basis" and "b1 \<in> Basis" and "b0 \<noteq> b1"
unfolding card_Suc_eq by auto
- then have "a + b0 - b1 \<in> ?A \<inter> ?B" by (auto simp: inner_simps inner_Basis)
- then have "?A \<inter> ?B \<noteq> {}" by fast
+ then have "a + b0 - b1 \<in> ?A \<inter> ?B"
+ by (auto simp: inner_simps inner_Basis)
+ then have "?A \<inter> ?B \<noteq> {}"
+ by fast
with A B have "path_connected (?A \<union> ?B)"
by (rule path_connected_Un)
also have "?A \<union> ?B = {x. \<exists>i\<in>Basis. x \<bullet> i \<noteq> a \<bullet> i}"
unfolding neq_iff bex_disj_distrib Collect_disj_eq ..
also have "\<dots> = {x. x \<noteq> a}"
- unfolding euclidean_eq_iff [where 'a='a] by (simp add: Bex_def)
- also have "\<dots> = UNIV - {a}" by auto
+ unfolding euclidean_eq_iff [where 'a='a]
+ by (simp add: Bex_def)
+ also have "\<dots> = UNIV - {a}"
+ by auto
finally show ?thesis .
qed
lemma path_connected_sphere:
assumes "2 \<le> DIM('a::euclidean_space)"
- shows "path_connected {x::'a::euclidean_space. norm(x - a) = r}"
+ shows "path_connected {x::'a. norm (x - a) = r}"
proof (rule linorder_cases [of r 0])
assume "r < 0"
- then have "{x::'a. norm(x - a) = r} = {}" by auto
- then show ?thesis using path_connected_empty by simp
+ then have "{x::'a. norm(x - a) = r} = {}"
+ by auto
+ then show ?thesis
+ using path_connected_empty by simp
next
assume "r = 0"
- then show ?thesis using path_connected_singleton by simp
+ then show ?thesis
+ using path_connected_singleton by simp
next
assume r: "0 < r"
- then have *: "{x::'a. norm(x - a) = r} = (\<lambda>x. a + r *\<^sub>R x) ` {x. norm x = 1}"
- apply -
- apply (rule set_eqI, rule)
+ have *: "{x::'a. norm(x - a) = r} = (\<lambda>x. a + r *\<^sub>R x) ` {x. norm x = 1}"
+ apply (rule set_eqI)
+ apply rule
unfolding image_iff
apply (rule_tac x="(1/r) *\<^sub>R (x - a)" in bexI)
unfolding mem_Collect_eq norm_scaleR
+ using r
apply (auto simp add: scaleR_right_diff_distrib)
done
have **: "{x::'a. norm x = 1} = (\<lambda>x. (1/norm x) *\<^sub>R x) ` (UNIV - {0})"
- apply (rule set_eqI,rule)
+ apply (rule set_eqI)
+ apply rule
unfolding image_iff
apply (rule_tac x=x in bexI)
unfolding mem_Collect_eq
- apply (auto split:split_if_asm)
+ apply (auto split: split_if_asm)
done
have "continuous_on (UNIV - {0}) (\<lambda>x::'a. 1 / norm x)"
- unfolding field_divide_inverse by (simp add: continuous_on_intros)
- then show ?thesis unfolding * ** using path_connected_punctured_universe[OF assms]
+ unfolding field_divide_inverse
+ by (simp add: continuous_on_intros)
+ then show ?thesis
+ unfolding * **
+ using path_connected_punctured_universe[OF assms]
by (auto intro!: path_connected_continuous_image continuous_on_intros)
qed
-lemma connected_sphere: "2 \<le> DIM('a::euclidean_space) \<Longrightarrow> connected {x::'a. norm(x - a) = r}"
- using path_connected_sphere path_connected_imp_connected by auto
+lemma connected_sphere: "2 \<le> DIM('a::euclidean_space) \<Longrightarrow> connected {x::'a. norm (x - a) = r}"
+ using path_connected_sphere path_connected_imp_connected
+ by auto
end
--- a/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Sat Sep 14 22:50:15 2013 +0200
+++ b/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Sat Sep 14 23:52:36 2013 +0200
@@ -25,7 +25,7 @@
using dist_triangle[of y z x] by (simp add: dist_commute)
(* LEGACY *)
-lemma lim_subseq: "subseq r \<Longrightarrow> s ----> l \<Longrightarrow> (s o r) ----> l"
+lemma lim_subseq: "subseq r \<Longrightarrow> s ----> l \<Longrightarrow> (s \<circ> r) ----> l"
by (rule LIMSEQ_subseq_LIMSEQ)
lemmas real_isGlb_unique = isGlb_unique[where 'a=real]
@@ -85,7 +85,7 @@
show "topological_basis B"
using assms unfolding topological_basis_def
proof safe
- fix O'::"'a set"
+ fix O' :: "'a set"
assume "open O'"
with H obtain f where "\<forall>x\<in>O'. f x \<in> B \<and> x \<in> f x \<and> f x \<subseteq> O'"
by (force intro: bchoice simp: Bex_def)
@@ -138,14 +138,14 @@
qed
lemma basis_dense:
- fixes B::"'a set set"
- and f::"'a set \<Rightarrow> 'a"
+ fixes B :: "'a set set"
+ and f :: "'a set \<Rightarrow> 'a"
assumes "topological_basis B"
and choosefrom_basis: "\<And>B'. B' \<noteq> {} \<Longrightarrow> f B' \<in> B'"
shows "(\<forall>X. open X \<longrightarrow> X \<noteq> {} \<longrightarrow> (\<exists>B' \<in> B. f B' \<in> X))"
proof (intro allI impI)
- fix X::"'a set"
- assume "open X" "X \<noteq> {}"
+ fix X :: "'a set"
+ assume "open X" and "X \<noteq> {}"
from topological_basisE[OF `topological_basis B` `open X` choosefrom_basis[OF `X \<noteq> {}`]]
guess B' . note B' = this
then show "\<exists>B'\<in>B. f B' \<in> X"
@@ -180,7 +180,7 @@
subsection {* Countable Basis *}
locale countable_basis =
- fixes B::"'a::topological_space set set"
+ fixes B :: "'a::topological_space set set"
assumes is_basis: "topological_basis B"
and countable_basis: "countable B"
begin
@@ -283,8 +283,9 @@
proof (rule first_countableI[of "(\<lambda>(a, b). a \<times> b) ` (A \<times> B)"], safe)
fix a b
assume x: "a \<in> A" "b \<in> B"
- with A(2, 3)[of a] B(2, 3)[of b] show "x \<in> a \<times> b" "open (a \<times> b)"
- unfolding mem_Times_iff by (auto intro: open_Times)
+ with A(2, 3)[of a] B(2, 3)[of b] show "x \<in> a \<times> b" and "open (a \<times> b)"
+ unfolding mem_Times_iff
+ by (auto intro: open_Times)
next
fix S
assume "open S" "x \<in> S"
@@ -418,7 +419,7 @@
text{* Infer the "universe" from union of all sets in the topology. *}
-definition "topspace T = \<Union>{S. openin T S}"
+definition "topspace T = \<Union>{S. openin T S}"
subsubsection {* Main properties of open sets *}
@@ -1007,7 +1008,7 @@
lemma islimpt_approachable_le:
fixes x :: "'a::metric_space"
- shows "x islimpt S \<longleftrightarrow> (\<forall>e>0. \<exists>x'\<in> S. x' \<noteq> x \<and> dist x' x <= e)"
+ shows "x islimpt S \<longleftrightarrow> (\<forall>e>0. \<exists>x'\<in> S. x' \<noteq> x \<and> dist x' x \<le> e)"
unfolding islimpt_approachable
using approachable_lt_le [where f="\<lambda>y. dist y x" and P="\<lambda>y. y \<notin> S \<or> y = x",
THEN arg_cong [where f=Not]]
@@ -1043,7 +1044,7 @@
lemma finite_set_avoid:
fixes a :: "'a::metric_space"
assumes fS: "finite S"
- shows "\<exists>d>0. \<forall>x\<in>S. x \<noteq> a \<longrightarrow> d <= dist a x"
+ shows "\<exists>d>0. \<forall>x\<in>S. x \<noteq> a \<longrightarrow> d \<le> dist a x"
proof (induct rule: finite_induct[OF fS])
case 1
then show ?case by (auto intro: zero_less_one)
@@ -1423,8 +1424,9 @@
apply (drule_tac x=UNIV in spec, simp)
done
-lemma not_trivial_limit_within: "~trivial_limit (at x within S) = (x:closure(S-{x}))"
- using islimpt_in_closure by (metis trivial_limit_within)
+lemma not_trivial_limit_within: "\<not> trivial_limit (at x within S) = (x \<in> closure (S - {x}))"
+ using islimpt_in_closure
+ by (metis trivial_limit_within)
text {* Some property holds "sufficiently close" to the limit point. *}
@@ -1463,19 +1465,19 @@
text{* Show that they yield usual definitions in the various cases. *}
lemma Lim_within_le: "(f ---> l)(at a within S) \<longleftrightarrow>
- (\<forall>e>0. \<exists>d>0. \<forall>x\<in>S. 0 < dist x a \<and> dist x a <= d \<longrightarrow> dist (f x) l < e)"
+ (\<forall>e>0. \<exists>d>0. \<forall>x\<in>S. 0 < dist x a \<and> dist x a \<le> d \<longrightarrow> dist (f x) l < e)"
by (auto simp add: tendsto_iff eventually_at_le dist_nz)
lemma Lim_within: "(f ---> l) (at a within S) \<longleftrightarrow>
- (\<forall>e >0. \<exists>d>0. \<forall>x \<in> S. 0 < dist x a \<and> dist x a < d \<longrightarrow> dist (f x) l < e)"
+ (\<forall>e >0. \<exists>d>0. \<forall>x \<in> S. 0 < dist x a \<and> dist x a < d \<longrightarrow> dist (f x) l < e)"
by (auto simp add: tendsto_iff eventually_at dist_nz)
lemma Lim_at: "(f ---> l) (at a) \<longleftrightarrow>
- (\<forall>e >0. \<exists>d>0. \<forall>x. 0 < dist x a \<and> dist x a < d \<longrightarrow> dist (f x) l < e)"
+ (\<forall>e >0. \<exists>d>0. \<forall>x. 0 < dist x a \<and> dist x a < d \<longrightarrow> dist (f x) l < e)"
by (auto simp add: tendsto_iff eventually_at2)
lemma Lim_at_infinity:
- "(f ---> l) at_infinity \<longleftrightarrow> (\<forall>e>0. \<exists>b. \<forall>x. norm x >= b \<longrightarrow> dist (f x) l < e)"
+ "(f ---> l) at_infinity \<longleftrightarrow> (\<forall>e>0. \<exists>b. \<forall>x. norm x \<ge> b \<longrightarrow> dist (f x) l < e)"
by (auto simp add: tendsto_iff eventually_at_infinity)
lemma Lim_eventually: "eventually (\<lambda>x. f x = l) net \<Longrightarrow> (f ---> l) net"
@@ -1489,7 +1491,8 @@
lemma Lim_Un:
assumes "(f ---> l) (at x within S)" "(f ---> l) (at x within T)"
shows "(f ---> l) (at x within (S \<union> T))"
- using assms unfolding tendsto_def eventually_at_filter
+ using assms
+ unfolding tendsto_def eventually_at_filter
apply clarify
apply (drule spec, drule (1) mp, drule (1) mp)
apply (drule spec, drule (1) mp, drule (1) mp)
@@ -1515,10 +1518,10 @@
from assms obtain T where T: "open T" "x \<in> T" "T \<subseteq> S" ..
{
assume "?lhs"
- then obtain A where "open A" "x \<in> A" "\<forall>y\<in>A. y \<noteq> x \<longrightarrow> y \<in> S \<longrightarrow> P y"
+ then obtain A where "open A" and "x \<in> A" and "\<forall>y\<in>A. y \<noteq> x \<longrightarrow> y \<in> S \<longrightarrow> P y"
unfolding eventually_at_topological
by auto
- with T have "open (A \<inter> T)" "x \<in> A \<inter> T" "\<forall>y\<in>(A \<inter> T). y \<noteq> x \<longrightarrow> P y"
+ with T have "open (A \<inter> T)" and "x \<in> A \<inter> T" and "\<forall>y \<in> A \<inter> T. y \<noteq> x \<longrightarrow> P y"
by auto
then show "?rhs"
unfolding eventually_at_topological by auto
@@ -1546,11 +1549,11 @@
assumes mono: "\<And>a b. a \<in> I \<Longrightarrow> b \<in> I \<Longrightarrow> x < a \<Longrightarrow> a \<le> b \<Longrightarrow> f a \<le> f b"
and bnd: "\<And>a. a \<in> I \<Longrightarrow> x < a \<Longrightarrow> K \<le> f a"
shows "(f ---> Inf (f ` ({x<..} \<inter> I))) (at x within ({x<..} \<inter> I))"
-proof cases
- assume "{x<..} \<inter> I = {}"
+proof (cases "{x<..} \<inter> I = {}")
+ case True
then show ?thesis by (simp add: Lim_within_empty)
next
- assume e: "{x<..} \<inter> I \<noteq> {}"
+ case False
show ?thesis
proof (rule order_tendstoI)
fix a
@@ -1558,7 +1561,7 @@
{
fix y
assume "y \<in> {x<..} \<inter> I"
- with e bnd have "Inf (f ` ({x<..} \<inter> I)) \<le> f y"
+ with False bnd have "Inf (f ` ({x<..} \<inter> I)) \<le> f y"
by (auto intro: cInf_lower)
with a have "a < f y"
by (blast intro: less_le_trans)
@@ -1568,7 +1571,7 @@
next
fix a
assume "Inf (f ` ({x<..} \<inter> I)) < a"
- from cInf_lessD[OF _ this] e obtain y where y: "x < y" "y \<in> I" "f y < a"
+ from cInf_lessD[OF _ this] False obtain y where y: "x < y" "y \<in> I" "f y < a"
by auto
then have "eventually (\<lambda>x. x \<in> I \<longrightarrow> f x < a) (at_right x)"
unfolding eventually_at_right by (metis less_imp_le le_less_trans mono)
@@ -1625,7 +1628,7 @@
and A :: "'a filter"
assumes "(f ---> l) A"
and "l \<noteq> 0"
- shows "((inverse o f) ---> inverse l) A"
+ shows "((inverse \<circ> f) ---> inverse l) A"
unfolding o_def using assms by (rule tendsto_inverse)
lemma Lim_null:
@@ -1646,7 +1649,7 @@
lemma Lim_transform_bound:
fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"
and g :: "'a \<Rightarrow> 'c::real_normed_vector"
- assumes "eventually (\<lambda>n. norm(f n) <= norm(g n)) net"
+ assumes "eventually (\<lambda>n. norm (f n) \<le> norm (g n)) net"
and "(g ---> 0) net"
shows "(f ---> 0) net"
using assms(1) tendsto_norm_zero [OF assms(2)]
@@ -1657,7 +1660,7 @@
lemma Lim_in_closed_set:
assumes "closed S"
and "eventually (\<lambda>x. f(x) \<in> S) net"
- and "\<not>(trivial_limit net)" "(f ---> l) net"
+ and "\<not> trivial_limit net" "(f ---> l) net"
shows "l \<in> S"
proof (rule ccontr)
assume "l \<notin> S"
@@ -1676,8 +1679,8 @@
lemma Lim_dist_ubound:
assumes "\<not>(trivial_limit net)"
and "(f ---> l) net"
- and "eventually (\<lambda>x. dist a (f x) <= e) net"
- shows "dist a l <= e"
+ and "eventually (\<lambda>x. dist a (f x) \<le> e) net"
+ shows "dist a l \<le> e"
proof -
have "dist a l \<in> {..e}"
proof (rule Lim_in_closed_set)
@@ -1714,7 +1717,9 @@
lemma Lim_norm_lbound:
fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"
- assumes "\<not> (trivial_limit net)" "(f ---> l) net" "eventually (\<lambda>x. e <= norm(f x)) net"
+ assumes "\<not> trivial_limit net"
+ and "(f ---> l) net"
+ and "eventually (\<lambda>x. e \<le> norm (f x)) net"
shows "e \<le> norm l"
proof -
have "norm l \<in> {e..}"
@@ -1946,7 +1951,7 @@
lemma not_trivial_limit_within_ball:
- "(\<not> trivial_limit (at x within S)) = (\<forall>e>0. S \<inter> ball x e - {x} \<noteq> {})"
+ "\<not> trivial_limit (at x within S) \<longleftrightarrow> (\<forall>e>0. S \<inter> ball x e - {x} \<noteq> {})"
(is "?lhs = ?rhs")
proof -
{
@@ -1954,12 +1959,12 @@
{
fix e :: real
assume "e > 0"
- then obtain y where "y:(S-{x}) & dist y x < e"
+ then obtain y where "y \<in> S - {x}" and "dist y x < e"
using `?lhs` not_trivial_limit_within[of x S] closure_approachable[of x "S - {x}"]
by auto
- then have "y : (S Int ball x e - {x})"
+ then have "y \<in> S \<inter> ball x e - {x}"
unfolding ball_def by (simp add: dist_commute)
- then have "S Int ball x e - {x} ~= {}" by blast
+ then have "S \<inter> ball x e - {x} \<noteq> {}" by blast
}
then have "?rhs" by auto
}
@@ -1969,11 +1974,11 @@
{
fix e :: real
assume "e > 0"
- then obtain y where "y : (S Int ball x e - {x})"
+ then obtain y where "y \<in> S \<inter> ball x e - {x}"
using `?rhs` by blast
- then have "y:(S-{x}) & dist y x < e"
- unfolding ball_def by (simp add: dist_commute)
- then have "EX y:(S-{x}). dist y x < e"
+ then have "y \<in> S - {x}" and "dist y x < e"
+ unfolding ball_def by (simp_all add: dist_commute)
+ then have "\<exists>y \<in> S - {x}. dist y x < e"
by auto
}
then have "?lhs"
@@ -2004,16 +2009,18 @@
assumes "a \<in> A"
shows "infdist a A = 0"
proof -
- from infdist_le[OF assms, of "dist a a"] have "infdist a A \<le> 0" by auto
- with infdist_nonneg[of a A] assms show "infdist a A = 0" by auto
+ from infdist_le[OF assms, of "dist a a"] have "infdist a A \<le> 0"
+ by auto
+ with infdist_nonneg[of a A] assms show "infdist a A = 0"
+ by auto
qed
lemma infdist_triangle: "infdist x A \<le> infdist y A + dist x y"
-proof cases
- assume "A = {}"
+proof (cases "A = {}")
+ case True
then show ?thesis by (simp add: infdist_def)
next
- assume "A \<noteq> {}"
+ case False
then obtain a where "a \<in> A" by auto
have "infdist x A \<le> Inf {dist x y + dist y a |a. a \<in> A}"
proof (rule cInf_greatest)
@@ -2202,7 +2209,7 @@
ultimately show "?rhs" by auto
next
assume "?rhs"
- then have "e>0" by auto
+ then have "e > 0" by auto
{
fix d :: real
assume "d > 0"
@@ -2340,7 +2347,7 @@
lemma interior_cball:
fixes x :: "'a::{real_normed_vector, perfect_space}"
shows "interior (cball x e) = ball x e"
-proof (cases "e\<ge>0")
+proof (cases "e \<ge> 0")
case False note cs = this
from cs have "ball x e = {}"
using ball_empty[of e x] by auto
@@ -2409,7 +2416,9 @@
then have "\<forall>S \<subseteq> cball x e. open S \<longrightarrow> S \<subseteq> ball x e"
by auto
ultimately show ?thesis
- using interior_unique[of "ball x e" "cball x e"] using open_ball[of x e] by auto
+ using interior_unique[of "ball x e" "cball x e"]
+ using open_ball[of x e]
+ by auto
qed
lemma frontier_ball:
@@ -2422,13 +2431,13 @@
lemma frontier_cball:
fixes a :: "'a::{real_normed_vector, perfect_space}"
- shows "frontier(cball a e) = {x. dist a x = e}"
+ shows "frontier (cball a e) = {x. dist a x = e}"
apply (simp add: frontier_def interior_cball closed_cball order_less_imp_le)
apply (simp add: set_eq_iff)
apply arith
done
-lemma cball_eq_empty: "(cball x e = {}) \<longleftrightarrow> e < 0"
+lemma cball_eq_empty: "cball x e = {} \<longleftrightarrow> e < 0"
apply (simp add: set_eq_iff not_le)
apply (metis zero_le_dist dist_self order_less_le_trans)
done
@@ -2438,7 +2447,7 @@
lemma cball_eq_sing:
fixes x :: "'a::{metric_space,perfect_space}"
- shows "(cball x e = {x}) \<longleftrightarrow> e = 0"
+ shows "cball x e = {x} \<longleftrightarrow> e = 0"
proof (rule linorder_cases)
assume e: "0 < e"
obtain a where "a \<noteq> x" "dist a x < e"
@@ -2466,7 +2475,8 @@
lemma bounded_any_center: "bounded S \<longleftrightarrow> (\<exists>e. \<forall>y\<in>S. dist a y \<le> e)"
unfolding bounded_def
apply safe
- apply (rule_tac x="dist a x + e" in exI, clarify)
+ apply (rule_tac x="dist a x + e" in exI)
+ apply clarify
apply (drule (1) bspec)
apply (erule order_trans [OF dist_triangle add_left_mono])
apply auto
@@ -2526,7 +2536,7 @@
apply auto
done
-lemma bounded_ball[simp,intro]: "bounded(ball x e)"
+lemma bounded_ball[simp,intro]: "bounded (ball x e)"
by (metis ball_subset_cball bounded_cball bounded_subset)
lemma bounded_Un[simp]: "bounded (S \<union> T) \<longleftrightarrow> bounded S \<and> bounded T"
@@ -2534,14 +2544,16 @@
apply (rename_tac x y r s)
apply (rule_tac x=x in exI)
apply (rule_tac x="max r (dist x y + s)" in exI)
- apply (rule ballI, rename_tac z, safe)
- apply (drule (1) bspec, simp)
+ apply (rule ballI)
+ apply safe
+ apply (drule (1) bspec)
+ apply simp
apply (drule (1) bspec)
apply (rule min_max.le_supI2)
apply (erule order_trans [OF dist_triangle add_left_mono])
done
-lemma bounded_Union[intro]: "finite F \<Longrightarrow> (\<forall>S\<in>F. bounded S) \<Longrightarrow> bounded(\<Union>F)"
+lemma bounded_Union[intro]: "finite F \<Longrightarrow> \<forall>S\<in>F. bounded S \<Longrightarrow> bounded (\<Union>F)"
by (induct rule: finite_induct[of F]) auto
lemma bounded_UN [intro]: "finite A \<Longrightarrow> \<forall>x\<in>A. bounded (B x) \<Longrightarrow> bounded (\<Union>x\<in>A. B x)"
@@ -2549,22 +2561,27 @@
lemma bounded_insert [simp]: "bounded (insert x S) \<longleftrightarrow> bounded S"
proof -
- have "\<forall>y\<in>{x}. dist x y \<le> 0" by simp
- then have "bounded {x}" unfolding bounded_def by fast
- then show ?thesis by (metis insert_is_Un bounded_Un)
+ have "\<forall>y\<in>{x}. dist x y \<le> 0"
+ by simp
+ then have "bounded {x}"
+ unfolding bounded_def by fast
+ then show ?thesis
+ by (metis insert_is_Un bounded_Un)
qed
lemma finite_imp_bounded [intro]: "finite S \<Longrightarrow> bounded S"
by (induct set: finite) simp_all
-lemma bounded_pos: "bounded S \<longleftrightarrow> (\<exists>b>0. \<forall>x\<in> S. norm x <= b)"
+lemma bounded_pos: "bounded S \<longleftrightarrow> (\<exists>b>0. \<forall>x\<in> S. norm x \<le> b)"
apply (simp add: bounded_iff)
- apply (subgoal_tac "\<And>x (y::real). 0 < 1 + abs y \<and> (x <= y \<longrightarrow> x <= 1 + abs y)")
+ apply (subgoal_tac "\<And>x (y::real). 0 < 1 + abs y \<and> (x \<le> y \<longrightarrow> x \<le> 1 + abs y)")
apply metis
apply arith
done
-lemma Bseq_eq_bounded: "Bseq f \<longleftrightarrow> bounded (range f::_::real_normed_vector set)"
+lemma Bseq_eq_bounded:
+ fixes f :: "nat \<Rightarrow> 'a::real_normed_vector"
+ shows "Bseq f \<longleftrightarrow> bounded (range f)"
unfolding Bseq_def bounded_pos by auto
lemma bounded_Int[intro]: "bounded S \<or> bounded T \<Longrightarrow> bounded (S \<inter> T)"
@@ -2575,11 +2592,13 @@
lemma not_bounded_UNIV[simp, intro]:
"\<not> bounded (UNIV :: 'a::{real_normed_vector, perfect_space} set)"
-proof(auto simp add: bounded_pos not_le)
+proof (auto simp add: bounded_pos not_le)
obtain x :: 'a where "x \<noteq> 0"
using perfect_choose_dist [OF zero_less_one] by fast
- fix b::real assume b: "b >0"
- have b1: "b +1 \<ge> 0" using b by simp
+ fix b :: real
+ assume b: "b >0"
+ have b1: "b +1 \<ge> 0"
+ using b by simp
with `x \<noteq> 0` have "b < norm (scaleR (b + 1) (sgn x))"
by (simp add: norm_sgn)
then show "\<exists>x::'a. b < norm x" ..
@@ -2590,15 +2609,17 @@
and "bounded_linear f"
shows "bounded (f ` S)"
proof -
- from assms(1) obtain b where b:"b>0" "\<forall>x\<in>S. norm x \<le> b"
+ from assms(1) obtain b where b: "b > 0" "\<forall>x\<in>S. norm x \<le> b"
unfolding bounded_pos by auto
- from assms(2) obtain B where B:"B>0" "\<forall>x. norm (f x) \<le> B * norm x"
+ from assms(2) obtain B where B: "B > 0" "\<forall>x. norm (f x) \<le> B * norm x"
using bounded_linear.pos_bounded by (auto simp add: mult_ac)
{
fix x
- assume "x\<in>S"
- then have "norm x \<le> b" using b by auto
- then have "norm (f x) \<le> B * b" using B(2)
+ assume "x \<in> S"
+ then have "norm x \<le> b"
+ using b by auto
+ then have "norm (f x) \<le> B * b"
+ using B(2)
apply (erule_tac x=x in allE)
apply (metis B(1) B(2) order_trans mult_le_cancel_left_pos)
done
@@ -2624,11 +2645,11 @@
assumes "bounded S"
shows "bounded ((\<lambda>x. a + x) ` S)"
proof -
- from assms obtain b where b:"b>0" "\<forall>x\<in>S. norm x \<le> b"
+ from assms obtain b where b: "b > 0" "\<forall>x\<in>S. norm x \<le> b"
unfolding bounded_pos by auto
{
fix x
- assume "x\<in>S"
+ assume "x \<in> S"
then have "norm (a + x) \<le> b + norm a"
using norm_triangle_ineq[of a x] b by auto
}
@@ -2648,7 +2669,8 @@
lemma bounded_has_Sup:
fixes S :: "real set"
- assumes "bounded S" "S \<noteq> {}"
+ assumes "bounded S"
+ and "S \<noteq> {}"
shows "\<forall>x\<in>S. x \<le> Sup S"
and "\<forall>b. (\<forall>x\<in>S. x \<le> b) \<longrightarrow> Sup S \<le> b"
proof
@@ -2679,18 +2701,19 @@
lemma bounded_has_Inf:
fixes S :: "real set"
- assumes "bounded S" "S \<noteq> {}"
+ assumes "bounded S"
+ and "S \<noteq> {}"
shows "\<forall>x\<in>S. x \<ge> Inf S"
and "\<forall>b. (\<forall>x\<in>S. x \<ge> b) \<longrightarrow> Inf S \<ge> b"
proof
fix x
- assume "x\<in>S"
+ assume "x \<in> S"
from assms(1) obtain a where a: "\<forall>x\<in>S. \<bar>x\<bar> \<le> a"
unfolding bounded_real by auto
- then show "x \<ge> Inf S" using `x\<in>S`
+ then show "x \<ge> Inf S" using `x \<in> S`
by (metis cInf_lower_EX abs_le_D2 minus_le_iff)
next
- show "\<forall>b. (\<forall>x\<in>S. x >= b) \<longrightarrow> Inf S \<ge> b"
+ show "\<forall>b. (\<forall>x\<in>S. x \<ge> b) \<longrightarrow> Inf S \<ge> b"
using assms by (metis cInf_greatest)
qed
@@ -2715,25 +2738,29 @@
subsubsection {* Bolzano-Weierstrass property *}
lemma heine_borel_imp_bolzano_weierstrass:
- assumes "compact s" and "infinite t" and "t \<subseteq> s"
+ assumes "compact s"
+ and "infinite t"
+ and "t \<subseteq> s"
shows "\<exists>x \<in> s. x islimpt t"
proof (rule ccontr)
assume "\<not> (\<exists>x \<in> s. x islimpt t)"
- then obtain f where f:"\<forall>x\<in>s. x \<in> f x \<and> open (f x) \<and> (\<forall>y\<in>t. y \<in> f x \<longrightarrow> y = x)"
+ then obtain f where f: "\<forall>x\<in>s. x \<in> f x \<and> open (f x) \<and> (\<forall>y\<in>t. y \<in> f x \<longrightarrow> y = x)"
unfolding islimpt_def
using bchoice[of s "\<lambda> x T. x \<in> T \<and> open T \<and> (\<forall>y\<in>t. y \<in> T \<longrightarrow> y = x)"]
by auto
- obtain g where g: "g\<subseteq>{t. \<exists>x. x \<in> s \<and> t = f x}" "finite g" "s \<subseteq> \<Union>g"
+ obtain g where g: "g \<subseteq> {t. \<exists>x. x \<in> s \<and> t = f x}" "finite g" "s \<subseteq> \<Union>g"
using assms(1)[unfolded compact_eq_heine_borel, THEN spec[where x="{t. \<exists>x. x\<in>s \<and> t = f x}"]]
using f by auto
- from g(1,3) have g':"\<forall>x\<in>g. \<exists>xa \<in> s. x = f xa" by auto
+ from g(1,3) have g':"\<forall>x\<in>g. \<exists>xa \<in> s. x = f xa"
+ by auto
{
fix x y
- assume "x\<in>t" "y\<in>t" "f x = f y"
+ assume "x \<in> t" "y \<in> t" "f x = f y"
then have "x \<in> f x" "y \<in> f x \<longrightarrow> y = x"
- using f[THEN bspec[where x=x]] and `t\<subseteq>s` by auto
+ using f[THEN bspec[where x=x]] and `t \<subseteq> s` by auto
then have "x = y"
- using `f x = f y` and f[THEN bspec[where x=y]] and `y\<in>t` and `t\<subseteq>s` by auto
+ using `f x = f y` and f[THEN bspec[where x=y]] and `y \<in> t` and `t \<subseteq> s`
+ by auto
}
then have "inj_on f t"
unfolding inj_on_def by simp
@@ -2742,14 +2769,17 @@
moreover
{
fix x
- assume "x\<in>t" "f x \<notin> g"
- from g(3) assms(3) `x\<in>t` obtain h where "h\<in>g" and "x\<in>h" by auto
- then obtain y where "y\<in>s" "h = f y"
+ assume "x \<in> t" "f x \<notin> g"
+ from g(3) assms(3) `x \<in> t` obtain h where "h \<in> g" and "x \<in> h"
+ by auto
+ then obtain y where "y \<in> s" "h = f y"
using g'[THEN bspec[where x=h]] by auto
then have "y = x"
- using f[THEN bspec[where x=y]] and `x\<in>t` and `x\<in>h`[unfolded `h = f y`] by auto
+ using f[THEN bspec[where x=y]] and `x\<in>t` and `x\<in>h`[unfolded `h = f y`]
+ by auto
then have False
- using `f x \<notin> g` `h\<in>g` unfolding `h = f y` by auto
+ using `f x \<notin> g` `h \<in> g` unfolding `h = f y`
+ by auto
}
then have "f ` t \<subseteq> g" by auto
ultimately show False
@@ -2786,7 +2816,8 @@
proof (rule topological_tendstoI)
fix S
assume "open S" "l \<in> S"
- with A(3) have "eventually (\<lambda>i. A i \<subseteq> S) sequentially" by auto
+ with A(3) have "eventually (\<lambda>i. A i \<subseteq> S) sequentially"
+ by auto
moreover
{
fix i
@@ -2810,12 +2841,18 @@
shows "infinite (range f)"
proof
assume "finite (range f)"
- then have "closed (range f)" by (rule finite_imp_closed)
- then have "open (- range f)" by (rule open_Compl)
- from assms(1) have "l \<in> - range f" by auto
+ then have "closed (range f)"
+ by (rule finite_imp_closed)
+ then have "open (- range f)"
+ by (rule open_Compl)
+ from assms(1) have "l \<in> - range f"
+ by auto
from assms(2) have "eventually (\<lambda>n. f n \<in> - range f) sequentially"
- using `open (- range f)` `l \<in> - range f` by (rule topological_tendstoD)
- then show False unfolding eventually_sequentially by auto
+ using `open (- range f)` `l \<in> - range f`
+ by (rule topological_tendstoD)
+ then show False
+ unfolding eventually_sequentially
+ by auto
qed
lemma closure_insert:
@@ -2928,7 +2965,7 @@
qed
lemma bolzano_weierstrass_imp_closed:
- fixes s :: "'a::{first_countable_topology, t2_space} set"
+ fixes s :: "'a::{first_countable_topology,t2_space} set"
assumes "\<forall>t. infinite t \<and> t \<subseteq> s --> (\<exists>x \<in> s. x islimpt t)"
shows "closed s"
proof -
@@ -3276,7 +3313,7 @@
definition seq_compact :: "'a::topological_space set \<Rightarrow> bool"
where "seq_compact S \<longleftrightarrow>
- (\<forall>f. (\<forall>n. f n \<in> S) \<longrightarrow> (\<exists>l\<in>S. \<exists>r. subseq r \<and> ((f o r) ---> l) sequentially))"
+ (\<forall>f. (\<forall>n. f n \<in> S) \<longrightarrow> (\<exists>l\<in>S. \<exists>r. subseq r \<and> ((f \<circ> r) ---> l) sequentially))"
lemma seq_compact_imp_countably_compact:
fixes U :: "'a :: first_countable_topology set"
@@ -3391,7 +3428,7 @@
qed
lemma seq_compactI:
- assumes "\<And>f. \<forall>n. f n \<in> S \<Longrightarrow> \<exists>l\<in>S. \<exists>r. subseq r \<and> ((f o r) ---> l) sequentially"
+ assumes "\<And>f. \<forall>n. f n \<in> S \<Longrightarrow> \<exists>l\<in>S. \<exists>r. subseq r \<and> ((f \<circ> r) ---> l) sequentially"
shows "seq_compact S"
unfolding seq_compact_def using assms by fast
@@ -3611,7 +3648,7 @@
lemma compact_def:
"compact (S :: 'a::metric_space set) \<longleftrightarrow>
- (\<forall>f. (\<forall>n. f n \<in> S) \<longrightarrow> (\<exists>l\<in>S. \<exists>r. subseq r \<and> (f o r) ----> l))"
+ (\<forall>f. (\<forall>n. f n \<in> S) \<longrightarrow> (\<exists>l\<in>S. \<exists>r. subseq r \<and> (f \<circ> r) ----> l))"
unfolding compact_eq_seq_compact_metric seq_compact_def by auto
subsubsection {* Complete the chain of compactness variants *}
@@ -4514,7 +4551,7 @@
fixes f :: "'a::metric_space \<Rightarrow> 'b::topological_space"
shows "continuous (at a within s) f \<longleftrightarrow>
(\<forall>x. (\<forall>n::nat. x n \<in> s) \<and> (x ---> a) sequentially
- \<longrightarrow> ((f o x) ---> f a) sequentially)"
+ \<longrightarrow> ((f \<circ> x) ---> f a) sequentially)"
(is "?lhs = ?rhs")
proof
assume ?lhs
@@ -4546,14 +4583,14 @@
lemma continuous_at_sequentially:
fixes f :: "'a::metric_space \<Rightarrow> 'b::topological_space"
shows "continuous (at a) f \<longleftrightarrow>
- (\<forall>x. (x ---> a) sequentially --> ((f o x) ---> f a) sequentially)"
+ (\<forall>x. (x ---> a) sequentially --> ((f \<circ> x) ---> f a) sequentially)"
using continuous_within_sequentially[of a UNIV f] by simp
lemma continuous_on_sequentially:
fixes f :: "'a::metric_space \<Rightarrow> 'b::topological_space"
shows "continuous_on s f \<longleftrightarrow>
(\<forall>x. \<forall>a \<in> s. (\<forall>n. x(n) \<in> s) \<and> (x ---> a) sequentially
- --> ((f o x) ---> f(a)) sequentially)"
+ --> ((f \<circ> x) ---> f a) sequentially)"
(is "?lhs = ?rhs")
proof
assume ?rhs
@@ -4804,8 +4841,8 @@
lemma uniformly_continuous_on_compose[continuous_on_intros]:
assumes "uniformly_continuous_on s f" "uniformly_continuous_on (f ` s) g"
- shows "uniformly_continuous_on s (g o f)"
-proof-
+ shows "uniformly_continuous_on s (g \<circ> f)"
+proof -
{
fix e :: real
assume "e > 0"
@@ -6818,7 +6855,7 @@
lemma Lim_component_eq:
fixes f :: "'a \<Rightarrow> 'b::euclidean_space"
- assumes net: "(f ---> l) net" "~(trivial_limit net)"
+ assumes net: "(f ---> l) net" "\<not> trivial_limit net"
and ev:"eventually (\<lambda>x. f(x)\<bullet>i = b) net"
shows "l\<bullet>i = b"
using ev[unfolded order_eq_iff eventually_conj_iff]
@@ -6887,8 +6924,7 @@
(\<forall>x\<in>s. (g(f x) = x)) \<and> (f ` s = t) \<and> continuous_on s f \<and>
(\<forall>y\<in>t. (f(g y) = y)) \<and> (g ` t = s) \<and> continuous_on t g"
-definition
- homeomorphic :: "'a::topological_space set \<Rightarrow> 'b::topological_space set \<Rightarrow> bool"
+definition homeomorphic :: "'a::topological_space set \<Rightarrow> 'b::topological_space set \<Rightarrow> bool"
(infixr "homeomorphic" 60)
where "s homeomorphic t \<equiv> (\<exists>f g. homeomorphism s t f g)"
@@ -7099,12 +7135,12 @@
lemma cauchy_isometric:
fixes x :: "nat \<Rightarrow> 'a::euclidean_space"
- assumes e: "0 < e"
+ assumes e: "e > 0"
and s: "subspace s"
and f: "bounded_linear f"
- and normf: "\<forall>x\<in>s. norm(f x) \<ge> e * norm(x)"
- and xs: "\<forall>n::nat. x n \<in> s"
- and cf: "Cauchy(f o x)"
+ and normf: "\<forall>x\<in>s. norm (f x) \<ge> e * norm x"
+ and xs: "\<forall>n. x n \<in> s"
+ and cf: "Cauchy (f \<circ> x)"
shows "Cauchy x"
proof -
interpret f: bounded_linear f by fact
@@ -7145,24 +7181,31 @@
fix g
assume as:"\<forall>n::nat. g n \<in> f ` s" and cfg:"Cauchy g"
then obtain x where "\<forall>n. x n \<in> s \<and> g n = f (x n)"
- using choice[of "\<lambda> n xa. xa \<in> s \<and> g n = f xa"] by auto
- then have x:"\<forall>n. x n \<in> s" "\<forall>n. g n = f (x n)" by auto
- then have "f \<circ> x = g" unfolding fun_eq_iff by auto
+ using choice[of "\<lambda> n xa. xa \<in> s \<and> g n = f xa"]
+ by auto
+ then have x:"\<forall>n. x n \<in> s" "\<forall>n. g n = f (x n)"
+ by auto
+ then have "f \<circ> x = g"
+ unfolding fun_eq_iff
+ by auto
then obtain l where "l\<in>s" and l:"(x ---> l) sequentially"
using cs[unfolded complete_def, THEN spec[where x="x"]]
- using cauchy_isometric[OF `0<e` s f normf] and cfg and x(1) by auto
+ using cauchy_isometric[OF `0<e` s f normf] and cfg and x(1)
+ by auto
then have "\<exists>l\<in>f ` s. (g ---> l) sequentially"
using linear_continuous_at[OF f, unfolded continuous_at_sequentially, THEN spec[where x=x], of l]
- unfolding `f \<circ> x = g` by auto
+ unfolding `f \<circ> x = g`
+ by auto
}
- then show ?thesis unfolding complete_def by auto
+ then show ?thesis
+ unfolding complete_def by auto
qed
lemma injective_imp_isometric:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes s: "closed s" "subspace s"
- and f: "bounded_linear f" "\<forall>x\<in>s. (f x = 0) \<longrightarrow> (x = 0)"
- shows "\<exists>e>0. \<forall>x\<in>s. norm (f x) \<ge> e * norm(x)"
+ and f: "bounded_linear f" "\<forall>x\<in>s. f x = 0 \<longrightarrow> x = 0"
+ shows "\<exists>e>0. \<forall>x\<in>s. norm (f x) \<ge> e * norm x"
proof (cases "s \<subseteq> {0::'a}")
case True
{
@@ -7175,8 +7218,10 @@
next
interpret f: bounded_linear f by fact
case False
- then obtain a where a:"a\<noteq>0" "a\<in>s" by auto
- from False have "s \<noteq> {}" by auto
+ then obtain a where a: "a \<noteq> 0" "a \<in> s"
+ by auto
+ from False have "s \<noteq> {}"
+ by auto
let ?S = "{f x| x. (x \<in> s \<and> norm x = norm a)}"
let ?S' = "{x::'a. x\<in>s \<and> norm x = norm a}"
let ?S'' = "{x::'a. norm x = norm a}"