distinguished session for multivariate analysis
authorhimmelma
Fri, 23 Oct 2009 13:23:18 +0200
changeset 33175 2083bde13ce1
parent 33083 1fad3160d873
child 33176 d6936fd7cda8
distinguished session for multivariate analysis
src/HOL/IsaMakefile
src/HOL/Library/#Topology_Euclidean_Space.thy#
src/HOL/Library/Convex_Euclidean_Space.thy
src/HOL/Library/Determinants.thy
src/HOL/Library/Euclidean_Space.thy
src/HOL/Library/Finite_Cartesian_Product.thy
src/HOL/Library/Library.thy
src/HOL/Library/Topology_Euclidean_Space.thy
src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy
src/HOL/Multivariate_Analysis/Determinants.thy
src/HOL/Multivariate_Analysis/Euclidean_Space.thy
src/HOL/Multivariate_Analysis/Finite_Cartesian_Product.thy
src/HOL/Multivariate_Analysis/Multivariate_Analysis.thy
src/HOL/Multivariate_Analysis/ROOT.ML
src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy
--- a/src/HOL/IsaMakefile	Fri Oct 23 14:33:07 2009 +0200
+++ b/src/HOL/IsaMakefile	Fri Oct 23 13:23:18 2009 +0200
@@ -323,15 +323,14 @@
 
 $(LOG)/HOL-Library.gz: $(OUT)/HOL Library/SetsAndFunctions.thy		\
   Library/Abstract_Rat.thy Library/BigO.thy Library/ContNotDenum.thy	\
-  Library/Efficient_Nat.thy Library/Euclidean_Space.thy			\
+  Library/Efficient_Nat.thy 			 			\
   Library/Sum_Of_Squares.thy Library/Sum_Of_Squares/sos_wrapper.ML	\
   Library/Sum_Of_Squares/sum_of_squares.ML Library/Fset.thy		\
-  Library/Convex_Euclidean_Space.thy Library/Glbs.thy			\
+  Library/Glbs.thy							\
   Library/normarith.ML Library/Executable_Set.thy			\
   Library/Infinite_Set.thy Library/FuncSet.thy				\
-  Library/Permutations.thy Library/Determinants.thy Library/Bit.thy	\
-  Library/Topology_Euclidean_Space.thy					\
-  Library/Finite_Cartesian_Product.thy Library/FrechetDeriv.thy		\
+  Library/Permutations.thy Library/Bit.thy				\
+  Library/FrechetDeriv.thy		\
   Library/Fraction_Field.thy Library/Fundamental_Theorem_Algebra.thy	\
   Library/Inner_Product.thy Library/Kleene_Algebra.thy			\
   Library/Lattice_Syntax.thy			\
@@ -1005,6 +1004,19 @@
 	@cd TLA; $(ISABELLE_TOOL) usedir $(OUT)/TLA Memory
 
 
+## HOL-Multivariate_Analysis
+
+HOL-Multivariate_Analysis: HOL $(OUT)/HOL-Multivariate_Analysis
+
+$(OUT)/HOL-Multivariate_Analysis: $(OUT)/HOL Multivariate_Analysis/ROOT.ML \
+  Multivariate_Analysis/Multivariate_Analysis.thy \
+  Multivariate_Analysis/Determinants.thy \
+  Multivariate_Analysis/Finite_Cartesian_Product.thy \
+  Multivariate_Analysis/Euclidean_Space.thy \
+  Multivariate_Analysis/Topology_Euclidean_Space.thy \
+  Multivariate_Analysis/Convex_Euclidean_Space.thy
+	@cd Multivariate_Analysis; $(ISABELLE_TOOL) usedir -b -g true $(OUT)/HOL HOL-Multivariate_Analysis
+
 ## HOL-Nominal
 
 HOL-Nominal: HOL $(OUT)/HOL-Nominal
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Library/#Topology_Euclidean_Space.thy#	Fri Oct 23 13:23:18 2009 +0200
@@ -0,0 +1,6029 @@
+(* Title:      Topology
+   Author:     Amine Chaieb, University of Cambridge
+   Author:     Robert Himmelmann, TU Muenchen
+*)
+
+header {* Elementary topology in Euclidean space. *}
+
+theory Topology_Euclidean_Space
+imports SEQ Euclidean_Space Product_Vector
+begin
+
+declare fstcart_pastecart[simp] sndcart_pastecart[simp]
+
+subsection{* General notion of a topology *}
+
+definition "istopology L \<longleftrightarrow> {} \<in> L \<and> (\<forall>S \<in>L. \<forall>T \<in>L. S \<inter> T \<in> L) \<and> (\<forall>K. K \<subseteq>L \<longrightarrow> \<Union> K \<in> L)"
+typedef (open) 'a topology = "{L::('a set) set. istopology L}"
+  morphisms "openin" "topology"
+  unfolding istopology_def by blast
+
+lemma istopology_open_in[intro]: "istopology(openin U)"
+  using openin[of U] by blast
+
+lemma topology_inverse': "istopology U \<Longrightarrow> openin (topology U) = U"
+  using topology_inverse[unfolded mem_def Collect_def] .
+
+lemma topology_inverse_iff: "istopology U \<longleftrightarrow> openin (topology U) = U"
+  using topology_inverse[of U] istopology_open_in[of "topology U"] by auto
+
+lemma topology_eq: "T1 = T2 \<longleftrightarrow> (\<forall>S. openin T1 S \<longleftrightarrow> openin T2 S)"
+proof-
+  {assume "T1=T2" hence "\<forall>S. openin T1 S \<longleftrightarrow> openin T2 S" by simp}
+  moreover
+  {assume H: "\<forall>S. openin T1 S \<longleftrightarrow> openin T2 S"
+    hence "openin T1 = openin T2" by (metis mem_def set_ext)
+    hence "topology (openin T1) = topology (openin T2)" by simp
+    hence "T1 = T2" unfolding openin_inverse .}
+  ultimately show ?thesis by blast
+qed
+
+text{* Infer the "universe" from union of all sets in the topology. *}
+
+definition "topspace T =  \<Union>{S. openin T S}"
+
+subsection{* Main properties of open sets *}
+
+lemma openin_clauses:
+  fixes U :: "'a topology"
+  shows "openin U {}"
+  "\<And>S T. openin U S \<Longrightarrow> openin U T \<Longrightarrow> openin U (S\<inter>T)"
+  "\<And>K. (\<forall>S \<in> K. openin U S) \<Longrightarrow> openin U (\<Union>K)"
+  using openin[of U] unfolding istopology_def Collect_def mem_def
+  by (metis mem_def subset_eq)+
+
+lemma openin_subset[intro]: "openin U S \<Longrightarrow> S \<subseteq> topspace U"
+  unfolding topspace_def by blast
+lemma openin_empty[simp]: "openin U {}" by (simp add: openin_clauses)
+
+lemma openin_Int[intro]: "openin U S \<Longrightarrow> openin U T \<Longrightarrow> openin U (S \<inter> T)"
+  by (simp add: openin_clauses)
+
+lemma openin_Union[intro]: "(\<forall>S \<in>K. openin U S) \<Longrightarrow> openin U (\<Union> K)" by (simp add: openin_clauses)
+
+lemma openin_Un[intro]: "openin U S \<Longrightarrow> openin U T \<Longrightarrow> openin U (S \<union> T)"
+  using openin_Union[of "{S,T}" U] by auto
+
+lemma openin_topspace[intro, simp]: "openin U (topspace U)" by (simp add: openin_Union topspace_def)
+
+lemma openin_subopen: "openin U S \<longleftrightarrow> (\<forall>x \<in> S. \<exists>T. openin U T \<and> x \<in> T \<and> T \<subseteq> S)" (is "?lhs \<longleftrightarrow> ?rhs")
+proof-
+  {assume ?lhs then have ?rhs by auto }
+  moreover
+  {assume H: ?rhs
+    then obtain t where t: "\<forall>x\<in>S. openin U (t x) \<and> x \<in> t x \<and> t x \<subseteq> S"
+      unfolding Ball_def ex_simps(6)[symmetric] choice_iff by blast
+    from t have th0: "\<forall>x\<in> t`S. openin U x" by auto
+    have "\<Union> t`S = S" using t by auto
+    with openin_Union[OF th0] have "openin U S" by simp }
+  ultimately show ?thesis by blast
+qed
+
+subsection{* Closed sets *}
+
+definition "closedin U S \<longleftrightarrow> S \<subseteq> topspace U \<and> openin U (topspace U - S)"
+
+lemma closedin_subset: "closedin U S \<Longrightarrow> S \<subseteq> topspace U" by (metis closedin_def)
+lemma closedin_empty[simp]: "closedin U {}" by (simp add: closedin_def)
+lemma closedin_topspace[intro,simp]:
+  "closedin U (topspace U)" by (simp add: closedin_def)
+lemma closedin_Un[intro]: "closedin U S \<Longrightarrow> closedin U T \<Longrightarrow> closedin U (S \<union> T)"
+  by (auto simp add: Diff_Un closedin_def)
+
+lemma Diff_Inter[intro]: "A - \<Inter>S = \<Union> {A - s|s. s\<in>S}" by auto
+lemma closedin_Inter[intro]: assumes Ke: "K \<noteq> {}" and Kc: "\<forall>S \<in>K. closedin U S"
+  shows "closedin U (\<Inter> K)"  using Ke Kc unfolding closedin_def Diff_Inter by auto
+
+lemma closedin_Int[intro]: "closedin U S \<Longrightarrow> closedin U T \<Longrightarrow> closedin U (S \<inter> T)"
+  using closedin_Inter[of "{S,T}" U] by auto
+
+lemma Diff_Diff_Int: "A - (A - B) = A \<inter> B" by blast
+lemma openin_closedin_eq: "openin U S \<longleftrightarrow> S \<subseteq> topspace U \<and> closedin U (topspace U - S)"
+  apply (auto simp add: closedin_def)
+  apply (metis openin_subset subset_eq)
+  apply (auto simp add: Diff_Diff_Int)
+  apply (subgoal_tac "topspace U \<inter> S = S")
+  by auto
+
+lemma openin_closedin:  "S \<subseteq> topspace U \<Longrightarrow> (openin U S \<longleftrightarrow> closedin U (topspace U - S))"
+  by (simp add: openin_closedin_eq)
+
+lemma openin_diff[intro]: assumes oS: "openin U S" and cT: "closedin U T" shows "openin U (S - T)"
+proof-
+  have "S - T = S \<inter> (topspace U - T)" using openin_subset[of U S]  oS cT
+    by (auto simp add: topspace_def openin_subset)
+  then show ?thesis using oS cT by (auto simp add: closedin_def)
+qed
+
+lemma closedin_diff[intro]: assumes oS: "closedin U S" and cT: "openin U T" shows "closedin U (S - T)"
+proof-
+  have "S - T = S \<inter> (topspace U - T)" using closedin_subset[of U S]  oS cT
+    by (auto simp add: topspace_def )
+  then show ?thesis using oS cT by (auto simp add: openin_closedin_eq)
+qed
+
+subsection{* Subspace topology. *}
+
+definition "subtopology U V = topology {S \<inter> V |S. openin U S}"
+
+lemma istopology_subtopology: "istopology {S \<inter> V |S. openin U S}" (is "istopology ?L")
+proof-
+  have "{} \<in> ?L" by blast
+  {fix A B assume A: "A \<in> ?L" and B: "B \<in> ?L"
+    from A B obtain Sa and Sb where Sa: "openin U Sa" "A = Sa \<inter> V" and Sb: "openin U Sb" "B = Sb \<inter> V" by blast
+    have "A\<inter>B = (Sa \<inter> Sb) \<inter> V" "openin U (Sa \<inter> Sb)"  using Sa Sb by blast+
+    then have "A \<inter> B \<in> ?L" by blast}
+  moreover
+  {fix K assume K: "K \<subseteq> ?L"
+    have th0: "?L = (\<lambda>S. S \<inter> V) ` openin U "
+      apply (rule set_ext)
+      apply (simp add: Ball_def image_iff)
+      by (metis mem_def)
+    from K[unfolded th0 subset_image_iff]
+    obtain Sk where Sk: "Sk \<subseteq> openin U" "K = (\<lambda>S. S \<inter> V) ` Sk" by blast
+    have "\<Union>K = (\<Union>Sk) \<inter> V" using Sk by auto
+    moreover have "openin U (\<Union> Sk)" using Sk by (auto simp add: subset_eq mem_def)
+    ultimately have "\<Union>K \<in> ?L" by blast}
+  ultimately show ?thesis unfolding istopology_def by blast
+qed
+
+lemma openin_subtopology:
+  "openin (subtopology U V) S \<longleftrightarrow> (\<exists> T. (openin U T) \<and> (S = T \<inter> V))"
+  unfolding subtopology_def topology_inverse'[OF istopology_subtopology]
+  by (auto simp add: Collect_def)
+
+lemma topspace_subtopology: "topspace(subtopology U V) = topspace U \<inter> V"
+  by (auto simp add: topspace_def openin_subtopology)
+
+lemma closedin_subtopology:
+  "closedin (subtopology U V) S \<longleftrightarrow> (\<exists>T. closedin U T \<and> S = T \<inter> V)"
+  unfolding closedin_def topspace_subtopology
+  apply (simp add: openin_subtopology)
+  apply (rule iffI)
+  apply clarify
+  apply (rule_tac x="topspace U - T" in exI)
+  by auto
+
+lemma openin_subtopology_refl: "openin (subtopology U V) V \<longleftrightarrow> V \<subseteq> topspace U"
+  unfolding openin_subtopology
+  apply (rule iffI, clarify)
+  apply (frule openin_subset[of U])  apply blast
+  apply (rule exI[where x="topspace U"])
+  by auto
+
+lemma subtopology_superset: assumes UV: "topspace U \<subseteq> V"
+  shows "subtopology U V = U"
+proof-
+  {fix S
+    {fix T assume T: "openin U T" "S = T \<inter> V"
+      from T openin_subset[OF T(1)] UV have eq: "S = T" by blast
+      have "openin U S" unfolding eq using T by blast}
+    moreover
+    {assume S: "openin U S"
+      hence "\<exists>T. openin U T \<and> S = T \<inter> V"
+	using openin_subset[OF S] UV by auto}
+    ultimately have "(\<exists>T. openin U T \<and> S = T \<inter> V) \<longleftrightarrow> openin U S" by blast}
+  then show ?thesis unfolding topology_eq openin_subtopology by blast
+qed
+
+
+lemma subtopology_topspace[simp]: "subtopology U (topspace U) = U"
+  by (simp add: subtopology_superset)
+
+lemma subtopology_UNIV[simp]: "subtopology U UNIV = U"
+  by (simp add: subtopology_superset)
+
+subsection{* The universal Euclidean versions are what we use most of the time *}
+
+definition
+  euclidean :: "'a::topological_space topology" where
+  "euclidean = topology open"
+
+lemma open_openin: "open S \<longleftrightarrow> openin euclidean S"
+  unfolding euclidean_def
+  apply (rule cong[where x=S and y=S])
+  apply (rule topology_inverse[symmetric])
+  apply (auto simp add: istopology_def)
+  by (auto simp add: mem_def subset_eq)
+
+lemma topspace_euclidean: "topspace euclidean = UNIV"
+  apply (simp add: topspace_def)
+  apply (rule set_ext)
+  by (auto simp add: open_openin[symmetric])
+
+lemma topspace_euclidean_subtopology[simp]: "topspace (subtopology euclidean S) = S"
+  by (simp add: topspace_euclidean topspace_subtopology)
+
+lemma closed_closedin: "closed S \<longleftrightarrow> closedin euclidean S"
+  by (simp add: closed_def closedin_def topspace_euclidean open_openin Compl_eq_Diff_UNIV)
+
+lemma open_subopen: "open S \<longleftrightarrow> (\<forall>x\<in>S. \<exists>T. open T \<and> x \<in> T \<and> T \<subseteq> S)"
+  by (simp add: open_openin openin_subopen[symmetric])
+
+subsection{* Open and closed balls. *}
+
+definition
+  ball :: "'a::metric_space \<Rightarrow> real \<Rightarrow> 'a set" where
+  "ball x e = {y. dist x y < e}"
+
+definition
+  cball :: "'a::metric_space \<Rightarrow> real \<Rightarrow> 'a set" where
+  "cball x e = {y. dist x y \<le> e}"
+
+lemma mem_ball[simp]: "y \<in> ball x e \<longleftrightarrow> dist x y < e" by (simp add: ball_def)
+lemma mem_cball[simp]: "y \<in> cball x e \<longleftrightarrow> dist x y \<le> e" by (simp add: cball_def)
+
+lemma mem_ball_0 [simp]:
+  fixes x :: "'a::real_normed_vecto"
+  shows "x \<in> ball 0 e \<longleftrightarrow> norm x < e"
+  by (simp add: dist_norm)
+
+lemma mem_cball_0 [simp]:
+  fixes x :: "'a::real_normed_vector"
+  shows "x \<in> cball 0 e \<longleftrightarrow> norm x \<le> e"
+  by (simp add: dist_norm)
+
+lemma centre_in_cball[simp]: "x \<in> cball x e \<longleftrightarrow> 0\<le> e"  by simp
+lemma ball_subset_cball[simp,intro]: "ball x e \<subseteq> cball x e" by (simp add: subset_eq)
+lemma subset_ball[intro]: "d <= e ==> ball x d \<subseteq> ball x e" by (simp add: subset_eq)
+lemma subset_cball[intro]: "d <= e ==> cball x d \<subseteq> cball x e" by (simp add: subset_eq)
+lemma ball_max_Un: "ball a (max r s) = ball a r \<union> ball a s"
+  by (simp add: expand_set_eq) arith
+
+lemma ball_min_Int: "ball a (min r s) = ball a r \<inter> ball a s"
+  by (simp add: expand_set_eq)
+
+subsection{* Topological properties of open balls *}
+
+lemma diff_less_iff: "(a::real) - b > 0 \<longleftrightarrow> a > b"
+  "(a::real) - b < 0 \<longleftrightarrow> a < b"
+  "a - b < c \<longleftrightarrow> a < c +b" "a - b > c \<longleftrightarrow> a > c +b" by arith+
+lemma diff_le_iff: "(a::real) - b \<ge> 0 \<longleftrightarrow> a \<ge> b" "(a::real) - b \<le> 0 \<longleftrightarrow> a \<le> b"
+  "a - b \<le> c \<longleftrightarrow> a \<le> c +b" "a - b \<ge> c \<longleftrightarrow> a \<ge> c +b"  by arith+
+
+lemma open_ball[intro, simp]: "open (ball x e)"
+  unfolding open_dist ball_def Collect_def Ball_def mem_def
+  unfolding dist_commute
+  apply clarify
+  apply (rule_tac x="e - dist xa x" in exI)
+  using dist_triangle_alt[where z=x]
+  apply (clarsimp simp add: diff_less_iff)
+  apply atomize
+  apply (erule_tac x="y" in allE)
+  apply (erule_tac x="xa" in allE)
+  by arith
+
+lemma centre_in_ball[simp]: "x \<in> ball x e \<longleftrightarrow> e > 0" by (metis mem_ball dist_self)
+lemma open_contains_ball: "open S \<longleftrightarrow> (\<forall>x\<in>S. \<exists>e>0. ball x e \<subseteq> S)"
+  unfolding open_dist subset_eq mem_ball Ball_def dist_commute ..
+
+lemma open_contains_ball_eq: "open S \<Longrightarrow> \<forall>x. x\<in>S \<longleftrightarrow> (\<exists>e>0. ball x e \<subseteq> S)"
+  by (metis open_contains_ball subset_eq centre_in_ball)
+
+lemma ball_eq_empty[simp]: "ball x e = {} \<longleftrightarrow> e \<le> 0"
+  unfolding mem_ball expand_set_eq
+  apply (simp add: not_less)
+  by (metis zero_le_dist order_trans dist_self)
+
+lemma ball_empty[intro]: "e \<le> 0 ==> ball x e = {}" by simp
+
+subsection{* Basic "localization" results are handy for connectedness. *}
+
+lemma openin_open: "openin (subtopology euclidean U) S \<longleftrightarrow> (\<exists>T. open T \<and> (S = U \<inter> T))"
+  by (auto simp add: openin_subtopology open_openin[symmetric])
+
+lemma openin_open_Int[intro]: "open S \<Longrightarrow> openin (subtopology euclidean U) (U \<inter> S)"
+  by (auto simp add: openin_open)
+
+lemma open_openin_trans[trans]:
+ "open S \<Longrightarrow> open T \<Longrightarrow> T \<subseteq> S \<Longrightarrow> openin (subtopology euclidean S) T"
+  by (metis Int_absorb1  openin_open_Int)
+
+lemma open_subset:  "S \<subseteq> T \<Longrightarrow> open S \<Longrightarrow> openin (subtopology euclidean T) S"
+  by (auto simp add: openin_open)
+
+lemma closedin_closed: "closedin (subtopology euclidean U) S \<longleftrightarrow> (\<exists>T. closed T \<and> S = U \<inter> T)"
+  by (simp add: closedin_subtopology closed_closedin Int_ac)
+
+lemma closedin_closed_Int: "closed S ==> closedin (subtopology euclidean U) (U \<inter> S)"
+  by (metis closedin_closed)
+
+lemma closed_closedin_trans: "closed S \<Longrightarrow> closed T \<Longrightarrow> T \<subseteq> S \<Longrightarrow> closedin (subtopology euclidean S) T"
+  apply (subgoal_tac "S \<inter> T = T" )
+  apply auto
+  apply (frule closedin_closed_Int[of T S])
+  by simp
+
+lemma closed_subset: "S \<subseteq> T \<Longrightarrow> closed S \<Longrightarrow> closedin (subtopology euclidean T) S"
+  by (auto simp add: closedin_closed)
+
+lemma openin_euclidean_subtopology_iff:
+  fixes S U :: "'a::metric_space set"
+  shows "openin (subtopology euclidean U) S
+  \<longleftrightarrow> S \<subseteq> U \<and> (\<forall>x\<in>S. \<exists>e>0. \<forall>x'\<in>U. dist x' x < e \<longrightarrow> x'\<in> S)" (is "?lhs \<longleftrightarrow> ?rhs")
+proof-
+  {assume ?lhs hence ?rhs unfolding openin_subtopology open_openin[symmetric]
+      by (simp add: open_dist) blast}
+  moreover
+  {assume SU: "S \<subseteq> U" and H: "\<And>x. x \<in> S \<Longrightarrow> \<exists>e>0. \<forall>x'\<in>U. dist x' x < e \<longrightarrow> x' \<in> S"
+    from H obtain d where d: "\<And>x . x\<in> S \<Longrightarrow> d x > 0 \<and> (\<forall>x' \<in> U. dist x' x < d x \<longrightarrow> x' \<in> S)"
+      by metis
+    let ?T = "\<Union>{B. \<exists>x\<in>S. B = ball x (d x)}"
+    have oT: "open ?T" by auto
+    { fix x assume "x\<in>S"
+      hence "x \<in> \<Union>{B. \<exists>x\<in>S. B = ball x (d x)}"
+	apply simp apply(rule_tac x="ball x(d x)" in exI) apply auto
+        by (rule d [THEN conjunct1])
+      hence "x\<in> ?T \<inter> U" using SU and `x\<in>S` by auto  }
+    moreover
+    { fix y assume "y\<in>?T"
+      then obtain B where "y\<in>B" "B\<in>{B. \<exists>x\<in>S. B = ball x (d x)}" by auto
+      then obtain x where "x\<in>S" and x:"y \<in> ball x (d x)" by auto
+      assume "y\<in>U"
+      hence "y\<in>S" using d[OF `x\<in>S`] and x by(auto simp add: dist_commute) }
+    ultimately have "S = ?T \<inter> U" by blast
+    with oT have ?lhs unfolding openin_subtopology open_openin[symmetric] by blast}
+  ultimately show ?thesis by blast
+qed
+
+text{* These "transitivity" results are handy too. *}
+
+lemma openin_trans[trans]: "openin (subtopology euclidean T) S \<Longrightarrow> openin (subtopology euclidean U) T
+  \<Longrightarrow> openin (subtopology euclidean U) S"
+  unfolding open_openin openin_open by blast
+
+lemma openin_open_trans: "openin (subtopology euclidean T) S \<Longrightarrow> open T \<Longrightarrow> open S"
+  by (auto simp add: openin_open intro: openin_trans)
+
+lemma closedin_trans[trans]:
+ "closedin (subtopology euclidean T) S \<Longrightarrow>
+           closedin (subtopology euclidean U) T
+           ==> closedin (subtopology euclidean U) S"
+  by (auto simp add: closedin_closed closed_closedin closed_Inter Int_assoc)
+
+lemma closedin_closed_trans: "closedin (subtopology euclidean T) S \<Longrightarrow> closed T \<Longrightarrow> closed S"
+  by (auto simp add: closedin_closed intro: closedin_trans)
+
+subsection{* Connectedness *}
+
+definition "connected S \<longleftrightarrow>
+  ~(\<exists>e1 e2. open e1 \<and> open e2 \<and> S \<subseteq> (e1 \<union> e2) \<and> (e1 \<inter> e2 \<inter> S = {})
+  \<and> ~(e1 \<inter> S = {}) \<and> ~(e2 \<inter> S = {}))"
+
+lemma connected_local:
+ "connected S \<longleftrightarrow> ~(\<exists>e1 e2.
+                 openin (subtopology euclidean S) e1 \<and>
+                 openin (subtopology euclidean S) e2 \<and>
+                 S \<subseteq> e1 \<union> e2 \<and>
+                 e1 \<inter> e2 = {} \<and>
+                 ~(e1 = {}) \<and>
+                 ~(e2 = {}))"
+unfolding connected_def openin_open by (safe, blast+)
+
+lemma exists_diff: "(\<exists>S. P(UNIV - S)) \<longleftrightarrow> (\<exists>S. P S)" (is "?lhs \<longleftrightarrow> ?rhs")
+proof-
+
+  {assume "?lhs" hence ?rhs by blast }
+  moreover
+  {fix S assume H: "P S"
+    have "S = UNIV - (UNIV - S)" by auto
+    with H have "P (UNIV - (UNIV - S))" by metis }
+  ultimately show ?thesis by metis
+qed
+
+lemma connected_clopen: "connected S \<longleftrightarrow>
+        (\<forall>T. openin (subtopology euclidean S) T \<and>
+            closedin (subtopology euclidean S) T \<longrightarrow> T = {} \<or> T = S)" (is "?lhs \<longleftrightarrow> ?rhs")
+proof-
+  have " \<not> connected S \<longleftrightarrow> (\<exists>e1 e2. open e1 \<and> open (UNIV - e2) \<and> S \<subseteq> e1 \<union> (UNIV - e2) \<and> e1 \<inter> (UNIV - e2) \<inter> S = {} \<and> e1 \<inter> S \<noteq> {} \<and> (UNIV - e2) \<inter> S \<noteq> {})"
+    unfolding connected_def openin_open closedin_closed
+    apply (subst exists_diff) by blast
+  hence th0: "connected S \<longleftrightarrow> \<not> (\<exists>e2 e1. closed e2 \<and> open e1 \<and> S \<subseteq> e1 \<union> (UNIV - e2) \<and> e1 \<inter> (UNIV - e2) \<inter> S = {} \<and> e1 \<inter> S \<noteq> {} \<and> (UNIV - e2) \<inter> S \<noteq> {})"
+    (is " _ \<longleftrightarrow> \<not> (\<exists>e2 e1. ?P e2 e1)") apply (simp add: closed_def Compl_eq_Diff_UNIV) by metis
+
+  have th1: "?rhs \<longleftrightarrow> \<not> (\<exists>t' t. closed t'\<and>t = S\<inter>t' \<and> t\<noteq>{} \<and> t\<noteq>S \<and> (\<exists>t'. open t' \<and> t = S \<inter> t'))"
+    (is "_ \<longleftrightarrow> \<not> (\<exists>t' t. ?Q t' t)")
+    unfolding connected_def openin_open closedin_closed by auto
+  {fix e2
+    {fix e1 have "?P e2 e1 \<longleftrightarrow> (\<exists>t.  closed e2 \<and> t = S\<inter>e2 \<and> open e1 \<and> t = S\<inter>e1 \<and> t\<noteq>{} \<and> t\<noteq>S)"
+	by auto}
+    then have "(\<exists>e1. ?P e2 e1) \<longleftrightarrow> (\<exists>t. ?Q e2 t)" by metis}
+  then have "\<forall>e2. (\<exists>e1. ?P e2 e1) \<longleftrightarrow> (\<exists>t. ?Q e2 t)" by blast
+  then show ?thesis unfolding th0 th1 by simp
+qed
+
+lemma connected_empty[simp, intro]: "connected {}"
+  by (simp add: connected_def)
+
+subsection{* Hausdorff and other separation properties *}
+
+class t0_space =
+  assumes t0_space: "x \<noteq> y \<Longrightarrow> \<exists>U. open U \<and> \<not> (x \<in> U \<longleftrightarrow> y \<in> U)"
+
+class t1_space =
+  assumes t1_space: "x \<noteq> y \<Longrightarrow> \<exists>U V. open U \<and> open V \<and> x \<in> U \<and> y \<notin> U \<and> x \<notin> V \<and> y \<in> V"
+begin
+
+subclass t0_space
+proof
+qed (fast dest: t1_space)
+
+end
+
+text {* T2 spaces are also known as Hausdorff spaces. *}
+
+class t2_space =
+  assumes hausdorff: "x \<noteq> y \<Longrightarrow> \<exists>U V. open U \<and> open V \<and> x \<in> U \<and> y \<in> V \<and> U \<inter> V = {}"
+begin
+
+subclass t1_space
+proof
+qed (fast dest: hausdorff)
+
+end
+
+instance metric_space \<subseteq> t2_space
+proof
+  fix x y :: "'a::metric_space"
+  assume xy: "x \<noteq> y"
+  let ?U = "ball x (dist x y / 2)"
+  let ?V = "ball y (dist x y / 2)"
+  have th0: "\<And>d x y z. (d x z :: real) <= d x y + d y z \<Longrightarrow> d y z = d z y
+               ==> ~(d x y * 2 < d x z \<and> d z y * 2 < d x z)" by arith
+  have "open ?U \<and> open ?V \<and> x \<in> ?U \<and> y \<in> ?V \<and> ?U \<inter> ?V = {}"
+    using dist_pos_lt[OF xy] th0[of dist,OF dist_triangle dist_commute]
+    by (auto simp add: expand_set_eq)
+  then show "\<exists>U V. open U \<and> open V \<and> x \<in> U \<and> y \<in> V \<and> U \<inter> V = {}"
+    by blast
+qed
+
+lemma separation_t2:
+  fixes x y :: "'a::t2_space"
+  shows "x \<noteq> y \<longleftrightarrow> (\<exists>U V. open U \<and> open V \<and> x \<in> U \<and> y \<in> V \<and> U \<inter> V = {})"
+  using hausdorff[of x y] by blast
+
+lemma separation_t1:
+  fixes x y :: "'a::t1_space"
+  shows "x \<noteq> y \<longleftrightarrow> (\<exists>U V. open U \<and> open V \<and> x \<in>U \<and> y\<notin> U \<and> x\<notin>V \<and> y\<in>V)"
+  using t1_space[of x y] by blast
+
+lemma separation_t0:
+  fixes x y :: "'a::t0_space"
+  shows "x \<noteq> y \<longleftrightarrow> (\<exists>U. open U \<and> ~(x\<in>U \<longleftrightarrow> y\<in>U))"
+  using t0_space[of x y] by blast
+
+subsection{* Limit points *}
+
+definition
+  islimpt:: "'a::topological_space \<Rightarrow> 'a set \<Rightarrow> bool"
+    (infixr "islimpt" 60) where
+  "x islimpt S \<longleftrightarrow> (\<forall>T. x\<in>T \<longrightarrow> open T \<longrightarrow> (\<exists>y\<in>S. y\<in>T \<and> y\<noteq>x))"
+
+lemma islimptI:
+  assumes "\<And>T. x \<in> T \<Longrightarrow> open T \<Longrightarrow> \<exists>y\<in>S. y \<in> T \<and> y \<noteq> x"
+  shows "x islimpt S"
+  using assms unfolding islimpt_def by auto
+
+lemma islimptE:
+  assumes "x islimpt S" and "x \<in> T" and "open T"
+  obtains y where "y \<in> S" and "y \<in> T" and "y \<noteq> x"
+  using assms unfolding islimpt_def by auto
+
+lemma islimpt_subset: "x islimpt S \<Longrightarrow> S \<subseteq> T ==> x islimpt T" by (auto simp add: islimpt_def)
+
+lemma islimpt_approachable:
+  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)"
+  unfolding islimpt_def
+  apply auto
+  apply(erule_tac x="ball x e" in allE)
+  apply auto
+  apply(rule_tac x=y in bexI)
+  apply (auto simp add: dist_commute)
+  apply (simp add: open_dist, drule (1) bspec)
+  apply (clarify, drule spec, drule (1) mp, auto)
+  done
+
+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)"
+  unfolding islimpt_approachable
+  using approachable_lt_le[where f="\<lambda>x'. dist x' x" and P="\<lambda>x'. \<not> (x'\<in>S \<and> x'\<noteq>x)"]
+  by metis (* FIXME: VERY slow! *)
+
+class perfect_space =
+  (* FIXME: perfect_space should inherit from topological_space *)
+  assumes islimpt_UNIV [simp, intro]: "(x::'a::metric_space) islimpt UNIV"
+
+lemma perfect_choose_dist:
+  fixes x :: "'a::perfect_space"
+  shows "0 < r \<Longrightarrow> \<exists>a. a \<noteq> x \<and> dist a x < r"
+using islimpt_UNIV [of x]
+by (simp add: islimpt_approachable)
+
+instance real :: perfect_space
+apply default
+apply (rule islimpt_approachable [THEN iffD2])
+apply (clarify, rule_tac x="x + e/2" in bexI)
+apply (auto simp add: dist_norm)
+done
+
+instance "^" :: (perfect_space, finite) perfect_space
+proof
+  fix x :: "'a ^ 'b"
+  {
+    fix e :: real assume "0 < e"
+    def a \<equiv> "x $ arbitrary"
+    have "a islimpt UNIV" by (rule islimpt_UNIV)
+    with `0 < e` obtain b where "b \<noteq> a" and "dist b a < e"
+      unfolding islimpt_approachable by auto
+    def y \<equiv> "Cart_lambda ((Cart_nth x)(arbitrary := b))"
+    from `b \<noteq> a` have "y \<noteq> x"
+      unfolding a_def y_def by (simp add: Cart_eq)
+    from `dist b a < e` have "dist y x < e"
+      unfolding dist_vector_def a_def y_def
+      apply simp
+      apply (rule le_less_trans [OF setL2_le_setsum [OF zero_le_dist]])
+      apply (subst setsum_diff1' [where a=arbitrary], simp, simp, simp)
+      done
+    from `y \<noteq> x` and `dist y x < e`
+    have "\<exists>y\<in>UNIV. y \<noteq> x \<and> dist y x < e" by auto
+  }
+  then show "x islimpt UNIV" unfolding islimpt_approachable by blast
+qed
+
+lemma closed_limpt: "closed S \<longleftrightarrow> (\<forall>x. x islimpt S \<longrightarrow> x \<in> S)"
+  unfolding closed_def
+  apply (subst open_subopen)
+  apply (simp add: islimpt_def subset_eq Compl_eq_Diff_UNIV)
+  by (metis DiffE DiffI UNIV_I insertCI insert_absorb mem_def)
+
+lemma islimpt_EMPTY[simp]: "\<not> x islimpt {}"
+  unfolding islimpt_def by auto
+
+lemma closed_positive_orthant: "closed {x::real^'n::finite. \<forall>i. 0 \<le>x$i}"
+proof-
+  let ?U = "UNIV :: 'n set"
+  let ?O = "{x::real^'n. \<forall>i. x$i\<ge>0}"
+  {fix x:: "real^'n" and i::'n assume H: "\<forall>e>0. \<exists>x'\<in>?O. x' \<noteq> x \<and> dist x' x < e"
+    and xi: "x$i < 0"
+    from xi have th0: "-x$i > 0" by arith
+    from H[rule_format, OF th0] obtain x' where x': "x' \<in>?O" "x' \<noteq> x" "dist x' x < -x $ i" by blast
+      have th:" \<And>b a (x::real). abs x <= b \<Longrightarrow> b <= a ==> ~(a + x < 0)" by arith
+      have th': "\<And>x (y::real). x < 0 \<Longrightarrow> 0 <= y ==> abs x <= abs (y - x)" by arith
+      have th1: "\<bar>x$i\<bar> \<le> \<bar>(x' - x)$i\<bar>" using x'(1) xi
+	apply (simp only: vector_component)
+	by (rule th') auto
+      have th2: "\<bar>dist x x'\<bar> \<ge> \<bar>(x' - x)$i\<bar>" using  component_le_norm[of "x'-x" i]
+	apply (simp add: dist_norm) by norm
+      from th[OF th1 th2] x'(3) have False by (simp add: dist_commute) }
+  then show ?thesis unfolding closed_limpt islimpt_approachable
+    unfolding not_le[symmetric] by blast
+qed
+
+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"
+proof(induct rule: finite_induct[OF fS])
+  case 1 thus ?case apply auto by ferrack
+next
+  case (2 x F)
+  from 2 obtain d where d: "d >0" "\<forall>x\<in>F. x\<noteq>a \<longrightarrow> d \<le> dist a x" by blast
+  {assume "x = a" hence ?case using d by auto  }
+  moreover
+  {assume xa: "x\<noteq>a"
+    let ?d = "min d (dist a x)"
+    have dp: "?d > 0" using xa d(1) using dist_nz by auto
+    from d have d': "\<forall>x\<in>F. x\<noteq>a \<longrightarrow> ?d \<le> dist a x" by auto
+    with dp xa have ?case by(auto intro!: exI[where x="?d"]) }
+  ultimately show ?case by blast
+qed
+
+lemma islimpt_finite:
+  fixes S :: "'a::metric_space set"
+  assumes fS: "finite S" shows "\<not> a islimpt S"
+  unfolding islimpt_approachable
+  using finite_set_avoid[OF fS, of a] by (metis dist_commute  not_le)
+
+lemma islimpt_Un: "x islimpt (S \<union> T) \<longleftrightarrow> x islimpt S \<or> x islimpt T"
+  apply (rule iffI)
+  defer
+  apply (metis Un_upper1 Un_upper2 islimpt_subset)
+  unfolding islimpt_def
+  apply (rule ccontr, clarsimp, rename_tac A B)
+  apply (drule_tac x="A \<inter> B" in spec)
+  apply (auto simp add: open_Int)
+  done
+
+lemma discrete_imp_closed:
+  fixes S :: "'a::metric_space set"
+  assumes e: "0 < e" and d: "\<forall>x \<in> S. \<forall>y \<in> S. dist y x < e \<longrightarrow> y = x"
+  shows "closed S"
+proof-
+  {fix x assume C: "\<forall>e>0. \<exists>x'\<in>S. x' \<noteq> x \<and> dist x' x < e"
+    from e have e2: "e/2 > 0" by arith
+    from C[rule_format, OF e2] obtain y where y: "y \<in> S" "y\<noteq>x" "dist y x < e/2" by blast
+    let ?m = "min (e/2) (dist x y) "
+    from e2 y(2) have mp: "?m > 0" by (simp add: dist_nz[THEN sym])
+    from C[rule_format, OF mp] obtain z where z: "z \<in> S" "z\<noteq>x" "dist z x < ?m" by blast
+    have th: "dist z y < e" using z y
+      by (intro dist_triangle_lt [where z=x], simp)
+    from d[rule_format, OF y(1) z(1) th] y z
+    have False by (auto simp add: dist_commute)}
+  then show ?thesis by (metis islimpt_approachable closed_limpt [where 'a='a])
+qed
+
+subsection{* Interior of a Set *}
+definition "interior S = {x. \<exists>T. open T \<and> x \<in> T \<and> T \<subseteq> S}"
+
+lemma interior_eq: "interior S = S \<longleftrightarrow> open S"
+  apply (simp add: expand_set_eq interior_def)
+  apply (subst (2) open_subopen) by (safe, blast+)
+
+lemma interior_open: "open S ==> (interior S = S)" by (metis interior_eq)
+
+lemma interior_empty[simp]: "interior {} = {}" by (simp add: interior_def)
+
+lemma open_interior[simp, intro]: "open(interior S)"
+  apply (simp add: interior_def)
+  apply (subst open_subopen) by blast
+
+lemma interior_interior[simp]: "interior(interior S) = interior S" by (metis interior_eq open_interior)
+lemma interior_subset: "interior S \<subseteq> S" by (auto simp add: interior_def)
+lemma subset_interior: "S \<subseteq> T ==> (interior S) \<subseteq> (interior T)" by (auto simp add: interior_def)
+lemma interior_maximal: "T \<subseteq> S \<Longrightarrow> open T ==> T \<subseteq> (interior S)" by (auto simp add: interior_def)
+lemma interior_unique: "T \<subseteq> S \<Longrightarrow> open T  \<Longrightarrow> (\<forall>T'. T' \<subseteq> S \<and> open T' \<longrightarrow> T' \<subseteq> T) \<Longrightarrow> interior S = T"
+  by (metis equalityI interior_maximal interior_subset open_interior)
+lemma mem_interior: "x \<in> interior S \<longleftrightarrow> (\<exists>e. 0 < e \<and> ball x e \<subseteq> S)"
+  apply (simp add: interior_def)
+  by (metis open_contains_ball centre_in_ball open_ball subset_trans)
+
+lemma open_subset_interior: "open S ==> S \<subseteq> interior T \<longleftrightarrow> S \<subseteq> T"
+  by (metis interior_maximal interior_subset subset_trans)
+
+lemma interior_inter[simp]: "interior(S \<inter> T) = interior S \<inter> interior T"
+  apply (rule equalityI, simp)
+  apply (metis Int_lower1 Int_lower2 subset_interior)
+  by (metis Int_mono interior_subset open_Int open_interior open_subset_interior)
+
+lemma interior_limit_point [intro]:
+  fixes x :: "'a::perfect_space"
+  assumes x: "x \<in> interior S" shows "x islimpt S"
+proof-
+  from x obtain e where e: "e>0" "\<forall>x'. dist x x' < e \<longrightarrow> x' \<in> S"
+    unfolding mem_interior subset_eq Ball_def mem_ball by blast
+  {
+    fix d::real assume d: "d>0"
+    let ?m = "min d e"
+    have mde2: "0 < ?m" using e(1) d(1) by simp
+    from perfect_choose_dist [OF mde2, of x]
+    obtain y where "y \<noteq> x" and "dist y x < ?m" by blast
+    then have "dist y x < e" "dist y x < d" by simp_all
+    from `dist y x < e` e(2) have "y \<in> S" by (simp add: dist_commute)
+    have "\<exists>x'\<in>S. x'\<noteq> x \<and> dist x' x < d"
+      using `y \<in> S` `y \<noteq> x` `dist y x < d` by fast
+  }
+  then show ?thesis unfolding islimpt_approachable by blast
+qed
+
+lemma interior_closed_Un_empty_interior:
+  assumes cS: "closed S" and iT: "interior T = {}"
+  shows "interior(S \<union> T) = interior S"
+proof
+  show "interior S \<subseteq> interior (S\<union>T)"
+    by (rule subset_interior, blast)
+next
+  show "interior (S \<union> T) \<subseteq> interior S"
+  proof
+    fix x assume "x \<in> interior (S \<union> T)"
+    then obtain R where "open R" "x \<in> R" "R \<subseteq> S \<union> T"
+      unfolding interior_def by fast
+    show "x \<in> interior S"
+    proof (rule ccontr)
+      assume "x \<notin> interior S"
+      with `x \<in> R` `open R` obtain y where "y \<in> R - S"
+        unfolding interior_def expand_set_eq by fast
+      from `open R` `closed S` have "open (R - S)" by (rule open_Diff)
+      from `R \<subseteq> S \<union> T` have "R - S \<subseteq> T" by fast
+      from `y \<in> R - S` `open (R - S)` `R - S \<subseteq> T` `interior T = {}`
+      show "False" unfolding interior_def by fast
+    qed
+  qed
+qed
+
+
+subsection{* Closure of a Set *}
+
+definition "closure S = S \<union> {x | x. x islimpt S}"
+
+lemma closure_interior: "closure S = UNIV - interior (UNIV - S)"
+proof-
+  { fix x
+    have "x\<in>UNIV - interior (UNIV - S) \<longleftrightarrow> x \<in> closure S"  (is "?lhs = ?rhs")
+    proof
+      let ?exT = "\<lambda> y. (\<exists>T. open T \<and> y \<in> T \<and> T \<subseteq> UNIV - S)"
+      assume "?lhs"
+      hence *:"\<not> ?exT x"
+	unfolding interior_def
+	by simp
+      { assume "\<not> ?rhs"
+	hence False using *
+	  unfolding closure_def islimpt_def
+	  by blast
+      }
+      thus "?rhs"
+	by blast
+    next
+      assume "?rhs" thus "?lhs"
+	unfolding closure_def interior_def islimpt_def
+	by blast
+    qed
+  }
+  thus ?thesis
+    by blast
+qed
+
+lemma interior_closure: "interior S = UNIV - (closure (UNIV - S))"
+proof-
+  { fix x
+    have "x \<in> interior S \<longleftrightarrow> x \<in> UNIV - (closure (UNIV - S))"
+      unfolding interior_def closure_def islimpt_def
+      by blast (* FIXME: VERY slow! *)
+  }
+  thus ?thesis
+    by blast
+qed
+
+lemma closed_closure[simp, intro]: "closed (closure S)"
+proof-
+  have "closed (UNIV - interior (UNIV -S))" by blast
+  thus ?thesis using closure_interior[of S] by simp
+qed
+
+lemma closure_hull: "closure S = closed hull S"
+proof-
+  have "S \<subseteq> closure S"
+    unfolding closure_def
+    by blast
+  moreover
+  have "closed (closure S)"
+    using closed_closure[of S]
+    by assumption
+  moreover
+  { fix t
+    assume *:"S \<subseteq> t" "closed t"
+    { fix x
+      assume "x islimpt S"
+      hence "x islimpt t" using *(1)
+	using islimpt_subset[of x, of S, of t]
+	by blast
+    }
+    with * have "closure S \<subseteq> t"
+      unfolding closure_def
+      using closed_limpt[of t]
+      by auto
+  }
+  ultimately show ?thesis
+    using hull_unique[of S, of "closure S", of closed]
+    unfolding mem_def
+    by simp
+qed
+
+lemma closure_eq: "closure S = S \<longleftrightarrow> closed S"
+  unfolding closure_hull
+  using hull_eq[of closed, unfolded mem_def, OF  closed_Inter, of S]
+  by (metis mem_def subset_eq)
+
+lemma closure_closed[simp]: "closed S \<Longrightarrow> closure S = S"
+  using closure_eq[of S]
+  by simp
+
+lemma closure_closure[simp]: "closure (closure S) = closure S"
+  unfolding closure_hull
+  using hull_hull[of closed S]
+  by assumption
+
+lemma closure_subset: "S \<subseteq> closure S"
+  unfolding closure_hull
+  using hull_subset[of S closed]
+  by assumption
+
+lemma subset_closure: "S \<subseteq> T \<Longrightarrow> closure S \<subseteq> closure T"
+  unfolding closure_hull
+  using hull_mono[of S T closed]
+  by assumption
+
+lemma closure_minimal: "S \<subseteq> T \<Longrightarrow>  closed T \<Longrightarrow> closure S \<subseteq> T"
+  using hull_minimal[of S T closed]
+  unfolding closure_hull mem_def
+  by simp
+
+lemma closure_unique: "S \<subseteq> T \<and> closed T \<and> (\<forall> T'. S \<subseteq> T' \<and> closed T' \<longrightarrow> T \<subseteq> T') \<Longrightarrow> closure S = T"
+  using hull_unique[of S T closed]
+  unfolding closure_hull mem_def
+  by simp
+
+lemma closure_empty[simp]: "closure {} = {}"
+  using closed_empty closure_closed[of "{}"]
+  by simp
+
+lemma closure_univ[simp]: "closure UNIV = UNIV"
+  using closure_closed[of UNIV]
+  by simp
+
+lemma closure_eq_empty: "closure S = {} \<longleftrightarrow> S = {}"
+  using closure_empty closure_subset[of S]
+  by blast
+
+lemma closure_subset_eq: "closure S \<subseteq> S \<longleftrightarrow> closed S"
+  using closure_eq[of S] closure_subset[of S]
+  by simp
+
+lemma open_inter_closure_eq_empty:
+  "open S \<Longrightarrow> (S \<inter> closure T) = {} \<longleftrightarrow> S \<inter> T = {}"
+  using open_subset_interior[of S "UNIV - T"]
+  using interior_subset[of "UNIV - T"]
+  unfolding closure_interior
+  by auto
+
+lemma open_inter_closure_subset:
+  "open S \<Longrightarrow> (S \<inter> (closure T)) \<subseteq> closure(S \<inter> T)"
+proof
+  fix x
+  assume as: "open S" "x \<in> S \<inter> closure T"
+  { assume *:"x islimpt T"
+    have "x islimpt (S \<inter> T)"
+    proof (rule islimptI)
+      fix A
+      assume "x \<in> A" "open A"
+      with as have "x \<in> A \<inter> S" "open (A \<inter> S)"
+        by (simp_all add: open_Int)
+      with * obtain y where "y \<in> T" "y \<in> A \<inter> S" "y \<noteq> x"
+        by (rule islimptE)
+      hence "y \<in> S \<inter> T" "y \<in> A \<and> y \<noteq> x"
+        by simp_all
+      thus "\<exists>y\<in>(S \<inter> T). y \<in> A \<and> y \<noteq> x" ..
+    qed
+  }
+  then show "x \<in> closure (S \<inter> T)" using as
+    unfolding closure_def
+    by blast
+qed
+
+lemma closure_complement: "closure(UNIV - S) = UNIV - interior(S)"
+proof-
+  have "S = UNIV - (UNIV - S)"
+    by auto
+  thus ?thesis
+    unfolding closure_interior
+    by auto
+qed
+
+lemma interior_complement: "interior(UNIV - S) = UNIV - closure(S)"
+  unfolding closure_interior
+  by blast
+
+subsection{* Frontier (aka boundary) *}
+
+definition "frontier S = closure S - interior S"
+
+lemma frontier_closed: "closed(frontier S)"
+  by (simp add: frontier_def closed_Diff)
+
+lemma frontier_closures: "frontier S = (closure S) \<inter> (closure(UNIV - S))"
+  by (auto simp add: frontier_def interior_closure)
+
+lemma frontier_straddle:
+  fixes a :: "'a::metric_space"
+  shows "a \<in> frontier S \<longleftrightarrow> (\<forall>e>0. (\<exists>x\<in>S. dist a x < e) \<and> (\<exists>x. x \<notin> S \<and> dist a x < e))" (is "?lhs \<longleftrightarrow> ?rhs")
+proof
+  assume "?lhs"
+  { fix e::real
+    assume "e > 0"
+    let ?rhse = "(\<exists>x\<in>S. dist a x < e) \<and> (\<exists>x. x \<notin> S \<and> dist a x < e)"
+    { assume "a\<in>S"
+      have "\<exists>x\<in>S. dist a x < e" using `e>0` `a\<in>S` by(rule_tac x=a in bexI) auto
+      moreover have "\<exists>x. x \<notin> S \<and> dist a x < e" using `?lhs` `a\<in>S`
+	unfolding frontier_closures closure_def islimpt_def using `e>0`
+	by (auto, erule_tac x="ball a e" in allE, auto)
+      ultimately have ?rhse by auto
+    }
+    moreover
+    { assume "a\<notin>S"
+      hence ?rhse using `?lhs`
+	unfolding frontier_closures closure_def islimpt_def
+	using open_ball[of a e] `e > 0`
+	by (auto, erule_tac x = "ball a e" in allE, auto) (* FIXME: VERY slow! *)
+    }
+    ultimately have ?rhse by auto
+  }
+  thus ?rhs by auto
+next
+  assume ?rhs
+  moreover
+  { fix T assume "a\<notin>S" and
+    as:"\<forall>e>0. (\<exists>x\<in>S. dist a x < e) \<and> (\<exists>x. x \<notin> S \<and> dist a x < e)" "a \<notin> S" "a \<in> T" "open T"
+    from `open T` `a \<in> T` have "\<exists>e>0. ball a e \<subseteq> T" unfolding open_contains_ball[of T] by auto
+    then obtain e where "e>0" "ball a e \<subseteq> T" by auto
+    then obtain y where y:"y\<in>S" "dist a y < e"  using as(1) by auto
+    have "\<exists>y\<in>S. y \<in> T \<and> y \<noteq> a"
+      using `dist a y < e` `ball a e \<subseteq> T` unfolding ball_def using `y\<in>S` `a\<notin>S` by auto
+  }
+  hence "a \<in> closure S" unfolding closure_def islimpt_def using `?rhs` by auto
+  moreover
+  { fix T assume "a \<in> T"  "open T" "a\<in>S"
+    then obtain e where "e>0" and balle: "ball a e \<subseteq> T" unfolding open_contains_ball using `?rhs` by auto
+    obtain x where "x \<notin> S" "dist a x < e" using `?rhs` using `e>0` by auto
+    hence "\<exists>y\<in>UNIV - S. y \<in> T \<and> y \<noteq> a" using balle `a\<in>S` unfolding ball_def by (rule_tac x=x in bexI)auto
+  }
+  hence "a islimpt (UNIV - S) \<or> a\<notin>S" unfolding islimpt_def by auto
+  ultimately show ?lhs unfolding frontier_closures using closure_def[of "UNIV - S"] by auto
+qed
+
+lemma frontier_subset_closed: "closed S \<Longrightarrow> frontier S \<subseteq> S"
+  by (metis frontier_def closure_closed Diff_subset)
+
+lemma frontier_empty: "frontier {} = {}"
+  by (simp add: frontier_def closure_empty)
+
+lemma frontier_subset_eq: "frontier S \<subseteq> S \<longleftrightarrow> closed S"
+proof-
+  { assume "frontier S \<subseteq> S"
+    hence "closure S \<subseteq> S" using interior_subset unfolding frontier_def by auto
+    hence "closed S" using closure_subset_eq by auto
+  }
+  thus ?thesis using frontier_subset_closed[of S] by auto
+qed
+
+lemma frontier_complement: "frontier(UNIV - S) = frontier S"
+  by (auto simp add: frontier_def closure_complement interior_complement)
+
+lemma frontier_disjoint_eq: "frontier S \<inter> S = {} \<longleftrightarrow> open S"
+  using frontier_complement frontier_subset_eq[of "UNIV - S"]
+  unfolding open_closed Compl_eq_Diff_UNIV by auto
+
+subsection{* Common nets and The "within" modifier for nets. *}
+
+definition
+  at_infinity :: "'a::real_normed_vector net" where
+  "at_infinity = Abs_net (range (\<lambda>r. {x. r \<le> norm x}))"
+
+definition
+  indirection :: "'a::real_normed_vector \<Rightarrow> 'a \<Rightarrow> 'a net" (infixr "indirection" 70) where
+  "a indirection v = (at a) within {b. \<exists>c\<ge>0. b - a = scaleR c v}"
+
+text{* Prove That They are all nets. *}
+
+lemma Rep_net_at_infinity:
+  "Rep_net at_infinity = range (\<lambda>r. {x. r \<le> norm x})"
+unfolding at_infinity_def
+apply (rule Abs_net_inverse')
+apply (rule image_nonempty, simp)
+apply (clarsimp, rename_tac r s)
+apply (rule_tac x="max r s" in exI, auto)
+done
+
+lemma within_UNIV: "net within UNIV = net"
+  by (simp add: Rep_net_inject [symmetric] Rep_net_within)
+
+subsection{* Identify Trivial limits, where we can't approach arbitrarily closely. *}
+
+definition
+  trivial_limit :: "'a net \<Rightarrow> bool" where
+  "trivial_limit net \<longleftrightarrow> {} \<in> Rep_net net"
+
+lemma trivial_limit_within:
+  shows "trivial_limit (at a within S) \<longleftrightarrow> \<not> a islimpt S"
+proof
+  assume "trivial_limit (at a within S)"
+  thus "\<not> a islimpt S"
+    unfolding trivial_limit_def
+    unfolding Rep_net_within Rep_net_at
+    unfolding islimpt_def
+    apply (clarsimp simp add: expand_set_eq)
+    apply (rename_tac T, rule_tac x=T in exI)
+    apply (clarsimp, drule_tac x=y in spec, simp)
+    done
+next
+  assume "\<not> a islimpt S"
+  thus "trivial_limit (at a within S)"
+    unfolding trivial_limit_def
+    unfolding Rep_net_within Rep_net_at
+    unfolding islimpt_def
+    apply (clarsimp simp add: image_image)
+    apply (rule_tac x=T in image_eqI)
+    apply (auto simp add: expand_set_eq)
+    done
+qed
+
+lemma trivial_limit_at_iff: "trivial_limit (at a) \<longleftrightarrow> \<not> a islimpt UNIV"
+  using trivial_limit_within [of a UNIV]
+  by (simp add: within_UNIV)
+
+lemma trivial_limit_at:
+  fixes a :: "'a::perfect_space"
+  shows "\<not> trivial_limit (at a)"
+  by (simp add: trivial_limit_at_iff)
+
+lemma trivial_limit_at_infinity:
+  "\<not> trivial_limit (at_infinity :: ('a::{real_normed_vector,zero_neq_one}) net)"
+  (* FIXME: find a more appropriate type class *)
+  unfolding trivial_limit_def Rep_net_at_infinity
+  apply (clarsimp simp add: expand_set_eq)
+  apply (drule_tac x="scaleR r (sgn 1)" in spec)
+  apply (simp add: norm_sgn)
+  done
+
+lemma trivial_limit_sequentially: "\<not> trivial_limit sequentially"
+  by (auto simp add: trivial_limit_def Rep_net_sequentially)
+
+subsection{* Some property holds "sufficiently close" to the limit point. *}
+
+lemma eventually_at: (* FIXME: this replaces Limits.eventually_at *)
+  "eventually P (at a) \<longleftrightarrow> (\<exists>d>0. \<forall>x. 0 < dist x a \<and> dist x a < d \<longrightarrow> P x)"
+unfolding eventually_at dist_nz by auto
+
+lemma eventually_at_infinity:
+  "eventually P at_infinity \<longleftrightarrow> (\<exists>b. \<forall>x. norm x >= b \<longrightarrow> P x)"
+unfolding eventually_def Rep_net_at_infinity by auto
+
+lemma eventually_within: "eventually P (at a within S) \<longleftrightarrow>
+        (\<exists>d>0. \<forall>x\<in>S. 0 < dist x a \<and> dist x a < d \<longrightarrow> P x)"
+unfolding eventually_within eventually_at dist_nz by auto
+
+lemma eventually_within_le: "eventually P (at a within S) \<longleftrightarrow>
+        (\<exists>d>0. \<forall>x\<in>S. 0 < dist x a \<and> dist x a <= d \<longrightarrow> P x)" (is "?lhs = ?rhs")
+unfolding eventually_within
+apply safe
+apply (rule_tac x="d/2" in exI, simp)
+apply (rule_tac x="d" in exI, simp)
+done
+
+lemma eventually_happens: "eventually P net ==> trivial_limit net \<or> (\<exists>x. P x)"
+  unfolding eventually_def trivial_limit_def
+  using Rep_net_nonempty [of net] by auto
+
+lemma always_eventually: "(\<forall>x. P x) ==> eventually P net"
+  unfolding eventually_def trivial_limit_def
+  using Rep_net_nonempty [of net] by auto
+
+lemma trivial_limit_eventually: "trivial_limit net \<Longrightarrow> eventually P net"
+  unfolding trivial_limit_def eventually_def by auto
+
+lemma eventually_False: "eventually (\<lambda>x. False) net \<longleftrightarrow> trivial_limit net"
+  unfolding trivial_limit_def eventually_def by auto
+
+lemma trivial_limit_eq: "trivial_limit net \<longleftrightarrow> (\<forall>P. eventually P net)"
+  apply (safe elim!: trivial_limit_eventually)
+  apply (simp add: eventually_False [symmetric])
+  done
+
+text{* Combining theorems for "eventually" *}
+
+lemma eventually_conjI:
+  "\<lbrakk>eventually (\<lambda>x. P x) net; eventually (\<lambda>x. Q x) net\<rbrakk>
+    \<Longrightarrow> eventually (\<lambda>x. P x \<and> Q x) net"
+by (rule eventually_conj)
+
+lemma eventually_rev_mono:
+  "eventually P net \<Longrightarrow> (\<forall>x. P x \<longrightarrow> Q x) \<Longrightarrow> eventually Q net"
+using eventually_mono [of P Q] by fast
+
+lemma eventually_and: " eventually (\<lambda>x. P x \<and> Q x) net \<longleftrightarrow> eventually P net \<and> eventually Q net"
+  by (auto intro!: eventually_conjI elim: eventually_rev_mono)
+
+lemma eventually_false: "eventually (\<lambda>x. False) net \<longleftrightarrow> trivial_limit net"
+  by (auto simp add: eventually_False)
+
+lemma not_eventually: "(\<forall>x. \<not> P x ) \<Longrightarrow> ~(trivial_limit net) ==> ~(eventually (\<lambda>x. P x) net)"
+  by (simp add: eventually_False)
+
+subsection{* Limits, defined as vacuously true when the limit is trivial. *}
+
+  text{* Notation Lim to avoid collition with lim defined in analysis *}
+definition
+  Lim :: "'a net \<Rightarrow> ('a \<Rightarrow> 'b::t2_space) \<Rightarrow> 'b" where
+  "Lim net f = (THE l. (f ---> l) net)"
+
+lemma Lim:
+ "(f ---> l) net \<longleftrightarrow>
+        trivial_limit net \<or>
+        (\<forall>e>0. eventually (\<lambda>x. dist (f x) l < e) net)"
+  unfolding tendsto_iff trivial_limit_eq by auto
+
+
+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)"
+  by (auto simp add: tendsto_iff eventually_within_le)
+
+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)"
+  by (auto simp add: tendsto_iff eventually_within)
+
+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)"
+  by (auto simp add: tendsto_iff eventually_at)
+
+lemma Lim_at_iff_LIM: "(f ---> l) (at a) \<longleftrightarrow> f -- a --> l"
+  unfolding Lim_at LIM_def by (simp only: zero_less_dist_iff)
+
+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)"
+  by (auto simp add: tendsto_iff eventually_at_infinity)
+
+lemma Lim_sequentially:
+ "(S ---> l) sequentially \<longleftrightarrow>
+          (\<forall>e>0. \<exists>N. \<forall>n\<ge>N. dist (S n) l < e)"
+  by (auto simp add: tendsto_iff eventually_sequentially)
+
+lemma Lim_sequentially_iff_LIMSEQ: "(S ---> l) sequentially \<longleftrightarrow> S ----> l"
+  unfolding Lim_sequentially LIMSEQ_def ..
+
+lemma Lim_eventually: "eventually (\<lambda>x. f x = l) net \<Longrightarrow> (f ---> l) net"
+  by (rule topological_tendstoI, auto elim: eventually_rev_mono)
+
+text{* The expected monotonicity property. *}
+
+lemma Lim_within_empty: "(f ---> l) (net within {})"
+  unfolding tendsto_def Limits.eventually_within by simp
+
+lemma Lim_within_subset: "(f ---> l) (net within S) \<Longrightarrow> T \<subseteq> S \<Longrightarrow> (f ---> l) (net within T)"
+  unfolding tendsto_def Limits.eventually_within
+  by (auto elim!: eventually_elim1)
+
+lemma Lim_Un: assumes "(f ---> l) (net within S)" "(f ---> l) (net within T)"
+  shows "(f ---> l) (net within (S \<union> T))"
+  using assms unfolding tendsto_def Limits.eventually_within
+  apply clarify
+  apply (drule spec, drule (1) mp, drule (1) mp)
+  apply (drule spec, drule (1) mp, drule (1) mp)
+  apply (auto elim: eventually_elim2)
+  done
+
+lemma Lim_Un_univ:
+ "(f ---> l) (net within S) \<Longrightarrow> (f ---> l) (net within T) \<Longrightarrow>  S \<union> T = UNIV
+        ==> (f ---> l) net"
+  by (metis Lim_Un within_UNIV)
+
+text{* Interrelations between restricted and unrestricted limits. *}
+
+lemma Lim_at_within: "(f ---> l) net ==> (f ---> l)(net within S)"
+  (* FIXME: rename *)
+  unfolding tendsto_def Limits.eventually_within
+  apply (clarify, drule spec, drule (1) mp, drule (1) mp)
+  by (auto elim!: eventually_elim1)
+
+lemma Lim_within_open:
+  fixes f :: "'a::topological_space \<Rightarrow> 'b::topological_space"
+  assumes"a \<in> S" "open S"
+  shows "(f ---> l)(at a within S) \<longleftrightarrow> (f ---> l)(at a)" (is "?lhs \<longleftrightarrow> ?rhs")
+proof
+  assume ?lhs
+  { fix A assume "open A" "l \<in> A"
+    with `?lhs` have "eventually (\<lambda>x. f x \<in> A) (at a within S)"
+      by (rule topological_tendstoD)
+    hence "eventually (\<lambda>x. x \<in> S \<longrightarrow> f x \<in> A) (at a)"
+      unfolding Limits.eventually_within .
+    then obtain T where "open T" "a \<in> T" "\<forall>x\<in>T. x \<noteq> a \<longrightarrow> x \<in> S \<longrightarrow> f x \<in> A"
+      unfolding eventually_at_topological by fast
+    hence "open (T \<inter> S)" "a \<in> T \<inter> S" "\<forall>x\<in>(T \<inter> S). x \<noteq> a \<longrightarrow> f x \<in> A"
+      using assms by auto
+    hence "\<exists>T. open T \<and> a \<in> T \<and> (\<forall>x\<in>T. x \<noteq> a \<longrightarrow> f x \<in> A)"
+      by fast
+    hence "eventually (\<lambda>x. f x \<in> A) (at a)"
+      unfolding eventually_at_topological .
+  }
+  thus ?rhs by (rule topological_tendstoI)
+next
+  assume ?rhs
+  thus ?lhs by (rule Lim_at_within)
+qed
+
+text{* Another limit point characterization. *}
+
+lemma islimpt_sequential:
+  fixes x :: "'a::metric_space" (* FIXME: generalize to topological_space *)
+  shows "x islimpt S \<longleftrightarrow> (\<exists>f. (\<forall>n::nat. f n \<in> S -{x}) \<and> (f ---> x) sequentially)"
+    (is "?lhs = ?rhs")
+proof
+  assume ?lhs
+  then obtain f where f:"\<forall>y. y>0 \<longrightarrow> f y \<in> S \<and> f y \<noteq> x \<and> dist (f y) x < y"
+    unfolding islimpt_approachable using choice[of "\<lambda>e y. e>0 \<longrightarrow> y\<in>S \<and> y\<noteq>x \<and> dist y x < e"] by auto
+  { fix n::nat
+    have "f (inverse (real n + 1)) \<in> S - {x}" using f by auto
+  }
+  moreover
+  { fix e::real assume "e>0"
+    hence "\<exists>N::nat. inverse (real (N + 1)) < e" using real_arch_inv[of e] apply (auto simp add: Suc_pred') apply(rule_tac x="n - 1" in exI) by auto
+    then obtain N::nat where "inverse (real (N + 1)) < e" by auto
+    hence "\<forall>n\<ge>N. inverse (real n + 1) < e" by (auto, metis Suc_le_mono le_SucE less_imp_inverse_less nat_le_real_less order_less_trans real_of_nat_Suc real_of_nat_Suc_gt_zero)
+    moreover have "\<forall>n\<ge>N. dist (f (inverse (real n + 1))) x < (inverse (real n + 1))" using f `e>0` by auto
+    ultimately have "\<exists>N::nat. \<forall>n\<ge>N. dist (f (inverse (real n + 1))) x < e" apply(rule_tac x=N in exI) apply auto apply(erule_tac x=n in allE)+ by auto
+  }
+  hence " ((\<lambda>n. f (inverse (real n + 1))) ---> x) sequentially"
+    unfolding Lim_sequentially using f by auto
+  ultimately show ?rhs apply (rule_tac x="(\<lambda>n::nat. f (inverse (real n + 1)))" in exI) by auto
+next
+  assume ?rhs
+  then obtain f::"nat\<Rightarrow>'a"  where f:"(\<forall>n. f n \<in> S - {x})" "(\<forall>e>0. \<exists>N. \<forall>n\<ge>N. dist (f n) x < e)" unfolding Lim_sequentially by auto
+  { fix e::real assume "e>0"
+    then obtain N where "dist (f N) x < e" using f(2) by auto
+    moreover have "f N\<in>S" "f N \<noteq> x" using f(1) by auto
+    ultimately have "\<exists>x'\<in>S. x' \<noteq> x \<and> dist x' x < e" by auto
+  }
+  thus ?lhs unfolding islimpt_approachable by auto
+qed
+
+text{* Basic arithmetical combining theorems for limits. *}
+
+lemma Lim_linear:
+  assumes "(f ---> l) net" "bounded_linear h"
+  shows "((\<lambda>x. h (f x)) ---> h l) net"
+using `bounded_linear h` `(f ---> l) net`
+by (rule bounded_linear.tendsto)
+
+lemma Lim_ident_at: "((\<lambda>x. x) ---> a) (at a)"
+  unfolding tendsto_def Limits.eventually_at_topological by fast
+
+lemma Lim_const: "((\<lambda>x. a) ---> a) net"
+  by (rule tendsto_const)
+
+lemma Lim_cmul:
+  fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"
+  shows "(f ---> l) net ==> ((\<lambda>x. c *\<^sub>R f x) ---> c *\<^sub>R l) net"
+  by (intro tendsto_intros)
+
+lemma Lim_neg:
+  fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"
+  shows "(f ---> l) net ==> ((\<lambda>x. -(f x)) ---> -l) net"
+  by (rule tendsto_minus)
+
+lemma Lim_add: fixes f :: "'a \<Rightarrow> 'b::real_normed_vector" shows
+ "(f ---> l) net \<Longrightarrow> (g ---> m) net \<Longrightarrow> ((\<lambda>x. f(x) + g(x)) ---> l + m) net"
+  by (rule tendsto_add)
+
+lemma Lim_sub:
+  fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"
+  shows "(f ---> l) net \<Longrightarrow> (g ---> m) net \<Longrightarrow> ((\<lambda>x. f(x) - g(x)) ---> l - m) net"
+  by (rule tendsto_diff)
+
+lemma Lim_null:
+  fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"
+  shows "(f ---> l) net \<longleftrightarrow> ((\<lambda>x. f(x) - l) ---> 0) net" by (simp add: Lim dist_norm)
+
+lemma Lim_null_norm:
+  fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"
+  shows "(f ---> 0) net \<longleftrightarrow> ((\<lambda>x. norm(f x)) ---> 0) net"
+  by (simp add: Lim dist_norm)
+
+lemma Lim_null_comparison:
+  fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"
+  assumes "eventually (\<lambda>x. norm (f x) \<le> g x) net" "(g ---> 0) net"
+  shows "(f ---> 0) net"
+proof(simp add: tendsto_iff, rule+)
+  fix e::real assume "0<e"
+  { fix x
+    assume "norm (f x) \<le> g x" "dist (g x) 0 < e"
+    hence "dist (f x) 0 < e" by (simp add: dist_norm)
+  }
+  thus "eventually (\<lambda>x. dist (f x) 0 < e) net"
+    using eventually_and[of "\<lambda>x. norm(f x) <= g x" "\<lambda>x. dist (g x) 0 < e" net]
+    using eventually_mono[of "(\<lambda>x. norm (f x) \<le> g x \<and> dist (g x) 0 < e)" "(\<lambda>x. dist (f x) 0 < e)" net]
+    using assms `e>0` unfolding tendsto_iff by auto
+qed
+
+lemma Lim_component:
+  fixes f :: "'a \<Rightarrow> 'b::metric_space ^ 'n::finite"
+  shows "(f ---> l) net \<Longrightarrow> ((\<lambda>a. f a $i) ---> l$i) net"
+  unfolding tendsto_iff
+  apply (clarify)
+  apply (drule spec, drule (1) mp)
+  apply (erule eventually_elim1)
+  apply (erule le_less_trans [OF dist_nth_le])
+  done
+
+lemma Lim_transform_bound:
+  fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"
+  fixes g :: "'a \<Rightarrow> 'c::real_normed_vector"
+  assumes "eventually (\<lambda>n. norm(f n) <= norm(g n)) net"  "(g ---> 0) net"
+  shows "(f ---> 0) net"
+proof (rule tendstoI)
+  fix e::real assume "e>0"
+  { fix x
+    assume "norm (f x) \<le> norm (g x)" "dist (g x) 0 < e"
+    hence "dist (f x) 0 < e" by (simp add: dist_norm)}
+  thus "eventually (\<lambda>x. dist (f x) 0 < e) net"
+    using eventually_and[of "\<lambda>x. norm (f x) \<le> norm (g x)" "\<lambda>x. dist (g x) 0 < e" net]
+    using eventually_mono[of "\<lambda>x. norm (f x) \<le> norm (g x) \<and> dist (g x) 0 < e" "\<lambda>x. dist (f x) 0 < e" net]
+    using assms `e>0` unfolding tendsto_iff by blast
+qed
+
+text{* Deducing things about the limit from the elements. *}
+
+lemma Lim_in_closed_set:
+  assumes "closed S" "eventually (\<lambda>x. f(x) \<in> S) net" "\<not>(trivial_limit net)" "(f ---> l) net"
+  shows "l \<in> S"
+proof (rule ccontr)
+  assume "l \<notin> S"
+  with `closed S` have "open (- S)" "l \<in> - S"
+    by (simp_all add: open_Compl)
+  with assms(4) have "eventually (\<lambda>x. f x \<in> - S) net"
+    by (rule topological_tendstoD)
+  with assms(2) have "eventually (\<lambda>x. False) net"
+    by (rule eventually_elim2) simp
+  with assms(3) show "False"
+    by (simp add: eventually_False)
+qed
+
+text{* Need to prove closed(cball(x,e)) before deducing this as a corollary. *}
+
+lemma Lim_dist_ubound:
+  assumes "\<not>(trivial_limit net)" "(f ---> l) net" "eventually (\<lambda>x. dist a (f x) <= e) net"
+  shows "dist a l <= e"
+proof (rule ccontr)
+  assume "\<not> dist a l \<le> e"
+  then have "0 < dist a l - e" by simp
+  with assms(2) have "eventually (\<lambda>x. dist (f x) l < dist a l - e) net"
+    by (rule tendstoD)
+  with assms(3) have "eventually (\<lambda>x. dist a (f x) \<le> e \<and> dist (f x) l < dist a l - e) net"
+    by (rule eventually_conjI)
+  then obtain w where "dist a (f w) \<le> 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)"
+    by (rule add_le_less_mono)
+  hence "dist a (f w) + dist (f w) l < dist a l"
+    by simp
+  also have "\<dots> \<le> dist a (f w) + dist (f w) l"
+    by (rule dist_triangle)
+  finally show False by simp
+qed
+
+lemma Lim_norm_ubound:
+  fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"
+  assumes "\<not>(trivial_limit net)" "(f ---> l) net" "eventually (\<lambda>x. norm(f x) <= e) net"
+  shows "norm(l) <= e"
+proof (rule ccontr)
+  assume "\<not> norm l \<le> e"
+  then have "0 < norm l - e" by simp
+  with assms(2) have "eventually (\<lambda>x. dist (f x) l < norm l - e) net"
+    by (rule tendstoD)
+  with assms(3) have "eventually (\<lambda>x. norm (f x) \<le> e \<and> dist (f x) l < norm l - e) net"
+    by (rule eventually_conjI)
+  then obtain w where "norm (f w) \<le> 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) \<le> e" by (simp_all add: dist_norm)
+  hence "norm (f w - l) + norm (f w) < norm l" by simp
+  hence "norm (f w - l - f w) < norm l" by (rule le_less_trans [OF norm_triangle_ineq4])
+  thus False using `\<not> norm l \<le> e` by simp
+qed
+
+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"
+  shows "e \<le> norm l"
+proof (rule ccontr)
+  assume "\<not> e \<le> norm l"
+  then have "0 < e - norm l" by simp
+  with assms(2) have "eventually (\<lambda>x. dist (f x) l < e - norm l) net"
+    by (rule tendstoD)
+  with assms(3) have "eventually (\<lambda>x. e \<le> norm (f x) \<and> dist (f x) l < e - norm l) net"
+    by (rule eventually_conjI)
+  then obtain w where "e \<le> 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 \<le> norm (f w)" by (simp_all add: dist_norm)
+  hence "norm (f w - l) + norm l < norm (f w)" by (rule less_le_trans)
+  hence "norm (f w - l + l) < norm (f w)" by (rule le_less_trans [OF norm_triangle_ineq])
+  thus False by simp
+qed
+
+text{* Uniqueness of the limit, when nontrivial. *}
+
+lemma Lim_unique:
+  fixes f :: "'a \<Rightarrow> 'b::t2_space"
+  assumes "\<not> trivial_limit net"  "(f ---> l) net"  "(f ---> l') net"
+  shows "l = l'"
+proof (rule ccontr)
+  assume "l \<noteq> l'"
+  obtain U V where "open U" "open V" "l \<in> U" "l' \<in> V" "U \<inter> V = {}"
+    using hausdorff [OF `l \<noteq> l'`] by fast
+  have "eventually (\<lambda>x. f x \<in> U) net"
+    using `(f ---> l) net` `open U` `l \<in> U` by (rule topological_tendstoD)
+  moreover
+  have "eventually (\<lambda>x. f x \<in> V) net"
+    using `(f ---> l') net` `open V` `l' \<in> V` by (rule topological_tendstoD)
+  ultimately
+  have "eventually (\<lambda>x. False) net"
+  proof (rule eventually_elim2)
+    fix x
+    assume "f x \<in> U" "f x \<in> V"
+    hence "f x \<in> U \<inter> V" by simp
+    with `U \<inter> V = {}` show "False" by simp
+  qed
+  with `\<not> trivial_limit net` show "False"
+    by (simp add: eventually_False)
+qed
+
+lemma tendsto_Lim:
+  fixes f :: "'a \<Rightarrow> 'b::t2_space"
+  shows "~(trivial_limit net) \<Longrightarrow> (f ---> l) net ==> Lim net f = l"
+  unfolding Lim_def using Lim_unique[of net f] by auto
+
+text{* Limit under bilinear function *}
+
+lemma Lim_bilinear:
+  assumes "(f ---> l) net" and "(g ---> m) net" and "bounded_bilinear h"
+  shows "((\<lambda>x. h (f x) (g x)) ---> (h l m)) net"
+using `bounded_bilinear h` `(f ---> l) net` `(g ---> m) net`
+by (rule bounded_bilinear.tendsto)
+
+text{* These are special for limits out of the same vector space. *}
+
+lemma Lim_within_id: "(id ---> a) (at a within s)"
+  unfolding tendsto_def Limits.eventually_within eventually_at_topological
+  by auto
+
+lemma Lim_at_id: "(id ---> a) (at a)"
+apply (subst within_UNIV[symmetric]) by (simp add: Lim_within_id)
+
+lemma Lim_at_zero:
+  fixes a :: "'a::real_normed_vector"
+  fixes l :: "'b::topological_space"
+  shows "(f ---> l) (at a) \<longleftrightarrow> ((\<lambda>x. f(a + x)) ---> l) (at 0)" (is "?lhs = ?rhs")
+proof
+  assume "?lhs"
+  { fix S assume "open S" "l \<in> S"
+    with `?lhs` have "eventually (\<lambda>x. f x \<in> S) (at a)"
+      by (rule topological_tendstoD)
+    then obtain d where d: "d>0" "\<forall>x. x \<noteq> a \<and> dist x a < d \<longrightarrow> f x \<in> S"
+      unfolding Limits.eventually_at by fast
+    { fix x::"'a" assume "x \<noteq> 0 \<and> dist x 0 < d"
+      hence "f (a + x) \<in> S" using d
+      apply(erule_tac x="x+a" in allE)
+      by(auto simp add: comm_monoid_add.mult_commute dist_norm dist_commute)
+    }
+    hence "\<exists>d>0. \<forall>x. x \<noteq> 0 \<and> dist x 0 < d \<longrightarrow> f (a + x) \<in> S"
+      using d(1) by auto
+    hence "eventually (\<lambda>x. f (a + x) \<in> S) (at 0)"
+      unfolding Limits.eventually_at .
+  }
+  thus "?rhs" by (rule topological_tendstoI)
+next
+  assume "?rhs"
+  { fix S assume "open S" "l \<in> S"
+    with `?rhs` have "eventually (\<lambda>x. f (a + x) \<in> S) (at 0)"
+      by (rule topological_tendstoD)
+    then obtain d where d: "d>0" "\<forall>x. x \<noteq> 0 \<and> dist x 0 < d \<longrightarrow> f (a + x) \<in> S"
+      unfolding Limits.eventually_at by fast
+    { fix x::"'a" assume "x \<noteq> a \<and> dist x a < d"
+      hence "f x \<in> S" using d apply(erule_tac x="x-a" in allE)
+	by(auto simp add: comm_monoid_add.mult_commute dist_norm dist_commute)
+    }
+    hence "\<exists>d>0. \<forall>x. x \<noteq> a \<and> dist x a < d \<longrightarrow> f x \<in> S" using d(1) by auto
+    hence "eventually (\<lambda>x. f x \<in> S) (at a)" unfolding Limits.eventually_at .
+  }
+  thus "?lhs" by (rule topological_tendstoI)
+qed
+
+text{* It's also sometimes useful to extract the limit point from the net.  *}
+
+definition
+  netlimit :: "'a::t2_space net \<Rightarrow> 'a" where
+  "netlimit net = (SOME a. ((\<lambda>x. x) ---> a) net)"
+
+lemma netlimit_within:
+  assumes "\<not> trivial_limit (at a within S)"
+  shows "netlimit (at a within S) = a"
+unfolding netlimit_def
+apply (rule some_equality)
+apply (rule Lim_at_within)
+apply (rule Lim_ident_at)
+apply (erule Lim_unique [OF assms])
+apply (rule Lim_at_within)
+apply (rule Lim_ident_at)
+done
+
+lemma netlimit_at:
+  fixes a :: "'a::perfect_space"
+  shows "netlimit (at a) = a"
+  apply (subst within_UNIV[symmetric])
+  using netlimit_within[of a UNIV]
+  by (simp add: trivial_limit_at within_UNIV)
+
+text{* Transformation of limit. *}
+
+lemma Lim_transform:
+  fixes f g :: "'a::type \<Rightarrow> 'b::real_normed_vector"
+  assumes "((\<lambda>x. f x - g x) ---> 0) net" "(f ---> l) net"
+  shows "(g ---> l) net"
+proof-
+  from assms have "((\<lambda>x. f x - g x - f x) ---> 0 - l) net" using Lim_sub[of "\<lambda>x. f x - g x" 0 net f l] by auto
+  thus "?thesis" using Lim_neg [of "\<lambda> x. - g x" "-l" net] by auto
+qed
+
+lemma Lim_transform_eventually:
+  "eventually (\<lambda>x. f x = g x) net \<Longrightarrow> (f ---> l) net ==> (g ---> l) net"
+  apply (rule topological_tendstoI)
+  apply (drule (2) topological_tendstoD)
+  apply (erule (1) eventually_elim2, simp)
+  done
+
+lemma Lim_transform_within:
+  fixes l :: "'b::metric_space" (* TODO: generalize *)
+  assumes "0 < d" "(\<forall>x'\<in>S. 0 < dist x' x \<and> dist x' x < d \<longrightarrow> f x' = g x')"
+          "(f ---> l) (at x within S)"
+  shows   "(g ---> l) (at x within S)"
+  using assms(1,3) unfolding Lim_within
+  apply -
+  apply (clarify, rename_tac e)
+  apply (drule_tac x=e in spec, clarsimp, rename_tac r)
+  apply (rule_tac x="min d r" in exI, clarsimp, rename_tac y)
+  apply (drule_tac x=y in bspec, assumption, clarsimp)
+  apply (simp add: assms(2))
+  done
+
+lemma Lim_transform_at:
+  fixes l :: "'b::metric_space" (* TODO: generalize *)
+  shows "0 < d \<Longrightarrow> (\<forall>x'. 0 < dist x' x \<and> dist x' x < d \<longrightarrow> f x' = g x') \<Longrightarrow>
+  (f ---> l) (at x) ==> (g ---> l) (at x)"
+  apply (subst within_UNIV[symmetric])
+  using Lim_transform_within[of d UNIV x f g l]
+  by (auto simp add: within_UNIV)
+
+text{* Common case assuming being away from some crucial point like 0. *}
+
+lemma Lim_transform_away_within:
+  fixes a b :: "'a::metric_space"
+  fixes l :: "'b::metric_space" (* TODO: generalize *)
+  assumes "a\<noteq>b" "\<forall>x\<in> S. x \<noteq> a \<and> x \<noteq> b \<longrightarrow> f x = g x"
+  and "(f ---> l) (at a within S)"
+  shows "(g ---> l) (at a within S)"
+proof-
+  have "\<forall>x'\<in>S. 0 < dist x' a \<and> dist x' a < dist a b \<longrightarrow> f x' = g x'" using assms(2)
+    apply auto apply(erule_tac x=x' in ballE) by (auto simp add: dist_commute)
+  thus ?thesis using Lim_transform_within[of "dist a b" S a f g l] using assms(1,3) unfolding dist_nz by auto
+qed
+
+lemma Lim_transform_away_at:
+  fixes a b :: "'a::metric_space"
+  fixes l :: "'b::metric_space" (* TODO: generalize *)
+  assumes ab: "a\<noteq>b" and fg: "\<forall>x. x \<noteq> a \<and> x \<noteq> b \<longrightarrow> f x = g x"
+  and fl: "(f ---> l) (at a)"
+  shows "(g ---> l) (at a)"
+  using Lim_transform_away_within[OF ab, of UNIV f g l] fg fl
+  by (auto simp add: within_UNIV)
+
+text{* Alternatively, within an open set. *}
+
+lemma Lim_transform_within_open:
+  fixes a :: "'a::metric_space"
+  fixes l :: "'b::metric_space" (* TODO: generalize *)
+  assumes "open S"  "a \<in> S"  "\<forall>x\<in>S. x \<noteq> a \<longrightarrow> f x = g x"  "(f ---> l) (at a)"
+  shows "(g ---> l) (at a)"
+proof-
+  from assms(1,2) obtain e::real where "e>0" and e:"ball a e \<subseteq> S" unfolding open_contains_ball by auto
+  hence "\<forall>x'. 0 < dist x' a \<and> dist x' a < e \<longrightarrow> f x' = g x'" using assms(3)
+    unfolding ball_def subset_eq apply auto apply(erule_tac x=x' in allE) apply(erule_tac x=x' in ballE) by(auto simp add: dist_commute)
+  thus ?thesis using Lim_transform_at[of e a f g l] `e>0` assms(4) by auto
+qed
+
+text{* A congruence rule allowing us to transform limits assuming not at point. *}
+
+(* FIXME: Only one congruence rule for tendsto can be used at a time! *)
+
+lemma Lim_cong_within[cong add]:
+  fixes a :: "'a::metric_space"
+  fixes l :: "'b::metric_space" (* TODO: generalize *)
+  shows "(\<And>x. x \<noteq> a \<Longrightarrow> f x = g x) ==> ((\<lambda>x. f x) ---> l) (at a within S) \<longleftrightarrow> ((g ---> l) (at a within S))"
+  by (simp add: Lim_within dist_nz[symmetric])
+
+lemma Lim_cong_at[cong add]:
+  fixes a :: "'a::metric_space"
+  fixes l :: "'b::metric_space" (* TODO: generalize *)
+  shows "(\<And>x. x \<noteq> a ==> f x = g x) ==> (((\<lambda>x. f x) ---> l) (at a) \<longleftrightarrow> ((g ---> l) (at a)))"
+  by (simp add: Lim_at dist_nz[symmetric])
+
+text{* Useful lemmas on closure and set of possible sequential limits.*}
+
+lemma closure_sequential:
+  fixes l :: "'a::metric_space" (* TODO: generalize *)
+  shows "l \<in> closure S \<longleftrightarrow> (\<exists>x. (\<forall>n. x n \<in> S) \<and> (x ---> l) sequentially)" (is "?lhs = ?rhs")
+proof
+  assume "?lhs" moreover
+  { assume "l \<in> S"
+    hence "?rhs" using Lim_const[of l sequentially] by auto
+  } moreover
+  { assume "l islimpt S"
+    hence "?rhs" unfolding islimpt_sequential by auto
+  } ultimately
+  show "?rhs" unfolding closure_def by auto
+next
+  assume "?rhs"
+  thus "?lhs" unfolding closure_def unfolding islimpt_sequential by auto
+qed
+
+lemma closed_sequential_limits:
+  fixes S :: "'a::metric_space set"
+  shows "closed S \<longleftrightarrow> (\<forall>x l. (\<forall>n. x n \<in> S) \<and> (x ---> l) sequentially \<longrightarrow> l \<in> S)"
+  unfolding closed_limpt
+  using closure_sequential [where 'a='a] closure_closed [where 'a='a] closed_limpt [where 'a='a] islimpt_sequential [where 'a='a] mem_delete [where 'a='a]
+  by metis
+
+lemma closure_approachable:
+  fixes S :: "'a::metric_space set"
+  shows "x \<in> closure S \<longleftrightarrow> (\<forall>e>0. \<exists>y\<in>S. dist y x < e)"
+  apply (auto simp add: closure_def islimpt_approachable)
+  by (metis dist_self)
+
+lemma closed_approachable:
+  fixes S :: "'a::metric_space set"
+  shows "closed S ==> (\<forall>e>0. \<exists>y\<in>S. dist y x < e) \<longleftrightarrow> x \<in> S"
+  by (metis closure_closed closure_approachable)
+
+text{* Some other lemmas about sequences. *}
+
+lemma seq_offset:
+  fixes l :: "'a::metric_space" (* TODO: generalize *)
+  shows "(f ---> l) sequentially ==> ((\<lambda>i. f( i + k)) ---> l) sequentially"
+  apply (auto simp add: Lim_sequentially)
+  by (metis trans_le_add1 )
+
+lemma seq_offset_neg:
+  "(f ---> l) sequentially ==> ((\<lambda>i. f(i - k)) ---> l) sequentially"
+  apply (rule topological_tendstoI)
+  apply (drule (2) topological_tendstoD)
+  apply (simp only: eventually_sequentially)
+  apply (subgoal_tac "\<And>N k (n::nat). N + k <= n ==> N <= n - k")
+  apply metis
+  by arith
+
+lemma seq_offset_rev:
+  "((\<lambda>i. f(i + k)) ---> l) sequentially ==> (f ---> l) sequentially"
+  apply (rule topological_tendstoI)
+  apply (drule (2) topological_tendstoD)
+  apply (simp only: eventually_sequentially)
+  apply (subgoal_tac "\<And>N k (n::nat). N + k <= n ==> N <= n - k \<and> (n - k) + k = n")
+  by metis arith
+
+lemma seq_harmonic: "((\<lambda>n. inverse (real n)) ---> 0) sequentially"
+proof-
+  { fix e::real assume "e>0"
+    hence "\<exists>N::nat. \<forall>n::nat\<ge>N. inverse (real n) < e"
+      using real_arch_inv[of e] apply auto apply(rule_tac x=n in exI)
+      by (metis not_le le_imp_inverse_le not_less real_of_nat_gt_zero_cancel_iff real_of_nat_less_iff xt1(7))
+  }
+  thus ?thesis unfolding Lim_sequentially dist_norm by simp
+qed
+
+text{* More properties of closed balls. *}
+
+lemma closed_cball: "closed (cball x e)"
+unfolding cball_def closed_def
+unfolding Collect_neg_eq [symmetric] not_le
+apply (clarsimp simp add: open_dist, rename_tac y)
+apply (rule_tac x="dist x y - e" in exI, clarsimp)
+apply (rename_tac x')
+apply (cut_tac x=x and y=x' and z=y in dist_triangle)
+apply simp
+done
+
+lemma open_contains_cball: "open S \<longleftrightarrow> (\<forall>x\<in>S. \<exists>e>0.  cball x e \<subseteq> S)"
+proof-
+  { fix x and e::real assume "x\<in>S" "e>0" "ball x e \<subseteq> S"
+    hence "\<exists>d>0. cball x d \<subseteq> S" unfolding subset_eq by (rule_tac x="e/2" in exI, auto)
+  } moreover
+  { fix x and e::real assume "x\<in>S" "e>0" "cball x e \<subseteq> S"
+    hence "\<exists>d>0. ball x d \<subseteq> S" unfolding subset_eq apply(rule_tac x="e/2" in exI) by auto
+  } ultimately
+  show ?thesis unfolding open_contains_ball by auto
+qed
+
+lemma open_contains_cball_eq: "open S ==> (\<forall>x. x \<in> S \<longleftrightarrow> (\<exists>e>0. cball x e \<subseteq> S))"
+  by (metis open_contains_cball subset_eq order_less_imp_le centre_in_cball mem_def)
+
+lemma mem_interior_cball: "x \<in> interior S \<longleftrightarrow> (\<exists>e>0. cball x e \<subseteq> S)"
+  apply (simp add: interior_def, safe)
+  apply (force simp add: open_contains_cball)
+  apply (rule_tac x="ball x e" in exI)
+  apply (simp add: open_ball centre_in_ball subset_trans [OF ball_subset_cball])
+  done
+
+lemma islimpt_ball:
+  fixes x y :: "'a::{real_normed_vector,perfect_space}"
+  shows "y islimpt ball x e \<longleftrightarrow> 0 < e \<and> y \<in> cball x e" (is "?lhs = ?rhs")
+proof
+  assume "?lhs"
+  { assume "e \<le> 0"
+    hence *:"ball x e = {}" using ball_eq_empty[of x e] by auto
+    have False using `?lhs` unfolding * using islimpt_EMPTY[of y] by auto
+  }
+  hence "e > 0" by (metis not_less)
+  moreover
+  have "y \<in> cball x e" using closed_cball[of x e] islimpt_subset[of y "ball x e" "cball x e"] ball_subset_cball[of x e] `?lhs` unfolding closed_limpt by auto
+  ultimately show "?rhs" by auto
+next
+  assume "?rhs" hence "e>0"  by auto
+  { fix d::real assume "d>0"
+    have "\<exists>x'\<in>ball x e. x' \<noteq> y \<and> dist x' y < d"
+    proof(cases "d \<le> dist x y")
+      case True thus "\<exists>x'\<in>ball x e. x' \<noteq> y \<and> dist x' y < d"
+      proof(cases "x=y")
+	case True hence False using `d \<le> dist x y` `d>0` by auto
+	thus "\<exists>x'\<in>ball x e. x' \<noteq> y \<and> dist x' y < d" by auto
+      next
+	case False
+
+	have "dist x (y - (d / (2 * dist y x)) *\<^sub>R (y - x))
+	      = norm (x - y + (d / (2 * norm (y - x))) *\<^sub>R (y - x))"
+	  unfolding mem_cball mem_ball dist_norm diff_diff_eq2 diff_add_eq[THEN sym] by auto
+	also have "\<dots> = \<bar>- 1 + d / (2 * norm (x - y))\<bar> * norm (x - y)"
+	  using scaleR_left_distrib[of "- 1" "d / (2 * norm (y - x))", THEN sym, of "y - x"]
+	  unfolding scaleR_minus_left scaleR_one
+	  by (auto simp add: norm_minus_commute)
+	also have "\<dots> = \<bar>- norm (x - y) + d / 2\<bar>"
+	  unfolding abs_mult_pos[of "norm (x - y)", OF norm_ge_zero[of "x - y"]]
+	  unfolding real_add_mult_distrib using `x\<noteq>y`[unfolded dist_nz, unfolded dist_norm] by auto
+	also have "\<dots> \<le> e - d/2" using `d \<le> dist x y` and `d>0` and `?rhs` by(auto simp add: dist_norm)
+	finally have "y - (d / (2 * dist y x)) *\<^sub>R (y - x) \<in> ball x e" using `d>0` by auto
+
+	moreover
+
+	have "(d / (2*dist y x)) *\<^sub>R (y - x) \<noteq> 0"
+	  using `x\<noteq>y`[unfolded dist_nz] `d>0` unfolding scaleR_eq_0_iff by (auto simp add: dist_commute)
+	moreover
+	have "dist (y - (d / (2 * dist y x)) *\<^sub>R (y - x)) y < d" unfolding dist_norm apply simp unfolding norm_minus_cancel
+	  using `d>0` `x\<noteq>y`[unfolded dist_nz] dist_commute[of x y]
+	  unfolding dist_norm by auto
+	ultimately show "\<exists>x'\<in>ball x e. x' \<noteq> y \<and> dist x' y < d" by (rule_tac  x="y - (d / (2*dist y x)) *\<^sub>R (y - x)" in bexI) auto
+      qed
+    next
+      case False hence "d > dist x y" by auto
+      show "\<exists>x'\<in>ball x e. x' \<noteq> y \<and> dist x' y < d"
+      proof(cases "x=y")
+	case True
+	obtain z where **: "z \<noteq> y" "dist z y < min e d"
+          using perfect_choose_dist[of "min e d" y]
+	  using `d > 0` `e>0` by auto
+	show "\<exists>x'\<in>ball x e. x' \<noteq> y \<and> dist x' y < d"
+          unfolding `x = y`
+          using `z \<noteq> y` **
+          by (rule_tac x=z in bexI, auto simp add: dist_commute)
+      next
+	case False thus "\<exists>x'\<in>ball x e. x' \<noteq> y \<and> dist x' y < d"
+	  using `d>0` `d > dist x y` `?rhs` by(rule_tac x=x in bexI, auto)
+      qed
+    qed  }
+  thus "?lhs" unfolding mem_cball islimpt_approachable mem_ball by auto
+qed
+
+lemma closure_ball_lemma:
+  fixes x y :: "'a::real_normed_vector"
+  assumes "x \<noteq> y" shows "y islimpt ball x (dist x y)"
+proof (rule islimptI)
+  fix T assume "y \<in> T" "open T"
+  then obtain r where "0 < r" "\<forall>z. dist z y < r \<longrightarrow> z \<in> T"
+    unfolding open_dist by fast
+  (* choose point between x and y, within distance r of y. *)
+  def k \<equiv> "min 1 (r / (2 * dist x y))"
+  def z \<equiv> "y + scaleR k (x - y)"
+  have z_def2: "z = x + scaleR (1 - k) (y - x)"
+    unfolding z_def by (simp add: algebra_simps)
+  have "dist z y < r"
+    unfolding z_def k_def using `0 < r`
+    by (simp add: dist_norm min_def)
+  hence "z \<in> T" using `\<forall>z. dist z y < r \<longrightarrow> z \<in> T` by simp
+  have "dist x z < dist x y"
+    unfolding z_def2 dist_norm
+    apply (simp add: norm_minus_commute)
+    apply (simp only: dist_norm [symmetric])
+    apply (subgoal_tac "\<bar>1 - k\<bar> * dist x y < 1 * dist x y", simp)
+    apply (rule mult_strict_right_mono)
+    apply (simp add: k_def divide_pos_pos zero_less_dist_iff `0 < r` `x \<noteq> y`)
+    apply (simp add: zero_less_dist_iff `x \<noteq> y`)
+    done
+  hence "z \<in> ball x (dist x y)" by simp
+  have "z \<noteq> y"
+    unfolding z_def k_def using `x \<noteq> y` `0 < r`
+    by (simp add: min_def)
+  show "\<exists>z\<in>ball x (dist x y). z \<in> T \<and> z \<noteq> y"
+    using `z \<in> ball x (dist x y)` `z \<in> T` `z \<noteq> y`
+    by fast
+qed
+
+lemma closure_ball:
+  fixes x :: "'a::real_normed_vector"
+  shows "0 < e \<Longrightarrow> closure (ball x e) = cball x e"
+apply (rule equalityI)
+apply (rule closure_minimal)
+apply (rule ball_subset_cball)
+apply (rule closed_cball)
+apply (rule subsetI, rename_tac y)
+apply (simp add: le_less [where 'a=real])
+apply (erule disjE)
+apply (rule subsetD [OF closure_subset], simp)
+apply (simp add: closure_def)
+apply clarify
+apply (rule closure_ball_lemma)
+apply (simp add: zero_less_dist_iff)
+done
+
+(* In a trivial vector space, this fails for e = 0. *)
+lemma interior_cball:
+  fixes x :: "'a::{real_normed_vector, perfect_space}"
+  shows "interior (cball x e) = ball x e"
+proof(cases "e\<ge>0")
+  case False note cs = this
+  from cs have "ball x e = {}" using ball_empty[of e x] by auto moreover
+  { fix y assume "y \<in> cball x e"
+    hence False unfolding mem_cball using dist_nz[of x y] cs by auto  }
+  hence "cball x e = {}" by auto
+  hence "interior (cball x e) = {}" using interior_empty by auto
+  ultimately show ?thesis by blast
+next
+  case True note cs = this
+  have "ball x e \<subseteq> cball x e" using ball_subset_cball by auto moreover
+  { fix S y assume as: "S \<subseteq> cball x e" "open S" "y\<in>S"
+    then obtain d where "d>0" and d:"\<forall>x'. dist x' y < d \<longrightarrow> x' \<in> S" unfolding open_dist by blast
+
+    then obtain xa where xa_y: "xa \<noteq> y" and xa: "dist xa y < d"
+      using perfect_choose_dist [of d] by auto
+    have "xa\<in>S" using d[THEN spec[where x=xa]] using xa by(auto simp add: dist_commute)
+    hence xa_cball:"xa \<in> cball x e" using as(1) by auto
+
+    hence "y \<in> ball x e" proof(cases "x = y")
+      case True
+      hence "e>0" using xa_y[unfolded dist_nz] xa_cball[unfolded mem_cball] by (auto simp add: dist_commute)
+      thus "y \<in> ball x e" using `x = y ` by simp
+    next
+      case False
+      have "dist (y + (d / 2 / dist y x) *\<^sub>R (y - x)) y < d" unfolding dist_norm
+	using `d>0` norm_ge_zero[of "y - x"] `x \<noteq> y` by auto
+      hence *:"y + (d / 2 / dist y x) *\<^sub>R (y - x) \<in> cball x e" using d as(1)[unfolded subset_eq] by blast
+      have "y - x \<noteq> 0" using `x \<noteq> y` by auto
+      hence **:"d / (2 * norm (y - x)) > 0" unfolding zero_less_norm_iff[THEN sym]
+	using `d>0` divide_pos_pos[of d "2*norm (y - x)"] by auto
+
+      have "dist (y + (d / 2 / dist y x) *\<^sub>R (y - x)) x = norm (y + (d / (2 * norm (y - x))) *\<^sub>R y - (d / (2 * norm (y - x))) *\<^sub>R x - x)"
+        by (auto simp add: dist_norm algebra_simps)
+      also have "\<dots> = norm ((1 + d / (2 * norm (y - x))) *\<^sub>R (y - x))"
+        by (auto simp add: algebra_simps)
+      also have "\<dots> = \<bar>1 + d / (2 * norm (y - x))\<bar> * norm (y - x)"
+        using ** by auto
+      also have "\<dots> = (dist y x) + d/2"using ** by (auto simp add: left_distrib dist_norm)
+      finally have "e \<ge> dist x y +d/2" using *[unfolded mem_cball] by (auto simp add: dist_commute)
+      thus "y \<in> ball x e" unfolding mem_ball using `d>0` by auto
+    qed  }
+  hence "\<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
+qed
+
+lemma frontier_ball:
+  fixes a :: "'a::real_normed_vector"
+  shows "0 < e ==> frontier(ball a e) = {x. dist a x = e}"
+  apply (simp add: frontier_def closure_ball interior_open open_ball order_less_imp_le)
+  apply (simp add: expand_set_eq)
+  by arith
+
+lemma frontier_cball:
+  fixes a :: "'a::{real_normed_vector, perfect_space}"
+  shows "frontier(cball a e) = {x. dist a x = e}"
+  apply (simp add: frontier_def interior_cball closed_cball closure_closed order_less_imp_le)
+  apply (simp add: expand_set_eq)
+  by arith
+
+lemma cball_eq_empty: "(cball x e = {}) \<longleftrightarrow> e < 0"
+  apply (simp add: expand_set_eq not_le)
+  by (metis zero_le_dist dist_self order_less_le_trans)
+lemma cball_empty: "e < 0 ==> cball x e = {}" by (simp add: cball_eq_empty)
+
+lemma cball_eq_sing:
+  fixes x :: "'a::perfect_space"
+  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"
+    using perfect_choose_dist [OF e] by auto
+  hence "a \<noteq> x" "dist x a \<le> e" by (auto simp add: dist_commute)
+  with e show ?thesis by (auto simp add: expand_set_eq)
+qed auto
+
+lemma cball_sing:
+  fixes x :: "'a::metric_space"
+  shows "e = 0 ==> cball x e = {x}"
+  by (auto simp add: expand_set_eq)
+
+text{* For points in the interior, localization of limits makes no difference.   *}
+
+lemma eventually_within_interior:
+  assumes "x \<in> interior S"
+  shows "eventually P (at x within S) \<longleftrightarrow> eventually P (at x)" (is "?lhs = ?rhs")
+proof-
+  from assms obtain T where T: "open T" "x \<in> T" "T \<subseteq> S"
+    unfolding interior_def by fast
+  { 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"
+      unfolding Limits.eventually_within Limits.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"
+      by auto
+    then have "?rhs"
+      unfolding Limits.eventually_at_topological by auto
+  } moreover
+  { assume "?rhs" hence "?lhs"
+      unfolding Limits.eventually_within
+      by (auto elim: eventually_elim1)
+  } ultimately
+  show "?thesis" ..
+qed
+
+lemma lim_within_interior:
+  "x \<in> interior S \<Longrightarrow> (f ---> l) (at x within S) \<longleftrightarrow> (f ---> l) (at x)"
+  unfolding tendsto_def by (simp add: eventually_within_interior)
+
+lemma netlimit_within_interior:
+  fixes x :: "'a::{perfect_space, real_normed_vector}"
+    (* FIXME: generalize to perfect_space *)
+  assumes "x \<in> interior S"
+  shows "netlimit(at x within S) = x" (is "?lhs = ?rhs")
+proof-
+  from assms obtain e::real where e:"e>0" "ball x e \<subseteq> S" using open_interior[of S] unfolding open_contains_ball using interior_subset[of S] by auto
+  hence "\<not> trivial_limit (at x within S)" using islimpt_subset[of x "ball x e" S] unfolding trivial_limit_within islimpt_ball centre_in_cball by auto
+  thus ?thesis using netlimit_within by auto
+qed
+
+subsection{* Boundedness. *}
+
+  (* FIXME: This has to be unified with BSEQ!! *)
+definition
+  bounded :: "'a::metric_space set \<Rightarrow> bool" where
+  "bounded S \<longleftrightarrow> (\<exists>x e. \<forall>y\<in>S. dist x y \<le> e)"
+
+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 (drule (1) bspec)
+apply (erule order_trans [OF dist_triangle add_left_mono])
+apply auto
+done
+
+lemma bounded_iff: "bounded S \<longleftrightarrow> (\<exists>a. \<forall>x\<in>S. norm x \<le> a)"
+unfolding bounded_any_center [where a=0]
+by (simp add: dist_norm)
+
+lemma bounded_empty[simp]: "bounded {}" by (simp add: bounded_def)
+lemma bounded_subset: "bounded T \<Longrightarrow> S \<subseteq> T ==> bounded S"
+  by (metis bounded_def subset_eq)
+
+lemma bounded_interior[intro]: "bounded S ==> bounded(interior S)"
+  by (metis bounded_subset interior_subset)
+
+lemma bounded_closure[intro]: assumes "bounded S" shows "bounded(closure S)"
+proof-
+  from assms obtain x and a where a: "\<forall>y\<in>S. dist x y \<le> a" unfolding bounded_def by auto
+  { fix y assume "y \<in> closure S"
+    then obtain f where f: "\<forall>n. f n \<in> S"  "(f ---> y) sequentially"
+      unfolding closure_sequential by auto
+    have "\<forall>n. f n \<in> S \<longrightarrow> dist x (f n) \<le> a" using a by simp
+    hence "eventually (\<lambda>n. dist x (f n) \<le> a) sequentially"
+      by (rule eventually_mono, simp add: f(1))
+    have "dist x y \<le> a"
+      apply (rule Lim_dist_ubound [of sequentially f])
+      apply (rule trivial_limit_sequentially)
+      apply (rule f(2))
+      apply fact
+      done
+  }
+  thus ?thesis unfolding bounded_def by auto
+qed
+
+lemma bounded_cball[simp,intro]: "bounded (cball x e)"
+  apply (simp add: bounded_def)
+  apply (rule_tac x=x in exI)
+  apply (rule_tac x=e in exI)
+  apply auto
+  done
+
+lemma bounded_ball[simp,intro]: "bounded(ball x e)"
+  by (metis ball_subset_cball bounded_cball bounded_subset)
+
+lemma finite_imp_bounded[intro]: assumes "finite S" shows "bounded S"
+proof-
+  { fix a F assume as:"bounded F"
+    then obtain x e where "\<forall>y\<in>F. dist x y \<le> e" unfolding bounded_def by auto
+    hence "\<forall>y\<in>(insert a F). dist x y \<le> max e (dist x a)" by auto
+    hence "bounded (insert a F)" unfolding bounded_def by (intro exI)
+  }
+  thus ?thesis using finite_induct[of S bounded]  using bounded_empty assms by auto
+qed
+
+lemma bounded_Un[simp]: "bounded (S \<union> T) \<longleftrightarrow> bounded S \<and> bounded T"
+  apply (auto simp add: bounded_def)
+  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 (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)"
+  by (induct rule: finite_induct[of F], auto)
+
+lemma bounded_pos: "bounded S \<longleftrightarrow> (\<exists>b>0. \<forall>x\<in> S. norm x <= 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)")
+  by metis arith
+
+lemma bounded_Int[intro]: "bounded S \<or> bounded T \<Longrightarrow> bounded (S \<inter> T)"
+  by (metis Int_lower1 Int_lower2 bounded_subset)
+
+lemma bounded_diff[intro]: "bounded S ==> bounded (S - T)"
+apply (metis Diff_subset bounded_subset)
+done
+
+lemma bounded_insert[intro]:"bounded(insert x S) \<longleftrightarrow> bounded S"
+  by (metis Diff_cancel Un_empty_right Un_insert_right bounded_Un bounded_subset finite.emptyI finite_imp_bounded infinite_remove subset_insertI)
+
+lemma not_bounded_UNIV[simp, intro]:
+  "\<not> bounded (UNIV :: 'a::{real_normed_vector, perfect_space} set)"
+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
+  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" ..
+qed
+
+lemma bounded_linear_image:
+  assumes "bounded S" "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" unfolding bounded_pos by auto
+  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"
+    hence "norm x \<le> b" using b by auto
+    hence "norm (f x) \<le> B * b" using B(2) apply(erule_tac x=x in allE)
+      by (metis B(1) B(2) real_le_trans real_mult_le_cancel_iff2)
+  }
+  thus ?thesis unfolding bounded_pos apply(rule_tac x="b*B" in exI)
+    using b B real_mult_order[of b B] by (auto simp add: real_mult_commute)
+qed
+
+lemma bounded_scaling:
+  fixes S :: "'a::real_normed_vector set"
+  shows "bounded S \<Longrightarrow> bounded ((\<lambda>x. c *\<^sub>R x) ` S)"
+  apply (rule bounded_linear_image, assumption)
+  apply (rule scaleR.bounded_linear_right)
+  done
+
+lemma bounded_translation:
+  fixes S :: "'a::real_normed_vector set"
+  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" unfolding bounded_pos by auto
+  { fix x assume "x\<in>S"
+    hence "norm (a + x) \<le> b + norm a" using norm_triangle_ineq[of a x] b by auto
+  }
+  thus ?thesis unfolding bounded_pos using norm_ge_zero[of a] b(1) using add_strict_increasing[of b 0 "norm a"]
+    by (auto intro!: add exI[of _ "b + norm a"])
+qed
+
+
+text{* Some theorems on sups and infs using the notion "bounded". *}
+
+lemma bounded_real:
+  fixes S :: "real set"
+  shows "bounded S \<longleftrightarrow>  (\<exists>a. \<forall>x\<in>S. abs x <= a)"
+  by (simp add: bounded_iff)
+
+lemma bounded_has_rsup: assumes "bounded S" "S \<noteq> {}"
+  shows "\<forall>x\<in>S. x <= rsup S" and "\<forall>b. (\<forall>x\<in>S. x <= b) \<longrightarrow> rsup S <= b"
+proof
+  fix x 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
+  hence *:"S *<= a" using setleI[of S a] by (metis abs_le_interval_iff mem_def)
+  thus "x \<le> rsup S" using rsup[OF `S\<noteq>{}`] using assms(1)[unfolded bounded_real] using isLubD2[of UNIV S "rsup S" x] using `x\<in>S` by auto
+next
+  show "\<forall>b. (\<forall>x\<in>S. x \<le> b) \<longrightarrow> rsup S \<le> b" using assms
+  using rsup[of S, unfolded isLub_def isUb_def leastP_def setle_def setge_def]
+  apply (auto simp add: bounded_real)
+  by (auto simp add: isLub_def isUb_def leastP_def setle_def setge_def)
+qed
+
+lemma rsup_insert: assumes "bounded S"
+  shows "rsup(insert x S) = (if S = {} then x else max x (rsup S))"
+proof(cases "S={}")
+  case True thus ?thesis using rsup_finite_in[of "{x}"] by auto
+next
+  let ?S = "insert x S"
+  case False
+  hence *:"\<forall>x\<in>S. x \<le> rsup S" using bounded_has_rsup(1)[of S] using assms by auto
+  hence "insert x S *<= max x (rsup S)" unfolding setle_def by auto
+  hence "isLub UNIV ?S (rsup ?S)" using rsup[of ?S] by auto
+  moreover
+  have **:"isUb UNIV ?S (max x (rsup S))" unfolding isUb_def setle_def using * by auto
+  { fix y assume as:"isUb UNIV (insert x S) y"
+    hence "max x (rsup S) \<le> y" unfolding isUb_def using rsup_le[OF `S\<noteq>{}`]
+      unfolding setle_def by auto  }
+  hence "max x (rsup S) <=* isUb UNIV (insert x S)" unfolding setge_def Ball_def mem_def by auto
+  hence "isLub UNIV ?S (max x (rsup S))" using ** isLubI2[of UNIV ?S "max x (rsup S)"] unfolding Collect_def by auto
+  ultimately show ?thesis using real_isLub_unique[of UNIV ?S] using `S\<noteq>{}` by auto
+qed
+
+lemma sup_insert_finite: "finite S \<Longrightarrow> rsup(insert x S) = (if S = {} then x else max x (rsup S))"
+  apply (rule rsup_insert)
+  apply (rule finite_imp_bounded)
+  by simp
+
+lemma bounded_has_rinf:
+  assumes "bounded S"  "S \<noteq> {}"
+  shows "\<forall>x\<in>S. x >= rinf S" and "\<forall>b. (\<forall>x\<in>S. x >= b) \<longrightarrow> rinf S >= b"
+proof
+  fix x 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
+  hence *:"- a <=* S" using setgeI[of S "-a"] unfolding abs_le_interval_iff by auto
+  thus "x \<ge> rinf S" using rinf[OF `S\<noteq>{}`] using isGlbD2[of UNIV S "rinf S" x] using `x\<in>S` by auto
+next
+  show "\<forall>b. (\<forall>x\<in>S. x >= b) \<longrightarrow> rinf S \<ge> b" using assms
+  using rinf[of S, unfolded isGlb_def isLb_def greatestP_def setle_def setge_def]
+  apply (auto simp add: bounded_real)
+  by (auto simp add: isGlb_def isLb_def greatestP_def setle_def setge_def)
+qed
+
+(* TODO: Move this to RComplete.thy -- would need to include Glb into RComplete *)
+lemma real_isGlb_unique: "[| isGlb R S x; isGlb R S y |] ==> x = (y::real)"
+  apply (frule isGlb_isLb)
+  apply (frule_tac x = y in isGlb_isLb)
+  apply (blast intro!: order_antisym dest!: isGlb_le_isLb)
+  done
+
+lemma rinf_insert: assumes "bounded S"
+  shows "rinf(insert x S) = (if S = {} then x else min x (rinf S))" (is "?lhs = ?rhs")
+proof(cases "S={}")
+  case True thus ?thesis using rinf_finite_in[of "{x}"] by auto
+next
+  let ?S = "insert x S"
+  case False
+  hence *:"\<forall>x\<in>S. x \<ge> rinf S" using bounded_has_rinf(1)[of S] using assms by auto
+  hence "min x (rinf S) <=* insert x S" unfolding setge_def by auto
+  hence "isGlb UNIV ?S (rinf ?S)" using rinf[of ?S] by auto
+  moreover
+  have **:"isLb UNIV ?S (min x (rinf S))" unfolding isLb_def setge_def using * by auto
+  { fix y assume as:"isLb UNIV (insert x S) y"
+    hence "min x (rinf S) \<ge> y" unfolding isLb_def using rinf_ge[OF `S\<noteq>{}`]
+      unfolding setge_def by auto  }
+  hence "isLb UNIV (insert x S) *<= min x (rinf S)" unfolding setle_def Ball_def mem_def by auto
+  hence "isGlb UNIV ?S (min x (rinf S))" using ** isGlbI2[of UNIV ?S "min x (rinf S)"] unfolding Collect_def by auto
+  ultimately show ?thesis using real_isGlb_unique[of UNIV ?S] using `S\<noteq>{}` by auto
+qed
+
+lemma inf_insert_finite: "finite S ==> rinf(insert x S) = (if S = {} then x else min x (rinf S))"
+  by (rule rinf_insert, rule finite_imp_bounded, simp)
+
+subsection{* Compactness (the definition is the one based on convegent subsequences). *}
+
+definition
+  compact :: "'a::metric_space set \<Rightarrow> bool" where (* TODO: generalize *)
+  "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))"
+
+text {*
+  A metric space (or topological vector space) is said to have the
+  Heine-Borel property if every closed and bounded subset is compact.
+*}
+
+class heine_borel =
+  assumes bounded_imp_convergent_subsequence:
+    "bounded s \<Longrightarrow> \<forall>n. f n \<in> s
+      \<Longrightarrow> \<exists>l r. subseq r \<and> ((f \<circ> r) ---> l) sequentially"
+
+lemma bounded_closed_imp_compact:
+  fixes s::"'a::heine_borel set"
+  assumes "bounded s" and "closed s" shows "compact s"
+proof (unfold compact_def, clarify)
+  fix f :: "nat \<Rightarrow> 'a" assume f: "\<forall>n. f n \<in> s"
+  obtain l r where r: "subseq r" and l: "((f \<circ> r) ---> l) sequentially"
+    using bounded_imp_convergent_subsequence [OF `bounded s` `\<forall>n. f n \<in> s`] by auto
+  from f have fr: "\<forall>n. (f \<circ> r) n \<in> s" by simp
+  have "l \<in> s" using `closed s` fr l
+    unfolding closed_sequential_limits by blast
+  show "\<exists>l\<in>s. \<exists>r. subseq r \<and> ((f \<circ> r) ---> l) sequentially"
+    using `l \<in> s` r l by blast
+qed
+
+lemma subseq_bigger: assumes "subseq r" shows "n \<le> r n"
+proof(induct n)
+  show "0 \<le> r 0" by auto
+next
+  fix n assume "n \<le> r n"
+  moreover have "r n < r (Suc n)"
+    using assms [unfolded subseq_def] by auto
+  ultimately show "Suc n \<le> r (Suc n)" by auto
+qed
+
+lemma eventually_subseq:
+  assumes r: "subseq r"
+  shows "eventually P sequentially \<Longrightarrow> eventually (\<lambda>n. P (r n)) sequentially"
+unfolding eventually_sequentially
+by (metis subseq_bigger [OF r] le_trans)
+
+lemma lim_subseq:
+  "subseq r \<Longrightarrow> (s ---> l) sequentially \<Longrightarrow> ((s o r) ---> l) sequentially"
+unfolding tendsto_def eventually_sequentially o_def
+by (metis subseq_bigger le_trans)
+
+lemma num_Axiom: "EX! g. g 0 = e \<and> (\<forall>n. g (Suc n) = f n (g n))"
+  unfolding Ex1_def
+  apply (rule_tac x="nat_rec e f" in exI)
+  apply (rule conjI)+
+apply (rule def_nat_rec_0, simp)
+apply (rule allI, rule def_nat_rec_Suc, simp)
+apply (rule allI, rule impI, rule ext)
+apply (erule conjE)
+apply (induct_tac x)
+apply (simp add: nat_rec_0)
+apply (erule_tac x="n" in allE)
+apply (simp)
+done
+
+lemma convergent_bounded_increasing: fixes s ::"nat\<Rightarrow>real"
+  assumes "incseq s" and "\<forall>n. abs(s n) \<le> b"
+  shows "\<exists> l. \<forall>e::real>0. \<exists> N. \<forall>n \<ge> N.  abs(s n - l) < e"
+proof-
+  have "isUb UNIV (range s) b" using assms(2) and abs_le_D1 unfolding isUb_def and setle_def by auto
+  then obtain t where t:"isLub UNIV (range s) t" using reals_complete[of "range s" ] by auto
+  { fix e::real assume "e>0" and as:"\<forall>N. \<exists>n\<ge>N. \<not> \<bar>s n - t\<bar> < e"
+    { fix n::nat
+      obtain N where "N\<ge>n" and n:"\<bar>s N - t\<bar> \<ge> e" using as[THEN spec[where x=n]] by auto
+      have "t \<ge> s N" using isLub_isUb[OF t, unfolded isUb_def setle_def] by auto
+      with n have "s N \<le> t - e" using `e>0` by auto
+      hence "s n \<le> t - e" using assms(1)[unfolded incseq_def, THEN spec[where x=n], THEN spec[where x=N]] using `n\<le>N` by auto  }
+    hence "isUb UNIV (range s) (t - e)" unfolding isUb_def and setle_def by auto
+    hence False using isLub_le_isUb[OF t, of "t - e"] and `e>0` by auto  }
+  thus ?thesis by blast
+qed
+
+lemma convergent_bounded_monotone: fixes s::"nat \<Rightarrow> real"
+  assumes "\<forall>n. abs(s n) \<le> b" and "monoseq s"
+  shows "\<exists>l. \<forall>e::real>0. \<exists>N. \<forall>n\<ge>N. abs(s n - l) < e"
+  using convergent_bounded_increasing[of s b] assms using convergent_bounded_increasing[of "\<lambda>n. - s n" b]
+  unfolding monoseq_def incseq_def
+  apply auto unfolding minus_add_distrib[THEN sym, unfolded diff_minus[THEN sym]]
+  unfolding abs_minus_cancel by(rule_tac x="-l" in exI)auto
+
+lemma compact_real_lemma:
+  assumes "\<forall>n::nat. abs(s n) \<le> b"
+  shows "\<exists>(l::real) r. subseq r \<and> ((s \<circ> r) ---> l) sequentially"
+proof-
+  obtain r where r:"subseq r" "monoseq (\<lambda>n. s (r n))"
+    using seq_monosub[of s] by auto
+  thus ?thesis using convergent_bounded_monotone[of "\<lambda>n. s (r n)" b] and assms
+    unfolding tendsto_iff dist_norm eventually_sequentially by auto
+qed
+
+instance real :: heine_borel
+proof
+  fix s :: "real set" and f :: "nat \<Rightarrow> real"
+  assume s: "bounded s" and f: "\<forall>n. f n \<in> s"
+  then obtain b where b: "\<forall>n. abs (f n) \<le> b"
+    unfolding bounded_iff by auto
+  obtain l :: real and r :: "nat \<Rightarrow> nat" where
+    r: "subseq r" and l: "((f \<circ> r) ---> l) sequentially"
+    using compact_real_lemma [OF b] by auto
+  thus "\<exists>l r. subseq r \<and> ((f \<circ> r) ---> l) sequentially"
+    by auto
+qed
+
+lemma bounded_component: "bounded s \<Longrightarrow> bounded ((\<lambda>x. x $ i) ` s)"
+unfolding bounded_def
+apply clarify
+apply (rule_tac x="x $ i" in exI)
+apply (rule_tac x="e" in exI)
+apply clarify
+apply (rule order_trans [OF dist_nth_le], simp)
+done
+
+lemma compact_lemma:
+  fixes f :: "nat \<Rightarrow> 'a::heine_borel ^ 'n::finite"
+  assumes "bounded s" and "\<forall>n. f n \<in> s"
+  shows "\<forall>d.
+        \<exists>l r. subseq r \<and>
+        (\<forall>e>0. eventually (\<lambda>n. \<forall>i\<in>d. dist (f (r n) $ i) (l $ i) < e) sequentially)"
+proof
+  fix d::"'n set" have "finite d" by simp
+  thus "\<exists>l::'a ^ 'n. \<exists>r. subseq r \<and>
+      (\<forall>e>0. eventually (\<lambda>n. \<forall>i\<in>d. dist (f (r n) $ i) (l $ i) < e) sequentially)"
+  proof(induct d) case empty thus ?case unfolding subseq_def by auto
+  next case (insert k d)
+    have s': "bounded ((\<lambda>x. x $ k) ` s)" using `bounded s` by (rule bounded_component)
+    obtain l1::"'a^'n" and r1 where r1:"subseq r1" and lr1:"\<forall>e>0. eventually (\<lambda>n. \<forall>i\<in>d. dist (f (r1 n) $ i) (l1 $ i) < e) sequentially"
+      using insert(3) by auto
+    have f': "\<forall>n. f (r1 n) $ k \<in> (\<lambda>x. x $ k) ` s" using `\<forall>n. f n \<in> s` by simp
+    obtain l2 r2 where r2:"subseq r2" and lr2:"((\<lambda>i. f (r1 (r2 i)) $ k) ---> l2) sequentially"
+      using bounded_imp_convergent_subsequence[OF s' f'] unfolding o_def by auto
+    def r \<equiv> "r1 \<circ> r2" have r:"subseq r"
+      using r1 and r2 unfolding r_def o_def subseq_def by auto
+    moreover
+    def l \<equiv> "(\<chi> i. if i = k then l2 else l1$i)::'a^'n"
+    { fix e::real assume "e>0"
+      from lr1 `e>0` have N1:"eventually (\<lambda>n. \<forall>i\<in>d. dist (f (r1 n) $ i) (l1 $ i) < e) sequentially" by blast
+      from lr2 `e>0` have N2:"eventually (\<lambda>n. dist (f (r1 (r2 n)) $ k) l2 < e) sequentially" by (rule tendstoD)
+      from r2 N1 have N1': "eventually (\<lambda>n. \<forall>i\<in>d. dist (f (r1 (r2 n)) $ i) (l1 $ i) < e) sequentially"
+        by (rule eventually_subseq)
+      have "eventually (\<lambda>n. \<forall>i\<in>(insert k d). dist (f (r n) $ i) (l $ i) < e) sequentially"
+        using N1' N2 by (rule eventually_elim2, simp add: l_def r_def)
+    }
+    ultimately show ?case by auto
+  qed
+qed
+
+instance "^" :: (heine_borel, finite) heine_borel
+proof
+  fix s :: "('a ^ 'b) set" and f :: "nat \<Rightarrow> 'a ^ 'b"
+  assume s: "bounded s" and f: "\<forall>n. f n \<in> s"
+  then obtain l r where r: "subseq r"
+    and l: "\<forall>e>0. eventually (\<lambda>n. \<forall>i\<in>UNIV. dist (f (r n) $ i) (l $ i) < e) sequentially"
+    using compact_lemma [OF s f] by blast
+  let ?d = "UNIV::'b set"
+  { fix e::real assume "e>0"
+    hence "0 < e / (real_of_nat (card ?d))"
+      using zero_less_card_finite using divide_pos_pos[of e, of "real_of_nat (card ?d)"] by auto
+    with l have "eventually (\<lambda>n. \<forall>i. dist (f (r n) $ i) (l $ i) < e / (real_of_nat (card ?d))) sequentially"
+      by simp
+    moreover
+    { fix n assume n: "\<forall>i. dist (f (r n) $ i) (l $ i) < e / (real_of_nat (card ?d))"
+      have "dist (f (r n)) l \<le> (\<Sum>i\<in>?d. dist (f (r n) $ i) (l $ i))"
+        unfolding dist_vector_def using zero_le_dist by (rule setL2_le_setsum)
+      also have "\<dots> < (\<Sum>i\<in>?d. e / (real_of_nat (card ?d)))"
+        by (rule setsum_strict_mono) (simp_all add: n)
+      finally have "dist (f (r n)) l < e" by simp
+    }
+    ultimately have "eventually (\<lambda>n. dist (f (r n)) l < e) sequentially"
+      by (rule eventually_elim1)
+  }
+  hence *:"((f \<circ> r) ---> l) sequentially" unfolding o_def tendsto_iff by simp
+  with r show "\<exists>l r. subseq r \<and> ((f \<circ> r) ---> l) sequentially" by auto
+qed
+
+lemma bounded_fst: "bounded s \<Longrightarrow> bounded (fst ` s)"
+unfolding bounded_def
+apply clarify
+apply (rule_tac x="a" in exI)
+apply (rule_tac x="e" in exI)
+apply clarsimp
+apply (drule (1) bspec)
+apply (simp add: dist_Pair_Pair)
+apply (erule order_trans [OF real_sqrt_sum_squares_ge1])
+done
+
+lemma bounded_snd: "bounded s \<Longrightarrow> bounded (snd ` s)"
+unfolding bounded_def
+apply clarify
+apply (rule_tac x="b" in exI)
+apply (rule_tac x="e" in exI)
+apply clarsimp
+apply (drule (1) bspec)
+apply (simp add: dist_Pair_Pair)
+apply (erule order_trans [OF real_sqrt_sum_squares_ge2])
+done
+
+instance "*" :: (heine_borel, heine_borel) heine_borel
+proof
+  fix s :: "('a * 'b) set" and f :: "nat \<Rightarrow> 'a * 'b"
+  assume s: "bounded s" and f: "\<forall>n. f n \<in> s"
+  from s have s1: "bounded (fst ` s)" by (rule bounded_fst)
+  from f have f1: "\<forall>n. fst (f n) \<in> fst ` s" by simp
+  obtain l1 r1 where r1: "subseq r1"
+    and l1: "((\<lambda>n. fst (f (r1 n))) ---> l1) sequentially"
+    using bounded_imp_convergent_subsequence [OF s1 f1]
+    unfolding o_def by fast
+  from s have s2: "bounded (snd ` s)" by (rule bounded_snd)
+  from f have f2: "\<forall>n. snd (f (r1 n)) \<in> snd ` s" by simp
+  obtain l2 r2 where r2: "subseq r2"
+    and l2: "((\<lambda>n. snd (f (r1 (r2 n)))) ---> l2) sequentially"
+    using bounded_imp_convergent_subsequence [OF s2 f2]
+    unfolding o_def by fast
+  have l1': "((\<lambda>n. fst (f (r1 (r2 n)))) ---> l1) sequentially"
+    using lim_subseq [OF r2 l1] unfolding o_def .
+  have l: "((f \<circ> (r1 \<circ> r2)) ---> (l1, l2)) sequentially"
+    using tendsto_Pair [OF l1' l2] unfolding o_def by simp
+  have r: "subseq (r1 \<circ> r2)"
+    using r1 r2 unfolding subseq_def by simp
+  show "\<exists>l r. subseq r \<and> ((f \<circ> r) ---> l) sequentially"
+    using l r by fast
+qed
+
+subsection{* Completeness. *}
+
+lemma cauchy_def:
+  "Cauchy s \<longleftrightarrow> (\<forall>e>0. \<exists>N. \<forall>m n. m \<ge> N \<and> n \<ge> N --> dist(s m)(s n) < e)"
+unfolding Cauchy_def by blast
+
+definition
+  complete :: "'a::metric_space set \<Rightarrow> bool" where
+  "complete s \<longleftrightarrow> (\<forall>f. (\<forall>n. f n \<in> s) \<and> Cauchy f
+                      --> (\<exists>l \<in> s. (f ---> l) sequentially))"
+
+lemma cauchy: "Cauchy s \<longleftrightarrow> (\<forall>e>0.\<exists> N::nat. \<forall>n\<ge>N. dist(s n)(s N) < e)" (is "?lhs = ?rhs")
+proof-
+  { assume ?rhs
+    { fix e::real
+      assume "e>0"
+      with `?rhs` obtain N where N:"\<forall>n\<ge>N. dist (s n) (s N) < e/2"
+	by (erule_tac x="e/2" in allE) auto
+      { fix n m
+	assume nm:"N \<le> m \<and> N \<le> n"
+	hence "dist (s m) (s n) < e" using N
+	  using dist_triangle_half_l[of "s m" "s N" "e" "s n"]
+	  by blast
+      }
+      hence "\<exists>N. \<forall>m n. N \<le> m \<and> N \<le> n \<longrightarrow> dist (s m) (s n) < e"
+	by blast
+    }
+    hence ?lhs
+      unfolding cauchy_def
+      by blast
+  }
+  thus ?thesis
+    unfolding cauchy_def
+    using dist_triangle_half_l
+    by blast
+qed
+
+lemma convergent_imp_cauchy:
+ "(s ---> l) sequentially ==> Cauchy s"
+proof(simp only: cauchy_def, rule, rule)
+  fix e::real assume "e>0" "(s ---> l) sequentially"
+  then obtain N::nat where N:"\<forall>n\<ge>N. dist (s n) l < e/2" unfolding Lim_sequentially by(erule_tac x="e/2" in allE) auto
+  thus "\<exists>N. \<forall>m n. N \<le> m \<and> N \<le> n \<longrightarrow> dist (s m) (s n) < e"  using dist_triangle_half_l[of _ l e _] by (rule_tac x=N in exI) auto
+qed
+
+lemma cauchy_imp_bounded: assumes "Cauchy s" shows "bounded {y. (\<exists>n::nat. y = s n)}"
+proof-
+  from assms obtain N::nat where "\<forall>m n. N \<le> m \<and> N \<le> n \<longrightarrow> dist (s m) (s n) < 1" unfolding cauchy_def apply(erule_tac x= 1 in allE) by auto
+  hence N:"\<forall>n. N \<le> n \<longrightarrow> dist (s N) (s n) < 1" by auto
+  moreover
+  have "bounded (s ` {0..N})" using finite_imp_bounded[of "s ` {1..N}"] by auto
+  then obtain a where a:"\<forall>x\<in>s ` {0..N}. dist (s N) x \<le> a"
+    unfolding bounded_any_center [where a="s N"] by auto
+  ultimately show "?thesis"
+    unfolding bounded_any_center [where a="s N"]
+    apply(rule_tac x="max a 1" in exI) apply auto
+    apply(erule_tac x=n in allE) apply(erule_tac x=n in ballE) by auto
+qed
+
+lemma compact_imp_complete: assumes "compact s" shows "complete s"
+proof-
+  { fix f assume as: "(\<forall>n::nat. f n \<in> s)" "Cauchy f"
+    from as(1) obtain l r where lr: "l\<in>s" "subseq r" "((f \<circ> r) ---> l) sequentially" using assms unfolding compact_def by blast
+
+    note lr' = subseq_bigger [OF lr(2)]
+
+    { fix e::real assume "e>0"
+      from as(2) obtain N where N:"\<forall>m n. N \<le> m \<and> N \<le> n \<longrightarrow> dist (f m) (f n) < e/2" unfolding cauchy_def using `e>0` apply (erule_tac x="e/2" in allE) by auto
+      from lr(3)[unfolded Lim_sequentially, THEN spec[where x="e/2"]] obtain M where M:"\<forall>n\<ge>M. dist ((f \<circ> r) n) l < e/2" using `e>0` by auto
+      { fix n::nat assume n:"n \<ge> max N M"
+	have "dist ((f \<circ> r) n) l < e/2" using n M by auto
+	moreover have "r n \<ge> N" using lr'[of n] n by auto
+	hence "dist (f n) ((f \<circ> r) n) < e / 2" using N using n by auto
+	ultimately have "dist (f n) l < e" using dist_triangle_half_r[of "f (r n)" "f n" e l] by (auto simp add: dist_commute)  }
+      hence "\<exists>N. \<forall>n\<ge>N. dist (f n) l < e" by blast  }
+    hence "\<exists>l\<in>s. (f ---> l) sequentially" using `l\<in>s` unfolding Lim_sequentially by auto  }
+  thus ?thesis unfolding complete_def by auto
+qed
+
+instance heine_borel < complete_space
+proof
+  fix f :: "nat \<Rightarrow> 'a" assume "Cauchy f"
+  hence "bounded (range f)" unfolding image_def
+    using cauchy_imp_bounded [of f] by auto
+  hence "compact (closure (range f))"
+    using bounded_closed_imp_compact [of "closure (range f)"] by auto
+  hence "complete (closure (range f))"
+    using compact_imp_complete by auto
+  moreover have "\<forall>n. f n \<in> closure (range f)"
+    using closure_subset [of "range f"] by auto
+  ultimately have "\<exists>l\<in>closure (range f). (f ---> l) sequentially"
+    using `Cauchy f` unfolding complete_def by auto
+  then show "convergent f"
+    unfolding convergent_def LIMSEQ_conv_tendsto [symmetric] by auto
+qed
+
+lemma complete_univ: "complete (UNIV :: 'a::complete_space set)"
+proof(simp add: complete_def, rule, rule)
+  fix f :: "nat \<Rightarrow> 'a" assume "Cauchy f"
+  hence "convergent f" by (rule Cauchy_convergent)
+  hence "\<exists>l. f ----> l" unfolding convergent_def .  
+  thus "\<exists>l. (f ---> l) sequentially" unfolding LIMSEQ_conv_tendsto .
+qed
+
+lemma complete_imp_closed: assumes "complete s" shows "closed s"
+proof -
+  { fix x assume "x islimpt s"
+    then obtain f where f: "\<forall>n. f n \<in> s - {x}" "(f ---> x) sequentially"
+      unfolding islimpt_sequential by auto
+    then obtain l where l: "l\<in>s" "(f ---> l) sequentially"
+      using `complete s`[unfolded complete_def] using convergent_imp_cauchy[of f x] by auto
+    hence "x \<in> s"  using Lim_unique[of sequentially f l x] trivial_limit_sequentially f(2) by auto
+  }
+  thus "closed s" unfolding closed_limpt by auto
+qed
+
+lemma complete_eq_closed:
+  fixes s :: "'a::complete_space set"
+  shows "complete s \<longleftrightarrow> closed s" (is "?lhs = ?rhs")
+proof
+  assume ?lhs thus ?rhs by (rule complete_imp_closed)
+next
+  assume ?rhs
+  { fix f assume as:"\<forall>n::nat. f n \<in> s" "Cauchy f"
+    then obtain l where "(f ---> l) sequentially" using complete_univ[unfolded complete_def, THEN spec[where x=f]] by auto
+    hence "\<exists>l\<in>s. (f ---> l) sequentially" using `?rhs`[unfolded closed_sequential_limits, THEN spec[where x=f], THEN spec[where x=l]] using as(1) by auto  }
+  thus ?lhs unfolding complete_def by auto
+qed
+
+lemma convergent_eq_cauchy:
+  fixes s :: "nat \<Rightarrow> 'a::complete_space"
+  shows "(\<exists>l. (s ---> l) sequentially) \<longleftrightarrow> Cauchy s" (is "?lhs = ?rhs")
+proof
+  assume ?lhs then obtain l where "(s ---> l) sequentially" by auto
+  thus ?rhs using convergent_imp_cauchy by auto
+next
+  assume ?rhs thus ?lhs using complete_univ[unfolded complete_def, THEN spec[where x=s]] by auto
+qed
+
+lemma convergent_imp_bounded:
+  fixes s :: "nat \<Rightarrow> 'a::metric_space"
+  shows "(s ---> l) sequentially ==> bounded (s ` (UNIV::(nat set)))"
+  using convergent_imp_cauchy[of s]
+  using cauchy_imp_bounded[of s]
+  unfolding image_def
+  by auto
+
+subsection{* Total boundedness. *}
+
+fun helper_1::"('a::metric_space set) \<Rightarrow> real \<Rightarrow> nat \<Rightarrow> 'a" where
+  "helper_1 s e n = (SOME y::'a. y \<in> s \<and> (\<forall>m<n. \<not> (dist (helper_1 s e m) y < e)))"
+declare helper_1.simps[simp del]
+
+lemma compact_imp_totally_bounded:
+  assumes "compact s"
+  shows "\<forall>e>0. \<exists>k. finite k \<and> k \<subseteq> s \<and> s \<subseteq> (\<Union>((\<lambda>x. ball x e) ` k))"
+proof(rule, rule, rule ccontr)
+  fix e::real assume "e>0" and assm:"\<not> (\<exists>k. finite k \<and> k \<subseteq> s \<and> s \<subseteq> \<Union>(\<lambda>x. ball x e) ` k)"
+  def x \<equiv> "helper_1 s e"
+  { fix n
+    have "x n \<in> s \<and> (\<forall>m<n. \<not> dist (x m) (x n) < e)"
+    proof(induct_tac rule:nat_less_induct)
+      fix n  def Q \<equiv> "(\<lambda>y. y \<in> s \<and> (\<forall>m<n. \<not> dist (x m) y < e))"
+      assume as:"\<forall>m<n. x m \<in> s \<and> (\<forall>ma<m. \<not> dist (x ma) (x m) < e)"
+      have "\<not> s \<subseteq> (\<Union>x\<in>x ` {0..<n}. ball x e)" using assm apply simp apply(erule_tac x="x ` {0 ..< n}" in allE) using as by auto
+      then obtain z where z:"z\<in>s" "z \<notin> (\<Union>x\<in>x ` {0..<n}. ball x e)" unfolding subset_eq by auto
+      have "Q (x n)" unfolding x_def and helper_1.simps[of s e n]
+	apply(rule someI2[where a=z]) unfolding x_def[symmetric] and Q_def using z by auto
+      thus "x n \<in> s \<and> (\<forall>m<n. \<not> dist (x m) (x n) < e)" unfolding Q_def by auto
+    qed }
+  hence "\<forall>n::nat. x n \<in> s" and x:"\<forall>n. \<forall>m < n. \<not> (dist (x m) (x n) < e)" by blast+
+  then obtain l r where "l\<in>s" and r:"subseq r" and "((x \<circ> r) ---> l) sequentially" using assms(1)[unfolded compact_def, THEN spec[where x=x]] by auto
+  from this(3) have "Cauchy (x \<circ> r)" using convergent_imp_cauchy by auto
+  then obtain N::nat where N:"\<forall>m n. N \<le> m \<and> N \<le> n \<longrightarrow> dist ((x \<circ> r) m) ((x \<circ> r) n) < e" unfolding cauchy_def using `e>0` by auto
+  show False
+    using N[THEN spec[where x=N], THEN spec[where x="N+1"]]
+    using r[unfolded subseq_def, THEN spec[where x=N], THEN spec[where x="N+1"]]
+    using x[THEN spec[where x="r (N+1)"], THEN spec[where x="r (N)"]] by auto
+qed
+
+subsection{* Heine-Borel theorem (following Burkill \& Burkill vol. 2) *}
+
+lemma heine_borel_lemma: fixes s::"'a::metric_space set"
+  assumes "compact s"  "s \<subseteq> (\<Union> t)"  "\<forall>b \<in> t. open b"
+  shows "\<exists>e>0. \<forall>x \<in> s. \<exists>b \<in> t. ball x e \<subseteq> b"
+proof(rule ccontr)
+  assume "\<not> (\<exists>e>0. \<forall>x\<in>s. \<exists>b\<in>t. ball x e \<subseteq> b)"
+  hence cont:"\<forall>e>0. \<exists>x\<in>s. \<forall>xa\<in>t. \<not> (ball x e \<subseteq> xa)" by auto
+  { fix n::nat
+    have "1 / real (n + 1) > 0" by auto
+    hence "\<exists>x. x\<in>s \<and> (\<forall>xa\<in>t. \<not> (ball x (inverse (real (n+1))) \<subseteq> xa))" using cont unfolding Bex_def by auto }
+  hence "\<forall>n::nat. \<exists>x. x \<in> s \<and> (\<forall>xa\<in>t. \<not> ball x (inverse (real (n + 1))) \<subseteq> xa)" by auto
+  then obtain f where f:"\<forall>n::nat. f n \<in> s \<and> (\<forall>xa\<in>t. \<not> ball (f n) (inverse (real (n + 1))) \<subseteq> xa)"
+    using choice[of "\<lambda>n::nat. \<lambda>x. x\<in>s \<and> (\<forall>xa\<in>t. \<not> ball x (inverse (real (n + 1))) \<subseteq> xa)"] by auto
+
+  then obtain l r where l:"l\<in>s" and r:"subseq r" and lr:"((f \<circ> r) ---> l) sequentially"
+    using assms(1)[unfolded compact_def, THEN spec[where x=f]] by auto
+
+  obtain b where "l\<in>b" "b\<in>t" using assms(2) and l by auto
+  then obtain e where "e>0" and e:"\<forall>z. dist z l < e \<longrightarrow> z\<in>b"
+    using assms(3)[THEN bspec[where x=b]] unfolding open_dist by auto
+
+  then obtain N1 where N1:"\<forall>n\<ge>N1. dist ((f \<circ> r) n) l < e / 2"
+    using lr[unfolded Lim_sequentially, THEN spec[where x="e/2"]] by auto
+
+  obtain N2::nat where N2:"N2>0" "inverse (real N2) < e /2" using real_arch_inv[of "e/2"] and `e>0` by auto
+  have N2':"inverse (real (r (N1 + N2) +1 )) < e/2"
+    apply(rule order_less_trans) apply(rule less_imp_inverse_less) using N2
+    using subseq_bigger[OF r, of "N1 + N2"] by auto
+
+  def x \<equiv> "(f (r (N1 + N2)))"
+  have x:"\<not> ball x (inverse (real (r (N1 + N2) + 1))) \<subseteq> b" unfolding x_def
+    using f[THEN spec[where x="r (N1 + N2)"]] using `b\<in>t` by auto
+  have "\<exists>y\<in>ball x (inverse (real (r (N1 + N2) + 1))). y\<notin>b" apply(rule ccontr) using x by auto
+  then obtain y where y:"y \<in> ball x (inverse (real (r (N1 + N2) + 1)))" "y \<notin> b" by auto
+
+  have "dist x l < e/2" using N1 unfolding x_def o_def by auto
+  hence "dist y l < e" using y N2' using dist_triangle[of y l x]by (auto simp add:dist_commute)
+
+  thus False using e and `y\<notin>b` by auto
+qed
+
+lemma compact_imp_heine_borel: "compact s ==> (\<forall>f. (\<forall>t \<in> f. open t) \<and> s \<subseteq> (\<Union> f)
+               \<longrightarrow> (\<exists>f'. f' \<subseteq> f \<and> finite f' \<and> s \<subseteq> (\<Union> f')))"
+proof clarify
+  fix f assume "compact s" " \<forall>t\<in>f. open t" "s \<subseteq> \<Union>f"
+  then obtain e::real where "e>0" and "\<forall>x\<in>s. \<exists>b\<in>f. ball x e \<subseteq> b" using heine_borel_lemma[of s f] by auto
+  hence "\<forall>x\<in>s. \<exists>b. b\<in>f \<and> ball x e \<subseteq> b" by auto
+  hence "\<exists>bb. \<forall>x\<in>s. bb x \<in>f \<and> ball x e \<subseteq> bb x" using bchoice[of s "\<lambda>x b. b\<in>f \<and> ball x e \<subseteq> b"] by auto
+  then obtain  bb where bb:"\<forall>x\<in>s. (bb x) \<in> f \<and> ball x e \<subseteq> (bb x)" by blast
+
+  from `compact s` have  "\<exists> k. finite k \<and> k \<subseteq> s \<and> s \<subseteq> \<Union>(\<lambda>x. ball x e) ` k" using compact_imp_totally_bounded[of s] `e>0` by auto
+  then obtain k where k:"finite k" "k \<subseteq> s" "s \<subseteq> \<Union>(\<lambda>x. ball x e) ` k" by auto
+
+  have "finite (bb ` k)" using k(1) by auto
+  moreover
+  { fix x assume "x\<in>s"
+    hence "x\<in>\<Union>(\<lambda>x. ball x e) ` k" using k(3)  unfolding subset_eq by auto
+    hence "\<exists>X\<in>bb ` k. x \<in> X" using bb k(2) by blast
+    hence "x \<in> \<Union>(bb ` k)" using  Union_iff[of x "bb ` k"] by auto
+  }
+  ultimately show "\<exists>f'\<subseteq>f. finite f' \<and> s \<subseteq> \<Union>f'" using bb k(2) by (rule_tac x="bb ` k" in exI) auto
+qed
+
+subsection{* Bolzano-Weierstrass property. *}
+
+lemma heine_borel_imp_bolzano_weierstrass:
+  assumes "\<forall>f. (\<forall>t \<in> f. open t) \<and> s \<subseteq> (\<Union> f) --> (\<exists>f'. f' \<subseteq> f \<and> finite f' \<and> s \<subseteq> (\<Union> f'))"
+          "infinite t"  "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)" 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"
+    using assms(1)[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
+  { fix x y assume "x\<in>t" "y\<in>t" "f x = f y"
+    hence "x \<in> f x"  "y \<in> f x \<longrightarrow> y = x" using f[THEN bspec[where x=x]] and `t\<subseteq>s` by auto
+    hence "x = y" using `f x = f y` and f[THEN bspec[where x=y]] and `y\<in>t` and `t\<subseteq>s` by auto  }
+  hence "infinite (f ` t)" using assms(2) using finite_imageD[unfolded inj_on_def, of f t] by auto
+  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" using g'[THEN bspec[where x=h]] by auto
+    hence "y = x" using f[THEN bspec[where x=y]] and `x\<in>t` and `x\<in>h`[unfolded `h = f y`] by auto
+    hence False using `f x \<notin> g` `h\<in>g` unfolding `h = f y` by auto  }
+  hence "f ` t \<subseteq> g" by auto
+  ultimately show False using g(2) using finite_subset by auto
+qed
+
+subsection{* Complete the chain of compactness variants. *}
+
+primrec helper_2::"(real \<Rightarrow> 'a::metric_space) \<Rightarrow> nat \<Rightarrow> 'a" where
+  "helper_2 beyond 0 = beyond 0" |
+  "helper_2 beyond (Suc n) = beyond (dist arbitrary (helper_2 beyond n) + 1 )"
+
+lemma bolzano_weierstrass_imp_bounded: fixes s::"'a::metric_space set"
+  assumes "\<forall>t. infinite t \<and> t \<subseteq> s --> (\<exists>x \<in> s. x islimpt t)"
+  shows "bounded s"
+proof(rule ccontr)
+  assume "\<not> bounded s"
+  then obtain beyond where "\<forall>a. beyond a \<in>s \<and> \<not> dist arbitrary (beyond a) \<le> a"
+    unfolding bounded_any_center [where a=arbitrary]
+    apply simp using choice[of "\<lambda>a x. x\<in>s \<and> \<not> dist arbitrary x \<le> a"] by auto
+  hence beyond:"\<And>a. beyond a \<in>s" "\<And>a. dist arbitrary (beyond a) > a"
+    unfolding linorder_not_le by auto
+  def x \<equiv> "helper_2 beyond"
+
+  { fix m n ::nat assume "m<n"
+    hence "dist arbitrary (x m) + 1 < dist arbitrary (x n)"
+    proof(induct n)
+      case 0 thus ?case by auto
+    next
+      case (Suc n)
+      have *:"dist arbitrary (x n) + 1 < dist arbitrary (x (Suc n))"
+        unfolding x_def and helper_2.simps
+	using beyond(2)[of "dist arbitrary (helper_2 beyond n) + 1"] by auto
+      thus ?case proof(cases "m < n")
+	case True thus ?thesis using Suc and * by auto
+      next
+	case False hence "m = n" using Suc(2) by auto
+	thus ?thesis using * by auto
+      qed
+    qed  } note * = this
+  { fix m n ::nat assume "m\<noteq>n"
+    have "1 < dist (x m) (x n)"
+    proof(cases "m<n")
+      case True
+      hence "1 < dist arbitrary (x n) - dist arbitrary (x m)" using *[of m n] by auto
+      thus ?thesis using dist_triangle [of arbitrary "x n" "x m"] by arith
+    next
+      case False hence "n<m" using `m\<noteq>n` by auto
+      hence "1 < dist arbitrary (x m) - dist arbitrary (x n)" using *[of n m] by auto
+      thus ?thesis using dist_triangle2 [of arbitrary "x m" "x n"] by arith
+    qed  } note ** = this
+  { fix a b assume "x a = x b" "a \<noteq> b"
+    hence False using **[of a b] by auto  }
+  hence "inj x" unfolding inj_on_def by auto
+  moreover
+  { fix n::nat
+    have "x n \<in> s"
+    proof(cases "n = 0")
+      case True thus ?thesis unfolding x_def using beyond by auto
+    next
+      case False then obtain z where "n = Suc z" using not0_implies_Suc by auto
+      thus ?thesis unfolding x_def using beyond by auto
+    qed  }
+  ultimately have "infinite (range x) \<and> range x \<subseteq> s" unfolding x_def using range_inj_infinite[of "helper_2 beyond"] using beyond(1) by auto
+
+  then obtain l where "l\<in>s" and l:"l islimpt range x" using assms[THEN spec[where x="range x"]] by auto
+  then obtain y where "x y \<noteq> l" and y:"dist (x y) l < 1/2" unfolding islimpt_approachable apply(erule_tac x="1/2" in allE) by auto
+  then obtain z where "x z \<noteq> l" and z:"dist (x z) l < dist (x y) l" using l[unfolded islimpt_approachable, THEN spec[where x="dist (x y) l"]]
+    unfolding dist_nz by auto
+  show False using y and z and dist_triangle_half_l[of "x y" l 1 "x z"] and **[of y z] by auto
+qed
+
+lemma sequence_infinite_lemma:
+  fixes l :: "'a::metric_space" (* TODO: generalize *)
+  assumes "\<forall>n::nat. (f n  \<noteq> l)"  "(f ---> l) sequentially"
+  shows "infinite {y. (\<exists> n. y = f n)}"
+proof(rule ccontr)
+  let ?A = "(\<lambda>x. dist x l) ` {y. \<exists>n. y = f n}"
+  assume "\<not> infinite {y. \<exists>n. y = f n}"
+  hence **:"finite ?A" "?A \<noteq> {}" by auto
+  obtain k where k:"dist (f k) l = Min ?A" using Min_in[OF **] by auto
+  have "0 < Min ?A" using assms(1) unfolding dist_nz unfolding Min_gr_iff[OF **] by auto
+  then obtain N where "dist (f N) l < Min ?A" using assms(2)[unfolded Lim_sequentially, THEN spec[where x="Min ?A"]] by auto
+  moreover have "dist (f N) l \<in> ?A" by auto
+  ultimately show False using Min_le[OF **(1), of "dist (f N) l"] by auto
+qed
+
+lemma sequence_unique_limpt:
+  fixes l :: "'a::metric_space" (* TODO: generalize *)
+  assumes "\<forall>n::nat. (f n \<noteq> l)"  "(f ---> l) sequentially"  "l' islimpt {y.  (\<exists>n. y = f n)}"
+  shows "l' = l"
+proof(rule ccontr)
+  def e \<equiv> "dist l' l"
+  assume "l' \<noteq> l" hence "e>0" unfolding dist_nz e_def by auto
+  then obtain N::nat where N:"\<forall>n\<ge>N. dist (f n) l < e / 2"
+    using assms(2)[unfolded Lim_sequentially, THEN spec[where x="e/2"]] by auto
+  def d \<equiv> "Min (insert (e/2) ((\<lambda>n. if dist (f n) l' = 0 then e/2 else dist (f n) l') ` {0 .. N}))"
+  have "d>0" using `e>0` unfolding d_def e_def using zero_le_dist[of _ l', unfolded order_le_less] by auto
+  obtain k where k:"f k \<noteq> l'"  "dist (f k) l' < d" using `d>0` and assms(3)[unfolded islimpt_approachable, THEN spec[where x="d"]] by auto
+  have "k\<ge>N" using k(1)[unfolded dist_nz] using k(2)[unfolded d_def]
+    by force
+  hence "dist l' l < e" using N[THEN spec[where x=k]] using k(2)[unfolded d_def] and dist_triangle_half_r[of "f k" l' e l] by auto
+  thus False unfolding e_def by auto
+qed
+
+lemma bolzano_weierstrass_imp_closed:
+  fixes s :: "'a::metric_space set" (* TODO: can this be generalized? *)
+  assumes "\<forall>t. infinite t \<and> t \<subseteq> s --> (\<exists>x \<in> s. x islimpt t)"
+  shows "closed s"
+proof-
+  { fix x l assume as: "\<forall>n::nat. x n \<in> s" "(x ---> l) sequentially"
+    hence "l \<in> s"
+    proof(cases "\<forall>n. x n \<noteq> l")
+      case False thus "l\<in>s" using as(1) by auto
+    next
+      case True note cas = this
+      with as(2) have "infinite {y. \<exists>n. y = x n}" using sequence_infinite_lemma[of x l] by auto
+      then obtain l' where "l'\<in>s" "l' islimpt {y. \<exists>n. y = x n}" using assms[THEN spec[where x="{y. \<exists>n. y = x n}"]] as(1) by auto
+      thus "l\<in>s" using sequence_unique_limpt[of x l l'] using as cas by auto
+    qed  }
+  thus ?thesis unfolding closed_sequential_limits by fast
+qed
+
+text{* Hence express everything as an equivalence.   *}
+
+lemma compact_eq_heine_borel:
+  fixes s :: "'a::heine_borel set"
+  shows "compact s \<longleftrightarrow>
+           (\<forall>f. (\<forall>t \<in> f. open t) \<and> s \<subseteq> (\<Union> f)
+               --> (\<exists>f'. f' \<subseteq> f \<and> finite f' \<and> s \<subseteq> (\<Union> f')))" (is "?lhs = ?rhs")
+proof
+  assume ?lhs thus ?rhs using compact_imp_heine_borel[of s] by blast
+next
+  assume ?rhs
+  hence "\<forall>t. infinite t \<and> t \<subseteq> s \<longrightarrow> (\<exists>x\<in>s. x islimpt t)"
+    by (blast intro: heine_borel_imp_bolzano_weierstrass[of s])
+  thus ?lhs using bolzano_weierstrass_imp_bounded[of s] bolzano_weierstrass_imp_closed[of s] bounded_closed_imp_compact[of s] by blast
+qed
+
+lemma compact_eq_bolzano_weierstrass:
+  fixes s :: "'a::heine_borel set"
+  shows "compact s \<longleftrightarrow> (\<forall>t. infinite t \<and> t \<subseteq> s --> (\<exists>x \<in> s. x islimpt t))" (is "?lhs = ?rhs")
+proof
+  assume ?lhs thus ?rhs unfolding compact_eq_heine_borel using heine_borel_imp_bolzano_weierstrass[of s] by auto
+next
+  assume ?rhs thus ?lhs using bolzano_weierstrass_imp_bounded bolzano_weierstrass_imp_closed bounded_closed_imp_compact by auto
+qed
+
+lemma compact_eq_bounded_closed:
+  fixes s :: "'a::heine_borel set"
+  shows "compact s \<longleftrightarrow> bounded s \<and> closed s"  (is "?lhs = ?rhs")
+proof
+  assume ?lhs thus ?rhs unfolding compact_eq_bolzano_weierstrass using bolzano_weierstrass_imp_bounded bolzano_weierstrass_imp_closed by auto
+next
+  assume ?rhs thus ?lhs using bounded_closed_imp_compact by auto
+qed
+
+lemma compact_imp_bounded:
+  fixes s :: "'a::metric_space set"
+  shows "compact s ==> bounded s"
+proof -
+  assume "compact s"
+  hence "\<forall>f. (\<forall>t\<in>f. open t) \<and> s \<subseteq> \<Union>f \<longrightarrow> (\<exists>f'\<subseteq>f. finite f' \<and> s \<subseteq> \<Union>f')"
+    by (rule compact_imp_heine_borel)
+  hence "\<forall>t. infinite t \<and> t \<subseteq> s \<longrightarrow> (\<exists>x \<in> s. x islimpt t)"
+    using heine_borel_imp_bolzano_weierstrass[of s] by auto
+  thus "bounded s"
+    by (rule bolzano_weierstrass_imp_bounded)
+qed
+
+lemma compact_imp_closed:
+  fixes s :: "'a::metric_space set"
+  shows "compact s ==> closed s"
+proof -
+  assume "compact s"
+  hence "\<forall>f. (\<forall>t\<in>f. open t) \<and> s \<subseteq> \<Union>f \<longrightarrow> (\<exists>f'\<subseteq>f. finite f' \<and> s \<subseteq> \<Union>f')"
+    by (rule compact_imp_heine_borel)
+  hence "\<forall>t. infinite t \<and> t \<subseteq> s \<longrightarrow> (\<exists>x \<in> s. x islimpt t)"
+    using heine_borel_imp_bolzano_weierstrass[of s] by auto
+  thus "closed s"
+    by (rule bolzano_weierstrass_imp_closed)
+qed
+
+text{* In particular, some common special cases. *}
+
+lemma compact_empty[simp]:
+ "compact {}"
+  unfolding compact_def
+  by simp
+
+(* TODO: can any of the next 3 lemmas be generalized to metric spaces? *)
+
+  (* FIXME : Rename *)
+lemma compact_union[intro]:
+  fixes s t :: "'a::heine_borel set"
+  shows "compact s \<Longrightarrow> compact t ==> compact (s \<union> t)"
+  unfolding compact_eq_bounded_closed
+  using bounded_Un[of s t]
+  using closed_Un[of s t]
+  by simp
+
+lemma compact_inter[intro]:
+  fixes s t :: "'a::heine_borel set"
+  shows "compact s \<Longrightarrow> compact t ==> compact (s \<inter> t)"
+  unfolding compact_eq_bounded_closed
+  using bounded_Int[of s t]
+  using closed_Int[of s t]
+  by simp
+
+lemma compact_inter_closed[intro]:
+  fixes s t :: "'a::heine_borel set"
+  shows "compact s \<Longrightarrow> closed t ==> compact (s \<inter> t)"
+  unfolding compact_eq_bounded_closed
+  using closed_Int[of s t]
+  using bounded_subset[of "s \<inter> t" s]
+  by blast
+
+lemma closed_inter_compact[intro]:
+  fixes s t :: "'a::heine_borel set"
+  shows "closed s \<Longrightarrow> compact t ==> compact (s \<inter> t)"
+proof-
+  assume "closed s" "compact t"
+  moreover
+  have "s \<inter> t = t \<inter> s" by auto ultimately
+  show ?thesis
+    using compact_inter_closed[of t s]
+    by auto
+qed
+
+lemma closed_sing [simp]:
+  fixes a :: "'a::metric_space"
+  shows "closed {a}"
+  apply (clarsimp simp add: closed_def open_dist)
+  apply (rule ccontr)
+  apply (drule_tac x="dist x a" in spec)
+  apply (simp add: dist_nz dist_commute)
+  done
+
+lemma finite_imp_closed:
+  fixes s :: "'a::metric_space set"
+  shows "finite s ==> closed s"
+proof (induct set: finite)
+  case empty show "closed {}" by simp
+next
+  case (insert x F)
+  hence "closed ({x} \<union> F)" by (simp only: closed_Un closed_sing)
+  thus "closed (insert x F)" by simp
+qed
+
+lemma finite_imp_compact:
+  fixes s :: "'a::heine_borel set"
+  shows "finite s ==> compact s"
+  unfolding compact_eq_bounded_closed
+  using finite_imp_closed finite_imp_bounded
+  by blast
+
+lemma compact_sing [simp]: "compact {a}"
+  unfolding compact_def o_def subseq_def
+  by (auto simp add: tendsto_const)
+
+lemma compact_cball[simp]:
+  fixes x :: "'a::heine_borel"
+  shows "compact(cball x e)"
+  using compact_eq_bounded_closed bounded_cball closed_cball
+  by blast
+
+lemma compact_frontier_bounded[intro]:
+  fixes s :: "'a::heine_borel set"
+  shows "bounded s ==> compact(frontier s)"
+  unfolding frontier_def
+  using compact_eq_bounded_closed
+  by blast
+
+lemma compact_frontier[intro]:
+  fixes s :: "'a::heine_borel set"
+  shows "compact s ==> compact (frontier s)"
+  using compact_eq_bounded_closed compact_frontier_bounded
+  by blast
+
+lemma frontier_subset_compact:
+  fixes s :: "'a::heine_borel set"
+  shows "compact s ==> frontier s \<subseteq> s"
+  using frontier_subset_closed compact_eq_bounded_closed
+  by blast
+
+lemma open_delete:
+  fixes s :: "'a::metric_space set"
+  shows "open s ==> open(s - {x})"
+  using open_Diff[of s "{x}"] closed_sing
+  by blast
+
+text{* Finite intersection property. I could make it an equivalence in fact. *}
+
+lemma compact_imp_fip:
+  fixes s :: "'a::heine_borel set"
+  assumes "compact s"  "\<forall>t \<in> f. closed t"
+        "\<forall>f'. finite f' \<and> f' \<subseteq> f --> (s \<inter> (\<Inter> f') \<noteq> {})"
+  shows "s \<inter> (\<Inter> f) \<noteq> {}"
+proof
+  assume as:"s \<inter> (\<Inter> f) = {}"
+  hence "s \<subseteq> \<Union>op - UNIV ` f" by auto
+  moreover have "Ball (op - UNIV ` f) open" using open_Diff closed_Diff using assms(2) by auto
+  ultimately obtain f' where f':"f' \<subseteq> op - UNIV ` f"  "finite f'"  "s \<subseteq> \<Union>f'" using assms(1)[unfolded compact_eq_heine_borel, THEN spec[where x="(\<lambda>t. UNIV - t) ` f"]] by auto
+  hence "finite (op - UNIV ` f') \<and> op - UNIV ` f' \<subseteq> f" by(auto simp add: Diff_Diff_Int)
+  hence "s \<inter> \<Inter>op - UNIV ` f' \<noteq> {}" using assms(3)[THEN spec[where x="op - UNIV ` f'"]] by auto
+  thus False using f'(3) unfolding subset_eq and Union_iff by blast
+qed
+
+subsection{* Bounded closed nest property (proof does not use Heine-Borel).            *}
+
+lemma bounded_closed_nest:
+  assumes "\<forall>n. closed(s n)" "\<forall>n. (s n \<noteq> {})"
+  "(\<forall>m n. m \<le> n --> s n \<subseteq> s m)"  "bounded(s 0)"
+  shows "\<exists>a::'a::heine_borel. \<forall>n::nat. a \<in> s(n)"
+proof-
+  from assms(2) obtain x where x:"\<forall>n::nat. x n \<in> s n" using choice[of "\<lambda>n x. x\<in> s n"] by auto
+  from assms(4,1) have *:"compact (s 0)" using bounded_closed_imp_compact[of "s 0"] by auto
+
+  then obtain l r where lr:"l\<in>s 0" "subseq r" "((x \<circ> r) ---> l) sequentially"
+    unfolding compact_def apply(erule_tac x=x in allE)  using x using assms(3) by blast
+
+  { fix n::nat
+    { fix e::real assume "e>0"
+      with lr(3) obtain N where N:"\<forall>m\<ge>N. dist ((x \<circ> r) m) l < e" unfolding Lim_sequentially by auto
+      hence "dist ((x \<circ> r) (max N n)) l < e" by auto
+      moreover
+      have "r (max N n) \<ge> n" using lr(2) using subseq_bigger[of r "max N n"] by auto
+      hence "(x \<circ> r) (max N n) \<in> s n"
+	using x apply(erule_tac x=n in allE)
+	using x apply(erule_tac x="r (max N n)" in allE)
+	using assms(3) apply(erule_tac x=n in allE)apply( erule_tac x="r (max N n)" in allE) by auto
+      ultimately have "\<exists>y\<in>s n. dist y l < e" by auto
+    }
+    hence "l \<in> s n" using closed_approachable[of "s n" l] assms(1) by blast
+  }
+  thus ?thesis by auto
+qed
+
+text{* Decreasing case does not even need compactness, just completeness.        *}
+
+lemma decreasing_closed_nest:
+  assumes "\<forall>n. closed(s n)"
+          "\<forall>n. (s n \<noteq> {})"
+          "\<forall>m n. m \<le> n --> s n \<subseteq> s m"
+          "\<forall>e>0. \<exists>n. \<forall>x \<in> (s n). \<forall> y \<in> (s n). dist x y < e"
+  shows "\<exists>a::'a::heine_borel. \<forall>n::nat. a \<in> s n"
+proof-
+  have "\<forall>n. \<exists> x. x\<in>s n" using assms(2) by auto
+  hence "\<exists>t. \<forall>n. t n \<in> s n" using choice[of "\<lambda> n x. x \<in> s n"] by auto
+  then obtain t where t: "\<forall>n. t n \<in> s n" by auto
+  { fix e::real assume "e>0"
+    then obtain N where N:"\<forall>x\<in>s N. \<forall>y\<in>s N. dist x y < e" using assms(4) by auto
+    { fix m n ::nat assume "N \<le> m \<and> N \<le> n"
+      hence "t m \<in> s N" "t n \<in> s N" using assms(3) t unfolding  subset_eq t by blast+
+      hence "dist (t m) (t n) < e" using N by auto
+    }
+    hence "\<exists>N. \<forall>m n. N \<le> m \<and> N \<le> n \<longrightarrow> dist (t m) (t n) < e" by auto
+  }
+  hence  "Cauchy t" unfolding cauchy_def by auto
+  then obtain l where l:"(t ---> l) sequentially" using complete_univ unfolding complete_def by auto
+  { fix n::nat
+    { fix e::real assume "e>0"
+      then obtain N::nat where N:"\<forall>n\<ge>N. dist (t n) l < e" using l[unfolded Lim_sequentially] by auto
+      have "t (max n N) \<in> s n" using assms(3) unfolding subset_eq apply(erule_tac x=n in allE) apply (erule_tac x="max n N" in allE) using t by auto
+      hence "\<exists>y\<in>s n. dist y l < e" apply(rule_tac x="t (max n N)" in bexI) using N by auto
+    }
+    hence "l \<in> s n" using closed_approachable[of "s n" l] assms(1) by auto
+  }
+  then show ?thesis by auto
+qed
+
+text{* Strengthen it to the intersection actually being a singleton.             *}
+
+lemma decreasing_closed_nest_sing:
+  assumes "\<forall>n. closed(s n)"
+          "\<forall>n. s n \<noteq> {}"
+          "\<forall>m n. m \<le> n --> s n \<subseteq> s m"
+          "\<forall>e>0. \<exists>n. \<forall>x \<in> (s n). \<forall> y\<in>(s n). dist x y < e"
+  shows "\<exists>a::'a::heine_borel. \<Inter> {t. (\<exists>n::nat. t = s n)} = {a}"
+proof-
+  obtain a where a:"\<forall>n. a \<in> s n" using decreasing_closed_nest[of s] using assms by auto
+  { fix b assume b:"b \<in> \<Inter>{t. \<exists>n. t = s n}"
+    { fix e::real assume "e>0"
+      hence "dist a b < e" using assms(4 )using b using a by blast
+    }
+    hence "dist a b = 0" by (metis dist_eq_0_iff dist_nz real_less_def)
+  }
+  with a have "\<Inter>{t. \<exists>n. t = s n} = {a}"  by auto
+  thus ?thesis by auto
+qed
+
+text{* Cauchy-type criteria for uniform convergence. *}
+
+lemma uniformly_convergent_eq_cauchy: fixes s::"nat \<Rightarrow> 'b \<Rightarrow> 'a::heine_borel" shows
+ "(\<exists>l. \<forall>e>0. \<exists>N. \<forall>n x. N \<le> n \<and> P x --> dist(s n x)(l x) < e) \<longleftrightarrow>
+  (\<forall>e>0. \<exists>N. \<forall>m n x. N \<le> m \<and> N \<le> n \<and> P x  --> dist (s m x) (s n x) < e)" (is "?lhs = ?rhs")
+proof(rule)
+  assume ?lhs
+  then obtain l where l:"\<forall>e>0. \<exists>N. \<forall>n x. N \<le> n \<and> P x \<longrightarrow> dist (s n x) (l x) < e" by auto
+  { fix e::real assume "e>0"
+    then obtain N::nat where N:"\<forall>n x. N \<le> n \<and> P x \<longrightarrow> dist (s n x) (l x) < e / 2" using l[THEN spec[where x="e/2"]] by auto
+    { fix n m::nat and x::"'b" assume "N \<le> m \<and> N \<le> n \<and> P x"
+      hence "dist (s m x) (s n x) < e"
+	using N[THEN spec[where x=m], THEN spec[where x=x]]
+	using N[THEN spec[where x=n], THEN spec[where x=x]]
+	using dist_triangle_half_l[of "s m x" "l x" e "s n x"] by auto  }
+    hence "\<exists>N. \<forall>m n x. N \<le> m \<and> N \<le> n \<and> P x  --> dist (s m x) (s n x) < e"  by auto  }
+  thus ?rhs by auto
+next
+  assume ?rhs
+  hence "\<forall>x. P x \<longrightarrow> Cauchy (\<lambda>n. s n x)" unfolding cauchy_def apply auto by (erule_tac x=e in allE)auto
+  then obtain l where l:"\<forall>x. P x \<longrightarrow> ((\<lambda>n. s n x) ---> l x) sequentially" unfolding convergent_eq_cauchy[THEN sym]
+    using choice[of "\<lambda>x l. P x \<longrightarrow> ((\<lambda>n. s n x) ---> l) sequentially"] by auto
+  { fix e::real assume "e>0"
+    then obtain N where N:"\<forall>m n x. N \<le> m \<and> N \<le> n \<and> P x \<longrightarrow> dist (s m x) (s n x) < e/2"
+      using `?rhs`[THEN spec[where x="e/2"]] by auto
+    { fix x assume "P x"
+      then obtain M where M:"\<forall>n\<ge>M. dist (s n x) (l x) < e/2"
+	using l[THEN spec[where x=x], unfolded Lim_sequentially] using `e>0` by(auto elim!: allE[where x="e/2"])
+      fix n::nat assume "n\<ge>N"
+      hence "dist(s n x)(l x) < e"  using `P x`and N[THEN spec[where x=n], THEN spec[where x="N+M"], THEN spec[where x=x]]
+	using M[THEN spec[where x="N+M"]] and dist_triangle_half_l[of "s n x" "s (N+M) x" e "l x"] by (auto simp add: dist_commute)  }
+    hence "\<exists>N. \<forall>n x. N \<le> n \<and> P x \<longrightarrow> dist(s n x)(l x) < e" by auto }
+  thus ?lhs by auto
+qed
+
+lemma uniformly_cauchy_imp_uniformly_convergent:
+  fixes s :: "nat \<Rightarrow> 'a \<Rightarrow> 'b::heine_borel"
+  assumes "\<forall>e>0.\<exists>N. \<forall>m (n::nat) x. N \<le> m \<and> N \<le> n \<and> P x --> dist(s m x)(s n x) < e"
+          "\<forall>x. P x --> (\<forall>e>0. \<exists>N. \<forall>n. N \<le> n --> dist(s n x)(l x) < e)"
+  shows "\<forall>e>0. \<exists>N. \<forall>n x. N \<le> n \<and> P x --> dist(s n x)(l x) < e"
+proof-
+  obtain l' where l:"\<forall>e>0. \<exists>N. \<forall>n x. N \<le> n \<and> P x \<longrightarrow> dist (s n x) (l' x) < e"
+    using assms(1) unfolding uniformly_convergent_eq_cauchy[THEN sym] by auto
+  moreover
+  { fix x assume "P x"
+    hence "l x = l' x" using Lim_unique[OF trivial_limit_sequentially, of "\<lambda>n. s n x" "l x" "l' x"]
+      using l and assms(2) unfolding Lim_sequentially by blast  }
+  ultimately show ?thesis by auto
+qed
+
+subsection{* Define continuity over a net to take in restrictions of the set. *}
+
+definition
+  continuous :: "'a::t2_space net \<Rightarrow> ('a \<Rightarrow> 'b::topological_space) \<Rightarrow> bool" where
+  "continuous net f \<longleftrightarrow> (f ---> f(netlimit net)) net"
+
+lemma continuous_trivial_limit:
+ "trivial_limit net ==> continuous net f"
+  unfolding continuous_def tendsto_def trivial_limit_eq by auto
+
+lemma continuous_within: "continuous (at x within s) f \<longleftrightarrow> (f ---> f(x)) (at x within s)"
+  unfolding continuous_def
+  unfolding tendsto_def
+  using netlimit_within[of x s]
+  by (cases "trivial_limit (at x within s)") (auto simp add: trivial_limit_eventually)
+
+lemma continuous_at: "continuous (at x) f \<longleftrightarrow> (f ---> f(x)) (at x)"
+  using continuous_within [of x UNIV f] by (simp add: within_UNIV)
+
+lemma continuous_at_within:
+  assumes "continuous (at x) f"  shows "continuous (at x within s) f"
+  using assms unfolding continuous_at continuous_within
+  by (rule Lim_at_within)
+
+text{* Derive the epsilon-delta forms, which we often use as "definitions" *}
+
+lemma continuous_within_eps_delta:
+  "continuous (at x within s) f \<longleftrightarrow> (\<forall>e>0. \<exists>d>0. \<forall>x'\<in> s.  dist x' x < d --> dist (f x') (f x) < e)"
+  unfolding continuous_within and Lim_within
+  apply auto unfolding dist_nz[THEN sym] apply(auto elim!:allE) apply(rule_tac x=d in exI) by auto
+
+lemma continuous_at_eps_delta: "continuous (at x) f \<longleftrightarrow>  (\<forall>e>0. \<exists>d>0.
+                           \<forall>x'. dist x' x < d --> dist(f x')(f x) < e)"
+  using continuous_within_eps_delta[of x UNIV f]
+  unfolding within_UNIV by blast
+
+text{* Versions in terms of open balls. *}
+
+lemma continuous_within_ball:
+ "continuous (at x within s) f \<longleftrightarrow> (\<forall>e>0. \<exists>d>0.
+                            f ` (ball x d \<inter> s) \<subseteq> ball (f x) e)" (is "?lhs = ?rhs")
+proof
+  assume ?lhs
+  { fix e::real assume "e>0"
+    then obtain d where d: "d>0" "\<forall>xa\<in>s. 0 < dist xa x \<and> dist xa x < d \<longrightarrow> dist (f xa) (f x) < e"
+      using `?lhs`[unfolded continuous_within Lim_within] by auto
+    { fix y assume "y\<in>f ` (ball x d \<inter> s)"
+      hence "y \<in> ball (f x) e" using d(2) unfolding dist_nz[THEN sym]
+	apply (auto simp add: dist_commute mem_ball) apply(erule_tac x=xa in ballE) apply auto using `e>0` by auto
+    }
+    hence "\<exists>d>0. f ` (ball x d \<inter> s) \<subseteq> ball (f x) e" using `d>0` unfolding subset_eq ball_def by (auto simp add: dist_commute)  }
+  thus ?rhs by auto
+next
+  assume ?rhs thus ?lhs unfolding continuous_within Lim_within ball_def subset_eq
+    apply (auto simp add: dist_commute) apply(erule_tac x=e in allE) by auto
+qed
+
+lemma continuous_at_ball:
+  "continuous (at x) f \<longleftrightarrow> (\<forall>e>0. \<exists>d>0. f ` (ball x d) \<subseteq> ball (f x) e)" (is "?lhs = ?rhs")
+proof
+  assume ?lhs thus ?rhs unfolding continuous_at Lim_at subset_eq Ball_def Bex_def image_iff mem_ball
+    apply auto apply(erule_tac x=e in allE) apply auto apply(rule_tac x=d in exI) apply auto apply(erule_tac x=xa in allE) apply (auto simp add: dist_commute dist_nz)
+    unfolding dist_nz[THEN sym] by auto
+next
+  assume ?rhs thus ?lhs unfolding continuous_at Lim_at subset_eq Ball_def Bex_def image_iff mem_ball
+    apply auto apply(erule_tac x=e in allE) apply auto apply(rule_tac x=d in exI) apply auto apply(erule_tac x="f xa" in allE) by (auto simp add: dist_commute dist_nz)
+qed
+
+text{* For setwise continuity, just start from the epsilon-delta definitions. *}
+
+definition
+  continuous_on :: "'a::metric_space set \<Rightarrow> ('a \<Rightarrow> 'b::metric_space) \<Rightarrow> bool" where
+  "continuous_on s f \<longleftrightarrow> (\<forall>x \<in> s. \<forall>e>0. \<exists>d::real>0. \<forall>x' \<in> s. dist x' x < d --> dist (f x') (f x) < e)"
+
+
+definition
+  uniformly_continuous_on ::
+    "'a::metric_space set \<Rightarrow> ('a \<Rightarrow> 'b::metric_space) \<Rightarrow> bool" where
+  "uniformly_continuous_on s f \<longleftrightarrow>
+        (\<forall>e>0. \<exists>d>0. \<forall>x\<in>s. \<forall> x'\<in>s. dist x' x < d
+                           --> dist (f x') (f x) < e)"
+
+text{* Some simple consequential lemmas. *}
+
+lemma uniformly_continuous_imp_continuous:
+ " uniformly_continuous_on s f ==> continuous_on s f"
+  unfolding uniformly_continuous_on_def continuous_on_def by blast
+
+lemma continuous_at_imp_continuous_within:
+ "continuous (at x) f ==> continuous (at x within s) f"
+  unfolding continuous_within continuous_at using Lim_at_within by auto
+
+lemma continuous_at_imp_continuous_on: assumes "(\<forall>x \<in> s. continuous (at x) f)"
+  shows "continuous_on s f"
+proof(simp add: continuous_at continuous_on_def, rule, rule, rule)
+  fix x and e::real assume "x\<in>s" "e>0"
+  hence "eventually (\<lambda>xa. dist (f xa) (f x) < e) (at x)" using assms unfolding continuous_at tendsto_iff by auto
+  then obtain d where d:"d>0" "\<forall>xa. 0 < dist xa x \<and> dist xa x < d \<longrightarrow> dist (f xa) (f x) < e" unfolding eventually_at by auto
+  { fix x' assume "\<not> 0 < dist x' x"
+    hence "x=x'"
+      using dist_nz[of x' x] by auto
+    hence "dist (f x') (f x) < e" using `e>0` by auto
+  }
+  thus "\<exists>d>0. \<forall>x'\<in>s. dist x' x < d \<longrightarrow> dist (f x') (f x) < e" using d by auto
+qed
+
+lemma continuous_on_eq_continuous_within:
+ "continuous_on s f \<longleftrightarrow> (\<forall>x \<in> s. continuous (at x within s) f)" (is "?lhs = ?rhs")
+proof
+  assume ?rhs
+  { fix x assume "x\<in>s"
+    fix e::real assume "e>0"
+    assume "\<exists>d>0. \<forall>xa\<in>s. 0 < dist xa x \<and> dist xa x < d \<longrightarrow> dist (f xa) (f x) < e"
+    then obtain d where "d>0" and d:"\<forall>xa\<in>s. 0 < dist xa x \<and> dist xa x < d \<longrightarrow> dist (f xa) (f x) < e" by auto
+    { fix x' assume as:"x'\<in>s" "dist x' x < d"
+      hence "dist (f x') (f x) < e" using `e>0` d `x'\<in>s` dist_eq_0_iff[of x' x] zero_le_dist[of x' x] as(2) by (metis dist_eq_0_iff dist_nz) }
+    hence "\<exists>d>0. \<forall>x'\<in>s. dist x' x < d \<longrightarrow> dist (f x') (f x) < e" using `d>0` by auto
+  }
+  thus ?lhs using `?rhs` unfolding continuous_on_def continuous_within Lim_within by auto
+next
+  assume ?lhs
+  thus ?rhs unfolding continuous_on_def continuous_within Lim_within by blast
+qed
+
+lemma continuous_on:
+ "continuous_on s f \<longleftrightarrow> (\<forall>x \<in> s. (f ---> f(x)) (at x within s))"
+  by (auto simp add: continuous_on_eq_continuous_within continuous_within)
+
+lemma continuous_on_eq_continuous_at:
+ "open s ==> (continuous_on s f \<longleftrightarrow> (\<forall>x \<in> s. continuous (at x) f))"
+  by (auto simp add: continuous_on continuous_at Lim_within_open)
+
+lemma continuous_within_subset:
+ "continuous (at x within s) f \<Longrightarrow> t \<subseteq> s
+             ==> continuous (at x within t) f"
+  unfolding continuous_within by(metis Lim_within_subset)
+
+lemma continuous_on_subset:
+ "continuous_on s f \<Longrightarrow> t \<subseteq> s ==> continuous_on t f"
+  unfolding continuous_on by (metis subset_eq Lim_within_subset)
+
+lemma continuous_on_interior:
+ "continuous_on s f \<Longrightarrow> x \<in> interior s ==> continuous (at x) f"
+unfolding interior_def
+apply simp
+by (meson continuous_on_eq_continuous_at continuous_on_subset)
+
+lemma continuous_on_eq:
+ "(\<forall>x \<in> s. f x = g x) \<Longrightarrow> continuous_on s f
+           ==> continuous_on s g"
+  by (simp add: continuous_on_def)
+
+text{* Characterization of various kinds of continuity in terms of sequences.  *}
+
+(* \<longrightarrow> could be generalized, but \<longleftarrow> requires metric space *)
+lemma continuous_within_sequentially:
+  fixes f :: "'a::metric_space \<Rightarrow> 'b::metric_space"
+  shows "continuous (at a within s) f \<longleftrightarrow>
+                (\<forall>x. (\<forall>n::nat. x n \<in> s) \<and> (x ---> a) sequentially
+                     --> ((f o x) ---> f a) sequentially)" (is "?lhs = ?rhs")
+proof
+  assume ?lhs
+  { fix x::"nat \<Rightarrow> 'a" assume x:"\<forall>n. x n \<in> s" "\<forall>e>0. \<exists>N. \<forall>n\<ge>N. dist (x n) a < e"
+    fix e::real assume "e>0"
+    from `?lhs` obtain d where "d>0" and d:"\<forall>x\<in>s. 0 < dist x a \<and> dist x a < d \<longrightarrow> dist (f x) (f a) < e" unfolding continuous_within Lim_within using `e>0` by auto
+    from x(2) `d>0` obtain N where N:"\<forall>n\<ge>N. dist (x n) a < d" by auto
+    hence "\<exists>N. \<forall>n\<ge>N. dist ((f \<circ> x) n) (f a) < e"
+      apply(rule_tac  x=N in exI) using N d  apply auto using x(1)
+      apply(erule_tac x=n in allE) apply(erule_tac x=n in allE)
+      apply(erule_tac x="x n" in ballE)  apply auto unfolding dist_nz[THEN sym] apply auto using `e>0` by auto
+  }
+  thus ?rhs unfolding continuous_within unfolding Lim_sequentially by simp
+next
+  assume ?rhs
+  { fix e::real assume "e>0"
+    assume "\<not> (\<exists>d>0. \<forall>x\<in>s. 0 < dist x a \<and> dist x a < d \<longrightarrow> dist (f x) (f a) < e)"
+    hence "\<forall>d. \<exists>x. d>0 \<longrightarrow> x\<in>s \<and> (0 < dist x a \<and> dist x a < d \<and> \<not> dist (f x) (f a) < e)" by blast
+    then obtain x where x:"\<forall>d>0. x d \<in> s \<and> (0 < dist (x d) a \<and> dist (x d) a < d \<and> \<not> dist (f (x d)) (f a) < e)"
+      using choice[of "\<lambda>d x.0<d \<longrightarrow> x\<in>s \<and> (0 < dist x a \<and> dist x a < d \<and> \<not> dist (f x) (f a) < e)"] by auto
+    { fix d::real assume "d>0"
+      hence "\<exists>N::nat. inverse (real (N + 1)) < d" using real_arch_inv[of d] by (auto, rule_tac x="n - 1" in exI)auto
+      then obtain N::nat where N:"inverse (real (N + 1)) < d" by auto
+      { fix n::nat assume n:"n\<ge>N"
+	hence "dist (x (inverse (real (n + 1)))) a < inverse (real (n + 1))" using x[THEN spec[where x="inverse (real (n + 1))"]] by auto
+	moreover have "inverse (real (n + 1)) < d" using N n by (auto, metis Suc_le_mono le_SucE less_imp_inverse_less nat_le_real_less order_less_trans real_of_nat_Suc real_of_nat_Suc_gt_zero)
+	ultimately have "dist (x (inverse (real (n + 1)))) a < d" by auto
+      }
+      hence "\<exists>N::nat. \<forall>n\<ge>N. dist (x (inverse (real (n + 1)))) a < d" by auto
+    }
+    hence "(\<forall>n::nat. x (inverse (real (n + 1))) \<in> s) \<and> (\<forall>e>0. \<exists>N::nat. \<forall>n\<ge>N. dist (x (inverse (real (n + 1)))) a < e)" using x by auto
+    hence "\<forall>e>0. \<exists>N::nat. \<forall>n\<ge>N. dist (f (x (inverse (real (n + 1))))) (f a) < e"  using `?rhs`[THEN spec[where x="\<lambda>n::nat. x (inverse (real (n+1)))"], unfolded Lim_sequentially] by auto
+    hence "False" apply(erule_tac x=e in allE) using `e>0` using x by auto
+  }
+  thus ?lhs  unfolding continuous_within unfolding Lim_within unfolding Lim_sequentially by blast
+qed
+
+lemma continuous_at_sequentially:
+  fixes f :: "'a::metric_space \<Rightarrow> 'b::metric_space"
+  shows "continuous (at a) f \<longleftrightarrow> (\<forall>x. (x ---> a) sequentially
+                  --> ((f o x) ---> f a) sequentially)"
+  using continuous_within_sequentially[of a UNIV f] unfolding within_UNIV by auto
+
+lemma continuous_on_sequentially:
+ "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)" (is "?lhs = ?rhs")
+proof
+  assume ?rhs thus ?lhs using continuous_within_sequentially[of _ s f] unfolding continuous_on_eq_continuous_within by auto
+next
+  assume ?lhs thus ?rhs unfolding continuous_on_eq_continuous_within using continuous_within_sequentially[of _ s f] by auto
+qed
+
+lemma uniformly_continuous_on_sequentially:
+  fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
+  shows "uniformly_continuous_on s f \<longleftrightarrow> (\<forall>x y. (\<forall>n. x n \<in> s) \<and> (\<forall>n. y n \<in> s) \<and>
+                    ((\<lambda>n. x n - y n) ---> 0) sequentially
+                    \<longrightarrow> ((\<lambda>n. f(x n) - f(y n)) ---> 0) sequentially)" (is "?lhs = ?rhs")
+proof
+  assume ?lhs
+  { fix x y assume x:"\<forall>n. x n \<in> s" and y:"\<forall>n. y n \<in> s" and xy:"((\<lambda>n. x n - y n) ---> 0) sequentially"
+    { fix e::real assume "e>0"
+      then obtain d where "d>0" and d:"\<forall>x\<in>s. \<forall>x'\<in>s. dist x' x < d \<longrightarrow> dist (f x') (f x) < e"
+	using `?lhs`[unfolded uniformly_continuous_on_def, THEN spec[where x=e]] by auto
+      obtain N where N:"\<forall>n\<ge>N. norm (x n - y n - 0) < d" using xy[unfolded Lim_sequentially dist_norm] and `d>0` by auto
+      { fix n assume "n\<ge>N"
+	hence "norm (f (x n) - f (y n) - 0) < e"
+	  using N[THEN spec[where x=n]] using d[THEN bspec[where x="x n"], THEN bspec[where x="y n"]] using x and y
+	  unfolding dist_commute and dist_norm by simp  }
+      hence "\<exists>N. \<forall>n\<ge>N. norm (f (x n) - f (y n) - 0) < e"  by auto  }
+    hence "((\<lambda>n. f(x n) - f(y n)) ---> 0) sequentially" unfolding Lim_sequentially and dist_norm by auto  }
+  thus ?rhs by auto
+next
+  assume ?rhs
+  { assume "\<not> ?lhs"
+    then obtain e where "e>0" "\<forall>d>0. \<exists>x\<in>s. \<exists>x'\<in>s. dist x' x < d \<and> \<not> dist (f x') (f x) < e" unfolding uniformly_continuous_on_def by auto
+    then obtain fa where fa:"\<forall>x.  0 < x \<longrightarrow> fst (fa x) \<in> s \<and> snd (fa x) \<in> s \<and> dist (fst (fa x)) (snd (fa x)) < x \<and> \<not> dist (f (fst (fa x))) (f (snd (fa x))) < e"
+      using choice[of "\<lambda>d x. d>0 \<longrightarrow> fst x \<in> s \<and> snd x \<in> s \<and> dist (snd x) (fst x) < d \<and> \<not> dist (f (snd x)) (f (fst x)) < e"] unfolding Bex_def
+      by (auto simp add: dist_commute)
+    def x \<equiv> "\<lambda>n::nat. fst (fa (inverse (real n + 1)))"
+    def y \<equiv> "\<lambda>n::nat. snd (fa (inverse (real n + 1)))"
+    have xyn:"\<forall>n. x n \<in> s \<and> y n \<in> s" and xy0:"\<forall>n. dist (x n) (y n) < inverse (real n + 1)" and fxy:"\<forall>n. \<not> dist (f (x n)) (f (y n)) < e"
+      unfolding x_def and y_def using fa by auto
+    have 1:"\<And>(x::'a) y. dist (x - y) 0 = dist x y" unfolding dist_norm by auto
+    have 2:"\<And>(x::'b) y. dist (x - y) 0 = dist x y" unfolding dist_norm by auto
+    { fix e::real assume "e>0"
+      then obtain N::nat where "N \<noteq> 0" and N:"0 < inverse (real N) \<and> inverse (real N) < e" unfolding real_arch_inv[of e]   by auto
+      { fix n::nat assume "n\<ge>N"
+	hence "inverse (real n + 1) < inverse (real N)" using real_of_nat_ge_zero and `N\<noteq>0` by auto
+	also have "\<dots> < e" using N by auto
+	finally have "inverse (real n + 1) < e" by auto
+	hence "dist (x n - y n) 0 < e" unfolding 1 using xy0[THEN spec[where x=n]] by auto  }
+      hence "\<exists>N. \<forall>n\<ge>N. dist (x n - y n) 0 < e" by auto  }
+    hence "\<forall>e>0. \<exists>N. \<forall>n\<ge>N. dist (f (x n) - f (y n)) 0 < e" using `?rhs`[THEN spec[where x=x], THEN spec[where x=y]] and xyn unfolding Lim_sequentially by auto
+    hence False unfolding 2 using fxy and `e>0` by auto  }
+  thus ?lhs unfolding uniformly_continuous_on_def by blast
+qed
+
+text{* The usual transformation theorems. *}
+
+lemma continuous_transform_within:
+  fixes f g :: "'a::metric_space \<Rightarrow> 'b::metric_space"
+  assumes "0 < d" "x \<in> s" "\<forall>x' \<in> s. dist x' x < d --> f x' = g x'"
+          "continuous (at x within s) f"
+  shows "continuous (at x within s) g"
+proof-
+  { fix e::real assume "e>0"
+    then obtain d' where d':"d'>0" "\<forall>xa\<in>s. 0 < dist xa x \<and> dist xa x < d' \<longrightarrow> dist (f xa) (f x) < e" using assms(4) unfolding continuous_within Lim_within by auto
+    { fix x' assume "x'\<in>s" "0 < dist x' x" "dist x' x < (min d d')"
+      hence "dist (f x') (g x) < e" using assms(2,3) apply(erule_tac x=x in ballE) using d' by auto  }
+    hence "\<forall>xa\<in>s. 0 < dist xa x \<and> dist xa x < (min d d') \<longrightarrow> dist (f xa) (g x) < e" by blast
+    hence "\<exists>d>0. \<forall>xa\<in>s. 0 < dist xa x \<and> dist xa x < d \<longrightarrow> dist (f xa) (g x) < e" using `d>0` `d'>0` by(rule_tac x="min d d'" in exI)auto  }
+  hence "(f ---> g x) (at x within s)" unfolding Lim_within using assms(1) by auto
+  thus ?thesis unfolding continuous_within using Lim_transform_within[of d s x f g "g x"] using assms by blast
+qed
+
+lemma continuous_transform_at:
+  fixes f g :: "'a::metric_space \<Rightarrow> 'b::metric_space"
+  assumes "0 < d" "\<forall>x'. dist x' x < d --> f x' = g x'"
+          "continuous (at x) f"
+  shows "continuous (at x) g"
+proof-
+  { fix e::real assume "e>0"
+    then obtain d' where d':"d'>0" "\<forall>xa. 0 < dist xa x \<and> dist xa x < d' \<longrightarrow> dist (f xa) (f x) < e" using assms(3) unfolding continuous_at Lim_at by auto
+    { fix x' assume "0 < dist x' x" "dist x' x < (min d d')"
+      hence "dist (f x') (g x) < e" using assms(2) apply(erule_tac x=x in allE) using d' by auto
+    }
+    hence "\<forall>xa. 0 < dist xa x \<and> dist xa x < (min d d') \<longrightarrow> dist (f xa) (g x) < e" by blast
+    hence "\<exists>d>0. \<forall>xa. 0 < dist xa x \<and> dist xa x < d \<longrightarrow> dist (f xa) (g x) < e" using `d>0` `d'>0` by(rule_tac x="min d d'" in exI)auto
+  }
+  hence "(f ---> g x) (at x)" unfolding Lim_at using assms(1) by auto
+  thus ?thesis unfolding continuous_at using Lim_transform_at[of d x f g "g x"] using assms by blast
+qed
+
+text{* Combination results for pointwise continuity. *}
+
+lemma continuous_const: "continuous net (\<lambda>x. c)"
+  by (auto simp add: continuous_def Lim_const)
+
+lemma continuous_cmul:
+  fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector"
+  shows "continuous net f ==> continuous net (\<lambda>x. c *\<^sub>R f x)"
+  by (auto simp add: continuous_def Lim_cmul)
+
+lemma continuous_neg:
+  fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector"
+  shows "continuous net f ==> continuous net (\<lambda>x. -(f x))"
+  by (auto simp add: continuous_def Lim_neg)
+
+lemma continuous_add:
+  fixes f g :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector"
+  shows "continuous net f \<Longrightarrow> continuous net g \<Longrightarrow> continuous net (\<lambda>x. f x + g x)"
+  by (auto simp add: continuous_def Lim_add)
+
+lemma continuous_sub:
+  fixes f g :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector"
+  shows "continuous net f \<Longrightarrow> continuous net g \<Longrightarrow> continuous net (\<lambda>x. f x - g x)"
+  by (auto simp add: continuous_def Lim_sub)
+
+text{* Same thing for setwise continuity. *}
+
+lemma continuous_on_const:
+ "continuous_on s (\<lambda>x. c)"
+  unfolding continuous_on_eq_continuous_within using continuous_const by blast
+
+lemma continuous_on_cmul:
+  fixes f :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector"
+  shows "continuous_on s f ==>  continuous_on s (\<lambda>x. c *\<^sub>R (f x))"
+  unfolding continuous_on_eq_continuous_within using continuous_cmul by blast
+
+lemma continuous_on_neg:
+  fixes f :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector"
+  shows "continuous_on s f \<Longrightarrow> continuous_on s (\<lambda>x. - f x)"
+  unfolding continuous_on_eq_continuous_within using continuous_neg by blast
+
+lemma continuous_on_add:
+  fixes f g :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector"
+  shows "continuous_on s f \<Longrightarrow> continuous_on s g
+           \<Longrightarrow> continuous_on s (\<lambda>x. f x + g x)"
+  unfolding continuous_on_eq_continuous_within using continuous_add by blast
+
+lemma continuous_on_sub:
+  fixes f g :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector"
+  shows "continuous_on s f \<Longrightarrow> continuous_on s g
+           \<Longrightarrow> continuous_on s (\<lambda>x. f x - g x)"
+  unfolding continuous_on_eq_continuous_within using continuous_sub by blast
+
+text{* Same thing for uniform continuity, using sequential formulations. *}
+
+lemma uniformly_continuous_on_const:
+ "uniformly_continuous_on s (\<lambda>x. c)"
+  unfolding uniformly_continuous_on_def by simp
+
+lemma uniformly_continuous_on_cmul:
+  fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
+    (* FIXME: generalize 'a to metric_space *)
+  assumes "uniformly_continuous_on s f"
+  shows "uniformly_continuous_on s (\<lambda>x. c *\<^sub>R f(x))"
+proof-
+  { fix x y assume "((\<lambda>n. f (x n) - f (y n)) ---> 0) sequentially"
+    hence "((\<lambda>n. c *\<^sub>R f (x n) - c *\<^sub>R f (y n)) ---> 0) sequentially"
+      using Lim_cmul[of "(\<lambda>n. f (x n) - f (y n))" 0 sequentially c]
+      unfolding scaleR_zero_right scaleR_right_diff_distrib by auto
+  }
+  thus ?thesis using assms unfolding uniformly_continuous_on_sequentially by auto
+qed
+
+lemma dist_minus:
+  fixes x y :: "'a::real_normed_vector"
+  shows "dist (- x) (- y) = dist x y"
+  unfolding dist_norm minus_diff_minus norm_minus_cancel ..
+
+lemma uniformly_continuous_on_neg:
+  fixes f :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector"
+  shows "uniformly_continuous_on s f
+         ==> uniformly_continuous_on s (\<lambda>x. -(f x))"
+  unfolding uniformly_continuous_on_def dist_minus .
+
+lemma uniformly_continuous_on_add:
+  fixes f g :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector" (* FIXME: generalize 'a *)
+  assumes "uniformly_continuous_on s f" "uniformly_continuous_on s g"
+  shows "uniformly_continuous_on s (\<lambda>x. f x + g x)"
+proof-
+  {  fix x y assume "((\<lambda>n. f (x n) - f (y n)) ---> 0) sequentially"
+                    "((\<lambda>n. g (x n) - g (y n)) ---> 0) sequentially"
+    hence "((\<lambda>xa. f (x xa) - f (y xa) + (g (x xa) - g (y xa))) ---> 0 + 0) sequentially"
+      using Lim_add[of "\<lambda> n. f (x n) - f (y n)" 0  sequentially "\<lambda> n. g (x n) - g (y n)" 0] by auto
+    hence "((\<lambda>n. f (x n) + g (x n) - (f (y n) + g (y n))) ---> 0) sequentially" unfolding Lim_sequentially and add_diff_add [symmetric] by auto  }
+  thus ?thesis using assms unfolding uniformly_continuous_on_sequentially by auto
+qed
+
+lemma uniformly_continuous_on_sub:
+  fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector" (* FIXME: generalize 'a *)
+  shows "uniformly_continuous_on s f \<Longrightarrow> uniformly_continuous_on s g
+           ==> uniformly_continuous_on s  (\<lambda>x. f x - g x)"
+  unfolding ab_diff_minus
+  using uniformly_continuous_on_add[of s f "\<lambda>x. - g x"]
+  using uniformly_continuous_on_neg[of s g] by auto
+
+text{* Identity function is continuous in every sense. *}
+
+lemma continuous_within_id:
+ "continuous (at a within s) (\<lambda>x. x)"
+  unfolding continuous_within by (rule Lim_at_within [OF Lim_ident_at])
+
+lemma continuous_at_id:
+ "continuous (at a) (\<lambda>x. x)"
+  unfolding continuous_at by (rule Lim_ident_at)
+
+lemma continuous_on_id:
+ "continuous_on s (\<lambda>x. x)"
+  unfolding continuous_on Lim_within by auto
+
+lemma uniformly_continuous_on_id:
+ "uniformly_continuous_on s (\<lambda>x. x)"
+  unfolding uniformly_continuous_on_def by auto
+
+text{* Continuity of all kinds is preserved under composition. *}
+
+lemma continuous_within_compose:
+  fixes f :: "'a::metric_space \<Rightarrow> 'b::metric_space" (* FIXME: generalize *)
+  fixes g :: "'b::metric_space \<Rightarrow> 'c::metric_space"
+  assumes "continuous (at x within s) f"   "continuous (at (f x) within f ` s) g"
+  shows "continuous (at x within s) (g o f)"
+proof-
+  { fix e::real assume "e>0"
+    with assms(2)[unfolded continuous_within Lim_within] obtain d  where "d>0" and d:"\<forall>xa\<in>f ` s. 0 < dist xa (f x) \<and> dist xa (f x) < d \<longrightarrow> dist (g xa) (g (f x)) < e" by auto
+    from assms(1)[unfolded continuous_within Lim_within] obtain d' where "d'>0" and d':"\<forall>xa\<in>s. 0 < dist xa x \<and> dist xa x < d' \<longrightarrow> dist (f xa) (f x) < d" using `d>0` by auto
+    { fix y assume as:"y\<in>s"  "0 < dist y x"  "dist y x < d'"
+      hence "dist (f y) (f x) < d" using d'[THEN bspec[where x=y]] by (auto simp add:dist_commute)
+      hence "dist (g (f y)) (g (f x)) < e" using as(1) d[THEN bspec[where x="f y"]] unfolding dist_nz[THEN sym] using `e>0` by auto   }
+    hence "\<exists>d>0. \<forall>xa\<in>s. 0 < dist xa x \<and> dist xa x < d \<longrightarrow> dist (g (f xa)) (g (f x)) < e" using `d'>0` by auto  }
+  thus ?thesis unfolding continuous_within Lim_within by auto
+qed
+
+lemma continuous_at_compose:
+  fixes f :: "'a::metric_space \<Rightarrow> 'b::metric_space" (* FIXME: generalize *)
+  fixes g :: "'b::metric_space \<Rightarrow> 'c::metric_space"
+  assumes "continuous (at x) f"  "continuous (at (f x)) g"
+  shows "continuous (at x) (g o f)"
+proof-
+  have " continuous (at (f x) within range f) g" using assms(2) using continuous_within_subset[of "f x" UNIV g "range f", unfolded within_UNIV] by auto
+  thus ?thesis using assms(1) using continuous_within_compose[of x UNIV f g, unfolded within_UNIV] by auto
+qed
+
+lemma continuous_on_compose:
+ "continuous_on s f \<Longrightarrow> continuous_on (f ` s) g \<Longrightarrow> continuous_on s (g o f)"
+  unfolding continuous_on_eq_continuous_within using continuous_within_compose[of _ s f g] by auto
+
+lemma uniformly_continuous_on_compose:
+  assumes "uniformly_continuous_on s f"  "uniformly_continuous_on (f ` s) g"
+  shows "uniformly_continuous_on s (g o f)"
+proof-
+  { fix e::real assume "e>0"
+    then obtain d where "d>0" and d:"\<forall>x\<in>f ` s. \<forall>x'\<in>f ` s. dist x' x < d \<longrightarrow> dist (g x') (g x) < e" using assms(2) unfolding uniformly_continuous_on_def by auto
+    obtain d' where "d'>0" "\<forall>x\<in>s. \<forall>x'\<in>s. dist x' x < d' \<longrightarrow> dist (f x') (f x) < d" using `d>0` using assms(1) unfolding uniformly_continuous_on_def by auto
+    hence "\<exists>d>0. \<forall>x\<in>s. \<forall>x'\<in>s. dist x' x < d \<longrightarrow> dist ((g \<circ> f) x') ((g \<circ> f) x) < e" using `d>0` using d by auto  }
+  thus ?thesis using assms unfolding uniformly_continuous_on_def by auto
+qed
+
+text{* Continuity in terms of open preimages. *}
+
+lemma continuous_at_open:
+  fixes f :: "'a::metric_space \<Rightarrow> 'b::metric_space" (* FIXME: generalize *)
+  shows "continuous (at x) f \<longleftrightarrow> (\<forall>t. open t \<and> f x \<in> t --> (\<exists>s. open s \<and> x \<in> s \<and> (\<forall>x' \<in> s. (f x') \<in> t)))" (is "?lhs = ?rhs")
+proof
+  assume ?lhs
+  { fix t assume as: "open t" "f x \<in> t"
+    then obtain e where "e>0" and e:"ball (f x) e \<subseteq> t" unfolding open_contains_ball by auto
+
+    obtain d where "d>0" and d:"\<forall>y. 0 < dist y x \<and> dist y x < d \<longrightarrow> dist (f y) (f x) < e" using `e>0` using `?lhs`[unfolded continuous_at Lim_at open_dist] by auto
+
+    have "open (ball x d)" using open_ball by auto
+    moreover have "x \<in> ball x d" unfolding centre_in_ball using `d>0` by simp
+    moreover
+    { fix x' assume "x'\<in>ball x d" hence "f x' \<in> t"
+	using e[unfolded subset_eq Ball_def mem_ball, THEN spec[where x="f x'"]]    d[THEN spec[where x=x']]
+	unfolding mem_ball apply (auto simp add: dist_commute)
+	unfolding dist_nz[THEN sym] using as(2) by auto  }
+    hence "\<forall>x'\<in>ball x d. f x' \<in> t" by auto
+    ultimately have "\<exists>s. open s \<and> x \<in> s \<and> (\<forall>x'\<in>s. f x' \<in> t)"
+      apply(rule_tac x="ball x d" in exI) by simp  }
+  thus ?rhs by auto
+next
+  assume ?rhs
+  { fix e::real assume "e>0"
+    then obtain s where s: "open s"  "x \<in> s"  "\<forall>x'\<in>s. f x' \<in> ball (f x) e" using `?rhs`[unfolded continuous_at Lim_at, THEN spec[where x="ball (f x) e"]]
+      unfolding centre_in_ball[of "f x" e, THEN sym] by auto
+    then obtain d where "d>0" and d:"ball x d \<subseteq> s" unfolding open_contains_ball by auto
+    { fix y assume "0 < dist y x \<and> dist y x < d"
+      hence "dist (f y) (f x) < e" using d[unfolded subset_eq Ball_def mem_ball, THEN spec[where x=y]]
+	using s(3)[THEN bspec[where x=y], unfolded mem_ball] by (auto simp add: dist_commute)  }
+    hence "\<exists>d>0. \<forall>xa. 0 < dist xa x \<and> dist xa x < d \<longrightarrow> dist (f xa) (f x) < e" using `d>0` by auto  }
+  thus ?lhs unfolding continuous_at Lim_at by auto
+qed
+
+lemma continuous_on_open:
+ "continuous_on s f \<longleftrightarrow>
+        (\<forall>t. openin (subtopology euclidean (f ` s)) t
+            --> openin (subtopology euclidean s) {x \<in> s. f x \<in> t})" (is "?lhs = ?rhs")
+proof
+  assume ?lhs
+  { fix t assume as:"openin (subtopology euclidean (f ` s)) t"
+    have "{x \<in> s. f x \<in> t} \<subseteq> s" using as[unfolded openin_euclidean_subtopology_iff] by auto
+    moreover
+    { fix x assume as':"x\<in>{x \<in> s. f x \<in> t}"
+      then obtain e where e: "e>0" "\<forall>x'\<in>f ` s. dist x' (f x) < e \<longrightarrow> x' \<in> t" using as[unfolded openin_euclidean_subtopology_iff, THEN conjunct2, THEN bspec[where x="f x"]] by auto
+      from this(1) obtain d where d: "d>0" "\<forall>xa\<in>s. 0 < dist xa x \<and> dist xa x < d \<longrightarrow> dist (f xa) (f x) < e" using `?lhs`[unfolded continuous_on Lim_within, THEN bspec[where x=x]] using as' by auto
+      have "\<exists>e>0. \<forall>x'\<in>s. dist x' x < e \<longrightarrow> x' \<in> {x \<in> s. f x \<in> t}" using d e unfolding dist_nz[THEN sym] by (rule_tac x=d in exI, auto)  }
+    ultimately have "openin (subtopology euclidean s) {x \<in> s. f x \<in> t}" unfolding openin_euclidean_subtopology_iff by auto  }
+  thus ?rhs unfolding continuous_on Lim_within using openin by auto
+next
+  assume ?rhs
+  { fix e::real and x assume "x\<in>s" "e>0"
+    { fix xa x' assume "dist (f xa) (f x) < e" "xa \<in> s" "x' \<in> s" "dist (f xa) (f x') < e - dist (f xa) (f x)"
+      hence "dist (f x') (f x) < e" using dist_triangle[of "f x'" "f x" "f xa"]
+	by (auto simp add: dist_commute)  }
+    hence "ball (f x) e \<inter> f ` s \<subseteq> f ` s \<and> (\<forall>xa\<in>ball (f x) e \<inter> f ` s. \<exists>ea>0. \<forall>x'\<in>f ` s. dist x' xa < ea \<longrightarrow> x' \<in> ball (f x) e \<inter> f ` s)" apply auto
+      apply(rule_tac x="e - dist (f xa) (f x)" in exI) using `e>0` by (auto simp add: dist_commute)
+    hence "\<forall>xa\<in>{xa \<in> s. f xa \<in> ball (f x) e \<inter> f ` s}. \<exists>ea>0. \<forall>x'\<in>s. dist x' xa < ea \<longrightarrow> x' \<in> {xa \<in> s. f xa \<in> ball (f x) e \<inter> f ` s}"
+      using `?rhs`[unfolded openin_euclidean_subtopology_iff, THEN spec[where x="ball (f x) e \<inter> f ` s"]] by auto
+    hence "\<exists>d>0. \<forall>xa\<in>s. 0 < dist xa x \<and> dist xa x < d \<longrightarrow> dist (f xa) (f x) < e" apply(erule_tac x=x in ballE) apply auto using `e>0` `x\<in>s` by (auto simp add: dist_commute)  }
+  thus ?lhs unfolding continuous_on Lim_within by auto
+qed
+
+(* ------------------------------------------------------------------------- *)
+(* Similarly in terms of closed sets.                                        *)
+(* ------------------------------------------------------------------------- *)
+
+lemma continuous_on_closed:
+ "continuous_on s f \<longleftrightarrow>  (\<forall>t. closedin (subtopology euclidean (f ` s)) t  --> closedin (subtopology euclidean s) {x \<in> s. f x \<in> t})" (is "?lhs = ?rhs")
+proof
+  assume ?lhs
+  { fix t
+    have *:"s - {x \<in> s. f x \<in> f ` s - t} = {x \<in> s. f x \<in> t}" by auto
+    have **:"f ` s - (f ` s - (f ` s - t)) = f ` s - t" by auto
+    assume as:"closedin (subtopology euclidean (f ` s)) t"
+    hence "closedin (subtopology euclidean (f ` s)) (f ` s - (f ` s - t))" unfolding closedin_def topspace_euclidean_subtopology unfolding ** by auto
+    hence "closedin (subtopology euclidean s) {x \<in> s. f x \<in> t}" using `?lhs`[unfolded continuous_on_open, THEN spec[where x="(f ` s) - t"]]
+      unfolding openin_closedin_eq topspace_euclidean_subtopology unfolding * by auto  }
+  thus ?rhs by auto
+next
+  assume ?rhs
+  { fix t
+    have *:"s - {x \<in> s. f x \<in> f ` s - t} = {x \<in> s. f x \<in> t}" by auto
+    assume as:"openin (subtopology euclidean (f ` s)) t"
+    hence "openin (subtopology euclidean s) {x \<in> s. f x \<in> t}" using `?rhs`[THEN spec[where x="(f ` s) - t"]]
+      unfolding openin_closedin_eq topspace_euclidean_subtopology *[THEN sym] closedin_subtopology by auto }
+  thus ?lhs unfolding continuous_on_open by auto
+qed
+
+text{* Half-global and completely global cases.                                  *}
+
+lemma continuous_open_in_preimage:
+  assumes "continuous_on s f"  "open t"
+  shows "openin (subtopology euclidean s) {x \<in> s. f x \<in> t}"
+proof-
+  have *:"\<forall>x. x \<in> s \<and> f x \<in> t \<longleftrightarrow> x \<in> s \<and> f x \<in> (t \<inter> f ` s)" by auto
+  have "openin (subtopology euclidean (f ` s)) (t \<inter> f ` s)"
+    using openin_open_Int[of t "f ` s", OF assms(2)] unfolding openin_open by auto
+  thus ?thesis using assms(1)[unfolded continuous_on_open, THEN spec[where x="t \<inter> f ` s"]] using * by auto
+qed
+
+lemma continuous_closed_in_preimage:
+  assumes "continuous_on s f"  "closed t"
+  shows "closedin (subtopology euclidean s) {x \<in> s. f x \<in> t}"
+proof-
+  have *:"\<forall>x. x \<in> s \<and> f x \<in> t \<longleftrightarrow> x \<in> s \<and> f x \<in> (t \<inter> f ` s)" by auto
+  have "closedin (subtopology euclidean (f ` s)) (t \<inter> f ` s)"
+    using closedin_closed_Int[of t "f ` s", OF assms(2)] unfolding Int_commute by auto
+  thus ?thesis
+    using assms(1)[unfolded continuous_on_closed, THEN spec[where x="t \<inter> f ` s"]] using * by auto
+qed
+
+lemma continuous_open_preimage:
+  assumes "continuous_on s f" "open s" "open t"
+  shows "open {x \<in> s. f x \<in> t}"
+proof-
+  obtain T where T: "open T" "{x \<in> s. f x \<in> t} = s \<inter> T"
+    using continuous_open_in_preimage[OF assms(1,3)] unfolding openin_open by auto
+  thus ?thesis using open_Int[of s T, OF assms(2)] by auto
+qed
+
+lemma continuous_closed_preimage:
+  assumes "continuous_on s f" "closed s" "closed t"
+  shows "closed {x \<in> s. f x \<in> t}"
+proof-
+  obtain T where T: "closed T" "{x \<in> s. f x \<in> t} = s \<inter> T"
+    using continuous_closed_in_preimage[OF assms(1,3)] unfolding closedin_closed by auto
+  thus ?thesis using closed_Int[of s T, OF assms(2)] by auto
+qed
+
+lemma continuous_open_preimage_univ:
+  fixes f :: "'a::metric_space \<Rightarrow> 'b::metric_space" (* FIXME: generalize *)
+  shows "\<forall>x. continuous (at x) f \<Longrightarrow> open s \<Longrightarrow> open {x. f x \<in> s}"
+  using continuous_open_preimage[of UNIV f s] open_UNIV continuous_at_imp_continuous_on by auto
+
+lemma continuous_closed_preimage_univ:
+  fixes f :: "'a::metric_space \<Rightarrow> 'b::metric_space" (* FIXME: generalize *)
+  shows "(\<forall>x. continuous (at x) f) \<Longrightarrow> closed s ==> closed {x. f x \<in> s}"
+  using continuous_closed_preimage[of UNIV f s] closed_UNIV continuous_at_imp_continuous_on by auto
+
+lemma continuous_open_vimage:
+  fixes f :: "'a::metric_space \<Rightarrow> 'b::metric_space" (* FIXME: generalize *)
+  shows "\<forall>x. continuous (at x) f \<Longrightarrow> open s \<Longrightarrow> open (f -` s)"
+  unfolding vimage_def by (rule continuous_open_preimage_univ)
+
+lemma continuous_closed_vimage:
+  fixes f :: "'a::metric_space \<Rightarrow> 'b::metric_space" (* FIXME: generalize *)
+  shows "\<forall>x. continuous (at x) f \<Longrightarrow> closed s \<Longrightarrow> closed (f -` s)"
+  unfolding vimage_def by (rule continuous_closed_preimage_univ)
+
+text{* Equality of continuous functions on closure and related results.          *}
+
+lemma continuous_closed_in_preimage_constant:
+ "continuous_on s f ==> closedin (subtopology euclidean s) {x \<in> s. f x = a}"
+  using continuous_closed_in_preimage[of s f "{a}"] closed_sing by auto
+
+lemma continuous_closed_preimage_constant:
+ "continuous_on s f \<Longrightarrow> closed s ==> closed {x \<in> s. f x = a}"
+  using continuous_closed_preimage[of s f "{a}"] closed_sing by auto
+
+lemma continuous_constant_on_closure:
+  assumes "continuous_on (closure s) f"
+          "\<forall>x \<in> s. f x = a"
+  shows "\<forall>x \<in> (closure s). f x = a"
+    using continuous_closed_preimage_constant[of "closure s" f a]
+    assms closure_minimal[of s "{x \<in> closure s. f x = a}"] closure_subset unfolding subset_eq by auto
+
+lemma image_closure_subset:
+  assumes "continuous_on (closure s) f"  "closed t"  "(f ` s) \<subseteq> t"
+  shows "f ` (closure s) \<subseteq> t"
+proof-
+  have "s \<subseteq> {x \<in> closure s. f x \<in> t}" using assms(3) closure_subset by auto
+  moreover have "closed {x \<in> closure s. f x \<in> t}"
+    using continuous_closed_preimage[OF assms(1)] and assms(2) by auto
+  ultimately have "closure s = {x \<in> closure s . f x \<in> t}"
+    using closure_minimal[of s "{x \<in> closure s. f x \<in> t}"] by auto
+  thus ?thesis by auto
+qed
+
+lemma continuous_on_closure_norm_le:
+  fixes f :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector"
+  assumes "continuous_on (closure s) f"  "\<forall>y \<in> s. norm(f y) \<le> b"  "x \<in> (closure s)"
+  shows "norm(f x) \<le> b"
+proof-
+  have *:"f ` s \<subseteq> cball 0 b" using assms(2)[unfolded mem_cball_0[THEN sym]] by auto
+  show ?thesis
+    using image_closure_subset[OF assms(1) closed_cball[of 0 b] *] assms(3)
+    unfolding subset_eq apply(erule_tac x="f x" in ballE) by (auto simp add: dist_norm)
+qed
+
+text{* Making a continuous function avoid some value in a neighbourhood.         *}
+
+lemma continuous_within_avoid:
+  fixes f :: "'a::metric_space \<Rightarrow> 'b::metric_space" (* FIXME: generalize *)
+  assumes "continuous (at x within s) f"  "x \<in> s"  "f x \<noteq> a"
+  shows "\<exists>e>0. \<forall>y \<in> s. dist x y < e --> f y \<noteq> a"
+proof-
+  obtain d where "d>0" and d:"\<forall>xa\<in>s. 0 < dist xa x \<and> dist xa x < d \<longrightarrow> dist (f xa) (f x) < dist (f x) a"
+    using assms(1)[unfolded continuous_within Lim_within, THEN spec[where x="dist (f x) a"]] assms(3)[unfolded dist_nz] by auto
+  { fix y assume " y\<in>s"  "dist x y < d"
+    hence "f y \<noteq> a" using d[THEN bspec[where x=y]] assms(3)[unfolded dist_nz]
+      apply auto unfolding dist_nz[THEN sym] by (auto simp add: dist_commute) }
+  thus ?thesis using `d>0` by auto
+qed
+
+lemma continuous_at_avoid:
+  fixes f :: "'a::metric_space \<Rightarrow> 'b::metric_space" (* FIXME: generalize *)
+  assumes "continuous (at x) f"  "f x \<noteq> a"
+  shows "\<exists>e>0. \<forall>y. dist x y < e \<longrightarrow> f y \<noteq> a"
+using assms using continuous_within_avoid[of x UNIV f a, unfolded within_UNIV] by auto
+
+lemma continuous_on_avoid:
+  assumes "continuous_on s f"  "x \<in> s"  "f x \<noteq> a"
+  shows "\<exists>e>0. \<forall>y \<in> s. dist x y < e \<longrightarrow> f y \<noteq> a"
+using assms(1)[unfolded continuous_on_eq_continuous_within, THEN bspec[where x=x], OF assms(2)]  continuous_within_avoid[of x s f a]  assms(2,3) by auto
+
+lemma continuous_on_open_avoid:
+  assumes "continuous_on s f"  "open s"  "x \<in> s"  "f x \<noteq> a"
+  shows "\<exists>e>0. \<forall>y. dist x y < e \<longrightarrow> f y \<noteq> a"
+using assms(1)[unfolded continuous_on_eq_continuous_at[OF assms(2)], THEN bspec[where x=x], OF assms(3)]  continuous_at_avoid[of x f a]  assms(3,4) by auto
+
+text{* Proving a function is constant by proving open-ness of level set.         *}
+
+lemma continuous_levelset_open_in_cases:
+ "connected s \<Longrightarrow> continuous_on s f \<Longrightarrow>
+        openin (subtopology euclidean s) {x \<in> s. f x = a}
+        ==> (\<forall>x \<in> s. f x \<noteq> a) \<or> (\<forall>x \<in> s. f x = a)"
+unfolding connected_clopen using continuous_closed_in_preimage_constant by auto
+
+lemma continuous_levelset_open_in:
+ "connected s \<Longrightarrow> continuous_on s f \<Longrightarrow>
+        openin (subtopology euclidean s) {x \<in> s. f x = a} \<Longrightarrow>
+        (\<exists>x \<in> s. f x = a)  ==> (\<forall>x \<in> s. f x = a)"
+using continuous_levelset_open_in_cases[of s f ]
+by meson
+
+lemma continuous_levelset_open:
+  assumes "connected s"  "continuous_on s f"  "open {x \<in> s. f x = a}"  "\<exists>x \<in> s.  f x = a"
+  shows "\<forall>x \<in> s. f x = a"
+using continuous_levelset_open_in[OF assms(1,2), of a, unfolded openin_open] using assms (3,4) by auto
+
+text{* Some arithmetical combinations (more to prove).                           *}
+
+lemma open_scaling[intro]:
+  fixes s :: "'a::real_normed_vector set"
+  assumes "c \<noteq> 0"  "open s"
+  shows "open((\<lambda>x. c *\<^sub>R x) ` s)"
+proof-
+  { fix x assume "x \<in> s"
+    then obtain e where "e>0" and e:"\<forall>x'. dist x' x < e \<longrightarrow> x' \<in> s" using assms(2)[unfolded open_dist, THEN bspec[where x=x]] by auto
+    have "e * abs c > 0" using assms(1)[unfolded zero_less_abs_iff[THEN sym]] using real_mult_order[OF `e>0`] by auto
+    moreover
+    { fix y assume "dist y (c *\<^sub>R x) < e * \<bar>c\<bar>"
+      hence "norm ((1 / c) *\<^sub>R y - x) < e" unfolding dist_norm
+	using norm_scaleR[of c "(1 / c) *\<^sub>R y - x", unfolded scaleR_right_diff_distrib, unfolded scaleR_scaleR] assms(1)
+	  assms(1)[unfolded zero_less_abs_iff[THEN sym]] by (simp del:zero_less_abs_iff)
+      hence "y \<in> op *\<^sub>R c ` s" using rev_image_eqI[of "(1 / c) *\<^sub>R y" s y "op *\<^sub>R c"]  e[THEN spec[where x="(1 / c) *\<^sub>R y"]]  assms(1) unfolding dist_norm scaleR_scaleR by auto  }
+    ultimately have "\<exists>e>0. \<forall>x'. dist x' (c *\<^sub>R x) < e \<longrightarrow> x' \<in> op *\<^sub>R c ` s" apply(rule_tac x="e * abs c" in exI) by auto  }
+  thus ?thesis unfolding open_dist by auto
+qed
+
+lemma minus_image_eq_vimage:
+  fixes A :: "'a::ab_group_add set"
+  shows "(\<lambda>x. - x) ` A = (\<lambda>x. - x) -` A"
+  by (auto intro!: image_eqI [where f="\<lambda>x. - x"])
+
+lemma open_negations:
+  fixes s :: "'a::real_normed_vector set"
+  shows "open s ==> open ((\<lambda> x. -x) ` s)"
+  unfolding scaleR_minus1_left [symmetric]
+  by (rule open_scaling, auto)
+
+lemma open_translation:
+  fixes s :: "'a::real_normed_vector set"
+  assumes "open s"  shows "open((\<lambda>x. a + x) ` s)"
+proof-
+  { fix x have "continuous (at x) (\<lambda>x. x - a)" using continuous_sub[of "at x" "\<lambda>x. x" "\<lambda>x. a"] continuous_at_id[of x] continuous_const[of "at x" a] by auto  }
+  moreover have "{x. x - a \<in> s}  = op + a ` s" apply auto unfolding image_iff apply(rule_tac x="x - a" in bexI) by auto
+  ultimately show ?thesis using continuous_open_preimage_univ[of "\<lambda>x. x - a" s] using assms by auto
+qed
+
+lemma open_affinity:
+  fixes s :: "'a::real_normed_vector set"
+  assumes "open s"  "c \<noteq> 0"
+  shows "open ((\<lambda>x. a + c *\<^sub>R x) ` s)"
+proof-
+  have *:"(\<lambda>x. a + c *\<^sub>R x) = (\<lambda>x. a + x) \<circ> (\<lambda>x. c *\<^sub>R x)" unfolding o_def ..
+  have "op + a ` op *\<^sub>R c ` s = (op + a \<circ> op *\<^sub>R c) ` s" by auto
+  thus ?thesis using assms open_translation[of "op *\<^sub>R c ` s" a] unfolding * by auto
+qed
+
+lemma interior_translation:
+  fixes s :: "'a::real_normed_vector set"
+  shows "interior ((\<lambda>x. a + x) ` s) = (\<lambda>x. a + x) ` (interior s)"
+proof (rule set_ext, rule)
+  fix x assume "x \<in> interior (op + a ` s)"
+  then obtain e where "e>0" and e:"ball x e \<subseteq> op + a ` s" unfolding mem_interior by auto
+  hence "ball (x - a) e \<subseteq> s" unfolding subset_eq Ball_def mem_ball dist_norm apply auto apply(erule_tac x="a + xa" in allE) unfolding ab_group_add_class.diff_diff_eq[THEN sym] by auto
+  thus "x \<in> op + a ` interior s" unfolding image_iff apply(rule_tac x="x - a" in bexI) unfolding mem_interior using `e > 0` by auto
+next
+  fix x assume "x \<in> op + a ` interior s"
+  then obtain y e where "e>0" and e:"ball y e \<subseteq> s" and y:"x = a + y" unfolding image_iff Bex_def mem_interior by auto
+  { fix z have *:"a + y - z = y + a - z" by auto
+    assume "z\<in>ball x e"
+    hence "z - a \<in> s" using e[unfolded subset_eq, THEN bspec[where x="z - a"]] unfolding mem_ball dist_norm y ab_group_add_class.diff_diff_eq2 * by auto
+    hence "z \<in> op + a ` s" unfolding image_iff by(auto intro!: bexI[where x="z - a"])  }
+  hence "ball x e \<subseteq> op + a ` s" unfolding subset_eq by auto
+  thus "x \<in> interior (op + a ` s)" unfolding mem_interior using `e>0` by auto
+qed
+
+subsection {* Preservation of compactness and connectedness under continuous function.  *}
+
+lemma compact_continuous_image:
+  assumes "continuous_on s f"  "compact s"
+  shows "compact(f ` s)"
+proof-
+  { fix x assume x:"\<forall>n::nat. x n \<in> f ` s"
+    then obtain y where y:"\<forall>n. y n \<in> s \<and> x n = f (y n)" unfolding image_iff Bex_def using choice[of "\<lambda>n xa. xa \<in> s \<and> x n = f xa"] by auto
+    then obtain l r where "l\<in>s" and r:"subseq r" and lr:"((y \<circ> r) ---> l) sequentially" using assms(2)[unfolded compact_def, THEN spec[where x=y]] by auto
+    { fix e::real assume "e>0"
+      then obtain d where "d>0" and d:"\<forall>x'\<in>s. dist x' l < d \<longrightarrow> dist (f x') (f l) < e" using assms(1)[unfolded continuous_on_def, THEN bspec[where x=l], OF `l\<in>s`] by auto
+      then obtain N::nat where N:"\<forall>n\<ge>N. dist ((y \<circ> r) n) l < d" using lr[unfolded Lim_sequentially, THEN spec[where x=d]] by auto
+      { fix n::nat assume "n\<ge>N" hence "dist ((x \<circ> r) n) (f l) < e" using N[THEN spec[where x=n]] d[THEN bspec[where x="y (r n)"]] y[THEN spec[where x="r n"]] by auto  }
+      hence "\<exists>N. \<forall>n\<ge>N. dist ((x \<circ> r) n) (f l) < e" by auto  }
+    hence "\<exists>l\<in>f ` s. \<exists>r. subseq r \<and> ((x \<circ> r) ---> l) sequentially" unfolding Lim_sequentially using r lr `l\<in>s` by auto  }
+  thus ?thesis unfolding compact_def by auto
+qed
+
+lemma connected_continuous_image:
+  assumes "continuous_on s f"  "connected s"
+  shows "connected(f ` s)"
+proof-
+  { fix T assume as: "T \<noteq> {}"  "T \<noteq> f ` s"  "openin (subtopology euclidean (f ` s)) T"  "closedin (subtopology euclidean (f ` s)) T"
+    have "{x \<in> s. f x \<in> T} = {} \<or> {x \<in> s. f x \<in> T} = s"
+      using assms(1)[unfolded continuous_on_open, THEN spec[where x=T]]
+      using assms(1)[unfolded continuous_on_closed, THEN spec[where x=T]]
+      using assms(2)[unfolded connected_clopen, THEN spec[where x="{x \<in> s. f x \<in> T}"]] as(3,4) by auto
+    hence False using as(1,2)
+      using as(4)[unfolded closedin_def topspace_euclidean_subtopology] by auto }
+  thus ?thesis unfolding connected_clopen by auto
+qed
+
+text{* Continuity implies uniform continuity on a compact domain.                *}
+
+lemma compact_uniformly_continuous:
+  assumes "continuous_on s f"  "compact s"
+  shows "uniformly_continuous_on s f"
+proof-
+    { fix x assume x:"x\<in>s"
+      hence "\<forall>xa. \<exists>y. 0 < xa \<longrightarrow> (y > 0 \<and> (\<forall>x'\<in>s. dist x' x < y \<longrightarrow> dist (f x') (f x) < xa))" using assms(1)[unfolded continuous_on_def, THEN bspec[where x=x]] by auto
+      hence "\<exists>fa. \<forall>xa>0. \<forall>x'\<in>s. fa xa > 0 \<and> (dist x' x < fa xa \<longrightarrow> dist (f x') (f x) < xa)" using choice[of "\<lambda>e d. e>0 \<longrightarrow> d>0 \<and>(\<forall>x'\<in>s. (dist x' x < d \<longrightarrow> dist (f x') (f x) < e))"] by auto  }
+    then have "\<forall>x\<in>s. \<exists>y. \<forall>xa. 0 < xa \<longrightarrow> (\<forall>x'\<in>s. y xa > 0 \<and> (dist x' x < y xa \<longrightarrow> dist (f x') (f x) < xa))" by auto
+    then obtain d where d:"\<forall>e>0. \<forall>x\<in>s. \<forall>x'\<in>s. d x e > 0 \<and> (dist x' x < d x e \<longrightarrow> dist (f x') (f x) < e)"
+      using bchoice[of s "\<lambda>x fa. \<forall>xa>0. \<forall>x'\<in>s. fa xa > 0 \<and> (dist x' x < fa xa \<longrightarrow> dist (f x') (f x) < xa)"] by blast
+
+  { fix e::real assume "e>0"
+
+    { fix x assume "x\<in>s" hence "x \<in> ball x (d x (e / 2))" unfolding centre_in_ball using d[THEN spec[where x="e/2"]] using `e>0` by auto  }
+    hence "s \<subseteq> \<Union>{ball x (d x (e / 2)) |x. x \<in> s}" unfolding subset_eq by auto
+    moreover
+    { fix b assume "b\<in>{ball x (d x (e / 2)) |x. x \<in> s}" hence "open b" by auto  }
+    ultimately obtain ea where "ea>0" and ea:"\<forall>x\<in>s. \<exists>b\<in>{ball x (d x (e / 2)) |x. x \<in> s}. ball x ea \<subseteq> b" using heine_borel_lemma[OF assms(2), of "{ball x (d x (e / 2)) | x. x\<in>s }"] by auto
+
+    { fix x y assume "x\<in>s" "y\<in>s" and as:"dist y x < ea"
+      obtain z where "z\<in>s" and z:"ball x ea \<subseteq> ball z (d z (e / 2))" using ea[THEN bspec[where x=x]] and `x\<in>s` by auto
+      hence "x\<in>ball z (d z (e / 2))" using `ea>0` unfolding subset_eq by auto
+      hence "dist (f z) (f x) < e / 2" using d[THEN spec[where x="e/2"]] and `e>0` and `x\<in>s` and `z\<in>s`
+	by (auto  simp add: dist_commute)
+      moreover have "y\<in>ball z (d z (e / 2))" using as and `ea>0` and z[unfolded subset_eq]
+	by (auto simp add: dist_commute)
+      hence "dist (f z) (f y) < e / 2" using d[THEN spec[where x="e/2"]] and `e>0` and `y\<in>s` and `z\<in>s`
+	by (auto  simp add: dist_commute)
+      ultimately have "dist (f y) (f x) < e" using dist_triangle_half_r[of "f z" "f x" e "f y"]
+	by (auto simp add: dist_commute)  }
+    then have "\<exists>d>0. \<forall>x\<in>s. \<forall>x'\<in>s. dist x' x < d \<longrightarrow> dist (f x') (f x) < e" using `ea>0` by auto  }
+  thus ?thesis unfolding uniformly_continuous_on_def by auto
+qed
+
+text{* Continuity of inverse function on compact domain. *}
+
+lemma continuous_on_inverse:
+  fixes f :: "'a::heine_borel \<Rightarrow> 'b::heine_borel"
+    (* TODO: can this be generalized more? *)
+  assumes "continuous_on s f"  "compact s"  "\<forall>x \<in> s. g (f x) = x"
+  shows "continuous_on (f ` s) g"
+proof-
+  have *:"g ` f ` s = s" using assms(3) by (auto simp add: image_iff)
+  { fix t assume t:"closedin (subtopology euclidean (g ` f ` s)) t"
+    then obtain T where T: "closed T" "t = s \<inter> T" unfolding closedin_closed unfolding * by auto
+    have "continuous_on (s \<inter> T) f" using continuous_on_subset[OF assms(1), of "s \<inter> t"]
+      unfolding T(2) and Int_left_absorb by auto
+    moreover have "compact (s \<inter> T)"
+      using assms(2) unfolding compact_eq_bounded_closed
+      using bounded_subset[of s "s \<inter> T"] and T(1) by auto
+    ultimately have "closed (f ` t)" using T(1) unfolding T(2)
+      using compact_continuous_image [of "s \<inter> T" f] unfolding compact_eq_bounded_closed by auto
+    moreover have "{x \<in> f ` s. g x \<in> t} = f ` s \<inter> f ` t" using assms(3) unfolding T(2) by auto
+    ultimately have "closedin (subtopology euclidean (f ` s)) {x \<in> f ` s. g x \<in> t}"
+      unfolding closedin_closed by auto  }
+  thus ?thesis unfolding continuous_on_closed by auto
+qed
+
+subsection{* A uniformly convergent limit of continuous functions is continuous.       *}
+
+lemma norm_triangle_lt:
+  fixes x y :: "'a::real_normed_vector"
+  shows "norm x + norm y < e \<Longrightarrow> norm (x + y) < e"
+by (rule le_less_trans [OF norm_triangle_ineq])
+
+lemma continuous_uniform_limit:
+  fixes f :: "'a \<Rightarrow> 'b::metric_space \<Rightarrow> 'c::real_normed_vector"
+  assumes "\<not> (trivial_limit net)"  "eventually (\<lambda>n. continuous_on s (f n)) net"
+  "\<forall>e>0. eventually (\<lambda>n. \<forall>x \<in> s. norm(f n x - g x) < e) net"
+  shows "continuous_on s g"
+proof-
+  { fix x and e::real assume "x\<in>s" "e>0"
+    have "eventually (\<lambda>n. \<forall>x\<in>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:"\<forall>xa\<in>s. norm (f n xa - g xa) < e / 3"  "continuous_on s (f n)"
+      using eventually_and[of "(\<lambda>n. \<forall>x\<in>s. norm (f n x - g x) < e / 3)" "(\<lambda>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:"\<forall>x'\<in>s. dist x' x < d \<longrightarrow> dist (f n x') (f n x) < e / 3"
+      using n(2)[unfolded continuous_on_def, THEN bspec[where x=x], OF `x\<in>s`, THEN spec[where x="e/3"]] by blast
+    { fix y assume "y\<in>s" "dist y x < d"
+      hence "dist (f n y) (f n x) < e / 3" using d[THEN bspec[where x=y]] by auto
+      hence "norm (f n y - g x) < 2 * e / 3" using norm_triangle_lt[of "f n y - f n x" "f n x - g x" "2*e/3"]
+	using n(1)[THEN bspec[where x=x], OF `x\<in>s`] unfolding dist_norm unfolding ab_group_add_class.ab_diff_minus by auto
+      hence "dist (g y) (g x) < e" unfolding dist_norm using n(1)[THEN bspec[where x=y], OF `y\<in>s`]
+	unfolding norm_minus_cancel[of "f n y - g y", THEN sym] using norm_triangle_lt[of "f n y - g x" "g y - f n y" e] by (auto simp add: uminus_add_conv_diff)  }
+    hence "\<exists>d>0. \<forall>x'\<in>s. dist x' x < d \<longrightarrow> dist (g x') (g x) < e" using `d>0` by auto  }
+  thus ?thesis unfolding continuous_on_def by auto
+qed
+
+subsection{* Topological properties of linear functions.                               *}
+
+lemma linear_lim_0:
+  assumes "bounded_linear f" shows "(f ---> 0) (at (0))"
+proof-
+  interpret f: bounded_linear f by fact
+  have "(f ---> f 0) (at 0)"
+    using tendsto_ident_at by (rule f.tendsto)
+  thus ?thesis unfolding f.zero .
+qed
+
+lemma linear_continuous_at:
+  assumes "bounded_linear f"  shows "continuous (at a) f"
+  unfolding continuous_at using assms
+  apply (rule bounded_linear.tendsto)
+  apply (rule tendsto_ident_at)
+  done
+
+lemma linear_continuous_within:
+  shows "bounded_linear f ==> continuous (at x within s) f"
+  using continuous_at_imp_continuous_within[of x f s] using linear_continuous_at[of f] by auto
+
+lemma linear_continuous_on:
+  shows "bounded_linear f ==> continuous_on s f"
+  using continuous_at_imp_continuous_on[of s f] using linear_continuous_at[of f] by auto
+
+text{* Also bilinear functions, in composition form.                             *}
+
+lemma bilinear_continuous_at_compose:
+  shows "continuous (at x) f \<Longrightarrow> continuous (at x) g \<Longrightarrow> bounded_bilinear h
+        ==> continuous (at x) (\<lambda>x. h (f x) (g x))"
+  unfolding continuous_at using Lim_bilinear[of f "f x" "(at x)" g "g x" h] by auto
+
+lemma bilinear_continuous_within_compose:
+  shows "continuous (at x within s) f \<Longrightarrow> continuous (at x within s) g \<Longrightarrow> bounded_bilinear h
+        ==> continuous (at x within s) (\<lambda>x. h (f x) (g x))"
+  unfolding continuous_within using Lim_bilinear[of f "f x"] by auto
+
+lemma bilinear_continuous_on_compose:
+  shows "continuous_on s f \<Longrightarrow> continuous_on s g \<Longrightarrow> bounded_bilinear h
+             ==> continuous_on s (\<lambda>x. h (f x) (g x))"
+  unfolding continuous_on_eq_continuous_within apply auto apply(erule_tac x=x in ballE) apply auto apply(erule_tac x=x in ballE) apply auto
+  using bilinear_continuous_within_compose[of _ s f g h] by auto
+
+subsection{* Topological stuff lifted from and dropped to R                            *}
+
+
+lemma open_real:
+  fixes s :: "real set" shows
+ "open s \<longleftrightarrow>
+        (\<forall>x \<in> s. \<exists>e>0. \<forall>x'. abs(x' - x) < e --> x' \<in> s)" (is "?lhs = ?rhs")
+  unfolding open_dist dist_norm by simp
+
+lemma islimpt_approachable_real:
+  fixes s :: "real set"
+  shows "x islimpt s \<longleftrightarrow> (\<forall>e>0.  \<exists>x'\<in> s. x' \<noteq> x \<and> abs(x' - x) < e)"
+  unfolding islimpt_approachable dist_norm by simp
+
+lemma closed_real:
+  fixes s :: "real set"
+  shows "closed s \<longleftrightarrow>
+        (\<forall>x. (\<forall>e>0.  \<exists>x' \<in> s. x' \<noteq> x \<and> abs(x' - x) < e)
+            --> x \<in> s)"
+  unfolding closed_limpt islimpt_approachable dist_norm by simp
+
+lemma continuous_at_real_range:
+  fixes f :: "'a::real_normed_vector \<Rightarrow> real"
+  shows "continuous (at x) f \<longleftrightarrow> (\<forall>e>0. \<exists>d>0.
+        \<forall>x'. norm(x' - x) < d --> abs(f x' - f x) < e)"
+  unfolding continuous_at unfolding Lim_at
+  unfolding dist_nz[THEN sym] unfolding dist_norm apply auto
+  apply(erule_tac x=e in allE) apply auto apply (rule_tac x=d in exI) apply auto apply (erule_tac x=x' in allE) apply auto
+  apply(erule_tac x=e in allE) by auto
+
+lemma continuous_on_real_range:
+  fixes f :: "'a::real_normed_vector \<Rightarrow> real"
+  shows "continuous_on s f \<longleftrightarrow> (\<forall>x \<in> s. \<forall>e>0. \<exists>d>0. (\<forall>x' \<in> s. norm(x' - x) < d --> abs(f x' - f x) < e))"
+  unfolding continuous_on_def dist_norm by simp
+
+lemma continuous_at_norm: "continuous (at x) norm"
+  unfolding continuous_at by (intro tendsto_intros)
+
+lemma continuous_on_norm: "continuous_on s norm"
+unfolding continuous_on by (intro ballI tendsto_intros)
+
+lemma continuous_at_component: "continuous (at a) (\<lambda>x. x $ i)"
+unfolding continuous_at by (intro tendsto_intros)
+
+lemma continuous_on_component: "continuous_on s (\<lambda>x. x $ i)"
+unfolding continuous_on by (intro ballI tendsto_intros)
+
+lemma continuous_at_infnorm: "continuous (at x) infnorm"
+  unfolding continuous_at Lim_at o_def unfolding dist_norm
+  apply auto apply (rule_tac x=e in exI) apply auto
+  using order_trans[OF real_abs_sub_infnorm infnorm_le_norm, of _ x] by (metis xt1(7))
+
+text{* Hence some handy theorems on distance, diameter etc. of/from a set.       *}
+
+lemma compact_attains_sup:
+  fixes s :: "real set"
+  assumes "compact s"  "s \<noteq> {}"
+  shows "\<exists>x \<in> s. \<forall>y \<in> s. y \<le> x"
+proof-
+  from assms(1) have a:"bounded s" "closed s" unfolding compact_eq_bounded_closed by auto
+  { fix e::real assume as: "\<forall>x\<in>s. x \<le> rsup s" "rsup s \<notin> s"  "0 < e" "\<forall>x'\<in>s. x' = rsup s \<or> \<not> rsup s - x' < e"
+    have "isLub UNIV s (rsup s)" using rsup[OF assms(2)] unfolding setle_def using as(1) by auto
+    moreover have "isUb UNIV s (rsup s - e)" unfolding isUb_def unfolding setle_def using as(4,2) by auto
+    ultimately have False using isLub_le_isUb[of UNIV s "rsup s" "rsup s - e"] using `e>0` by auto  }
+  thus ?thesis using bounded_has_rsup(1)[OF a(1) assms(2)] using a(2)[unfolded closed_real, THEN spec[where x="rsup s"]]
+    apply(rule_tac x="rsup s" in bexI) by auto
+qed
+
+lemma compact_attains_inf:
+  fixes s :: "real set"
+  assumes "compact s" "s \<noteq> {}"  shows "\<exists>x \<in> s. \<forall>y \<in> s. x \<le> y"
+proof-
+  from assms(1) have a:"bounded s" "closed s" unfolding compact_eq_bounded_closed by auto
+  { fix e::real assume as: "\<forall>x\<in>s. x \<ge> rinf s"  "rinf s \<notin> s"  "0 < e"
+      "\<forall>x'\<in>s. x' = rinf s \<or> \<not> abs (x' - rinf s) < e"
+    have "isGlb UNIV s (rinf s)" using rinf[OF assms(2)] unfolding setge_def using as(1) by auto
+    moreover
+    { fix x assume "x \<in> s"
+      hence *:"abs (x - rinf s) = x - rinf s" using as(1)[THEN bspec[where x=x]] by auto
+      have "rinf s + e \<le> x" using as(4)[THEN bspec[where x=x]] using as(2) `x\<in>s` unfolding * by auto }
+    hence "isLb UNIV s (rinf s + e)" unfolding isLb_def and setge_def by auto
+    ultimately have False using isGlb_le_isLb[of UNIV s "rinf s" "rinf s + e"] using `e>0` by auto  }
+  thus ?thesis using bounded_has_rinf(1)[OF a(1) assms(2)] using a(2)[unfolded closed_real, THEN spec[where x="rinf s"]]
+    apply(rule_tac x="rinf s" in bexI) by auto
+qed
+
+lemma continuous_attains_sup:
+  fixes f :: "'a::metric_space \<Rightarrow> real"
+  shows "compact s \<Longrightarrow> s \<noteq> {} \<Longrightarrow> continuous_on s f
+        ==> (\<exists>x \<in> s. \<forall>y \<in> s.  f y \<le> f x)"
+  using compact_attains_sup[of "f ` s"]
+  using compact_continuous_image[of s f] by auto
+
+lemma continuous_attains_inf:
+  fixes f :: "'a::metric_space \<Rightarrow> real"
+  shows "compact s \<Longrightarrow> s \<noteq> {} \<Longrightarrow> continuous_on s f
+        \<Longrightarrow> (\<exists>x \<in> s. \<forall>y \<in> s. f x \<le> f y)"
+  using compact_attains_inf[of "f ` s"]
+  using compact_continuous_image[of s f] by auto
+
+lemma distance_attains_sup:
+  assumes "compact s" "s \<noteq> {}"
+  shows "\<exists>x \<in> s. \<forall>y \<in> s. dist a y \<le> dist a x"
+proof (rule continuous_attains_sup [OF assms])
+  { fix x assume "x\<in>s"
+    have "(dist a ---> dist a x) (at x within s)"
+      by (intro tendsto_dist tendsto_const Lim_at_within Lim_ident_at)
+  }
+  thus "continuous_on s (dist a)"
+    unfolding continuous_on ..
+qed
+
+text{* For *minimal* distance, we only need closure, not compactness.            *}
+
+lemma distance_attains_inf:
+  fixes a :: "'a::heine_borel"
+  assumes "closed s"  "s \<noteq> {}"
+  shows "\<exists>x \<in> s. \<forall>y \<in> s. dist a x \<le> dist a y"
+proof-
+  from assms(2) obtain b where "b\<in>s" by auto
+  let ?B = "cball a (dist b a) \<inter> s"
+  have "b \<in> ?B" using `b\<in>s` by (simp add: dist_commute)
+  hence "?B \<noteq> {}" by auto
+  moreover
+  { fix x assume "x\<in>?B"
+    fix e::real assume "e>0"
+    { fix x' assume "x'\<in>?B" and as:"dist x' x < e"
+      from as have "\<bar>dist a x' - dist a x\<bar> < e"
+        unfolding abs_less_iff minus_diff_eq
+        using dist_triangle2 [of a x' x]
+        using dist_triangle [of a x x']
+        by arith
+    }
+    hence "\<exists>d>0. \<forall>x'\<in>?B. dist x' x < d \<longrightarrow> \<bar>dist a x' - dist a x\<bar> < e"
+      using `e>0` by auto
+  }
+  hence "continuous_on (cball a (dist b a) \<inter> s) (dist a)"
+    unfolding continuous_on Lim_within dist_norm real_norm_def
+    by fast
+  moreover have "compact ?B"
+    using compact_cball[of a "dist b a"]
+    unfolding compact_eq_bounded_closed
+    using bounded_Int and closed_Int and assms(1) by auto
+  ultimately obtain x where "x\<in>cball a (dist b a) \<inter> s" "\<forall>y\<in>cball a (dist b a) \<inter> s. dist a x \<le> dist a y"
+    using continuous_attains_inf[of ?B "dist a"] by fastsimp
+  thus ?thesis by fastsimp
+qed
+
+subsection{* We can now extend limit compositions to consider the scalar multiplier.   *}
+
+lemma Lim_mul:
+  fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"
+  assumes "(c ---> d) net"  "(f ---> l) net"
+  shows "((\<lambda>x. c(x) *\<^sub>R f x) ---> (d *\<^sub>R l)) net"
+  using assms by (rule scaleR.tendsto)
+
+lemma Lim_vmul:
+  fixes c :: "'a \<Rightarrow> real" and v :: "'b::real_normed_vector"
+  shows "(c ---> d) net ==> ((\<lambda>x. c(x) *\<^sub>R v) ---> d *\<^sub>R v) net"
+  by (intro tendsto_intros)
+
+lemma continuous_vmul:
+  fixes c :: "'a::metric_space \<Rightarrow> real" and v :: "'b::real_normed_vector"
+  shows "continuous net c ==> continuous net (\<lambda>x. c(x) *\<^sub>R v)"
+  unfolding continuous_def using Lim_vmul[of c] by auto
+
+lemma continuous_mul:
+  fixes c :: "'a::metric_space \<Rightarrow> real"
+  fixes f :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector"
+  shows "continuous net c \<Longrightarrow> continuous net f
+             ==> continuous net (\<lambda>x. c(x) *\<^sub>R f x) "
+  unfolding continuous_def by (intro tendsto_intros)
+
+lemma continuous_on_vmul:
+  fixes c :: "'a::metric_space \<Rightarrow> real" and v :: "'b::real_normed_vector"
+  shows "continuous_on s c ==> continuous_on s (\<lambda>x. c(x) *\<^sub>R v)"
+  unfolding continuous_on_eq_continuous_within using continuous_vmul[of _ c] by auto
+
+lemma continuous_on_mul:
+  fixes c :: "'a::metric_space \<Rightarrow> real"
+  fixes f :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector"
+  shows "continuous_on s c \<Longrightarrow> continuous_on s f
+             ==> continuous_on s (\<lambda>x. c(x) *\<^sub>R f x)"
+  unfolding continuous_on_eq_continuous_within using continuous_mul[of _ c] by auto
+
+text{* And so we have continuity of inverse.                                     *}
+
+lemma Lim_inv:
+  fixes f :: "'a \<Rightarrow> real"
+  assumes "(f ---> l) (net::'a net)"  "l \<noteq> 0"
+  shows "((inverse o f) ---> inverse l) net"
+  unfolding o_def using assms by (rule tendsto_inverse)
+
+lemma continuous_inv:
+  fixes f :: "'a::metric_space \<Rightarrow> real"
+  shows "continuous net f \<Longrightarrow> f(netlimit net) \<noteq> 0
+           ==> continuous net (inverse o f)"
+  unfolding continuous_def using Lim_inv by auto
+
+lemma continuous_at_within_inv:
+  fixes f :: "'a::metric_space \<Rightarrow> 'b::real_normed_field"
+  assumes "continuous (at a within s) f" "f a \<noteq> 0"
+  shows "continuous (at a within s) (inverse o f)"
+  using assms unfolding continuous_within o_def
+  by (intro tendsto_intros)
+
+lemma continuous_at_inv:
+  fixes f :: "'a::metric_space \<Rightarrow> 'b::real_normed_field"
+  shows "continuous (at a) f \<Longrightarrow> f a \<noteq> 0
+         ==> continuous (at a) (inverse o f) "
+  using within_UNIV[THEN sym, of "at a"] using continuous_at_within_inv[of a UNIV] by auto
+
+subsection{* Preservation properties for pasted sets.                                  *}
+
+lemma bounded_pastecart:
+  fixes s :: "('a::real_normed_vector ^ _) set" (* FIXME: generalize to metric_space *)
+  assumes "bounded s" "bounded t"
+  shows "bounded { pastecart x y | x y . (x \<in> s \<and> y \<in> t)}"
+proof-
+  obtain a b where ab:"\<forall>x\<in>s. norm x \<le> a" "\<forall>x\<in>t. norm x \<le> b" using assms[unfolded bounded_iff] by auto
+  { fix x y assume "x\<in>s" "y\<in>t"
+    hence "norm x \<le> a" "norm y \<le> b" using ab by auto
+    hence "norm (pastecart x y) \<le> a + b" using norm_pastecart[of x y] by auto }
+  thus ?thesis unfolding bounded_iff by auto
+qed
+
+lemma bounded_Times:
+  assumes "bounded s" "bounded t" shows "bounded (s \<times> t)"
+proof-
+  obtain x y a b where "\<forall>z\<in>s. dist x z \<le> a" "\<forall>z\<in>t. dist y z \<le> b"
+    using assms [unfolded bounded_def] by auto
+  then have "\<forall>z\<in>s \<times> t. dist (x, y) z \<le> sqrt (a\<twosuperior> + b\<twosuperior>)"
+    by (auto simp add: dist_Pair_Pair real_sqrt_le_mono add_mono power_mono)
+  thus ?thesis unfolding bounded_any_center [where a="(x, y)"] by auto
+qed
+
+lemma closed_pastecart:
+  fixes s :: "(real ^ 'a::finite) set" (* FIXME: generalize *)
+  assumes "closed s"  "closed t"
+  shows "closed {pastecart x y | x y . x \<in> s \<and> y \<in> t}"
+proof-
+  { fix x l assume as:"\<forall>n::nat. x n \<in> {pastecart x y |x y. x \<in> s \<and> y \<in> t}"  "(x ---> l) sequentially"
+    { fix n::nat have "fstcart (x n) \<in> s" "sndcart (x n) \<in> t" using as(1)[THEN spec[where x=n]] by auto } note * = this
+    moreover
+    { fix e::real assume "e>0"
+      then obtain N::nat where N:"\<forall>n\<ge>N. dist (x n) l < e" using as(2)[unfolded Lim_sequentially, THEN spec[where x=e]] by auto
+      { fix n::nat assume "n\<ge>N"
+	hence "dist (fstcart (x n)) (fstcart l) < e" "dist (sndcart (x n)) (sndcart l) < e"
+	  using N[THEN spec[where x=n]] dist_fstcart[of "x n" l] dist_sndcart[of "x n" l] by auto   }
+      hence "\<exists>N. \<forall>n\<ge>N. dist (fstcart (x n)) (fstcart l) < e" "\<exists>N. \<forall>n\<ge>N. dist (sndcart (x n)) (sndcart l) < e" by auto  }
+    ultimately have "fstcart l \<in> s" "sndcart l \<in> t"
+      using assms(1)[unfolded closed_sequential_limits, THEN spec[where x="\<lambda>n. fstcart (x n)"], THEN spec[where x="fstcart l"]]
+      using assms(2)[unfolded closed_sequential_limits, THEN spec[where x="\<lambda>n. sndcart (x n)"], THEN spec[where x="sndcart l"]]
+      unfolding Lim_sequentially by auto
+    hence "l \<in> {pastecart x y |x y. x \<in> s \<and> y \<in> t}" using pastecart_fst_snd[THEN sym, of l] by auto  }
+  thus ?thesis unfolding closed_sequential_limits by auto
+qed
+
+lemma compact_pastecart:
+  fixes s t :: "(real ^ _) set"
+  shows "compact s \<Longrightarrow> compact t ==> compact {pastecart x y | x y . x \<in> s \<and> y \<in> t}"
+  unfolding compact_eq_bounded_closed using bounded_pastecart[of s t] closed_pastecart[of s t] by auto
+
+lemma mem_Times_iff: "x \<in> A \<times> B \<longleftrightarrow> fst x \<in> A \<and> snd x \<in> B"
+by (induct x) simp
+
+lemma compact_Times: "compact s \<Longrightarrow> compact t \<Longrightarrow> compact (s \<times> t)"
+unfolding compact_def
+apply clarify
+apply (drule_tac x="fst \<circ> f" in spec)
+apply (drule mp, simp add: mem_Times_iff)
+apply (clarify, rename_tac l1 r1)
+apply (drule_tac x="snd \<circ> f \<circ> r1" in spec)
+apply (drule mp, simp add: mem_Times_iff)
+apply (clarify, rename_tac l2 r2)
+apply (rule_tac x="(l1, l2)" in rev_bexI, simp)
+apply (rule_tac x="r1 \<circ> r2" in exI)
+apply (rule conjI, simp add: subseq_def)
+apply (drule_tac r=r2 in lim_subseq [COMP swap_prems_rl], assumption)
+apply (drule (1) tendsto_Pair) back
+apply (simp add: o_def)
+done
+
+text{* Hence some useful properties follow quite easily.                         *}
+
+lemma compact_scaling:
+  fixes s :: "'a::real_normed_vector set"
+  assumes "compact s"  shows "compact ((\<lambda>x. c *\<^sub>R x) ` s)"
+proof-
+  let ?f = "\<lambda>x. scaleR c x"
+  have *:"bounded_linear ?f" by (rule scaleR.bounded_linear_right)
+  show ?thesis using compact_continuous_image[of s ?f] continuous_at_imp_continuous_on[of s ?f]
+    using linear_continuous_at[OF *] assms by auto
+qed
+
+lemma compact_negations:
+  fixes s :: "'a::real_normed_vector set"
+  assumes "compact s"  shows "compact ((\<lambda>x. -x) ` s)"
+  using compact_scaling [OF assms, of "- 1"] by auto
+
+lemma compact_sums:
+  fixes s t :: "'a::real_normed_vector set"
+  assumes "compact s"  "compact t"  shows "compact {x + y | x y. x \<in> s \<and> y \<in> t}"
+proof-
+  have *:"{x + y | x y. x \<in> s \<and> y \<in> t} = (\<lambda>z. fst z + snd z) ` (s \<times> t)"
+    apply auto unfolding image_iff apply(rule_tac x="(xa, y)" in bexI) by auto
+  have "continuous_on (s \<times> t) (\<lambda>z. fst z + snd z)"
+    unfolding continuous_on by (rule ballI) (intro tendsto_intros)
+  thus ?thesis unfolding * using compact_continuous_image compact_Times [OF assms] by auto
+qed
+
+lemma compact_differences:
+  fixes s t :: "'a::real_normed_vector set"
+  assumes "compact s" "compact t"  shows "compact {x - y | x y. x \<in> s \<and> y \<in> t}"
+proof-
+  have "{x - y | x y. x\<in>s \<and> y \<in> t} =  {x + y | x y. x \<in> s \<and> y \<in> (uminus ` t)}"
+    apply auto apply(rule_tac x= xa in exI) apply auto apply(rule_tac x=xa in exI) by auto
+  thus ?thesis using compact_sums[OF assms(1) compact_negations[OF assms(2)]] by auto
+qed
+
+lemma compact_translation:
+  fixes s :: "'a::real_normed_vector set"
+  assumes "compact s"  shows "compact ((\<lambda>x. a + x) ` s)"
+proof-
+  have "{x + y |x y. x \<in> s \<and> y \<in> {a}} = (\<lambda>x. a + x) ` s" by auto
+  thus ?thesis using compact_sums[OF assms compact_sing[of a]] by auto
+qed
+
+lemma compact_affinity:
+  fixes s :: "'a::real_normed_vector set"
+  assumes "compact s"  shows "compact ((\<lambda>x. a + c *\<^sub>R x) ` s)"
+proof-
+  have "op + a ` op *\<^sub>R c ` s = (\<lambda>x. a + c *\<^sub>R x) ` s" by auto
+  thus ?thesis using compact_translation[OF compact_scaling[OF assms], of a c] by auto
+qed
+
+text{* Hence we get the following.                                               *}
+
+lemma compact_sup_maxdistance:
+  fixes s :: "'a::real_normed_vector set"
+  assumes "compact s"  "s \<noteq> {}"
+  shows "\<exists>x\<in>s. \<exists>y\<in>s. \<forall>u\<in>s. \<forall>v\<in>s. norm(u - v) \<le> norm(x - y)"
+proof-
+  have "{x - y | x y . x\<in>s \<and> y\<in>s} \<noteq> {}" using `s \<noteq> {}` by auto
+  then obtain x where x:"x\<in>{x - y |x y. x \<in> s \<and> y \<in> s}"  "\<forall>y\<in>{x - y |x y. x \<in> s \<and> y \<in> s}. norm y \<le> norm x"
+    using compact_differences[OF assms(1) assms(1)]
+    using distance_attains_sup[where 'a="'a", unfolded dist_norm, of "{x - y | x y . x\<in>s \<and> y\<in>s}" 0] by(auto simp add: norm_minus_cancel)
+  from x(1) obtain a b where "a\<in>s" "b\<in>s" "x = a - b" by auto
+  thus ?thesis using x(2)[unfolded `x = a - b`] by blast
+qed
+
+text{* We can state this in terms of diameter of a set.                          *}
+
+definition "diameter s = (if s = {} then 0::real else rsup {norm(x - y) | x y. x \<in> s \<and> y \<in> s})"
+  (* TODO: generalize to class metric_space *)
+
+lemma diameter_bounded:
+  assumes "bounded s"
+  shows "\<forall>x\<in>s. \<forall>y\<in>s. norm(x - y) \<le> diameter s"
+        "\<forall>d>0. d < diameter s --> (\<exists>x\<in>s. \<exists>y\<in>s. norm(x - y) > d)"
+proof-
+  let ?D = "{norm (x - y) |x y. x \<in> s \<and> y \<in> s}"
+  obtain a where a:"\<forall>x\<in>s. norm x \<le> a" using assms[unfolded bounded_iff] by auto
+  { fix x y assume "x \<in> s" "y \<in> s"
+    hence "norm (x - y) \<le> 2 * a" using norm_triangle_ineq[of x "-y", unfolded norm_minus_cancel] a[THEN bspec[where x=x]] a[THEN bspec[where x=y]] by (auto simp add: ring_simps)  }
+  note * = this
+  { fix x y assume "x\<in>s" "y\<in>s"  hence "s \<noteq> {}" by auto
+    have lub:"isLub UNIV ?D (rsup ?D)" using * rsup[of ?D] using `s\<noteq>{}` unfolding setle_def by auto
+    have "norm(x - y) \<le> diameter s" unfolding diameter_def using `s\<noteq>{}` *[OF `x\<in>s` `y\<in>s`] `x\<in>s` `y\<in>s` isLubD1[OF lub] unfolding setle_def by auto  }
+  moreover
+  { fix d::real assume "d>0" "d < diameter s"
+    hence "s\<noteq>{}" unfolding diameter_def by auto
+    hence lub:"isLub UNIV ?D (rsup ?D)" using * rsup[of ?D] unfolding setle_def by auto
+    have "\<exists>d' \<in> ?D. d' > d"
+    proof(rule ccontr)
+      assume "\<not> (\<exists>d'\<in>{norm (x - y) |x y. x \<in> s \<and> y \<in> s}. d < d')"
+      hence as:"\<forall>d'\<in>?D. d' \<le> d" apply auto apply(erule_tac x="norm (x - y)" in allE) by auto
+      hence "isUb UNIV ?D d" unfolding isUb_def unfolding setle_def by auto
+      thus False using `d < diameter s` `s\<noteq>{}` isLub_le_isUb[OF lub, of d] unfolding diameter_def  by auto
+    qed
+    hence "\<exists>x\<in>s. \<exists>y\<in>s. norm(x - y) > d" by auto  }
+  ultimately show "\<forall>x\<in>s. \<forall>y\<in>s. norm(x - y) \<le> diameter s"
+        "\<forall>d>0. d < diameter s --> (\<exists>x\<in>s. \<exists>y\<in>s. norm(x - y) > d)" by auto
+qed
+
+lemma diameter_bounded_bound:
+ "bounded s \<Longrightarrow> x \<in> s \<Longrightarrow> y \<in> s ==> norm(x - y) \<le> diameter s"
+  using diameter_bounded by blast
+
+lemma diameter_compact_attained:
+  fixes s :: "'a::real_normed_vector set"
+  assumes "compact s"  "s \<noteq> {}"
+  shows "\<exists>x\<in>s. \<exists>y\<in>s. (norm(x - y) = diameter s)"
+proof-
+  have b:"bounded s" using assms(1) by (rule compact_imp_bounded)
+  then obtain x y where xys:"x\<in>s" "y\<in>s" and xy:"\<forall>u\<in>s. \<forall>v\<in>s. norm (u - v) \<le> norm (x - y)" using compact_sup_maxdistance[OF assms] by auto
+  hence "diameter s \<le> norm (x - y)" using rsup_le[of "{norm (x - y) |x y. x \<in> s \<and> y \<in> s}" "norm (x - y)"]
+    unfolding setle_def and diameter_def by auto
+  thus ?thesis using diameter_bounded(1)[OF b, THEN bspec[where x=x], THEN bspec[where x=y], OF xys] and xys by auto
+qed
+
+text{* Related results with closure as the conclusion.                           *}
+
+lemma closed_scaling:
+  fixes s :: "'a::real_normed_vector set"
+  assumes "closed s" shows "closed ((\<lambda>x. c *\<^sub>R x) ` s)"
+proof(cases "s={}")
+  case True thus ?thesis by auto
+next
+  case False
+  show ?thesis
+  proof(cases "c=0")
+    have *:"(\<lambda>x. 0) ` s = {0}" using `s\<noteq>{}` by auto
+    case True thus ?thesis apply auto unfolding * using closed_sing by auto
+  next
+    case False
+    { fix x l assume as:"\<forall>n::nat. x n \<in> scaleR c ` s"  "(x ---> l) sequentially"
+      { fix n::nat have "scaleR (1 / c) (x n) \<in> s"
+          using as(1)[THEN spec[where x=n]]
+          using `c\<noteq>0` by (auto simp add: vector_smult_assoc)
+      }
+      moreover
+      { fix e::real assume "e>0"
+	hence "0 < e *\<bar>c\<bar>"  using `c\<noteq>0` mult_pos_pos[of e "abs c"] by auto
+	then obtain N where "\<forall>n\<ge>N. dist (x n) l < e * \<bar>c\<bar>"
+          using as(2)[unfolded Lim_sequentially, THEN spec[where x="e * abs c"]] by auto
+	hence "\<exists>N. \<forall>n\<ge>N. dist (scaleR (1 / c) (x n)) (scaleR (1 / c) l) < e"
+          unfolding dist_norm unfolding scaleR_right_diff_distrib[THEN sym]
+	  using mult_imp_div_pos_less[of "abs c" _ e] `c\<noteq>0` by auto  }
+      hence "((\<lambda>n. scaleR (1 / c) (x n)) ---> scaleR (1 / c) l) sequentially" unfolding Lim_sequentially by auto
+      ultimately have "l \<in> scaleR c ` s"
+        using assms[unfolded closed_sequential_limits, THEN spec[where x="\<lambda>n. scaleR (1/c) (x n)"], THEN spec[where x="scaleR (1/c) l"]]
+	unfolding image_iff using `c\<noteq>0` apply(rule_tac x="scaleR (1 / c) l" in bexI) by auto  }
+    thus ?thesis unfolding closed_sequential_limits by fast
+  qed
+qed
+
+lemma closed_negations:
+  fixes s :: "'a::real_normed_vector set"
+  assumes "closed s"  shows "closed ((\<lambda>x. -x) ` s)"
+  using closed_scaling[OF assms, of "- 1"] by simp
+
+lemma compact_closed_sums:
+  fixes s :: "'a::real_normed_vector set"
+  assumes "compact s"  "closed t"  shows "closed {x + y | x y. x \<in> s \<and> y \<in> t}"
+proof-
+  let ?S = "{x + y |x y. x \<in> s \<and> y \<in> t}"
+  { fix x l assume as:"\<forall>n. x n \<in> ?S"  "(x ---> l) sequentially"
+    from as(1) obtain f where f:"\<forall>n. x n = fst (f n) + snd (f n)"  "\<forall>n. fst (f n) \<in> s"  "\<forall>n. snd (f n) \<in> t"
+      using choice[of "\<lambda>n y. x n = (fst y) + (snd y) \<and> fst y \<in> s \<and> snd y \<in> t"] by auto
+    obtain l' r where "l'\<in>s" and r:"subseq r" and lr:"(((\<lambda>n. fst (f n)) \<circ> r) ---> l') sequentially"
+      using assms(1)[unfolded compact_def, THEN spec[where x="\<lambda> n. fst (f n)"]] using f(2) by auto
+    have "((\<lambda>n. snd (f (r n))) ---> l - l') sequentially"
+      using Lim_sub[OF lim_subseq[OF r as(2)] lr] and f(1) unfolding o_def by auto
+    hence "l - l' \<in> t"
+      using assms(2)[unfolded closed_sequential_limits, THEN spec[where x="\<lambda> n. snd (f (r n))"], THEN spec[where x="l - l'"]]
+      using f(3) by auto
+    hence "l \<in> ?S" using `l' \<in> s` apply auto apply(rule_tac x=l' in exI) apply(rule_tac x="l - l'" in exI) by auto
+  }
+  thus ?thesis unfolding closed_sequential_limits by fast
+qed
+
+lemma closed_compact_sums:
+  fixes s t :: "'a::real_normed_vector set"
+  assumes "closed s"  "compact t"
+  shows "closed {x + y | x y. x \<in> s \<and> y \<in> t}"
+proof-
+  have "{x + y |x y. x \<in> t \<and> y \<in> s} = {x + y |x y. x \<in> s \<and> y \<in> t}" apply auto
+    apply(rule_tac x=y in exI) apply auto apply(rule_tac x=y in exI) by auto
+  thus ?thesis using compact_closed_sums[OF assms(2,1)] by simp
+qed
+
+lemma compact_closed_differences:
+  fixes s t :: "'a::real_normed_vector set"
+  assumes "compact s"  "closed t"
+  shows "closed {x - y | x y. x \<in> s \<and> y \<in> t}"
+proof-
+  have "{x + y |x y. x \<in> s \<and> y \<in> uminus ` t} =  {x - y |x y. x \<in> s \<and> y \<in> t}"
+    apply auto apply(rule_tac x=xa in exI) apply auto apply(rule_tac x=xa in exI) by auto
+  thus ?thesis using compact_closed_sums[OF assms(1) closed_negations[OF assms(2)]] by auto
+qed
+
+lemma closed_compact_differences:
+  fixes s t :: "'a::real_normed_vector set"
+  assumes "closed s" "compact t"
+  shows "closed {x - y | x y. x \<in> s \<and> y \<in> t}"
+proof-
+  have "{x + y |x y. x \<in> s \<and> y \<in> uminus ` t} = {x - y |x y. x \<in> s \<and> y \<in> t}"
+    apply auto apply(rule_tac x=xa in exI) apply auto apply(rule_tac x=xa in exI) by auto
+ thus ?thesis using closed_compact_sums[OF assms(1) compact_negations[OF assms(2)]] by simp
+qed
+
+lemma closed_translation:
+  fixes a :: "'a::real_normed_vector"
+  assumes "closed s"  shows "closed ((\<lambda>x. a + x) ` s)"
+proof-
+  have "{a + y |y. y \<in> s} = (op + a ` s)" by auto
+  thus ?thesis using compact_closed_sums[OF compact_sing[of a] assms] by auto
+qed
+
+lemma translation_UNIV:
+  fixes a :: "'a::ab_group_add" shows "range (\<lambda>x. a + x) = UNIV"
+  apply (auto simp add: image_iff) apply(rule_tac x="x - a" in exI) by auto
+
+lemma translation_diff:
+  fixes a :: "'a::ab_group_add"
+  shows "(\<lambda>x. a + x) ` (s - t) = ((\<lambda>x. a + x) ` s) - ((\<lambda>x. a + x) ` t)"
+  by auto
+
+lemma closure_translation:
+  fixes a :: "'a::real_normed_vector"
+  shows "closure ((\<lambda>x. a + x) ` s) = (\<lambda>x. a + x) ` (closure s)"
+proof-
+  have *:"op + a ` (UNIV - s) = UNIV - op + a ` s"
+    apply auto unfolding image_iff apply(rule_tac x="x - a" in bexI) by auto
+  show ?thesis unfolding closure_interior translation_diff translation_UNIV
+    using interior_translation[of a "UNIV - s"] unfolding * by auto
+qed
+
+lemma frontier_translation:
+  fixes a :: "'a::real_normed_vector"
+  shows "frontier((\<lambda>x. a + x) ` s) = (\<lambda>x. a + x) ` (frontier s)"
+  unfolding frontier_def translation_diff interior_translation closure_translation by auto
+
+subsection{* Separation between points and sets.                                       *}
+
+lemma separate_point_closed:
+  fixes s :: "'a::heine_borel set"
+  shows "closed s \<Longrightarrow> a \<notin> s  ==> (\<exists>d>0. \<forall>x\<in>s. d \<le> dist a x)"
+proof(cases "s = {}")
+  case True
+  thus ?thesis by(auto intro!: exI[where x=1])
+next
+  case False
+  assume "closed s" "a \<notin> s"
+  then obtain x where "x\<in>s" "\<forall>y\<in>s. dist a x \<le> dist a y" using `s \<noteq> {}` distance_attains_inf [of s a] by blast
+  with `x\<in>s` show ?thesis using dist_pos_lt[of a x] and`a \<notin> s` by blast
+qed
+
+lemma separate_compact_closed:
+  fixes s t :: "'a::{heine_borel, real_normed_vector} set"
+    (* TODO: does this generalize to heine_borel? *)
+  assumes "compact s" and "closed t" and "s \<inter> t = {}"
+  shows "\<exists>d>0. \<forall>x\<in>s. \<forall>y\<in>t. d \<le> dist x y"
+proof-
+  have "0 \<notin> {x - y |x y. x \<in> s \<and> y \<in> t}" using assms(3) by auto
+  then obtain d where "d>0" and d:"\<forall>x\<in>{x - y |x y. x \<in> s \<and> y \<in> t}. d \<le> dist 0 x"
+    using separate_point_closed[OF compact_closed_differences[OF assms(1,2)], of 0] by auto
+  { fix x y assume "x\<in>s" "y\<in>t"
+    hence "x - y \<in> {x - y |x y. x \<in> s \<and> y \<in> t}" by auto
+    hence "d \<le> dist (x - y) 0" using d[THEN bspec[where x="x - y"]] using dist_commute
+      by (auto  simp add: dist_commute)
+    hence "d \<le> dist x y" unfolding dist_norm by auto  }
+  thus ?thesis using `d>0` by auto
+qed
+
+lemma separate_closed_compact:
+  fixes s t :: "'a::{heine_borel, real_normed_vector} set"
+  assumes "closed s" and "compact t" and "s \<inter> t = {}"
+  shows "\<exists>d>0. \<forall>x\<in>s. \<forall>y\<in>t. d \<le> dist x y"
+proof-
+  have *:"t \<inter> s = {}" using assms(3) by auto
+  show ?thesis using separate_compact_closed[OF assms(2,1) *]
+    apply auto apply(rule_tac x=d in exI) apply auto apply (erule_tac x=y in ballE)
+    by (auto simp add: dist_commute)
+qed
+
+(* A cute way of denoting open and closed intervals using overloading.       *)
+
+lemma interval: fixes a :: "'a::ord^'n::finite" shows
+  "{a <..< b} = {x::'a^'n. \<forall>i. a$i < x$i \<and> x$i < b$i}" and
+  "{a .. b} = {x::'a^'n. \<forall>i. a$i \<le> x$i \<and> x$i \<le> b$i}"
+  by (auto simp add: expand_set_eq vector_less_def vector_less_eq_def)
+
+lemma mem_interval: fixes a :: "'a::ord^'n::finite" shows
+  "x \<in> {a<..<b} \<longleftrightarrow> (\<forall>i. a$i < x$i \<and> x$i < b$i)"
+  "x \<in> {a .. b} \<longleftrightarrow> (\<forall>i. a$i \<le> x$i \<and> x$i \<le> b$i)"
+  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 \<in> {a .. b} \<longleftrightarrow> dest_vec1 a \<le> dest_vec1 x \<and> dest_vec1 x \<le> dest_vec1 b)"
+ "(x \<in> {a<..<b} \<longleftrightarrow> dest_vec1 a < dest_vec1 x \<and> 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} = {} \<longleftrightarrow> (\<exists>i. b$i \<le> a$i))" (is ?th1) and
+ "({a  ..  b} = {} \<longleftrightarrow> (\<exists>i. b$i < a$i))" (is ?th2)
+proof-
+  { fix i x assume as:"b$i \<le> a$i" and x:"x\<in>{a <..< b}"
+    hence "a $ i < x $ i \<and> x $ i < b $ i" unfolding mem_interval by auto
+    hence "a$i < b$i" by auto
+    hence False using as by auto  }
+  moreover
+  { assume as:"\<forall>i. \<not> (b$i \<le> a$i)"
+    let ?x = "(1/2) *\<^sub>R (a + b)"
+    { fix i
+      have "a$i < b$i" using as[THEN spec[where x=i]] by auto
+      hence "a$i < ((1/2) *\<^sub>R (a+b)) $ i" "((1/2) *\<^sub>R (a+b)) $ i < b$i"
+	unfolding vector_smult_component and vector_add_component
+	by (auto simp add: less_divide_eq_number_of1)  }
+    hence "{a <..< b} \<noteq> {}" using mem_interval(1)[of "?x" a b] by auto  }
+  ultimately show ?th1 by blast
+
+  { fix i x assume as:"b$i < a$i" and x:"x\<in>{a .. b}"
+    hence "a $ i \<le> x $ i \<and> x $ i \<le> b $ i" unfolding mem_interval by auto
+    hence "a$i \<le> b$i" by auto
+    hence False using as by auto  }
+  moreover
+  { assume as:"\<forall>i. \<not> (b$i < a$i)"
+    let ?x = "(1/2) *\<^sub>R (a + b)"
+    { fix i
+      have "a$i \<le> b$i" using as[THEN spec[where x=i]] by auto
+      hence "a$i \<le> ((1/2) *\<^sub>R (a+b)) $ i" "((1/2) *\<^sub>R (a+b)) $ i \<le> b$i"
+	unfolding vector_smult_component and vector_add_component
+	by (auto simp add: less_divide_eq_number_of1)  }
+    hence "{a .. b} \<noteq> {}" using mem_interval(2)[of "?x" a b] by auto  }
+  ultimately show ?th2 by blast
+qed
+
+lemma interval_ne_empty: fixes a :: "real^'n::finite" shows
+  "{a  ..  b} \<noteq> {} \<longleftrightarrow> (\<forall>i. a$i \<le> b$i)" and
+  "{a <..< b} \<noteq> {} \<longleftrightarrow> (\<forall>i. a$i < b$i)"
+  unfolding interval_eq_empty[of a b] by (auto simp add: not_less not_le) (* BH: Why doesn't just "auto" work here? *)
+
+lemma subset_interval_imp: fixes a :: "real^'n::finite" shows
+ "(\<forall>i. a$i \<le> c$i \<and> d$i \<le> b$i) \<Longrightarrow> {c .. d} \<subseteq> {a .. b}" and
+ "(\<forall>i. a$i < c$i \<and> d$i < b$i) \<Longrightarrow> {c .. d} \<subseteq> {a<..<b}" and
+ "(\<forall>i. a$i \<le> c$i \<and> d$i \<le> b$i) \<Longrightarrow> {c<..<d} \<subseteq> {a .. b}" and
+ "(\<forall>i. a$i \<le> c$i \<and> d$i \<le> b$i) \<Longrightarrow> {c<..<d} \<subseteq> {a<..<b}"
+  unfolding subset_eq[unfolded Ball_def] unfolding mem_interval
+  by (auto intro: order_trans less_le_trans le_less_trans less_imp_le) (* BH: Why doesn't just "auto" work here? *)
+
+lemma interval_sing: fixes a :: "'a::linorder^'n::finite" shows
+ "{a .. a} = {a} \<and> {a<..<a} = {}"
+apply(auto simp add: expand_set_eq vector_less_def vector_less_eq_def Cart_eq)
+apply (simp add: order_eq_iff)
+apply (auto simp add: not_less less_imp_le)
+done
+
+lemma interval_open_subset_closed:  fixes a :: "'a::preorder^'n::finite" shows
+ "{a<..<b} \<subseteq> {a .. b}"
+proof(simp add: subset_eq, rule)
+  fix x
+  assume x:"x \<in>{a<..<b}"
+  { fix i
+    have "a $ i \<le> x $ i"
+      using x order_less_imp_le[of "a$i" "x$i"]
+      by(simp add: expand_set_eq vector_less_def vector_less_eq_def Cart_eq)
+  }
+  moreover
+  { fix i
+    have "x $ i \<le> b $ i"
+      using x order_less_imp_le[of "x$i" "b$i"]
+      by(simp add: expand_set_eq vector_less_def vector_less_eq_def Cart_eq)
+  }
+  ultimately
+  show "a \<le> x \<and> x \<le> b"
+    by(simp add: expand_set_eq vector_less_def vector_less_eq_def Cart_eq)
+qed
+
+lemma subset_interval: fixes a :: "real^'n::finite" shows
+ "{c .. d} \<subseteq> {a .. b} \<longleftrightarrow> (\<forall>i. c$i \<le> d$i) --> (\<forall>i. a$i \<le> c$i \<and> d$i \<le> b$i)" (is ?th1) and
+ "{c .. d} \<subseteq> {a<..<b} \<longleftrightarrow> (\<forall>i. c$i \<le> d$i) --> (\<forall>i. a$i < c$i \<and> d$i < b$i)" (is ?th2) and
+ "{c<..<d} \<subseteq> {a .. b} \<longleftrightarrow> (\<forall>i. c$i < d$i) --> (\<forall>i. a$i \<le> c$i \<and> d$i \<le> b$i)" (is ?th3) and
+ "{c<..<d} \<subseteq> {a<..<b} \<longleftrightarrow> (\<forall>i. c$i < d$i) --> (\<forall>i. a$i \<le> c$i \<and> d$i \<le> b$i)" (is ?th4)
+proof-
+  show ?th1 unfolding subset_eq and Ball_def and mem_interval by (auto intro: order_trans)
+  show ?th2 unfolding subset_eq and Ball_def and mem_interval by (auto intro: le_less_trans less_le_trans order_trans less_imp_le)
+  { assume as: "{c<..<d} \<subseteq> {a .. b}" "\<forall>i. c$i < d$i"
+    hence "{c<..<d} \<noteq> {}" unfolding interval_eq_empty by (auto, drule_tac x=i in spec, simp) (* BH: Why doesn't just "auto" work? *)
+    fix i
+    (** TODO combine the following two parts as done in the HOL_light version. **)
+    { let ?x = "(\<chi> j. (if j=i then ((min (a$j) (d$j))+c$j)/2 else (c$j+d$j)/2))::real^'n"
+      assume as2: "a$i > c$i"
+      { fix j
+	have "c $ j < ?x $ j \<and> ?x $ j < d $ j" unfolding Cart_lambda_beta
+	  apply(cases "j=i") using as(2)[THEN spec[where x=j]]
+	  by (auto simp add: less_divide_eq_number_of1 as2)  }
+      hence "?x\<in>{c<..<d}" unfolding mem_interval by auto
+      moreover
+      have "?x\<notin>{a .. b}"
+	unfolding mem_interval apply auto apply(rule_tac x=i in exI)
+	using as(2)[THEN spec[where x=i]] and as2
+	by (auto simp add: less_divide_eq_number_of1)
+      ultimately have False using as by auto  }
+    hence "a$i \<le> c$i" by(rule ccontr)auto
+    moreover
+    { let ?x = "(\<chi> j. (if j=i then ((max (b$j) (c$j))+d$j)/2 else (c$j+d$j)/2))::real^'n"
+      assume as2: "b$i < d$i"
+      { fix j
+	have "d $ j > ?x $ j \<and> ?x $ j > c $ j" unfolding Cart_lambda_beta
+	  apply(cases "j=i") using as(2)[THEN spec[where x=j]]
+	  by (auto simp add: less_divide_eq_number_of1 as2)  }
+      hence "?x\<in>{c<..<d}" unfolding mem_interval by auto
+      moreover
+      have "?x\<notin>{a .. b}"
+	unfolding mem_interval apply auto apply(rule_tac x=i in exI)
+	using as(2)[THEN spec[where x=i]] and as2
+	by (auto simp add: less_divide_eq_number_of1)
+      ultimately have False using as by auto  }
+    hence "b$i \<ge> d$i" by(rule ccontr)auto
+    ultimately
+    have "a$i \<le> c$i \<and> d$i \<le> b$i" by auto
+  } note part1 = this
+  thus ?th3 unfolding subset_eq and Ball_def and mem_interval apply auto apply (erule_tac x=ia in allE, simp)+ by (erule_tac x=i in allE, erule_tac x=i in allE, simp)+
+  { assume as:"{c<..<d} \<subseteq> {a<..<b}" "\<forall>i. c$i < d$i"
+    fix i
+    from as(1) have "{c<..<d} \<subseteq> {a..b}" using interval_open_subset_closed[of a b] by auto
+    hence "a$i \<le> c$i \<and> d$i \<le> b$i" using part1 and as(2) by auto  } note * = this
+  thus ?th4 unfolding subset_eq and Ball_def and mem_interval apply auto apply (erule_tac x=ia in allE, simp)+ by (erule_tac x=i in allE, erule_tac x=i in allE, simp)+
+qed
+
+lemma disjoint_interval: fixes a::"real^'n::finite" shows
+  "{a .. b} \<inter> {c .. d} = {} \<longleftrightarrow> (\<exists>i. (b$i < a$i \<or> d$i < c$i \<or> b$i < c$i \<or> d$i < a$i))" (is ?th1) and
+  "{a .. b} \<inter> {c<..<d} = {} \<longleftrightarrow> (\<exists>i. (b$i < a$i \<or> d$i \<le> c$i \<or> b$i \<le> c$i \<or> d$i \<le> a$i))" (is ?th2) and
+  "{a<..<b} \<inter> {c .. d} = {} \<longleftrightarrow> (\<exists>i. (b$i \<le> a$i \<or> d$i < c$i \<or> b$i \<le> c$i \<or> d$i \<le> a$i))" (is ?th3) and
+  "{a<..<b} \<inter> {c<..<d} = {} \<longleftrightarrow> (\<exists>i. (b$i \<le> a$i \<or> d$i \<le> c$i \<or> b$i \<le> c$i \<or> d$i \<le> a$i))" (is ?th4)
+proof-
+  let ?z = "(\<chi> i. ((max (a$i) (c$i)) + (min (b$i) (d$i))) / 2)::real^'n"
+  show ?th1 ?th2 ?th3 ?th4
+  unfolding expand_set_eq and Int_iff and empty_iff and mem_interval and all_conj_distrib[THEN sym] and eq_False
+  apply (auto elim!: allE[where x="?z"])
+  apply ((rule_tac x=x in exI, force) | (rule_tac x=i in exI, force))+
+  done
+qed
+
+lemma inter_interval: fixes a :: "'a::linorder^'n::finite" shows
+ "{a .. b} \<inter> {c .. d} =  {(\<chi> i. max (a$i) (c$i)) .. (\<chi> i. min (b$i) (d$i))}"
+  unfolding expand_set_eq and Int_iff and mem_interval
+  by (auto simp add: less_divide_eq_number_of1 intro!: bexI)
+
+(* Moved interval_open_subset_closed a bit upwards *)
+
+lemma open_interval_lemma: fixes x :: "real" shows
+ "a < x \<Longrightarrow> x < b ==> (\<exists>d>0. \<forall>x'. abs(x' - x) < d --> a < x' \<and> x' < b)"
+  by(rule_tac x="min (x - a) (b - x)" in exI, auto)
+
+lemma open_interval: fixes a :: "real^'n::finite" shows "open {a<..<b}"
+proof-
+  { fix x assume x:"x\<in>{a<..<b}"
+    { fix i
+      have "\<exists>d>0. \<forall>x'. abs (x' - (x$i)) < d \<longrightarrow> a$i < x' \<and> x' < b$i"
+	using x[unfolded mem_interval, THEN spec[where x=i]]
+	using open_interval_lemma[of "a$i" "x$i" "b$i"] by auto  }
+
+    hence "\<forall>i. \<exists>d>0. \<forall>x'. abs (x' - (x$i)) < d \<longrightarrow> a$i < x' \<and> x' < b$i" by auto
+    then obtain d where d:"\<forall>i. 0 < d i \<and> (\<forall>x'. \<bar>x' - x $ i\<bar> < d i \<longrightarrow> a $ i < x' \<and> x' < b $ i)"
+      using bchoice[of "UNIV" "\<lambda>i d. d>0 \<and> (\<forall>x'. \<bar>x' - x $ i\<bar> < d \<longrightarrow> a $ i < x' \<and> x' < b $ i)"] by auto
+
+    let ?d = "Min (range d)"
+    have **:"finite (range d)" "range d \<noteq> {}" by auto
+    have "?d>0" unfolding Min_gr_iff[OF **] using d by auto
+    moreover
+    { fix x' assume as:"dist x' x < ?d"
+      { fix i
+	have "\<bar>x'$i - x $ i\<bar> < d i"
+	  using norm_bound_component_lt[OF as[unfolded dist_norm], of i]
+	  unfolding vector_minus_component and Min_gr_iff[OF **] by auto
+	hence "a $ i < x' $ i" "x' $ i < b $ i" using d[THEN spec[where x=i]] by auto  }
+      hence "a < x' \<and> x' < b" unfolding vector_less_def by auto  }
+    ultimately have "\<exists>e>0. \<forall>x'. dist x' x < e \<longrightarrow> x' \<in> {a<..<b}" by (auto, rule_tac x="?d" in exI, simp)
+  }
+  thus ?thesis unfolding open_dist using open_interval_lemma by auto
+qed
+
+lemma closed_interval: fixes a :: "real^'n::finite" shows "closed {a .. b}"
+proof-
+  { fix x i assume as:"\<forall>e>0. \<exists>x'\<in>{a..b}. x' \<noteq> x \<and> dist x' x < e"(* and xab:"a$i > x$i \<or> b$i < x$i"*)
+    { assume xa:"a$i > x$i"
+      with as obtain y where y:"y\<in>{a..b}" "y \<noteq> x" "dist y x < a$i - x$i" by(erule_tac x="a$i - x$i" in allE)auto
+      hence False unfolding mem_interval and dist_norm
+	using component_le_norm[of "y-x" i, unfolded vector_minus_component] and xa by(auto elim!: allE[where x=i])
+    } hence "a$i \<le> x$i" by(rule ccontr)auto
+    moreover
+    { assume xb:"b$i < x$i"
+      with as obtain y where y:"y\<in>{a..b}" "y \<noteq> x" "dist y x < x$i - b$i" by(erule_tac x="x$i - b$i" in allE)auto
+      hence False unfolding mem_interval and dist_norm
+	using component_le_norm[of "y-x" i, unfolded vector_minus_component] and xb by(auto elim!: allE[where x=i])
+    } hence "x$i \<le> b$i" by(rule ccontr)auto
+    ultimately
+    have "a $ i \<le> x $ i \<and> x $ i \<le> b $ i" by auto }
+  thus ?thesis unfolding closed_limpt islimpt_approachable mem_interval by auto
+qed
+
+lemma interior_closed_interval: fixes a :: "real^'n::finite" shows
+ "interior {a .. b} = {a<..<b}" (is "?L = ?R")
+proof(rule subset_antisym)
+  show "?R \<subseteq> ?L" using interior_maximal[OF interval_open_subset_closed open_interval] by auto
+next
+  { fix x assume "\<exists>T. open T \<and> x \<in> T \<and> T \<subseteq> {a..b}"
+    then obtain s where s:"open s" "x \<in> s" "s \<subseteq> {a..b}" by auto
+    then obtain e where "e>0" and e:"\<forall>x'. dist x' x < e \<longrightarrow> x' \<in> {a..b}" unfolding open_dist and subset_eq by auto
+    { fix i
+      have "dist (x - (e / 2) *\<^sub>R basis i) x < e"
+	   "dist (x + (e / 2) *\<^sub>R basis i) x < e"
+	unfolding dist_norm apply auto
+	unfolding norm_minus_cancel using norm_basis[of i] and `e>0` by auto
+      hence "a $ i \<le> (x - (e / 2) *\<^sub>R basis i) $ i"
+                    "(x + (e / 2) *\<^sub>R basis i) $ i \<le> b $ i"
+	using e[THEN spec[where x="x - (e/2) *\<^sub>R basis i"]]
+	and   e[THEN spec[where x="x + (e/2) *\<^sub>R basis i"]]
+	unfolding mem_interval by (auto elim!: allE[where x=i])
+      hence "a $ i < x $ i" and "x $ i < b $ i"
+	unfolding vector_minus_component and vector_add_component
+	unfolding vector_smult_component and basis_component using `e>0` by auto   }
+    hence "x \<in> {a<..<b}" unfolding mem_interval by auto  }
+  thus "?L \<subseteq> ?R" unfolding interior_def and subset_eq by auto
+qed
+
+lemma bounded_closed_interval: fixes a :: "real^'n::finite" shows
+ "bounded {a .. b}"
+proof-
+  let ?b = "\<Sum>i\<in>UNIV. \<bar>a$i\<bar> + \<bar>b$i\<bar>"
+  { fix x::"real^'n" assume x:"\<forall>i. a $ i \<le> x $ i \<and> x $ i \<le> b $ i"
+    { fix i
+      have "\<bar>x$i\<bar> \<le> \<bar>a$i\<bar> + \<bar>b$i\<bar>" using x[THEN spec[where x=i]] by auto  }
+    hence "(\<Sum>i\<in>UNIV. \<bar>x $ i\<bar>) \<le> ?b" by(rule setsum_mono)
+    hence "norm x \<le> ?b" using norm_le_l1[of x] by auto  }
+  thus ?thesis unfolding interval and bounded_iff by auto
+qed
+
+lemma bounded_interval: fixes a :: "real^'n::finite" shows
+ "bounded {a .. b} \<and> bounded {a<..<b}"
+  using bounded_closed_interval[of a b]
+  using interval_open_subset_closed[of a b]
+  using bounded_subset[of "{a..b}" "{a<..<b}"]
+  by simp
+
+lemma not_interval_univ: fixes a :: "real^'n::finite" shows
+ "({a .. b} \<noteq> UNIV) \<and> ({a<..<b} \<noteq> UNIV)"
+  using bounded_interval[of a b]
+  by auto
+
+lemma compact_interval: fixes a :: "real^'n::finite" shows
+ "compact {a .. b}"
+  using bounded_closed_imp_compact using bounded_interval[of a b] using closed_interval[of a b] by auto
+
+lemma open_interval_midpoint: fixes a :: "real^'n::finite"
+  assumes "{a<..<b} \<noteq> {}" shows "((1/2) *\<^sub>R (a + b)) \<in> {a<..<b}"
+proof-
+  { fix i
+    have "a $ i < ((1 / 2) *\<^sub>R (a + b)) $ i \<and> ((1 / 2) *\<^sub>R (a + b)) $ i < b $ i"
+      using assms[unfolded interval_ne_empty, THEN spec[where x=i]]
+      unfolding vector_smult_component and vector_add_component
+      by(auto simp add: less_divide_eq_number_of1)  }
+  thus ?thesis unfolding mem_interval by auto
+qed
+
+lemma open_closed_interval_convex: fixes x :: "real^'n::finite"
+  assumes x:"x \<in> {a<..<b}" and y:"y \<in> {a .. b}" and e:"0 < e" "e \<le> 1"
+  shows "(e *\<^sub>R x + (1 - e) *\<^sub>R y) \<in> {a<..<b}"
+proof-
+  { fix i
+    have "a $ i = e * a$i + (1 - e) * a$i" unfolding left_diff_distrib by simp
+    also have "\<dots> < e * x $ i + (1 - e) * y $ i" apply(rule add_less_le_mono)
+      using e unfolding mult_less_cancel_left and mult_le_cancel_left apply simp_all
+      using x unfolding mem_interval  apply simp
+      using y unfolding mem_interval  apply simp
+      done
+    finally have "a $ i < (e *\<^sub>R x + (1 - e) *\<^sub>R y) $ i" by auto
+    moreover {
+    have "b $ i = e * b$i + (1 - e) * b$i" unfolding left_diff_distrib by simp
+    also have "\<dots> > e * x $ i + (1 - e) * y $ i" apply(rule add_less_le_mono)
+      using e unfolding mult_less_cancel_left and mult_le_cancel_left apply simp_all
+      using x unfolding mem_interval  apply simp
+      using y unfolding mem_interval  apply simp
+      done
+    finally have "(e *\<^sub>R x + (1 - e) *\<^sub>R y) $ i < b $ i" by auto
+    } ultimately have "a $ i < (e *\<^sub>R x + (1 - e) *\<^sub>R y) $ i \<and> (e *\<^sub>R x + (1 - e) *\<^sub>R y) $ i < b $ i" by auto }
+  thus ?thesis unfolding mem_interval by auto
+qed
+
+lemma closure_open_interval: fixes a :: "real^'n::finite"
+  assumes "{a<..<b} \<noteq> {}"
+  shows "closure {a<..<b} = {a .. b}"
+proof-
+  have ab:"a < b" using assms[unfolded interval_ne_empty] unfolding vector_less_def by auto
+  let ?c = "(1 / 2) *\<^sub>R (a + b)"
+  { fix x assume as:"x \<in> {a .. b}"
+    def f == "\<lambda>n::nat. x + (inverse (real n + 1)) *\<^sub>R (?c - x)"
+    { fix n assume fn:"f n < b \<longrightarrow> a < f n \<longrightarrow> f n = x" and xc:"x \<noteq> ?c"
+      have *:"0 < inverse (real n + 1)" "inverse (real n + 1) \<le> 1" unfolding inverse_le_1_iff by auto
+      have "(inverse (real n + 1)) *\<^sub>R ((1 / 2) *\<^sub>R (a + b)) + (1 - inverse (real n + 1)) *\<^sub>R x =
+	x + (inverse (real n + 1)) *\<^sub>R (((1 / 2) *\<^sub>R (a + b)) - x)"
+        by (auto simp add: algebra_simps)
+      hence "f n < b" and "a < f n" using open_closed_interval_convex[OF open_interval_midpoint[OF assms] as *] unfolding f_def by auto
+      hence False using fn unfolding f_def using xc by(auto simp add: vector_mul_lcancel vector_ssub_ldistrib)  }
+    moreover
+    { assume "\<not> (f ---> x) sequentially"
+      { fix e::real assume "e>0"
+	hence "\<exists>N::nat. inverse (real (N + 1)) < e" using real_arch_inv[of e] apply (auto simp add: Suc_pred') apply(rule_tac x="n - 1" in exI) by auto
+	then obtain N::nat where "inverse (real (N + 1)) < e" by auto
+	hence "\<forall>n\<ge>N. inverse (real n + 1) < e" by (auto, metis Suc_le_mono le_SucE less_imp_inverse_less nat_le_real_less order_less_trans real_of_nat_Suc real_of_nat_Suc_gt_zero)
+	hence "\<exists>N::nat. \<forall>n\<ge>N. inverse (real n + 1) < e" by auto  }
+      hence "((\<lambda>n. inverse (real n + 1)) ---> 0) sequentially"
+	unfolding Lim_sequentially by(auto simp add: dist_norm)
+      hence "(f ---> x) sequentially" unfolding f_def
+	using Lim_add[OF Lim_const, of "\<lambda>n::nat. (inverse (real n + 1)) *\<^sub>R ((1 / 2) *\<^sub>R (a + b) - x)" 0 sequentially x]
+	using Lim_vmul[of "\<lambda>n::nat. inverse (real n + 1)" 0 sequentially "((1 / 2) *\<^sub>R (a + b) - x)"] by auto  }
+    ultimately have "x \<in> closure {a<..<b}"
+      using as and open_interval_midpoint[OF assms] unfolding closure_def unfolding islimpt_sequential by(cases "x=?c")auto  }
+  thus ?thesis using closure_minimal[OF interval_open_subset_closed closed_interval, of a b] by blast
+qed
+
+lemma bounded_subset_open_interval_symmetric: fixes s::"(real^'n::finite) set"
+  assumes "bounded s"  shows "\<exists>a. s \<subseteq> {-a<..<a}"
+proof-
+  obtain b where "b>0" and b:"\<forall>x\<in>s. norm x \<le> b" using assms[unfolded bounded_pos] by auto
+  def a \<equiv> "(\<chi> i. b+1)::real^'n"
+  { fix x assume "x\<in>s"
+    fix i
+    have "(-a)$i < x$i" and "x$i < a$i" using b[THEN bspec[where x=x], OF `x\<in>s`] and component_le_norm[of x i]
+      unfolding vector_uminus_component and a_def and Cart_lambda_beta by auto
+  }
+  thus ?thesis by(auto intro: exI[where x=a] simp add: vector_less_def)
+qed
+
+lemma bounded_subset_open_interval:
+  fixes s :: "(real ^ 'n::finite) set"
+  shows "bounded s ==> (\<exists>a b. s \<subseteq> {a<..<b})"
+  by (auto dest!: bounded_subset_open_interval_symmetric)
+
+lemma bounded_subset_closed_interval_symmetric:
+  fixes s :: "(real ^ 'n::finite) set"
+  assumes "bounded s" shows "\<exists>a. s \<subseteq> {-a .. a}"
+proof-
+  obtain a where "s \<subseteq> {- a<..<a}" using bounded_subset_open_interval_symmetric[OF assms] by auto
+  thus ?thesis using interval_open_subset_closed[of "-a" a] by auto
+qed
+
+lemma bounded_subset_closed_interval:
+  fixes s :: "(real ^ 'n::finite) set"
+  shows "bounded s ==> (\<exists>a b. s \<subseteq> {a .. b})"
+  using bounded_subset_closed_interval_symmetric[of s] by auto
+
+lemma frontier_closed_interval:
+  fixes a b :: "real ^ _"
+  shows "frontier {a .. b} = {a .. b} - {a<..<b}"
+  unfolding frontier_def unfolding interior_closed_interval and closure_closed[OF closed_interval] ..
+
+lemma frontier_open_interval:
+  fixes a b :: "real ^ _"
+  shows "frontier {a<..<b} = (if {a<..<b} = {} then {} else {a .. b} - {a<..<b})"
+proof(cases "{a<..<b} = {}")
+  case True thus ?thesis using frontier_empty by auto
+next
+  case False thus ?thesis unfolding frontier_def and closure_open_interval[OF False] and interior_open[OF open_interval] by auto
+qed
+
+lemma inter_interval_mixed_eq_empty: fixes a :: "real^'n::finite"
+  assumes "{c<..<d} \<noteq> {}"  shows "{a<..<b} \<inter> {c .. d} = {} \<longleftrightarrow> {a<..<b} \<inter> {c<..<d} = {}"
+  unfolding closure_open_interval[OF assms, THEN sym] unfolding open_inter_closure_eq_empty[OF open_interval] ..
+
+
+(* Some special cases for intervals in R^1.                                  *)
+
+lemma all_1: "(\<forall>x::1. P x) \<longleftrightarrow> P 1"
+  by (metis num1_eq_iff)
+
+lemma ex_1: "(\<exists>x::1. P x) \<longleftrightarrow> P 1"
+  by auto (metis num1_eq_iff)
+
+lemma interval_cases_1: fixes x :: "real^1" shows
+ "x \<in> {a .. b} ==> x \<in> {a<..<b} \<or> (x = a) \<or> (x = b)"
+  by(simp add:  Cart_eq vector_less_def vector_less_eq_def all_1, auto)
+
+lemma in_interval_1: fixes x :: "real^1" shows
+ "(x \<in> {a .. b} \<longleftrightarrow> dest_vec1 a \<le> dest_vec1 x \<and> dest_vec1 x \<le> dest_vec1 b) \<and>
+  (x \<in> {a<..<b} \<longleftrightarrow> dest_vec1 a < dest_vec1 x \<and> dest_vec1 x < dest_vec1 b)"
+by(simp add: Cart_eq vector_less_def vector_less_eq_def all_1 dest_vec1_def)
+
+lemma interval_eq_empty_1: fixes a :: "real^1" shows
+  "{a .. b} = {} \<longleftrightarrow> dest_vec1 b < dest_vec1 a"
+  "{a<..<b} = {} \<longleftrightarrow> dest_vec1 b \<le> dest_vec1 a"
+  unfolding interval_eq_empty and ex_1 and dest_vec1_def by auto
+
+lemma subset_interval_1: fixes a :: "real^1" shows
+ "({a .. b} \<subseteq> {c .. d} \<longleftrightarrow>  dest_vec1 b < dest_vec1 a \<or>
+                dest_vec1 c \<le> dest_vec1 a \<and> dest_vec1 a \<le> dest_vec1 b \<and> dest_vec1 b \<le> dest_vec1 d)"
+ "({a .. b} \<subseteq> {c<..<d} \<longleftrightarrow>  dest_vec1 b < dest_vec1 a \<or>
+                dest_vec1 c < dest_vec1 a \<and> dest_vec1 a \<le> dest_vec1 b \<and> dest_vec1 b < dest_vec1 d)"
+ "({a<..<b} \<subseteq> {c .. d} \<longleftrightarrow>  dest_vec1 b \<le> dest_vec1 a \<or>
+                dest_vec1 c \<le> dest_vec1 a \<and> dest_vec1 a < dest_vec1 b \<and> dest_vec1 b \<le> dest_vec1 d)"
+ "({a<..<b} \<subseteq> {c<..<d} \<longleftrightarrow> dest_vec1 b \<le> dest_vec1 a \<or>
+                dest_vec1 c \<le> dest_vec1 a \<and> dest_vec1 a < dest_vec1 b \<and> dest_vec1 b \<le> dest_vec1 d)"
+  unfolding subset_interval[of a b c d] unfolding all_1 and dest_vec1_def by auto
+
+lemma eq_interval_1: fixes a :: "real^1" shows
+ "{a .. b} = {c .. d} \<longleftrightarrow>
+          dest_vec1 b < dest_vec1 a \<and> dest_vec1 d < dest_vec1 c \<or>
+          dest_vec1 a = dest_vec1 c \<and> dest_vec1 b = dest_vec1 d"
+using set_eq_subset[of "{a .. b}" "{c .. d}"]
+using subset_interval_1(1)[of a b c d]
+using subset_interval_1(1)[of c d a b]
+by auto (* FIXME: slow *)
+
+lemma disjoint_interval_1: fixes a :: "real^1" shows
+  "{a .. b} \<inter> {c .. d} = {} \<longleftrightarrow> dest_vec1 b < dest_vec1 a \<or> dest_vec1 d < dest_vec1 c  \<or>  dest_vec1 b < dest_vec1 c \<or> dest_vec1 d < dest_vec1 a"
+  "{a .. b} \<inter> {c<..<d} = {} \<longleftrightarrow> dest_vec1 b < dest_vec1 a \<or> dest_vec1 d \<le> dest_vec1 c  \<or>  dest_vec1 b \<le> dest_vec1 c \<or> dest_vec1 d \<le> dest_vec1 a"
+  "{a<..<b} \<inter> {c .. d} = {} \<longleftrightarrow> dest_vec1 b \<le> dest_vec1 a \<or> dest_vec1 d < dest_vec1 c  \<or>  dest_vec1 b \<le> dest_vec1 c \<or> dest_vec1 d \<le> dest_vec1 a"
+  "{a<..<b} \<inter> {c<..<d} = {} \<longleftrightarrow> dest_vec1 b \<le> dest_vec1 a \<or> dest_vec1 d \<le> dest_vec1 c  \<or>  dest_vec1 b \<le> dest_vec1 c \<or> dest_vec1 d \<le> dest_vec1 a"
+  unfolding disjoint_interval and dest_vec1_def ex_1 by auto
+
+lemma open_closed_interval_1: fixes a :: "real^1" shows
+ "{a<..<b} = {a .. b} - {a, b}"
+  unfolding expand_set_eq apply simp unfolding vector_less_def and vector_less_eq_def and all_1 and dest_vec1_eq[THEN sym] and dest_vec1_def by auto
+
+lemma closed_open_interval_1: "dest_vec1 (a::real^1) \<le> dest_vec1 b ==> {a .. b} = {a<..<b} \<union> {a,b}"
+  unfolding expand_set_eq apply simp unfolding vector_less_def and vector_less_eq_def and all_1 and dest_vec1_eq[THEN sym] and dest_vec1_def by auto
+
+(* Some stuff for half-infinite intervals too; FIXME: notation?  *)
+
+lemma closed_interval_left: fixes b::"real^'n::finite"
+  shows "closed {x::real^'n. \<forall>i. x$i \<le> b$i}"
+proof-
+  { fix i
+    fix x::"real^'n" assume x:"\<forall>e>0. \<exists>x'\<in>{x. \<forall>i. x $ i \<le> b $ i}. x' \<noteq> x \<and> dist x' x < e"
+    { assume "x$i > b$i"
+      then obtain y where "y $ i \<le> b $ i"  "y \<noteq> x"  "dist y x < x$i - b$i" using x[THEN spec[where x="x$i - b$i"]] by auto
+      hence False using component_le_norm[of "y - x" i] unfolding dist_norm and vector_minus_component by auto   }
+    hence "x$i \<le> b$i" by(rule ccontr)auto  }
+  thus ?thesis unfolding closed_limpt unfolding islimpt_approachable by blast
+qed
+
+lemma closed_interval_right: fixes a::"real^'n::finite"
+  shows "closed {x::real^'n. \<forall>i. a$i \<le> x$i}"
+proof-
+  { fix i
+    fix x::"real^'n" assume x:"\<forall>e>0. \<exists>x'\<in>{x. \<forall>i. a $ i \<le> x $ i}. x' \<noteq> x \<and> dist x' x < e"
+    { assume "a$i > x$i"
+      then obtain y where "a $ i \<le> y $ i"  "y \<noteq> x"  "dist y x < a$i - x$i" using x[THEN spec[where x="a$i - x$i"]] by auto
+      hence False using component_le_norm[of "y - x" i] unfolding dist_norm and vector_minus_component by auto   }
+    hence "a$i \<le> x$i" by(rule ccontr)auto  }
+  thus ?thesis unfolding closed_limpt unfolding islimpt_approachable by blast
+qed
+
+subsection{* Intervals in general, including infinite and mixtures of open and closed. *}
+
+definition "is_interval s \<longleftrightarrow> (\<forall>a\<in>s. \<forall>b\<in>s. \<forall>x. (\<forall>i. ((a$i \<le> x$i \<and> x$i \<le> b$i) \<or> (b$i \<le> x$i \<and> x$i \<le> a$i)))  \<longrightarrow> x \<in> s)"
+
+lemma is_interval_interval: "is_interval {a .. b::real^'n::finite}" (is ?th1) "is_interval {a<..<b}" (is ?th2) proof - 
+  have *:"\<And>x y z::real. x < y \<Longrightarrow> y < z \<Longrightarrow> 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 {}"
+  unfolding is_interval_def
+  by simp
+
+lemma is_interval_univ:
+ "is_interval UNIV"
+  unfolding is_interval_def
+  by simp
+
+subsection{* Closure of halfspaces and hyperplanes.                                    *}
+
+lemma Lim_inner:
+  assumes "(f ---> l) net"  shows "((\<lambda>y. inner a (f y)) ---> inner a l) net"
+  by (intro tendsto_intros assms)
+
+lemma continuous_at_inner: "continuous (at x) (inner a)"
+  unfolding continuous_at by (intro tendsto_intros)
+
+lemma continuous_on_inner:
+  fixes s :: "'a::real_inner set"
+  shows "continuous_on s (inner a)"
+  unfolding continuous_on by (rule ballI) (intro tendsto_intros)
+
+lemma closed_halfspace_le: "closed {x. inner a x \<le> b}"
+proof-
+  have "\<forall>x. continuous (at x) (inner a)"
+    unfolding continuous_at by (rule allI) (intro tendsto_intros)
+  hence "closed (inner a -` {..b})"
+    using closed_real_atMost by (rule continuous_closed_vimage)
+  moreover have "{x. inner a x \<le> b} = inner a -` {..b}" by auto
+  ultimately show ?thesis by simp
+qed
+
+lemma closed_halfspace_ge: "closed {x. inner a x \<ge> b}"
+  using closed_halfspace_le[of "-a" "-b"] unfolding inner_minus_left by auto
+
+lemma closed_hyperplane: "closed {x. inner a x = b}"
+proof-
+  have "{x. inner a x = b} = {x. inner a x \<ge> b} \<inter> {x. inner a x \<le> b}" by auto
+  thus ?thesis using closed_halfspace_le[of a b] and closed_halfspace_ge[of b a] using closed_Int by auto
+qed
+
+lemma closed_halfspace_component_le:
+  shows "closed {x::real^'n::finite. x$i \<le> a}"
+  using closed_halfspace_le[of "(basis i)::real^'n" a] unfolding inner_basis[OF assms] by auto
+
+lemma closed_halfspace_component_ge:
+  shows "closed {x::real^'n::finite. x$i \<ge> a}"
+  using closed_halfspace_ge[of a "(basis i)::real^'n"] unfolding inner_basis[OF assms] by auto
+
+text{* Openness of halfspaces.                                                   *}
+
+lemma open_halfspace_lt: "open {x. inner a x < b}"
+proof-
+  have "UNIV - {x. b \<le> inner a x} = {x. inner a x < b}" by auto
+  thus ?thesis using closed_halfspace_ge[unfolded closed_def Compl_eq_Diff_UNIV, of b a] by auto
+qed
+
+lemma open_halfspace_gt: "open {x. inner a x > b}"
+proof-
+  have "UNIV - {x. b \<ge> inner a x} = {x. inner a x > b}" by auto
+  thus ?thesis using closed_halfspace_le[unfolded closed_def Compl_eq_Diff_UNIV, of a b] by auto
+qed
+
+lemma open_halfspace_component_lt:
+  shows "open {x::real^'n::finite. x$i < a}"
+  using open_halfspace_lt[of "(basis i)::real^'n" a] unfolding inner_basis[OF assms] by auto
+
+lemma open_halfspace_component_gt:
+  shows "open {x::real^'n::finite. x$i  > a}"
+  using open_halfspace_gt[of a "(basis i)::real^'n"] unfolding inner_basis[OF assms] by auto
+
+text{* This gives a simple derivation of limit component bounds.                 *}
+
+lemma Lim_component_le: fixes f :: "'a \<Rightarrow> real^'n::finite"
+  assumes "(f ---> l) net" "\<not> (trivial_limit net)"  "eventually (\<lambda>x. f(x)$i \<le> b) net"
+  shows "l$i \<le> b"
+proof-
+  { fix x have "x \<in> {x::real^'n. inner (basis i) x \<le> b} \<longleftrightarrow> x$i \<le> b" unfolding inner_basis by auto } note * = this
+  show ?thesis using Lim_in_closed_set[of "{x. inner (basis i) x \<le> b}" f net l] unfolding *
+    using closed_halfspace_le[of "(basis i)::real^'n" b] and assms(1,2,3) by auto
+qed
+
+lemma Lim_component_ge: fixes f :: "'a \<Rightarrow> real^'n::finite"
+  assumes "(f ---> l) net"  "\<not> (trivial_limit net)"  "eventually (\<lambda>x. b \<le> (f x)$i) net"
+  shows "b \<le> l$i"
+proof-
+  { fix x have "x \<in> {x::real^'n. inner (basis i) x \<ge> b} \<longleftrightarrow> x$i \<ge> b" unfolding inner_basis by auto } note * = this
+  show ?thesis using Lim_in_closed_set[of "{x. inner (basis i) x \<ge> b}" f net l] unfolding *
+    using closed_halfspace_ge[of b "(basis i)::real^'n"] and assms(1,2,3) by auto
+qed
+
+lemma Lim_component_eq: fixes f :: "'a \<Rightarrow> real^'n::finite"
+  assumes net:"(f ---> l) net" "~(trivial_limit net)" and ev:"eventually (\<lambda>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
+
+lemma Lim_drop_le: fixes f :: "'a \<Rightarrow> real^1" shows
+  "(f ---> l) net \<Longrightarrow> ~(trivial_limit net) \<Longrightarrow> eventually (\<lambda>x. dest_vec1 (f x) \<le> b) net ==> dest_vec1 l \<le> b"
+  using Lim_component_le[of f l net 1 b] unfolding dest_vec1_def by auto
+
+lemma Lim_drop_ge: fixes f :: "'a \<Rightarrow> real^1" shows
+ "(f ---> l) net \<Longrightarrow> ~(trivial_limit net) \<Longrightarrow> eventually (\<lambda>x. b \<le> dest_vec1 (f x)) net ==> b \<le> dest_vec1 l"
+  using Lim_component_ge[of f l net b 1] unfolding dest_vec1_def by auto
+
+text{* Limits relative to a union.                                               *}
+
+lemma eventually_within_Un:
+  "eventually P (net within (s \<union> t)) \<longleftrightarrow>
+    eventually P (net within s) \<and> eventually P (net within t)"
+  unfolding Limits.eventually_within
+  by (auto elim!: eventually_rev_mp)
+
+lemma Lim_within_union:
+ "(f ---> l) (net within (s \<union> t)) \<longleftrightarrow>
+  (f ---> l) (net within s) \<and> (f ---> l) (net within t)"
+  unfolding tendsto_def
+  by (auto simp add: eventually_within_Un)
+
+lemma continuous_on_union:
+  assumes "closed s" "closed t" "continuous_on s f" "continuous_on t f"
+  shows "continuous_on (s \<union> t) f"
+  using assms unfolding continuous_on unfolding Lim_within_union
+  unfolding Lim unfolding trivial_limit_within unfolding closed_limpt by auto
+
+lemma continuous_on_cases:
+  assumes "closed s" "closed t" "continuous_on s f" "continuous_on t g"
+          "\<forall>x. (x\<in>s \<and> \<not> P x) \<or> (x \<in> t \<and> P x) \<longrightarrow> f x = g x"
+  shows "continuous_on (s \<union> t) (\<lambda>x. if P x then f x else g x)"
+proof-
+  let ?h = "(\<lambda>x. if P x then f x else g x)"
+  have "\<forall>x\<in>s. f x = (if P x then f x else g x)" using assms(5) by auto
+  hence "continuous_on s ?h" using continuous_on_eq[of s f ?h] using assms(3) by auto
+  moreover
+  have "\<forall>x\<in>t. g x = (if P x then f x else g x)" using assms(5) by auto
+  hence "continuous_on t ?h" using continuous_on_eq[of t g ?h] using assms(4) by auto
+  ultimately show ?thesis using continuous_on_union[OF assms(1,2), of ?h] by auto
+qed
+
+
+text{* Some more convenient intermediate-value theorem formulations.             *}
+
+lemma connected_ivt_hyperplane:
+  assumes "connected s" "x \<in> s" "y \<in> s" "inner a x \<le> b" "b \<le> inner a y"
+  shows "\<exists>z \<in> s. inner a z = b"
+proof(rule ccontr)
+  assume as:"\<not> (\<exists>z\<in>s. inner a z = b)"
+  let ?A = "{x. inner a x < b}"
+  let ?B = "{x. inner a x > b}"
+  have "open ?A" "open ?B" using open_halfspace_lt and open_halfspace_gt by auto
+  moreover have "?A \<inter> ?B = {}" by auto
+  moreover have "s \<subseteq> ?A \<union> ?B" using as by auto
+  ultimately show False using assms(1)[unfolded connected_def not_ex, THEN spec[where x="?A"], THEN spec[where x="?B"]] and assms(2-5) by auto
+qed
+
+lemma connected_ivt_component: fixes x::"real^'n::finite" shows
+ "connected s \<Longrightarrow> x \<in> s \<Longrightarrow> y \<in> s \<Longrightarrow> x$k \<le> a \<Longrightarrow> a \<le> y$k \<Longrightarrow> (\<exists>z\<in>s.  z$k = a)"
+  using connected_ivt_hyperplane[of s x y "(basis k)::real^'n" a] by (auto simp add: inner_basis)
+
+text{* Also more convenient formulations of monotone convergence.                *}
+
+lemma bounded_increasing_convergent: fixes s::"nat \<Rightarrow> real^1"
+  assumes "bounded {s n| n::nat. True}"  "\<forall>n. dest_vec1(s n) \<le> dest_vec1(s(Suc n))"
+  shows "\<exists>l. (s ---> l) sequentially"
+proof-
+  obtain a where a:"\<forall>n. \<bar>dest_vec1 (s n)\<bar> \<le>  a" using assms(1)[unfolded bounded_iff abs_dest_vec1] by auto
+  { fix m::nat
+    have "\<And> n. n\<ge>m \<longrightarrow> dest_vec1 (s m) \<le> dest_vec1 (s n)"
+      apply(induct_tac n) apply simp using assms(2) apply(erule_tac x="na" in allE) by(auto simp add: not_less_eq_eq)  }
+  hence "\<forall>m n. m \<le> n \<longrightarrow> dest_vec1 (s m) \<le> dest_vec1 (s n)" by auto
+  then obtain l where "\<forall>e>0. \<exists>N. \<forall>n\<ge>N. \<bar>dest_vec1 (s n) - l\<bar> < e" using convergent_bounded_monotone[OF a] unfolding monoseq_def by auto
+  thus ?thesis unfolding Lim_sequentially apply(rule_tac x="vec1 l" in exI)
+    unfolding dist_norm unfolding abs_dest_vec1 and dest_vec1_sub by auto
+qed
+
+subsection{* Basic homeomorphism definitions.                                          *}
+
+definition "homeomorphism s t f g \<equiv>
+     (\<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::metric_space set \<Rightarrow> 'b::metric_space set \<Rightarrow> bool"
+    (infixr "homeomorphic" 60) where
+  homeomorphic_def: "s homeomorphic t \<equiv> (\<exists>f g. homeomorphism s t f g)"
+
+lemma homeomorphic_refl: "s homeomorphic s"
+  unfolding homeomorphic_def
+  unfolding homeomorphism_def
+  using continuous_on_id
+  apply(rule_tac x = "(\<lambda>x. x)" in exI)
+  apply(rule_tac x = "(\<lambda>x. x)" in exI)
+  by blast
+
+lemma homeomorphic_sym:
+ "s homeomorphic t \<longleftrightarrow> t homeomorphic s"
+unfolding homeomorphic_def
+unfolding homeomorphism_def
+by blast (* FIXME: slow *)
+
+lemma homeomorphic_trans:
+  assumes "s homeomorphic t" "t homeomorphic u" shows "s homeomorphic u"
+proof-
+  obtain f1 g1 where fg1:"\<forall>x\<in>s. g1 (f1 x) = x"  "f1 ` s = t" "continuous_on s f1" "\<forall>y\<in>t. f1 (g1 y) = y" "g1 ` t = s" "continuous_on t g1"
+    using assms(1) unfolding homeomorphic_def homeomorphism_def by auto
+  obtain f2 g2 where fg2:"\<forall>x\<in>t. g2 (f2 x) = x"  "f2 ` t = u" "continuous_on t f2" "\<forall>y\<in>u. f2 (g2 y) = y" "g2 ` u = t" "continuous_on u g2"
+    using assms(2) unfolding homeomorphic_def homeomorphism_def by auto
+
+  { fix x assume "x\<in>s" hence "(g1 \<circ> g2) ((f2 \<circ> f1) x) = x" using fg1(1)[THEN bspec[where x=x]] and fg2(1)[THEN bspec[where x="f1 x"]] and fg1(2) by auto }
+  moreover have "(f2 \<circ> f1) ` s = u" using fg1(2) fg2(2) by auto
+  moreover have "continuous_on s (f2 \<circ> f1)" using continuous_on_compose[OF fg1(3)] and fg2(3) unfolding fg1(2) by auto
+  moreover { fix y assume "y\<in>u" hence "(f2 \<circ> f1) ((g1 \<circ> g2) y) = y" using fg2(4)[THEN bspec[where x=y]] and fg1(4)[THEN bspec[where x="g2 y"]] and fg2(5) by auto }
+  moreover have "(g1 \<circ> g2) ` u = s" using fg1(5) fg2(5) by auto
+  moreover have "continuous_on u (g1 \<circ> g2)" using continuous_on_compose[OF fg2(6)] and fg1(6)  unfolding fg2(5) by auto
+  ultimately show ?thesis unfolding homeomorphic_def homeomorphism_def apply(rule_tac x="f2 \<circ> f1" in exI) apply(rule_tac x="g1 \<circ> g2" in exI) by auto
+qed
+
+lemma homeomorphic_minimal:
+ "s homeomorphic t \<longleftrightarrow>
+    (\<exists>f g. (\<forall>x\<in>s. f(x) \<in> t \<and> (g(f(x)) = x)) \<and>
+           (\<forall>y\<in>t. g(y) \<in> s \<and> (f(g(y)) = y)) \<and>
+           continuous_on s f \<and> continuous_on t g)"
+unfolding homeomorphic_def homeomorphism_def
+apply auto apply (rule_tac x=f in exI) apply (rule_tac x=g in exI)
+apply auto apply (rule_tac x=f in exI) apply (rule_tac x=g in exI) apply auto
+unfolding image_iff
+apply(erule_tac x="g x" in ballE) apply(erule_tac x="x" in ballE)
+apply auto apply(rule_tac x="g x" in bexI) apply auto
+apply(erule_tac x="f x" in ballE) apply(erule_tac x="x" in ballE)
+apply auto apply(rule_tac x="f x" in bexI) by auto
+
+subsection{* Relatively weak hypotheses if a set is compact.                           *}
+
+definition "inv_on f s = (\<lambda>x. SOME y. y\<in>s \<and> f y = x)"
+
+lemma assumes "inj_on f s" "x\<in>s"
+  shows "inv_on f s (f x) = x"
+ using assms unfolding inj_on_def inv_on_def by auto
+
+lemma homeomorphism_compact:
+  fixes f :: "'a::heine_borel \<Rightarrow> 'b::heine_borel"
+    (* class constraint due to continuous_on_inverse *)
+  assumes "compact s" "continuous_on s f"  "f ` s = t"  "inj_on f s"
+  shows "\<exists>g. homeomorphism s t f g"
+proof-
+  def g \<equiv> "\<lambda>x. SOME y. y\<in>s \<and> f y = x"
+  have g:"\<forall>x\<in>s. g (f x) = x" using assms(3) assms(4)[unfolded inj_on_def] unfolding g_def by auto
+  { fix y assume "y\<in>t"
+    then obtain x where x:"f x = y" "x\<in>s" using assms(3) by auto
+    hence "g (f x) = x" using g by auto
+    hence "f (g y) = y" unfolding x(1)[THEN sym] by auto  }
+  hence g':"\<forall>x\<in>t. f (g x) = x" by auto
+  moreover
+  { fix x
+    have "x\<in>s \<Longrightarrow> x \<in> g ` t" using g[THEN bspec[where x=x]] unfolding image_iff using assms(3) by(auto intro!: bexI[where x="f x"])
+    moreover
+    { assume "x\<in>g ` t"
+      then obtain y where y:"y\<in>t" "g y = x" by auto
+      then obtain x' where x':"x'\<in>s" "f x' = y" using assms(3) by auto
+      hence "x \<in> s" unfolding g_def using someI2[of "\<lambda>b. b\<in>s \<and> f b = y" x' "\<lambda>x. x\<in>s"] unfolding y(2)[THEN sym] and g_def by auto }
+    ultimately have "x\<in>s \<longleftrightarrow> x \<in> g ` t" by auto  }
+  hence "g ` t = s" by auto
+  ultimately
+  show ?thesis unfolding homeomorphism_def homeomorphic_def
+    apply(rule_tac x=g in exI) using g and assms(3) and continuous_on_inverse[OF assms(2,1), of g, unfolded assms(3)] and assms(2) by auto
+qed
+
+lemma homeomorphic_compact:
+  fixes f :: "'a::heine_borel \<Rightarrow> 'b::heine_borel"
+    (* class constraint due to continuous_on_inverse *)
+  shows "compact s \<Longrightarrow> continuous_on s f \<Longrightarrow> (f ` s = t) \<Longrightarrow> inj_on f s
+          \<Longrightarrow> s homeomorphic t"
+  unfolding homeomorphic_def by(metis homeomorphism_compact)
+
+text{* Preservation of topological properties.                                   *}
+
+lemma homeomorphic_compactness:
+ "s homeomorphic t ==> (compact s \<longleftrightarrow> compact t)"
+unfolding homeomorphic_def homeomorphism_def
+by (metis compact_continuous_image)
+
+text{* Results on translation, scaling etc.                                      *}
+
+lemma homeomorphic_scaling:
+  fixes s :: "'a::real_normed_vector set"
+  assumes "c \<noteq> 0"  shows "s homeomorphic ((\<lambda>x. c *\<^sub>R x) ` s)"
+  unfolding homeomorphic_minimal
+  apply(rule_tac x="\<lambda>x. c *\<^sub>R x" in exI)
+  apply(rule_tac x="\<lambda>x. (1 / c) *\<^sub>R x" in exI)
+  using assms apply auto
+  using continuous_on_cmul[OF continuous_on_id] by auto
+
+lemma homeomorphic_translation:
+  fixes s :: "'a::real_normed_vector set"
+  shows "s homeomorphic ((\<lambda>x. a + x) ` s)"
+  unfolding homeomorphic_minimal
+  apply(rule_tac x="\<lambda>x. a + x" in exI)
+  apply(rule_tac x="\<lambda>x. -a + x" in exI)
+  using continuous_on_add[OF continuous_on_const continuous_on_id] by auto
+
+lemma homeomorphic_affinity:
+  fixes s :: "'a::real_normed_vector set"
+  assumes "c \<noteq> 0"  shows "s homeomorphic ((\<lambda>x. a + c *\<^sub>R x) ` s)"
+proof-
+  have *:"op + a ` op *\<^sub>R c ` s = (\<lambda>x. a + c *\<^sub>R x) ` s" by auto
+  show ?thesis
+    using homeomorphic_trans
+    using homeomorphic_scaling[OF assms, of s]
+    using homeomorphic_translation[of "(\<lambda>x. c *\<^sub>R x) ` s" a] unfolding * by auto
+qed
+
+lemma homeomorphic_balls:
+  fixes a b ::"'a::real_normed_vector" (* FIXME: generalize to metric_space *)
+  assumes "0 < d"  "0 < e"
+  shows "(ball a d) homeomorphic  (ball b e)" (is ?th)
+        "(cball a d) homeomorphic (cball b e)" (is ?cth)
+proof-
+  have *:"\<bar>e / d\<bar> > 0" "\<bar>d / e\<bar> >0" using assms using divide_pos_pos by auto
+  show ?th unfolding homeomorphic_minimal
+    apply(rule_tac x="\<lambda>x. b + (e/d) *\<^sub>R (x - a)" in exI)
+    apply(rule_tac x="\<lambda>x. a + (d/e) *\<^sub>R (x - b)" in exI)
+    using assms apply (auto simp add: dist_commute)
+    unfolding dist_norm
+    apply (auto simp add: pos_divide_less_eq mult_strict_left_mono)
+    unfolding continuous_on
+    by (intro ballI tendsto_intros, simp, assumption)+
+next
+  have *:"\<bar>e / d\<bar> > 0" "\<bar>d / e\<bar> >0" using assms using divide_pos_pos by auto
+  show ?cth unfolding homeomorphic_minimal
+    apply(rule_tac x="\<lambda>x. b + (e/d) *\<^sub>R (x - a)" in exI)
+    apply(rule_tac x="\<lambda>x. a + (d/e) *\<^sub>R (x - b)" in exI)
+    using assms apply (auto simp add: dist_commute)
+    unfolding dist_norm
+    apply (auto simp add: pos_divide_le_eq)
+    unfolding continuous_on
+    by (intro ballI tendsto_intros, simp, assumption)+
+qed
+
+text{* "Isometry" (up to constant bounds) of injective linear map etc.           *}
+
+lemma cauchy_isometric:
+  fixes x :: "nat \<Rightarrow> real ^ 'n::finite"
+  assumes e:"0 < e" 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)"
+  shows "Cauchy x"
+proof-
+  interpret f: bounded_linear f by fact
+  { fix d::real assume "d>0"
+    then obtain N where N:"\<forall>n\<ge>N. norm (f (x n) - f (x N)) < e * d"
+      using cf[unfolded cauchy o_def dist_norm, THEN spec[where x="e*d"]] and e and mult_pos_pos[of e d] by auto
+    { fix n assume "n\<ge>N"
+      hence "norm (f (x n - x N)) < e * d" using N[THEN spec[where x=n]] unfolding f.diff[THEN sym] by auto
+      moreover have "e * norm (x n - x N) \<le> norm (f (x n - x N))"
+	using subspace_sub[OF s, of "x n" "x N"] using xs[THEN spec[where x=N]] and xs[THEN spec[where x=n]]
+	using normf[THEN bspec[where x="x n - x N"]] by auto
+      ultimately have "norm (x n - x N) < d" using `e>0`
+	using mult_left_less_imp_less[of e "norm (x n - x N)" d] by auto   }
+    hence "\<exists>N. \<forall>n\<ge>N. norm (x n - x N) < d" by auto }
+  thus ?thesis unfolding cauchy and dist_norm by auto
+qed
+
+lemma complete_isometric_image:
+  fixes f :: "real ^ _ \<Rightarrow> real ^ _"
+  assumes "0 < e" and s:"subspace s" and f:"bounded_linear f" and normf:"\<forall>x\<in>s. norm(f x) \<ge> e * norm(x)" and cs:"complete s"
+  shows "complete(f ` s)"
+proof-
+  { 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)" unfolding image_iff and Bex_def
+      using choice[of "\<lambda> n xa. xa \<in> s \<and> g n = f xa"] by auto
+    hence x:"\<forall>n. x n \<in> s"  "\<forall>n. g n = f (x n)" by auto
+    hence "f \<circ> x = g" unfolding expand_fun_eq 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
+    hence "\<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  }
+  thus ?thesis unfolding complete_def by auto
+qed
+
+lemma dist_0_norm:
+  fixes x :: "'a::real_normed_vector"
+  shows "dist 0 x = norm x"
+unfolding dist_norm by simp
+
+lemma injective_imp_isometric: fixes f::"real^'m::finite \<Rightarrow> real^'n::finite"
+  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)"
+proof(cases "s \<subseteq> {0::real^'m}")
+  case True
+  { fix x assume "x \<in> s"
+    hence "x = 0" using True by auto
+    hence "norm x \<le> norm (f x)" by auto  }
+  thus ?thesis by(auto intro!: exI[where x=1])
+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
+  let ?S = "{f x| x. (x \<in> s \<and> norm x = norm a)}"
+  let ?S' = "{x::real^'m. x\<in>s \<and> norm x = norm a}"
+  let ?S'' = "{x::real^'m. norm x = norm a}"
+
+  have "?S'' = frontier(cball 0 (norm a))" unfolding frontier_cball and dist_norm by (auto simp add: norm_minus_cancel)
+  hence "compact ?S''" using compact_frontier[OF compact_cball, of 0 "norm a"] by auto
+  moreover have "?S' = s \<inter> ?S''" by auto
+  ultimately have "compact ?S'" using closed_inter_compact[of s ?S''] using s(1) by auto
+  moreover have *:"f ` ?S' = ?S" by auto
+  ultimately have "compact ?S" using compact_continuous_image[OF linear_continuous_on[OF f(1)], of ?S'] by auto
+  hence "closed ?S" using compact_imp_closed by auto
+  moreover have "?S \<noteq> {}" using a by auto
+  ultimately obtain b' where "b'\<in>?S" "\<forall>y\<in>?S. norm b' \<le> norm y" using distance_attains_inf[of ?S 0] unfolding dist_0_norm by auto
+  then obtain b where "b\<in>s" and ba:"norm b = norm a" and b:"\<forall>x\<in>{x \<in> s. norm x = norm a}. norm (f b) \<le> norm (f x)" unfolding *[THEN sym] unfolding image_iff by auto
+
+  let ?e = "norm (f b) / norm b"
+  have "norm b > 0" using ba and a and norm_ge_zero by auto
+  moreover have "norm (f b) > 0" using f(2)[THEN bspec[where x=b], OF `b\<in>s`] using `norm b >0` unfolding zero_less_norm_iff by auto
+  ultimately have "0 < norm (f b) / norm b" by(simp only: divide_pos_pos)
+  moreover
+  { fix x assume "x\<in>s"
+    hence "norm (f b) / norm b * norm x \<le> norm (f x)"
+    proof(cases "x=0")
+      case True thus "norm (f b) / norm b * norm x \<le> norm (f x)" by auto
+    next
+      case False
+      hence *:"0 < norm a / norm x" using `a\<noteq>0` unfolding zero_less_norm_iff[THEN sym] by(simp only: divide_pos_pos)
+      have "\<forall>c. \<forall>x\<in>s. c *\<^sub>R x \<in> s" using s[unfolded subspace_def smult_conv_scaleR] by auto
+      hence "(norm a / norm x) *\<^sub>R x \<in> {x \<in> s. norm x = norm a}" using `x\<in>s` and `x\<noteq>0` by auto
+      thus "norm (f b) / norm b * norm x \<le> norm (f x)" using b[THEN bspec[where x="(norm a / norm x) *\<^sub>R x"]]
+	unfolding f.scaleR and ba using `x\<noteq>0` `a\<noteq>0`
+	by (auto simp add: real_mult_commute pos_le_divide_eq pos_divide_le_eq)
+    qed }
+  ultimately
+  show ?thesis by auto
+qed
+
+lemma closed_injective_image_subspace:
+  fixes f :: "real ^ _ \<Rightarrow> real ^ _"
+  assumes "subspace s" "bounded_linear f" "\<forall>x\<in>s. f x = 0 --> x = 0" "closed s"
+  shows "closed(f ` s)"
+proof-
+  obtain e where "e>0" and e:"\<forall>x\<in>s. e * norm x \<le> norm (f x)" using injective_imp_isometric[OF assms(4,1,2,3)] by auto
+  show ?thesis using complete_isometric_image[OF `e>0` assms(1,2) e] and assms(4)
+    unfolding complete_eq_closed[THEN sym] by auto
+qed
+
+subsection{* Some properties of a canonical subspace.                                  *}
+
+lemma subspace_substandard:
+ "subspace {x::real^'n. (\<forall>i. P i \<longrightarrow> x$i = 0)}"
+  unfolding subspace_def by(auto simp add: vector_add_component vector_smult_component elim!: ballE)
+
+lemma closed_substandard:
+ "closed {x::real^'n::finite. \<forall>i. P i --> x$i = 0}" (is "closed ?A")
+proof-
+  let ?D = "{i. P i}"
+  let ?Bs = "{{x::real^'n. inner (basis i) x = 0}| i. i \<in> ?D}"
+  { fix x
+    { assume "x\<in>?A"
+      hence x:"\<forall>i\<in>?D. x $ i = 0" by auto
+      hence "x\<in> \<Inter> ?Bs" by(auto simp add: inner_basis x) }
+    moreover
+    { assume x:"x\<in>\<Inter>?Bs"
+      { fix i assume i:"i \<in> ?D"
+	then obtain B where BB:"B \<in> ?Bs" and B:"B = {x::real^'n. inner (basis i) x = 0}" by auto
+	hence "x $ i = 0" unfolding B using x unfolding inner_basis by auto  }
+      hence "x\<in>?A" by auto }
+    ultimately have "x\<in>?A \<longleftrightarrow> x\<in> \<Inter>?Bs" by auto }
+  hence "?A = \<Inter> ?Bs" by auto
+  thus ?thesis by(auto simp add: closed_Inter closed_hyperplane)
+qed
+
+lemma dim_substandard:
+  shows "dim {x::real^'n::finite. \<forall>i. i \<notin> d \<longrightarrow> x$i = 0} = card d" (is "dim ?A = _")
+proof-
+  let ?D = "UNIV::'n set"
+  let ?B = "(basis::'n\<Rightarrow>real^'n) ` d"
+
+    let ?bas = "basis::'n \<Rightarrow> real^'n"
+
+  have "?B \<subseteq> ?A" by auto
+
+  moreover
+  { fix x::"real^'n" assume "x\<in>?A"
+    with finite[of d]
+    have "x\<in> span ?B"
+    proof(induct d arbitrary: x)
+      case empty hence "x=0" unfolding Cart_eq by auto
+      thus ?case using subspace_0[OF subspace_span[of "{}"]] by auto
+    next
+      case (insert k F)
+      hence *:"\<forall>i. i \<notin> insert k F \<longrightarrow> x $ i = 0" by auto
+      have **:"F \<subseteq> insert k F" by auto
+      def y \<equiv> "x - x$k *\<^sub>R basis k"
+      have y:"x = y + (x$k) *\<^sub>R basis k" unfolding y_def by auto
+      { fix i assume i':"i \<notin> F"
+	hence "y $ i = 0" unfolding y_def unfolding vector_minus_component
+	  and vector_smult_component and basis_component
+	  using *[THEN spec[where x=i]] by auto }
+      hence "y \<in> span (basis ` (insert k F))" using insert(3)
+	using span_mono[of "?bas ` F" "?bas ` (insert k F)"]
+	using image_mono[OF **, of basis] by auto
+      moreover
+      have "basis k \<in> span (?bas ` (insert k F))" by(rule span_superset, auto)
+      hence "x$k *\<^sub>R basis k \<in> span (?bas ` (insert k F))"
+        using span_mul [where 'a=real, unfolded smult_conv_scaleR] by auto
+      ultimately
+      have "y + x$k *\<^sub>R basis k \<in> span (?bas ` (insert k F))"
+	using span_add by auto
+      thus ?case using y by auto
+    qed
+  }
+  hence "?A \<subseteq> span ?B" by auto
+
+  moreover
+  { fix x assume "x \<in> ?B"
+    hence "x\<in>{(basis i)::real^'n |i. i \<in> ?D}" using assms by auto  }
+  hence "independent ?B" using independent_mono[OF independent_stdbasis, of ?B] and assms by auto
+
+  moreover
+  have "d \<subseteq> ?D" unfolding subset_eq using assms by auto
+  hence *:"inj_on (basis::'n\<Rightarrow>real^'n) d" using subset_inj_on[OF basis_inj, of "d"] by auto
+  have "?B hassize (card d)" unfolding hassize_def and card_image[OF *] by auto
+
+  ultimately show ?thesis using dim_unique[of "basis ` d" ?A] by auto
+qed
+
+text{* Hence closure and completeness of all subspaces.                          *}
+
+lemma closed_subspace_lemma: "n \<le> card (UNIV::'n::finite set) \<Longrightarrow> \<exists>A::'n set. card A = n"
+apply (induct n)
+apply (rule_tac x="{}" in exI, simp)
+apply clarsimp
+apply (subgoal_tac "\<exists>x. x \<notin> A")
+apply (erule exE)
+apply (rule_tac x="insert x A" in exI, simp)
+apply (subgoal_tac "A \<noteq> UNIV", auto)
+done
+
+lemma closed_subspace: fixes s::"(real^'n::finite) set"
+  assumes "subspace s" shows "closed s"
+proof-
+  have "dim s \<le> card (UNIV :: 'n set)" using dim_subset_univ by auto
+  then obtain d::"'n set" where t: "card d = dim s"
+    using closed_subspace_lemma by auto
+  let ?t = "{x::real^'n. \<forall>i. i \<notin> d \<longrightarrow> x$i = 0}"
+  obtain f where f:"bounded_linear f"  "f ` ?t = s" "inj_on f ?t"
+    using subspace_isomorphism[unfolded linear_conv_bounded_linear, OF subspace_substandard[of "\<lambda>i. i \<notin> d"] assms]
+    using dim_substandard[of d] and t by auto
+  interpret f: bounded_linear f by fact
+  have "\<forall>x\<in>?t. f x = 0 \<longrightarrow> x = 0" using f.zero using f(3)[unfolded inj_on_def]
+    by(erule_tac x=0 in ballE) auto
+  moreover have "closed ?t" using closed_substandard .
+  moreover have "subspace ?t" using subspace_substandard .
+  ultimately show ?thesis using closed_injective_image_subspace[of ?t f]
+    unfolding f(2) using f(1) by auto
+qed
+
+lemma complete_subspace:
+  fixes s :: "(real ^ _) set" shows "subspace s ==> complete s"
+  using complete_eq_closed closed_subspace
+  by auto
+
+lemma dim_closure:
+  fixes s :: "(real ^ _) set"
+  shows "dim(closure s) = dim s" (is "?dc = ?d")
+proof-
+  have "?dc \<le> ?d" using closure_minimal[OF span_inc, of s]
+    using closed_subspace[OF subspace_span, of s]
+    using dim_subset[of "closure s" "span s"] unfolding dim_span by auto
+  thus ?thesis using dim_subset[OF closure_subset, of s] by auto
+qed
+
+text{* Affine transformations of intervals.                                      *}
+
+lemma affinity_inverses:
+  assumes m0: "m \<noteq> (0::'a::field)"
+  shows "(\<lambda>x. m *s x + c) o (\<lambda>x. inverse(m) *s x + (-(inverse(m) *s c))) = id"
+  "(\<lambda>x. inverse(m) *s x + (-(inverse(m) *s c))) o (\<lambda>x. m *s x + c) = id"
+  using m0
+apply (auto simp add: expand_fun_eq vector_add_ldistrib vector_smult_assoc)
+by (simp add: vector_smult_lneg[symmetric] vector_smult_assoc vector_sneg_minus1[symmetric])
+
+lemma real_affinity_le:
+ "0 < (m::'a::ordered_field) ==> (m * x + c \<le> y \<longleftrightarrow> x \<le> inverse(m) * y + -(c / m))"
+  by (simp add: field_simps inverse_eq_divide)
+
+lemma real_le_affinity:
+ "0 < (m::'a::ordered_field) ==> (y \<le> m * x + c \<longleftrightarrow> inverse(m) * y + -(c / m) \<le> x)"
+  by (simp add: field_simps inverse_eq_divide)
+
+lemma real_affinity_lt:
+ "0 < (m::'a::ordered_field) ==> (m * x + c < y \<longleftrightarrow> x < inverse(m) * y + -(c / m))"
+  by (simp add: field_simps inverse_eq_divide)
+
+lemma real_lt_affinity:
+ "0 < (m::'a::ordered_field) ==> (y < m * x + c \<longleftrightarrow> inverse(m) * y + -(c / m) < x)"
+  by (simp add: field_simps inverse_eq_divide)
+
+lemma real_affinity_eq:
+ "(m::'a::ordered_field) \<noteq> 0 ==> (m * x + c = y \<longleftrightarrow> x = inverse(m) * y + -(c / m))"
+  by (simp add: field_simps inverse_eq_divide)
+
+lemma real_eq_affinity:
+ "(m::'a::ordered_field) \<noteq> 0 ==> (y = m * x + c  \<longleftrightarrow> inverse(m) * y + -(c / m) = x)"
+  by (simp add: field_simps inverse_eq_divide)
+
+lemma vector_affinity_eq:
+  assumes m0: "(m::'a::field) \<noteq> 0"
+  shows "m *s x + c = y \<longleftrightarrow> x = inverse m *s y + -(inverse m *s c)"
+proof
+  assume h: "m *s x + c = y"
+  hence "m *s x = y - c" by (simp add: ring_simps)
+  hence "inverse m *s (m *s x) = inverse m *s (y - c)" by simp
+  then show "x = inverse m *s y + - (inverse m *s c)"
+    using m0 by (simp add: vector_smult_assoc vector_ssub_ldistrib)
+next
+  assume h: "x = inverse m *s y + - (inverse m *s c)"
+  show "m *s x + c = y" unfolding h diff_minus[symmetric]
+    using m0 by (simp add: vector_smult_assoc vector_ssub_ldistrib)
+qed
+
+lemma vector_eq_affinity:
+ "(m::'a::field) \<noteq> 0 ==> (y = m *s x + c \<longleftrightarrow> inverse(m) *s y + -(inverse(m) *s c) = x)"
+  using vector_affinity_eq[where m=m and x=x and y=y and c=c]
+  by metis
+
+lemma image_affinity_interval: fixes m::real
+  fixes a b c :: "real^'n::finite"
+  shows "(\<lambda>x. m *\<^sub>R x + c) ` {a .. b} =
+            (if {a .. b} = {} then {}
+            else (if 0 \<le> m then {m *\<^sub>R a + c .. m *\<^sub>R b + c}
+            else {m *\<^sub>R b + c .. m *\<^sub>R a + c}))"
+proof(cases "m=0")
+  { fix x assume "x \<le> c" "c \<le> x"
+    hence "x=c" unfolding vector_less_eq_def and Cart_eq by (auto intro: order_antisym) }
+  moreover case True
+  moreover have "c \<in> {m *\<^sub>R a + c..m *\<^sub>R b + c}" unfolding True by(auto simp add: vector_less_eq_def)
+  ultimately show ?thesis by auto
+next
+  case False
+  { fix y assume "a \<le> y" "y \<le> b" "m > 0"
+    hence "m *\<^sub>R a + c \<le> m *\<^sub>R y + c"  "m *\<^sub>R y + c \<le> m *\<^sub>R b + c"
+      unfolding vector_less_eq_def by(auto simp add: vector_smult_component vector_add_component)
+  } moreover
+  { fix y assume "a \<le> y" "y \<le> b" "m < 0"
+    hence "m *\<^sub>R b + c \<le> m *\<^sub>R y + c"  "m *\<^sub>R y + c \<le> m *\<^sub>R a + c"
+      unfolding vector_less_eq_def by(auto simp add: vector_smult_component vector_add_component mult_left_mono_neg elim!:ballE)
+  } moreover
+  { fix y assume "m > 0"  "m *\<^sub>R a + c \<le> y"  "y \<le> m *\<^sub>R b + c"
+    hence "y \<in> (\<lambda>x. m *\<^sub>R x + c) ` {a..b}"
+      unfolding image_iff Bex_def mem_interval vector_less_eq_def
+      apply(auto simp add: vector_smult_component vector_add_component vector_minus_component vector_smult_assoc pth_3[symmetric]
+	intro!: exI[where x="(1 / m) *\<^sub>R (y - c)"])
+      by(auto simp add: pos_le_divide_eq pos_divide_le_eq real_mult_commute diff_le_iff)
+  } moreover
+  { fix y assume "m *\<^sub>R b + c \<le> y" "y \<le> m *\<^sub>R a + c" "m < 0"
+    hence "y \<in> (\<lambda>x. m *\<^sub>R x + c) ` {a..b}"
+      unfolding image_iff Bex_def mem_interval vector_less_eq_def
+      apply(auto simp add: vector_smult_component vector_add_component vector_minus_component vector_smult_assoc pth_3[symmetric]
+	intro!: exI[where x="(1 / m) *\<^sub>R (y - c)"])
+      by(auto simp add: neg_le_divide_eq neg_divide_le_eq real_mult_commute diff_le_iff)
+  }
+  ultimately show ?thesis using False by auto
+qed
+
+lemma image_smult_interval:"(\<lambda>x. m *\<^sub>R (x::real^'n::finite)) ` {a..b} =
+  (if {a..b} = {} then {} else if 0 \<le> m then {m *\<^sub>R a..m *\<^sub>R b} else {m *\<^sub>R b..m *\<^sub>R 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 \<noteq> {}" and c:"0 \<le> c" "c < 1" and f:"(f ` s) \<subseteq> s" and
+          lipschitz:"\<forall>x\<in>s. \<forall>y\<in>s. dist (f x) (f y) \<le> c * dist x y"
+  shows "\<exists>! x\<in>s. (f x = x)"
+proof-
+  have "1 - c > 0" using c by auto
+
+  from s(2) obtain z0 where "z0 \<in> s" by auto
+  def z \<equiv> "\<lambda>n. (f ^^ n) z0"
+  { fix n::nat
+    have "z n \<in> s" unfolding z_def
+    proof(induct n) case 0 thus ?case using `z0 \<in>s` by auto
+    next case Suc thus ?case using f by auto qed }
+  note z_in_s = this
+
+  def d \<equiv> "dist (z 0) (z 1)"
+
+  have fzn:"\<And>n. f (z n) = z (Suc n)" unfolding z_def by auto
+  { fix n::nat
+    have "dist (z n) (z (Suc n)) \<le> (c ^ n) * d"
+    proof(induct n)
+      case 0 thus ?case unfolding d_def by auto
+    next
+      case (Suc m)
+      hence "c * dist (z m) (z (Suc m)) \<le> c ^ Suc m * d"
+	using `0 \<le> c` using mult_mono1_class.mult_mono1[of "dist (z m) (z (Suc m))" "c ^ m * d" c] by auto
+      thus ?case using lipschitz[THEN bspec[where x="z m"], OF z_in_s, THEN bspec[where x="z (Suc m)"], OF z_in_s]
+	unfolding fzn and mult_le_cancel_left by auto
+    qed
+  } note cf_z = this
+
+  { fix n m::nat
+    have "(1 - c) * dist (z m) (z (m+n)) \<le> (c ^ m) * d * (1 - c ^ n)"
+    proof(induct n)
+      case 0 show ?case by auto
+    next
+      case (Suc k)
+      have "(1 - c) * dist (z m) (z (m + Suc k)) \<le> (1 - c) * (dist (z m) (z (m + k)) + dist (z (m + k)) (z (Suc (m + k))))"
+	using dist_triangle and c by(auto simp add: dist_triangle)
+      also have "\<dots> \<le> (1 - c) * (dist (z m) (z (m + k)) + c ^ (m + k) * d)"
+	using cf_z[of "m + k"] and c by auto
+      also have "\<dots> \<le> c ^ m * d * (1 - c ^ k) + (1 - c) * c ^ (m + k) * d"
+	using Suc by (auto simp add: ring_simps)
+      also have "\<dots> = (c ^ m) * (d * (1 - c ^ k) + (1 - c) * c ^ k * d)"
+	unfolding power_add by (auto simp add: ring_simps)
+      also have "\<dots> \<le> (c ^ m) * d * (1 - c ^ Suc k)"
+	using c by (auto simp add: ring_simps)
+      finally show ?case by auto
+    qed
+  } note cf_z2 = this
+  { fix e::real assume "e>0"
+    hence "\<exists>N. \<forall>m n. N \<le> m \<and> N \<le> n \<longrightarrow> dist (z m) (z n) < e"
+    proof(cases "d = 0")
+      case True
+      hence "\<And>n. z n = z0" using cf_z2[of 0] and c unfolding z_def by (auto simp add: pos_prod_le[OF `1 - c > 0`])
+      thus ?thesis using `e>0` by auto
+    next
+      case False hence "d>0" unfolding d_def using zero_le_dist[of "z 0" "z 1"]
+	by (metis False d_def real_less_def)
+      hence "0 < e * (1 - c) / d" using `e>0` and `1-c>0`
+	using divide_pos_pos[of "e * (1 - c)" d] and mult_pos_pos[of e "1 - c"] by auto
+      then obtain N where N:"c ^ N < e * (1 - c) / d" using real_arch_pow_inv[of "e * (1 - c) / d" c] and c by auto
+      { fix m n::nat assume "m>n" and as:"m\<ge>N" "n\<ge>N"
+	have *:"c ^ n \<le> c ^ N" using `n\<ge>N` and c using power_decreasing[OF `n\<ge>N`, of c] by auto
+	have "1 - c ^ (m - n) > 0" using c and power_strict_mono[of c 1 "m - n"] using `m>n` by auto
+	hence **:"d * (1 - c ^ (m - n)) / (1 - c) > 0"
+	  using real_mult_order[OF `d>0`, of "1 - c ^ (m - n)"]
+	  using divide_pos_pos[of "d * (1 - c ^ (m - n))" "1 - c"]
+	  using `0 < 1 - c` by auto
+
+	have "dist (z m) (z n) \<le> c ^ n * d * (1 - c ^ (m - n)) / (1 - c)"
+	  using cf_z2[of n "m - n"] and `m>n` unfolding pos_le_divide_eq[OF `1-c>0`]
+	  by (auto simp add: real_mult_commute dist_commute)
+	also have "\<dots> \<le> c ^ N * d * (1 - c ^ (m - n)) / (1 - c)"
+	  using mult_right_mono[OF * order_less_imp_le[OF **]]
+	  unfolding real_mult_assoc by auto
+	also have "\<dots> < (e * (1 - c) / d) * d * (1 - c ^ (m - n)) / (1 - c)"
+	  using mult_strict_right_mono[OF N **] unfolding real_mult_assoc by auto
+	also have "\<dots> = e * (1 - c ^ (m - n))" using c and `d>0` and `1 - c > 0` by auto
+	also have "\<dots> \<le> e" using c and `1 - c ^ (m - n) > 0` and `e>0` using mult_right_le_one_le[of e "1 - c ^ (m - n)"] by auto
+	finally have  "dist (z m) (z n) < e" by auto
+      } note * = this
+      { fix m n::nat assume as:"N\<le>m" "N\<le>n"
+	hence "dist (z n) (z m) < e"
+	proof(cases "n = m")
+	  case True thus ?thesis using `e>0` by auto
+	next
+	  case False thus ?thesis using as and *[of n m] *[of m n] unfolding nat_neq_iff by (auto simp add: dist_commute)
+	qed }
+      thus ?thesis by auto
+    qed
+  }
+  hence "Cauchy z" unfolding cauchy_def by auto
+  then obtain x where "x\<in>s" and x:"(z ---> x) sequentially" using s(1)[unfolded compact_def complete_def, THEN spec[where x=z]] and z_in_s by auto
+
+  def e \<equiv> "dist (f x) x"
+  have "e = 0" proof(rule ccontr)
+    assume "e \<noteq> 0" hence "e>0" unfolding e_def using zero_le_dist[of "f x" x]
+      by (metis dist_eq_0_iff dist_nz e_def)
+    then obtain N where N:"\<forall>n\<ge>N. dist (z n) x < e / 2"
+      using x[unfolded Lim_sequentially, THEN spec[where x="e/2"]] by auto
+    hence N':"dist (z N) x < e / 2" by auto
+
+    have *:"c * dist (z N) x \<le> dist (z N) x" unfolding mult_le_cancel_right2
+      using zero_le_dist[of "z N" x] and c
+      by (metis dist_eq_0_iff dist_nz order_less_asym real_less_def)
+    have "dist (f (z N)) (f x) \<le> c * dist (z N) x" using lipschitz[THEN bspec[where x="z N"], THEN bspec[where x=x]]
+      using z_in_s[of N] `x\<in>s` using c by auto
+    also have "\<dots> < e / 2" using N' and c using * by auto
+    finally show False unfolding fzn
+      using N[THEN spec[where x="Suc N"]] and dist_triangle_half_r[of "z (Suc N)" "f x" e x]
+      unfolding e_def by auto
+  qed
+  hence "f x = x" unfolding e_def by auto
+  moreover
+  { fix y assume "f y = y" "y\<in>s"
+    hence "dist x y \<le> c * dist x y" using lipschitz[THEN bspec[where x=x], THEN bspec[where x=y]]
+      using `x\<in>s` and `f x = x` by auto
+    hence "dist x y = 0" unfolding mult_le_cancel_right1
+      using c and zero_le_dist[of x y] by auto
+    hence "y = x" by auto
+  }
+  ultimately show ?thesis unfolding Bex1_def using `x\<in>s` by blast+
+qed
+
+subsection{* Edelstein fixed point theorem.                                            *}
+
+lemma edelstein_fix:
+  fixes s :: "'a::real_normed_vector set"
+  assumes s:"compact s" "s \<noteq> {}" and gs:"(g ` s) \<subseteq> s"
+      and dist:"\<forall>x\<in>s. \<forall>y\<in>s. x \<noteq> y \<longrightarrow> dist (g x) (g y) < dist x y"
+  shows "\<exists>! x\<in>s. g x = x"
+proof(cases "\<exists>x\<in>s. g x \<noteq> x")
+  obtain x where "x\<in>s" using s(2) by auto
+  case False hence g:"\<forall>x\<in>s. g x = x" by auto
+  { fix y assume "y\<in>s"
+    hence "x = y" using `x\<in>s` and dist[THEN bspec[where x=x], THEN bspec[where x=y]]
+      unfolding g[THEN bspec[where x=x], OF `x\<in>s`]
+      unfolding g[THEN bspec[where x=y], OF `y\<in>s`] by auto  }
+  thus ?thesis unfolding Bex1_def using `x\<in>s` and g by blast+
+next
+  case True
+  then obtain x where [simp]:"x\<in>s" and "g x \<noteq> x" by auto
+  { fix x y assume "x \<in> s" "y \<in> s"
+    hence "dist (g x) (g y) \<le> dist x y"
+      using dist[THEN bspec[where x=x], THEN bspec[where x=y]] by auto } note dist' = this
+  def y \<equiv> "g x"
+  have [simp]:"y\<in>s" unfolding y_def using gs[unfolded image_subset_iff] and `x\<in>s` by blast
+  def f \<equiv> "\<lambda>n. g ^^ n"
+  have [simp]:"\<And>n z. g (f n z) = f (Suc n) z" unfolding f_def by auto
+  have [simp]:"\<And>z. f 0 z = z" unfolding f_def by auto
+  { fix n::nat and z assume "z\<in>s"
+    have "f n z \<in> s" unfolding f_def
+    proof(induct n)
+      case 0 thus ?case using `z\<in>s` by simp
+    next
+      case (Suc n) thus ?case using gs[unfolded image_subset_iff] by auto
+    qed } note fs = this
+  { fix m n ::nat assume "m\<le>n"
+    fix w z assume "w\<in>s" "z\<in>s"
+    have "dist (f n w) (f n z) \<le> dist (f m w) (f m z)" using `m\<le>n`
+    proof(induct n)
+      case 0 thus ?case by auto
+    next
+      case (Suc n)
+      thus ?case proof(cases "m\<le>n")
+	case True thus ?thesis using Suc(1)
+	  using dist'[OF fs fs, OF `w\<in>s` `z\<in>s`, of n n] by auto
+      next
+	case False hence mn:"m = Suc n" using Suc(2) by simp
+	show ?thesis unfolding mn  by auto
+      qed
+    qed } note distf = this
+
+  def h \<equiv> "\<lambda>n. (f n x, f n y)"
+  let ?s2 = "s \<times> s"
+  obtain l r where "l\<in>?s2" and r:"subseq r" and lr:"((h \<circ> r) ---> l) sequentially"
+    using compact_Times [OF s(1) s(1), unfolded compact_def, THEN spec[where x=h]] unfolding  h_def
+    using fs[OF `x\<in>s`] and fs[OF `y\<in>s`] by blast
+  def a \<equiv> "fst l" def b \<equiv> "snd l"
+  have lab:"l = (a, b)" unfolding a_def b_def by simp
+  have [simp]:"a\<in>s" "b\<in>s" unfolding a_def b_def using `l\<in>?s2` by auto
+
+  have lima:"((fst \<circ> (h \<circ> r)) ---> a) sequentially"
+   and limb:"((snd \<circ> (h \<circ> r)) ---> b) sequentially"
+    using lr
+    unfolding o_def a_def b_def by (simp_all add: tendsto_intros)
+
+  { fix n::nat
+    have *:"\<And>fx fy (x::'a) y. dist fx fy \<le> dist x y \<Longrightarrow> \<not> (dist (fx - fy) (a - b) < dist a b - dist x y)" unfolding dist_norm by norm
+    { fix x y :: 'a
+      have "dist (-x) (-y) = dist x y" unfolding dist_norm
+	using norm_minus_cancel[of "x - y"] by (auto simp add: uminus_add_conv_diff) } note ** = this
+
+    { assume as:"dist a b > dist (f n x) (f n y)"
+      then obtain Na Nb where "\<forall>m\<ge>Na. dist (f (r m) x) a < (dist a b - dist (f n x) (f n y)) / 2"
+	and "\<forall>m\<ge>Nb. dist (f (r m) y) b < (dist a b - dist (f n x) (f n y)) / 2"
+	using lima limb unfolding h_def Lim_sequentially by (fastsimp simp del: less_divide_eq_number_of1)
+      hence "dist (f (r (Na + Nb + n)) x - f (r (Na + Nb + n)) y) (a - b) < dist a b - dist (f n x) (f n y)"
+	apply(erule_tac x="Na+Nb+n" in allE)
+	apply(erule_tac x="Na+Nb+n" in allE) apply simp
+	using dist_triangle_add_half[of a "f (r (Na + Nb + n)) x" "dist a b - dist (f n x) (f n y)"
+          "-b"  "- f (r (Na + Nb + n)) y"]
+	unfolding ** unfolding group_simps(12) by (auto simp add: dist_commute)
+      moreover
+      have "dist (f (r (Na + Nb + n)) x - f (r (Na + Nb + n)) y) (a - b) \<ge> dist a b - dist (f n x) (f n y)"
+	using distf[of n "r (Na+Nb+n)", OF _ `x\<in>s` `y\<in>s`]
+	using subseq_bigger[OF r, of "Na+Nb+n"]
+	using *[of "f (r (Na + Nb + n)) x" "f (r (Na + Nb + n)) y" "f n x" "f n y"] by auto
+      ultimately have False by simp
+    }
+    hence "dist a b \<le> dist (f n x) (f n y)" by(rule ccontr)auto }
+  note ab_fn = this
+
+  have [simp]:"a = b" proof(rule ccontr)
+    def e \<equiv> "dist a b - dist (g a) (g b)"
+    assume "a\<noteq>b" hence "e > 0" unfolding e_def using dist by fastsimp
+    hence "\<exists>n. dist (f n x) a < e/2 \<and> dist (f n y) b < e/2"
+      using lima limb unfolding Lim_sequentially
+      apply (auto elim!: allE[where x="e/2"]) apply(rule_tac x="r (max N Na)" in exI) unfolding h_def by fastsimp
+    then obtain n where n:"dist (f n x) a < e/2 \<and> dist (f n y) b < e/2" by auto
+    have "dist (f (Suc n) x) (g a) \<le> dist (f n x) a"
+      using dist[THEN bspec[where x="f n x"], THEN bspec[where x="a"]] and fs by auto
+    moreover have "dist (f (Suc n) y) (g b) \<le> dist (f n y) b"
+      using dist[THEN bspec[where x="f n y"], THEN bspec[where x="b"]] and fs by auto
+    ultimately have "dist (f (Suc n) x) (g a) + dist (f (Suc n) y) (g b) < e" using n by auto
+    thus False unfolding e_def using ab_fn[of "Suc n"] by norm
+  qed
+
+  have [simp]:"\<And>n. f (Suc n) x = f n y" unfolding f_def y_def by(induct_tac n)auto
+  { fix x y assume "x\<in>s" "y\<in>s" moreover
+    fix e::real assume "e>0" ultimately
+    have "dist y x < e \<longrightarrow> dist (g y) (g x) < e" using dist by fastsimp }
+  hence "continuous_on s g" unfolding continuous_on_def by auto
+
+  hence "((snd \<circ> h \<circ> r) ---> g a) sequentially" unfolding continuous_on_sequentially
+    apply (rule allE[where x="\<lambda>n. (fst \<circ> h \<circ> r) n"]) apply (erule ballE[where x=a])
+    using lima unfolding h_def o_def using fs[OF `x\<in>s`] by (auto simp add: y_def)
+  hence "g a = a" using Lim_unique[OF trivial_limit_sequentially limb, of "g a"]
+    unfolding `a=b` and o_assoc by auto
+  moreover
+  { fix x assume "x\<in>s" "g x = x" "x\<noteq>a"
+    hence "False" using dist[THEN bspec[where x=a], THEN bspec[where x=x]]
+      using `g a = a` and `a\<in>s` by auto  }
+  ultimately show "\<exists>!x\<in>s. g x = x" unfolding Bex1_def using `a\<in>s` by blast
+qed
+
+end
--- a/src/HOL/Library/Convex_Euclidean_Space.thy	Fri Oct 23 14:33:07 2009 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,3371 +0,0 @@
-(*  Title:      HOL/Library/Convex_Euclidean_Space.thy
-    Author:     Robert Himmelmann, TU Muenchen
-*)
-
-header {* Convex sets, functions and related things. *}
-
-theory Convex_Euclidean_Space
-imports Topology_Euclidean_Space
-begin
-
-
-(* ------------------------------------------------------------------------- *)
-(* To be moved elsewhere                                                     *)
-(* ------------------------------------------------------------------------- *)
-
-declare vector_add_ldistrib[simp] vector_ssub_ldistrib[simp] vector_smult_assoc[simp] vector_smult_rneg[simp]
-declare vector_sadd_rdistrib[simp] vector_sub_rdistrib[simp]
-declare dot_ladd[simp] dot_radd[simp] dot_lsub[simp] dot_rsub[simp]
-declare dot_lmult[simp] dot_rmult[simp] dot_lneg[simp] dot_rneg[simp]
-declare UNIV_1[simp]
-
-term "(x::real^'n \<Rightarrow> real) 0"
-
-lemma dim1in[intro]:"Suc 0 \<in> {1::nat .. CARD(1)}" by auto
-
-lemmas vector_component_simps = vector_minus_component vector_smult_component vector_add_component vector_less_eq_def Cart_lambda_beta dest_vec1_def basis_component vector_uminus_component
-
-lemmas continuous_intros = continuous_add continuous_vmul continuous_cmul continuous_const continuous_sub continuous_at_id continuous_within_id
-
-lemmas continuous_on_intros = continuous_on_add continuous_on_const continuous_on_id continuous_on_compose continuous_on_cmul continuous_on_neg continuous_on_sub
-  uniformly_continuous_on_add uniformly_continuous_on_const uniformly_continuous_on_id uniformly_continuous_on_compose uniformly_continuous_on_cmul uniformly_continuous_on_neg uniformly_continuous_on_sub
-
-lemma dest_vec1_simps[simp]: fixes a::"real^1"
-  shows "a$1 = 0 \<longleftrightarrow> a = 0" (*"a \<le> 1 \<longleftrightarrow> dest_vec1 a \<le> 1" "0 \<le> a \<longleftrightarrow> 0 \<le> dest_vec1 a"*)
-  "a \<le> b \<longleftrightarrow> dest_vec1 a \<le> dest_vec1 b" "dest_vec1 (1::real^1) = 1"
-  by(auto simp add:vector_component_simps all_1 Cart_eq)
-
-lemma nequals0I:"x\<in>A \<Longrightarrow> A \<noteq> {}" by auto
-
-lemma norm_not_0:"(x::real^'n::finite)\<noteq>0 \<Longrightarrow> norm x \<noteq> 0" by auto
-
-lemma setsum_delta_notmem: assumes "x\<notin>s"
-  shows "setsum (\<lambda>y. if (y = x) then P x else Q y) s = setsum Q s"
-        "setsum (\<lambda>y. if (x = y) then P x else Q y) s = setsum Q s"
-        "setsum (\<lambda>y. if (y = x) then P y else Q y) s = setsum Q s"
-        "setsum (\<lambda>y. if (x = y) then P y else Q y) s = setsum Q s"
-  apply(rule_tac [!] setsum_cong2) using assms by auto
-
-lemma setsum_delta'':
-  fixes s::"'a::real_vector set" assumes "finite s"
-  shows "(\<Sum>x\<in>s. (if y = x then f x else 0) *\<^sub>R x) = (if y\<in>s then (f y) *\<^sub>R y else 0)"
-proof-
-  have *:"\<And>x y. (if y = x then f x else (0::real)) *\<^sub>R x = (if x=y then (f x) *\<^sub>R x else 0)" by auto
-  show ?thesis unfolding * using setsum_delta[OF assms, of y "\<lambda>x. f x *\<^sub>R x"] by auto
-qed
-
-lemma not_disjointI:"x\<in>A \<Longrightarrow> x\<in>B \<Longrightarrow> A \<inter> B \<noteq> {}" by blast
-
-lemma if_smult:"(if P then x else (y::real)) *\<^sub>R v = (if P then x *\<^sub>R v else y *\<^sub>R v)" by auto
-
-lemma mem_interval_1: fixes x :: "real^1" shows
- "(x \<in> {a .. b} \<longleftrightarrow> dest_vec1 a \<le> dest_vec1 x \<and> dest_vec1 x \<le> dest_vec1 b)"
- "(x \<in> {a<..<b} \<longleftrightarrow> dest_vec1 a < dest_vec1 x \<and> dest_vec1 x < dest_vec1 b)"
-by(simp_all add: Cart_eq vector_less_def vector_less_eq_def dest_vec1_def all_1)
-
-lemma image_smult_interval:"(\<lambda>x. m *\<^sub>R (x::real^'n::finite)) ` {a..b} =
-  (if {a..b} = {} then {} else if 0 \<le> m then {m *\<^sub>R a..m *\<^sub>R b} else {m *\<^sub>R b..m *\<^sub>R a})"
-  using image_affinity_interval[of m 0 a b] by auto
-
-lemma dest_vec1_inverval:
-  "dest_vec1 ` {a .. b} = {dest_vec1 a .. dest_vec1 b}"
-  "dest_vec1 ` {a<.. b} = {dest_vec1 a<.. dest_vec1 b}"
-  "dest_vec1 ` {a ..<b} = {dest_vec1 a ..<dest_vec1 b}"
-  "dest_vec1 ` {a<..<b} = {dest_vec1 a<..<dest_vec1 b}"
-  apply(rule_tac [!] equalityI)
-  unfolding subset_eq Ball_def Bex_def mem_interval_1 image_iff
-  apply(rule_tac [!] allI)apply(rule_tac [!] impI)
-  apply(rule_tac[2] x="vec1 x" in exI)apply(rule_tac[4] x="vec1 x" in exI)
-  apply(rule_tac[6] x="vec1 x" in exI)apply(rule_tac[8] x="vec1 x" in exI)
-  by (auto simp add: vector_less_def vector_less_eq_def all_1 dest_vec1_def
-    vec1_dest_vec1[unfolded dest_vec1_def One_nat_def])
-
-lemma dest_vec1_setsum: assumes "finite S"
-  shows " dest_vec1 (setsum f S) = setsum (\<lambda>x. dest_vec1 (f x)) S"
-  using dest_vec1_sum[OF assms] by auto
-
-lemma dist_triangle_eq:
-  fixes x y z :: "real ^ _"
-  shows "dist x z = dist x y + dist y z \<longleftrightarrow> norm (x - y) *\<^sub>R (y - z) = norm (y - z) *\<^sub>R (x - y)"
-proof- have *:"x - y + (y - z) = x - z" by auto
-  show ?thesis unfolding dist_norm norm_triangle_eq[of "x - y" "y - z", unfolded smult_conv_scaleR *]
-    by(auto simp add:norm_minus_commute) qed
-
-lemma norm_eqI:"x = y \<Longrightarrow> norm x = norm y" by auto 
-lemma norm_minus_eqI:"(x::real^'n::finite) = - y \<Longrightarrow> norm x = norm y" by auto
-
-lemma Min_grI: assumes "finite A" "A \<noteq> {}" "\<forall>a\<in>A. x < a" shows "x < Min A"
-  unfolding Min_gr_iff[OF assms(1,2)] using assms(3) by auto
-
-lemma dimindex_ge_1:"CARD(_::finite) \<ge> 1"
-  using one_le_card_finite by auto
-
-lemma real_dimindex_ge_1:"real (CARD('n::finite)) \<ge> 1" 
-  by(metis dimindex_ge_1 linorder_not_less real_eq_of_nat real_le_trans real_of_nat_1 real_of_nat_le_iff) 
-
-lemma real_dimindex_gt_0:"real (CARD('n::finite)) > 0" apply(rule less_le_trans[OF _ real_dimindex_ge_1]) by auto
-
-subsection {* Affine set and affine hull.*}
-
-definition
-  affine :: "'a::real_vector set \<Rightarrow> bool" where
-  "affine s \<longleftrightarrow> (\<forall>x\<in>s. \<forall>y\<in>s. \<forall>u v. u + v = 1 \<longrightarrow> u *\<^sub>R x + v *\<^sub>R y \<in> s)"
-
-lemma affine_alt: "affine s \<longleftrightarrow> (\<forall>x\<in>s. \<forall>y\<in>s. \<forall>u::real. (1 - u) *\<^sub>R x + u *\<^sub>R y \<in> s)"
-proof- have *:"\<And>u v ::real. u + v = 1 \<longleftrightarrow> v = 1 - u" by auto
-  { fix x y assume "x\<in>s" "y\<in>s"
-    hence "(\<forall>u v::real. u + v = 1 \<longrightarrow> u *\<^sub>R x + v *\<^sub>R y \<in> s) \<longleftrightarrow> (\<forall>u::real. (1 - u) *\<^sub>R x + u *\<^sub>R y \<in> s)" apply auto 
-      apply(erule_tac[!] x="1 - u" in allE) unfolding * by auto  }
-  thus ?thesis unfolding affine_def by auto qed
-
-lemma affine_empty[intro]: "affine {}"
-  unfolding affine_def by auto
-
-lemma affine_sing[intro]: "affine {x}"
-  unfolding affine_alt by (auto simp add: scaleR_left_distrib [symmetric])
-
-lemma affine_UNIV[intro]: "affine UNIV"
-  unfolding affine_def by auto
-
-lemma affine_Inter: "(\<forall>s\<in>f. affine s) \<Longrightarrow> affine (\<Inter> f)"
-  unfolding affine_def by auto 
-
-lemma affine_Int: "affine s \<Longrightarrow> affine t \<Longrightarrow> affine (s \<inter> t)"
-  unfolding affine_def by auto
-
-lemma affine_affine_hull: "affine(affine hull s)"
-  unfolding hull_def using affine_Inter[of "{t \<in> affine. s \<subseteq> t}"]
-  unfolding mem_def by auto
-
-lemma affine_hull_eq[simp]: "(affine hull s = s) \<longleftrightarrow> affine s"
-proof-
-  { fix f assume "f \<subseteq> affine"
-    hence "affine (\<Inter>f)" using affine_Inter[of f] unfolding subset_eq mem_def by auto  }
-  thus ?thesis using hull_eq[unfolded mem_def, of affine s] by auto
-qed
-
-lemma setsum_restrict_set'': assumes "finite A"
-  shows "setsum f {x \<in> A. P x} = (\<Sum>x\<in>A. if P x  then f x else 0)"
-  unfolding mem_def[of _ P, symmetric] unfolding setsum_restrict_set'[OF assms] ..
-
-subsection {* Some explicit formulations (from Lars Schewe). *}
-
-lemma affine: fixes V::"'a::real_vector set"
-  shows "affine V \<longleftrightarrow> (\<forall>s u. finite s \<and> s \<noteq> {} \<and> s \<subseteq> V \<and> setsum u s = 1 \<longrightarrow> (setsum (\<lambda>x. (u x) *\<^sub>R x)) s \<in> V)"
-unfolding affine_def apply rule apply(rule, rule, rule) apply(erule conjE)+ 
-defer apply(rule, rule, rule, rule, rule) proof-
-  fix x y u v assume as:"x \<in> V" "y \<in> V" "u + v = (1::real)"
-    "\<forall>s u. finite s \<and> s \<noteq> {} \<and> s \<subseteq> V \<and> setsum u s = 1 \<longrightarrow> (\<Sum>x\<in>s. u x *\<^sub>R x) \<in> V"
-  thus "u *\<^sub>R x + v *\<^sub>R y \<in> V" apply(cases "x=y")
-    using as(4)[THEN spec[where x="{x,y}"], THEN spec[where x="\<lambda>w. if w = x then u else v"]] and as(1-3) 
-    by(auto simp add: scaleR_left_distrib[THEN sym])
-next
-  fix s u assume as:"\<forall>x\<in>V. \<forall>y\<in>V. \<forall>u v. u + v = 1 \<longrightarrow> u *\<^sub>R x + v *\<^sub>R y \<in> V"
-    "finite s" "s \<noteq> {}" "s \<subseteq> V" "setsum u s = (1::real)"
-  def n \<equiv> "card s"
-  have "card s = 0 \<or> card s = 1 \<or> card s = 2 \<or> card s > 2" by auto
-  thus "(\<Sum>x\<in>s. u x *\<^sub>R x) \<in> V" proof(auto simp only: disjE)
-    assume "card s = 2" hence "card s = Suc (Suc 0)" by auto
-    then obtain a b where "s = {a, b}" unfolding card_Suc_eq by auto
-    thus ?thesis using as(1)[THEN bspec[where x=a], THEN bspec[where x=b]] using as(4,5)
-      by(auto simp add: setsum_clauses(2))
-  next assume "card s > 2" thus ?thesis using as and n_def proof(induct n arbitrary: u s)
-      case (Suc n) fix s::"'a set" and u::"'a \<Rightarrow> real"
-      assume IA:"\<And>u s.  \<lbrakk>2 < card s; \<forall>x\<in>V. \<forall>y\<in>V. \<forall>u v. u + v = 1 \<longrightarrow> u *\<^sub>R x + v *\<^sub>R y \<in> V; finite s;
-               s \<noteq> {}; s \<subseteq> V; setsum u s = 1; n \<equiv> card s \<rbrakk> \<Longrightarrow> (\<Sum>x\<in>s. u x *\<^sub>R x) \<in> V" and
-        as:"Suc n \<equiv> card s" "2 < card s" "\<forall>x\<in>V. \<forall>y\<in>V. \<forall>u v. u + v = 1 \<longrightarrow> u *\<^sub>R x + v *\<^sub>R y \<in> V"
-           "finite s" "s \<noteq> {}" "s \<subseteq> V" "setsum u s = 1"
-      have "\<exists>x\<in>s. u x \<noteq> 1" proof(rule_tac ccontr)
-        assume " \<not> (\<exists>x\<in>s. u x \<noteq> 1)" hence "setsum u s = real_of_nat (card s)" unfolding card_eq_setsum by auto
-        thus False using as(7) and `card s > 2` by (metis Numeral1_eq1_nat less_0_number_of less_int_code(15)
-          less_nat_number_of not_less_iff_gr_or_eq of_nat_1 of_nat_eq_iff pos2 rel_simps(4)) qed
-      then obtain x where x:"x\<in>s" "u x \<noteq> 1" by auto
-
-      have c:"card (s - {x}) = card s - 1" apply(rule card_Diff_singleton) using `x\<in>s` as(4) by auto
-      have *:"s = insert x (s - {x})" "finite (s - {x})" using `x\<in>s` and as(4) by auto
-      have **:"setsum u (s - {x}) = 1 - u x"
-        using setsum_clauses(2)[OF *(2), of u x, unfolded *(1)[THEN sym] as(7)] by auto
-      have ***:"inverse (1 - u x) * setsum u (s - {x}) = 1" unfolding ** using `u x \<noteq> 1` by auto
-      have "(\<Sum>xa\<in>s - {x}. (inverse (1 - u x) * u xa) *\<^sub>R xa) \<in> V" proof(cases "card (s - {x}) > 2")
-        case True hence "s - {x} \<noteq> {}" "card (s - {x}) = n" unfolding c and as(1)[symmetric] proof(rule_tac ccontr) 
-          assume "\<not> s - {x} \<noteq> {}" hence "card (s - {x}) = 0" unfolding card_0_eq[OF *(2)] by simp 
-          thus False using True by auto qed auto
-        thus ?thesis apply(rule_tac IA[of "s - {x}" "\<lambda>y. (inverse (1 - u x) * u y)"])
-        unfolding setsum_right_distrib[THEN sym] using as and *** and True by auto
-      next case False hence "card (s - {x}) = Suc (Suc 0)" using as(2) and c by auto
-        then obtain a b where "(s - {x}) = {a, b}" "a\<noteq>b" unfolding card_Suc_eq by auto
-        thus ?thesis using as(3)[THEN bspec[where x=a], THEN bspec[where x=b]]
-          using *** *(2) and `s \<subseteq> V` unfolding setsum_right_distrib by(auto simp add: setsum_clauses(2)) qed
-      thus "(\<Sum>x\<in>s. u x *\<^sub>R x) \<in> V" unfolding scaleR_scaleR[THEN sym] and scaleR_right.setsum [symmetric]
-         apply(subst *) unfolding setsum_clauses(2)[OF *(2)]
-         using as(3)[THEN bspec[where x=x], THEN bspec[where x="(inverse (1 - u x)) *\<^sub>R (\<Sum>xa\<in>s - {x}. u xa *\<^sub>R xa)"], 
-         THEN spec[where x="u x"], THEN spec[where x="1 - u x"]] and rev_subsetD[OF `x\<in>s` `s\<subseteq>V`] and `u x \<noteq> 1` by auto
-    qed auto
-  next assume "card s = 1" then obtain a where "s={a}" by(auto simp add: card_Suc_eq)
-    thus ?thesis using as(4,5) by simp
-  qed(insert `s\<noteq>{}` `finite s`, auto)
-qed
-
-lemma affine_hull_explicit:
-  "affine hull p = {y. \<exists>s u. finite s \<and> s \<noteq> {} \<and> s \<subseteq> p \<and> setsum u s = 1 \<and> setsum (\<lambda>v. (u v) *\<^sub>R v) s = y}"
-  apply(rule hull_unique) apply(subst subset_eq) prefer 3 apply rule unfolding mem_Collect_eq and mem_def[of _ affine]
-  apply (erule exE)+ apply(erule conjE)+ prefer 2 apply rule proof-
-  fix x assume "x\<in>p" thus "\<exists>s u. finite s \<and> s \<noteq> {} \<and> s \<subseteq> p \<and> setsum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = x"
-    apply(rule_tac x="{x}" in exI, rule_tac x="\<lambda>x. 1" in exI) by auto
-next
-  fix t x s u assume as:"p \<subseteq> t" "affine t" "finite s" "s \<noteq> {}" "s \<subseteq> p" "setsum u s = 1" "(\<Sum>v\<in>s. u v *\<^sub>R v) = x" 
-  thus "x \<in> t" using as(2)[unfolded affine, THEN spec[where x=s], THEN spec[where x=u]] by auto
-next
-  show "affine {y. \<exists>s u. finite s \<and> s \<noteq> {} \<and> s \<subseteq> p \<and> setsum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = y}" unfolding affine_def
-    apply(rule,rule,rule,rule,rule) unfolding mem_Collect_eq proof-
-    fix u v ::real assume uv:"u + v = 1"
-    fix x assume "\<exists>s u. finite s \<and> s \<noteq> {} \<and> s \<subseteq> p \<and> setsum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = x"
-    then obtain sx ux where x:"finite sx" "sx \<noteq> {}" "sx \<subseteq> p" "setsum ux sx = 1" "(\<Sum>v\<in>sx. ux v *\<^sub>R v) = x" by auto
-    fix y assume "\<exists>s u. finite s \<and> s \<noteq> {} \<and> s \<subseteq> p \<and> setsum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = y"
-    then obtain sy uy where y:"finite sy" "sy \<noteq> {}" "sy \<subseteq> p" "setsum uy sy = 1" "(\<Sum>v\<in>sy. uy v *\<^sub>R v) = y" by auto
-    have xy:"finite (sx \<union> sy)" using x(1) y(1) by auto
-    have **:"(sx \<union> sy) \<inter> sx = sx" "(sx \<union> sy) \<inter> sy = sy" by auto
-    show "\<exists>s ua. finite s \<and> s \<noteq> {} \<and> s \<subseteq> p \<and> setsum ua s = 1 \<and> (\<Sum>v\<in>s. ua v *\<^sub>R v) = u *\<^sub>R x + v *\<^sub>R y"
-      apply(rule_tac x="sx \<union> sy" in exI)
-      apply(rule_tac x="\<lambda>a. (if a\<in>sx then u * ux a else 0) + (if a\<in>sy then v * uy a else 0)" in exI)
-      unfolding scaleR_left_distrib setsum_addf if_smult scaleR_zero_left  ** setsum_restrict_set[OF xy, THEN sym]
-      unfolding scaleR_scaleR[THEN sym] scaleR_right.setsum [symmetric] and setsum_right_distrib[THEN sym]
-      unfolding x y using x(1-3) y(1-3) uv by simp qed qed
-
-lemma affine_hull_finite:
-  assumes "finite s"
-  shows "affine hull s = {y. \<exists>u. setsum u s = 1 \<and> setsum (\<lambda>v. u v *\<^sub>R v) s = y}"
-  unfolding affine_hull_explicit and expand_set_eq and mem_Collect_eq apply (rule,rule)
-  apply(erule exE)+ apply(erule conjE)+ defer apply(erule exE) apply(erule conjE) proof-
-  fix x u assume "setsum u s = 1" "(\<Sum>v\<in>s. u v *\<^sub>R v) = x"
-  thus "\<exists>sa u. finite sa \<and> \<not> (\<forall>x. (x \<in> sa) = (x \<in> {})) \<and> sa \<subseteq> s \<and> setsum u sa = 1 \<and> (\<Sum>v\<in>sa. u v *\<^sub>R v) = x"
-    apply(rule_tac x=s in exI, rule_tac x=u in exI) using assms by auto
-next
-  fix x t u assume "t \<subseteq> s" hence *:"s \<inter> t = t" by auto
-  assume "finite t" "\<not> (\<forall>x. (x \<in> t) = (x \<in> {}))" "setsum u t = 1" "(\<Sum>v\<in>t. u v *\<^sub>R v) = x"
-  thus "\<exists>u. setsum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = x" apply(rule_tac x="\<lambda>x. if x\<in>t then u x else 0" in exI)
-    unfolding if_smult scaleR_zero_left and setsum_restrict_set[OF assms, THEN sym] and * by auto qed
-
-subsection {* Stepping theorems and hence small special cases. *}
-
-lemma affine_hull_empty[simp]: "affine hull {} = {}"
-  apply(rule hull_unique) unfolding mem_def by auto
-
-lemma affine_hull_finite_step:
-  fixes y :: "'a::real_vector"
-  shows "(\<exists>u. setsum u {} = w \<and> setsum (\<lambda>x. u x *\<^sub>R x) {} = y) \<longleftrightarrow> w = 0 \<and> y = 0" (is ?th1)
-  "finite s \<Longrightarrow> (\<exists>u. setsum u (insert a s) = w \<and> setsum (\<lambda>x. u x *\<^sub>R x) (insert a s) = y) \<longleftrightarrow>
-                (\<exists>v u. setsum u s = w - v \<and> setsum (\<lambda>x. u x *\<^sub>R x) s = y - v *\<^sub>R a)" (is "?as \<Longrightarrow> (?lhs = ?rhs)")
-proof-
-  show ?th1 by simp
-  assume ?as 
-  { assume ?lhs
-    then obtain u where u:"setsum u (insert a s) = w \<and> (\<Sum>x\<in>insert a s. u x *\<^sub>R x) = y" by auto
-    have ?rhs proof(cases "a\<in>s")
-      case True hence *:"insert a s = s" by auto
-      show ?thesis using u[unfolded *] apply(rule_tac x=0 in exI) by auto
-    next
-      case False thus ?thesis apply(rule_tac x="u a" in exI) using u and `?as` by auto 
-    qed  } moreover
-  { assume ?rhs
-    then obtain v u where vu:"setsum u s = w - v"  "(\<Sum>x\<in>s. u x *\<^sub>R x) = y - v *\<^sub>R a" by auto
-    have *:"\<And>x M. (if x = a then v else M) *\<^sub>R x = (if x = a then v *\<^sub>R x else M *\<^sub>R x)" by auto
-    have ?lhs proof(cases "a\<in>s")
-      case True thus ?thesis
-        apply(rule_tac x="\<lambda>x. (if x=a then v else 0) + u x" in exI)
-        unfolding setsum_clauses(2)[OF `?as`]  apply simp
-        unfolding scaleR_left_distrib and setsum_addf 
-        unfolding vu and * and scaleR_zero_left
-        by (auto simp add: setsum_delta[OF `?as`])
-    next
-      case False 
-      hence **:"\<And>x. x \<in> s \<Longrightarrow> u x = (if x = a then v else u x)"
-               "\<And>x. x \<in> s \<Longrightarrow> u x *\<^sub>R x = (if x = a then v *\<^sub>R x else u x *\<^sub>R x)" by auto
-      from False show ?thesis
-        apply(rule_tac x="\<lambda>x. if x=a then v else u x" in exI)
-        unfolding setsum_clauses(2)[OF `?as`] and * using vu
-        using setsum_cong2[of s "\<lambda>x. u x *\<^sub>R x" "\<lambda>x. if x = a then v *\<^sub>R x else u x *\<^sub>R x", OF **(2)]
-        using setsum_cong2[of s u "\<lambda>x. if x = a then v else u x", OF **(1)] by auto  
-    qed }
-  ultimately show "?lhs = ?rhs" by blast
-qed
-
-lemma affine_hull_2:
-  fixes a b :: "'a::real_vector"
-  shows "affine hull {a,b} = {u *\<^sub>R a + v *\<^sub>R b| u v. (u + v = 1)}" (is "?lhs = ?rhs")
-proof-
-  have *:"\<And>x y z. z = x - y \<longleftrightarrow> y + z = (x::real)" 
-         "\<And>x y z. z = x - y \<longleftrightarrow> y + z = (x::'a)" by auto
-  have "?lhs = {y. \<exists>u. setsum u {a, b} = 1 \<and> (\<Sum>v\<in>{a, b}. u v *\<^sub>R v) = y}"
-    using affine_hull_finite[of "{a,b}"] by auto
-  also have "\<dots> = {y. \<exists>v u. u b = 1 - v \<and> u b *\<^sub>R b = y - v *\<^sub>R a}"
-    by(simp add: affine_hull_finite_step(2)[of "{b}" a]) 
-  also have "\<dots> = ?rhs" unfolding * by auto
-  finally show ?thesis by auto
-qed
-
-lemma affine_hull_3:
-  fixes a b c :: "'a::real_vector"
-  shows "affine hull {a,b,c} = { u *\<^sub>R a + v *\<^sub>R b + w *\<^sub>R c| u v w. u + v + w = 1}" (is "?lhs = ?rhs")
-proof-
-  have *:"\<And>x y z. z = x - y \<longleftrightarrow> y + z = (x::real)" 
-         "\<And>x y z. z = x - y \<longleftrightarrow> y + z = (x::'a)" by auto
-  show ?thesis apply(simp add: affine_hull_finite affine_hull_finite_step)
-    unfolding * apply auto
-    apply(rule_tac x=v in exI) apply(rule_tac x=va in exI) apply auto
-    apply(rule_tac x=u in exI) by(auto intro!: exI)
-qed
-
-subsection {* Some relations between affine hull and subspaces. *}
-
-lemma affine_hull_insert_subset_span:
-  fixes a :: "real ^ _"
-  shows "affine hull (insert a s) \<subseteq> {a + v| v . v \<in> span {x - a | x . x \<in> s}}"
-  unfolding subset_eq Ball_def unfolding affine_hull_explicit span_explicit mem_Collect_eq smult_conv_scaleR
-  apply(rule,rule) apply(erule exE)+ apply(erule conjE)+ proof-
-  fix x t u assume as:"finite t" "t \<noteq> {}" "t \<subseteq> insert a s" "setsum u t = 1" "(\<Sum>v\<in>t. u v *\<^sub>R v) = x"
-  have "(\<lambda>x. x - a) ` (t - {a}) \<subseteq> {x - a |x. x \<in> s}" using as(3) by auto
-  thus "\<exists>v. x = a + v \<and> (\<exists>S u. finite S \<and> S \<subseteq> {x - a |x. x \<in> s} \<and> (\<Sum>v\<in>S. u v *\<^sub>R v) = v)"
-    apply(rule_tac x="x - a" in exI)
-    apply (rule conjI, simp)
-    apply(rule_tac x="(\<lambda>x. x - a) ` (t - {a})" in exI)
-    apply(rule_tac x="\<lambda>x. u (x + a)" in exI)
-    apply (rule conjI) using as(1) apply simp
-    apply (erule conjI)
-    using as(1)
-    apply (simp add: setsum_reindex[unfolded inj_on_def] scaleR_right_diff_distrib setsum_subtractf scaleR_left.setsum[THEN sym] setsum_diff1 scaleR_left_diff_distrib)
-    unfolding as by simp qed
-
-lemma affine_hull_insert_span:
-  fixes a :: "real ^ _"
-  assumes "a \<notin> s"
-  shows "affine hull (insert a s) =
-            {a + v | v . v \<in> span {x - a | x.  x \<in> s}}"
-  apply(rule, rule affine_hull_insert_subset_span) unfolding subset_eq Ball_def
-  unfolding affine_hull_explicit and mem_Collect_eq proof(rule,rule,erule exE,erule conjE)
-  fix y v assume "y = a + v" "v \<in> span {x - a |x. x \<in> s}"
-  then obtain t u where obt:"finite t" "t \<subseteq> {x - a |x. x \<in> s}" "a + (\<Sum>v\<in>t. u v *\<^sub>R v) = y" unfolding span_explicit smult_conv_scaleR by auto
-  def f \<equiv> "(\<lambda>x. x + a) ` t"
-  have f:"finite f" "f \<subseteq> s" "(\<Sum>v\<in>f. u (v - a) *\<^sub>R (v - a)) = y - a" unfolding f_def using obt 
-    by(auto simp add: setsum_reindex[unfolded inj_on_def])
-  have *:"f \<inter> {a} = {}" "f \<inter> - {a} = f" using f(2) assms by auto
-  show "\<exists>sa u. finite sa \<and> sa \<noteq> {} \<and> sa \<subseteq> insert a s \<and> setsum u sa = 1 \<and> (\<Sum>v\<in>sa. u v *\<^sub>R v) = y"
-    apply(rule_tac x="insert a f" in exI)
-    apply(rule_tac x="\<lambda>x. if x=a then 1 - setsum (\<lambda>x. u (x - a)) f else u (x - a)" in exI)
-    using assms and f unfolding setsum_clauses(2)[OF f(1)] and if_smult
-    unfolding setsum_cases[OF f(1), of "{a}", unfolded singleton_iff] and *
-    by (auto simp add: setsum_subtractf scaleR_left.setsum algebra_simps) qed
-
-lemma affine_hull_span:
-  fixes a :: "real ^ _"
-  assumes "a \<in> s"
-  shows "affine hull s = {a + v | v. v \<in> span {x - a | x. x \<in> s - {a}}}"
-  using affine_hull_insert_span[of a "s - {a}", unfolded insert_Diff[OF assms]] by auto
-
-subsection {* Convexity. *}
-
-definition
-  convex :: "'a::real_vector set \<Rightarrow> bool" where
-  "convex s \<longleftrightarrow> (\<forall>x\<in>s. \<forall>y\<in>s. \<forall>u\<ge>0. \<forall>v\<ge>0. u + v = 1 \<longrightarrow> u *\<^sub>R x + v *\<^sub>R y \<in> s)"
-
-lemma convex_alt: "convex s \<longleftrightarrow> (\<forall>x\<in>s. \<forall>y\<in>s. \<forall>u. 0 \<le> u \<and> u \<le> 1 \<longrightarrow> ((1 - u) *\<^sub>R x + u *\<^sub>R y) \<in> s)"
-proof- have *:"\<And>u v::real. u + v = 1 \<longleftrightarrow> u = 1 - v" by auto
-  show ?thesis unfolding convex_def apply auto
-    apply(erule_tac x=x in ballE) apply(erule_tac x=y in ballE) apply(erule_tac x="1 - u" in allE)
-    by (auto simp add: *) qed
-
-lemma mem_convex:
-  assumes "convex s" "a \<in> s" "b \<in> s" "0 \<le> u" "u \<le> 1"
-  shows "((1 - u) *\<^sub>R a + u *\<^sub>R b) \<in> s"
-  using assms unfolding convex_alt by auto
-
-lemma convex_empty[intro]: "convex {}"
-  unfolding convex_def by simp
-
-lemma convex_singleton[intro]: "convex {a}"
-  unfolding convex_def by (auto simp add:scaleR_left_distrib[THEN sym])
-
-lemma convex_UNIV[intro]: "convex UNIV"
-  unfolding convex_def by auto
-
-lemma convex_Inter: "(\<forall>s\<in>f. convex s) ==> convex(\<Inter> f)"
-  unfolding convex_def by auto
-
-lemma convex_Int: "convex s \<Longrightarrow> convex t \<Longrightarrow> convex (s \<inter> t)"
-  unfolding convex_def by auto
-
-lemma convex_halfspace_le: "convex {x. inner a x \<le> b}"
-  unfolding convex_def apply auto
-  unfolding inner_add inner_scaleR
-  by (metis real_convex_bound_le)
-
-lemma convex_halfspace_ge: "convex {x. inner a x \<ge> b}"
-proof- have *:"{x. inner a x \<ge> b} = {x. inner (-a) x \<le> -b}" by auto
-  show ?thesis apply(unfold *) using convex_halfspace_le[of "-a" "-b"] by auto qed
-
-lemma convex_hyperplane: "convex {x. inner a x = b}"
-proof-
-  have *:"{x. inner a x = b} = {x. inner a x \<le> b} \<inter> {x. inner a x \<ge> b}" by auto
-  show ?thesis unfolding * apply(rule convex_Int)
-    using convex_halfspace_le convex_halfspace_ge by auto
-qed
-
-lemma convex_halfspace_lt: "convex {x. inner a x < b}"
-  unfolding convex_def
-  by(auto simp add: real_convex_bound_lt inner_add)
-
-lemma convex_halfspace_gt: "convex {x. inner a x > b}"
-   using convex_halfspace_lt[of "-a" "-b"] by auto
-
-lemma convex_positive_orthant: "convex {x::real^'n::finite. (\<forall>i. 0 \<le> x$i)}"
-  unfolding convex_def apply auto apply(erule_tac x=i in allE)+
-  apply(rule add_nonneg_nonneg) by(auto simp add: mult_nonneg_nonneg)
-
-subsection {* Explicit expressions for convexity in terms of arbitrary sums. *}
-
-lemma convex: "convex s \<longleftrightarrow>
-  (\<forall>(k::nat) u x. (\<forall>i. 1\<le>i \<and> i\<le>k \<longrightarrow> 0 \<le> u i \<and> x i \<in>s) \<and> (setsum u {1..k} = 1)
-           \<longrightarrow> setsum (\<lambda>i. u i *\<^sub>R x i) {1..k} \<in> s)"
-  unfolding convex_def apply rule apply(rule allI)+ defer apply(rule ballI)+ apply(rule allI)+ proof(rule,rule,rule,rule)
-  fix x y u v assume as:"\<forall>(k::nat) u x. (\<forall>i. 1 \<le> i \<and> i \<le> k \<longrightarrow> 0 \<le> u i \<and> x i \<in> s) \<and> setsum u {1..k} = 1 \<longrightarrow> (\<Sum>i = 1..k. u i *\<^sub>R x i) \<in> s"
-    "x \<in> s" "y \<in> s" "0 \<le> u" "0 \<le> v" "u + v = (1::real)"
-  show "u *\<^sub>R x + v *\<^sub>R y \<in> s" using as(1)[THEN spec[where x=2], THEN spec[where x="\<lambda>n. if n=1 then u else v"], THEN spec[where x="\<lambda>n. if n=1 then x else y"]] and as(2-)
-    by (auto simp add: setsum_head_Suc) 
-next
-  fix k u x assume as:"\<forall>x\<in>s. \<forall>y\<in>s. \<forall>u\<ge>0. \<forall>v\<ge>0. u + v = 1 \<longrightarrow> u *\<^sub>R x + v *\<^sub>R y \<in> s" 
-  show "(\<forall>i::nat. 1 \<le> i \<and> i \<le> k \<longrightarrow> 0 \<le> u i \<and> x i \<in> s) \<and> setsum u {1..k} = 1 \<longrightarrow> (\<Sum>i = 1..k. u i *\<^sub>R x i) \<in> s" apply(rule,erule conjE) proof(induct k arbitrary: u)
-  case (Suc k) show ?case proof(cases "u (Suc k) = 1")
-    case True hence "(\<Sum>i = Suc 0..k. u i *\<^sub>R x i) = 0" apply(rule_tac setsum_0') apply(rule ccontr) unfolding ball_simps apply(erule bexE) proof-
-      fix i assume i:"i \<in> {Suc 0..k}" "u i *\<^sub>R x i \<noteq> 0"
-      hence ui:"u i \<noteq> 0" by auto
-      hence "setsum (\<lambda>k. if k=i then u i else 0) {1 .. k} \<le> setsum u {1 .. k}" apply(rule_tac setsum_mono) using Suc(2) by auto
-      hence "setsum u {1 .. k} \<ge> u i" using i(1) by(auto simp add: setsum_delta) 
-      hence "setsum u {1 .. k} > 0"  using ui apply(rule_tac less_le_trans[of _ "u i"]) using Suc(2)[THEN spec[where x=i]] and i(1) by auto
-      thus False using Suc(3) unfolding setsum_cl_ivl_Suc and True by simp qed
-    thus ?thesis unfolding setsum_cl_ivl_Suc using True and Suc(2) by auto
-  next
-    have *:"setsum u {1..k} = 1 - u (Suc k)" using Suc(3)[unfolded setsum_cl_ivl_Suc] by auto
-    have **:"u (Suc k) \<le> 1" apply(rule ccontr) unfolding not_le using Suc(3) using setsum_nonneg[of "{1..k}" u] using Suc(2) by auto
-    have ***:"\<And>i k. (u i / (1 - u (Suc k))) *\<^sub>R x i = (inverse (1 - u (Suc k))) *\<^sub>R (u i *\<^sub>R x i)" unfolding real_divide_def by (auto simp add: algebra_simps)
-    case False hence nn:"1 - u (Suc k) \<noteq> 0" by auto
-    have "(\<Sum>i = 1..k. (u i / (1 - u (Suc k))) *\<^sub>R x i) \<in> s" apply(rule Suc(1)) unfolding setsum_divide_distrib[THEN sym] and *
-      apply(rule_tac allI) apply(rule,rule) apply(rule divide_nonneg_pos) using nn Suc(2) ** by auto
-    hence "(1 - u (Suc k)) *\<^sub>R (\<Sum>i = 1..k. (u i / (1 - u (Suc k))) *\<^sub>R x i) + u (Suc k) *\<^sub>R x (Suc k) \<in> s"
-      apply(rule as[THEN bspec, THEN bspec, THEN spec, THEN mp, THEN spec, THEN mp, THEN mp]) using Suc(2)[THEN spec[where x="Suc k"]] and ** by auto
-    thus ?thesis unfolding setsum_cl_ivl_Suc and *** and scaleR_right.setsum [symmetric] using nn by auto qed qed auto qed
-
-
-lemma convex_explicit:
-  fixes s :: "'a::real_vector set"
-  shows "convex s \<longleftrightarrow>
-  (\<forall>t u. finite t \<and> t \<subseteq> s \<and> (\<forall>x\<in>t. 0 \<le> u x) \<and> setsum u t = 1 \<longrightarrow> setsum (\<lambda>x. u x *\<^sub>R x) t \<in> s)"
-  unfolding convex_def apply(rule,rule,rule) apply(subst imp_conjL,rule) defer apply(rule,rule,rule,rule,rule,rule,rule) proof-
-  fix x y u v assume as:"\<forall>t u. finite t \<and> t \<subseteq> s \<and> (\<forall>x\<in>t. 0 \<le> u x) \<and> setsum u t = 1 \<longrightarrow> (\<Sum>x\<in>t. u x *\<^sub>R x) \<in> s" "x \<in> s" "y \<in> s" "0 \<le> u" "0 \<le> v" "u + v = (1::real)"
-  show "u *\<^sub>R x + v *\<^sub>R y \<in> s" proof(cases "x=y")
-    case True show ?thesis unfolding True and scaleR_left_distrib[THEN sym] using as(3,6) by auto next
-    case False thus ?thesis using as(1)[THEN spec[where x="{x,y}"], THEN spec[where x="\<lambda>z. if z=x then u else v"]] and as(2-) by auto qed
-next 
-  fix t u assume asm:"\<forall>x\<in>s. \<forall>y\<in>s. \<forall>u\<ge>0. \<forall>v\<ge>0. u + v = 1 \<longrightarrow> u *\<^sub>R x + v *\<^sub>R y \<in> s" "finite (t::'a set)"
-  (*"finite t" "t \<subseteq> s" "\<forall>x\<in>t. (0::real) \<le> u x" "setsum u t = 1"*)
-  from this(2) have "\<forall>u. t \<subseteq> s \<and> (\<forall>x\<in>t. 0 \<le> u x) \<and> setsum u t = 1 \<longrightarrow> (\<Sum>x\<in>t. u x *\<^sub>R x) \<in> s" apply(induct_tac t rule:finite_induct)
-    prefer 3 apply (rule,rule) apply(erule conjE)+ proof-
-    fix x f u assume ind:"\<forall>u. f \<subseteq> s \<and> (\<forall>x\<in>f. 0 \<le> u x) \<and> setsum u f = 1 \<longrightarrow> (\<Sum>x\<in>f. u x *\<^sub>R x) \<in> s"
-    assume as:"finite f" "x \<notin> f" "insert x f \<subseteq> s" "\<forall>x\<in>insert x f. 0 \<le> u x" "setsum u (insert x f) = (1::real)"
-    show "(\<Sum>x\<in>insert x f. u x *\<^sub>R x) \<in> s" proof(cases "u x = 1")
-      case True hence "setsum (\<lambda>x. u x *\<^sub>R x) f = 0" apply(rule_tac setsum_0') apply(rule ccontr) unfolding ball_simps apply(erule bexE) proof-
-        fix y assume y:"y \<in> f" "u y *\<^sub>R y \<noteq> 0"
-        hence uy:"u y \<noteq> 0" by auto
-        hence "setsum (\<lambda>k. if k=y then u y else 0) f \<le> setsum u f" apply(rule_tac setsum_mono) using as(4) by auto
-        hence "setsum u f \<ge> u y" using y(1) and as(1) by(auto simp add: setsum_delta) 
-        hence "setsum u f > 0" using uy apply(rule_tac less_le_trans[of _ "u y"]) using as(4) and y(1) by auto
-        thus False using as(2,5) unfolding setsum_clauses(2)[OF as(1)] and True by auto qed
-      thus ?thesis unfolding setsum_clauses(2)[OF as(1)] using as(2,3) unfolding True by auto
-    next
-      have *:"setsum u f = setsum u (insert x f) - u x" using as(2) unfolding setsum_clauses(2)[OF as(1)] by auto
-      have **:"u x \<le> 1" apply(rule ccontr) unfolding not_le using as(5)[unfolded setsum_clauses(2)[OF as(1)]] and as(2)
-        using setsum_nonneg[of f u] and as(4) by auto
-      case False hence "inverse (1 - u x) *\<^sub>R (\<Sum>x\<in>f. u x *\<^sub>R x) \<in> s" unfolding scaleR_right.setsum and scaleR_scaleR
-        apply(rule_tac ind[THEN spec, THEN mp]) apply rule defer apply rule apply rule apply(rule mult_nonneg_nonneg)
-        unfolding setsum_right_distrib[THEN sym] and * using as and ** by auto
-      hence "u x *\<^sub>R x + (1 - u x) *\<^sub>R ((inverse (1 - u x)) *\<^sub>R setsum (\<lambda>x. u x *\<^sub>R x) f) \<in>s" 
-        apply(rule_tac asm(1)[THEN bspec, THEN bspec, THEN spec, THEN mp, THEN spec, THEN mp, THEN mp]) using as and ** False by auto 
-      thus ?thesis unfolding setsum_clauses(2)[OF as(1)] using as(2) and False by auto qed
-  qed auto thus "t \<subseteq> s \<and> (\<forall>x\<in>t. 0 \<le> u x) \<and> setsum u t = 1 \<longrightarrow> (\<Sum>x\<in>t. u x *\<^sub>R x) \<in> s" by auto
-qed
-
-lemma convex_finite: assumes "finite s"
-  shows "convex s \<longleftrightarrow> (\<forall>u. (\<forall>x\<in>s. 0 \<le> u x) \<and> setsum u s = 1
-                      \<longrightarrow> setsum (\<lambda>x. u x *\<^sub>R x) s \<in> s)"
-  unfolding convex_explicit apply(rule, rule, rule) defer apply(rule,rule,rule)apply(erule conjE)+ proof-
-  fix t u assume as:"\<forall>u. (\<forall>x\<in>s. 0 \<le> u x) \<and> setsum u s = 1 \<longrightarrow> (\<Sum>x\<in>s. u x *\<^sub>R x) \<in> s" " finite t" "t \<subseteq> s" "\<forall>x\<in>t. 0 \<le> u x" "setsum u t = (1::real)"
-  have *:"s \<inter> t = t" using as(3) by auto
-  show "(\<Sum>x\<in>t. u x *\<^sub>R x) \<in> s" using as(1)[THEN spec[where x="\<lambda>x. if x\<in>t then u x else 0"]]
-    unfolding if_smult and setsum_cases[OF assms] and * using as(2-) by auto
-qed (erule_tac x=s in allE, erule_tac x=u in allE, auto)
-
-subsection {* Cones. *}
-
-definition
-  cone :: "'a::real_vector set \<Rightarrow> bool" where
-  "cone s \<longleftrightarrow> (\<forall>x\<in>s. \<forall>c\<ge>0. (c *\<^sub>R x) \<in> s)"
-
-lemma cone_empty[intro, simp]: "cone {}"
-  unfolding cone_def by auto
-
-lemma cone_univ[intro, simp]: "cone UNIV"
-  unfolding cone_def by auto
-
-lemma cone_Inter[intro]: "(\<forall>s\<in>f. cone s) \<Longrightarrow> cone(\<Inter> f)"
-  unfolding cone_def by auto
-
-subsection {* Conic hull. *}
-
-lemma cone_cone_hull: "cone (cone hull s)"
-  unfolding hull_def using cone_Inter[of "{t \<in> conic. s \<subseteq> t}"] 
-  by (auto simp add: mem_def)
-
-lemma cone_hull_eq: "(cone hull s = s) \<longleftrightarrow> cone s"
-  apply(rule hull_eq[unfolded mem_def])
-  using cone_Inter unfolding subset_eq by (auto simp add: mem_def)
-
-subsection {* Affine dependence and consequential theorems (from Lars Schewe). *}
-
-definition
-  affine_dependent :: "'a::real_vector set \<Rightarrow> bool" where
-  "affine_dependent s \<longleftrightarrow> (\<exists>x\<in>s. x \<in> (affine hull (s - {x})))"
-
-lemma affine_dependent_explicit:
-  "affine_dependent p \<longleftrightarrow>
-    (\<exists>s u. finite s \<and> s \<subseteq> p \<and> setsum u s = 0 \<and>
-    (\<exists>v\<in>s. u v \<noteq> 0) \<and> setsum (\<lambda>v. u v *\<^sub>R v) s = 0)"
-  unfolding affine_dependent_def affine_hull_explicit mem_Collect_eq apply(rule)
-  apply(erule bexE,erule exE,erule exE) apply(erule conjE)+ defer apply(erule exE,erule exE) apply(erule conjE)+ apply(erule bexE)
-proof-
-  fix x s u assume as:"x \<in> p" "finite s" "s \<noteq> {}" "s \<subseteq> p - {x}" "setsum u s = 1" "(\<Sum>v\<in>s. u v *\<^sub>R v) = x"
-  have "x\<notin>s" using as(1,4) by auto
-  show "\<exists>s u. finite s \<and> s \<subseteq> p \<and> setsum u s = 0 \<and> (\<exists>v\<in>s. u v \<noteq> 0) \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = 0"
-    apply(rule_tac x="insert x s" in exI, rule_tac x="\<lambda>v. if v = x then - 1 else u v" in exI)
-    unfolding if_smult and setsum_clauses(2)[OF as(2)] and setsum_delta_notmem[OF `x\<notin>s`] and as using as by auto 
-next
-  fix s u v assume as:"finite s" "s \<subseteq> p" "setsum u s = 0" "(\<Sum>v\<in>s. u v *\<^sub>R v) = 0" "v \<in> s" "u v \<noteq> 0"
-  have "s \<noteq> {v}" using as(3,6) by auto
-  thus "\<exists>x\<in>p. \<exists>s u. finite s \<and> s \<noteq> {} \<and> s \<subseteq> p - {x} \<and> setsum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = x" 
-    apply(rule_tac x=v in bexI, rule_tac x="s - {v}" in exI, rule_tac x="\<lambda>x. - (1 / u v) * u x" in exI)
-    unfolding scaleR_scaleR[THEN sym] and scaleR_right.setsum [symmetric] unfolding setsum_right_distrib[THEN sym] and setsum_diff1[OF as(1)] using as by auto
-qed
-
-lemma affine_dependent_explicit_finite:
-  fixes s :: "'a::real_vector set" assumes "finite s"
-  shows "affine_dependent s \<longleftrightarrow> (\<exists>u. setsum u s = 0 \<and> (\<exists>v\<in>s. u v \<noteq> 0) \<and> setsum (\<lambda>v. u v *\<^sub>R v) s = 0)"
-  (is "?lhs = ?rhs")
-proof
-  have *:"\<And>vt u v. (if vt then u v else 0) *\<^sub>R v = (if vt then (u v) *\<^sub>R v else (0::'a))" by auto
-  assume ?lhs
-  then obtain t u v where "finite t" "t \<subseteq> s" "setsum u t = 0" "v\<in>t" "u v \<noteq> 0"  "(\<Sum>v\<in>t. u v *\<^sub>R v) = 0"
-    unfolding affine_dependent_explicit by auto
-  thus ?rhs apply(rule_tac x="\<lambda>x. if x\<in>t then u x else 0" in exI)
-    apply auto unfolding * and setsum_restrict_set[OF assms, THEN sym]
-    unfolding Int_absorb1[OF `t\<subseteq>s`] by auto
-next
-  assume ?rhs
-  then obtain u v where "setsum u s = 0"  "v\<in>s" "u v \<noteq> 0" "(\<Sum>v\<in>s. u v *\<^sub>R v) = 0" by auto
-  thus ?lhs unfolding affine_dependent_explicit using assms by auto
-qed
-
-subsection {* A general lemma. *}
-
-lemma convex_connected:
-  fixes s :: "'a::real_normed_vector set"
-  assumes "convex s" shows "connected s"
-proof-
-  { fix e1 e2 assume as:"open e1" "open e2" "e1 \<inter> e2 \<inter> s = {}" "s \<subseteq> e1 \<union> e2" 
-    assume "e1 \<inter> s \<noteq> {}" "e2 \<inter> s \<noteq> {}"
-    then obtain x1 x2 where x1:"x1\<in>e1" "x1\<in>s" and x2:"x2\<in>e2" "x2\<in>s" by auto
-    hence n:"norm (x1 - x2) > 0" unfolding zero_less_norm_iff using as(3) by auto
-
-    { fix x e::real assume as:"0 \<le> x" "x \<le> 1" "0 < e"
-      { fix y have *:"(1 - x) *\<^sub>R x1 + x *\<^sub>R x2 - ((1 - y) *\<^sub>R x1 + y *\<^sub>R x2) = (y - x) *\<^sub>R x1 - (y - x) *\<^sub>R x2"
-          by (simp add: algebra_simps)
-        assume "\<bar>y - x\<bar> < e / norm (x1 - x2)"
-        hence "norm ((1 - x) *\<^sub>R x1 + x *\<^sub>R x2 - ((1 - y) *\<^sub>R x1 + y *\<^sub>R x2)) < e"
-          unfolding * and scaleR_right_diff_distrib[THEN sym]
-          unfolding less_divide_eq using n by auto  }
-      hence "\<exists>d>0. \<forall>y. \<bar>y - x\<bar> < d \<longrightarrow> norm ((1 - x) *\<^sub>R x1 + x *\<^sub>R x2 - ((1 - y) *\<^sub>R x1 + y *\<^sub>R x2)) < e"
-        apply(rule_tac x="e / norm (x1 - x2)" in exI) using as
-        apply auto unfolding zero_less_divide_iff using n by simp  }  note * = this
-
-    have "\<exists>x\<ge>0. x \<le> 1 \<and> (1 - x) *\<^sub>R x1 + x *\<^sub>R x2 \<notin> e1 \<and> (1 - x) *\<^sub>R x1 + x *\<^sub>R x2 \<notin> e2"
-      apply(rule connected_real_lemma) apply (simp add: `x1\<in>e1` `x2\<in>e2` dist_commute)+
-      using * apply(simp add: dist_norm)
-      using as(1,2)[unfolded open_dist] apply simp
-      using as(1,2)[unfolded open_dist] apply simp
-      using assms[unfolded convex_alt, THEN bspec[where x=x1], THEN bspec[where x=x2]] using x1 x2
-      using as(3) by auto
-    then obtain x where "x\<ge>0" "x\<le>1" "(1 - x) *\<^sub>R x1 + x *\<^sub>R x2 \<notin> e1"  "(1 - x) *\<^sub>R x1 + x *\<^sub>R x2 \<notin> e2" by auto
-    hence False using as(4) 
-      using assms[unfolded convex_alt, THEN bspec[where x=x1], THEN bspec[where x=x2]]
-      using x1(2) x2(2) by auto  }
-  thus ?thesis unfolding connected_def by auto
-qed
-
-subsection {* One rather trivial consequence. *}
-
-lemma connected_UNIV: "connected (UNIV :: 'a::real_normed_vector set)"
-  by(simp add: convex_connected convex_UNIV)
-
-subsection {* Convex functions into the reals. *}
-
-definition
-  convex_on :: "'a::real_vector set \<Rightarrow> ('a \<Rightarrow> real) \<Rightarrow> bool" where
-  "convex_on s f \<longleftrightarrow>
-  (\<forall>x\<in>s. \<forall>y\<in>s. \<forall>u\<ge>0. \<forall>v\<ge>0. u + v = 1 \<longrightarrow> f (u *\<^sub>R x + v *\<^sub>R y) \<le> u * f x + v * f y)"
-
-lemma convex_on_subset: "convex_on t f \<Longrightarrow> s \<subseteq> t \<Longrightarrow> convex_on s f"
-  unfolding convex_on_def by auto
-
-lemma convex_add:
-  assumes "convex_on s f" "convex_on s g"
-  shows "convex_on s (\<lambda>x. f x + g x)"
-proof-
-  { fix x y assume "x\<in>s" "y\<in>s" moreover
-    fix u v ::real assume "0 \<le> u" "0 \<le> v" "u + v = 1"
-    ultimately have "f (u *\<^sub>R x + v *\<^sub>R y) + g (u *\<^sub>R x + v *\<^sub>R y) \<le> (u * f x + v * f y) + (u * g x + v * g y)"
-      using assms(1)[unfolded convex_on_def, THEN bspec[where x=x], THEN bspec[where x=y], THEN spec[where x=u]]
-      using assms(2)[unfolded convex_on_def, THEN bspec[where x=x], THEN bspec[where x=y], THEN spec[where x=u]]
-      apply - apply(rule add_mono) by auto
-    hence "f (u *\<^sub>R x + v *\<^sub>R y) + g (u *\<^sub>R x + v *\<^sub>R y) \<le> u * (f x + g x) + v * (f y + g y)" by (simp add: ring_simps)  }
-  thus ?thesis unfolding convex_on_def by auto 
-qed
-
-lemma convex_cmul: 
-  assumes "0 \<le> (c::real)" "convex_on s f"
-  shows "convex_on s (\<lambda>x. c * f x)"
-proof-
-  have *:"\<And>u c fx v fy ::real. u * (c * fx) + v * (c * fy) = c * (u * fx + v * fy)" by (simp add: ring_simps)
-  show ?thesis using assms(2) and mult_mono1[OF _ assms(1)] unfolding convex_on_def and * by auto
-qed
-
-lemma convex_lower:
-  assumes "convex_on s f"  "x\<in>s"  "y \<in> s"  "0 \<le> u"  "0 \<le> v"  "u + v = 1"
-  shows "f (u *\<^sub>R x + v *\<^sub>R y) \<le> max (f x) (f y)"
-proof-
-  let ?m = "max (f x) (f y)"
-  have "u * f x + v * f y \<le> u * max (f x) (f y) + v * max (f x) (f y)" apply(rule add_mono) 
-    using assms(4,5) by(auto simp add: mult_mono1)
-  also have "\<dots> = max (f x) (f y)" using assms(6) unfolding distrib[THEN sym] by auto
-  finally show ?thesis using assms(1)[unfolded convex_on_def, THEN bspec[where x=x], THEN bspec[where x=y], THEN spec[where x=u]]
-    using assms(2-6) by auto 
-qed
-
-lemma convex_local_global_minimum:
-  fixes s :: "'a::real_normed_vector set"
-  assumes "0<e" "convex_on s f" "ball x e \<subseteq> s" "\<forall>y\<in>ball x e. f x \<le> f y"
-  shows "\<forall>y\<in>s. f x \<le> f y"
-proof(rule ccontr)
-  have "x\<in>s" using assms(1,3) by auto
-  assume "\<not> (\<forall>y\<in>s. f x \<le> f y)"
-  then obtain y where "y\<in>s" and y:"f x > f y" by auto
-  hence xy:"0 < dist x y" by (auto simp add: dist_nz[THEN sym])
-
-  then obtain u where "0 < u" "u \<le> 1" and u:"u < e / dist x y"
-    using real_lbound_gt_zero[of 1 "e / dist x y"] using xy `e>0` and divide_pos_pos[of e "dist x y"] by auto
-  hence "f ((1-u) *\<^sub>R x + u *\<^sub>R y) \<le> (1-u) * f x + u * f y" using `x\<in>s` `y\<in>s`
-    using assms(2)[unfolded convex_on_def, THEN bspec[where x=x], THEN bspec[where x=y], THEN spec[where x="1-u"]] by auto
-  moreover
-  have *:"x - ((1 - u) *\<^sub>R x + u *\<^sub>R y) = u *\<^sub>R (x - y)" by (simp add: algebra_simps)
-  have "(1 - u) *\<^sub>R x + u *\<^sub>R y \<in> ball x e" unfolding mem_ball dist_norm unfolding * and norm_scaleR and abs_of_pos[OF `0<u`] unfolding dist_norm[THEN sym]
-    using u unfolding pos_less_divide_eq[OF xy] by auto
-  hence "f x \<le> f ((1 - u) *\<^sub>R x + u *\<^sub>R y)" using assms(4) by auto
-  ultimately show False using mult_strict_left_mono[OF y `u>0`] unfolding left_diff_distrib by auto
-qed
-
-lemma convex_distance:
-  fixes s :: "'a::real_normed_vector set"
-  shows "convex_on s (\<lambda>x. dist a x)"
-proof(auto simp add: convex_on_def dist_norm)
-  fix x y assume "x\<in>s" "y\<in>s"
-  fix u v ::real assume "0 \<le> u" "0 \<le> v" "u + v = 1"
-  have "a = u *\<^sub>R a + v *\<^sub>R a" unfolding scaleR_left_distrib[THEN sym] and `u+v=1` by simp
-  hence *:"a - (u *\<^sub>R x + v *\<^sub>R y) = (u *\<^sub>R (a - x)) + (v *\<^sub>R (a - y))"
-    by (auto simp add: algebra_simps)
-  show "norm (a - (u *\<^sub>R x + v *\<^sub>R y)) \<le> u * norm (a - x) + v * norm (a - y)"
-    unfolding * using norm_triangle_ineq[of "u *\<^sub>R (a - x)" "v *\<^sub>R (a - y)"]
-    using `0 \<le> u` `0 \<le> v` by auto
-qed
-
-subsection {* Arithmetic operations on sets preserve convexity. *}
-
-lemma convex_scaling: "convex s \<Longrightarrow> convex ((\<lambda>x. c *\<^sub>R x) ` s)"
-  unfolding convex_def and image_iff apply auto
-  apply (rule_tac x="u *\<^sub>R x+v *\<^sub>R y" in bexI) by (auto simp add: algebra_simps)
-
-lemma convex_negations: "convex s \<Longrightarrow> convex ((\<lambda>x. -x)` s)"
-  unfolding convex_def and image_iff apply auto
-  apply (rule_tac x="u *\<^sub>R x+v *\<^sub>R y" in bexI) by auto
-
-lemma convex_sums:
-  assumes "convex s" "convex t"
-  shows "convex {x + y| x y. x \<in> s \<and> y \<in> t}"
-proof(auto simp add: convex_def image_iff scaleR_right_distrib)
-  fix xa xb ya yb assume xy:"xa\<in>s" "xb\<in>s" "ya\<in>t" "yb\<in>t"
-  fix u v ::real assume uv:"0 \<le> u" "0 \<le> v" "u + v = 1"
-  show "\<exists>x y. u *\<^sub>R xa + u *\<^sub>R ya + (v *\<^sub>R xb + v *\<^sub>R yb) = x + y \<and> x \<in> s \<and> y \<in> t"
-    apply(rule_tac x="u *\<^sub>R xa + v *\<^sub>R xb" in exI) apply(rule_tac x="u *\<^sub>R ya + v *\<^sub>R yb" in exI)
-    using assms(1)[unfolded convex_def, THEN bspec[where x=xa], THEN bspec[where x=xb]]
-    using assms(2)[unfolded convex_def, THEN bspec[where x=ya], THEN bspec[where x=yb]]
-    using uv xy by auto
-qed
-
-lemma convex_differences: 
-  assumes "convex s" "convex t"
-  shows "convex {x - y| x y. x \<in> s \<and> y \<in> t}"
-proof-
-  have "{x - y| x y. x \<in> s \<and> y \<in> t} = {x + y |x y. x \<in> s \<and> y \<in> uminus ` t}" unfolding image_iff apply auto
-    apply(rule_tac x=xa in exI) apply(rule_tac x="-y" in exI) apply simp
-    apply(rule_tac x=xa in exI) apply(rule_tac x=xb in exI) by simp
-  thus ?thesis using convex_sums[OF assms(1)  convex_negations[OF assms(2)]] by auto
-qed
-
-lemma convex_translation: assumes "convex s" shows "convex ((\<lambda>x. a + x) ` s)"
-proof- have "{a + y |y. y \<in> s} = (\<lambda>x. a + x) ` s" by auto
-  thus ?thesis using convex_sums[OF convex_singleton[of a] assms] by auto qed
-
-lemma convex_affinity: assumes "convex s" shows "convex ((\<lambda>x. a + c *\<^sub>R x) ` s)"
-proof- have "(\<lambda>x. a + c *\<^sub>R x) ` s = op + a ` op *\<^sub>R c ` s" by auto
-  thus ?thesis using convex_translation[OF convex_scaling[OF assms], of a c] by auto qed
-
-lemma convex_linear_image:
-  assumes c:"convex s" and l:"bounded_linear f"
-  shows "convex(f ` s)"
-proof(auto simp add: convex_def)
-  interpret f: bounded_linear f by fact
-  fix x y assume xy:"x \<in> s" "y \<in> s"
-  fix u v ::real assume uv:"0 \<le> u" "0 \<le> v" "u + v = 1"
-  show "u *\<^sub>R f x + v *\<^sub>R f y \<in> f ` s" unfolding image_iff
-    apply(rule_tac x="u *\<^sub>R x + v *\<^sub>R y" in bexI)
-    unfolding f.add f.scaleR
-    using c[unfolded convex_def] xy uv by auto
-qed
-
-subsection {* Balls, being convex, are connected. *}
-
-lemma convex_ball:
-  fixes x :: "'a::real_normed_vector"
-  shows "convex (ball x e)" 
-proof(auto simp add: convex_def)
-  fix y z assume yz:"dist x y < e" "dist x z < e"
-  fix u v ::real assume uv:"0 \<le> u" "0 \<le> v" "u + v = 1"
-  have "dist x (u *\<^sub>R y + v *\<^sub>R z) \<le> u * dist x y + v * dist x z" using uv yz
-    using convex_distance[of "ball x e" x, unfolded convex_on_def, THEN bspec[where x=y], THEN bspec[where x=z]] by auto
-  thus "dist x (u *\<^sub>R y + v *\<^sub>R z) < e" using real_convex_bound_lt[OF yz uv] by auto 
-qed
-
-lemma convex_cball:
-  fixes x :: "'a::real_normed_vector"
-  shows "convex(cball x e)"
-proof(auto simp add: convex_def Ball_def mem_cball)
-  fix y z assume yz:"dist x y \<le> e" "dist x z \<le> e"
-  fix u v ::real assume uv:" 0 \<le> u" "0 \<le> v" "u + v = 1"
-  have "dist x (u *\<^sub>R y + v *\<^sub>R z) \<le> u * dist x y + v * dist x z" using uv yz
-    using convex_distance[of "cball x e" x, unfolded convex_on_def, THEN bspec[where x=y], THEN bspec[where x=z]] by auto
-  thus "dist x (u *\<^sub>R y + v *\<^sub>R z) \<le> e" using real_convex_bound_le[OF yz uv] by auto 
-qed
-
-lemma connected_ball:
-  fixes x :: "'a::real_normed_vector"
-  shows "connected (ball x e)"
-  using convex_connected convex_ball by auto
-
-lemma connected_cball:
-  fixes x :: "'a::real_normed_vector"
-  shows "connected(cball x e)"
-  using convex_connected convex_cball by auto
-
-subsection {* Convex hull. *}
-
-lemma convex_convex_hull: "convex(convex hull s)"
-  unfolding hull_def using convex_Inter[of "{t\<in>convex. s\<subseteq>t}"]
-  unfolding mem_def by auto
-
-lemma convex_hull_eq: "(convex hull s = s) \<longleftrightarrow> convex s" apply(rule hull_eq[unfolded mem_def])
-  using convex_Inter[unfolded Ball_def mem_def] by auto
-
-lemma bounded_convex_hull:
-  fixes s :: "'a::real_normed_vector set"
-  assumes "bounded s" shows "bounded(convex hull s)"
-proof- from assms obtain B where B:"\<forall>x\<in>s. norm x \<le> B" unfolding bounded_iff by auto
-  show ?thesis apply(rule bounded_subset[OF bounded_cball, of _ 0 B])
-    unfolding subset_hull[unfolded mem_def, of convex, OF convex_cball]
-    unfolding subset_eq mem_cball dist_norm using B by auto qed
-
-lemma finite_imp_bounded_convex_hull:
-  fixes s :: "'a::real_normed_vector set"
-  shows "finite s \<Longrightarrow> bounded(convex hull s)"
-  using bounded_convex_hull finite_imp_bounded by auto
-
-subsection {* Stepping theorems for convex hulls of finite sets. *}
-
-lemma convex_hull_empty[simp]: "convex hull {} = {}"
-  apply(rule hull_unique) unfolding mem_def by auto
-
-lemma convex_hull_singleton[simp]: "convex hull {a} = {a}"
-  apply(rule hull_unique) unfolding mem_def by auto
-
-lemma convex_hull_insert:
-  fixes s :: "'a::real_vector set"
-  assumes "s \<noteq> {}"
-  shows "convex hull (insert a s) = {x. \<exists>u\<ge>0. \<exists>v\<ge>0. \<exists>b. (u + v = 1) \<and>
-                                    b \<in> (convex hull s) \<and> (x = u *\<^sub>R a + v *\<^sub>R b)}" (is "?xyz = ?hull")
- apply(rule,rule hull_minimal,rule) unfolding mem_def[of _ convex] and insert_iff prefer 3 apply rule proof-
- fix x assume x:"x = a \<or> x \<in> s"
- thus "x\<in>?hull" apply rule unfolding mem_Collect_eq apply(rule_tac x=1 in exI) defer 
-   apply(rule_tac x=0 in exI) using assms hull_subset[of s convex] by auto
-next
-  fix x assume "x\<in>?hull"
-  then obtain u v b where obt:"u\<ge>0" "v\<ge>0" "u + v = 1" "b \<in> convex hull s" "x = u *\<^sub>R a + v *\<^sub>R b" by auto
-  have "a\<in>convex hull insert a s" "b\<in>convex hull insert a s"
-    using hull_mono[of s "insert a s" convex] hull_mono[of "{a}" "insert a s" convex] and obt(4) by auto
-  thus "x\<in> convex hull insert a s" unfolding obt(5) using convex_convex_hull[of "insert a s", unfolded convex_def]
-    apply(erule_tac x=a in ballE) apply(erule_tac x=b in ballE) apply(erule_tac x=u in allE) using obt by auto
-next
-  show "convex ?hull" unfolding convex_def apply(rule,rule,rule,rule,rule,rule,rule) proof-
-    fix x y u v assume as:"(0::real) \<le> u" "0 \<le> v" "u + v = 1" "x\<in>?hull" "y\<in>?hull"
-    from as(4) obtain u1 v1 b1 where obt1:"u1\<ge>0" "v1\<ge>0" "u1 + v1 = 1" "b1 \<in> convex hull s" "x = u1 *\<^sub>R a + v1 *\<^sub>R b1" by auto
-    from as(5) obtain u2 v2 b2 where obt2:"u2\<ge>0" "v2\<ge>0" "u2 + v2 = 1" "b2 \<in> convex hull s" "y = u2 *\<^sub>R a + v2 *\<^sub>R b2" by auto
-    have *:"\<And>(x::'a) s1 s2. x - s1 *\<^sub>R x - s2 *\<^sub>R x = ((1::real) - (s1 + s2)) *\<^sub>R x" by (auto simp add: algebra_simps)
-    have "\<exists>b \<in> convex hull s. u *\<^sub>R x + v *\<^sub>R y = (u * u1) *\<^sub>R a + (v * u2) *\<^sub>R a + (b - (u * u1) *\<^sub>R b - (v * u2) *\<^sub>R b)"
-    proof(cases "u * v1 + v * v2 = 0")
-      have *:"\<And>(x::'a) s1 s2. x - s1 *\<^sub>R x - s2 *\<^sub>R x = ((1::real) - (s1 + s2)) *\<^sub>R x" by (auto simp add: algebra_simps)
-      case True hence **:"u * v1 = 0" "v * v2 = 0" apply- apply(rule_tac [!] ccontr)
-        using mult_nonneg_nonneg[OF `u\<ge>0` `v1\<ge>0`] mult_nonneg_nonneg[OF `v\<ge>0` `v2\<ge>0`] by auto
-      hence "u * u1 + v * u2 = 1" using as(3) obt1(3) obt2(3) by auto
-      thus ?thesis unfolding obt1(5) obt2(5) * using assms hull_subset[of s convex] by(auto simp add: ** scaleR_right_distrib)
-    next
-      have "1 - (u * u1 + v * u2) = (u + v) - (u * u1 + v * u2)" using as(3) obt1(3) obt2(3) by (auto simp add: field_simps)
-      also have "\<dots> = u * (v1 + u1 - u1) + v * (v2 + u2 - u2)" using as(3) obt1(3) obt2(3) by (auto simp add: field_simps) 
-      also have "\<dots> = u * v1 + v * v2" by simp finally have **:"1 - (u * u1 + v * u2) = u * v1 + v * v2" by auto
-      case False have "0 \<le> u * v1 + v * v2" "0 \<le> u * v1" "0 \<le> u * v1 + v * v2" "0 \<le> v * v2" apply -
-        apply(rule add_nonneg_nonneg) prefer 4 apply(rule add_nonneg_nonneg) apply(rule_tac [!] mult_nonneg_nonneg)
-        using as(1,2) obt1(1,2) obt2(1,2) by auto 
-      thus ?thesis unfolding obt1(5) obt2(5) unfolding * and ** using False
-        apply(rule_tac x="((u * v1) / (u * v1 + v * v2)) *\<^sub>R b1 + ((v * v2) / (u * v1 + v * v2)) *\<^sub>R b2" in bexI) defer
-        apply(rule convex_convex_hull[of s, unfolded convex_def, rule_format]) using obt1(4) obt2(4)
-        unfolding add_divide_distrib[THEN sym] and real_0_le_divide_iff
-        by (auto simp add: scaleR_left_distrib scaleR_right_distrib)
-    qed note * = this
-    have u1:"u1 \<le> 1" apply(rule ccontr) unfolding obt1(3)[THEN sym] and not_le using obt1(2) by auto
-    have u2:"u2 \<le> 1" apply(rule ccontr) unfolding obt2(3)[THEN sym] and not_le using obt2(2) by auto
-    have "u1 * u + u2 * v \<le> (max u1 u2) * u + (max u1 u2) * v" apply(rule add_mono)
-      apply(rule_tac [!] mult_right_mono) using as(1,2) obt1(1,2) obt2(1,2) by auto
-    also have "\<dots> \<le> 1" unfolding mult.add_right[THEN sym] and as(3) using u1 u2 by auto
-    finally 
-    show "u *\<^sub>R x + v *\<^sub>R y \<in> ?hull" unfolding mem_Collect_eq apply(rule_tac x="u * u1 + v * u2" in exI)
-      apply(rule conjI) defer apply(rule_tac x="1 - u * u1 - v * u2" in exI) unfolding Bex_def
-      using as(1,2) obt1(1,2) obt2(1,2) * by(auto intro!: mult_nonneg_nonneg add_nonneg_nonneg simp add: algebra_simps)
-  qed
-qed
-
-
-subsection {* Explicit expression for convex hull. *}
-
-lemma convex_hull_indexed:
-  fixes s :: "'a::real_vector set"
-  shows "convex hull s = {y. \<exists>k u x. (\<forall>i\<in>{1::nat .. k}. 0 \<le> u i \<and> x i \<in> s) \<and>
-                            (setsum u {1..k} = 1) \<and>
-                            (setsum (\<lambda>i. u i *\<^sub>R x i) {1..k} = y)}" (is "?xyz = ?hull")
-  apply(rule hull_unique) unfolding mem_def[of _ convex] apply(rule) defer
-  apply(subst convex_def) apply(rule,rule,rule,rule,rule,rule,rule)
-proof-
-  fix x assume "x\<in>s"
-  thus "x \<in> ?hull" unfolding mem_Collect_eq apply(rule_tac x=1 in exI, rule_tac x="\<lambda>x. 1" in exI) by auto
-next
-  fix t assume as:"s \<subseteq> t" "convex t"
-  show "?hull \<subseteq> t" apply(rule) unfolding mem_Collect_eq apply(erule exE | erule conjE)+ proof-
-    fix x k u y assume assm:"\<forall>i\<in>{1::nat..k}. 0 \<le> u i \<and> y i \<in> s" "setsum u {1..k} = 1" "(\<Sum>i = 1..k. u i *\<^sub>R y i) = x"
-    show "x\<in>t" unfolding assm(3)[THEN sym] apply(rule as(2)[unfolded convex, rule_format])
-      using assm(1,2) as(1) by auto qed
-next
-  fix x y u v assume uv:"0\<le>u" "0\<le>v" "u+v=(1::real)" and xy:"x\<in>?hull" "y\<in>?hull"
-  from xy obtain k1 u1 x1 where x:"\<forall>i\<in>{1::nat..k1}. 0\<le>u1 i \<and> x1 i \<in> s" "setsum u1 {Suc 0..k1} = 1" "(\<Sum>i = Suc 0..k1. u1 i *\<^sub>R x1 i) = x" by auto
-  from xy obtain k2 u2 x2 where y:"\<forall>i\<in>{1::nat..k2}. 0\<le>u2 i \<and> x2 i \<in> s" "setsum u2 {Suc 0..k2} = 1" "(\<Sum>i = Suc 0..k2. u2 i *\<^sub>R x2 i) = y" by auto
-  have *:"\<And>P (x1::'a) x2 s1 s2 i.(if P i then s1 else s2) *\<^sub>R (if P i then x1 else x2) = (if P i then s1 *\<^sub>R x1 else s2 *\<^sub>R x2)"
-    "{1..k1 + k2} \<inter> {1..k1} = {1..k1}" "{1..k1 + k2} \<inter> - {1..k1} = (\<lambda>i. i + k1) ` {1..k2}"
-    prefer 3 apply(rule,rule) unfolding image_iff apply(rule_tac x="x - k1" in bexI) by(auto simp add: not_le)
-  have inj:"inj_on (\<lambda>i. i + k1) {1..k2}" unfolding inj_on_def by auto  
-  show "u *\<^sub>R x + v *\<^sub>R y \<in> ?hull" apply(rule)
-    apply(rule_tac x="k1 + k2" in exI, rule_tac x="\<lambda>i. if i \<in> {1..k1} then u * u1 i else v * u2 (i - k1)" in exI)
-    apply(rule_tac x="\<lambda>i. if i \<in> {1..k1} then x1 i else x2 (i - k1)" in exI) apply(rule,rule) defer apply(rule)
-    unfolding * and setsum_cases[OF finite_atLeastAtMost[of 1 "k1 + k2"]] and setsum_reindex[OF inj] and o_def
-    unfolding scaleR_scaleR[THEN sym] scaleR_right.setsum [symmetric] setsum_right_distrib[THEN sym] proof-
-    fix i assume i:"i \<in> {1..k1+k2}"
-    show "0 \<le> (if i \<in> {1..k1} then u * u1 i else v * u2 (i - k1)) \<and> (if i \<in> {1..k1} then x1 i else x2 (i - k1)) \<in> s"
-    proof(cases "i\<in>{1..k1}")
-      case True thus ?thesis using mult_nonneg_nonneg[of u "u1 i"] and uv(1) x(1)[THEN bspec[where x=i]] by auto
-    next def j \<equiv> "i - k1"
-      case False with i have "j \<in> {1..k2}" unfolding j_def by auto
-      thus ?thesis unfolding j_def[symmetric] using False
-        using mult_nonneg_nonneg[of v "u2 j"] and uv(2) y(1)[THEN bspec[where x=j]] by auto qed
-  qed(auto simp add: not_le x(2,3) y(2,3) uv(3))
-qed
-
-lemma convex_hull_finite:
-  fixes s :: "'a::real_vector set"
-  assumes "finite s"
-  shows "convex hull s = {y. \<exists>u. (\<forall>x\<in>s. 0 \<le> u x) \<and>
-         setsum u s = 1 \<and> setsum (\<lambda>x. u x *\<^sub>R x) s = y}" (is "?HULL = ?set")
-proof(rule hull_unique, auto simp add: mem_def[of _ convex] convex_def[of ?set])
-  fix x assume "x\<in>s" thus " \<exists>u. (\<forall>x\<in>s. 0 \<le> u x) \<and> setsum u s = 1 \<and> (\<Sum>x\<in>s. u x *\<^sub>R x) = x" 
-    apply(rule_tac x="\<lambda>y. if x=y then 1 else 0" in exI) apply auto
-    unfolding setsum_delta'[OF assms] and setsum_delta''[OF assms] by auto 
-next
-  fix u v ::real assume uv:"0 \<le> u" "0 \<le> v" "u + v = 1"
-  fix ux assume ux:"\<forall>x\<in>s. 0 \<le> ux x" "setsum ux s = (1::real)"
-  fix uy assume uy:"\<forall>x\<in>s. 0 \<le> uy x" "setsum uy s = (1::real)"
-  { fix x assume "x\<in>s"
-    hence "0 \<le> u * ux x + v * uy x" using ux(1)[THEN bspec[where x=x]] uy(1)[THEN bspec[where x=x]] and uv(1,2)
-      by (auto, metis add_nonneg_nonneg mult_nonneg_nonneg uv(1) uv(2))  }
-  moreover have "(\<Sum>x\<in>s. u * ux x + v * uy x) = 1"
-    unfolding setsum_addf and setsum_right_distrib[THEN sym] and ux(2) uy(2) using uv(3) by auto
-  moreover have "(\<Sum>x\<in>s. (u * ux x + v * uy x) *\<^sub>R x) = u *\<^sub>R (\<Sum>x\<in>s. ux x *\<^sub>R x) + v *\<^sub>R (\<Sum>x\<in>s. uy x *\<^sub>R x)"
-    unfolding scaleR_left_distrib and setsum_addf and scaleR_scaleR[THEN sym] and scaleR_right.setsum [symmetric] by auto
-  ultimately show "\<exists>uc. (\<forall>x\<in>s. 0 \<le> uc x) \<and> setsum uc s = 1 \<and> (\<Sum>x\<in>s. uc x *\<^sub>R x) = u *\<^sub>R (\<Sum>x\<in>s. ux x *\<^sub>R x) + v *\<^sub>R (\<Sum>x\<in>s. uy x *\<^sub>R x)"
-    apply(rule_tac x="\<lambda>x. u * ux x + v * uy x" in exI) by auto 
-next
-  fix t assume t:"s \<subseteq> t" "convex t" 
-  fix u assume u:"\<forall>x\<in>s. 0 \<le> u x" "setsum u s = (1::real)"
-  thus "(\<Sum>x\<in>s. u x *\<^sub>R x) \<in> t" using t(2)[unfolded convex_explicit, THEN spec[where x=s], THEN spec[where x=u]]
-    using assms and t(1) by auto
-qed
-
-subsection {* Another formulation from Lars Schewe. *}
-
-lemma setsum_constant_scaleR:
-  fixes y :: "'a::real_vector"
-  shows "(\<Sum>x\<in>A. y) = of_nat (card A) *\<^sub>R y"
-apply (cases "finite A")
-apply (induct set: finite)
-apply (simp_all add: algebra_simps)
-done
-
-lemma convex_hull_explicit:
-  fixes p :: "'a::real_vector set"
-  shows "convex hull p = {y. \<exists>s u. finite s \<and> s \<subseteq> p \<and>
-             (\<forall>x\<in>s. 0 \<le> u x) \<and> setsum u s = 1 \<and> setsum (\<lambda>v. u v *\<^sub>R v) s = y}" (is "?lhs = ?rhs")
-proof-
-  { fix x assume "x\<in>?lhs"
-    then obtain k u y where obt:"\<forall>i\<in>{1::nat..k}. 0 \<le> u i \<and> y i \<in> p" "setsum u {1..k} = 1" "(\<Sum>i = 1..k. u i *\<^sub>R y i) = x"
-      unfolding convex_hull_indexed by auto
-
-    have fin:"finite {1..k}" by auto
-    have fin':"\<And>v. finite {i \<in> {1..k}. y i = v}" by auto
-    { fix j assume "j\<in>{1..k}"
-      hence "y j \<in> p" "0 \<le> setsum u {i. Suc 0 \<le> i \<and> i \<le> k \<and> y i = y j}"
-        using obt(1)[THEN bspec[where x=j]] and obt(2) apply simp
-        apply(rule setsum_nonneg) using obt(1) by auto } 
-    moreover
-    have "(\<Sum>v\<in>y ` {1..k}. setsum u {i \<in> {1..k}. y i = v}) = 1"  
-      unfolding setsum_image_gen[OF fin, THEN sym] using obt(2) by auto
-    moreover have "(\<Sum>v\<in>y ` {1..k}. setsum u {i \<in> {1..k}. y i = v} *\<^sub>R v) = x"
-      using setsum_image_gen[OF fin, of "\<lambda>i. u i *\<^sub>R y i" y, THEN sym]
-      unfolding scaleR_left.setsum using obt(3) by auto
-    ultimately have "\<exists>s u. finite s \<and> s \<subseteq> p \<and> (\<forall>x\<in>s. 0 \<le> u x) \<and> setsum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = x"
-      apply(rule_tac x="y ` {1..k}" in exI)
-      apply(rule_tac x="\<lambda>v. setsum u {i\<in>{1..k}. y i = v}" in exI) by auto
-    hence "x\<in>?rhs" by auto  }
-  moreover
-  { fix y assume "y\<in>?rhs"
-    then obtain s u where obt:"finite s" "s \<subseteq> p" "\<forall>x\<in>s. 0 \<le> u x" "setsum u s = 1" "(\<Sum>v\<in>s. u v *\<^sub>R v) = y" by auto
-
-    obtain f where f:"inj_on f {1..card s}" "f ` {1..card s} = s" using ex_bij_betw_nat_finite_1[OF obt(1)] unfolding bij_betw_def by auto
-    
-    { fix i::nat assume "i\<in>{1..card s}"
-      hence "f i \<in> s"  apply(subst f(2)[THEN sym]) by auto
-      hence "0 \<le> u (f i)" "f i \<in> p" using obt(2,3) by auto  }
-    moreover have *:"finite {1..card s}" by auto
-    { fix y assume "y\<in>s"
-      then obtain i where "i\<in>{1..card s}" "f i = y" using f using image_iff[of y f "{1..card s}"] by auto
-      hence "{x. Suc 0 \<le> x \<and> x \<le> card s \<and> f x = y} = {i}" apply auto using f(1)[unfolded inj_on_def] apply(erule_tac x=x in ballE) by auto
-      hence "card {x. Suc 0 \<le> x \<and> x \<le> card s \<and> f x = y} = 1" by auto
-      hence "(\<Sum>x\<in>{x \<in> {1..card s}. f x = y}. u (f x)) = u y"
-            "(\<Sum>x\<in>{x \<in> {1..card s}. f x = y}. u (f x) *\<^sub>R f x) = u y *\<^sub>R y"
-        by (auto simp add: setsum_constant_scaleR)   }
-
-    hence "(\<Sum>x = 1..card s. u (f x)) = 1" "(\<Sum>i = 1..card s. u (f i) *\<^sub>R f i) = y"
-      unfolding setsum_image_gen[OF *(1), of "\<lambda>x. u (f x) *\<^sub>R f x" f] and setsum_image_gen[OF *(1), of "\<lambda>x. u (f x)" f] 
-      unfolding f using setsum_cong2[of s "\<lambda>y. (\<Sum>x\<in>{x \<in> {1..card s}. f x = y}. u (f x) *\<^sub>R f x)" "\<lambda>v. u v *\<^sub>R v"]
-      using setsum_cong2 [of s "\<lambda>y. (\<Sum>x\<in>{x \<in> {1..card s}. f x = y}. u (f x))" u] unfolding obt(4,5) by auto
-    
-    ultimately have "\<exists>k u x. (\<forall>i\<in>{1..k}. 0 \<le> u i \<and> x i \<in> p) \<and> setsum u {1..k} = 1 \<and> (\<Sum>i::nat = 1..k. u i *\<^sub>R x i) = y"
-      apply(rule_tac x="card s" in exI) apply(rule_tac x="u \<circ> f" in exI) apply(rule_tac x=f in exI) by fastsimp
-    hence "y \<in> ?lhs" unfolding convex_hull_indexed by auto  }
-  ultimately show ?thesis unfolding expand_set_eq by blast
-qed
-
-subsection {* A stepping theorem for that expansion. *}
-
-lemma convex_hull_finite_step:
-  fixes s :: "'a::real_vector set" assumes "finite s"
-  shows "(\<exists>u. (\<forall>x\<in>insert a s. 0 \<le> u x) \<and> setsum u (insert a s) = w \<and> setsum (\<lambda>x. u x *\<^sub>R x) (insert a s) = y)
-     \<longleftrightarrow> (\<exists>v\<ge>0. \<exists>u. (\<forall>x\<in>s. 0 \<le> u x) \<and> setsum u s = w - v \<and> setsum (\<lambda>x. u x *\<^sub>R x) s = y - v *\<^sub>R a)" (is "?lhs = ?rhs")
-proof(rule, case_tac[!] "a\<in>s")
-  assume "a\<in>s" hence *:"insert a s = s" by auto
-  assume ?lhs thus ?rhs unfolding * apply(rule_tac x=0 in exI) by auto
-next
-  assume ?lhs then obtain u where u:"\<forall>x\<in>insert a s. 0 \<le> u x" "setsum u (insert a s) = w" "(\<Sum>x\<in>insert a s. u x *\<^sub>R x) = y" by auto
-  assume "a\<notin>s" thus ?rhs apply(rule_tac x="u a" in exI) using u(1)[THEN bspec[where x=a]] apply simp
-    apply(rule_tac x=u in exI) using u[unfolded setsum_clauses(2)[OF assms]] and `a\<notin>s` by auto
-next
-  assume "a\<in>s" hence *:"insert a s = s" by auto
-  have fin:"finite (insert a s)" using assms by auto
-  assume ?rhs then obtain v u where uv:"v\<ge>0" "\<forall>x\<in>s. 0 \<le> u x" "setsum u s = w - v" "(\<Sum>x\<in>s. u x *\<^sub>R x) = y - v *\<^sub>R a" by auto
-  show ?lhs apply(rule_tac x="\<lambda>x. (if a = x then v else 0) + u x" in exI) unfolding scaleR_left_distrib and setsum_addf and setsum_delta''[OF fin] and setsum_delta'[OF fin]
-    unfolding setsum_clauses(2)[OF assms] using uv and uv(2)[THEN bspec[where x=a]] and `a\<in>s` by auto
-next
-  assume ?rhs then obtain v u where uv:"v\<ge>0" "\<forall>x\<in>s. 0 \<le> u x" "setsum u s = w - v" "(\<Sum>x\<in>s. u x *\<^sub>R x) = y - v *\<^sub>R a" by auto
-  moreover assume "a\<notin>s" moreover have "(\<Sum>x\<in>s. if a = x then v else u x) = setsum u s" "(\<Sum>x\<in>s. (if a = x then v else u x) *\<^sub>R x) = (\<Sum>x\<in>s. u x *\<^sub>R x)"
-    apply(rule_tac setsum_cong2) defer apply(rule_tac setsum_cong2) using `a\<notin>s` by auto
-  ultimately show ?lhs apply(rule_tac x="\<lambda>x. if a = x then v else u x" in exI)  unfolding setsum_clauses(2)[OF assms] by auto
-qed
-
-subsection {* Hence some special cases. *}
-
-lemma convex_hull_2:
-  "convex hull {a,b} = {u *\<^sub>R a + v *\<^sub>R b | u v. 0 \<le> u \<and> 0 \<le> v \<and> u + v = 1}"
-proof- have *:"\<And>u. (\<forall>x\<in>{a, b}. 0 \<le> u x) \<longleftrightarrow> 0 \<le> u a \<and> 0 \<le> u b" by auto have **:"finite {b}" by auto
-show ?thesis apply(simp add: convex_hull_finite) unfolding convex_hull_finite_step[OF **, of a 1, unfolded * conj_assoc]
-  apply auto apply(rule_tac x=v in exI) apply(rule_tac x="1 - v" in exI) apply simp
-  apply(rule_tac x=u in exI) apply simp apply(rule_tac x="\<lambda>x. v" in exI) by simp qed
-
-lemma convex_hull_2_alt: "convex hull {a,b} = {a + u *\<^sub>R (b - a) | u.  0 \<le> u \<and> u \<le> 1}"
-  unfolding convex_hull_2 unfolding Collect_def 
-proof(rule ext) have *:"\<And>x y ::real. x + y = 1 \<longleftrightarrow> x = 1 - y" by auto
-  fix x show "(\<exists>v u. x = v *\<^sub>R a + u *\<^sub>R b \<and> 0 \<le> v \<and> 0 \<le> u \<and> v + u = 1) = (\<exists>u. x = a + u *\<^sub>R (b - a) \<and> 0 \<le> u \<and> u \<le> 1)"
-    unfolding * apply auto apply(rule_tac[!] x=u in exI) by (auto simp add: algebra_simps) qed
-
-lemma convex_hull_3:
-  "convex hull {a,b,c} = { u *\<^sub>R a + v *\<^sub>R b + w *\<^sub>R c | u v w. 0 \<le> u \<and> 0 \<le> v \<and> 0 \<le> w \<and> u + v + w = 1}"
-proof-
-  have fin:"finite {a,b,c}" "finite {b,c}" "finite {c}" by auto
-  have *:"\<And>x y z ::real. x + y + z = 1 \<longleftrightarrow> x = 1 - y - z"
-         "\<And>x y z ::real^'n. x + y + z = 1 \<longleftrightarrow> x = 1 - y - z" by (auto simp add: ring_simps)
-  show ?thesis unfolding convex_hull_finite[OF fin(1)] and Collect_def and convex_hull_finite_step[OF fin(2)] and *
-    unfolding convex_hull_finite_step[OF fin(3)] apply(rule ext) apply simp apply auto
-    apply(rule_tac x=va in exI) apply (rule_tac x="u c" in exI) apply simp
-    apply(rule_tac x="1 - v - w" in exI) apply simp apply(rule_tac x=v in exI) apply simp apply(rule_tac x="\<lambda>x. w" in exI) by simp qed
-
-lemma convex_hull_3_alt:
-  "convex hull {a,b,c} = {a + u *\<^sub>R (b - a) + v *\<^sub>R (c - a) | u v.  0 \<le> u \<and> 0 \<le> v \<and> u + v \<le> 1}"
-proof- have *:"\<And>x y z ::real. x + y + z = 1 \<longleftrightarrow> x = 1 - y - z" by auto
-  show ?thesis unfolding convex_hull_3 apply (auto simp add: *) apply(rule_tac x=v in exI) apply(rule_tac x=w in exI) apply (simp add: algebra_simps)
-    apply(rule_tac x=u in exI) apply(rule_tac x=v in exI) by (simp add: algebra_simps) qed
-
-subsection {* Relations among closure notions and corresponding hulls. *}
-
-text {* TODO: Generalize linear algebra concepts defined in @{text
-Euclidean_Space.thy} so that we can generalize these lemmas. *}
-
-lemma subspace_imp_affine:
-  fixes s :: "(real ^ _) set" shows "subspace s \<Longrightarrow> affine s"
-  unfolding subspace_def affine_def smult_conv_scaleR by auto
-
-lemma affine_imp_convex: "affine s \<Longrightarrow> convex s"
-  unfolding affine_def convex_def by auto
-
-lemma subspace_imp_convex:
-  fixes s :: "(real ^ _) set" shows "subspace s \<Longrightarrow> convex s"
-  using subspace_imp_affine affine_imp_convex by auto
-
-lemma affine_hull_subset_span:
-  fixes s :: "(real ^ _) set" shows "(affine hull s) \<subseteq> (span s)"
-  unfolding span_def apply(rule hull_antimono) unfolding subset_eq Ball_def mem_def
-  using subspace_imp_affine  by auto
-
-lemma convex_hull_subset_span:
-  fixes s :: "(real ^ _) set" shows "(convex hull s) \<subseteq> (span s)"
-  unfolding span_def apply(rule hull_antimono) unfolding subset_eq Ball_def mem_def
-  using subspace_imp_convex by auto
-
-lemma convex_hull_subset_affine_hull: "(convex hull s) \<subseteq> (affine hull s)"
-  unfolding span_def apply(rule hull_antimono) unfolding subset_eq Ball_def mem_def
-  using affine_imp_convex by auto
-
-lemma affine_dependent_imp_dependent:
-  fixes s :: "(real ^ _) set" shows "affine_dependent s \<Longrightarrow> dependent s"
-  unfolding affine_dependent_def dependent_def 
-  using affine_hull_subset_span by auto
-
-lemma dependent_imp_affine_dependent:
-  fixes s :: "(real ^ _) set"
-  assumes "dependent {x - a| x . x \<in> s}" "a \<notin> s"
-  shows "affine_dependent (insert a s)"
-proof-
-  from assms(1)[unfolded dependent_explicit smult_conv_scaleR] obtain S u v 
-    where obt:"finite S" "S \<subseteq> {x - a |x. x \<in> s}" "v\<in>S" "u v  \<noteq> 0" "(\<Sum>v\<in>S. u v *\<^sub>R v) = 0" by auto
-  def t \<equiv> "(\<lambda>x. x + a) ` S"
-
-  have inj:"inj_on (\<lambda>x. x + a) S" unfolding inj_on_def by auto
-  have "0\<notin>S" using obt(2) assms(2) unfolding subset_eq by auto
-  have fin:"finite t" and  "t\<subseteq>s" unfolding t_def using obt(1,2) by auto 
-
-  hence "finite (insert a t)" and "insert a t \<subseteq> insert a s" by auto 
-  moreover have *:"\<And>P Q. (\<Sum>x\<in>t. (if x = a then P x else Q x)) = (\<Sum>x\<in>t. Q x)"
-    apply(rule setsum_cong2) using `a\<notin>s` `t\<subseteq>s` by auto
-  have "(\<Sum>x\<in>insert a t. if x = a then - (\<Sum>x\<in>t. u (x - a)) else u (x - a)) = 0"
-    unfolding setsum_clauses(2)[OF fin] using `a\<notin>s` `t\<subseteq>s` apply auto unfolding * by auto
-  moreover have "\<exists>v\<in>insert a t. (if v = a then - (\<Sum>x\<in>t. u (x - a)) else u (v - a)) \<noteq> 0"
-    apply(rule_tac x="v + a" in bexI) using obt(3,4) and `0\<notin>S` unfolding t_def by auto
-  moreover have *:"\<And>P Q. (\<Sum>x\<in>t. (if x = a then P x else Q x) *\<^sub>R x) = (\<Sum>x\<in>t. Q x *\<^sub>R x)"
-    apply(rule setsum_cong2) using `a\<notin>s` `t\<subseteq>s` by auto
-  have "(\<Sum>x\<in>t. u (x - a)) *\<^sub>R a = (\<Sum>v\<in>t. u (v - a) *\<^sub>R v)" 
-    unfolding scaleR_left.setsum unfolding t_def and setsum_reindex[OF inj] and o_def
-    using obt(5) by (auto simp add: setsum_addf scaleR_right_distrib)
-  hence "(\<Sum>v\<in>insert a t. (if v = a then - (\<Sum>x\<in>t. u (x - a)) else u (v - a)) *\<^sub>R v) = 0"
-    unfolding setsum_clauses(2)[OF fin] using `a\<notin>s` `t\<subseteq>s` by (auto simp add: *  vector_smult_lneg) 
-  ultimately show ?thesis unfolding affine_dependent_explicit
-    apply(rule_tac x="insert a t" in exI) by auto 
-qed
-
-lemma convex_cone:
-  "convex s \<and> cone s \<longleftrightarrow> (\<forall>x\<in>s. \<forall>y\<in>s. (x + y) \<in> s) \<and> (\<forall>x\<in>s. \<forall>c\<ge>0. (c *\<^sub>R x) \<in> s)" (is "?lhs = ?rhs")
-proof-
-  { fix x y assume "x\<in>s" "y\<in>s" and ?lhs
-    hence "2 *\<^sub>R x \<in>s" "2 *\<^sub>R y \<in> s" unfolding cone_def by auto
-    hence "x + y \<in> s" using `?lhs`[unfolded convex_def, THEN conjunct1]
-      apply(erule_tac x="2*\<^sub>R x" in ballE) apply(erule_tac x="2*\<^sub>R y" in ballE)
-      apply(erule_tac x="1/2" in allE) apply simp apply(erule_tac x="1/2" in allE) by auto  }
-  thus ?thesis unfolding convex_def cone_def by blast
-qed
-
-lemma affine_dependent_biggerset: fixes s::"(real^'n::finite) set"
-  assumes "finite s" "card s \<ge> CARD('n) + 2"
-  shows "affine_dependent s"
-proof-
-  have "s\<noteq>{}" using assms by auto then obtain a where "a\<in>s" by auto
-  have *:"{x - a |x. x \<in> s - {a}} = (\<lambda>x. x - a) ` (s - {a})" by auto
-  have "card {x - a |x. x \<in> s - {a}} = card (s - {a})" unfolding * 
-    apply(rule card_image) unfolding inj_on_def by auto
-  also have "\<dots> > CARD('n)" using assms(2)
-    unfolding card_Diff_singleton[OF assms(1) `a\<in>s`] by auto
-  finally show ?thesis apply(subst insert_Diff[OF `a\<in>s`, THEN sym])
-    apply(rule dependent_imp_affine_dependent)
-    apply(rule dependent_biggerset) by auto qed
-
-lemma affine_dependent_biggerset_general:
-  assumes "finite (s::(real^'n::finite) set)" "card s \<ge> dim s + 2"
-  shows "affine_dependent s"
-proof-
-  from assms(2) have "s \<noteq> {}" by auto
-  then obtain a where "a\<in>s" by auto
-  have *:"{x - a |x. x \<in> s - {a}} = (\<lambda>x. x - a) ` (s - {a})" by auto
-  have **:"card {x - a |x. x \<in> s - {a}} = card (s - {a})" unfolding * 
-    apply(rule card_image) unfolding inj_on_def by auto
-  have "dim {x - a |x. x \<in> s - {a}} \<le> dim s"
-    apply(rule subset_le_dim) unfolding subset_eq
-    using `a\<in>s` by (auto simp add:span_superset span_sub)
-  also have "\<dots> < dim s + 1" by auto
-  also have "\<dots> \<le> card (s - {a})" using assms
-    using card_Diff_singleton[OF assms(1) `a\<in>s`] by auto
-  finally show ?thesis apply(subst insert_Diff[OF `a\<in>s`, THEN sym])
-    apply(rule dependent_imp_affine_dependent) apply(rule dependent_biggerset_general) unfolding ** by auto qed
-
-subsection {* Caratheodory's theorem. *}
-
-lemma convex_hull_caratheodory: fixes p::"(real^'n::finite) set"
-  shows "convex hull p = {y. \<exists>s u. finite s \<and> s \<subseteq> p \<and> card s \<le> CARD('n) + 1 \<and>
-  (\<forall>x\<in>s. 0 \<le> u x) \<and> setsum u s = 1 \<and> setsum (\<lambda>v. u v *\<^sub>R v) s = y}"
-  unfolding convex_hull_explicit expand_set_eq mem_Collect_eq
-proof(rule,rule)
-  fix y let ?P = "\<lambda>n. \<exists>s u. finite s \<and> card s = n \<and> s \<subseteq> p \<and> (\<forall>x\<in>s. 0 \<le> u x) \<and> setsum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = y"
-  assume "\<exists>s u. finite s \<and> s \<subseteq> p \<and> (\<forall>x\<in>s. 0 \<le> u x) \<and> setsum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = y"
-  then obtain N where "?P N" by auto
-  hence "\<exists>n\<le>N. (\<forall>k<n. \<not> ?P k) \<and> ?P n" apply(rule_tac ex_least_nat_le) by auto
-  then obtain n where "?P n" and smallest:"\<forall>k<n. \<not> ?P k" by blast
-  then obtain s u where obt:"finite s" "card s = n" "s\<subseteq>p" "\<forall>x\<in>s. 0 \<le> u x" "setsum u s = 1"  "(\<Sum>v\<in>s. u v *\<^sub>R v) = y" by auto
-
-  have "card s \<le> CARD('n) + 1" proof(rule ccontr, simp only: not_le)
-    assume "CARD('n) + 1 < card s"
-    hence "affine_dependent s" using affine_dependent_biggerset[OF obt(1)] by auto
-    then obtain w v where wv:"setsum w s = 0" "v\<in>s" "w v \<noteq> 0" "(\<Sum>v\<in>s. w v *\<^sub>R v) = 0"
-      using affine_dependent_explicit_finite[OF obt(1)] by auto
-    def i \<equiv> "(\<lambda>v. (u v) / (- w v)) ` {v\<in>s. w v < 0}"  def t \<equiv> "Min i"
-    have "\<exists>x\<in>s. w x < 0" proof(rule ccontr, simp add: not_less)
-      assume as:"\<forall>x\<in>s. 0 \<le> w x"
-      hence "setsum w (s - {v}) \<ge> 0" apply(rule_tac setsum_nonneg) by auto
-      hence "setsum w s > 0" unfolding setsum_diff1'[OF obt(1) `v\<in>s`]
-        using as[THEN bspec[where x=v]] and `v\<in>s` using `w v \<noteq> 0` by auto
-      thus False using wv(1) by auto
-    qed hence "i\<noteq>{}" unfolding i_def by auto
-
-    hence "t \<ge> 0" using Min_ge_iff[of i 0 ] and obt(1) unfolding t_def i_def
-      using obt(4)[unfolded le_less] apply auto unfolding divide_le_0_iff by auto 
-    have t:"\<forall>v\<in>s. u v + t * w v \<ge> 0" proof
-      fix v assume "v\<in>s" hence v:"0\<le>u v" using obt(4)[THEN bspec[where x=v]] by auto
-      show"0 \<le> u v + t * w v" proof(cases "w v < 0")
-        case False thus ?thesis apply(rule_tac add_nonneg_nonneg) 
-          using v apply simp apply(rule mult_nonneg_nonneg) using `t\<ge>0` by auto next
-        case True hence "t \<le> u v / (- w v)" using `v\<in>s`
-          unfolding t_def i_def apply(rule_tac Min_le) using obt(1) by auto 
-        thus ?thesis unfolding real_0_le_add_iff
-          using pos_le_divide_eq[OF True[unfolded neg_0_less_iff_less[THEN sym]]] by auto
-      qed qed
-
-    obtain a where "a\<in>s" and "t = (\<lambda>v. (u v) / (- w v)) a" and "w a < 0"
-      using Min_in[OF _ `i\<noteq>{}`] and obt(1) unfolding i_def t_def by auto
-    hence a:"a\<in>s" "u a + t * w a = 0" by auto
-    have *:"\<And>f. setsum f (s - {a}) = setsum f s - ((f a)::'a::ring)" unfolding setsum_diff1'[OF obt(1) `a\<in>s`] by auto 
-    have "(\<Sum>v\<in>s. u v + t * w v) = 1"
-      unfolding setsum_addf wv(1) setsum_right_distrib[THEN sym] obt(5) by auto
-    moreover have "(\<Sum>v\<in>s. u v *\<^sub>R v + (t * w v) *\<^sub>R v) - (u a *\<^sub>R a + (t * w a) *\<^sub>R a) = y" 
-      unfolding setsum_addf obt(6) scaleR_scaleR[THEN sym] scaleR_right.setsum [symmetric] wv(4)
-      using a(2) [THEN eq_neg_iff_add_eq_0 [THEN iffD2]]
-      by (simp add: vector_smult_lneg)
-    ultimately have "?P (n - 1)" apply(rule_tac x="(s - {a})" in exI)
-      apply(rule_tac x="\<lambda>v. u v + t * w v" in exI) using obt(1-3) and t and a by (auto simp add: * scaleR_left_distrib)
-    thus False using smallest[THEN spec[where x="n - 1"]] by auto qed
-  thus "\<exists>s u. finite s \<and> s \<subseteq> p \<and> card s \<le> CARD('n) + 1
-    \<and> (\<forall>x\<in>s. 0 \<le> u x) \<and> setsum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = y" using obt by auto
-qed auto
-
-lemma caratheodory:
- "convex hull p = {x::real^'n::finite. \<exists>s. finite s \<and> s \<subseteq> p \<and>
-      card s \<le> CARD('n) + 1 \<and> x \<in> convex hull s}"
-  unfolding expand_set_eq apply(rule, rule) unfolding mem_Collect_eq proof-
-  fix x assume "x \<in> convex hull p"
-  then obtain s u where "finite s" "s \<subseteq> p" "card s \<le> CARD('n) + 1"
-     "\<forall>x\<in>s. 0 \<le> u x" "setsum u s = 1" "(\<Sum>v\<in>s. u v *\<^sub>R v) = x"unfolding convex_hull_caratheodory by auto
-  thus "\<exists>s. finite s \<and> s \<subseteq> p \<and> card s \<le> CARD('n) + 1 \<and> x \<in> convex hull s"
-    apply(rule_tac x=s in exI) using hull_subset[of s convex]
-  using convex_convex_hull[unfolded convex_explicit, of s, THEN spec[where x=s], THEN spec[where x=u]] by auto
-next
-  fix x assume "\<exists>s. finite s \<and> s \<subseteq> p \<and> card s \<le> CARD('n) + 1 \<and> x \<in> convex hull s"
-  then obtain s where "finite s" "s \<subseteq> p" "card s \<le> CARD('n) + 1" "x \<in> convex hull s" by auto
-  thus "x \<in> convex hull p" using hull_mono[OF `s\<subseteq>p`] by auto
-qed
-
-subsection {* Openness and compactness are preserved by convex hull operation. *}
-
-lemma open_convex_hull:
-  fixes s :: "'a::real_normed_vector set"
-  assumes "open s"
-  shows "open(convex hull s)"
-  unfolding open_contains_cball convex_hull_explicit unfolding mem_Collect_eq ball_simps(10) 
-proof(rule, rule) fix a
-  assume "\<exists>sa u. finite sa \<and> sa \<subseteq> s \<and> (\<forall>x\<in>sa. 0 \<le> u x) \<and> setsum u sa = 1 \<and> (\<Sum>v\<in>sa. u v *\<^sub>R v) = a"
-  then obtain t u where obt:"finite t" "t\<subseteq>s" "\<forall>x\<in>t. 0 \<le> u x" "setsum u t = 1" "(\<Sum>v\<in>t. u v *\<^sub>R v) = a" by auto
-
-  from assms[unfolded open_contains_cball] obtain b where b:"\<forall>x\<in>s. 0 < b x \<and> cball x (b x) \<subseteq> s"
-    using bchoice[of s "\<lambda>x e. e>0 \<and> cball x e \<subseteq> s"] by auto
-  have "b ` t\<noteq>{}" unfolding i_def using obt by auto  def i \<equiv> "b ` t"
-
-  show "\<exists>e>0. cball a e \<subseteq> {y. \<exists>sa u. finite sa \<and> sa \<subseteq> s \<and> (\<forall>x\<in>sa. 0 \<le> u x) \<and> setsum u sa = 1 \<and> (\<Sum>v\<in>sa. u v *\<^sub>R v) = y}"
-    apply(rule_tac x="Min i" in exI) unfolding subset_eq apply rule defer apply rule unfolding mem_Collect_eq
-  proof-
-    show "0 < Min i" unfolding i_def and Min_gr_iff[OF finite_imageI[OF obt(1)] `b \` t\<noteq>{}`]
-      using b apply simp apply rule apply(erule_tac x=x in ballE) using `t\<subseteq>s` by auto
-  next  fix y assume "y \<in> cball a (Min i)"
-    hence y:"norm (a - y) \<le> Min i" unfolding dist_norm[THEN sym] by auto
-    { fix x assume "x\<in>t"
-      hence "Min i \<le> b x" unfolding i_def apply(rule_tac Min_le) using obt(1) by auto
-      hence "x + (y - a) \<in> cball x (b x)" using y unfolding mem_cball dist_norm by auto
-      moreover from `x\<in>t` have "x\<in>s" using obt(2) by auto
-      ultimately have "x + (y - a) \<in> s" using y and b[THEN bspec[where x=x]] unfolding subset_eq by auto }
-    moreover
-    have *:"inj_on (\<lambda>v. v + (y - a)) t" unfolding inj_on_def by auto
-    have "(\<Sum>v\<in>(\<lambda>v. v + (y - a)) ` t. u (v - (y - a))) = 1"
-      unfolding setsum_reindex[OF *] o_def using obt(4) by auto
-    moreover have "(\<Sum>v\<in>(\<lambda>v. v + (y - a)) ` t. u (v - (y - a)) *\<^sub>R v) = y"
-      unfolding setsum_reindex[OF *] o_def using obt(4,5)
-      by (simp add: setsum_addf setsum_subtractf scaleR_left.setsum[THEN sym] scaleR_right_distrib)
-    ultimately show "\<exists>sa u. finite sa \<and> (\<forall>x\<in>sa. x \<in> s) \<and> (\<forall>x\<in>sa. 0 \<le> u x) \<and> setsum u sa = 1 \<and> (\<Sum>v\<in>sa. u v *\<^sub>R v) = y"
-      apply(rule_tac x="(\<lambda>v. v + (y - a)) ` t" in exI) apply(rule_tac x="\<lambda>v. u (v - (y - a))" in exI)
-      using obt(1, 3) by auto
-  qed
-qed
-
-lemma open_dest_vec1_vimage: "open S \<Longrightarrow> open (dest_vec1 -` S)"
-unfolding open_vector_def all_1
-by (auto simp add: dest_vec1_def)
-
-lemma tendsto_dest_vec1 [tendsto_intros]:
-  "(f ---> l) net \<Longrightarrow> ((\<lambda>x. dest_vec1 (f x)) ---> dest_vec1 l) net"
-  unfolding tendsto_def
-  apply clarify
-  apply (drule_tac x="dest_vec1 -` S" in spec)
-  apply (simp add: open_dest_vec1_vimage)
-  done
-
-lemma continuous_dest_vec1: "continuous net f \<Longrightarrow> continuous net (\<lambda>x. dest_vec1 (f x))"
-  unfolding continuous_def by (rule tendsto_dest_vec1)
-
-(* TODO: move *)
-lemma compact_real_interval:
-  fixes a b :: real shows "compact {a..b}"
-proof -
-  have "continuous_on {vec1 a .. vec1 b} dest_vec1"
-    unfolding continuous_on
-    by (simp add: tendsto_dest_vec1 Lim_at_within Lim_ident_at)
-  moreover have "compact {vec1 a .. vec1 b}" by (rule compact_interval)
-  ultimately have "compact (dest_vec1 ` {vec1 a .. vec1 b})"
-    by (rule compact_continuous_image)
-  also have "dest_vec1 ` {vec1 a .. vec1 b} = {a..b}"
-    by (auto simp add: image_def Bex_def exists_vec1)
-  finally show ?thesis .
-qed
-
-lemma compact_convex_combinations:
-  fixes s t :: "'a::real_normed_vector set"
-  assumes "compact s" "compact t"
-  shows "compact { (1 - u) *\<^sub>R x + u *\<^sub>R y | x y u. 0 \<le> u \<and> u \<le> 1 \<and> x \<in> s \<and> y \<in> t}"
-proof-
-  let ?X = "{0..1} \<times> s \<times> t"
-  let ?h = "(\<lambda>z. (1 - fst z) *\<^sub>R fst (snd z) + fst z *\<^sub>R snd (snd z))"
-  have *:"{ (1 - u) *\<^sub>R x + u *\<^sub>R y | x y u. 0 \<le> u \<and> u \<le> 1 \<and> x \<in> s \<and> y \<in> t} = ?h ` ?X"
-    apply(rule set_ext) unfolding image_iff mem_Collect_eq
-    apply rule apply auto
-    apply (rule_tac x=u in rev_bexI, simp)
-    apply (erule rev_bexI, erule rev_bexI, simp)
-    by auto
-  have "continuous_on ({0..1} \<times> s \<times> t)
-     (\<lambda>z. (1 - fst z) *\<^sub>R fst (snd z) + fst z *\<^sub>R snd (snd z))"
-    unfolding continuous_on by (rule ballI) (intro tendsto_intros)
-  thus ?thesis unfolding *
-    apply (rule compact_continuous_image)
-    apply (intro compact_Times compact_real_interval assms)
-    done
-qed
-
-lemma compact_convex_hull: fixes s::"(real^'n::finite) set"
-  assumes "compact s"  shows "compact(convex hull s)"
-proof(cases "s={}")
-  case True thus ?thesis using compact_empty by simp
-next
-  case False then obtain w where "w\<in>s" by auto
-  show ?thesis unfolding caratheodory[of s]
-  proof(induct "CARD('n) + 1")
-    have *:"{x.\<exists>sa. finite sa \<and> sa \<subseteq> s \<and> card sa \<le> 0 \<and> x \<in> convex hull sa} = {}" 
-      using compact_empty by (auto simp add: convex_hull_empty)
-    case 0 thus ?case unfolding * by simp
-  next
-    case (Suc n)
-    show ?case proof(cases "n=0")
-      case True have "{x. \<exists>t. finite t \<and> t \<subseteq> s \<and> card t \<le> Suc n \<and> x \<in> convex hull t} = s"
-        unfolding expand_set_eq and mem_Collect_eq proof(rule, rule)
-        fix x assume "\<exists>t. finite t \<and> t \<subseteq> s \<and> card t \<le> Suc n \<and> x \<in> convex hull t"
-        then obtain t where t:"finite t" "t \<subseteq> s" "card t \<le> Suc n" "x \<in> convex hull t" by auto
-        show "x\<in>s" proof(cases "card t = 0")
-          case True thus ?thesis using t(4) unfolding card_0_eq[OF t(1)] by(simp add: convex_hull_empty)
-        next
-          case False hence "card t = Suc 0" using t(3) `n=0` by auto
-          then obtain a where "t = {a}" unfolding card_Suc_eq by auto
-          thus ?thesis using t(2,4) by (simp add: convex_hull_singleton)
-        qed
-      next
-        fix x assume "x\<in>s"
-        thus "\<exists>t. finite t \<and> t \<subseteq> s \<and> card t \<le> Suc n \<and> x \<in> convex hull t"
-          apply(rule_tac x="{x}" in exI) unfolding convex_hull_singleton by auto 
-      qed thus ?thesis using assms by simp
-    next
-      case False have "{x. \<exists>t. finite t \<and> t \<subseteq> s \<and> card t \<le> Suc n \<and> x \<in> convex hull t} =
-        { (1 - u) *\<^sub>R x + u *\<^sub>R y | x y u. 
-        0 \<le> u \<and> u \<le> 1 \<and> x \<in> s \<and> y \<in> {x. \<exists>t. finite t \<and> t \<subseteq> s \<and> card t \<le> n \<and> x \<in> convex hull t}}"
-        unfolding expand_set_eq and mem_Collect_eq proof(rule,rule)
-        fix x assume "\<exists>u v c. x = (1 - c) *\<^sub>R u + c *\<^sub>R v \<and>
-          0 \<le> c \<and> c \<le> 1 \<and> u \<in> s \<and> (\<exists>t. finite t \<and> t \<subseteq> s \<and> card t \<le> n \<and> v \<in> convex hull t)"
-        then obtain u v c t where obt:"x = (1 - c) *\<^sub>R u + c *\<^sub>R v"
-          "0 \<le> c \<and> c \<le> 1" "u \<in> s" "finite t" "t \<subseteq> s" "card t \<le> n"  "v \<in> convex hull t" by auto
-        moreover have "(1 - c) *\<^sub>R u + c *\<^sub>R v \<in> convex hull insert u t"
-          apply(rule mem_convex) using obt(2) and convex_convex_hull and hull_subset[of "insert u t" convex]
-          using obt(7) and hull_mono[of t "insert u t"] by auto
-        ultimately show "\<exists>t. finite t \<and> t \<subseteq> s \<and> card t \<le> Suc n \<and> x \<in> convex hull t"
-          apply(rule_tac x="insert u t" in exI) by (auto simp add: card_insert_if)
-      next
-        fix x assume "\<exists>t. finite t \<and> t \<subseteq> s \<and> card t \<le> Suc n \<and> x \<in> convex hull t"
-        then obtain t where t:"finite t" "t \<subseteq> s" "card t \<le> Suc n" "x \<in> convex hull t" by auto
-        let ?P = "\<exists>u v c. x = (1 - c) *\<^sub>R u + c *\<^sub>R v \<and>
-          0 \<le> c \<and> c \<le> 1 \<and> u \<in> s \<and> (\<exists>t. finite t \<and> t \<subseteq> s \<and> card t \<le> n \<and> v \<in> convex hull t)"
-        show ?P proof(cases "card t = Suc n")
-          case False hence "card t \<le> n" using t(3) by auto
-          thus ?P apply(rule_tac x=w in exI, rule_tac x=x in exI, rule_tac x=1 in exI) using `w\<in>s` and t
-            by(auto intro!: exI[where x=t])
-        next
-          case True then obtain a u where au:"t = insert a u" "a\<notin>u" apply(drule_tac card_eq_SucD) by auto
-          show ?P proof(cases "u={}")
-            case True hence "x=a" using t(4)[unfolded au] by auto
-            show ?P unfolding `x=a` apply(rule_tac x=a in exI, rule_tac x=a in exI, rule_tac x=1 in exI)
-              using t and `n\<noteq>0` unfolding au by(auto intro!: exI[where x="{a}"] simp add: convex_hull_singleton)
-          next
-            case False obtain ux vx b where obt:"ux\<ge>0" "vx\<ge>0" "ux + vx = 1" "b \<in> convex hull u" "x = ux *\<^sub>R a + vx *\<^sub>R b"
-              using t(4)[unfolded au convex_hull_insert[OF False]] by auto
-            have *:"1 - vx = ux" using obt(3) by auto
-            show ?P apply(rule_tac x=a in exI, rule_tac x=b in exI, rule_tac x=vx in exI)
-              using obt and t(1-3) unfolding au and * using card_insert_disjoint[OF _ au(2)]
-              by(auto intro!: exI[where x=u])
-          qed
-        qed
-      qed
-      thus ?thesis using compact_convex_combinations[OF assms Suc] by simp 
-    qed
-  qed 
-qed
-
-lemma finite_imp_compact_convex_hull:
-  fixes s :: "(real ^ _) set"
-  shows "finite s \<Longrightarrow> compact(convex hull s)"
-  apply(drule finite_imp_compact, drule compact_convex_hull) by assumption
-
-subsection {* Extremal points of a simplex are some vertices. *}
-
-lemma dist_increases_online:
-  fixes a b d :: "'a::real_inner"
-  assumes "d \<noteq> 0"
-  shows "dist a (b + d) > dist a b \<or> dist a (b - d) > dist a b"
-proof(cases "inner a d - inner b d > 0")
-  case True hence "0 < inner d d + (inner a d * 2 - inner b d * 2)" 
-    apply(rule_tac add_pos_pos) using assms by auto
-  thus ?thesis apply(rule_tac disjI2) unfolding dist_norm and norm_eq_sqrt_inner and real_sqrt_less_iff
-    by (simp add: algebra_simps inner_commute)
-next
-  case False hence "0 < inner d d + (inner b d * 2 - inner a d * 2)" 
-    apply(rule_tac add_pos_nonneg) using assms by auto
-  thus ?thesis apply(rule_tac disjI1) unfolding dist_norm and norm_eq_sqrt_inner and real_sqrt_less_iff
-    by (simp add: algebra_simps inner_commute)
-qed
-
-lemma norm_increases_online:
-  fixes d :: "'a::real_inner"
-  shows "d \<noteq> 0 \<Longrightarrow> norm(a + d) > norm a \<or> norm(a - d) > norm a"
-  using dist_increases_online[of d a 0] unfolding dist_norm by auto
-
-lemma simplex_furthest_lt:
-  fixes s::"'a::real_inner set" assumes "finite s"
-  shows "\<forall>x \<in> (convex hull s).  x \<notin> s \<longrightarrow> (\<exists>y\<in>(convex hull s). norm(x - a) < norm(y - a))"
-proof(induct_tac rule: finite_induct[of s])
-  fix x s assume as:"finite s" "x\<notin>s" "\<forall>x\<in>convex hull s. x \<notin> s \<longrightarrow> (\<exists>y\<in>convex hull s. norm (x - a) < norm (y - a))"
-  show "\<forall>xa\<in>convex hull insert x s. xa \<notin> insert x s \<longrightarrow> (\<exists>y\<in>convex hull insert x s. norm (xa - a) < norm (y - a))"
-  proof(rule,rule,cases "s = {}")
-    case False fix y assume y:"y \<in> convex hull insert x s" "y \<notin> insert x s"
-    obtain u v b where obt:"u\<ge>0" "v\<ge>0" "u + v = 1" "b \<in> convex hull s" "y = u *\<^sub>R x + v *\<^sub>R b"
-      using y(1)[unfolded convex_hull_insert[OF False]] by auto
-    show "\<exists>z\<in>convex hull insert x s. norm (y - a) < norm (z - a)"
-    proof(cases "y\<in>convex hull s")
-      case True then obtain z where "z\<in>convex hull s" "norm (y - a) < norm (z - a)"
-        using as(3)[THEN bspec[where x=y]] and y(2) by auto
-      thus ?thesis apply(rule_tac x=z in bexI) unfolding convex_hull_insert[OF False] by auto
-    next
-      case False show ?thesis  using obt(3) proof(cases "u=0", case_tac[!] "v=0")
-        assume "u=0" "v\<noteq>0" hence "y = b" using obt by auto
-        thus ?thesis using False and obt(4) by auto
-      next
-        assume "u\<noteq>0" "v=0" hence "y = x" using obt by auto
-        thus ?thesis using y(2) by auto
-      next
-        assume "u\<noteq>0" "v\<noteq>0"
-        then obtain w where w:"w>0" "w<u" "w<v" using real_lbound_gt_zero[of u v] and obt(1,2) by auto
-        have "x\<noteq>b" proof(rule ccontr) 
-          assume "\<not> x\<noteq>b" hence "y=b" unfolding obt(5)
-            using obt(3) by(auto simp add: scaleR_left_distrib[THEN sym])
-          thus False using obt(4) and False by simp qed
-        hence *:"w *\<^sub>R (x - b) \<noteq> 0" using w(1) by auto
-        show ?thesis using dist_increases_online[OF *, of a y]
-        proof(erule_tac disjE)
-          assume "dist a y < dist a (y + w *\<^sub>R (x - b))"
-          hence "norm (y - a) < norm ((u + w) *\<^sub>R x + (v - w) *\<^sub>R b - a)"
-            unfolding dist_commute[of a] unfolding dist_norm obt(5) by (simp add: algebra_simps)
-          moreover have "(u + w) *\<^sub>R x + (v - w) *\<^sub>R b \<in> convex hull insert x s"
-            unfolding convex_hull_insert[OF `s\<noteq>{}`] and mem_Collect_eq
-            apply(rule_tac x="u + w" in exI) apply rule defer 
-            apply(rule_tac x="v - w" in exI) using `u\<ge>0` and w and obt(3,4) by auto
-          ultimately show ?thesis by auto
-        next
-          assume "dist a y < dist a (y - w *\<^sub>R (x - b))"
-          hence "norm (y - a) < norm ((u - w) *\<^sub>R x + (v + w) *\<^sub>R b - a)"
-            unfolding dist_commute[of a] unfolding dist_norm obt(5) by (simp add: algebra_simps)
-          moreover have "(u - w) *\<^sub>R x + (v + w) *\<^sub>R b \<in> convex hull insert x s"
-            unfolding convex_hull_insert[OF `s\<noteq>{}`] and mem_Collect_eq
-            apply(rule_tac x="u - w" in exI) apply rule defer 
-            apply(rule_tac x="v + w" in exI) using `u\<ge>0` and w and obt(3,4) by auto
-          ultimately show ?thesis by auto
-        qed
-      qed auto
-    qed
-  qed auto
-qed (auto simp add: assms)
-
-lemma simplex_furthest_le:
-  fixes s :: "(real ^ _) set"
-  assumes "finite s" "s \<noteq> {}"
-  shows "\<exists>y\<in>s. \<forall>x\<in>(convex hull s). norm(x - a) \<le> norm(y - a)"
-proof-
-  have "convex hull s \<noteq> {}" using hull_subset[of s convex] and assms(2) by auto
-  then obtain x where x:"x\<in>convex hull s" "\<forall>y\<in>convex hull s. norm (y - a) \<le> norm (x - a)"
-    using distance_attains_sup[OF finite_imp_compact_convex_hull[OF assms(1)], of a]
-    unfolding dist_commute[of a] unfolding dist_norm by auto
-  thus ?thesis proof(cases "x\<in>s")
-    case False then obtain y where "y\<in>convex hull s" "norm (x - a) < norm (y - a)"
-      using simplex_furthest_lt[OF assms(1), THEN bspec[where x=x]] and x(1) by auto
-    thus ?thesis using x(2)[THEN bspec[where x=y]] by auto
-  qed auto
-qed
-
-lemma simplex_furthest_le_exists:
-  fixes s :: "(real ^ _) set"
-  shows "finite s \<Longrightarrow> (\<forall>x\<in>(convex hull s). \<exists>y\<in>s. norm(x - a) \<le> norm(y - a))"
-  using simplex_furthest_le[of s] by (cases "s={}")auto
-
-lemma simplex_extremal_le:
-  fixes s :: "(real ^ _) set"
-  assumes "finite s" "s \<noteq> {}"
-  shows "\<exists>u\<in>s. \<exists>v\<in>s. \<forall>x\<in>convex hull s. \<forall>y \<in> convex hull s. norm(x - y) \<le> norm(u - v)"
-proof-
-  have "convex hull s \<noteq> {}" using hull_subset[of s convex] and assms(2) by auto
-  then obtain u v where obt:"u\<in>convex hull s" "v\<in>convex hull s"
-    "\<forall>x\<in>convex hull s. \<forall>y\<in>convex hull s. norm (x - y) \<le> norm (u - v)"
-    using compact_sup_maxdistance[OF finite_imp_compact_convex_hull[OF assms(1)]] by auto
-  thus ?thesis proof(cases "u\<notin>s \<or> v\<notin>s", erule_tac disjE)
-    assume "u\<notin>s" then obtain y where "y\<in>convex hull s" "norm (u - v) < norm (y - v)"
-      using simplex_furthest_lt[OF assms(1), THEN bspec[where x=u]] and obt(1) by auto
-    thus ?thesis using obt(3)[THEN bspec[where x=y], THEN bspec[where x=v]] and obt(2) by auto
-  next
-    assume "v\<notin>s" then obtain y where "y\<in>convex hull s" "norm (v - u) < norm (y - u)"
-      using simplex_furthest_lt[OF assms(1), THEN bspec[where x=v]] and obt(2) by auto
-    thus ?thesis using obt(3)[THEN bspec[where x=u], THEN bspec[where x=y]] and obt(1)
-      by (auto simp add: norm_minus_commute)
-  qed auto
-qed 
-
-lemma simplex_extremal_le_exists:
-  fixes s :: "(real ^ _) set"
-  shows "finite s \<Longrightarrow> x \<in> convex hull s \<Longrightarrow> y \<in> convex hull s
-  \<Longrightarrow> (\<exists>u\<in>s. \<exists>v\<in>s. norm(x - y) \<le> norm(u - v))"
-  using convex_hull_empty simplex_extremal_le[of s] by(cases "s={}")auto
-
-subsection {* Closest point of a convex set is unique, with a continuous projection. *}
-
-definition
-  closest_point :: "(real ^ 'n::finite) set \<Rightarrow> real ^ 'n \<Rightarrow> real ^ 'n" where
- "closest_point s a = (SOME x. x \<in> s \<and> (\<forall>y\<in>s. dist a x \<le> dist a y))"
-
-lemma closest_point_exists:
-  assumes "closed s" "s \<noteq> {}"
-  shows  "closest_point s a \<in> s" "\<forall>y\<in>s. dist a (closest_point s a) \<le> dist a y"
-  unfolding closest_point_def apply(rule_tac[!] someI2_ex) 
-  using distance_attains_inf[OF assms(1,2), of a] by auto
-
-lemma closest_point_in_set:
-  "closed s \<Longrightarrow> s \<noteq> {} \<Longrightarrow> (closest_point s a) \<in> s"
-  by(meson closest_point_exists)
-
-lemma closest_point_le:
-  "closed s \<Longrightarrow> x \<in> s \<Longrightarrow> dist a (closest_point s a) \<le> dist a x"
-  using closest_point_exists[of s] by auto
-
-lemma closest_point_self:
-  assumes "x \<in> s"  shows "closest_point s x = x"
-  unfolding closest_point_def apply(rule some1_equality, rule ex1I[of _ x]) 
-  using assms by auto
-
-lemma closest_point_refl:
- "closed s \<Longrightarrow> s \<noteq> {} \<Longrightarrow> (closest_point s x = x \<longleftrightarrow> x \<in> s)"
-  using closest_point_in_set[of s x] closest_point_self[of x s] by auto
-
-(* TODO: move *)
-lemma norm_lt: "norm x < norm y \<longleftrightarrow> inner x x < inner y y"
-  unfolding norm_eq_sqrt_inner by simp
-
-(* TODO: move *)
-lemma norm_le: "norm x \<le> norm y \<longleftrightarrow> inner x x \<le> inner y y"
-  unfolding norm_eq_sqrt_inner by simp
-
-lemma closer_points_lemma: fixes y::"real^'n::finite"
-  assumes "inner y z > 0"
-  shows "\<exists>u>0. \<forall>v>0. v \<le> u \<longrightarrow> norm(v *\<^sub>R z - y) < norm y"
-proof- have z:"inner z z > 0" unfolding inner_gt_zero_iff using assms by auto
-  thus ?thesis using assms apply(rule_tac x="inner y z / inner z z" in exI) apply(rule) defer proof(rule+)
-    fix v assume "0<v" "v \<le> inner y z / inner z z"
-    thus "norm (v *\<^sub>R z - y) < norm y" unfolding norm_lt using z and assms
-      by (simp add: field_simps inner_diff inner_commute mult_strict_left_mono[OF _ `0<v`])
-  qed(rule divide_pos_pos, auto) qed
-
-lemma closer_point_lemma:
-  fixes x y z :: "real ^ 'n::finite"
-  assumes "inner (y - x) (z - x) > 0"
-  shows "\<exists>u>0. u \<le> 1 \<and> dist (x + u *\<^sub>R (z - x)) y < dist x y"
-proof- obtain u where "u>0" and u:"\<forall>v>0. v \<le> u \<longrightarrow> norm (v *\<^sub>R (z - x) - (y - x)) < norm (y - x)"
-    using closer_points_lemma[OF assms] by auto
-  show ?thesis apply(rule_tac x="min u 1" in exI) using u[THEN spec[where x="min u 1"]] and `u>0`
-    unfolding dist_norm by(auto simp add: norm_minus_commute field_simps) qed
-
-lemma any_closest_point_dot:
-  fixes s :: "(real ^ _) set"
-  assumes "convex s" "closed s" "x \<in> s" "y \<in> s" "\<forall>z\<in>s. dist a x \<le> dist a z"
-  shows "inner (a - x) (y - x) \<le> 0"
-proof(rule ccontr) assume "\<not> inner (a - x) (y - x) \<le> 0"
-  then obtain u where u:"u>0" "u\<le>1" "dist (x + u *\<^sub>R (y - x)) a < dist x a" using closer_point_lemma[of a x y] by auto
-  let ?z = "(1 - u) *\<^sub>R x + u *\<^sub>R y" have "?z \<in> s" using mem_convex[OF assms(1,3,4), of u] using u by auto
-  thus False using assms(5)[THEN bspec[where x="?z"]] and u(3) by (auto simp add: dist_commute algebra_simps) qed
-
-(* TODO: move *)
-lemma norm_le_square: "norm x \<le> a \<longleftrightarrow> 0 \<le> a \<and> inner x x \<le> a\<twosuperior>"
-proof -
-  have "norm x \<le> a \<longleftrightarrow> 0 \<le> a \<and> norm x \<le> a"
-    using norm_ge_zero [of x] by arith
-  also have "\<dots> \<longleftrightarrow> 0 \<le> a \<and> (norm x)\<twosuperior> \<le> a\<twosuperior>"
-    by (auto intro: power_mono dest: power2_le_imp_le)
-  also have "\<dots> \<longleftrightarrow> 0 \<le> a \<and> inner x x \<le> a\<twosuperior>"
-    unfolding power2_norm_eq_inner ..
-  finally show ?thesis .
-qed
-
-lemma any_closest_point_unique:
-  fixes s :: "(real ^ _) set"
-  assumes "convex s" "closed s" "x \<in> s" "y \<in> s"
-  "\<forall>z\<in>s. dist a x \<le> dist a z" "\<forall>z\<in>s. dist a y \<le> dist a z"
-  shows "x = y" using any_closest_point_dot[OF assms(1-4,5)] and any_closest_point_dot[OF assms(1-2,4,3,6)]
-  unfolding norm_pths(1) and norm_le_square
-  by (auto simp add: algebra_simps)
-
-lemma closest_point_unique:
-  assumes "convex s" "closed s" "x \<in> s" "\<forall>z\<in>s. dist a x \<le> dist a z"
-  shows "x = closest_point s a"
-  using any_closest_point_unique[OF assms(1-3) _ assms(4), of "closest_point s a"] 
-  using closest_point_exists[OF assms(2)] and assms(3) by auto
-
-lemma closest_point_dot:
-  assumes "convex s" "closed s" "x \<in> s"
-  shows "inner (a - closest_point s a) (x - closest_point s a) \<le> 0"
-  apply(rule any_closest_point_dot[OF assms(1,2) _ assms(3)])
-  using closest_point_exists[OF assms(2)] and assms(3) by auto
-
-lemma closest_point_lt:
-  assumes "convex s" "closed s" "x \<in> s" "x \<noteq> closest_point s a"
-  shows "dist a (closest_point s a) < dist a x"
-  apply(rule ccontr) apply(rule_tac notE[OF assms(4)])
-  apply(rule closest_point_unique[OF assms(1-3), of a])
-  using closest_point_le[OF assms(2), of _ a] by fastsimp
-
-lemma closest_point_lipschitz:
-  assumes "convex s" "closed s" "s \<noteq> {}"
-  shows "dist (closest_point s x) (closest_point s y) \<le> dist x y"
-proof-
-  have "inner (x - closest_point s x) (closest_point s y - closest_point s x) \<le> 0"
-       "inner (y - closest_point s y) (closest_point s x - closest_point s y) \<le> 0"
-    apply(rule_tac[!] any_closest_point_dot[OF assms(1-2)])
-    using closest_point_exists[OF assms(2-3)] by auto
-  thus ?thesis unfolding dist_norm and norm_le
-    using inner_ge_zero[of "(x - closest_point s x) - (y - closest_point s y)"]
-    by (simp add: inner_add inner_diff inner_commute) qed
-
-lemma continuous_at_closest_point:
-  assumes "convex s" "closed s" "s \<noteq> {}"
-  shows "continuous (at x) (closest_point s)"
-  unfolding continuous_at_eps_delta 
-  using le_less_trans[OF closest_point_lipschitz[OF assms]] by auto
-
-lemma continuous_on_closest_point:
-  assumes "convex s" "closed s" "s \<noteq> {}"
-  shows "continuous_on t (closest_point s)"
-  apply(rule continuous_at_imp_continuous_on) using continuous_at_closest_point[OF assms] by auto
-
-subsection {* Various point-to-set separating/supporting hyperplane theorems. *}
-
-lemma supporting_hyperplane_closed_point:
-  fixes s :: "(real ^ _) set"
-  assumes "convex s" "closed s" "s \<noteq> {}" "z \<notin> s"
-  shows "\<exists>a b. \<exists>y\<in>s. inner a z < b \<and> (inner a y = b) \<and> (\<forall>x\<in>s. inner a x \<ge> b)"
-proof-
-  from distance_attains_inf[OF assms(2-3)] obtain y where "y\<in>s" and y:"\<forall>x\<in>s. dist z y \<le> dist z x" by auto
-  show ?thesis apply(rule_tac x="y - z" in exI, rule_tac x="inner (y - z) y" in exI, rule_tac x=y in bexI)
-    apply rule defer apply rule defer apply(rule, rule ccontr) using `y\<in>s` proof-
-    show "inner (y - z) z < inner (y - z) y" apply(subst diff_less_iff(1)[THEN sym])
-      unfolding inner_diff_right[THEN sym] and inner_gt_zero_iff using `y\<in>s` `z\<notin>s` by auto
-  next
-    fix x assume "x\<in>s" have *:"\<forall>u. 0 \<le> u \<and> u \<le> 1 \<longrightarrow> dist z y \<le> dist z ((1 - u) *\<^sub>R y + u *\<^sub>R x)"
-      using assms(1)[unfolded convex_alt] and y and `x\<in>s` and `y\<in>s` by auto
-    assume "\<not> inner (y - z) y \<le> inner (y - z) x" then obtain v where
-      "v>0" "v\<le>1" "dist (y + v *\<^sub>R (x - y)) z < dist y z" using closer_point_lemma[of z y x] apply - by (auto simp add: inner_diff)
-    thus False using *[THEN spec[where x=v]] by(auto simp add: dist_commute algebra_simps)
-  qed auto
-qed
-
-lemma separating_hyperplane_closed_point:
-  fixes s :: "(real ^ _) set"
-  assumes "convex s" "closed s" "z \<notin> s"
-  shows "\<exists>a b. inner a z < b \<and> (\<forall>x\<in>s. inner a x > b)"
-proof(cases "s={}")
-  case True thus ?thesis apply(rule_tac x="-z" in exI, rule_tac x=1 in exI)
-    using less_le_trans[OF _ inner_ge_zero[of z]] by auto
-next
-  case False obtain y where "y\<in>s" and y:"\<forall>x\<in>s. dist z y \<le> dist z x"
-    using distance_attains_inf[OF assms(2) False] by auto
-  show ?thesis apply(rule_tac x="y - z" in exI, rule_tac x="inner (y - z) z + (norm(y - z))\<twosuperior> / 2" in exI)
-    apply rule defer apply rule proof-
-    fix x assume "x\<in>s"
-    have "\<not> 0 < inner (z - y) (x - y)" apply(rule_tac notI) proof(drule closer_point_lemma)
-      assume "\<exists>u>0. u \<le> 1 \<and> dist (y + u *\<^sub>R (x - y)) z < dist y z"
-      then obtain u where "u>0" "u\<le>1" "dist (y + u *\<^sub>R (x - y)) z < dist y z" by auto
-      thus False using y[THEN bspec[where x="y + u *\<^sub>R (x - y)"]]
-        using assms(1)[unfolded convex_alt, THEN bspec[where x=y]]
-        using `x\<in>s` `y\<in>s` by (auto simp add: dist_commute algebra_simps) qed
-    moreover have "0 < norm (y - z) ^ 2" using `y\<in>s` `z\<notin>s` by auto
-    hence "0 < inner (y - z) (y - z)" unfolding power2_norm_eq_inner by simp
-    ultimately show "inner (y - z) z + (norm (y - z))\<twosuperior> / 2 < inner (y - z) x"
-      unfolding power2_norm_eq_inner and not_less by (auto simp add: field_simps inner_commute inner_diff)
-  qed(insert `y\<in>s` `z\<notin>s`, auto)
-qed
-
-lemma separating_hyperplane_closed_0:
-  assumes "convex (s::(real^'n::finite) set)" "closed s" "0 \<notin> s"
-  shows "\<exists>a b. a \<noteq> 0 \<and> 0 < b \<and> (\<forall>x\<in>s. inner a x > b)"
-  proof(cases "s={}") guess a using UNIV_witness[where 'a='n] ..
-  case True have "norm ((basis a)::real^'n::finite) = 1" 
-    using norm_basis and dimindex_ge_1 by auto
-  thus ?thesis apply(rule_tac x="basis a" in exI, rule_tac x=1 in exI) using True by auto
-next case False thus ?thesis using False using separating_hyperplane_closed_point[OF assms]
-    apply - apply(erule exE)+ unfolding dot_rzero apply(rule_tac x=a in exI, rule_tac x=b in exI) by auto qed
-
-subsection {* Now set-to-set for closed/compact sets. *}
-
-lemma separating_hyperplane_closed_compact:
-  assumes "convex (s::(real^'n::finite) set)" "closed s" "convex t" "compact t" "t \<noteq> {}" "s \<inter> t = {}"
-  shows "\<exists>a b. (\<forall>x\<in>s. inner a x < b) \<and> (\<forall>x\<in>t. inner a x > b)"
-proof(cases "s={}")
-  case True
-  obtain b where b:"b>0" "\<forall>x\<in>t. norm x \<le> b" using compact_imp_bounded[OF assms(4)] unfolding bounded_pos by auto
-  obtain z::"real^'n" where z:"norm z = b + 1" using vector_choose_size[of "b + 1"] and b(1) by auto
-  hence "z\<notin>t" using b(2)[THEN bspec[where x=z]] by auto
-  then obtain a b where ab:"inner a z < b" "\<forall>x\<in>t. b < inner a x"
-    using separating_hyperplane_closed_point[OF assms(3) compact_imp_closed[OF assms(4)], of z] by auto
-  thus ?thesis using True by auto
-next
-  case False then obtain y where "y\<in>s" by auto
-  obtain a b where "0 < b" "\<forall>x\<in>{x - y |x y. x \<in> s \<and> y \<in> t}. b < inner a x"
-    using separating_hyperplane_closed_point[OF convex_differences[OF assms(1,3)], of 0]
-    using closed_compact_differences[OF assms(2,4)] using assms(6) by(auto, blast)
-  hence ab:"\<forall>x\<in>s. \<forall>y\<in>t. b + inner a y < inner a x" apply- apply(rule,rule) apply(erule_tac x="x - y" in ballE) by (auto simp add: inner_diff)
-  def k \<equiv> "rsup ((\<lambda>x. inner a x) ` t)"
-  show ?thesis apply(rule_tac x="-a" in exI, rule_tac x="-(k + b / 2)" in exI)
-    apply(rule,rule) defer apply(rule) unfolding inner_minus_left and neg_less_iff_less proof-
-    from ab have "((\<lambda>x. inner a x) ` t) *<= (inner a y - b)"
-      apply(erule_tac x=y in ballE) apply(rule setleI) using `y\<in>s` by auto
-    hence k:"isLub UNIV ((\<lambda>x. inner a x) ` t) k" unfolding k_def apply(rule_tac rsup) using assms(5) by auto
-    fix x assume "x\<in>t" thus "inner a x < (k + b / 2)" using `0<b` and isLubD2[OF k, of "inner a x"] by auto
-  next
-    fix x assume "x\<in>s" 
-    hence "k \<le> inner a x - b" unfolding k_def apply(rule_tac rsup_le) using assms(5)
-      unfolding setle_def
-      using ab[THEN bspec[where x=x]] by auto
-    thus "k + b / 2 < inner a x" using `0 < b` by auto
-  qed
-qed
-
-lemma separating_hyperplane_compact_closed:
-  fixes s :: "(real ^ _) set"
-  assumes "convex s" "compact s" "s \<noteq> {}" "convex t" "closed t" "s \<inter> t = {}"
-  shows "\<exists>a b. (\<forall>x\<in>s. inner a x < b) \<and> (\<forall>x\<in>t. inner a x > b)"
-proof- obtain a b where "(\<forall>x\<in>t. inner a x < b) \<and> (\<forall>x\<in>s. b < inner a x)"
-    using separating_hyperplane_closed_compact[OF assms(4-5,1-2,3)] and assms(6) by auto
-  thus ?thesis apply(rule_tac x="-a" in exI, rule_tac x="-b" in exI) by auto qed
-
-subsection {* General case without assuming closure and getting non-strict separation. *}
-
-lemma separating_hyperplane_set_0:
-  assumes "convex s" "(0::real^'n::finite) \<notin> s"
-  shows "\<exists>a. a \<noteq> 0 \<and> (\<forall>x\<in>s. 0 \<le> inner a x)"
-proof- let ?k = "\<lambda>c. {x::real^'n. 0 \<le> inner c x}"
-  have "frontier (cball 0 1) \<inter> (\<Inter> (?k ` s)) \<noteq> {}"
-    apply(rule compact_imp_fip) apply(rule compact_frontier[OF compact_cball])
-    defer apply(rule,rule,erule conjE) proof-
-    fix f assume as:"f \<subseteq> ?k ` s" "finite f"
-    obtain c where c:"f = ?k ` c" "c\<subseteq>s" "finite c" using finite_subset_image[OF as(2,1)] by auto
-    then obtain a b where ab:"a \<noteq> 0" "0 < b"  "\<forall>x\<in>convex hull c. b < inner a x"
-      using separating_hyperplane_closed_0[OF convex_convex_hull, of c]
-      using finite_imp_compact_convex_hull[OF c(3), THEN compact_imp_closed] and assms(2)
-      using subset_hull[unfolded mem_def, of convex, OF assms(1), THEN sym, of c] by auto
-    hence "\<exists>x. norm x = 1 \<and> (\<forall>y\<in>c. 0 \<le> inner y x)" apply(rule_tac x="inverse(norm a) *\<^sub>R a" in exI)
-       using hull_subset[of c convex] unfolding subset_eq and inner_scaleR
-       apply- apply rule defer apply rule apply(rule mult_nonneg_nonneg)
-       by(auto simp add: inner_commute elim!: ballE)
-    thus "frontier (cball 0 1) \<inter> \<Inter>f \<noteq> {}" unfolding c(1) frontier_cball dist_norm by auto
-  qed(insert closed_halfspace_ge, auto)
-  then obtain x where "norm x = 1" "\<forall>y\<in>s. x\<in>?k y" unfolding frontier_cball dist_norm by auto
-  thus ?thesis apply(rule_tac x=x in exI) by(auto simp add: inner_commute) qed
-
-lemma separating_hyperplane_sets:
-  assumes "convex s" "convex (t::(real^'n::finite) set)" "s \<noteq> {}" "t \<noteq> {}" "s \<inter> t = {}"
-  shows "\<exists>a b. a \<noteq> 0 \<and> (\<forall>x\<in>s. inner a x \<le> b) \<and> (\<forall>x\<in>t. inner a x \<ge> b)"
-proof- from separating_hyperplane_set_0[OF convex_differences[OF assms(2,1)]]
-  obtain a where "a\<noteq>0" "\<forall>x\<in>{x - y |x y. x \<in> t \<and> y \<in> s}. 0 \<le> inner a x"  using assms(3-5) by auto 
-  hence "\<forall>x\<in>t. \<forall>y\<in>s. inner a y \<le> inner a x" apply- apply(rule, rule) apply(erule_tac x="x - y" in ballE) by (auto simp add: inner_diff)
-  thus ?thesis apply(rule_tac x=a in exI, rule_tac x="rsup ((\<lambda>x. inner a x) ` s)" in exI) using `a\<noteq>0`
-    apply(rule) apply(rule,rule) apply(rule rsup[THEN isLubD2]) prefer 4 apply(rule,rule rsup_le) unfolding setle_def
-    prefer 4 using assms(3-5) by blast+ qed
-
-subsection {* More convexity generalities. *}
-
-lemma convex_closure:
-  fixes s :: "'a::real_normed_vector set"
-  assumes "convex s" shows "convex(closure s)"
-  unfolding convex_def Ball_def closure_sequential
-  apply(rule,rule,rule,rule,rule,rule,rule,rule,rule) apply(erule_tac exE)+
-  apply(rule_tac x="\<lambda>n. u *\<^sub>R xb n + v *\<^sub>R xc n" in exI) apply(rule,rule)
-  apply(rule assms[unfolded convex_def, rule_format]) prefer 6
-  apply(rule Lim_add) apply(rule_tac [1-2] Lim_cmul) by auto
-
-lemma convex_interior:
-  fixes s :: "'a::real_normed_vector set"
-  assumes "convex s" shows "convex(interior s)"
-  unfolding convex_alt Ball_def mem_interior apply(rule,rule,rule,rule,rule,rule) apply(erule exE | erule conjE)+ proof-
-  fix x y u assume u:"0 \<le> u" "u \<le> (1::real)"
-  fix e d assume ed:"ball x e \<subseteq> s" "ball y d \<subseteq> s" "0<d" "0<e" 
-  show "\<exists>e>0. ball ((1 - u) *\<^sub>R x + u *\<^sub>R y) e \<subseteq> s" apply(rule_tac x="min d e" in exI)
-    apply rule unfolding subset_eq defer apply rule proof-
-    fix z assume "z \<in> ball ((1 - u) *\<^sub>R x + u *\<^sub>R y) (min d e)"
-    hence "(1- u) *\<^sub>R (z - u *\<^sub>R (y - x)) + u *\<^sub>R (z + (1 - u) *\<^sub>R (y - x)) \<in> s"
-      apply(rule_tac assms[unfolded convex_alt, rule_format])
-      using ed(1,2) and u unfolding subset_eq mem_ball Ball_def dist_norm by(auto simp add: algebra_simps)
-    thus "z \<in> s" using u by (auto simp add: algebra_simps) qed(insert u ed(3-4), auto) qed
-
-lemma convex_hull_eq_empty: "convex hull s = {} \<longleftrightarrow> s = {}"
-  using hull_subset[of s convex] convex_hull_empty by auto
-
-subsection {* Moving and scaling convex hulls. *}
-
-lemma convex_hull_translation_lemma:
-  "convex hull ((\<lambda>x. a + x) ` s) \<subseteq> (\<lambda>x. a + x) ` (convex hull s)"
-  apply(rule hull_minimal, rule image_mono, rule hull_subset) unfolding mem_def
-  using convex_translation[OF convex_convex_hull, of a s] by assumption
-
-lemma convex_hull_bilemma: fixes neg
-  assumes "(\<forall>s a. (convex hull (up a s)) \<subseteq> up a (convex hull s))"
-  shows "(\<forall>s. up a (up (neg a) s) = s) \<and> (\<forall>s. up (neg a) (up a s) = s) \<and> (\<forall>s t a. s \<subseteq> t \<longrightarrow> up a s \<subseteq> up a t)
-  \<Longrightarrow> \<forall>s. (convex hull (up a s)) = up a (convex hull s)"
-  using assms by(metis subset_antisym) 
-
-lemma convex_hull_translation:
-  "convex hull ((\<lambda>x. a + x) ` s) = (\<lambda>x. a + x) ` (convex hull s)"
-  apply(rule convex_hull_bilemma[rule_format, of _ _ "\<lambda>a. -a"], rule convex_hull_translation_lemma) unfolding image_image by auto
-
-lemma convex_hull_scaling_lemma:
- "(convex hull ((\<lambda>x. c *\<^sub>R x) ` s)) \<subseteq> (\<lambda>x. c *\<^sub>R x) ` (convex hull s)"
-  apply(rule hull_minimal, rule image_mono, rule hull_subset)
-  unfolding mem_def by(rule convex_scaling, rule convex_convex_hull)
-
-lemma convex_hull_scaling:
-  "convex hull ((\<lambda>x. c *\<^sub>R x) ` s) = (\<lambda>x. c *\<^sub>R x) ` (convex hull s)"
-  apply(cases "c=0") defer apply(rule convex_hull_bilemma[rule_format, of _ _ inverse]) apply(rule convex_hull_scaling_lemma)
-  unfolding image_image scaleR_scaleR by(auto simp add:image_constant_conv convex_hull_eq_empty)
-
-lemma convex_hull_affinity:
-  "convex hull ((\<lambda>x. a + c *\<^sub>R x) ` s) = (\<lambda>x. a + c *\<^sub>R x) ` (convex hull s)"
-  unfolding image_image[THEN sym] convex_hull_scaling convex_hull_translation  ..
-
-subsection {* Convex set as intersection of halfspaces. *}
-
-lemma convex_halfspace_intersection:
-  fixes s :: "(real ^ _) set"
-  assumes "closed s" "convex s"
-  shows "s = \<Inter> {h. s \<subseteq> h \<and> (\<exists>a b. h = {x. inner a x \<le> b})}"
-  apply(rule set_ext, rule) unfolding Inter_iff Ball_def mem_Collect_eq apply(rule,rule,erule conjE) proof- 
-  fix x  assume "\<forall>xa. s \<subseteq> xa \<and> (\<exists>a b. xa = {x. inner a x \<le> b}) \<longrightarrow> x \<in> xa"
-  hence "\<forall>a b. s \<subseteq> {x. inner a x \<le> b} \<longrightarrow> x \<in> {x. inner a x \<le> b}" by blast
-  thus "x\<in>s" apply(rule_tac ccontr) apply(drule separating_hyperplane_closed_point[OF assms(2,1)])
-    apply(erule exE)+ apply(erule_tac x="-a" in allE, erule_tac x="-b" in allE) by auto
-qed auto
-
-subsection {* Radon's theorem (from Lars Schewe). *}
-
-lemma radon_ex_lemma:
-  assumes "finite c" "affine_dependent c"
-  shows "\<exists>u. setsum u c = 0 \<and> (\<exists>v\<in>c. u v \<noteq> 0) \<and> setsum (\<lambda>v. u v *\<^sub>R v) c = 0"
-proof- from assms(2)[unfolded affine_dependent_explicit] guess s .. then guess u ..
-  thus ?thesis apply(rule_tac x="\<lambda>v. if v\<in>s then u v else 0" in exI) unfolding if_smult scaleR_zero_left
-    and setsum_restrict_set[OF assms(1), THEN sym] by(auto simp add: Int_absorb1) qed
-
-lemma radon_s_lemma:
-  assumes "finite s" "setsum f s = (0::real)"
-  shows "setsum f {x\<in>s. 0 < f x} = - setsum f {x\<in>s. f x < 0}"
-proof- have *:"\<And>x. (if f x < 0 then f x else 0) + (if 0 < f x then f x else 0) = f x" by auto
-  show ?thesis unfolding real_add_eq_0_iff[THEN sym] and setsum_restrict_set''[OF assms(1)] and setsum_addf[THEN sym] and *
-    using assms(2) by assumption qed
-
-lemma radon_v_lemma:
-  assumes "finite s" "setsum f s = 0" "\<forall>x. g x = (0::real) \<longrightarrow> f x = (0::real^'n)"
-  shows "(setsum f {x\<in>s. 0 < g x}) = - setsum f {x\<in>s. g x < 0}"
-proof-
-  have *:"\<And>x. (if 0 < g x then f x else 0) + (if g x < 0 then f x else 0) = f x" using assms(3) by auto 
-  show ?thesis unfolding eq_neg_iff_add_eq_0 and setsum_restrict_set''[OF assms(1)] and setsum_addf[THEN sym] and *
-    using assms(2) by assumption qed
-
-lemma radon_partition:
-  assumes "finite c" "affine_dependent c"
-  shows "\<exists>m p. m \<inter> p = {} \<and> m \<union> p = c \<and> (convex hull m) \<inter> (convex hull p) \<noteq> {}" proof-
-  obtain u v where uv:"setsum u c = 0" "v\<in>c" "u v \<noteq> 0"  "(\<Sum>v\<in>c. u v *\<^sub>R v) = 0" using radon_ex_lemma[OF assms] by auto
-  have fin:"finite {x \<in> c. 0 < u x}" "finite {x \<in> c. 0 > u x}" using assms(1) by auto
-  def z \<equiv> "(inverse (setsum u {x\<in>c. u x > 0})) *\<^sub>R setsum (\<lambda>x. u x *\<^sub>R x) {x\<in>c. u x > 0}"
-  have "setsum u {x \<in> c. 0 < u x} \<noteq> 0" proof(cases "u v \<ge> 0")
-    case False hence "u v < 0" by auto
-    thus ?thesis proof(cases "\<exists>w\<in>{x \<in> c. 0 < u x}. u w > 0") 
-      case True thus ?thesis using setsum_nonneg_eq_0_iff[of _ u, OF fin(1)] by auto
-    next
-      case False hence "setsum u c \<le> setsum (\<lambda>x. if x=v then u v else 0) c" apply(rule_tac setsum_mono) by auto
-      thus ?thesis unfolding setsum_delta[OF assms(1)] using uv(2) and `u v < 0` and uv(1) by auto qed
-  qed (insert setsum_nonneg_eq_0_iff[of _ u, OF fin(1)] uv(2-3), auto)
-
-  hence *:"setsum u {x\<in>c. u x > 0} > 0" unfolding real_less_def apply(rule_tac conjI, rule_tac setsum_nonneg) by auto
-  moreover have "setsum u ({x \<in> c. 0 < u x} \<union> {x \<in> c. u x < 0}) = setsum u c"
-    "(\<Sum>x\<in>{x \<in> c. 0 < u x} \<union> {x \<in> c. u x < 0}. u x *\<^sub>R x) = (\<Sum>x\<in>c. u x *\<^sub>R x)"
-    using assms(1) apply(rule_tac[!] setsum_mono_zero_left) by auto
-  hence "setsum u {x \<in> c. 0 < u x} = - setsum u {x \<in> c. 0 > u x}"
-   "(\<Sum>x\<in>{x \<in> c. 0 < u x}. u x *\<^sub>R x) = - (\<Sum>x\<in>{x \<in> c. 0 > u x}. u x *\<^sub>R x)" 
-    unfolding eq_neg_iff_add_eq_0 using uv(1,4) by (auto simp add:  setsum_Un_zero[OF fin, THEN sym]) 
-  moreover have "\<forall>x\<in>{v \<in> c. u v < 0}. 0 \<le> inverse (setsum u {x \<in> c. 0 < u x}) * - u x" 
-    apply (rule) apply (rule mult_nonneg_nonneg) using * by auto
-
-  ultimately have "z \<in> convex hull {v \<in> c. u v \<le> 0}" unfolding convex_hull_explicit mem_Collect_eq
-    apply(rule_tac x="{v \<in> c. u v < 0}" in exI, rule_tac x="\<lambda>y. inverse (setsum u {x\<in>c. u x > 0}) * - u y" in exI)
-    using assms(1) unfolding scaleR_scaleR[THEN sym] scaleR_right.setsum [symmetric] and z_def
-    by(auto simp add: setsum_negf vector_smult_lneg mult_right.setsum[THEN sym])
-  moreover have "\<forall>x\<in>{v \<in> c. 0 < u v}. 0 \<le> inverse (setsum u {x \<in> c. 0 < u x}) * u x" 
-    apply (rule) apply (rule mult_nonneg_nonneg) using * by auto 
-  hence "z \<in> convex hull {v \<in> c. u v > 0}" unfolding convex_hull_explicit mem_Collect_eq
-    apply(rule_tac x="{v \<in> c. 0 < u v}" in exI, rule_tac x="\<lambda>y. inverse (setsum u {x\<in>c. u x > 0}) * u y" in exI)
-    using assms(1) unfolding scaleR_scaleR[THEN sym] scaleR_right.setsum [symmetric] and z_def using *
-    by(auto simp add: setsum_negf vector_smult_lneg mult_right.setsum[THEN sym])
-  ultimately show ?thesis apply(rule_tac x="{v\<in>c. u v \<le> 0}" in exI, rule_tac x="{v\<in>c. u v > 0}" in exI) by auto
-qed
-
-lemma radon: assumes "affine_dependent c"
-  obtains m p where "m\<subseteq>c" "p\<subseteq>c" "m \<inter> p = {}" "(convex hull m) \<inter> (convex hull p) \<noteq> {}"
-proof- from assms[unfolded affine_dependent_explicit] guess s .. then guess u ..
-  hence *:"finite s" "affine_dependent s" and s:"s \<subseteq> c" unfolding affine_dependent_explicit by auto
-  from radon_partition[OF *] guess m .. then guess p ..
-  thus ?thesis apply(rule_tac that[of p m]) using s by auto qed
-
-subsection {* Helly's theorem. *}
-
-lemma helly_induct: fixes f::"(real^'n::finite) set set"
-  assumes "f hassize n" "n \<ge> CARD('n) + 1"
-  "\<forall>s\<in>f. convex s" "\<forall>t\<subseteq>f. card t = CARD('n) + 1 \<longrightarrow> \<Inter> t \<noteq> {}"
-  shows "\<Inter> f \<noteq> {}"
-  using assms unfolding hassize_def apply(erule_tac conjE) proof(induct n arbitrary: f)
-case (Suc n)
-show "\<Inter> f \<noteq> {}" apply(cases "n = CARD('n)") apply(rule Suc(4)[rule_format])
-  unfolding card_Diff_singleton_if[OF Suc(5)] and Suc(6) proof-
-  assume ng:"n \<noteq> CARD('n)" hence "\<exists>X. \<forall>s\<in>f. X s \<in> \<Inter>(f - {s})" apply(rule_tac bchoice) unfolding ex_in_conv
-    apply(rule, rule Suc(1)[rule_format])  unfolding card_Diff_singleton_if[OF Suc(5)] and Suc(6)
-    defer apply(rule Suc(3)[rule_format]) defer apply(rule Suc(4)[rule_format]) using Suc(2,5) by auto
-  then obtain X where X:"\<forall>s\<in>f. X s \<in> \<Inter>(f - {s})" by auto
-  show ?thesis proof(cases "inj_on X f")
-    case False then obtain s t where st:"s\<noteq>t" "s\<in>f" "t\<in>f" "X s = X t" unfolding inj_on_def by auto
-    hence *:"\<Inter> f = \<Inter> (f - {s}) \<inter> \<Inter> (f - {t})" by auto
-    show ?thesis unfolding * unfolding ex_in_conv[THEN sym] apply(rule_tac x="X s" in exI)
-      apply(rule, rule X[rule_format]) using X st by auto
-  next case True then obtain m p where mp:"m \<inter> p = {}" "m \<union> p = X ` f" "convex hull m \<inter> convex hull p \<noteq> {}"
-      using radon_partition[of "X ` f"] and affine_dependent_biggerset[of "X ` f"]
-      unfolding card_image[OF True] and Suc(6) using Suc(2,5) and ng by auto
-    have "m \<subseteq> X ` f" "p \<subseteq> X ` f" using mp(2) by auto
-    then obtain g h where gh:"m = X ` g" "p = X ` h" "g \<subseteq> f" "h \<subseteq> f" unfolding subset_image_iff by auto 
-    hence "f \<union> (g \<union> h) = f" by auto
-    hence f:"f = g \<union> h" using inj_on_Un_image_eq_iff[of X f "g \<union> h"] and True
-      unfolding mp(2)[unfolded image_Un[THEN sym] gh] by auto
-    have *:"g \<inter> h = {}" using mp(1) unfolding gh using inj_on_image_Int[OF True gh(3,4)] by auto
-    have "convex hull (X ` h) \<subseteq> \<Inter> g" "convex hull (X ` g) \<subseteq> \<Inter> h"
-      apply(rule_tac [!] hull_minimal) using Suc(3) gh(3-4)  unfolding mem_def unfolding subset_eq
-      apply(rule_tac [2] convex_Inter, rule_tac [4] convex_Inter) apply rule prefer 3 apply rule proof-
-      fix x assume "x\<in>X ` g" then guess y unfolding image_iff ..
-      thus "x\<in>\<Inter>h" using X[THEN bspec[where x=y]] using * f by auto next
-      fix x assume "x\<in>X ` h" then guess y unfolding image_iff ..
-      thus "x\<in>\<Inter>g" using X[THEN bspec[where x=y]] using * f by auto
-    qed(auto)
-    thus ?thesis unfolding f using mp(3)[unfolded gh] by blast qed
-qed(insert dimindex_ge_1, auto) qed(auto)
-
-lemma helly: fixes f::"(real^'n::finite) set set"
-  assumes "finite f" "card f \<ge> CARD('n) + 1" "\<forall>s\<in>f. convex s"
-          "\<forall>t\<subseteq>f. card t = CARD('n) + 1 \<longrightarrow> \<Inter> t \<noteq> {}"
-  shows "\<Inter> f \<noteq>{}"
-  apply(rule helly_induct) unfolding hassize_def using assms by auto
-
-subsection {* Convex hull is "preserved" by a linear function. *}
-
-lemma convex_hull_linear_image:
-  assumes "bounded_linear f"
-  shows "f ` (convex hull s) = convex hull (f ` s)"
-  apply rule unfolding subset_eq ball_simps apply(rule_tac[!] hull_induct, rule hull_inc) prefer 3  
-  apply(erule imageE)apply(rule_tac x=xa in image_eqI) apply assumption
-  apply(rule hull_subset[unfolded subset_eq, rule_format]) apply assumption
-proof-
-  interpret f: bounded_linear f by fact
-  show "convex {x. f x \<in> convex hull f ` s}" 
-  unfolding convex_def by(auto simp add: f.scaleR f.add convex_convex_hull[unfolded convex_def, rule_format]) next
-  interpret f: bounded_linear f by fact
-  show "convex {x. x \<in> f ` (convex hull s)}" using  convex_convex_hull[unfolded convex_def, of s] 
-    unfolding convex_def by (auto simp add: f.scaleR [symmetric] f.add [symmetric])
-qed auto
-
-lemma in_convex_hull_linear_image:
-  assumes "bounded_linear f" "x \<in> convex hull s"
-  shows "(f x) \<in> convex hull (f ` s)"
-using convex_hull_linear_image[OF assms(1)] assms(2) by auto
-
-subsection {* Homeomorphism of all convex compact sets with nonempty interior. *}
-
-lemma compact_frontier_line_lemma:
-  fixes s :: "(real ^ _) set"
-  assumes "compact s" "0 \<in> s" "x \<noteq> 0" 
-  obtains u where "0 \<le> u" "(u *\<^sub>R x) \<in> frontier s" "\<forall>v>u. (v *\<^sub>R x) \<notin> s"
-proof-
-  obtain b where b:"b>0" "\<forall>x\<in>s. norm x \<le> b" using compact_imp_bounded[OF assms(1), unfolded bounded_pos] by auto
-  let ?A = "{y. \<exists>u. 0 \<le> u \<and> u \<le> b / norm(x) \<and> (y = u *\<^sub>R x)}"
-  have A:"?A = (\<lambda>u. dest_vec1 u *\<^sub>R x) ` {0 .. vec1 (b / norm x)}"
-    unfolding image_image[of "\<lambda>u. u *\<^sub>R x" "\<lambda>x. dest_vec1 x", THEN sym]
-    unfolding dest_vec1_inverval vec1_dest_vec1 by auto
-  have "compact ?A" unfolding A apply(rule compact_continuous_image, rule continuous_at_imp_continuous_on)
-    apply(rule, rule continuous_vmul)
-    apply (rule continuous_dest_vec1)
-    apply(rule continuous_at_id) by(rule compact_interval)
-  moreover have "{y. \<exists>u\<ge>0. u \<le> b / norm x \<and> y = u *\<^sub>R x} \<inter> s \<noteq> {}" apply(rule not_disjointI[OF _ assms(2)])
-    unfolding mem_Collect_eq using `b>0` assms(3) by(auto intro!: divide_nonneg_pos)
-  ultimately obtain u y where obt: "u\<ge>0" "u \<le> b / norm x" "y = u *\<^sub>R x"
-    "y\<in>?A" "y\<in>s" "\<forall>z\<in>?A \<inter> s. dist 0 z \<le> dist 0 y" using distance_attains_sup[OF compact_inter[OF _ assms(1), of ?A], of 0] by auto
-
-  have "norm x > 0" using assms(3)[unfolded zero_less_norm_iff[THEN sym]] by auto
-  { fix v assume as:"v > u" "v *\<^sub>R x \<in> s"
-    hence "v \<le> b / norm x" using b(2)[rule_format, OF as(2)] 
-      using `u\<ge>0` unfolding pos_le_divide_eq[OF `norm x > 0`] by auto
-    hence "norm (v *\<^sub>R x) \<le> norm y" apply(rule_tac obt(6)[rule_format, unfolded dist_0_norm]) apply(rule IntI) defer 
-      apply(rule as(2)) unfolding mem_Collect_eq apply(rule_tac x=v in exI) 
-      using as(1) `u\<ge>0` by(auto simp add:field_simps) 
-    hence False unfolding obt(3) using `u\<ge>0` `norm x > 0` `v>u` by(auto simp add:field_simps)
-  } note u_max = this
-
-  have "u *\<^sub>R x \<in> frontier s" unfolding frontier_straddle apply(rule,rule,rule) apply(rule_tac x="u *\<^sub>R x" in bexI) unfolding obt(3)[THEN sym]
-    prefer 3 apply(rule_tac x="(u + (e / 2) / norm x) *\<^sub>R x" in exI) apply(rule, rule) proof-
-    fix e  assume "0 < e" and as:"(u + e / 2 / norm x) *\<^sub>R x \<in> s"
-    hence "u + e / 2 / norm x > u" using`norm x > 0` by(auto simp del:zero_less_norm_iff intro!: divide_pos_pos)
-    thus False using u_max[OF _ as] by auto
-  qed(insert `y\<in>s`, auto simp add: dist_norm scaleR_left_distrib obt(3))
-  thus ?thesis apply(rule_tac that[of u]) apply(rule obt(1), assumption)
-    apply(rule,rule,rule ccontr) apply(rule u_max) by auto qed
-
-lemma starlike_compact_projective:
-  assumes "compact s" "cball (0::real^'n::finite) 1 \<subseteq> s "
-  "\<forall>x\<in>s. \<forall>u. 0 \<le> u \<and> u < 1 \<longrightarrow> (u *\<^sub>R x) \<in> (s - frontier s )"
-  shows "s homeomorphic (cball (0::real^'n::finite) 1)"
-proof-
-  have fs:"frontier s \<subseteq> s" apply(rule frontier_subset_closed) using compact_imp_closed[OF assms(1)] by simp
-  def pi \<equiv> "\<lambda>x::real^'n. inverse (norm x) *\<^sub>R x"
-  have "0 \<notin> frontier s" unfolding frontier_straddle apply(rule ccontr) unfolding not_not apply(erule_tac x=1 in allE)
-    using assms(2)[unfolded subset_eq Ball_def mem_cball] by auto
-  have injpi:"\<And>x y. pi x = pi y \<and> norm x = norm y \<longleftrightarrow> x = y" unfolding pi_def by auto
-
-  have contpi:"continuous_on (UNIV - {0}) pi" apply(rule continuous_at_imp_continuous_on)
-    apply rule unfolding pi_def
-    apply (rule continuous_mul)
-    apply (rule continuous_at_inv[unfolded o_def])
-    apply (rule continuous_at_norm)
-    apply simp
-    apply (rule continuous_at_id)
-    done
-  def sphere \<equiv> "{x::real^'n. norm x = 1}"
-  have pi:"\<And>x. x \<noteq> 0 \<Longrightarrow> pi x \<in> sphere" "\<And>x u. u>0 \<Longrightarrow> pi (u *\<^sub>R x) = pi x" unfolding pi_def sphere_def by auto
-
-  have "0\<in>s" using assms(2) and centre_in_cball[of 0 1] by auto
-  have front_smul:"\<forall>x\<in>frontier s. \<forall>u\<ge>0. u *\<^sub>R x \<in> s \<longleftrightarrow> u \<le> 1" proof(rule,rule,rule)
-    fix x u assume x:"x\<in>frontier s" and "(0::real)\<le>u"
-    hence "x\<noteq>0" using `0\<notin>frontier s` by auto
-    obtain v where v:"0 \<le> v" "v *\<^sub>R x \<in> frontier s" "\<forall>w>v. w *\<^sub>R x \<notin> s"
-      using compact_frontier_line_lemma[OF assms(1) `0\<in>s` `x\<noteq>0`] by auto
-    have "v=1" apply(rule ccontr) unfolding neq_iff apply(erule disjE) proof-
-      assume "v<1" thus False using v(3)[THEN spec[where x=1]] using x and fs by auto next
-      assume "v>1" thus False using assms(3)[THEN bspec[where x="v *\<^sub>R x"], THEN spec[where x="inverse v"]]
-        using v and x and fs unfolding inverse_less_1_iff by auto qed
-    show "u *\<^sub>R x \<in> s \<longleftrightarrow> u \<le> 1" apply rule  using v(3)[unfolded `v=1`, THEN spec[where x=u]] proof-
-      assume "u\<le>1" thus "u *\<^sub>R x \<in> s" apply(cases "u=1")
-        using assms(3)[THEN bspec[where x=x], THEN spec[where x=u]] using `0\<le>u` and x and fs by auto qed auto qed
-
-  have "\<exists>surf. homeomorphism (frontier s) sphere pi surf"
-    apply(rule homeomorphism_compact) apply(rule compact_frontier[OF assms(1)])
-    apply(rule continuous_on_subset[OF contpi]) defer apply(rule set_ext,rule) 
-    unfolding inj_on_def prefer 3 apply(rule,rule,rule)
-  proof- fix x assume "x\<in>pi ` frontier s" then obtain y where "y\<in>frontier s" "x = pi y" by auto
-    thus "x \<in> sphere" using pi(1)[of y] and `0 \<notin> frontier s` by auto
-  next fix x assume "x\<in>sphere" hence "norm x = 1" "x\<noteq>0" unfolding sphere_def by auto
-    then obtain u where "0 \<le> u" "u *\<^sub>R x \<in> frontier s" "\<forall>v>u. v *\<^sub>R x \<notin> s"
-      using compact_frontier_line_lemma[OF assms(1) `0\<in>s`, of x] by auto
-    thus "x \<in> pi ` frontier s" unfolding image_iff le_less pi_def apply(rule_tac x="u *\<^sub>R x" in bexI) using `norm x = 1` `0\<notin>frontier s` by auto
-  next fix x y assume as:"x \<in> frontier s" "y \<in> frontier s" "pi x = pi y"
-    hence xys:"x\<in>s" "y\<in>s" using fs by auto
-    from as(1,2) have nor:"norm x \<noteq> 0" "norm y \<noteq> 0" using `0\<notin>frontier s` by auto 
-    from nor have x:"x = norm x *\<^sub>R ((inverse (norm y)) *\<^sub>R y)" unfolding as(3)[unfolded pi_def, THEN sym] by auto 
-    from nor have y:"y = norm y *\<^sub>R ((inverse (norm x)) *\<^sub>R x)" unfolding as(3)[unfolded pi_def] by auto 
-    have "0 \<le> norm y * inverse (norm x)" "0 \<le> norm x * inverse (norm y)"
-      unfolding divide_inverse[THEN sym] apply(rule_tac[!] divide_nonneg_pos) using nor by auto
-    hence "norm x = norm y" apply(rule_tac ccontr) unfolding neq_iff
-      using x y and front_smul[THEN bspec, OF as(1), THEN spec[where x="norm y * (inverse (norm x))"]]
-      using front_smul[THEN bspec, OF as(2), THEN spec[where x="norm x * (inverse (norm y))"]]
-      using xys nor by(auto simp add:field_simps divide_le_eq_1 divide_inverse[THEN sym])
-    thus "x = y" apply(subst injpi[THEN sym]) using as(3) by auto
-  qed(insert `0 \<notin> frontier s`, auto)
-  then obtain surf where surf:"\<forall>x\<in>frontier s. surf (pi x) = x"  "pi ` frontier s = sphere" "continuous_on (frontier s) pi"
-    "\<forall>y\<in>sphere. pi (surf y) = y" "surf ` sphere = frontier s" "continuous_on sphere surf" unfolding homeomorphism_def by auto
-  
-  have cont_surfpi:"continuous_on (UNIV -  {0}) (surf \<circ> pi)" apply(rule continuous_on_compose, rule contpi)
-    apply(rule continuous_on_subset[of sphere], rule surf(6)) using pi(1) by auto
-
-  { fix x assume as:"x \<in> cball (0::real^'n) 1"
-    have "norm x *\<^sub>R surf (pi x) \<in> s" proof(cases "x=0 \<or> norm x = 1") 
-      case False hence "pi x \<in> sphere" "norm x < 1" using pi(1)[of x] as by(auto simp add: dist_norm)
-      thus ?thesis apply(rule_tac assms(3)[rule_format, THEN DiffD1])
-        apply(rule_tac fs[unfolded subset_eq, rule_format])
-        unfolding surf(5)[THEN sym] by auto
-    next case True thus ?thesis apply rule defer unfolding pi_def apply(rule fs[unfolded subset_eq, rule_format])
-        unfolding  surf(5)[unfolded sphere_def, THEN sym] using `0\<in>s` by auto qed } note hom = this
-
-  { fix x assume "x\<in>s"
-    hence "x \<in> (\<lambda>x. norm x *\<^sub>R surf (pi x)) ` cball 0 1" proof(cases "x=0")
-      case True show ?thesis unfolding image_iff True apply(rule_tac x=0 in bexI) by auto
-    next let ?a = "inverse (norm (surf (pi x)))"
-      case False hence invn:"inverse (norm x) \<noteq> 0" by auto
-      from False have pix:"pi x\<in>sphere" using pi(1) by auto
-      hence "pi (surf (pi x)) = pi x" apply(rule_tac surf(4)[rule_format]) by assumption
-      hence **:"norm x *\<^sub>R (?a *\<^sub>R surf (pi x)) = x" apply(rule_tac scaleR_left_imp_eq[OF invn]) unfolding pi_def using invn by auto
-      hence *:"?a * norm x > 0" and"?a > 0" "?a \<noteq> 0" using surf(5) `0\<notin>frontier s` apply -
-        apply(rule_tac mult_pos_pos) using False[unfolded zero_less_norm_iff[THEN sym]] by auto
-      have "norm (surf (pi x)) \<noteq> 0" using ** False by auto
-      hence "norm x = norm ((?a * norm x) *\<^sub>R surf (pi x))"
-        unfolding norm_scaleR abs_mult abs_norm_cancel abs_of_pos[OF `?a > 0`] by auto
-      moreover have "pi x = pi ((inverse (norm (surf (pi x))) * norm x) *\<^sub>R surf (pi x))" 
-        unfolding pi(2)[OF *] surf(4)[rule_format, OF pix] ..
-      moreover have "surf (pi x) \<in> frontier s" using surf(5) pix by auto
-      hence "dist 0 (inverse (norm (surf (pi x))) *\<^sub>R x) \<le> 1" unfolding dist_norm
-        using ** and * using front_smul[THEN bspec[where x="surf (pi x)"], THEN spec[where x="norm x * ?a"]]
-        using False `x\<in>s` by(auto simp add:field_simps)
-      ultimately show ?thesis unfolding image_iff apply(rule_tac x="inverse (norm (surf(pi x))) *\<^sub>R x" in bexI)
-        apply(subst injpi[THEN sym]) unfolding abs_mult abs_norm_cancel abs_of_pos[OF `?a > 0`]
-        unfolding pi(2)[OF `?a > 0`] by auto
-    qed } note hom2 = this
-
-  show ?thesis apply(subst homeomorphic_sym) apply(rule homeomorphic_compact[where f="\<lambda>x. norm x *\<^sub>R surf (pi x)"])
-    apply(rule compact_cball) defer apply(rule set_ext, rule, erule imageE, drule hom)
-    prefer 4 apply(rule continuous_at_imp_continuous_on, rule) apply(rule_tac [3] hom2) proof-
-    fix x::"real^'n" assume as:"x \<in> cball 0 1"
-    thus "continuous (at x) (\<lambda>x. norm x *\<^sub>R surf (pi x))" proof(cases "x=0")
-      case False thus ?thesis apply(rule_tac continuous_mul, rule_tac continuous_at_norm)
-        using cont_surfpi unfolding continuous_on_eq_continuous_at[OF open_delete[OF open_UNIV]] o_def by auto
-    next guess a using UNIV_witness[where 'a = 'n] ..
-      obtain B where B:"\<forall>x\<in>s. norm x \<le> B" using compact_imp_bounded[OF assms(1)] unfolding bounded_iff by auto
-      hence "B > 0" using assms(2) unfolding subset_eq apply(erule_tac x="basis a" in ballE) defer apply(erule_tac x="basis a" in ballE)
-        unfolding Ball_def mem_cball dist_norm by (auto simp add: norm_basis[unfolded One_nat_def])
-      case True show ?thesis unfolding True continuous_at Lim_at apply(rule,rule) apply(rule_tac x="e / B" in exI)
-        apply(rule) apply(rule divide_pos_pos) prefer 3 apply(rule,rule,erule conjE)
-        unfolding norm_0 scaleR_zero_left dist_norm diff_0_right norm_scaleR abs_norm_cancel proof-
-        fix e and x::"real^'n" assume as:"norm x < e / B" "0 < norm x" "0<e"
-        hence "surf (pi x) \<in> frontier s" using pi(1)[of x] unfolding surf(5)[THEN sym] by auto
-        hence "norm (surf (pi x)) \<le> B" using B fs by auto
-        hence "norm x * norm (surf (pi x)) \<le> norm x * B" using as(2) by auto
-        also have "\<dots> < e / B * B" apply(rule mult_strict_right_mono) using as(1) `B>0` by auto
-        also have "\<dots> = e" using `B>0` by auto
-        finally show "norm x * norm (surf (pi x)) < e" by assumption
-      qed(insert `B>0`, auto) qed
-  next { fix x assume as:"surf (pi x) = 0"
-      have "x = 0" proof(rule ccontr)
-        assume "x\<noteq>0" hence "pi x \<in> sphere" using pi(1) by auto
-        hence "surf (pi x) \<in> frontier s" using surf(5) by auto
-        thus False using `0\<notin>frontier s` unfolding as by simp qed
-    } note surf_0 = this
-    show "inj_on (\<lambda>x. norm x *\<^sub>R surf (pi x)) (cball 0 1)" unfolding inj_on_def proof(rule,rule,rule)
-      fix x y assume as:"x \<in> cball 0 1" "y \<in> cball 0 1" "norm x *\<^sub>R surf (pi x) = norm y *\<^sub>R surf (pi y)"
-      thus "x=y" proof(cases "x=0 \<or> y=0") 
-        case True thus ?thesis using as by(auto elim: surf_0) next
-        case False
-        hence "pi (surf (pi x)) = pi (surf (pi y))" using as(3)
-          using pi(2)[of "norm x" "surf (pi x)"] pi(2)[of "norm y" "surf (pi y)"] by auto
-        moreover have "pi x \<in> sphere" "pi y \<in> sphere" using pi(1) False by auto
-        ultimately have *:"pi x = pi y" using surf(4)[THEN bspec[where x="pi x"]] surf(4)[THEN bspec[where x="pi y"]] by auto 
-        moreover have "norm x = norm y" using as(3)[unfolded *] using False by(auto dest:surf_0)
-        ultimately show ?thesis using injpi by auto qed qed
-  qed auto qed
-
-lemma homeomorphic_convex_compact_lemma: fixes s::"(real^'n::finite) set"
-  assumes "convex s" "compact s" "cball 0 1 \<subseteq> s"
-  shows "s homeomorphic (cball (0::real^'n) 1)"
-  apply(rule starlike_compact_projective[OF assms(2-3)]) proof(rule,rule,rule,erule conjE)
-  fix x u assume as:"x \<in> s" "0 \<le> u" "u < (1::real)"
-  hence "u *\<^sub>R x \<in> interior s" unfolding interior_def mem_Collect_eq
-    apply(rule_tac x="ball (u *\<^sub>R x) (1 - u)" in exI) apply(rule, rule open_ball)
-    unfolding centre_in_ball apply rule defer apply(rule) unfolding mem_ball proof-
-    fix y assume "dist (u *\<^sub>R x) y < 1 - u"
-    hence "inverse (1 - u) *\<^sub>R (y - u *\<^sub>R x) \<in> s"
-      using assms(3) apply(erule_tac subsetD) unfolding mem_cball dist_commute dist_norm
-      unfolding group_add_class.diff_0 group_add_class.diff_0_right norm_minus_cancel norm_scaleR
-      apply (rule mult_left_le_imp_le[of "1 - u"])
-      unfolding class_semiring.mul_a using `u<1` by auto
-    thus "y \<in> s" using assms(1)[unfolded convex_def, rule_format, of "inverse(1 - u) *\<^sub>R (y - u *\<^sub>R x)" x "1 - u" u]
-      using as unfolding scaleR_scaleR by auto qed auto
-  thus "u *\<^sub>R x \<in> s - frontier s" using frontier_def and interior_subset by auto qed
-
-lemma homeomorphic_convex_compact_cball: fixes e::real and s::"(real^'n::finite) set"
-  assumes "convex s" "compact s" "interior s \<noteq> {}" "0 < e"
-  shows "s homeomorphic (cball (b::real^'n::finite) e)"
-proof- obtain a where "a\<in>interior s" using assms(3) by auto
-  then obtain d where "d>0" and d:"cball a d \<subseteq> s" unfolding mem_interior_cball by auto
-  let ?d = "inverse d" and ?n = "0::real^'n"
-  have "cball ?n 1 \<subseteq> (\<lambda>x. inverse d *\<^sub>R (x - a)) ` s"
-    apply(rule, rule_tac x="d *\<^sub>R x + a" in image_eqI) defer
-    apply(rule d[unfolded subset_eq, rule_format]) using `d>0` unfolding mem_cball dist_norm
-    by(auto simp add: mult_right_le_one_le)
-  hence "(\<lambda>x. inverse d *\<^sub>R (x - a)) ` s homeomorphic cball ?n 1"
-    using homeomorphic_convex_compact_lemma[of "(\<lambda>x. ?d *\<^sub>R -a + ?d *\<^sub>R x) ` s", OF convex_affinity compact_affinity]
-    using assms(1,2) by(auto simp add: uminus_add_conv_diff scaleR_right_diff_distrib)
-  thus ?thesis apply(rule_tac homeomorphic_trans[OF _ homeomorphic_balls(2)[of 1 _ ?n]])
-    apply(rule homeomorphic_trans[OF homeomorphic_affinity[of "?d" s "?d *\<^sub>R -a"]])
-    using `d>0` `e>0` by(auto simp add: uminus_add_conv_diff scaleR_right_diff_distrib) qed
-
-lemma homeomorphic_convex_compact: fixes s::"(real^'n::finite) set" and t::"(real^'n) set"
-  assumes "convex s" "compact s" "interior s \<noteq> {}"
-          "convex t" "compact t" "interior t \<noteq> {}"
-  shows "s homeomorphic t"
-  using assms by(meson zero_less_one homeomorphic_trans homeomorphic_convex_compact_cball homeomorphic_sym)
-
-subsection {* Epigraphs of convex functions. *}
-
-definition "epigraph s (f::real^'n \<Rightarrow> real) = {xy. fstcart xy \<in> s \<and> f(fstcart xy) \<le> dest_vec1 (sndcart xy)}"
-
-lemma mem_epigraph: "(pastecart x (vec1 y)) \<in> epigraph s f \<longleftrightarrow> x \<in> s \<and> f x \<le> y" unfolding epigraph_def by auto
-
-lemma convex_epigraph: 
-  "convex(epigraph s f) \<longleftrightarrow> convex_on s f \<and> convex s"
-  unfolding convex_def convex_on_def unfolding Ball_def forall_pastecart epigraph_def
-  unfolding mem_Collect_eq fstcart_pastecart sndcart_pastecart sndcart_add sndcart_cmul [where 'a=real, unfolded smult_conv_scaleR] fstcart_add fstcart_cmul [where 'a=real, unfolded smult_conv_scaleR]
-  unfolding Ball_def[symmetric] unfolding dest_vec1_add dest_vec1_cmul [where 'a=real, unfolded smult_conv_scaleR]
-  apply(subst forall_dest_vec1[THEN sym])+ by(meson real_le_refl real_le_trans add_mono mult_left_mono) 
-
-lemma convex_epigraphI: assumes "convex_on s f" "convex s"
-  shows "convex(epigraph s f)" using assms unfolding convex_epigraph by auto
-
-lemma convex_epigraph_convex: "convex s \<Longrightarrow> (convex_on s f \<longleftrightarrow> convex(epigraph s f))"
-  using convex_epigraph by auto
-
-subsection {* Use this to derive general bound property of convex function. *}
-
-lemma forall_of_pastecart:
-  "(\<forall>p. P (\<lambda>x. fstcart (p x)) (\<lambda>x. sndcart (p x))) \<longleftrightarrow> (\<forall>x y. P x y)" apply meson
-  apply(erule_tac x="\<lambda>a. pastecart (x a) (y a)" in allE) unfolding o_def by auto
-
-lemma forall_of_pastecart':
-  "(\<forall>p. P (fstcart p) (sndcart p)) \<longleftrightarrow> (\<forall>x y. P x y)" apply meson
-  apply(erule_tac x="pastecart x y" in allE) unfolding o_def by auto
-
-lemma forall_of_dest_vec1: "(\<forall>v. P (\<lambda>x. dest_vec1 (v x))) \<longleftrightarrow> (\<forall>x. P x)"
-  apply rule apply rule apply(erule_tac x="(vec1 \<circ> x)" in allE) unfolding o_def vec1_dest_vec1 by auto 
-
-lemma forall_of_dest_vec1': "(\<forall>v. P (dest_vec1 v)) \<longleftrightarrow> (\<forall>x. P x)"
-  apply rule apply rule apply(erule_tac x="(vec1 x)" in allE) defer apply rule 
-  apply(erule_tac x="dest_vec1 v" in allE) unfolding o_def vec1_dest_vec1 by auto
-
-lemma convex_on:
-  fixes s :: "(real ^ _) set"
-  assumes "convex s"
-  shows "convex_on s f \<longleftrightarrow> (\<forall>k u x. (\<forall>i\<in>{1..k::nat}. 0 \<le> u i \<and> x i \<in> s) \<and> setsum u {1..k} = 1 \<longrightarrow>
-   f (setsum (\<lambda>i. u i *\<^sub>R x i) {1..k} ) \<le> setsum (\<lambda>i. u i * f(x i)) {1..k} ) "
-  unfolding convex_epigraph_convex[OF assms] convex epigraph_def Ball_def mem_Collect_eq
-  unfolding sndcart_setsum[OF finite_atLeastAtMost] fstcart_setsum[OF finite_atLeastAtMost] dest_vec1_setsum[OF finite_atLeastAtMost]
-  unfolding fstcart_pastecart sndcart_pastecart sndcart_add sndcart_cmul [where 'a=real, unfolded smult_conv_scaleR] fstcart_add fstcart_cmul [where 'a=real, unfolded smult_conv_scaleR]
-  unfolding dest_vec1_add dest_vec1_cmul [where 'a=real, unfolded smult_conv_scaleR] apply(subst forall_of_pastecart)+ apply(subst forall_of_dest_vec1)+ apply rule
-  using assms[unfolded convex] apply simp apply(rule,rule,rule)
-  apply(erule_tac x=k in allE, erule_tac x=u in allE, erule_tac x=x in allE) apply rule apply rule apply rule defer
-  apply(rule_tac j="\<Sum>i = 1..k. u i * f (x i)" in real_le_trans)
-  defer apply(rule setsum_mono) apply(erule conjE)+ apply(erule_tac x=i in allE)apply(rule mult_left_mono)
-  using assms[unfolded convex] by auto
-
-subsection {* Convexity of general and special intervals. *}
-
-lemma is_interval_convex:
-  fixes s :: "(real ^ _) set"
-  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 \<in> s" "y \<in> s" "0 \<le> u" "0 \<le> v" "u + v = (1::real)"
-  hence *:"u = 1 - v" "1 - v \<ge> 0" and **:"v = 1 - u" "1 - u \<ge> 0" by auto
-  { fix a b assume "\<not> b \<le> u * a + v * b"
-    hence "u * a < (1 - v) * b" unfolding not_le using as(4) by(auto simp add: field_simps)
-    hence "a < b" unfolding * using as(4) *(2) apply(rule_tac mult_left_less_imp_less[of "1 - v"]) by(auto simp add: field_simps)
-    hence "a \<le> u * a + v * b" unfolding * using as(4) by (auto simp add: field_simps intro!:mult_right_mono)
-  } moreover
-  { fix a b assume "\<not> u * a + v * b \<le> a"
-    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 \<le> b" unfolding ** using **(2) as(3) by(auto simp add: field_simps intro!:mult_right_mono) }
-  ultimately show "u *\<^sub>R x + v *\<^sub>R y \<in> 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:
-  fixes s :: "(real ^ _) set"
-  shows "is_interval s \<Longrightarrow> connected s"
-  using is_interval_convex convex_connected by auto
-
-lemma convex_interval: "convex {a .. b}" "convex {a<..<b::real^'n::finite}"
-  apply(rule_tac[!] is_interval_convex) using is_interval_interval by auto
-
-subsection {* On @{text "real^1"}, @{text "is_interval"}, @{text "convex"} and @{text "connected"} are all equivalent. *}
-
-lemma is_interval_1:
-  "is_interval s \<longleftrightarrow> (\<forall>a\<in>s. \<forall>b\<in>s. \<forall> x. dest_vec1 a \<le> dest_vec1 x \<and> dest_vec1 x \<le> dest_vec1 b \<longrightarrow> x \<in> s)"
-  unfolding is_interval_def dest_vec1_def forall_1 by auto
-
-lemma is_interval_connected_1: "is_interval s \<longleftrightarrow> 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 \<in> s" "b \<in> s" "dest_vec1 a \<le> dest_vec1 x" "dest_vec1 x \<le> dest_vec1 b" "x\<notin>s"
-  hence *:"dest_vec1 a < dest_vec1 x" "dest_vec1 x < dest_vec1 b" apply(rule_tac [!] ccontr) unfolding not_less by auto
-  let ?halfl = "{z. inner (basis 1) z < dest_vec1 x} " and ?halfr = "{z. inner (basis 1) z > dest_vec1 x} "
-  { fix y assume "y \<in> s" have "y \<in> ?halfr \<union> ?halfl" apply(rule ccontr)
-    using as(6) `y\<in>s` by (auto simp add: inner_vector_def dest_vec1_eq [unfolded dest_vec1_def] dest_vec1_def) }
-  moreover have "a\<in>?halfl" "b\<in>?halfr" using * by (auto simp add: inner_vector_def dest_vec1_def)
-  hence "?halfl \<inter> s \<noteq> {}" "?halfr \<inter> s \<noteq> {}"  using as(2-3) by auto
-  ultimately show False apply(rule_tac notE[OF as(1)[unfolded connected_def]])
-    apply(rule_tac x="?halfl" in exI, rule_tac x="?halfr" in exI) 
-    apply(rule, rule open_halfspace_lt, rule, rule open_halfspace_gt) apply(rule, rule, rule ccontr)
-    by(auto simp add: basis_component field_simps) qed 
-
-lemma is_interval_convex_1:
-  "is_interval s \<longleftrightarrow> convex (s::(real^1) set)" 
-  using is_interval_convex convex_connected is_interval_connected_1 by auto
-
-lemma convex_connected_1:
-  "connected s \<longleftrightarrow> convex (s::(real^1) set)" 
-  using is_interval_convex convex_connected is_interval_connected_1 by auto
-
-subsection {* Another intermediate value theorem formulation. *}
-
-lemma ivt_increasing_component_on_1: fixes f::"real^1 \<Rightarrow> real^'n::finite"
-  assumes "dest_vec1 a \<le> dest_vec1 b" "continuous_on {a .. b} f" "(f a)$k \<le> y" "y \<le> (f b)$k"
-  shows "\<exists>x\<in>{a..b}. (f x)$k = y"
-proof- have "f a \<in> f ` {a..b}" "f b \<in> f ` {a..b}" apply(rule_tac[!] imageI) 
-    using assms(1) by(auto simp add: vector_less_eq_def dest_vec1_def)
-  thus ?thesis using connected_ivt_component[of "f ` {a..b}" "f a" "f b" k y]
-    using connected_continuous_image[OF assms(2) convex_connected[OF convex_interval(1)]]
-    using assms by(auto intro!: imageI) qed
-
-lemma ivt_increasing_component_1: fixes f::"real^1 \<Rightarrow> real^'n::finite"
-  assumes "dest_vec1 a \<le> dest_vec1 b"
-  "\<forall>x\<in>{a .. b}. continuous (at x) f" "f a$k \<le> y" "y \<le> f b$k"
-  shows "\<exists>x\<in>{a..b}. (f x)$k = y"
-  apply(rule ivt_increasing_component_on_1) using assms using continuous_at_imp_continuous_on by auto
-
-lemma ivt_decreasing_component_on_1: fixes f::"real^1 \<Rightarrow> real^'n::finite"
-  assumes "dest_vec1 a \<le> dest_vec1 b" "continuous_on {a .. b} f" "(f b)$k \<le> y" "y \<le> (f a)$k"
-  shows "\<exists>x\<in>{a..b}. (f x)$k = y"
-  apply(subst neg_equal_iff_equal[THEN sym]) unfolding vector_uminus_component[THEN sym]
-  apply(rule ivt_increasing_component_on_1) using assms using continuous_on_neg
-  by(auto simp add:vector_uminus_component)
-
-lemma ivt_decreasing_component_1: fixes f::"real^1 \<Rightarrow> real^'n::finite"
-  assumes "dest_vec1 a \<le> dest_vec1 b" "\<forall>x\<in>{a .. b}. continuous (at x) f" "f b$k \<le> y" "y \<le> f a$k"
-  shows "\<exists>x\<in>{a..b}. (f x)$k = y"
-  apply(rule ivt_decreasing_component_on_1) using assms using continuous_at_imp_continuous_on by auto
-
-subsection {* A bound within a convex hull, and so an interval. *}
-
-lemma convex_on_convex_hull_bound:
-  fixes s :: "(real ^ _) set"
-  assumes "convex_on (convex hull s) f" "\<forall>x\<in>s. f x \<le> b"
-  shows "\<forall>x\<in> convex hull s. f x \<le> b" proof
-  fix x assume "x\<in>convex hull s"
-  then obtain k u v where obt:"\<forall>i\<in>{1..k::nat}. 0 \<le> u i \<and> v i \<in> s" "setsum u {1..k} = 1" "(\<Sum>i = 1..k. u i *\<^sub>R v i) = x"
-    unfolding convex_hull_indexed mem_Collect_eq by auto
-  have "(\<Sum>i = 1..k. u i * f (v i)) \<le> b" using setsum_mono[of "{1..k}" "\<lambda>i. u i * f (v i)" "\<lambda>i. u i * b"]
-    unfolding setsum_left_distrib[THEN sym] obt(2) mult_1 apply(drule_tac meta_mp) apply(rule mult_left_mono)
-    using assms(2) obt(1) by auto
-  thus "f x \<le> b" using assms(1)[unfolded convex_on[OF convex_convex_hull], rule_format, of k u v]
-    unfolding obt(2-3) using obt(1) and hull_subset[unfolded subset_eq, rule_format, of _ s] by auto qed
-
-lemma unit_interval_convex_hull:
-  "{0::real^'n::finite .. 1} = convex hull {x. \<forall>i. (x$i = 0) \<or> (x$i = 1)}" (is "?int = convex hull ?points")
-proof- have 01:"{0,1} \<subseteq> convex hull ?points" apply rule apply(rule_tac hull_subset[unfolded subset_eq, rule_format]) by auto
-  { fix n x assume "x\<in>{0::real^'n .. 1}" "n \<le> CARD('n)" "card {i. x$i \<noteq> 0} \<le> n" 
-  hence "x\<in>convex hull ?points" proof(induct n arbitrary: x)
-    case 0 hence "x = 0" apply(subst Cart_eq) apply rule by auto
-    thus "x\<in>convex hull ?points" using 01 by auto
-  next
-    case (Suc n) show "x\<in>convex hull ?points" proof(cases "{i. x$i \<noteq> 0} = {}")
-      case True hence "x = 0" unfolding Cart_eq by auto
-      thus "x\<in>convex hull ?points" using 01 by auto
-    next
-      case False def xi \<equiv> "Min ((\<lambda>i. x$i) ` {i. x$i \<noteq> 0})"
-      have "xi \<in> (\<lambda>i. x$i) ` {i. x$i \<noteq> 0}" unfolding xi_def apply(rule Min_in) using False by auto
-      then obtain i where i':"x$i = xi" "x$i \<noteq> 0" by auto
-      have i:"\<And>j. x$j > 0 \<Longrightarrow> x$i \<le> x$j"
-        unfolding i'(1) xi_def apply(rule_tac Min_le) unfolding image_iff
-        defer apply(rule_tac x=j in bexI) using i' by auto
-      have i01:"x$i \<le> 1" "x$i > 0" using Suc(2)[unfolded mem_interval,rule_format,of i] using i'(2) `x$i \<noteq> 0`
-        by(auto simp add: Cart_lambda_beta) 
-      show ?thesis proof(cases "x$i=1")
-        case True have "\<forall>j\<in>{i. x$i \<noteq> 0}. x$j = 1" apply(rule, rule ccontr) unfolding mem_Collect_eq proof-
-          fix j assume "x $ j \<noteq> 0" "x $ j \<noteq> 1"
-          hence j:"x$j \<in> {0<..<1}" using Suc(2) by(auto simp add: vector_less_eq_def elim!:allE[where x=j])
-          hence "x$j \<in> op $ x ` {i. x $ i \<noteq> 0}" by auto 
-          hence "x$j \<ge> x$i" unfolding i'(1) xi_def apply(rule_tac Min_le) by auto
-          thus False using True Suc(2) j by(auto simp add: vector_less_eq_def elim!:ballE[where x=j]) qed
-        thus "x\<in>convex hull ?points" apply(rule_tac hull_subset[unfolded subset_eq, rule_format])
-          by(auto simp add: Cart_lambda_beta)
-      next let ?y = "\<lambda>j. if x$j = 0 then 0 else (x$j - x$i) / (1 - x$i)"
-        case False hence *:"x = x$i *\<^sub>R (\<chi> j. if x$j = 0 then 0 else 1) + (1 - x$i) *\<^sub>R (\<chi> j. ?y j)" unfolding Cart_eq
-          by(auto simp add: Cart_lambda_beta vector_add_component vector_smult_component vector_minus_component field_simps)
-        { fix j have "x$j \<noteq> 0 \<Longrightarrow> 0 \<le> (x $ j - x $ i) / (1 - x $ i)" "(x $ j - x $ i) / (1 - x $ i) \<le> 1"
-            apply(rule_tac divide_nonneg_pos) using i(1)[of j] using False i01
-            using Suc(2)[unfolded mem_interval, rule_format, of j] by(auto simp add:field_simps Cart_lambda_beta) 
-          hence "0 \<le> ?y j \<and> ?y j \<le> 1" by auto }
-        moreover have "i\<in>{j. x$j \<noteq> 0} - {j. ((\<chi> j. ?y j)::real^'n) $ j \<noteq> 0}" using i01 by(auto simp add: Cart_lambda_beta)
-        hence "{j. x$j \<noteq> 0} \<noteq> {j. ((\<chi> j. ?y j)::real^'n::finite) $ j \<noteq> 0}" by auto
-        hence **:"{j. ((\<chi> j. ?y j)::real^'n::finite) $ j \<noteq> 0} \<subset> {j. x$j \<noteq> 0}" apply - apply rule by(auto simp add: Cart_lambda_beta)  
-        have "card {j. ((\<chi> j. ?y j)::real^'n) $ j \<noteq> 0} \<le> n" using less_le_trans[OF psubset_card_mono[OF _ **] Suc(4)] by auto
-        ultimately show ?thesis apply(subst *) apply(rule convex_convex_hull[unfolded convex_def, rule_format])
-          apply(rule_tac hull_subset[unfolded subset_eq, rule_format]) defer apply(rule Suc(1))
-          unfolding mem_interval using i01 Suc(3) by (auto simp add: Cart_lambda_beta)
-      qed qed qed } note * = this
-  show ?thesis apply rule defer apply(rule hull_minimal) unfolding subset_eq prefer 3 apply rule 
-    apply(rule_tac n2="CARD('n)" in *) prefer 3 apply(rule card_mono) using 01 and convex_interval(1) prefer 5 apply - apply rule
-    unfolding mem_interval apply rule unfolding mem_Collect_eq apply(erule_tac x=i in allE)
-    by(auto simp add: vector_less_eq_def mem_def[of _ convex]) qed
-
-subsection {* And this is a finite set of vertices. *}
-
-lemma unit_cube_convex_hull: obtains s where "finite s" "{0 .. 1::real^'n::finite} = convex hull s"
-  apply(rule that[of "{x::real^'n::finite. \<forall>i. x$i=0 \<or> x$i=1}"])
-  apply(rule finite_subset[of _ "(\<lambda>s. (\<chi> i. if i\<in>s then 1::real else 0)::real^'n::finite) ` UNIV"])
-  prefer 3 apply(rule unit_interval_convex_hull) apply rule unfolding mem_Collect_eq proof-
-  fix x::"real^'n" assume as:"\<forall>i. x $ i = 0 \<or> x $ i = 1"
-  show "x \<in> (\<lambda>s. \<chi> i. if i \<in> s then 1 else 0) ` UNIV" apply(rule image_eqI[where x="{i. x$i = 1}"])
-    unfolding Cart_eq using as by(auto simp add:Cart_lambda_beta) qed auto
-
-subsection {* Hence any cube (could do any nonempty interval). *}
-
-lemma cube_convex_hull:
-  assumes "0 < d" obtains s::"(real^'n::finite) set" where "finite s" "{x - (\<chi> i. d) .. x + (\<chi> i. d)} = convex hull s" proof-
-  let ?d = "(\<chi> i. d)::real^'n"
-  have *:"{x - ?d .. x + ?d} = (\<lambda>y. x - ?d + (2 * d) *\<^sub>R y) ` {0 .. 1}" apply(rule set_ext, rule)
-    unfolding image_iff defer apply(erule bexE) proof-
-    fix y assume as:"y\<in>{x - ?d .. x + ?d}"
-    { fix i::'n have "x $ i \<le> d + y $ i" "y $ i \<le> d + x $ i" using as[unfolded mem_interval, THEN spec[where x=i]]
-        by(auto simp add: vector_component)
-      hence "1 \<ge> inverse d * (x $ i - y $ i)" "1 \<ge> inverse d * (y $ i - x $ i)"
-        apply(rule_tac[!] mult_left_le_imp_le[OF _ assms]) unfolding mult_assoc[THEN sym]
-        using assms by(auto simp add: field_simps right_inverse) 
-      hence "inverse d * (x $ i * 2) \<le> 2 + inverse d * (y $ i * 2)"
-            "inverse d * (y $ i * 2) \<le> 2 + inverse d * (x $ i * 2)" by(auto simp add:field_simps) }
-    hence "inverse (2 * d) *\<^sub>R (y - (x - ?d)) \<in> {0..1}" unfolding mem_interval using assms
-      by(auto simp add: Cart_eq vector_component_simps field_simps)
-    thus "\<exists>z\<in>{0..1}. y = x - ?d + (2 * d) *\<^sub>R z" apply- apply(rule_tac x="inverse (2 * d) *\<^sub>R (y - (x - ?d))" in bexI) 
-      using assms by(auto simp add: Cart_eq vector_less_eq_def Cart_lambda_beta)
-  next
-    fix y z assume as:"z\<in>{0..1}" "y = x - ?d + (2*d) *\<^sub>R z" 
-    have "\<And>i. 0 \<le> d * z $ i \<and> d * z $ i \<le> d" using assms as(1)[unfolded mem_interval] apply(erule_tac x=i in allE)
-      apply rule apply(rule mult_nonneg_nonneg) prefer 3 apply(rule mult_right_le_one_le)
-      using assms by(auto simp add: vector_component_simps Cart_eq)
-    thus "y \<in> {x - ?d..x + ?d}" unfolding as(2) mem_interval apply- apply rule using as(1)[unfolded mem_interval]
-      apply(erule_tac x=i in allE) using assms by(auto simp add:  vector_component_simps Cart_eq) qed
-  obtain s where "finite s" "{0..1::real^'n} = convex hull s" using unit_cube_convex_hull by auto
-  thus ?thesis apply(rule_tac that[of "(\<lambda>y. x - ?d + (2 * d) *\<^sub>R y)` s"]) unfolding * and convex_hull_affinity by auto qed
-
-subsection {* Bounded convex function on open set is continuous. *}
-
-lemma convex_on_bounded_continuous:
-  fixes s :: "(real ^ _) set"
-  assumes "open s" "convex_on s f" "\<forall>x\<in>s. abs(f x) \<le> b"
-  shows "continuous_on s f"
-  apply(rule continuous_at_imp_continuous_on) unfolding continuous_at_real_range proof(rule,rule,rule)
-  fix x e assume "x\<in>s" "(0::real) < e"
-  def B \<equiv> "abs b + 1"
-  have B:"0 < B" "\<And>x. x\<in>s \<Longrightarrow> abs (f x) \<le> B"
-    unfolding B_def defer apply(drule assms(3)[rule_format]) by auto
-  obtain k where "k>0"and k:"cball x k \<subseteq> s" using assms(1)[unfolded open_contains_cball, THEN bspec[where x=x]] using `x\<in>s` by auto
-  show "\<exists>d>0. \<forall>x'. norm (x' - x) < d \<longrightarrow> \<bar>f x' - f x\<bar> < e"
-    apply(rule_tac x="min (k / 2) (e / (2 * B) * k)" in exI) apply rule defer proof(rule,rule)
-    fix y assume as:"norm (y - x) < min (k / 2) (e / (2 * B) * k)" 
-    show "\<bar>f y - f x\<bar> < e" proof(cases "y=x")
-      case False def t \<equiv> "k / norm (y - x)"
-      have "2 < t" "0<t" unfolding t_def using as False and `k>0` by(auto simp add:field_simps)
-      have "y\<in>s" apply(rule k[unfolded subset_eq,rule_format]) unfolding mem_cball dist_norm
-        apply(rule order_trans[of _ "2 * norm (x - y)"]) using as by(auto simp add: field_simps norm_minus_commute) 
-      { def w \<equiv> "x + t *\<^sub>R (y - x)"
-        have "w\<in>s" unfolding w_def apply(rule k[unfolded subset_eq,rule_format]) unfolding mem_cball dist_norm 
-          unfolding t_def using `k>0` by auto
-        have "(1 / t) *\<^sub>R x + - x + ((t - 1) / t) *\<^sub>R x = (1 / t - 1 + (t - 1) / t) *\<^sub>R x" by (auto simp add: algebra_simps)
-        also have "\<dots> = 0"  using `t>0` by(auto simp add:field_simps)
-        finally have w:"(1 / t) *\<^sub>R w + ((t - 1) / t) *\<^sub>R x = y" unfolding w_def using False and `t>0` by (auto simp add: algebra_simps)
-        have  "2 * B < e * t" unfolding t_def using `0<e` `0<k` `B>0` and as and False by (auto simp add:field_simps) 
-        hence "(f w - f x) / t < e"
-          using B(2)[OF `w\<in>s`] and B(2)[OF `x\<in>s`] using `t>0` by(auto simp add:field_simps) 
-        hence th1:"f y - f x < e" apply- apply(rule le_less_trans) defer apply assumption
-          using assms(2)[unfolded convex_on_def,rule_format,of w x "1/t" "(t - 1)/t", unfolded w]
-          using `0<t` `2<t` and `x\<in>s` `w\<in>s` by(auto simp add:field_simps) }
-      moreover 
-      { def w \<equiv> "x - t *\<^sub>R (y - x)"
-        have "w\<in>s" unfolding w_def apply(rule k[unfolded subset_eq,rule_format]) unfolding mem_cball dist_norm 
-          unfolding t_def using `k>0` by auto
-        have "(1 / (1 + t)) *\<^sub>R x + (t / (1 + t)) *\<^sub>R x = (1 / (1 + t) + t / (1 + t)) *\<^sub>R x" by (auto simp add: algebra_simps)
-        also have "\<dots>=x" using `t>0` by (auto simp add:field_simps)
-        finally have w:"(1 / (1+t)) *\<^sub>R w + (t / (1 + t)) *\<^sub>R y = x" unfolding w_def using False and `t>0` by (auto simp add: algebra_simps)
-        have  "2 * B < e * t" unfolding t_def using `0<e` `0<k` `B>0` and as and False by (auto simp add:field_simps) 
-        hence *:"(f w - f y) / t < e" using B(2)[OF `w\<in>s`] and B(2)[OF `y\<in>s`] using `t>0` by(auto simp add:field_simps) 
-        have "f x \<le> 1 / (1 + t) * f w + (t / (1 + t)) * f y" 
-          using assms(2)[unfolded convex_on_def,rule_format,of w y "1/(1+t)" "t / (1+t)",unfolded w]
-          using `0<t` `2<t` and `y\<in>s` `w\<in>s` by (auto simp add:field_simps)
-        also have "\<dots> = (f w + t * f y) / (1 + t)" using `t>0` unfolding real_divide_def by (auto simp add:field_simps)
-        also have "\<dots> < e + f y" using `t>0` * `e>0` by(auto simp add:field_simps)
-        finally have "f x - f y < e" by auto }
-      ultimately show ?thesis by auto 
-    qed(insert `0<e`, auto) 
-  qed(insert `0<e` `0<k` `0<B`, auto simp add:field_simps intro!:mult_pos_pos) qed
-
-subsection {* Upper bound on a ball implies upper and lower bounds. *}
-
-lemma convex_bounds_lemma:
-  fixes x :: "real ^ _"
-  assumes "convex_on (cball x e) f"  "\<forall>y \<in> cball x e. f y \<le> b"
-  shows "\<forall>y \<in> cball x e. abs(f y) \<le> b + 2 * abs(f x)"
-  apply(rule) proof(cases "0 \<le> e") case True
-  fix y assume y:"y\<in>cball x e" def z \<equiv> "2 *\<^sub>R x - y"
-  have *:"x - (2 *\<^sub>R x - y) = y - x" by vector
-  have z:"z\<in>cball x e" using y unfolding z_def mem_cball dist_norm * by(auto simp add: norm_minus_commute)
-  have "(1 / 2) *\<^sub>R y + (1 / 2) *\<^sub>R z = x" unfolding z_def by (auto simp add: algebra_simps)
-  thus "\<bar>f y\<bar> \<le> b + 2 * \<bar>f x\<bar>" using assms(1)[unfolded convex_on_def,rule_format, OF y z, of "1/2" "1/2"]
-    using assms(2)[rule_format,OF y] assms(2)[rule_format,OF z] by(auto simp add:field_simps)
-next case False fix y assume "y\<in>cball x e" 
-  hence "dist x y < 0" using False unfolding mem_cball not_le by (auto simp del: dist_not_less_zero)
-  thus "\<bar>f y\<bar> \<le> b + 2 * \<bar>f x\<bar>" using zero_le_dist[of x y] by auto qed
-
-subsection {* Hence a convex function on an open set is continuous. *}
-
-lemma convex_on_continuous:
-  assumes "open (s::(real^'n::finite) set)" "convex_on s f" 
-  shows "continuous_on s f"
-  unfolding continuous_on_eq_continuous_at[OF assms(1)] proof
-  note dimge1 = dimindex_ge_1[where 'a='n]
-  fix x assume "x\<in>s"
-  then obtain e where e:"cball x e \<subseteq> s" "e>0" using assms(1) unfolding open_contains_cball by auto
-  def d \<equiv> "e / real CARD('n)"
-  have "0 < d" unfolding d_def using `e>0` dimge1 by(rule_tac divide_pos_pos, auto) 
-  let ?d = "(\<chi> i. d)::real^'n"
-  obtain c where c:"finite c" "{x - ?d..x + ?d} = convex hull c" using cube_convex_hull[OF `d>0`, of x] by auto
-  have "x\<in>{x - ?d..x + ?d}" using `d>0` unfolding mem_interval by(auto simp add:vector_component_simps)
-  hence "c\<noteq>{}" apply(rule_tac ccontr) using c by(auto simp add:convex_hull_empty)
-  def k \<equiv> "Max (f ` c)"
-  have "convex_on {x - ?d..x + ?d} f" apply(rule convex_on_subset[OF assms(2)])
-    apply(rule subset_trans[OF _ e(1)]) unfolding subset_eq mem_cball proof 
-    fix z assume z:"z\<in>{x - ?d..x + ?d}"
-    have e:"e = setsum (\<lambda>i. d) (UNIV::'n set)" unfolding setsum_constant d_def using dimge1
-      by (metis card_enum field_simps d_def not_one_le_zero of_nat_le_iff real_eq_of_nat real_of_nat_1)
-    show "dist x z \<le> e" unfolding dist_norm e apply(rule_tac order_trans[OF norm_le_l1], rule setsum_mono)
-      using z[unfolded mem_interval] apply(erule_tac x=i in allE) by(auto simp add:field_simps vector_component_simps) qed
-  hence k:"\<forall>y\<in>{x - ?d..x + ?d}. f y \<le> k" unfolding c(2) apply(rule_tac convex_on_convex_hull_bound) apply assumption
-    unfolding k_def apply(rule, rule Max_ge) using c(1) by auto
-  have "d \<le> e" unfolding d_def apply(rule mult_imp_div_pos_le) using `e>0` dimge1 unfolding mult_le_cancel_left1 using real_dimindex_ge_1 by auto
-  hence dsube:"cball x d \<subseteq> cball x e" unfolding subset_eq Ball_def mem_cball by auto
-  have conv:"convex_on (cball x d) f" apply(rule convex_on_subset, rule convex_on_subset[OF assms(2)]) apply(rule e(1)) using dsube by auto
-  hence "\<forall>y\<in>cball x d. abs (f y) \<le> k + 2 * abs (f x)" apply(rule_tac convex_bounds_lemma) apply assumption proof
-    fix y assume y:"y\<in>cball x d"
-    { fix i::'n have "x $ i - d \<le> y $ i"  "y $ i \<le> x $ i + d" 
-        using order_trans[OF component_le_norm y[unfolded mem_cball dist_norm], of i] by(auto simp add: vector_component)  }
-    thus "f y \<le> k" apply(rule_tac k[rule_format]) unfolding mem_cball mem_interval dist_norm 
-      by(auto simp add: vector_component_simps) qed
-  hence "continuous_on (ball x d) f" apply(rule_tac convex_on_bounded_continuous)
-    apply(rule open_ball, rule convex_on_subset[OF conv], rule ball_subset_cball) by auto
-  thus "continuous (at x) f" unfolding continuous_on_eq_continuous_at[OF open_ball] using `d>0` by auto qed
-
-subsection {* Line segments, starlike sets etc.                                         *)
-(* Use the same overloading tricks as for intervals, so that                 *)
-(* segment[a,b] is closed and segment(a,b) is open relative to affine hull. *}
-
-definition
-  midpoint :: "real ^ 'n::finite \<Rightarrow> real ^ 'n \<Rightarrow> real ^ 'n" where
-  "midpoint a b = (inverse (2::real)) *\<^sub>R (a + b)"
-
-definition
-  open_segment :: "real ^ 'n::finite \<Rightarrow> real ^ 'n \<Rightarrow> (real ^ 'n) set" where
-  "open_segment a b = {(1 - u) *\<^sub>R a + u *\<^sub>R b | u::real.  0 < u \<and> u < 1}"
-
-definition
-  closed_segment :: "real ^ 'n::finite \<Rightarrow> real ^ 'n \<Rightarrow> (real ^ 'n) set" where
-  "closed_segment a b = {(1 - u) *\<^sub>R a + u *\<^sub>R b | u::real. 0 \<le> u \<and> u \<le> 1}"
-
-definition "between = (\<lambda> (a,b). closed_segment a b)"
-
-lemmas segment = open_segment_def closed_segment_def
-
-definition "starlike s \<longleftrightarrow> (\<exists>a\<in>s. \<forall>x\<in>s. closed_segment a x \<subseteq> s)"
-
-lemma midpoint_refl: "midpoint x x = x"
-  unfolding midpoint_def unfolding scaleR_right_distrib unfolding scaleR_left_distrib[THEN sym] by auto
-
-lemma midpoint_sym: "midpoint a b = midpoint b a" unfolding midpoint_def by (auto simp add: scaleR_right_distrib)
-
-lemma dist_midpoint:
-  "dist a (midpoint a b) = (dist a b) / 2" (is ?t1)
-  "dist b (midpoint a b) = (dist a b) / 2" (is ?t2)
-  "dist (midpoint a b) a = (dist a b) / 2" (is ?t3)
-  "dist (midpoint a b) b = (dist a b) / 2" (is ?t4)
-proof-
-  have *: "\<And>x y::real^'n::finite. 2 *\<^sub>R x = - y \<Longrightarrow> norm x = (norm y) / 2" unfolding equation_minus_iff by auto
-  have **:"\<And>x y::real^'n::finite. 2 *\<^sub>R x =   y \<Longrightarrow> norm x = (norm y) / 2" by auto
-  note scaleR_right_distrib [simp]
-  show ?t1 unfolding midpoint_def dist_norm apply (rule **) by(auto,vector)
-  show ?t2 unfolding midpoint_def dist_norm apply (rule *)  by(auto,vector)
-  show ?t3 unfolding midpoint_def dist_norm apply (rule *)  by(auto,vector)
-  show ?t4 unfolding midpoint_def dist_norm apply (rule **) by(auto,vector) qed
-
-lemma midpoint_eq_endpoint:
-  "midpoint a b = a \<longleftrightarrow> a = (b::real^'n::finite)"
-  "midpoint a b = b \<longleftrightarrow> a = b"
-  unfolding dist_eq_0_iff[where 'a="real^'n", THEN sym] dist_midpoint by auto
-
-lemma convex_contains_segment:
-  "convex s \<longleftrightarrow> (\<forall>a\<in>s. \<forall>b\<in>s. closed_segment a b \<subseteq> s)"
-  unfolding convex_alt closed_segment_def by auto
-
-lemma convex_imp_starlike:
-  "convex s \<Longrightarrow> s \<noteq> {} \<Longrightarrow> starlike s"
-  unfolding convex_contains_segment starlike_def by auto
-
-lemma segment_convex_hull:
- "closed_segment a b = convex hull {a,b}" proof-
-  have *:"\<And>x. {x} \<noteq> {}" by auto
-  have **:"\<And>u v. u + v = 1 \<longleftrightarrow> u = 1 - (v::real)" by auto
-  show ?thesis unfolding segment convex_hull_insert[OF *] convex_hull_singleton apply(rule set_ext)
-    unfolding mem_Collect_eq apply(rule,erule exE) 
-    apply(rule_tac x="1 - u" in exI) apply rule defer apply(rule_tac x=u in exI) defer
-    apply(erule exE, (erule conjE)?)+ apply(rule_tac x="1 - u" in exI) unfolding ** by auto qed
-
-lemma convex_segment: "convex (closed_segment a b)"
-  unfolding segment_convex_hull by(rule convex_convex_hull)
-
-lemma ends_in_segment: "a \<in> closed_segment a b" "b \<in> closed_segment a b"
-  unfolding segment_convex_hull apply(rule_tac[!] hull_subset[unfolded subset_eq, rule_format]) by auto
-
-lemma segment_furthest_le:
-  assumes "x \<in> closed_segment a b" shows "norm(y - x) \<le> norm(y - a) \<or>  norm(y - x) \<le> norm(y - b)" proof-
-  obtain z where "z\<in>{a, b}" "norm (x - y) \<le> norm (z - y)" using simplex_furthest_le[of "{a, b}" y]
-    using assms[unfolded segment_convex_hull] by auto
-  thus ?thesis by(auto simp add:norm_minus_commute) qed
-
-lemma segment_bound:
-  assumes "x \<in> closed_segment a b"
-  shows "norm(x - a) \<le> norm(b - a)" "norm(x - b) \<le> norm(b - a)"
-  using segment_furthest_le[OF assms, of a]
-  using segment_furthest_le[OF assms, of b]
-  by (auto simp add:norm_minus_commute) 
-
-lemma segment_refl:"closed_segment a a = {a}" unfolding segment by (auto simp add: algebra_simps)
-
-lemma between_mem_segment: "between (a,b) x \<longleftrightarrow> x \<in> closed_segment a b"
-  unfolding between_def mem_def by auto
-
-lemma between:"between (a,b) (x::real^'n::finite) \<longleftrightarrow> dist a b = (dist a x) + (dist x b)"
-proof(cases "a = b")
-  case True thus ?thesis unfolding between_def split_conv mem_def[of x, symmetric]
-    by(auto simp add:segment_refl dist_commute) next
-  case False hence Fal:"norm (a - b) \<noteq> 0" and Fal2: "norm (a - b) > 0" by auto 
-  have *:"\<And>u. a - ((1 - u) *\<^sub>R a + u *\<^sub>R b) = u *\<^sub>R (a - b)" by (auto simp add: algebra_simps)
-  show ?thesis unfolding between_def split_conv mem_def[of x, symmetric] closed_segment_def mem_Collect_eq
-    apply rule apply(erule exE, (erule conjE)+) apply(subst dist_triangle_eq) proof-
-      fix u assume as:"x = (1 - u) *\<^sub>R a + u *\<^sub>R b" "0 \<le> u" "u \<le> 1" 
-      hence *:"a - x = u *\<^sub>R (a - b)" "x - b = (1 - u) *\<^sub>R (a - b)"
-        unfolding as(1) by(auto simp add:algebra_simps)
-      show "norm (a - x) *\<^sub>R (x - b) = norm (x - b) *\<^sub>R (a - x)"
-        unfolding norm_minus_commute[of x a] * Cart_eq using as(2,3)
-        by(auto simp add: vector_component_simps field_simps)
-    next assume as:"dist a b = dist a x + dist x b"
-      have "norm (a - x) / norm (a - b) \<le> 1" unfolding divide_le_eq_1_pos[OF Fal2] unfolding as[unfolded dist_norm] norm_ge_zero by auto 
-      thus "\<exists>u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \<and> 0 \<le> u \<and> u \<le> 1" apply(rule_tac x="dist a x / dist a b" in exI)
-        unfolding dist_norm Cart_eq apply- apply rule defer apply(rule, rule divide_nonneg_pos) prefer 4 proof rule
-          fix i::'n have "((1 - norm (a - x) / norm (a - b)) *\<^sub>R a + (norm (a - x) / norm (a - b)) *\<^sub>R b) $ i =
-            ((norm (a - b) - norm (a - x)) * (a $ i) + norm (a - x) * (b $ i)) / norm (a - b)"
-            using Fal by(auto simp add:vector_component_simps field_simps)
-          also have "\<dots> = x$i" apply(rule divide_eq_imp[OF Fal])
-            unfolding as[unfolded dist_norm] using as[unfolded dist_triangle_eq Cart_eq,rule_format, of i]
-            by(auto simp add:field_simps vector_component_simps)
-          finally show "x $ i = ((1 - norm (a - x) / norm (a - b)) *\<^sub>R a + (norm (a - x) / norm (a - b)) *\<^sub>R b) $ i" by auto
-        qed(insert Fal2, auto) qed qed
-
-lemma between_midpoint: fixes a::"real^'n::finite" shows
-  "between (a,b) (midpoint a b)" (is ?t1) 
-  "between (b,a) (midpoint a b)" (is ?t2)
-proof- have *:"\<And>x y z. x = (1/2::real) *\<^sub>R z \<Longrightarrow> y = (1/2) *\<^sub>R z \<Longrightarrow> norm z = norm x + norm y" by auto
-  show ?t1 ?t2 unfolding between midpoint_def dist_norm apply(rule_tac[!] *)
-    by(auto simp add:field_simps Cart_eq vector_component_simps) qed
-
-lemma between_mem_convex_hull:
-  "between (a,b) x \<longleftrightarrow> x \<in> convex hull {a,b}"
-  unfolding between_mem_segment segment_convex_hull ..
-
-subsection {* Shrinking towards the interior of a convex set. *}
-
-lemma mem_interior_convex_shrink:
-  fixes s :: "(real ^ _) set"
-  assumes "convex s" "c \<in> interior s" "x \<in> s" "0 < e" "e \<le> 1"
-  shows "x - e *\<^sub>R (x - c) \<in> interior s"
-proof- obtain d where "d>0" and d:"ball c d \<subseteq> s" using assms(2) unfolding mem_interior by auto
-  show ?thesis unfolding mem_interior apply(rule_tac x="e*d" in exI)
-    apply(rule) defer unfolding subset_eq Ball_def mem_ball proof(rule,rule)
-    fix y assume as:"dist (x - e *\<^sub>R (x - c)) y < e * d"
-    have *:"y = (1 - (1 - e)) *\<^sub>R ((1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x) + (1 - e) *\<^sub>R x" using `e>0` by (auto simp add: scaleR_left_diff_distrib scaleR_right_diff_distrib)
-    have "dist c ((1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x) = abs(1/e) * norm (e *\<^sub>R c - y + (1 - e) *\<^sub>R x)"
-      unfolding dist_norm unfolding norm_scaleR[THEN sym] apply(rule norm_eqI) using `e>0`
-      by(auto simp add:vector_component_simps Cart_eq field_simps) 
-    also have "\<dots> = abs(1/e) * norm (x - e *\<^sub>R (x - c) - y)" by(auto intro!:norm_eqI simp add: algebra_simps)
-    also have "\<dots> < d" using as[unfolded dist_norm] and `e>0`
-      by(auto simp add:pos_divide_less_eq[OF `e>0`] real_mult_commute)
-    finally show "y \<in> s" apply(subst *) apply(rule assms(1)[unfolded convex_alt,rule_format])
-      apply(rule d[unfolded subset_eq,rule_format]) unfolding mem_ball using assms(3-5) by auto
-  qed(rule mult_pos_pos, insert `e>0` `d>0`, auto) qed
-
-lemma mem_interior_closure_convex_shrink:
-  fixes s :: "(real ^ _) set"
-  assumes "convex s" "c \<in> interior s" "x \<in> closure s" "0 < e" "e \<le> 1"
-  shows "x - e *\<^sub>R (x - c) \<in> interior s"
-proof- obtain d where "d>0" and d:"ball c d \<subseteq> s" using assms(2) unfolding mem_interior by auto
-  have "\<exists>y\<in>s. norm (y - x) * (1 - e) < e * d" proof(cases "x\<in>s")
-    case True thus ?thesis using `e>0` `d>0` by(rule_tac bexI[where x=x], auto intro!: mult_pos_pos) next
-    case False hence x:"x islimpt s" using assms(3)[unfolded closure_def] by auto
-    show ?thesis proof(cases "e=1")
-      case True obtain y where "y\<in>s" "y \<noteq> x" "dist y x < 1"
-        using x[unfolded islimpt_approachable,THEN spec[where x=1]] by auto
-      thus ?thesis apply(rule_tac x=y in bexI) unfolding True using `d>0` by auto next
-      case False hence "0 < e * d / (1 - e)" and *:"1 - e > 0"
-        using `e\<le>1` `e>0` `d>0` by(auto intro!:mult_pos_pos divide_pos_pos)
-      then obtain y where "y\<in>s" "y \<noteq> x" "dist y x < e * d / (1 - e)"
-        using x[unfolded islimpt_approachable,THEN spec[where x="e*d / (1 - e)"]] by auto
-      thus ?thesis apply(rule_tac x=y in bexI) unfolding dist_norm using pos_less_divide_eq[OF *] by auto qed qed
-  then obtain y where "y\<in>s" and y:"norm (y - x) * (1 - e) < e * d" by auto
-  def z \<equiv> "c + ((1 - e) / e) *\<^sub>R (x - y)"
-  have *:"x - e *\<^sub>R (x - c) = y - e *\<^sub>R (y - z)" unfolding z_def using `e>0` by (auto simp add: scaleR_right_diff_distrib scaleR_right_distrib scaleR_left_diff_distrib)
-  have "z\<in>interior s" apply(rule subset_interior[OF d,unfolded subset_eq,rule_format])
-    unfolding interior_open[OF open_ball] mem_ball z_def dist_norm using y and assms(4,5)
-    by(auto simp add:field_simps norm_minus_commute)
-  thus ?thesis unfolding * apply - apply(rule mem_interior_convex_shrink) 
-    using assms(1,4-5) `y\<in>s` by auto qed
-
-subsection {* Some obvious but surprisingly hard simplex lemmas. *}
-
-lemma simplex:
-  assumes "finite s" "0 \<notin> s"
-  shows "convex hull (insert 0 s) =  { y. (\<exists>u. (\<forall>x\<in>s. 0 \<le> u x) \<and> setsum u s \<le> 1 \<and> setsum (\<lambda>x. u x *\<^sub>R x) s = y)}"
-  unfolding convex_hull_finite[OF finite.insertI[OF assms(1)]] apply(rule set_ext, rule) unfolding mem_Collect_eq
-  apply(erule_tac[!] exE) apply(erule_tac[!] conjE)+ unfolding setsum_clauses(2)[OF assms(1)]
-  apply(rule_tac x=u in exI) defer apply(rule_tac x="\<lambda>x. if x = 0 then 1 - setsum u s else u x" in exI) using assms(2)
-  unfolding if_smult and setsum_delta_notmem[OF assms(2)] by auto
-
-lemma std_simplex:
-  "convex hull (insert 0 { basis i | i. i\<in>UNIV}) =
-        {x::real^'n::finite . (\<forall>i. 0 \<le> x$i) \<and> setsum (\<lambda>i. x$i) UNIV \<le> 1 }" (is "convex hull (insert 0 ?p) = ?s")
-proof- let ?D = "UNIV::'n set"
-  have "0\<notin>?p" by(auto simp add: basis_nonzero)
-  have "{(basis i)::real^'n |i. i \<in> ?D} = basis ` ?D" by auto
-  note sumbas = this  setsum_reindex[OF basis_inj, unfolded o_def]
-  show ?thesis unfolding simplex[OF finite_stdbasis `0\<notin>?p`] apply(rule set_ext) unfolding mem_Collect_eq apply rule
-    apply(erule exE, (erule conjE)+) apply(erule_tac[2] conjE)+ proof-
-    fix x::"real^'n" and u assume as: "\<forall>x\<in>{basis i |i. i \<in>?D}. 0 \<le> u x" "setsum u {basis i |i. i \<in> ?D} \<le> 1" "(\<Sum>x\<in>{basis i |i. i \<in>?D}. u x *\<^sub>R x) = x"
-    have *:"\<forall>i. u (basis i) = x$i" using as(3) unfolding sumbas and basis_expansion_unique [where 'a=real, unfolded smult_conv_scaleR] by auto
-    hence **:"setsum u {basis i |i. i \<in> ?D} = setsum (op $ x) ?D" unfolding sumbas by(rule_tac setsum_cong, auto)
-    show " (\<forall>i. 0 \<le> x $ i) \<and> setsum (op $ x) ?D \<le> 1" apply - proof(rule,rule)
-      fix i::'n show "0 \<le> x$i" unfolding *[rule_format,of i,THEN sym] apply(rule_tac as(1)[rule_format]) by auto
-    qed(insert as(2)[unfolded **], auto)
-  next fix x::"real^'n" assume as:"\<forall>i. 0 \<le> x $ i" "setsum (op $ x) ?D \<le> 1"
-    show "\<exists>u. (\<forall>x\<in>{basis i |i. i \<in> ?D}. 0 \<le> u x) \<and> setsum u {basis i |i. i \<in> ?D} \<le> 1 \<and> (\<Sum>x\<in>{basis i |i. i \<in> ?D}. u x *\<^sub>R x) = x"
-      apply(rule_tac x="\<lambda>y. inner y x" in exI) apply(rule,rule) unfolding mem_Collect_eq apply(erule exE) using as(1) apply(erule_tac x=i in allE) 
-      unfolding sumbas using as(2) and basis_expansion_unique [where 'a=real, unfolded smult_conv_scaleR] by(auto simp add:inner_basis) qed qed 
-
-lemma interior_std_simplex:
-  "interior (convex hull (insert 0 { basis i| i. i\<in>UNIV})) =
-  {x::real^'n::finite. (\<forall>i. 0 < x$i) \<and> setsum (\<lambda>i. x$i) UNIV < 1 }"
-  apply(rule set_ext) unfolding mem_interior std_simplex unfolding subset_eq mem_Collect_eq Ball_def mem_ball
-  unfolding Ball_def[symmetric] apply rule apply(erule exE, (erule conjE)+) defer apply(erule conjE) proof-
-  fix x::"real^'n" and e assume "0<e" and as:"\<forall>xa. dist x xa < e \<longrightarrow> (\<forall>x. 0 \<le> xa $ x) \<and> setsum (op $ xa) UNIV \<le> 1"
-  show "(\<forall>xa. 0 < x $ xa) \<and> setsum (op $ x) UNIV < 1" apply(rule,rule) proof-
-    fix i::'n show "0 < x $ i" using as[THEN spec[where x="x - (e / 2) *\<^sub>R basis i"]] and `e>0`
-      unfolding dist_norm by(auto simp add: norm_basis vector_component_simps basis_component elim:allE[where x=i])
-  next guess a using UNIV_witness[where 'a='n] ..
-    have **:"dist x (x + (e / 2) *\<^sub>R basis a) < e" using  `e>0` and norm_basis[of a]
-      unfolding dist_norm by(auto simp add: vector_component_simps basis_component intro!: mult_strict_left_mono_comm)
-    have "\<And>i. (x + (e / 2) *\<^sub>R basis a) $ i = x$i + (if i = a then e/2 else 0)" by(auto simp add:vector_component_simps)
-    hence *:"setsum (op $ (x + (e / 2) *\<^sub>R basis a)) UNIV = setsum (\<lambda>i. x$i + (if a = i then e/2 else 0)) UNIV" by(rule_tac setsum_cong, auto) 
-    have "setsum (op $ x) UNIV < setsum (op $ (x + (e / 2) *\<^sub>R basis a)) UNIV" unfolding * setsum_addf
-      using `0<e` dimindex_ge_1 by(auto simp add: setsum_delta')
-    also have "\<dots> \<le> 1" using ** apply(drule_tac as[rule_format]) by auto
-    finally show "setsum (op $ x) UNIV < 1" by auto qed
-next
-  fix x::"real^'n::finite" assume as:"\<forall>i. 0 < x $ i" "setsum (op $ x) UNIV < 1"
-  guess a using UNIV_witness[where 'a='b] ..
-  let ?d = "(1 - setsum (op $ x) UNIV) / real (CARD('n))"
-  have "Min ((op $ x) ` UNIV) > 0" apply(rule Min_grI) using as(1) dimindex_ge_1 by auto
-  moreover have"?d > 0" apply(rule divide_pos_pos) using as(2) using dimindex_ge_1 by(auto simp add: Suc_le_eq)
-  ultimately show "\<exists>e>0. \<forall>y. dist x y < e \<longrightarrow> (\<forall>i. 0 \<le> y $ i) \<and> setsum (op $ y) UNIV \<le> 1"
-    apply(rule_tac x="min (Min ((op $ x) ` UNIV)) ?D" in exI) apply rule defer apply(rule,rule) proof-
-    fix y assume y:"dist x y < min (Min (op $ x ` UNIV)) ?d"
-    have "setsum (op $ y) UNIV \<le> setsum (\<lambda>i. x$i + ?d) UNIV" proof(rule setsum_mono)
-      fix i::'n have "abs (y$i - x$i) < ?d" apply(rule le_less_trans) using component_le_norm[of "y - x" i]
-        using y[unfolded min_less_iff_conj dist_norm, THEN conjunct2] by(auto simp add:vector_component_simps norm_minus_commute)
-      thus "y $ i \<le> x $ i + ?d" by auto qed
-    also have "\<dots> \<le> 1" unfolding setsum_addf setsum_constant card_enum real_eq_of_nat using dimindex_ge_1 by(auto simp add: Suc_le_eq)
-    finally show "(\<forall>i. 0 \<le> y $ i) \<and> setsum (op $ y) UNIV \<le> 1" apply- proof(rule,rule)
-      fix i::'n have "norm (x - y) < x$i" using y[unfolded min_less_iff_conj dist_norm, THEN conjunct1]
-        using Min_gr_iff[of "op $ x ` dimset x"] dimindex_ge_1 by auto
-      thus "0 \<le> y$i" using component_le_norm[of "x - y" i] and as(1)[rule_format, of i] by(auto simp add: vector_component_simps)
-    qed auto qed auto qed
-
-lemma interior_std_simplex_nonempty: obtains a::"real^'n::finite" where
-  "a \<in> interior(convex hull (insert 0 {basis i | i . i \<in> UNIV}))" proof-
-  let ?D = "UNIV::'n set" let ?a = "setsum (\<lambda>b::real^'n. inverse (2 * real CARD('n)) *\<^sub>R b) {(basis i) | i. i \<in> ?D}"
-  have *:"{basis i :: real ^ 'n | i. i \<in> ?D} = basis ` ?D" by auto
-  { fix i have "?a $ i = inverse (2 * real CARD('n))"
-    unfolding setsum_component vector_smult_component and * and setsum_reindex[OF basis_inj] and o_def
-    apply(rule trans[of _ "setsum (\<lambda>j. if i = j then inverse (2 * real CARD('n)) else 0) ?D"]) apply(rule setsum_cong2)
-      unfolding setsum_delta'[OF finite_UNIV[where 'a='n]] and real_dimindex_ge_1[where 'n='n] by(auto simp add: basis_component[of i]) }
-  note ** = this
-  show ?thesis apply(rule that[of ?a]) unfolding interior_std_simplex mem_Collect_eq proof(rule,rule)
-    fix i::'n show "0 < ?a $ i" unfolding ** using dimindex_ge_1 by(auto simp add: Suc_le_eq) next
-    have "setsum (op $ ?a) ?D = setsum (\<lambda>i. inverse (2 * real CARD('n))) ?D" by(rule setsum_cong2, rule **) 
-    also have "\<dots> < 1" unfolding setsum_constant card_enum real_eq_of_nat real_divide_def[THEN sym] by (auto simp add:field_simps)
-    finally show "setsum (op $ ?a) ?D < 1" by auto qed qed
-
-subsection {* Paths. *}
-
-definition "path (g::real^1 \<Rightarrow> real^'n::finite) \<longleftrightarrow> continuous_on {0 .. 1} g"
-
-definition "pathstart (g::real^1 \<Rightarrow> real^'n) = g 0"
-
-definition "pathfinish (g::real^1 \<Rightarrow> real^'n) = g 1"
-
-definition "path_image (g::real^1 \<Rightarrow> real^'n) = g ` {0 .. 1}"
-
-definition "reversepath (g::real^1 \<Rightarrow> real^'n) = (\<lambda>x. g(1 - x))"
-
-definition joinpaths:: "(real^1 \<Rightarrow> real^'n) \<Rightarrow> (real^1 \<Rightarrow> real^'n) \<Rightarrow> (real^1 \<Rightarrow> real^'n)" (infixr "+++" 75)
-  where "joinpaths g1 g2 = (\<lambda>x. if dest_vec1 x \<le> ((1 / 2)::real) then g1 (2 *\<^sub>R x) else g2(2 *\<^sub>R x - 1))"
-definition "simple_path (g::real^1 \<Rightarrow> real^'n) \<longleftrightarrow>
-  (\<forall>x\<in>{0..1}. \<forall>y\<in>{0..1}. g x = g y \<longrightarrow> x = y \<or> x = 0 \<and> y = 1 \<or> x = 1 \<and> y = 0)"
-
-definition "injective_path (g::real^1 \<Rightarrow> real^'n) \<longleftrightarrow>
-  (\<forall>x\<in>{0..1}. \<forall>y\<in>{0..1}. g x = g y \<longrightarrow> x = y)"
-
-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
-
-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 apply(rule imageI)
-  unfolding mem_interval_1 vec_1[THEN sym] dest_vec1_vec by auto
-
-lemma pathfinish_in_path_image[intro]: "(pathfinish g) \<in> path_image g"
-  unfolding pathfinish_def path_image_def apply(rule imageI)
-  unfolding mem_interval_1 vec_1[THEN sym] dest_vec1_vec by auto
-
-lemma connected_path_image[intro]: "path g \<Longrightarrow> connected(path_image g)"
-  unfolding path_def path_image_def apply(rule connected_continuous_image, assumption)
-  by(rule convex_connected, rule convex_interval)
-
-lemma compact_path_image[intro]: "path g \<Longrightarrow> compact(path_image g)"
-  unfolding path_def path_image_def apply(rule compact_continuous_image, assumption)
-  by(rule compact_interval)
-
-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 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
-
-lemma pathfinish_join[simp]:"pathfinish(g1 +++ g2) = pathfinish g2" proof-
-  have "2 *\<^sub>R 1 - 1 = (1::real^1)" unfolding Cart_eq by(auto simp add:vector_component_simps)
-  thus ?thesis unfolding pathstart_def joinpaths_def pathfinish_def
-    unfolding vec_1[THEN sym] dest_vec1_vec by auto qed
-
-lemma path_image_reversepath[simp]: "path_image(reversepath g) = path_image g" proof-
-  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) by(auto simp add:vector_less_eq_def vector_component_simps elim!:ballE)
-  show ?thesis using *[of g] *[of "reversepath g"] unfolding reversepath_reversepath by auto qed
-
-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(rule continuous_on_sub, rule continuous_on_const, rule continuous_on_id)
-    apply(rule continuous_on_subset[of "{0..1}"], assumption)
-    by (auto, auto simp add:vector_less_eq_def vector_component_simps elim!:ballE)
-  show ?thesis using *[of g] *[of "reversepath g"] unfolding reversepath_reversepath by auto qed
-
-lemmas reversepath_simps = path_reversepath path_image_reversepath pathstart_reversepath pathfinish_reversepath
-
-lemma path_join[simp]: assumes "pathfinish g1 = pathstart g2" shows "path (g1 +++ g2) \<longleftrightarrow>  path g1 \<and> path g2"
-  unfolding path_def pathfinish_def pathstart_def apply rule defer apply(erule conjE) proof-
-  assume as:"continuous_on {0..1} (g1 +++ g2)"
-  have *:"g1 = (\<lambda>x. g1 (2 *\<^sub>R x)) \<circ> (\<lambda>x. (1/2) *\<^sub>R x)" 
-         "g2 = (\<lambda>x. g2 (2 *\<^sub>R x - 1)) \<circ> (\<lambda>x. (1/2) *\<^sub>R (x + 1))" unfolding o_def by auto
-  have "op *\<^sub>R (1 / 2) ` {0::real^1..1} \<subseteq> {0..1}"  "(\<lambda>x. (1 / 2) *\<^sub>R (x + 1)) ` {(0::real^1)..1} \<subseteq> {0..1}"
-    unfolding image_smult_interval by (auto, auto simp add:vector_less_eq_def vector_component_simps elim!:ballE)
-  thus "continuous_on {0..1} g1 \<and> continuous_on {0..1} g2" apply -apply rule
-    apply(subst *) defer apply(subst *) apply (rule_tac[!] continuous_on_compose)
-    apply (rule continuous_on_cmul, rule continuous_on_add, rule continuous_on_id, rule continuous_on_const) defer
-    apply (rule continuous_on_cmul, rule continuous_on_id) apply(rule_tac[!] continuous_on_eq[of _ "g1 +++ g2"]) defer prefer 3
-    apply(rule_tac[1-2] continuous_on_subset[of "{0 .. 1}"]) apply(rule as, assumption, rule as, assumption)
-    apply(rule) defer apply rule proof-
-    fix x assume "x \<in> op *\<^sub>R (1 / 2) ` {0::real^1..1}"
-    hence "dest_vec1 x \<le> 1 / 2" unfolding image_iff by(auto simp add: vector_component_simps)
-    thus "(g1 +++ g2) x = g1 (2 *\<^sub>R x)" unfolding joinpaths_def by auto next
-    fix x assume "x \<in> (\<lambda>x. (1 / 2) *\<^sub>R (x + 1)) ` {0::real^1..1}"
-    hence "dest_vec1 x \<ge> 1 / 2" unfolding image_iff by(auto simp add: vector_component_simps)
-    thus "(g1 +++ g2) x = g2 (2 *\<^sub>R x - 1)" proof(cases "dest_vec1 x = 1 / 2")
-      case True hence "x = (1/2) *\<^sub>R 1" unfolding Cart_eq by(auto simp add: forall_1 vector_component_simps)
-      thus ?thesis unfolding joinpaths_def using assms[unfolded pathstart_def pathfinish_def] by auto
-    qed (auto simp add:le_less joinpaths_def) qed
-next assume as:"continuous_on {0..1} g1" "continuous_on {0..1} g2"
-  have *:"{0 .. 1::real^1} = {0.. (1/2)*\<^sub>R 1} \<union> {(1/2) *\<^sub>R 1 .. 1}" by(auto simp add: vector_component_simps) 
-  have **:"op *\<^sub>R 2 ` {0..(1 / 2) *\<^sub>R 1} = {0..1::real^1}" apply(rule set_ext, rule) unfolding image_iff 
-    defer apply(rule_tac x="(1/2)*\<^sub>R x" in bexI) by(auto simp add: vector_component_simps)
-  have ***:"(\<lambda>x. 2 *\<^sub>R x - 1) ` {(1 / 2) *\<^sub>R 1..1} = {0..1::real^1}"
-    unfolding image_affinity_interval[of _ "- 1", unfolded diff_def[symmetric]] and interval_eq_empty_1
-    by(auto simp add: vector_component_simps)
-  have ****:"\<And>x::real^1. x $ 1 * 2 = 1 \<longleftrightarrow> x = (1/2) *\<^sub>R 1" unfolding Cart_eq by(auto simp add: forall_1 vector_component_simps)
-  show "continuous_on {0..1} (g1 +++ g2)" unfolding * apply(rule continuous_on_union) apply(rule closed_interval)+ proof-
-    show "continuous_on {0..(1 / 2) *\<^sub>R 1} (g1 +++ g2)" apply(rule continuous_on_eq[of _ "\<lambda>x. g1 (2 *\<^sub>R x)"]) defer
-      unfolding o_def[THEN sym] apply(rule continuous_on_compose) apply(rule continuous_on_cmul, rule continuous_on_id)
-      unfolding ** apply(rule as(1)) unfolding joinpaths_def by(auto simp add: vector_component_simps) next
-    show "continuous_on {(1/2)*\<^sub>R1..1} (g1 +++ g2)" apply(rule continuous_on_eq[of _ "g2 \<circ> (\<lambda>x. 2 *\<^sub>R x - 1)"]) defer
-      apply(rule continuous_on_compose) apply(rule continuous_on_sub, rule continuous_on_cmul, rule continuous_on_id, rule continuous_on_const)
-      unfolding *** o_def joinpaths_def apply(rule as(2)) using assms[unfolded pathstart_def pathfinish_def]
-      by(auto simp add: vector_component_simps ****) qed qed
-
-lemma path_image_join_subset: "path_image(g1 +++ g2) \<subseteq> (path_image g1 \<union> path_image g2)" proof
-  fix x assume "x \<in> path_image (g1 +++ g2)"
-  then obtain y where y:"y\<in>{0..1}" "x = (if dest_vec1 y \<le> 1 / 2 then g1 (2 *\<^sub>R y) else g2 (2 *\<^sub>R y - 1))"
-    unfolding path_image_def image_iff joinpaths_def by auto
-  thus "x \<in> path_image g1 \<union> path_image g2" apply(cases "dest_vec1 y \<le> 1/2")
-    apply(rule_tac UnI1) defer apply(rule_tac UnI2) unfolding y(2) path_image_def using y(1)
-    by(auto intro!: imageI simp add: vector_component_simps) qed
-
-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
-
-lemma path_image_join:
-  assumes "path g1" "path g2" "pathfinish g1 = pathstart g2"
-  shows "path_image(g1 +++ g2) = (path_image g1) \<union> (path_image g2)"
-apply(rule, rule path_image_join_subset, 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" unfolding path_image_def image_iff by auto
-  thus "x \<in> path_image (g1 +++ g2)" unfolding joinpaths_def path_image_def image_iff
-    apply(rule_tac x="(1/2) *\<^sub>R y" in bexI) by(auto simp add: vector_component_simps) next
-  fix x assume "x \<in> path_image g2"
-  then obtain y where y:"y\<in>{0..1}" "x = g2 y" unfolding path_image_def image_iff by auto
-  moreover have *:"y $ 1 = 0 \<Longrightarrow> y = 0" unfolding Cart_eq by auto
-  ultimately 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) using assms(3)[unfolded pathfinish_def pathstart_def]
-    by(auto simp add: vector_component_simps) 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
-
-lemma simple_path_reversepath: assumes "simple_path g" shows "simple_path (reversepath g)"
-  using assms 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)
-  unfolding mem_interval_1 by(auto simp add:vector_component_simps)
-
-lemma dest_vec1_scaleR [simp]:
-  "dest_vec1 (scaleR a x) = scaleR a (dest_vec1 x)"
-unfolding dest_vec1_def by simp
-
-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)"
-unfolding simple_path_def proof((rule ballI)+, rule impI) let ?g = "g1 +++ g2"
-  note inj = assms(1,2)[unfolded injective_path_def, rule_format]
-  fix x y::"real^1" 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$1 \<le> 1/2",case_tac[!] "y$1 \<le> 1/2", unfold not_le)
-    assume as:"x $ 1 \<le> 1 / 2" "y $ 1 \<le> 1 / 2"
-    hence "g1 (2 *\<^sub>R x) = g1 (2 *\<^sub>R y)" using xy(3) unfolding joinpaths_def dest_vec1_def by auto
-    moreover have "2 *\<^sub>R x \<in> {0..1}" "2 *\<^sub>R y \<in> {0..1}" using xy(1,2) as
-      unfolding mem_interval_1 dest_vec1_def by(auto simp add:vector_component_simps)
-    ultimately show ?thesis using inj(1)[of "2*\<^sub>R x" "2*\<^sub>R y"] by auto
-  next assume as:"x $ 1 > 1 / 2" "y $ 1 > 1 / 2"
-    hence "g2 (2 *\<^sub>R x - 1) = g2 (2 *\<^sub>R y - 1)" using xy(3) unfolding joinpaths_def dest_vec1_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
-      unfolding mem_interval_1 dest_vec1_def by(auto simp add:vector_component_simps)
-    ultimately show ?thesis using inj(2)[of "2*\<^sub>R x - 1" "2*\<^sub>R y - 1"] by auto
-  next assume as:"x $ 1 \<le> 1 / 2" "y $ 1 > 1 / 2"
-    hence "?g x \<in> path_image g1" "?g y \<in> path_image g2" unfolding path_image_def joinpaths_def
-      using xy(1,2)[unfolded mem_interval_1] by(auto simp add:vector_component_simps intro!: imageI)
-    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)[unfolded mem_interval_1]
-      apply(rule_tac ccontr) by(auto simp add:vector_component_simps field_simps Cart_eq)
-    ultimately have *:"?g x = pathstart g1" using assms(4) unfolding xy(3) by auto
-    hence "x = 0" unfolding pathstart_def joinpaths_def using as(1) and xy(1)[unfolded mem_interval_1]
-      using inj(1)[of "2 *\<^sub>R x" 0] by(auto simp add:vector_component_simps)
-    moreover have "y = 1" using * unfolding xy(3) assms(3)[THEN sym]
-      unfolding joinpaths_def pathfinish_def using as(2) and xy(2)[unfolded mem_interval_1]
-      using inj(2)[of "2 *\<^sub>R y - 1" 1] by (auto simp add:vector_component_simps Cart_eq)
-    ultimately show ?thesis by auto 
-  next assume as:"x $ 1 > 1 / 2" "y $ 1 \<le> 1 / 2"
-    hence "?g x \<in> path_image g2" "?g y \<in> path_image g1" unfolding path_image_def joinpaths_def
-      using xy(1,2)[unfolded mem_interval_1] by(auto simp add:vector_component_simps intro!: imageI)
-    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)[unfolded mem_interval_1]
-      apply(rule_tac ccontr) by(auto simp add:vector_component_simps field_simps Cart_eq)
-    ultimately have *:"?g y = pathstart g1" using assms(4) unfolding xy(3) by auto
-    hence "y = 0" unfolding pathstart_def joinpaths_def using as(2) and xy(2)[unfolded mem_interval_1]
-      using inj(1)[of "2 *\<^sub>R y" 0] by(auto simp add:vector_component_simps)
-    moreover have "x = 1" using * unfolding xy(3)[THEN sym] assms(3)[THEN sym]
-      unfolding joinpaths_def pathfinish_def using as(1) and xy(1)[unfolded mem_interval_1]
-      using inj(2)[of "2 *\<^sub>R x - 1" 1] by(auto simp add:vector_component_simps Cart_eq)
-    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)"
-  unfolding injective_path_def proof(rule,rule,rule) let ?g = "g1 +++ g2"
-  note inj = assms(1,2)[unfolded injective_path_def, rule_format]
-  fix x y assume xy:"x \<in> {0..1}" "y \<in> {0..1}" "(g1 +++ g2) x = (g1 +++ g2) y"
-  show "x = y" proof(cases "x$1 \<le> 1/2", case_tac[!] "y$1 \<le> 1/2", unfold not_le)
-    assume "x $ 1 \<le> 1 / 2" "y $ 1 \<le> 1 / 2" thus ?thesis using inj(1)[of "2*\<^sub>R x" "2*\<^sub>R y"] and xy
-      unfolding mem_interval_1 joinpaths_def by(auto simp add:vector_component_simps)
-  next assume "x $ 1 > 1 / 2" "y $ 1 > 1 / 2" thus ?thesis using inj(2)[of "2*\<^sub>R x - 1" "2*\<^sub>R y - 1"] and xy
-      unfolding mem_interval_1 joinpaths_def by(auto simp add:vector_component_simps)
-  next assume as:"x $ 1 \<le> 1 / 2" "y $ 1 > 1 / 2" 
-    hence "?g x \<in> path_image g1" "?g y \<in> path_image g2" unfolding path_image_def joinpaths_def
-      using xy(1,2)[unfolded mem_interval_1] by(auto simp add:vector_component_simps intro!: imageI)
-    hence "?g x = pathfinish g1" "?g y = pathstart g2" using assms(4) unfolding assms(3) xy(3) by auto
-    thus ?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 mem_interval_1
-      by(auto simp add:vector_component_simps Cart_eq forall_1)
-  next assume as:"x $ 1 > 1 / 2" "y $ 1 \<le> 1 / 2" 
-    hence "?g x \<in> path_image g2" "?g y \<in> path_image g1" unfolding path_image_def joinpaths_def
-      using xy(1,2)[unfolded mem_interval_1] by(auto simp add:vector_component_simps intro!: imageI)
-    hence "?g x = pathstart g2" "?g y = pathfinish g1" using assms(4) unfolding assms(3) xy(3) by auto
-    thus ?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 mem_interval_1
-      by(auto simp add:vector_component_simps forall_1 Cart_eq) qed 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. *}
-
-definition "shiftpath a (f::real^1 \<Rightarrow> real^'n) =
-  (\<lambda>x. if dest_vec1 (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"
-  unfolding pathstart_def shiftpath_def by auto
-
-(** move this **)
-declare forall_1[simp] ex_1[simp]
-
-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
-  by(auto simp add: vector_component_simps)
-
-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)
-
-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
-
-lemma path_shiftpath:
-  assumes "path g" "pathfinish g = pathstart g" "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 simp add: vector_component_simps)
-  have **:"\<And>x. x + a = 1 \<Longrightarrow> g (x + a - 1) = g (x + a)"
-    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_interval)+ 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 ** by(auto simp add:vector_component_simps field_simps Cart_eq) qed
-
-lemma shiftpath_shiftpath: assumes "pathfinish g = pathstart g" "a \<in> {0..1}" "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 simp add: vector_component_simps)
-
-lemma path_image_shiftpath:
-  assumes "a \<in> {0..1}" "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^1}" " \<forall>y\<in>{0..1} \<inter> {x. \<not> a $ 1 + x $ 1 \<le> 1}. g x \<noteq> g (a + y - 1)" 
-    hence "\<exists>y\<in>{0..1} \<inter> {x. a $ 1 + x $ 1 \<le> 1}. g x = g (a + y)" proof(cases "a \<le> x")
-      case False thus ?thesis apply(rule_tac x="1 + x - a" in bexI)
-        using as(1,2) and as(3)[THEN bspec[where x="1 + x - a"]] and assms(1)
-        by(auto simp add:vector_component_simps field_simps atomize_not) next
-      case True thus ?thesis using as(1-2) and assms(1) apply(rule_tac x="x - a" in bexI)
-        by(auto simp add:vector_component_simps field_simps) qed }
-  thus ?thesis using assms unfolding shiftpath_def path_image_def pathfinish_def pathstart_def 
-    by(auto simp add:vector_component_simps image_iff) qed
-
-subsection {* Special case of straight-line paths. *}
-
-definition
-  linepath :: "real ^ 'n::finite \<Rightarrow> real ^ 'n \<Rightarrow> real ^ 1 \<Rightarrow> real ^ 'n" where
-  "linepath a b = (\<lambda>x. (1 - dest_vec1 x) *\<^sub>R a + dest_vec1 x *\<^sub>R b)"
-
-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 continuous_linepath_at[intro]: "continuous (at x) (linepath a b)"
-  unfolding linepath_def
-  by (intro continuous_intros continuous_dest_vec1)
-
-lemma continuous_on_linepath[intro]: "continuous_on s (linepath a b)"
-  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_image_linepath[simp]: "path_image(linepath a b) = (closed_segment a b)"
-  unfolding path_image_def segment linepath_def apply (rule set_ext, rule) defer
-  unfolding mem_Collect_eq image_iff apply(erule exE) apply(rule_tac x="u *\<^sub>R 1" in bexI)
-  by(auto simp add:vector_component_simps)
-
-lemma reversepath_linepath[simp]:  "reversepath(linepath a b) = linepath b a"
-  unfolding reversepath_def linepath_def by(rule ext, auto simp add:vector_component_simps)
-
-lemma injective_path_linepath: assumes "a \<noteq> b" shows "injective_path(linepath a b)" proof- 
-  { obtain i where i:"a$i \<noteq> b$i" using assms[unfolded Cart_eq] by auto
-    fix x y::"real^1" assume "x $ 1 *\<^sub>R b + y $ 1 *\<^sub>R a = x $ 1 *\<^sub>R a + y $ 1 *\<^sub>R b"
-    hence "x$1 * (b$i - a$i) = y$1 * (b$i - a$i)" unfolding Cart_eq by(auto simp add:field_simps vector_component_simps)
-    hence "x = y" unfolding mult_cancel_right Cart_eq using i(1) by(auto simp add:field_simps) }
-  thus ?thesis unfolding injective_path_def linepath_def by(auto simp add:vector_component_simps 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)
-
-subsection {* Bounding a point away from a path. *}
-
-lemma not_on_path_ball: assumes "path g" "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]
-    using compact_path_image[THEN compact_imp_closed, OF assms(1)] by auto
-  thus ?thesis apply(rule_tac x="dist z a" in exI) using assms(2) by(auto intro!: dist_pos_lt) qed
-
-lemma not_on_path_cball: assumes "path g" "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" 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 qed
-
-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 
-
-lemma path_component_mem: assumes "path_component s x y" shows "x \<in> s" "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 
-  by(auto intro!:continuous_on_intros)    
-
-lemma path_component_refl_eq: "path_component s x x \<longleftrightarrow> x \<in> s"
-  by(auto intro!: path_component_mem path_component_refl) 
-
-lemma path_component_sym: "path_component s x y \<Longrightarrow> path_component s y x"
-  using assms unfolding path_component_def apply(erule exE) apply(rule_tac x="reversepath g" in exI) 
-  by(auto simp add: reversepath_simps)
-
-lemma path_component_trans: assumes "path_component s x y" "path_component s y z" shows "path_component s x z"
-  using assms unfolding path_component_def apply- apply(erule exE)+ apply(rule_tac x="g +++ ga" in exI) by(auto simp add: path_image_join)
-
-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. *}
-
-lemma path_component_set: "path_component s x = { y. (\<exists>g. path g \<and> path_image g \<subseteq> s \<and> pathstart g = x \<and> pathfinish g = y )}"
-  apply(rule set_ext) unfolding mem_Collect_eq unfolding mem_def path_component_def by auto
-
-lemma mem_path_component_set:"x \<in> path_component s y \<longleftrightarrow> path_component s y x" unfolding mem_def by auto
-
-lemma path_component_subset: "(path_component s x) \<subseteq> s"
-  apply(rule, rule path_component_mem(2)) by(auto simp add:mem_def)
-
-lemma path_component_eq_empty: "path_component s x = {} \<longleftrightarrow> x \<notin> s"
-  apply rule apply(drule equals0D[of _ x]) defer apply(rule equals0I) unfolding mem_path_component_set
-  apply(drule path_component_mem(1)) using path_component_refl by auto
-
-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)"
-
-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. path_component s x = s)" 
-  unfolding path_connected_component apply(rule, rule, rule, rule path_component_subset) 
-  unfolding subset_eq mem_path_component_set Ball_def mem_def by auto
-
-subsection {* Some useful lemmas about path-connectedness. *}
-
-lemma convex_imp_path_connected: assumes "convex s" shows "path_connected s"
-  unfolding path_connected_def apply(rule,rule,rule_tac x="linepath x y" in exI)
-  unfolding path_image_linepath using assms[unfolded convex_contains_segment] by auto
-
-lemma path_connected_imp_connected: assumes "path_connected s" shows "connected s"
-  unfolding connected_def not_ex apply(rule,rule,rule ccontr) unfolding not_not apply(erule 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"
-    using assms[unfolded path_connected_def,rule_format,of x1 x2] by auto
-  have *:"connected {0..1::real^1}" by(auto intro!: convex_connected convex_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 
-  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 by(auto intro!: exI)
-  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 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 qed
-
-lemma open_path_component: assumes "open s" shows "open(path_component s x)"
-  unfolding open_contains_ball proof
-  fix y assume as:"y \<in> path_component s x"
-  hence "y\<in>s" apply- apply(rule path_component_mem(2)) unfolding mem_def 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> path_component s x" apply(rule_tac x=e in exI) apply(rule,rule `e>0`,rule) unfolding mem_ball mem_path_component_set proof-
-    fix z assume "dist y z < e" thus "path_component s x z" 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`
-      using as[unfolded mem_def] by auto qed qed
-
-lemma open_non_path_component: assumes "open s" shows "open(s - path_component s x)" unfolding open_contains_ball proof
-  fix y assume as:"y\<in>s - path_component s x" 
-  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 - path_component s x" apply(rule_tac x=e in exI) apply(rule,rule `e>0`,rule,rule) defer proof(rule ccontr)
-    fix z assume "z\<in>ball y e" "\<not> z \<notin> path_component s x" 
-    hence "y \<in> path_component s x" unfolding not_not mem_path_component_set using `e>0` 
-      apply- apply(rule path_component_trans,assumption) apply(rule path_component_of_subset[OF e(2)])
-      apply(rule convex_imp_path_connected[OF convex_ball, unfolded path_connected_component, rule_format]) by auto
-    thus False using as by auto qed(insert e(2), auto) qed
-
-lemma connected_open_path_connected: assumes "open s" "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" show "y \<in> path_component s x" proof(rule ccontr)
-    assume "y \<notin> path_component s x" moreover
-    have "path_component s x \<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"path_component s x" "s - path_component s x"] by auto
-qed qed
-
-lemma path_connected_continuous_image:
-  assumes "continuous_on s f" "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)] ..
-  thus "\<exists>g. path g \<and> path_image g \<subseteq> f ` s \<and> pathstart g = x' \<and> pathfinish g = y'"
-    unfolding xy apply(rule_tac x="f \<circ> g" in exI) unfolding path_defs
-    using assms(1) by(auto intro!: continuous_on_compose continuous_on_subset[of _ _ "g ` {0..1}"]) qed
-
-lemma homeomorphic_path_connectedness:
-  "s homeomorphic t \<Longrightarrow> (path_connected s \<longleftrightarrow> path_connected t)"
-  unfolding homeomorphic_def homeomorphism_def apply(erule exE|erule conjE)+ apply rule
-  apply(drule_tac f=f in path_connected_continuous_image) prefer 3
-  apply(drule_tac f=g in path_connected_continuous_image) by auto
-
-lemma path_connected_empty: "path_connected {}"
-  unfolding path_connected_def by auto
-
-lemma path_connected_singleton: "path_connected {a}"
-  unfolding path_connected_def apply(rule,rule)
-  apply(rule_tac x="linepath a a" in exI) by(auto simp add:segment scaleR_left_diff_distrib)
-
-lemma path_connected_Un: assumes "path_connected s" "path_connected t" "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
-  thus "path_component (s \<union> t) x y" using as using assms(1-2)[unfolded path_connected_component] apply- 
-    apply(erule_tac[!] UnE)+ apply(rule_tac[2-3] path_component_trans[of _ _ z])
-    by(auto simp add:path_component_of_subset [OF Un_upper1] path_component_of_subset[OF Un_upper2]) qed
-
-subsection {* sphere is path-connected. *}
-
-lemma path_connected_punctured_universe:
- assumes "2 \<le> CARD('n::finite)" shows "path_connected((UNIV::(real^'n::finite) set) - {a})" proof-
-  obtain \<psi> where \<psi>:"bij_betw \<psi> {1..CARD('n)} (UNIV::'n set)" using ex_bij_betw_nat_finite_1[OF finite_UNIV] by auto
-  let ?U = "UNIV::(real^'n) set" let ?u = "?U - {0}"
-  let ?basis = "\<lambda>k. basis (\<psi> k)"
-  let ?A = "\<lambda>k. {x::real^'n. \<exists>i\<in>{1..k}. inner (basis (\<psi> i)) x \<noteq> 0}"
-  have "\<forall>k\<in>{2..CARD('n)}. path_connected (?A k)" proof
-    have *:"\<And>k. ?A (Suc k) = {x. inner (?basis (Suc k)) x < 0} \<union> {x. inner (?basis (Suc k)) x > 0} \<union> ?A k" apply(rule set_ext,rule) defer
-      apply(erule UnE)+  unfolding mem_Collect_eq apply(rule_tac[1-2] x="Suc k" in bexI)
-      by(auto elim!: ballE simp add: not_less le_Suc_eq)
-    fix k assume "k \<in> {2..CARD('n)}" thus "path_connected (?A k)" proof(induct k)
-      case (Suc k) show ?case proof(cases "k = 1")
-        case False from Suc have d:"k \<in> {1..CARD('n)}" "Suc k \<in> {1..CARD('n)}" by auto
-        hence "\<psi> k \<noteq> \<psi> (Suc k)" using \<psi>[unfolded bij_betw_def inj_on_def, THEN conjunct1, THEN bspec[where x=k]] by auto
-        hence **:"?basis k + ?basis (Suc k) \<in> {x. 0 < inner (?basis (Suc k)) x} \<inter> (?A k)" 
-          "?basis k - ?basis (Suc k) \<in> {x. 0 > inner (?basis (Suc k)) x} \<inter> ({x. 0 < inner (?basis (Suc k)) x} \<union> (?A k))" using d
-          by(auto simp add: inner_basis vector_component_simps intro!:bexI[where x=k])
-        show ?thesis unfolding * Un_assoc apply(rule path_connected_Un) defer apply(rule path_connected_Un) 
-          prefer 5 apply(rule_tac[1-2] convex_imp_path_connected, rule convex_halfspace_lt, rule convex_halfspace_gt)
-          apply(rule Suc(1)) apply(rule_tac[2-3] ccontr) using d ** False by auto
-      next case True hence d:"1\<in>{1..CARD('n)}" "2\<in>{1..CARD('n)}" using Suc(2) by auto
-        have ***:"Suc 1 = 2" by auto
-        have **:"\<And>s t P Q. s \<union> t \<union> {x. P x \<or> Q x} = (s \<union> {x. P x}) \<union> (t \<union> {x. Q x})" by auto
-        have "\<psi> 2 \<noteq> \<psi> (Suc 0)" apply(rule ccontr) using \<psi>[unfolded bij_betw_def inj_on_def, THEN conjunct1, THEN bspec[where x=2]] using assms by auto
-        thus ?thesis unfolding * True unfolding ** neq_iff bex_disj_distrib apply -
-          apply(rule path_connected_Un, rule_tac[1-2] path_connected_Un) defer 3 apply(rule_tac[1-4] convex_imp_path_connected) 
-          apply(rule_tac[5] x=" ?basis 1 + ?basis 2" in nequals0I)
-          apply(rule_tac[6] x="-?basis 1 + ?basis 2" in nequals0I)
-          apply(rule_tac[7] x="-?basis 1 - ?basis 2" in nequals0I)
-          using d unfolding *** by(auto intro!: convex_halfspace_gt convex_halfspace_lt, auto simp add:vector_component_simps inner_basis)
-  qed qed auto qed note lem = this
-
-  have ***:"\<And>x::real^'n. (\<exists>i\<in>{1..CARD('n)}. inner (basis (\<psi> i)) x \<noteq> 0) \<longleftrightarrow> (\<exists>i. inner (basis i) x \<noteq> 0)"
-    apply rule apply(erule bexE) apply(rule_tac x="\<psi> i" in exI) defer apply(erule exE) proof- 
-    fix x::"real^'n" and i assume as:"inner (basis i) x \<noteq> 0"
-    have "i\<in>\<psi> ` {1..CARD('n)}" using \<psi>[unfolded bij_betw_def, THEN conjunct2] by auto
-    then obtain j where "j\<in>{1..CARD('n)}" "\<psi> j = i" by auto
-    thus "\<exists>i\<in>{1..CARD('n)}. inner (basis (\<psi> i)) x \<noteq> 0" apply(rule_tac x=j in bexI) using as by auto qed auto
-  have *:"?U - {a} = (\<lambda>x. x + a) ` {x. x \<noteq> 0}" apply(rule set_ext) unfolding image_iff 
-    apply rule apply(rule_tac x="x - a" in bexI) by auto
-  have **:"\<And>x::real^'n. x\<noteq>0 \<longleftrightarrow> (\<exists>i. inner (basis i) x \<noteq> 0)" unfolding Cart_eq by(auto simp add: inner_basis)
-  show ?thesis unfolding * apply(rule path_connected_continuous_image) apply(rule continuous_on_intros)+ 
-    unfolding ** apply(rule lem[THEN bspec[where x="CARD('n)"], unfolded ***]) using assms by auto qed
-
-lemma path_connected_sphere: assumes "2 \<le> CARD('n::finite)" shows "path_connected {x::real^'n::finite. norm(x - a) = r}" proof(cases "r\<le>0")
-  case True thus ?thesis proof(cases "r=0") 
-    case False hence "{x::real^'n. norm(x - a) = r} = {}" using True by auto
-    thus ?thesis using path_connected_empty by auto
-  qed(auto intro!:path_connected_singleton) next
-  case False hence *:"{x::real^'n. norm(x - a) = r} = (\<lambda>x. a + r *\<^sub>R x) ` {x. norm x = 1}" unfolding not_le apply -apply(rule set_ext,rule)
-    unfolding image_iff apply(rule_tac x="(1/r) *\<^sub>R (x - a)" in bexI) unfolding mem_Collect_eq norm_scaleR by (auto simp add: scaleR_right_diff_distrib)
-  have **:"{x::real^'n. norm x = 1} = (\<lambda>x. (1/norm x) *\<^sub>R x) ` (UNIV - {0})" apply(rule set_ext,rule)
-    unfolding image_iff apply(rule_tac x=x in bexI) unfolding mem_Collect_eq by(auto split:split_if_asm)
-  have "continuous_on (UNIV - {0}) (\<lambda>x::real^'n. 1 / norm x)" unfolding o_def continuous_on_eq_continuous_within
-    apply(rule, rule continuous_at_within_inv[unfolded o_def inverse_eq_divide]) apply(rule continuous_at_within)
-    apply(rule continuous_at_norm[unfolded o_def]) by auto
-  thus ?thesis unfolding * ** using path_connected_punctured_universe[OF assms]
-    by(auto intro!: path_connected_continuous_image continuous_on_intros continuous_on_mul) qed
-
-lemma connected_sphere: "2 \<le> CARD('n) \<Longrightarrow> connected {x::real^'n::finite. norm(x - a) = r}"
-  using path_connected_sphere path_connected_imp_connected by auto
- 
-(** In continuous_at_vec1_norm : Use \<And> instead of \<forall>. **)
-
-end
--- a/src/HOL/Library/Determinants.thy	Fri Oct 23 14:33:07 2009 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,1087 +0,0 @@
-(* Title:      Determinants
-   Author:     Amine Chaieb, University of Cambridge
-*)
-
-header {* Traces, Determinant of square matrices and some properties *}
-
-theory Determinants
-imports Euclidean_Space Permutations
-begin
-
-subsection{* First some facts about products*}
-lemma setprod_insert_eq: "finite A \<Longrightarrow> setprod f (insert a A) = (if a \<in> A then setprod f A else f a * setprod f A)"
-apply clarsimp
-by(subgoal_tac "insert a A = A", auto)
-
-lemma setprod_add_split:
-  assumes mn: "(m::nat) <= n + 1"
-  shows "setprod f {m.. n+p} = setprod f {m .. n} * setprod f {n+1..n+p}"
-proof-
-  let ?A = "{m .. n+p}"
-  let ?B = "{m .. n}"
-  let ?C = "{n+1..n+p}"
-  from mn have un: "?B \<union> ?C = ?A" by auto
-  from mn have dj: "?B \<inter> ?C = {}" by auto
-  have f: "finite ?B" "finite ?C" by simp_all
-  from setprod_Un_disjoint[OF f dj, of f, unfolded un] show ?thesis .
-qed
-
-
-lemma setprod_offset: "setprod f {(m::nat) + p .. n + p} = setprod (\<lambda>i. f (i + p)) {m..n}"
-apply (rule setprod_reindex_cong[where f="op + p"])
-apply (auto simp add: image_iff Bex_def inj_on_def)
-apply arith
-apply (rule ext)
-apply (simp add: add_commute)
-done
-
-lemma setprod_singleton: "setprod f {x} = f x" by simp
-
-lemma setprod_singleton_nat_seg: "setprod f {n..n} = f (n::'a::order)" by simp
-
-lemma setprod_numseg: "setprod f {m..0} = (if m=0 then f 0 else 1)"
-  "setprod f {m .. Suc n} = (if m \<le> Suc n then f (Suc n) * setprod f {m..n}
-                             else setprod f {m..n})"
-  by (auto simp add: atLeastAtMostSuc_conv)
-
-lemma setprod_le: assumes fS: "finite S" and fg: "\<forall>x\<in>S. f x \<ge> 0 \<and> f x \<le> (g x :: 'a::ordered_idom)"
-  shows "setprod f S \<le> setprod g S"
-using fS fg
-apply(induct S)
-apply simp
-apply auto
-apply (rule mult_mono)
-apply (auto intro: setprod_nonneg)
-done
-
-  (* FIXME: In Finite_Set there is a useless further assumption *)
-lemma setprod_inversef: "finite A ==> setprod (inverse \<circ> f) A = (inverse (setprod f A) :: 'a:: {division_by_zero, field})"
-  apply (erule finite_induct)
-  apply (simp)
-  apply simp
-  done
-
-lemma setprod_le_1: assumes fS: "finite S" and f: "\<forall>x\<in>S. f x \<ge> 0 \<and> f x \<le> (1::'a::ordered_idom)"
-  shows "setprod f S \<le> 1"
-using setprod_le[OF fS f] unfolding setprod_1 .
-
-subsection{* Trace *}
-
-definition trace :: "'a::semiring_1^'n^'n \<Rightarrow> 'a" where
-  "trace A = setsum (\<lambda>i. ((A$i)$i)) (UNIV::'n set)"
-
-lemma trace_0: "trace(mat 0) = 0"
-  by (simp add: trace_def mat_def)
-
-lemma trace_I: "trace(mat 1 :: 'a::semiring_1^'n^'n) = of_nat(CARD('n))"
-  by (simp add: trace_def mat_def)
-
-lemma trace_add: "trace ((A::'a::comm_semiring_1^'n^'n) + B) = trace A + trace B"
-  by (simp add: trace_def setsum_addf)
-
-lemma trace_sub: "trace ((A::'a::comm_ring_1^'n^'n) - B) = trace A - trace B"
-  by (simp add: trace_def setsum_subtractf)
-
-lemma trace_mul_sym:"trace ((A::'a::comm_semiring_1^'n^'n) ** B) = trace (B**A)"
-  apply (simp add: trace_def matrix_matrix_mult_def)
-  apply (subst setsum_commute)
-  by (simp add: mult_commute)
-
-(* ------------------------------------------------------------------------- *)
-(* Definition of determinant.                                                *)
-(* ------------------------------------------------------------------------- *)
-
-definition det:: "'a::comm_ring_1^'n^'n \<Rightarrow> 'a" where
-  "det A = setsum (\<lambda>p. of_int (sign p) * setprod (\<lambda>i. A$i$p i) (UNIV :: 'n set)) {p. p permutes (UNIV :: 'n set)}"
-
-(* ------------------------------------------------------------------------- *)
-(* A few general lemmas we need below.                                       *)
-(* ------------------------------------------------------------------------- *)
-
-lemma setprod_permute:
-  assumes p: "p permutes S"
-  shows "setprod f S = setprod (f o p) S"
-proof-
-  {assume "\<not> finite S" hence ?thesis by simp}
-  moreover
-  {assume fS: "finite S"
-    then have ?thesis
-      apply (simp add: setprod_def cong del:strong_setprod_cong)
-      apply (rule ab_semigroup_mult.fold_image_permute)
-      apply (auto simp add: p)
-      apply unfold_locales
-      done}
-  ultimately show ?thesis by blast
-qed
-
-lemma setproduct_permute_nat_interval: "p permutes {m::nat .. n} ==> setprod f {m..n} = setprod (f o p) {m..n}"
-  by (blast intro!: setprod_permute)
-
-(* ------------------------------------------------------------------------- *)
-(* Basic determinant properties.                                             *)
-(* ------------------------------------------------------------------------- *)
-
-lemma det_transp: "det (transp A) = det (A::'a::comm_ring_1 ^'n^'n::finite)"
-proof-
-  let ?di = "\<lambda>A i j. A$i$j"
-  let ?U = "(UNIV :: 'n set)"
-  have fU: "finite ?U" by simp
-  {fix p assume p: "p \<in> {p. p permutes ?U}"
-    from p have pU: "p permutes ?U" by blast
-    have sth: "sign (inv p) = sign p"
-      by (metis sign_inverse fU p mem_def Collect_def permutation_permutes)
-    from permutes_inj[OF pU]
-    have pi: "inj_on p ?U" by (blast intro: subset_inj_on)
-    from permutes_image[OF pU]
-    have "setprod (\<lambda>i. ?di (transp A) i (inv p i)) ?U = setprod (\<lambda>i. ?di (transp A) i (inv p i)) (p ` ?U)" by simp
-    also have "\<dots> = setprod ((\<lambda>i. ?di (transp A) i (inv p i)) o p) ?U"
-      unfolding setprod_reindex[OF pi] ..
-    also have "\<dots> = setprod (\<lambda>i. ?di A i (p i)) ?U"
-    proof-
-      {fix i assume i: "i \<in> ?U"
-        from i permutes_inv_o[OF pU] permutes_in_image[OF pU]
-        have "((\<lambda>i. ?di (transp A) i (inv p i)) o p) i = ?di A i (p i)"
-          unfolding transp_def by (simp add: expand_fun_eq)}
-      then show "setprod ((\<lambda>i. ?di (transp A) i (inv p i)) o p) ?U = setprod (\<lambda>i. ?di A i (p i)) ?U" by (auto intro: setprod_cong)
-    qed
-    finally have "of_int (sign (inv p)) * (setprod (\<lambda>i. ?di (transp A) i (inv p i)) ?U) = of_int (sign p) * (setprod (\<lambda>i. ?di A i (p i)) ?U)" using sth
-      by simp}
-  then show ?thesis unfolding det_def apply (subst setsum_permutations_inverse)
-  apply (rule setsum_cong2) by blast
-qed
-
-lemma det_lowerdiagonal:
-  fixes A :: "'a::comm_ring_1^'n^'n::{finite,wellorder}"
-  assumes ld: "\<And>i j. i < j \<Longrightarrow> A$i$j = 0"
-  shows "det A = setprod (\<lambda>i. A$i$i) (UNIV:: 'n set)"
-proof-
-  let ?U = "UNIV:: 'n set"
-  let ?PU = "{p. p permutes ?U}"
-  let ?pp = "\<lambda>p. of_int (sign p) * setprod (\<lambda>i. A$i$p i) (UNIV :: 'n set)"
-  have fU: "finite ?U" by simp
-  from finite_permutations[OF fU] have fPU: "finite ?PU" .
-  have id0: "{id} \<subseteq> ?PU" by (auto simp add: permutes_id)
-  {fix p assume p: "p \<in> ?PU -{id}"
-    from p have pU: "p permutes ?U" and pid: "p \<noteq> id" by blast+
-    from permutes_natset_le[OF pU] pid obtain i where
-      i: "p i > i" by (metis not_le)
-    from ld[OF i] have ex:"\<exists>i \<in> ?U. A$i$p i = 0" by blast
-    from setprod_zero[OF fU ex] have "?pp p = 0" by simp}
-  then have p0: "\<forall>p \<in> ?PU -{id}. ?pp p = 0"  by blast
-  from setsum_mono_zero_cong_left[OF fPU id0 p0] show ?thesis
-    unfolding det_def by (simp add: sign_id)
-qed
-
-lemma det_upperdiagonal:
-  fixes A :: "'a::comm_ring_1^'n^'n::{finite,wellorder}"
-  assumes ld: "\<And>i j. i > j \<Longrightarrow> A$i$j = 0"
-  shows "det A = setprod (\<lambda>i. A$i$i) (UNIV:: 'n set)"
-proof-
-  let ?U = "UNIV:: 'n set"
-  let ?PU = "{p. p permutes ?U}"
-  let ?pp = "(\<lambda>p. of_int (sign p) * setprod (\<lambda>i. A$i$p i) (UNIV :: 'n set))"
-  have fU: "finite ?U" by simp
-  from finite_permutations[OF fU] have fPU: "finite ?PU" .
-  have id0: "{id} \<subseteq> ?PU" by (auto simp add: permutes_id)
-  {fix p assume p: "p \<in> ?PU -{id}"
-    from p have pU: "p permutes ?U" and pid: "p \<noteq> id" by blast+
-    from permutes_natset_ge[OF pU] pid obtain i where
-      i: "p i < i" by (metis not_le)
-    from ld[OF i] have ex:"\<exists>i \<in> ?U. A$i$p i = 0" by blast
-    from setprod_zero[OF fU ex] have "?pp p = 0" by simp}
-  then have p0: "\<forall>p \<in> ?PU -{id}. ?pp p = 0"  by blast
-  from   setsum_mono_zero_cong_left[OF fPU id0 p0] show ?thesis
-    unfolding det_def by (simp add: sign_id)
-qed
-
-lemma det_diagonal:
-  fixes A :: "'a::comm_ring_1^'n^'n::finite"
-  assumes ld: "\<And>i j. i \<noteq> j \<Longrightarrow> A$i$j = 0"
-  shows "det A = setprod (\<lambda>i. A$i$i) (UNIV::'n set)"
-proof-
-  let ?U = "UNIV:: 'n set"
-  let ?PU = "{p. p permutes ?U}"
-  let ?pp = "\<lambda>p. of_int (sign p) * setprod (\<lambda>i. A$i$p i) (UNIV :: 'n set)"
-  have fU: "finite ?U" by simp
-  from finite_permutations[OF fU] have fPU: "finite ?PU" .
-  have id0: "{id} \<subseteq> ?PU" by (auto simp add: permutes_id)
-  {fix p assume p: "p \<in> ?PU - {id}"
-    then have "p \<noteq> id" by simp
-    then obtain i where i: "p i \<noteq> i" unfolding expand_fun_eq by auto
-    from ld [OF i [symmetric]] have ex:"\<exists>i \<in> ?U. A$i$p i = 0" by blast
-    from setprod_zero [OF fU ex] have "?pp p = 0" by simp}
-  then have p0: "\<forall>p \<in> ?PU - {id}. ?pp p = 0"  by blast
-  from setsum_mono_zero_cong_left[OF fPU id0 p0] show ?thesis
-    unfolding det_def by (simp add: sign_id)
-qed
-
-lemma det_I: "det (mat 1 :: 'a::comm_ring_1^'n^'n::finite) = 1"
-proof-
-  let ?A = "mat 1 :: 'a::comm_ring_1^'n^'n"
-  let ?U = "UNIV :: 'n set"
-  let ?f = "\<lambda>i j. ?A$i$j"
-  {fix i assume i: "i \<in> ?U"
-    have "?f i i = 1" using i by (vector mat_def)}
-  hence th: "setprod (\<lambda>i. ?f i i) ?U = setprod (\<lambda>x. 1) ?U"
-    by (auto intro: setprod_cong)
-  {fix i j assume i: "i \<in> ?U" and j: "j \<in> ?U" and ij: "i \<noteq> j"
-    have "?f i j = 0" using i j ij by (vector mat_def) }
-  then have "det ?A = setprod (\<lambda>i. ?f i i) ?U" using det_diagonal
-    by blast
-  also have "\<dots> = 1" unfolding th setprod_1 ..
-  finally show ?thesis .
-qed
-
-lemma det_0: "det (mat 0 :: 'a::comm_ring_1^'n^'n::finite) = 0"
-  by (simp add: det_def setprod_zero)
-
-lemma det_permute_rows:
-  fixes A :: "'a::comm_ring_1^'n^'n::finite"
-  assumes p: "p permutes (UNIV :: 'n::finite set)"
-  shows "det(\<chi> i. A$p i :: 'a^'n^'n) = of_int (sign p) * det A"
-  apply (simp add: det_def setsum_right_distrib mult_assoc[symmetric])
-  apply (subst sum_permutations_compose_right[OF p])
-proof(rule setsum_cong2)
-  let ?U = "UNIV :: 'n set"
-  let ?PU = "{p. p permutes ?U}"
-  fix q assume qPU: "q \<in> ?PU"
-  have fU: "finite ?U" by simp
-  from qPU have q: "q permutes ?U" by blast
-  from p q have pp: "permutation p" and qp: "permutation q"
-    by (metis fU permutation_permutes)+
-  from permutes_inv[OF p] have ip: "inv p permutes ?U" .
-    have "setprod (\<lambda>i. A$p i$ (q o p) i) ?U = setprod ((\<lambda>i. A$p i$(q o p) i) o inv p) ?U"
-      by (simp only: setprod_permute[OF ip, symmetric])
-    also have "\<dots> = setprod (\<lambda>i. A $ (p o inv p) i $ (q o (p o inv p)) i) ?U"
-      by (simp only: o_def)
-    also have "\<dots> = setprod (\<lambda>i. A$i$q i) ?U" by (simp only: o_def permutes_inverses[OF p])
-    finally   have thp: "setprod (\<lambda>i. A$p i$ (q o p) i) ?U = setprod (\<lambda>i. A$i$q i) ?U"
-      by blast
-  show "of_int (sign (q o p)) * setprod (\<lambda>i. A$ p i$ (q o p) i) ?U = of_int (sign p) * of_int (sign q) * setprod (\<lambda>i. A$i$q i) ?U"
-    by (simp only: thp sign_compose[OF qp pp] mult_commute of_int_mult)
-qed
-
-lemma det_permute_columns:
-  fixes A :: "'a::comm_ring_1^'n^'n::finite"
-  assumes p: "p permutes (UNIV :: 'n set)"
-  shows "det(\<chi> i j. A$i$ p j :: 'a^'n^'n) = of_int (sign p) * det A"
-proof-
-  let ?Ap = "\<chi> i j. A$i$ p j :: 'a^'n^'n"
-  let ?At = "transp A"
-  have "of_int (sign p) * det A = det (transp (\<chi> i. transp A $ p i))"
-    unfolding det_permute_rows[OF p, of ?At] det_transp ..
-  moreover
-  have "?Ap = transp (\<chi> i. transp A $ p i)"
-    by (simp add: transp_def Cart_eq)
-  ultimately show ?thesis by simp
-qed
-
-lemma det_identical_rows:
-  fixes A :: "'a::ordered_idom^'n^'n::finite"
-  assumes ij: "i \<noteq> j"
-  and r: "row i A = row j A"
-  shows "det A = 0"
-proof-
-  have tha: "\<And>(a::'a) b. a = b ==> b = - a ==> a = 0"
-    by simp
-  have th1: "of_int (-1) = - 1" by (metis of_int_1 of_int_minus number_of_Min)
-  let ?p = "Fun.swap i j id"
-  let ?A = "\<chi> i. A $ ?p i"
-  from r have "A = ?A" by (simp add: Cart_eq row_def swap_def)
-  hence "det A = det ?A" by simp
-  moreover have "det A = - det ?A"
-    by (simp add: det_permute_rows[OF permutes_swap_id] sign_swap_id ij th1)
-  ultimately show "det A = 0" by (metis tha)
-qed
-
-lemma det_identical_columns:
-  fixes A :: "'a::ordered_idom^'n^'n::finite"
-  assumes ij: "i \<noteq> j"
-  and r: "column i A = column j A"
-  shows "det A = 0"
-apply (subst det_transp[symmetric])
-apply (rule det_identical_rows[OF ij])
-by (metis row_transp r)
-
-lemma det_zero_row:
-  fixes A :: "'a::{idom, ring_char_0}^'n^'n::finite"
-  assumes r: "row i A = 0"
-  shows "det A = 0"
-using r
-apply (simp add: row_def det_def Cart_eq)
-apply (rule setsum_0')
-apply (auto simp: sign_nz)
-done
-
-lemma det_zero_column:
-  fixes A :: "'a::{idom,ring_char_0}^'n^'n::finite"
-  assumes r: "column i A = 0"
-  shows "det A = 0"
-  apply (subst det_transp[symmetric])
-  apply (rule det_zero_row [of i])
-  by (metis row_transp r)
-
-lemma det_row_add:
-  fixes a b c :: "'n::finite \<Rightarrow> _ ^ 'n"
-  shows "det((\<chi> i. if i = k then a i + b i else c i)::'a::comm_ring_1^'n^'n) =
-             det((\<chi> i. if i = k then a i else c i)::'a::comm_ring_1^'n^'n) +
-             det((\<chi> i. if i = k then b i else c i)::'a::comm_ring_1^'n^'n)"
-unfolding det_def Cart_lambda_beta setsum_addf[symmetric]
-proof (rule setsum_cong2)
-  let ?U = "UNIV :: 'n set"
-  let ?pU = "{p. p permutes ?U}"
-  let ?f = "(\<lambda>i. if i = k then a i + b i else c i)::'n \<Rightarrow> 'a::comm_ring_1^'n"
-  let ?g = "(\<lambda> i. if i = k then a i else c i)::'n \<Rightarrow> 'a::comm_ring_1^'n"
-  let ?h = "(\<lambda> i. if i = k then b i else c i)::'n \<Rightarrow> 'a::comm_ring_1^'n"
-  fix p assume p: "p \<in> ?pU"
-  let ?Uk = "?U - {k}"
-  from p have pU: "p permutes ?U" by blast
-  have kU: "?U = insert k ?Uk" by blast
-  {fix j assume j: "j \<in> ?Uk"
-    from j have "?f j $ p j = ?g j $ p j" and "?f j $ p j= ?h j $ p j"
-      by simp_all}
-  then have th1: "setprod (\<lambda>i. ?f i $ p i) ?Uk = setprod (\<lambda>i. ?g i $ p i) ?Uk"
-    and th2: "setprod (\<lambda>i. ?f i $ p i) ?Uk = setprod (\<lambda>i. ?h i $ p i) ?Uk"
-    apply -
-    apply (rule setprod_cong, simp_all)+
-    done
-  have th3: "finite ?Uk" "k \<notin> ?Uk" by auto
-  have "setprod (\<lambda>i. ?f i $ p i) ?U = setprod (\<lambda>i. ?f i $ p i) (insert k ?Uk)"
-    unfolding kU[symmetric] ..
-  also have "\<dots> = ?f k $ p k  * setprod (\<lambda>i. ?f i $ p i) ?Uk"
-    apply (rule setprod_insert)
-    apply simp
-    by blast
-  also have "\<dots> = (a k $ p k * setprod (\<lambda>i. ?f i $ p i) ?Uk) + (b k$ p k * setprod (\<lambda>i. ?f i $ p i) ?Uk)" by (simp add: ring_simps)
-  also have "\<dots> = (a k $ p k * setprod (\<lambda>i. ?g i $ p i) ?Uk) + (b k$ p k * setprod (\<lambda>i. ?h i $ p i) ?Uk)" by (metis th1 th2)
-  also have "\<dots> = setprod (\<lambda>i. ?g i $ p i) (insert k ?Uk) + setprod (\<lambda>i. ?h i $ p i) (insert k ?Uk)"
-    unfolding  setprod_insert[OF th3] by simp
-  finally have "setprod (\<lambda>i. ?f i $ p i) ?U = setprod (\<lambda>i. ?g i $ p i) ?U + setprod (\<lambda>i. ?h i $ p i) ?U" unfolding kU[symmetric] .
-  then show "of_int (sign p) * setprod (\<lambda>i. ?f i $ p i) ?U = of_int (sign p) * setprod (\<lambda>i. ?g i $ p i) ?U + of_int (sign p) * setprod (\<lambda>i. ?h i $ p i) ?U"
-    by (simp add: ring_simps)
-qed
-
-lemma det_row_mul:
-  fixes a b :: "'n::finite \<Rightarrow> _ ^ 'n"
-  shows "det((\<chi> i. if i = k then c *s a i else b i)::'a::comm_ring_1^'n^'n) =
-             c* det((\<chi> i. if i = k then a i else b i)::'a::comm_ring_1^'n^'n)"
-
-unfolding det_def Cart_lambda_beta setsum_right_distrib
-proof (rule setsum_cong2)
-  let ?U = "UNIV :: 'n set"
-  let ?pU = "{p. p permutes ?U}"
-  let ?f = "(\<lambda>i. if i = k then c*s a i else b i)::'n \<Rightarrow> 'a::comm_ring_1^'n"
-  let ?g = "(\<lambda> i. if i = k then a i else b i)::'n \<Rightarrow> 'a::comm_ring_1^'n"
-  fix p assume p: "p \<in> ?pU"
-  let ?Uk = "?U - {k}"
-  from p have pU: "p permutes ?U" by blast
-  have kU: "?U = insert k ?Uk" by blast
-  {fix j assume j: "j \<in> ?Uk"
-    from j have "?f j $ p j = ?g j $ p j" by simp}
-  then have th1: "setprod (\<lambda>i. ?f i $ p i) ?Uk = setprod (\<lambda>i. ?g i $ p i) ?Uk"
-    apply -
-    apply (rule setprod_cong, simp_all)
-    done
-  have th3: "finite ?Uk" "k \<notin> ?Uk" by auto
-  have "setprod (\<lambda>i. ?f i $ p i) ?U = setprod (\<lambda>i. ?f i $ p i) (insert k ?Uk)"
-    unfolding kU[symmetric] ..
-  also have "\<dots> = ?f k $ p k  * setprod (\<lambda>i. ?f i $ p i) ?Uk"
-    apply (rule setprod_insert)
-    apply simp
-    by blast
-  also have "\<dots> = (c*s a k) $ p k * setprod (\<lambda>i. ?f i $ p i) ?Uk" by (simp add: ring_simps)
-  also have "\<dots> = c* (a k $ p k * setprod (\<lambda>i. ?g i $ p i) ?Uk)"
-    unfolding th1 by (simp add: mult_ac)
-  also have "\<dots> = c* (setprod (\<lambda>i. ?g i $ p i) (insert k ?Uk))"
-    unfolding  setprod_insert[OF th3] by simp
-  finally have "setprod (\<lambda>i. ?f i $ p i) ?U = c* (setprod (\<lambda>i. ?g i $ p i) ?U)" unfolding kU[symmetric] .
-  then show "of_int (sign p) * setprod (\<lambda>i. ?f i $ p i) ?U = c * (of_int (sign p) * setprod (\<lambda>i. ?g i $ p i) ?U)"
-    by (simp add: ring_simps)
-qed
-
-lemma det_row_0:
-  fixes b :: "'n::finite \<Rightarrow> _ ^ 'n"
-  shows "det((\<chi> i. if i = k then 0 else b i)::'a::comm_ring_1^'n^'n) = 0"
-using det_row_mul[of k 0 "\<lambda>i. 1" b]
-apply (simp)
-  unfolding vector_smult_lzero .
-
-lemma det_row_operation:
-  fixes A :: "'a::ordered_idom^'n^'n::finite"
-  assumes ij: "i \<noteq> j"
-  shows "det (\<chi> k. if k = i then row i A + c *s row j A else row k A) = det A"
-proof-
-  let ?Z = "(\<chi> k. if k = i then row j A else row k A) :: 'a ^'n^'n"
-  have th: "row i ?Z = row j ?Z" by (vector row_def)
-  have th2: "((\<chi> k. if k = i then row i A else row k A) :: 'a^'n^'n) = A"
-    by (vector row_def)
-  show ?thesis
-    unfolding det_row_add [of i] det_row_mul[of i] det_identical_rows[OF ij th] th2
-    by simp
-qed
-
-lemma det_row_span:
-  fixes A :: "'a:: ordered_idom^'n^'n::finite"
-  assumes x: "x \<in> span {row j A |j. j \<noteq> i}"
-  shows "det (\<chi> k. if k = i then row i A + x else row k A) = det A"
-proof-
-  let ?U = "UNIV :: 'n set"
-  let ?S = "{row j A |j. j \<noteq> i}"
-  let ?d = "\<lambda>x. det (\<chi> k. if k = i then x else row k A)"
-  let ?P = "\<lambda>x. ?d (row i A + x) = det A"
-  {fix k
-
-    have "(if k = i then row i A + 0 else row k A) = row k A" by simp}
-  then have P0: "?P 0"
-    apply -
-    apply (rule cong[of det, OF refl])
-    by (vector row_def)
-  moreover
-  {fix c z y assume zS: "z \<in> ?S" and Py: "?P y"
-    from zS obtain j where j: "z = row j A" "i \<noteq> j" by blast
-    let ?w = "row i A + y"
-    have th0: "row i A + (c*s z + y) = ?w + c*s z" by vector
-    have thz: "?d z = 0"
-      apply (rule det_identical_rows[OF j(2)])
-      using j by (vector row_def)
-    have "?d (row i A + (c*s z + y)) = ?d (?w + c*s z)" unfolding th0 ..
-    then have "?P (c*s z + y)" unfolding thz Py det_row_mul[of i] det_row_add[of i]
-      by simp }
-
-  ultimately show ?thesis
-    apply -
-    apply (rule span_induct_alt[of ?P ?S, OF P0])
-    apply blast
-    apply (rule x)
-    done
-qed
-
-(* ------------------------------------------------------------------------- *)
-(* May as well do this, though it's a bit unsatisfactory since it ignores    *)
-(* exact duplicates by considering the rows/columns as a set.                *)
-(* ------------------------------------------------------------------------- *)
-
-lemma det_dependent_rows:
-  fixes A:: "'a::ordered_idom^'n^'n::finite"
-  assumes d: "dependent (rows A)"
-  shows "det A = 0"
-proof-
-  let ?U = "UNIV :: 'n set"
-  from d obtain i where i: "row i A \<in> span (rows A - {row i A})"
-    unfolding dependent_def rows_def by blast
-  {fix j k assume jk: "j \<noteq> k"
-    and c: "row j A = row k A"
-    from det_identical_rows[OF jk c] have ?thesis .}
-  moreover
-  {assume H: "\<And> i j. i \<noteq> j \<Longrightarrow> row i A \<noteq> row j A"
-    have th0: "- row i A \<in> span {row j A|j. j \<noteq> i}"
-      apply (rule span_neg)
-      apply (rule set_rev_mp)
-      apply (rule i)
-      apply (rule span_mono)
-      using H i by (auto simp add: rows_def)
-    from det_row_span[OF th0]
-    have "det A = det (\<chi> k. if k = i then 0 *s 1 else row k A)"
-      unfolding right_minus vector_smult_lzero ..
-    with det_row_mul[of i "0::'a" "\<lambda>i. 1"]
-    have "det A = 0" by simp}
-  ultimately show ?thesis by blast
-qed
-
-lemma det_dependent_columns: assumes d: "dependent(columns (A::'a::ordered_idom^'n^'n::finite))" shows "det A = 0"
-by (metis d det_dependent_rows rows_transp det_transp)
-
-(* ------------------------------------------------------------------------- *)
-(* Multilinearity and the multiplication formula.                            *)
-(* ------------------------------------------------------------------------- *)
-
-lemma Cart_lambda_cong: "(\<And>x. f x = g x) \<Longrightarrow> (Cart_lambda f::'a^'n) = (Cart_lambda g :: 'a^'n)"
-  apply (rule iffD1[OF Cart_lambda_unique]) by vector
-
-lemma det_linear_row_setsum:
-  assumes fS: "finite S"
-  shows "det ((\<chi> i. if i = k then setsum (a i) S else c i)::'a::comm_ring_1^'n^'n::finite) = setsum (\<lambda>j. det ((\<chi> i. if i = k then a  i j else c i)::'a^'n^'n)) S"
-proof(induct rule: finite_induct[OF fS])
-  case 1 thus ?case apply simp  unfolding setsum_empty det_row_0[of k] ..
-next
-  case (2 x F)
-  then  show ?case by (simp add: det_row_add cong del: if_weak_cong)
-qed
-
-lemma finite_bounded_functions:
-  assumes fS: "finite S"
-  shows "finite {f. (\<forall>i \<in> {1.. (k::nat)}. f i \<in> S) \<and> (\<forall>i. i \<notin> {1 .. k} \<longrightarrow> f i = i)}"
-proof(induct k)
-  case 0
-  have th: "{f. \<forall>i. f i = i} = {id}" by (auto intro: ext)
-  show ?case by (auto simp add: th)
-next
-  case (Suc k)
-  let ?f = "\<lambda>(y::nat,g) i. if i = Suc k then y else g i"
-  let ?S = "?f ` (S \<times> {f. (\<forall>i\<in>{1..k}. f i \<in> S) \<and> (\<forall>i. i \<notin> {1..k} \<longrightarrow> f i = i)})"
-  have "?S = {f. (\<forall>i\<in>{1.. Suc k}. f i \<in> S) \<and> (\<forall>i. i \<notin> {1.. Suc k} \<longrightarrow> f i = i)}"
-    apply (auto simp add: image_iff)
-    apply (rule_tac x="x (Suc k)" in bexI)
-    apply (rule_tac x = "\<lambda>i. if i = Suc k then i else x i" in exI)
-    apply (auto intro: ext)
-    done
-  with finite_imageI[OF finite_cartesian_product[OF fS Suc.hyps(1)], of ?f]
-  show ?case by metis
-qed
-
-
-lemma eq_id_iff[simp]: "(\<forall>x. f x = x) = (f = id)" by (auto intro: ext)
-
-lemma det_linear_rows_setsum_lemma:
-  assumes fS: "finite S" and fT: "finite T"
-  shows "det((\<chi> i. if i \<in> T then setsum (a i) S else c i):: 'a::comm_ring_1^'n^'n::finite) =
-             setsum (\<lambda>f. det((\<chi> i. if i \<in> T then a i (f i) else c i)::'a^'n^'n))
-                 {f. (\<forall>i \<in> T. f i \<in> S) \<and> (\<forall>i. i \<notin> T \<longrightarrow> f i = i)}"
-using fT
-proof(induct T arbitrary: a c set: finite)
-  case empty
-  have th0: "\<And>x y. (\<chi> i. if i \<in> {} then x i else y i) = (\<chi> i. y i)" by vector
-  from "empty.prems"  show ?case unfolding th0 by simp
-next
-  case (insert z T a c)
-  let ?F = "\<lambda>T. {f. (\<forall>i \<in> T. f i \<in> S) \<and> (\<forall>i. i \<notin> T \<longrightarrow> f i = i)}"
-  let ?h = "\<lambda>(y,g) i. if i = z then y else g i"
-  let ?k = "\<lambda>h. (h(z),(\<lambda>i. if i = z then i else h i))"
-  let ?s = "\<lambda> k a c f. det((\<chi> i. if i \<in> T then a i (f i) else c i)::'a^'n^'n)"
-  let ?c = "\<lambda>i. if i = z then a i j else c i"
-  have thif: "\<And>a b c d. (if a \<or> b then c else d) = (if a then c else if b then c else d)" by simp
-  have thif2: "\<And>a b c d e. (if a then b else if c then d else e) =
-     (if c then (if a then b else d) else (if a then b else e))" by simp
-  from `z \<notin> T` have nz: "\<And>i. i \<in> T \<Longrightarrow> i = z \<longleftrightarrow> False" by auto
-  have "det (\<chi> i. if i \<in> insert z T then setsum (a i) S else c i) =
-        det (\<chi> i. if i = z then setsum (a i) S
-                 else if i \<in> T then setsum (a i) S else c i)"
-    unfolding insert_iff thif ..
-  also have "\<dots> = (\<Sum>j\<in>S. det (\<chi> i. if i \<in> T then setsum (a i) S
-                    else if i = z then a i j else c i))"
-    unfolding det_linear_row_setsum[OF fS]
-    apply (subst thif2)
-    using nz by (simp cong del: if_weak_cong cong add: if_cong)
-  finally have tha:
-    "det (\<chi> i. if i \<in> insert z T then setsum (a i) S else c i) =
-     (\<Sum>(j, f)\<in>S \<times> ?F T. det (\<chi> i. if i \<in> T then a i (f i)
-                                else if i = z then a i j
-                                else c i))"
-    unfolding  insert.hyps unfolding setsum_cartesian_product by blast
-  show ?case unfolding tha
-    apply(rule setsum_eq_general_reverses[where h= "?h" and k= "?k"],
-      blast intro: finite_cartesian_product fS finite,
-      blast intro: finite_cartesian_product fS finite)
-    using `z \<notin> T`
-    apply (auto intro: ext)
-    apply (rule cong[OF refl[of det]])
-    by vector
-qed
-
-lemma det_linear_rows_setsum:
-  assumes fS: "finite (S::'n::finite set)"
-  shows "det (\<chi> i. setsum (a i) S) = setsum (\<lambda>f. det (\<chi> i. a i (f i) :: 'a::comm_ring_1 ^ 'n^'n::finite)) {f. \<forall>i. f i \<in> S}"
-proof-
-  have th0: "\<And>x y. ((\<chi> i. if i \<in> (UNIV:: 'n set) then x i else y i) :: 'a^'n^'n) = (\<chi> i. x i)" by vector
-
-  from det_linear_rows_setsum_lemma[OF fS, of "UNIV :: 'n set" a, unfolded th0, OF finite] show ?thesis by simp
-qed
-
-lemma matrix_mul_setsum_alt:
-  fixes A B :: "'a::comm_ring_1^'n^'n::finite"
-  shows "A ** B = (\<chi> i. setsum (\<lambda>k. A$i$k *s B $ k) (UNIV :: 'n set))"
-  by (vector matrix_matrix_mult_def setsum_component)
-
-lemma det_rows_mul:
-  "det((\<chi> i. c i *s a i)::'a::comm_ring_1^'n^'n::finite) =
-  setprod (\<lambda>i. c i) (UNIV:: 'n set) * det((\<chi> i. a i)::'a^'n^'n)"
-proof (simp add: det_def setsum_right_distrib cong add: setprod_cong, rule setsum_cong2)
-  let ?U = "UNIV :: 'n set"
-  let ?PU = "{p. p permutes ?U}"
-  fix p assume pU: "p \<in> ?PU"
-  let ?s = "of_int (sign p)"
-  from pU have p: "p permutes ?U" by blast
-  have "setprod (\<lambda>i. c i * a i $ p i) ?U = setprod c ?U * setprod (\<lambda>i. a i $ p i) ?U"
-    unfolding setprod_timesf ..
-  then show "?s * (\<Prod>xa\<in>?U. c xa * a xa $ p xa) =
-        setprod c ?U * (?s* (\<Prod>xa\<in>?U. a xa $ p xa))" by (simp add: ring_simps)
-qed
-
-lemma det_mul:
-  fixes A B :: "'a::ordered_idom^'n^'n::finite"
-  shows "det (A ** B) = det A * det B"
-proof-
-  let ?U = "UNIV :: 'n set"
-  let ?F = "{f. (\<forall>i\<in> ?U. f i \<in> ?U) \<and> (\<forall>i. i \<notin> ?U \<longrightarrow> f i = i)}"
-  let ?PU = "{p. p permutes ?U}"
-  have fU: "finite ?U" by simp
-  have fF: "finite ?F" by (rule finite)
-  {fix p assume p: "p permutes ?U"
-
-    have "p \<in> ?F" unfolding mem_Collect_eq permutes_in_image[OF p]
-      using p[unfolded permutes_def] by simp}
-  then have PUF: "?PU \<subseteq> ?F"  by blast
-  {fix f assume fPU: "f \<in> ?F - ?PU"
-    have fUU: "f ` ?U \<subseteq> ?U" using fPU by auto
-    from fPU have f: "\<forall>i \<in> ?U. f i \<in> ?U"
-      "\<forall>i. i \<notin> ?U \<longrightarrow> f i = i" "\<not>(\<forall>y. \<exists>!x. f x = y)" unfolding permutes_def
-      by auto
-
-    let ?A = "(\<chi> i. A$i$f i *s B$f i) :: 'a^'n^'n"
-    let ?B = "(\<chi> i. B$f i) :: 'a^'n^'n"
-    {assume fni: "\<not> inj_on f ?U"
-      then obtain i j where ij: "f i = f j" "i \<noteq> j"
-        unfolding inj_on_def by blast
-      from ij
-      have rth: "row i ?B = row j ?B" by (vector row_def)
-      from det_identical_rows[OF ij(2) rth]
-      have "det (\<chi> i. A$i$f i *s B$f i) = 0"
-        unfolding det_rows_mul by simp}
-    moreover
-    {assume fi: "inj_on f ?U"
-      from f fi have fith: "\<And>i j. f i = f j \<Longrightarrow> i = j"
-        unfolding inj_on_def by metis
-      note fs = fi[unfolded surjective_iff_injective_gen[OF fU fU refl fUU, symmetric]]
-
-      {fix y
-        from fs f have "\<exists>x. f x = y" by blast
-        then obtain x where x: "f x = y" by blast
-        {fix z assume z: "f z = y" from fith x z have "z = x" by metis}
-        with x have "\<exists>!x. f x = y" by blast}
-      with f(3) have "det (\<chi> i. A$i$f i *s B$f i) = 0" by blast}
-    ultimately have "det (\<chi> i. A$i$f i *s B$f i) = 0" by blast}
-  hence zth: "\<forall> f\<in> ?F - ?PU. det (\<chi> i. A$i$f i *s B$f i) = 0" by simp
-  {fix p assume pU: "p \<in> ?PU"
-    from pU have p: "p permutes ?U" by blast
-    let ?s = "\<lambda>p. of_int (sign p)"
-    let ?f = "\<lambda>q. ?s p * (\<Prod>i\<in> ?U. A $ i $ p i) *
-               (?s q * (\<Prod>i\<in> ?U. B $ i $ q i))"
-    have "(setsum (\<lambda>q. ?s q *
-            (\<Prod>i\<in> ?U. (\<chi> i. A $ i $ p i *s B $ p i :: 'a^'n^'n) $ i $ q i)) ?PU) =
-        (setsum (\<lambda>q. ?s p * (\<Prod>i\<in> ?U. A $ i $ p i) *
-               (?s q * (\<Prod>i\<in> ?U. B $ i $ q i))) ?PU)"
-      unfolding sum_permutations_compose_right[OF permutes_inv[OF p], of ?f]
-    proof(rule setsum_cong2)
-      fix q assume qU: "q \<in> ?PU"
-      hence q: "q permutes ?U" by blast
-      from p q have pp: "permutation p" and pq: "permutation q"
-        unfolding permutation_permutes by auto
-      have th00: "of_int (sign p) * of_int (sign p) = (1::'a)"
-        "\<And>a. of_int (sign p) * (of_int (sign p) * a) = a"
-        unfolding mult_assoc[symmetric] unfolding of_int_mult[symmetric]
-        by (simp_all add: sign_idempotent)
-      have ths: "?s q = ?s p * ?s (q o inv p)"
-        using pp pq permutation_inverse[OF pp] sign_inverse[OF pp]
-        by (simp add:  th00 mult_ac sign_idempotent sign_compose)
-      have th001: "setprod (\<lambda>i. B$i$ q (inv p i)) ?U = setprod ((\<lambda>i. B$i$ q (inv p i)) o p) ?U"
-        by (rule setprod_permute[OF p])
-      have thp: "setprod (\<lambda>i. (\<chi> i. A$i$p i *s B$p i :: 'a^'n^'n) $i $ q i) ?U = setprod (\<lambda>i. A$i$p i) ?U * setprod (\<lambda>i. B$i$ q (inv p i)) ?U"
-        unfolding th001 setprod_timesf[symmetric] o_def permutes_inverses[OF p]
-        apply (rule setprod_cong[OF refl])
-        using permutes_in_image[OF q] by vector
-      show "?s q * setprod (\<lambda>i. (((\<chi> i. A$i$p i *s B$p i) :: 'a^'n^'n)$i$q i)) ?U = ?s p * (setprod (\<lambda>i. A$i$p i) ?U) * (?s (q o inv p) * setprod (\<lambda>i. B$i$(q o inv p) i) ?U)"
-        using ths thp pp pq permutation_inverse[OF pp] sign_inverse[OF pp]
-        by (simp add: sign_nz th00 ring_simps sign_idempotent sign_compose)
-    qed
-  }
-  then have th2: "setsum (\<lambda>f. det (\<chi> i. A$i$f i *s B$f i)) ?PU = det A * det B"
-    unfolding det_def setsum_product
-    by (rule setsum_cong2)
-  have "det (A**B) = setsum (\<lambda>f.  det (\<chi> i. A $ i $ f i *s B $ f i)) ?F"
-    unfolding matrix_mul_setsum_alt det_linear_rows_setsum[OF fU] by simp
-  also have "\<dots> = setsum (\<lambda>f. det (\<chi> i. A$i$f i *s B$f i)) ?PU"
-    using setsum_mono_zero_cong_left[OF fF PUF zth, symmetric]
-    unfolding det_rows_mul by auto
-  finally show ?thesis unfolding th2 .
-qed
-
-(* ------------------------------------------------------------------------- *)
-(* Relation to invertibility.                                                *)
-(* ------------------------------------------------------------------------- *)
-
-lemma invertible_left_inverse:
-  fixes A :: "real^'n^'n::finite"
-  shows "invertible A \<longleftrightarrow> (\<exists>(B::real^'n^'n). B** A = mat 1)"
-  by (metis invertible_def matrix_left_right_inverse)
-
-lemma invertible_righ_inverse:
-  fixes A :: "real^'n^'n::finite"
-  shows "invertible A \<longleftrightarrow> (\<exists>(B::real^'n^'n). A** B = mat 1)"
-  by (metis invertible_def matrix_left_right_inverse)
-
-lemma invertible_det_nz:
-  fixes A::"real ^'n^'n::finite"
-  shows "invertible A \<longleftrightarrow> det A \<noteq> 0"
-proof-
-  {assume "invertible A"
-    then obtain B :: "real ^'n^'n" where B: "A ** B = mat 1"
-      unfolding invertible_righ_inverse by blast
-    hence "det (A ** B) = det (mat 1 :: real ^'n^'n)" by simp
-    hence "det A \<noteq> 0"
-      apply (simp add: det_mul det_I) by algebra }
-  moreover
-  {assume H: "\<not> invertible A"
-    let ?U = "UNIV :: 'n set"
-    have fU: "finite ?U" by simp
-    from H obtain c i where c: "setsum (\<lambda>i. c i *s row i A) ?U = 0"
-      and iU: "i \<in> ?U" and ci: "c i \<noteq> 0"
-      unfolding invertible_righ_inverse
-      unfolding matrix_right_invertible_independent_rows by blast
-    have stupid: "\<And>(a::real^'n) b. a + b = 0 \<Longrightarrow> -a = b"
-      apply (drule_tac f="op + (- a)" in cong[OF refl])
-      apply (simp only: ab_left_minus add_assoc[symmetric])
-      apply simp
-      done
-    from c ci
-    have thr0: "- row i A = setsum (\<lambda>j. (1/ c i) *s (c j *s row j A)) (?U - {i})"
-      unfolding setsum_diff1'[OF fU iU] setsum_cmul
-      apply -
-      apply (rule vector_mul_lcancel_imp[OF ci])
-      apply (auto simp add: vector_smult_assoc vector_smult_rneg field_simps)
-      unfolding stupid ..
-    have thr: "- row i A \<in> span {row j A| j. j \<noteq> i}"
-      unfolding thr0
-      apply (rule span_setsum)
-      apply simp
-      apply (rule ballI)
-      apply (rule span_mul)+
-      apply (rule span_superset)
-      apply auto
-      done
-    let ?B = "(\<chi> k. if k = i then 0 else row k A) :: real ^'n^'n"
-    have thrb: "row i ?B = 0" using iU by (vector row_def)
-    have "det A = 0"
-      unfolding det_row_span[OF thr, symmetric] right_minus
-      unfolding  det_zero_row[OF thrb]  ..}
-  ultimately show ?thesis by blast
-qed
-
-(* ------------------------------------------------------------------------- *)
-(* Cramer's rule.                                                            *)
-(* ------------------------------------------------------------------------- *)
-
-lemma cramer_lemma_transp:
-  fixes A:: "'a::ordered_idom^'n^'n::finite" and x :: "'a ^'n::finite"
-  shows "det ((\<chi> i. if i = k then setsum (\<lambda>i. x$i *s row i A) (UNIV::'n set)
-                           else row i A)::'a^'n^'n) = x$k * det A"
-  (is "?lhs = ?rhs")
-proof-
-  let ?U = "UNIV :: 'n set"
-  let ?Uk = "?U - {k}"
-  have U: "?U = insert k ?Uk" by blast
-  have fUk: "finite ?Uk" by simp
-  have kUk: "k \<notin> ?Uk" by simp
-  have th00: "\<And>k s. x$k *s row k A + s = (x$k - 1) *s row k A + row k A + s"
-    by (vector ring_simps)
-  have th001: "\<And>f k . (\<lambda>x. if x = k then f k else f x) = f" by (auto intro: ext)
-  have "(\<chi> i. row i A) = A" by (vector row_def)
-  then have thd1: "det (\<chi> i. row i A) = det A"  by simp
-  have thd0: "det (\<chi> i. if i = k then row k A + (\<Sum>i \<in> ?Uk. x $ i *s row i A) else row i A) = det A"
-    apply (rule det_row_span)
-    apply (rule span_setsum[OF fUk])
-    apply (rule ballI)
-    apply (rule span_mul)
-    apply (rule span_superset)
-    apply auto
-    done
-  show "?lhs = x$k * det A"
-    apply (subst U)
-    unfolding setsum_insert[OF fUk kUk]
-    apply (subst th00)
-    unfolding add_assoc
-    apply (subst det_row_add)
-    unfolding thd0
-    unfolding det_row_mul
-    unfolding th001[of k "\<lambda>i. row i A"]
-    unfolding thd1  by (simp add: ring_simps)
-qed
-
-lemma cramer_lemma:
-  fixes A :: "'a::ordered_idom ^'n^'n::finite"
-  shows "det((\<chi> i j. if j = k then (A *v x)$i else A$i$j):: 'a^'n^'n) = x$k * det A"
-proof-
-  let ?U = "UNIV :: 'n set"
-  have stupid: "\<And>c. setsum (\<lambda>i. c i *s row i (transp A)) ?U = setsum (\<lambda>i. c i *s column i A) ?U"
-    by (auto simp add: row_transp intro: setsum_cong2)
-  show ?thesis  unfolding matrix_mult_vsum
-  unfolding cramer_lemma_transp[of k x "transp A", unfolded det_transp, symmetric]
-  unfolding stupid[of "\<lambda>i. x$i"]
-  apply (subst det_transp[symmetric])
-  apply (rule cong[OF refl[of det]]) by (vector transp_def column_def row_def)
-qed
-
-lemma cramer:
-  fixes A ::"real^'n^'n::finite"
-  assumes d0: "det A \<noteq> 0"
-  shows "A *v x = b \<longleftrightarrow> x = (\<chi> k. det(\<chi> i j. if j=k then b$i else A$i$j :: real^'n^'n) / det A)"
-proof-
-  from d0 obtain B where B: "A ** B = mat 1" "B ** A = mat 1"
-    unfolding invertible_det_nz[symmetric] invertible_def by blast
-  have "(A ** B) *v b = b" by (simp add: B matrix_vector_mul_lid)
-  hence "A *v (B *v b) = b" by (simp add: matrix_vector_mul_assoc)
-  then have xe: "\<exists>x. A*v x = b" by blast
-  {fix x assume x: "A *v x = b"
-  have "x = (\<chi> k. det(\<chi> i j. if j=k then b$i else A$i$j :: real^'n^'n) / det A)"
-    unfolding x[symmetric]
-    using d0 by (simp add: Cart_eq cramer_lemma field_simps)}
-  with xe show ?thesis by auto
-qed
-
-(* ------------------------------------------------------------------------- *)
-(* Orthogonality of a transformation and matrix.                             *)
-(* ------------------------------------------------------------------------- *)
-
-definition "orthogonal_transformation f \<longleftrightarrow> linear f \<and> (\<forall>v w. f v \<bullet> f w = v \<bullet> w)"
-
-lemma orthogonal_transformation: "orthogonal_transformation f \<longleftrightarrow> linear f \<and> (\<forall>(v::real ^_). norm (f v) = norm v)"
-  unfolding orthogonal_transformation_def
-  apply auto
-  apply (erule_tac x=v in allE)+
-  apply (simp add: real_vector_norm_def)
-  by (simp add: dot_norm  linear_add[symmetric])
-
-definition "orthogonal_matrix (Q::'a::semiring_1^'n^'n) \<longleftrightarrow> transp Q ** Q = mat 1 \<and> Q ** transp Q = mat 1"
-
-lemma orthogonal_matrix: "orthogonal_matrix (Q:: real ^'n^'n::finite)  \<longleftrightarrow> transp Q ** Q = mat 1"
-  by (metis matrix_left_right_inverse orthogonal_matrix_def)
-
-lemma orthogonal_matrix_id: "orthogonal_matrix (mat 1 :: _^'n^'n::finite)"
-  by (simp add: orthogonal_matrix_def transp_mat matrix_mul_lid)
-
-lemma orthogonal_matrix_mul:
-  fixes A :: "real ^'n^'n::finite"
-  assumes oA : "orthogonal_matrix A"
-  and oB: "orthogonal_matrix B"
-  shows "orthogonal_matrix(A ** B)"
-  using oA oB
-  unfolding orthogonal_matrix matrix_transp_mul
-  apply (subst matrix_mul_assoc)
-  apply (subst matrix_mul_assoc[symmetric])
-  by (simp add: matrix_mul_rid)
-
-lemma orthogonal_transformation_matrix:
-  fixes f:: "real^'n \<Rightarrow> real^'n::finite"
-  shows "orthogonal_transformation f \<longleftrightarrow> linear f \<and> orthogonal_matrix(matrix f)"
-  (is "?lhs \<longleftrightarrow> ?rhs")
-proof-
-  let ?mf = "matrix f"
-  let ?ot = "orthogonal_transformation f"
-  let ?U = "UNIV :: 'n set"
-  have fU: "finite ?U" by simp
-  let ?m1 = "mat 1 :: real ^'n^'n"
-  {assume ot: ?ot
-    from ot have lf: "linear f" and fd: "\<forall>v w. f v \<bullet> f w = v \<bullet> w"
-      unfolding  orthogonal_transformation_def orthogonal_matrix by blast+
-    {fix i j
-      let ?A = "transp ?mf ** ?mf"
-      have th0: "\<And>b (x::'a::comm_ring_1). (if b then 1 else 0)*x = (if b then x else 0)"
-        "\<And>b (x::'a::comm_ring_1). x*(if b then 1 else 0) = (if b then x else 0)"
-        by simp_all
-      from fd[rule_format, of "basis i" "basis j", unfolded matrix_works[OF lf, symmetric] dot_matrix_vector_mul]
-      have "?A$i$j = ?m1 $ i $ j"
-        by (simp add: dot_def matrix_matrix_mult_def columnvector_def rowvector_def basis_def th0 setsum_delta[OF fU] mat_def)}
-    hence "orthogonal_matrix ?mf" unfolding orthogonal_matrix by vector
-    with lf have ?rhs by blast}
-  moreover
-  {assume lf: "linear f" and om: "orthogonal_matrix ?mf"
-    from lf om have ?lhs
-      unfolding orthogonal_matrix_def norm_eq orthogonal_transformation
-      unfolding matrix_works[OF lf, symmetric]
-      apply (subst dot_matrix_vector_mul)
-      by (simp add: dot_matrix_product matrix_mul_lid)}
-  ultimately show ?thesis by blast
-qed
-
-lemma det_orthogonal_matrix:
-  fixes Q:: "'a::ordered_idom^'n^'n::finite"
-  assumes oQ: "orthogonal_matrix Q"
-  shows "det Q = 1 \<or> det Q = - 1"
-proof-
-
-  have th: "\<And>x::'a. x = 1 \<or> x = - 1 \<longleftrightarrow> x*x = 1" (is "\<And>x::'a. ?ths x")
-  proof-
-    fix x:: 'a
-    have th0: "x*x - 1 = (x - 1)*(x + 1)" by (simp add: ring_simps)
-    have th1: "\<And>(x::'a) y. x = - y \<longleftrightarrow> x + y = 0"
-      apply (subst eq_iff_diff_eq_0) by simp
-    have "x*x = 1 \<longleftrightarrow> x*x - 1 = 0" by simp
-    also have "\<dots> \<longleftrightarrow> x = 1 \<or> x = - 1" unfolding th0 th1 by simp
-    finally show "?ths x" ..
-  qed
-  from oQ have "Q ** transp Q = mat 1" by (metis orthogonal_matrix_def)
-  hence "det (Q ** transp Q) = det (mat 1:: 'a^'n^'n)" by simp
-  hence "det Q * det Q = 1" by (simp add: det_mul det_I det_transp)
-  then show ?thesis unfolding th .
-qed
-
-(* ------------------------------------------------------------------------- *)
-(* Linearity of scaling, and hence isometry, that preserves origin.          *)
-(* ------------------------------------------------------------------------- *)
-lemma scaling_linear:
-  fixes f :: "real ^'n \<Rightarrow> real ^'n::finite"
-  assumes f0: "f 0 = 0" and fd: "\<forall>x y. dist (f x) (f y) = c * dist x y"
-  shows "linear f"
-proof-
-  {fix v w
-    {fix x note fd[rule_format, of x 0, unfolded dist_norm f0 diff_0_right] }
-    note th0 = this
-    have "f v \<bullet> f w = c^2 * (v \<bullet> w)"
-      unfolding dot_norm_neg dist_norm[symmetric]
-      unfolding th0 fd[rule_format] by (simp add: power2_eq_square field_simps)}
-  note fc = this
-  show ?thesis unfolding linear_def vector_eq
-    by (simp add: dot_lmult dot_ladd dot_rmult dot_radd fc ring_simps)
-qed
-
-lemma isometry_linear:
-  "f (0:: real^'n) = (0:: real^'n::finite) \<Longrightarrow> \<forall>x y. dist(f x) (f y) = dist x y
-        \<Longrightarrow> linear f"
-by (rule scaling_linear[where c=1]) simp_all
-
-(* ------------------------------------------------------------------------- *)
-(* Hence another formulation of orthogonal transformation.                   *)
-(* ------------------------------------------------------------------------- *)
-
-lemma orthogonal_transformation_isometry:
-  "orthogonal_transformation f \<longleftrightarrow> f(0::real^'n) = (0::real^'n::finite) \<and> (\<forall>x y. dist(f x) (f y) = dist x y)"
-  unfolding orthogonal_transformation
-  apply (rule iffI)
-  apply clarify
-  apply (clarsimp simp add: linear_0 linear_sub[symmetric] dist_norm)
-  apply (rule conjI)
-  apply (rule isometry_linear)
-  apply simp
-  apply simp
-  apply clarify
-  apply (erule_tac x=v in allE)
-  apply (erule_tac x=0 in allE)
-  by (simp add: dist_norm)
-
-(* ------------------------------------------------------------------------- *)
-(* Can extend an isometry from unit sphere.                                  *)
-(* ------------------------------------------------------------------------- *)
-
-lemma isometry_sphere_extend:
-  fixes f:: "real ^'n \<Rightarrow> real ^'n::finite"
-  assumes f1: "\<forall>x. norm x = 1 \<longrightarrow> norm (f x) = 1"
-  and fd1: "\<forall> x y. norm x = 1 \<longrightarrow> norm y = 1 \<longrightarrow> dist (f x) (f y) = dist x y"
-  shows "\<exists>g. orthogonal_transformation g \<and> (\<forall>x. norm x = 1 \<longrightarrow> g x = f x)"
-proof-
-  {fix x y x' y' x0 y0 x0' y0' :: "real ^'n"
-    assume H: "x = norm x *s x0" "y = norm y *s y0"
-    "x' = norm x *s x0'" "y' = norm y *s y0'"
-    "norm x0 = 1" "norm x0' = 1" "norm y0 = 1" "norm y0' = 1"
-    "norm(x0' - y0') = norm(x0 - y0)"
-
-    have "norm(x' - y') = norm(x - y)"
-      apply (subst H(1))
-      apply (subst H(2))
-      apply (subst H(3))
-      apply (subst H(4))
-      using H(5-9)
-      apply (simp add: norm_eq norm_eq_1)
-      apply (simp add: dot_lsub dot_rsub dot_lmult dot_rmult)
-      apply (simp add: ring_simps)
-      by (simp only: right_distrib[symmetric])}
-  note th0 = this
-  let ?g = "\<lambda>x. if x = 0 then 0 else norm x *s f (inverse (norm x) *s x)"
-  {fix x:: "real ^'n" assume nx: "norm x = 1"
-    have "?g x = f x" using nx by auto}
-  hence thfg: "\<forall>x. norm x = 1 \<longrightarrow> ?g x = f x" by blast
-  have g0: "?g 0 = 0" by simp
-  {fix x y :: "real ^'n"
-    {assume "x = 0" "y = 0"
-      then have "dist (?g x) (?g y) = dist x y" by simp }
-    moreover
-    {assume "x = 0" "y \<noteq> 0"
-      then have "dist (?g x) (?g y) = dist x y"
-        apply (simp add: dist_norm norm_mul)
-        apply (rule f1[rule_format])
-        by(simp add: norm_mul field_simps)}
-    moreover
-    {assume "x \<noteq> 0" "y = 0"
-      then have "dist (?g x) (?g y) = dist x y"
-        apply (simp add: dist_norm norm_mul)
-        apply (rule f1[rule_format])
-        by(simp add: norm_mul field_simps)}
-    moreover
-    {assume z: "x \<noteq> 0" "y \<noteq> 0"
-      have th00: "x = norm x *s (inverse (norm x) *s x)" "y = norm y *s (inverse (norm y) *s y)" "norm x *s f ((inverse (norm x) *s x)) = norm x *s f (inverse (norm x) *s x)"
-        "norm y *s f (inverse (norm y) *s y) = norm y *s f (inverse (norm y) *s y)"
-        "norm (inverse (norm x) *s x) = 1"
-        "norm (f (inverse (norm x) *s x)) = 1"
-        "norm (inverse (norm y) *s y) = 1"
-        "norm (f (inverse (norm y) *s y)) = 1"
-        "norm (f (inverse (norm x) *s x) - f (inverse (norm y) *s y)) =
-        norm (inverse (norm x) *s x - inverse (norm y) *s y)"
-        using z
-        by (auto simp add: vector_smult_assoc field_simps norm_mul intro: f1[rule_format] fd1[rule_format, unfolded dist_norm])
-      from z th0[OF th00] have "dist (?g x) (?g y) = dist x y"
-        by (simp add: dist_norm)}
-    ultimately have "dist (?g x) (?g y) = dist x y" by blast}
-  note thd = this
-    show ?thesis
-    apply (rule exI[where x= ?g])
-    unfolding orthogonal_transformation_isometry
-      using  g0 thfg thd by metis
-qed
-
-(* ------------------------------------------------------------------------- *)
-(* Rotation, reflection, rotoinversion.                                      *)
-(* ------------------------------------------------------------------------- *)
-
-definition "rotation_matrix Q \<longleftrightarrow> orthogonal_matrix Q \<and> det Q = 1"
-definition "rotoinversion_matrix Q \<longleftrightarrow> orthogonal_matrix Q \<and> det Q = - 1"
-
-lemma orthogonal_rotation_or_rotoinversion:
-  fixes Q :: "'a::ordered_idom^'n^'n::finite"
-  shows " orthogonal_matrix Q \<longleftrightarrow> rotation_matrix Q \<or> rotoinversion_matrix Q"
-  by (metis rotoinversion_matrix_def rotation_matrix_def det_orthogonal_matrix)
-(* ------------------------------------------------------------------------- *)
-(* Explicit formulas for low dimensions.                                     *)
-(* ------------------------------------------------------------------------- *)
-
-lemma setprod_1: "setprod f {(1::nat)..1} = f 1" by simp
-
-lemma setprod_2: "setprod f {(1::nat)..2} = f 1 * f 2"
-  by (simp add: nat_number setprod_numseg mult_commute)
-lemma setprod_3: "setprod f {(1::nat)..3} = f 1 * f 2 * f 3"
-  by (simp add: nat_number setprod_numseg mult_commute)
-
-lemma det_1: "det (A::'a::comm_ring_1^1^1) = A$1$1"
-  by (simp add: det_def permutes_sing sign_id UNIV_1)
-
-lemma det_2: "det (A::'a::comm_ring_1^2^2) = A$1$1 * A$2$2 - A$1$2 * A$2$1"
-proof-
-  have f12: "finite {2::2}" "1 \<notin> {2::2}" by auto
-  show ?thesis
-  unfolding det_def UNIV_2
-  unfolding setsum_over_permutations_insert[OF f12]
-  unfolding permutes_sing
-  apply (simp add: sign_swap_id sign_id swap_id_eq)
-  by (simp add: arith_simps(31)[symmetric] of_int_minus of_int_1 del: arith_simps(31))
-qed
-
-lemma det_3: "det (A::'a::comm_ring_1^3^3) =
-  A$1$1 * A$2$2 * A$3$3 +
-  A$1$2 * A$2$3 * A$3$1 +
-  A$1$3 * A$2$1 * A$3$2 -
-  A$1$1 * A$2$3 * A$3$2 -
-  A$1$2 * A$2$1 * A$3$3 -
-  A$1$3 * A$2$2 * A$3$1"
-proof-
-  have f123: "finite {2::3, 3}" "1 \<notin> {2::3, 3}" by auto
-  have f23: "finite {3::3}" "2 \<notin> {3::3}" by auto
-
-  show ?thesis
-  unfolding det_def UNIV_3
-  unfolding setsum_over_permutations_insert[OF f123]
-  unfolding setsum_over_permutations_insert[OF f23]
-
-  unfolding permutes_sing
-  apply (simp add: sign_swap_id permutation_swap_id sign_compose sign_id swap_id_eq)
-  apply (simp add: arith_simps(31)[symmetric] of_int_minus of_int_1 del: arith_simps(31))
-  by (simp add: ring_simps)
-qed
-
-end
--- a/src/HOL/Library/Euclidean_Space.thy	Fri Oct 23 14:33:07 2009 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,5372 +0,0 @@
-(*  Title:      Library/Euclidean_Space
-    Author:     Amine Chaieb, University of Cambridge
-*)
-
-header {* (Real) Vectors in Euclidean space, and elementary linear algebra.*}
-
-theory Euclidean_Space
-imports
-  Complex_Main "~~/src/HOL/Decision_Procs/Dense_Linear_Order"
-  Finite_Cartesian_Product Glbs Infinite_Set Numeral_Type
-  Inner_Product
-uses "positivstellensatz.ML" ("normarith.ML")
-begin
-
-text{* Some common special cases.*}
-
-lemma forall_1: "(\<forall>i::1. P i) \<longleftrightarrow> P 1"
-  by (metis num1_eq_iff)
-
-lemma exhaust_2:
-  fixes x :: 2 shows "x = 1 \<or> x = 2"
-proof (induct x)
-  case (of_int z)
-  then have "0 <= z" and "z < 2" by simp_all
-  then have "z = 0 | z = 1" by arith
-  then show ?case by auto
-qed
-
-lemma forall_2: "(\<forall>i::2. P i) \<longleftrightarrow> P 1 \<and> P 2"
-  by (metis exhaust_2)
-
-lemma exhaust_3:
-  fixes x :: 3 shows "x = 1 \<or> x = 2 \<or> x = 3"
-proof (induct x)
-  case (of_int z)
-  then have "0 <= z" and "z < 3" by simp_all
-  then have "z = 0 \<or> z = 1 \<or> z = 2" by arith
-  then show ?case by auto
-qed
-
-lemma forall_3: "(\<forall>i::3. P i) \<longleftrightarrow> P 1 \<and> P 2 \<and> P 3"
-  by (metis exhaust_3)
-
-lemma UNIV_1: "UNIV = {1::1}"
-  by (auto simp add: num1_eq_iff)
-
-lemma UNIV_2: "UNIV = {1::2, 2::2}"
-  using exhaust_2 by auto
-
-lemma UNIV_3: "UNIV = {1::3, 2::3, 3::3}"
-  using exhaust_3 by auto
-
-lemma setsum_1: "setsum f (UNIV::1 set) = f 1"
-  unfolding UNIV_1 by simp
-
-lemma setsum_2: "setsum f (UNIV::2 set) = f 1 + f 2"
-  unfolding UNIV_2 by simp
-
-lemma setsum_3: "setsum f (UNIV::3 set) = f 1 + f 2 + f 3"
-  unfolding UNIV_3 by (simp add: add_ac)
-
-subsection{* Basic componentwise operations on vectors. *}
-
-instantiation "^" :: (plus,type) plus
-begin
-definition  vector_add_def : "op + \<equiv> (\<lambda> x y.  (\<chi> i. (x$i) + (y$i)))"
-instance ..
-end
-
-instantiation "^" :: (times,type) times
-begin
-  definition vector_mult_def : "op * \<equiv> (\<lambda> x y.  (\<chi> i. (x$i) * (y$i)))"
-  instance ..
-end
-
-instantiation "^" :: (minus,type) minus begin
-  definition vector_minus_def : "op - \<equiv> (\<lambda> x y.  (\<chi> i. (x$i) - (y$i)))"
-instance ..
-end
-
-instantiation "^" :: (uminus,type) uminus begin
-  definition vector_uminus_def : "uminus \<equiv> (\<lambda> x.  (\<chi> i. - (x$i)))"
-instance ..
-end
-instantiation "^" :: (zero,type) zero begin
-  definition vector_zero_def : "0 \<equiv> (\<chi> i. 0)"
-instance ..
-end
-
-instantiation "^" :: (one,type) one begin
-  definition vector_one_def : "1 \<equiv> (\<chi> i. 1)"
-instance ..
-end
-
-instantiation "^" :: (ord,type) ord
- begin
-definition vector_less_eq_def:
-  "less_eq (x :: 'a ^'b) y = (ALL i. x$i <= y$i)"
-definition vector_less_def: "less (x :: 'a ^'b) y = (ALL i. x$i < y$i)"
-
-instance by (intro_classes)
-end
-
-instantiation "^" :: (scaleR, type) scaleR
-begin
-definition vector_scaleR_def: "scaleR = (\<lambda> r x.  (\<chi> i. scaleR r (x$i)))"
-instance ..
-end
-
-text{* Also the scalar-vector multiplication. *}
-
-definition vector_scalar_mult:: "'a::times \<Rightarrow> 'a ^'n \<Rightarrow> 'a ^ 'n" (infixl "*s" 70)
-  where "c *s x = (\<chi> i. c * (x$i))"
-
-text{* Constant Vectors *} 
-
-definition "vec x = (\<chi> i. x)"
-
-text{* Dot products. *}
-
-definition dot :: "'a::{comm_monoid_add, times} ^ 'n \<Rightarrow> 'a ^ 'n \<Rightarrow> 'a" (infix "\<bullet>" 70) where
-  "x \<bullet> y = setsum (\<lambda>i. x$i * y$i) UNIV"
-
-lemma dot_1[simp]: "(x::'a::{comm_monoid_add, times}^1) \<bullet> y = (x$1) * (y$1)"
-  by (simp add: dot_def setsum_1)
-
-lemma dot_2[simp]: "(x::'a::{comm_monoid_add, times}^2) \<bullet> y = (x$1) * (y$1) + (x$2) * (y$2)"
-  by (simp add: dot_def setsum_2)
-
-lemma dot_3[simp]: "(x::'a::{comm_monoid_add, times}^3) \<bullet> y = (x$1) * (y$1) + (x$2) * (y$2) + (x$3) * (y$3)"
-  by (simp add: dot_def setsum_3)
-
-subsection {* A naive proof procedure to lift really trivial arithmetic stuff from the basis of the vector space. *}
-
-method_setup vector = {*
-let
-  val ss1 = HOL_basic_ss addsimps [@{thm dot_def}, @{thm setsum_addf} RS sym,
-  @{thm setsum_subtractf} RS sym, @{thm setsum_right_distrib},
-  @{thm setsum_left_distrib}, @{thm setsum_negf} RS sym]
-  val ss2 = @{simpset} addsimps
-             [@{thm vector_add_def}, @{thm vector_mult_def},
-              @{thm vector_minus_def}, @{thm vector_uminus_def},
-              @{thm vector_one_def}, @{thm vector_zero_def}, @{thm vec_def},
-              @{thm vector_scaleR_def},
-              @{thm Cart_lambda_beta}, @{thm vector_scalar_mult_def}]
- fun vector_arith_tac ths =
-   simp_tac ss1
-   THEN' (fn i => rtac @{thm setsum_cong2} i
-         ORELSE rtac @{thm setsum_0'} i
-         ORELSE simp_tac (HOL_basic_ss addsimps [@{thm "Cart_eq"}]) i)
-   (* THEN' TRY o clarify_tac HOL_cs  THEN' (TRY o rtac @{thm iffI}) *)
-   THEN' asm_full_simp_tac (ss2 addsimps ths)
- in
-  Attrib.thms >> (fn ths => K (SIMPLE_METHOD' (vector_arith_tac ths)))
- end
-*} "Lifts trivial vector statements to real arith statements"
-
-lemma vec_0[simp]: "vec 0 = 0" by (vector vector_zero_def)
-lemma vec_1[simp]: "vec 1 = 1" by (vector vector_one_def)
-
-
-
-text{* Obvious "component-pushing". *}
-
-lemma vec_component [simp]: "(vec x :: 'a ^ 'n)$i = x"
-  by (vector vec_def)
-
-lemma vector_add_component [simp]:
-  fixes x y :: "'a::{plus} ^ 'n"
-  shows "(x + y)$i = x$i + y$i"
-  by vector
-
-lemma vector_minus_component [simp]:
-  fixes x y :: "'a::{minus} ^ 'n"
-  shows "(x - y)$i = x$i - y$i"
-  by vector
-
-lemma vector_mult_component [simp]:
-  fixes x y :: "'a::{times} ^ 'n"
-  shows "(x * y)$i = x$i * y$i"
-  by vector
-
-lemma vector_smult_component [simp]:
-  fixes y :: "'a::{times} ^ 'n"
-  shows "(c *s y)$i = c * (y$i)"
-  by vector
-
-lemma vector_uminus_component [simp]:
-  fixes x :: "'a::{uminus} ^ 'n"
-  shows "(- x)$i = - (x$i)"
-  by vector
-
-lemma vector_scaleR_component [simp]:
-  fixes x :: "'a::scaleR ^ 'n"
-  shows "(scaleR r x)$i = scaleR r (x$i)"
-  by vector
-
-lemma cond_component: "(if b then x else y)$i = (if b then x$i else y$i)" by vector
-
-lemmas vector_component =
-  vec_component vector_add_component vector_mult_component
-  vector_smult_component vector_minus_component vector_uminus_component
-  vector_scaleR_component cond_component
-
-subsection {* Some frequently useful arithmetic lemmas over vectors. *}
-
-instance "^" :: (semigroup_add,type) semigroup_add
-  apply (intro_classes) by (vector add_assoc)
-
-
-instance "^" :: (monoid_add,type) monoid_add
-  apply (intro_classes) by vector+
-
-instance "^" :: (group_add,type) group_add
-  apply (intro_classes) by (vector algebra_simps)+
-
-instance "^" :: (ab_semigroup_add,type) ab_semigroup_add
-  apply (intro_classes) by (vector add_commute)
-
-instance "^" :: (comm_monoid_add,type) comm_monoid_add
-  apply (intro_classes) by vector
-
-instance "^" :: (ab_group_add,type) ab_group_add
-  apply (intro_classes) by vector+
-
-instance "^" :: (cancel_semigroup_add,type) cancel_semigroup_add
-  apply (intro_classes)
-  by (vector Cart_eq)+
-
-instance "^" :: (cancel_ab_semigroup_add,type) cancel_ab_semigroup_add
-  apply (intro_classes)
-  by (vector Cart_eq)
-
-instance "^" :: (real_vector, type) real_vector
-  by default (vector scaleR_left_distrib scaleR_right_distrib)+
-
-instance "^" :: (semigroup_mult,type) semigroup_mult
-  apply (intro_classes) by (vector mult_assoc)
-
-instance "^" :: (monoid_mult,type) monoid_mult
-  apply (intro_classes) by vector+
-
-instance "^" :: (ab_semigroup_mult,type) ab_semigroup_mult
-  apply (intro_classes) by (vector mult_commute)
-
-instance "^" :: (ab_semigroup_idem_mult,type) ab_semigroup_idem_mult
-  apply (intro_classes) by (vector mult_idem)
-
-instance "^" :: (comm_monoid_mult,type) comm_monoid_mult
-  apply (intro_classes) by vector
-
-fun vector_power :: "('a::{one,times} ^'n) \<Rightarrow> nat \<Rightarrow> 'a^'n" where
-  "vector_power x 0 = 1"
-  | "vector_power x (Suc n) = x * vector_power x n"
-
-instance "^" :: (semiring,type) semiring
-  apply (intro_classes) by (vector ring_simps)+
-
-instance "^" :: (semiring_0,type) semiring_0
-  apply (intro_classes) by (vector ring_simps)+
-instance "^" :: (semiring_1,type) semiring_1
-  apply (intro_classes) by vector
-instance "^" :: (comm_semiring,type) comm_semiring
-  apply (intro_classes) by (vector ring_simps)+
-
-instance "^" :: (comm_semiring_0,type) comm_semiring_0 by (intro_classes)
-instance "^" :: (cancel_comm_monoid_add, type) cancel_comm_monoid_add ..
-instance "^" :: (semiring_0_cancel,type) semiring_0_cancel by (intro_classes)
-instance "^" :: (comm_semiring_0_cancel,type) comm_semiring_0_cancel by (intro_classes)
-instance "^" :: (ring,type) ring by (intro_classes)
-instance "^" :: (semiring_1_cancel,type) semiring_1_cancel by (intro_classes)
-instance "^" :: (comm_semiring_1,type) comm_semiring_1 by (intro_classes)
-
-instance "^" :: (ring_1,type) ring_1 ..
-
-instance "^" :: (real_algebra,type) real_algebra
-  apply intro_classes
-  apply (simp_all add: vector_scaleR_def ring_simps)
-  apply vector
-  apply vector
-  done
-
-instance "^" :: (real_algebra_1,type) real_algebra_1 ..
-
-lemma of_nat_index:
-  "(of_nat n :: 'a::semiring_1 ^'n)$i = of_nat n"
-  apply (induct n)
-  apply vector
-  apply vector
-  done
-lemma zero_index[simp]:
-  "(0 :: 'a::zero ^'n)$i = 0" by vector
-
-lemma one_index[simp]:
-  "(1 :: 'a::one ^'n)$i = 1" by vector
-
-lemma one_plus_of_nat_neq_0: "(1::'a::semiring_char_0) + of_nat n \<noteq> 0"
-proof-
-  have "(1::'a) + of_nat n = 0 \<longleftrightarrow> of_nat 1 + of_nat n = (of_nat 0 :: 'a)" by simp
-  also have "\<dots> \<longleftrightarrow> 1 + n = 0" by (simp only: of_nat_add[symmetric] of_nat_eq_iff)
-  finally show ?thesis by simp
-qed
-
-instance "^" :: (semiring_char_0,type) semiring_char_0
-proof (intro_classes)
-  fix m n ::nat
-  show "(of_nat m :: 'a^'b) = of_nat n \<longleftrightarrow> m = n"
-    by (simp add: Cart_eq of_nat_index)
-qed
-
-instance "^" :: (comm_ring_1,type) comm_ring_1 by intro_classes
-instance "^" :: (ring_char_0,type) ring_char_0 by intro_classes
-
-lemma vector_smult_assoc: "a *s (b *s x) = ((a::'a::semigroup_mult) * b) *s x"
-  by (vector mult_assoc)
-lemma vector_sadd_rdistrib: "((a::'a::semiring) + b) *s x = a *s x + b *s x"
-  by (vector ring_simps)
-lemma vector_add_ldistrib: "(c::'a::semiring) *s (x + y) = c *s x + c *s y"
-  by (vector ring_simps)
-lemma vector_smult_lzero[simp]: "(0::'a::mult_zero) *s x = 0" by vector
-lemma vector_smult_lid[simp]: "(1::'a::monoid_mult) *s x = x" by vector
-lemma vector_ssub_ldistrib: "(c::'a::ring) *s (x - y) = c *s x - c *s y"
-  by (vector ring_simps)
-lemma vector_smult_rneg: "(c::'a::ring) *s -x = -(c *s x)" by vector
-lemma vector_smult_lneg: "- (c::'a::ring) *s x = -(c *s x)" by vector
-lemma vector_sneg_minus1: "-x = (- (1::'a::ring_1)) *s x" by vector
-lemma vector_smult_rzero[simp]: "c *s 0 = (0::'a::mult_zero ^ 'n)" by vector
-lemma vector_sub_rdistrib: "((a::'a::ring) - b) *s x = a *s x - b *s x"
-  by (vector ring_simps)
-
-lemma vec_eq[simp]: "(vec m = vec n) \<longleftrightarrow> (m = n)"
-  by (simp add: Cart_eq)
-
-subsection {* Topological space *}
-
-instantiation "^" :: (topological_space, finite) topological_space
-begin
-
-definition open_vector_def:
-  "open (S :: ('a ^ 'b) set) \<longleftrightarrow>
-    (\<forall>x\<in>S. \<exists>A. (\<forall>i. open (A i) \<and> x$i \<in> A i) \<and>
-      (\<forall>y. (\<forall>i. y$i \<in> A i) \<longrightarrow> y \<in> S))"
-
-instance proof
-  show "open (UNIV :: ('a ^ 'b) set)"
-    unfolding open_vector_def by auto
-next
-  fix S T :: "('a ^ 'b) set"
-  assume "open S" "open T" thus "open (S \<inter> T)"
-    unfolding open_vector_def
-    apply clarify
-    apply (drule (1) bspec)+
-    apply (clarify, rename_tac Sa Ta)
-    apply (rule_tac x="\<lambda>i. Sa i \<inter> Ta i" in exI)
-    apply (simp add: open_Int)
-    done
-next
-  fix K :: "('a ^ 'b) set set"
-  assume "\<forall>S\<in>K. open S" thus "open (\<Union>K)"
-    unfolding open_vector_def
-    apply clarify
-    apply (drule (1) bspec)
-    apply (drule (1) bspec)
-    apply clarify
-    apply (rule_tac x=A in exI)
-    apply fast
-    done
-qed
-
-end
-
-lemma open_vector_box: "\<forall>i. open (S i) \<Longrightarrow> open {x. \<forall>i. x $ i \<in> S i}"
-unfolding open_vector_def by auto
-
-lemma open_vimage_Cart_nth: "open S \<Longrightarrow> open ((\<lambda>x. x $ i) -` S)"
-unfolding open_vector_def
-apply clarify
-apply (rule_tac x="\<lambda>k. if k = i then S else UNIV" in exI, simp)
-done
-
-lemma closed_vimage_Cart_nth: "closed S \<Longrightarrow> closed ((\<lambda>x. x $ i) -` S)"
-unfolding closed_open vimage_Compl [symmetric]
-by (rule open_vimage_Cart_nth)
-
-lemma closed_vector_box: "\<forall>i. closed (S i) \<Longrightarrow> closed {x. \<forall>i. x $ i \<in> S i}"
-proof -
-  have "{x. \<forall>i. x $ i \<in> S i} = (\<Inter>i. (\<lambda>x. x $ i) -` S i)" by auto
-  thus "\<forall>i. closed (S i) \<Longrightarrow> closed {x. \<forall>i. x $ i \<in> S i}"
-    by (simp add: closed_INT closed_vimage_Cart_nth)
-qed
-
-lemma tendsto_Cart_nth [tendsto_intros]:
-  assumes "((\<lambda>x. f x) ---> a) net"
-  shows "((\<lambda>x. f x $ i) ---> a $ i) net"
-proof (rule topological_tendstoI)
-  fix S assume "open S" "a $ i \<in> S"
-  then have "open ((\<lambda>y. y $ i) -` S)" "a \<in> ((\<lambda>y. y $ i) -` S)"
-    by (simp_all add: open_vimage_Cart_nth)
-  with assms have "eventually (\<lambda>x. f x \<in> (\<lambda>y. y $ i) -` S) net"
-    by (rule topological_tendstoD)
-  then show "eventually (\<lambda>x. f x $ i \<in> S) net"
-    by simp
-qed
-
-subsection {* Square root of sum of squares *}
-
-definition
-  "setL2 f A = sqrt (\<Sum>i\<in>A. (f i)\<twosuperior>)"
-
-lemma setL2_cong:
-  "\<lbrakk>A = B; \<And>x. x \<in> B \<Longrightarrow> f x = g x\<rbrakk> \<Longrightarrow> setL2 f A = setL2 g B"
-  unfolding setL2_def by simp
-
-lemma strong_setL2_cong:
-  "\<lbrakk>A = B; \<And>x. x \<in> B =simp=> f x = g x\<rbrakk> \<Longrightarrow> setL2 f A = setL2 g B"
-  unfolding setL2_def simp_implies_def by simp
-
-lemma setL2_infinite [simp]: "\<not> finite A \<Longrightarrow> setL2 f A = 0"
-  unfolding setL2_def by simp
-
-lemma setL2_empty [simp]: "setL2 f {} = 0"
-  unfolding setL2_def by simp
-
-lemma setL2_insert [simp]:
-  "\<lbrakk>finite F; a \<notin> F\<rbrakk> \<Longrightarrow>
-    setL2 f (insert a F) = sqrt ((f a)\<twosuperior> + (setL2 f F)\<twosuperior>)"
-  unfolding setL2_def by (simp add: setsum_nonneg)
-
-lemma setL2_nonneg [simp]: "0 \<le> setL2 f A"
-  unfolding setL2_def by (simp add: setsum_nonneg)
-
-lemma setL2_0': "\<forall>a\<in>A. f a = 0 \<Longrightarrow> setL2 f A = 0"
-  unfolding setL2_def by simp
-
-lemma setL2_constant: "setL2 (\<lambda>x. y) A = sqrt (of_nat (card A)) * \<bar>y\<bar>"
-  unfolding setL2_def by (simp add: real_sqrt_mult)
-
-lemma setL2_mono:
-  assumes "\<And>i. i \<in> K \<Longrightarrow> f i \<le> g i"
-  assumes "\<And>i. i \<in> K \<Longrightarrow> 0 \<le> f i"
-  shows "setL2 f K \<le> setL2 g K"
-  unfolding setL2_def
-  by (simp add: setsum_nonneg setsum_mono power_mono prems)
-
-lemma setL2_strict_mono:
-  assumes "finite K" and "K \<noteq> {}"
-  assumes "\<And>i. i \<in> K \<Longrightarrow> f i < g i"
-  assumes "\<And>i. i \<in> K \<Longrightarrow> 0 \<le> f i"
-  shows "setL2 f K < setL2 g K"
-  unfolding setL2_def
-  by (simp add: setsum_strict_mono power_strict_mono assms)
-
-lemma setL2_right_distrib:
-  "0 \<le> r \<Longrightarrow> r * setL2 f A = setL2 (\<lambda>x. r * f x) A"
-  unfolding setL2_def
-  apply (simp add: power_mult_distrib)
-  apply (simp add: setsum_right_distrib [symmetric])
-  apply (simp add: real_sqrt_mult setsum_nonneg)
-  done
-
-lemma setL2_left_distrib:
-  "0 \<le> r \<Longrightarrow> setL2 f A * r = setL2 (\<lambda>x. f x * r) A"
-  unfolding setL2_def
-  apply (simp add: power_mult_distrib)
-  apply (simp add: setsum_left_distrib [symmetric])
-  apply (simp add: real_sqrt_mult setsum_nonneg)
-  done
-
-lemma setsum_nonneg_eq_0_iff:
-  fixes f :: "'a \<Rightarrow> 'b::pordered_ab_group_add"
-  shows "\<lbrakk>finite A; \<forall>x\<in>A. 0 \<le> f x\<rbrakk> \<Longrightarrow> setsum f A = 0 \<longleftrightarrow> (\<forall>x\<in>A. f x = 0)"
-  apply (induct set: finite, simp)
-  apply (simp add: add_nonneg_eq_0_iff setsum_nonneg)
-  done
-
-lemma setL2_eq_0_iff: "finite A \<Longrightarrow> setL2 f A = 0 \<longleftrightarrow> (\<forall>x\<in>A. f x = 0)"
-  unfolding setL2_def
-  by (simp add: setsum_nonneg setsum_nonneg_eq_0_iff)
-
-lemma setL2_triangle_ineq:
-  shows "setL2 (\<lambda>i. f i + g i) A \<le> setL2 f A + setL2 g A"
-proof (cases "finite A")
-  case False
-  thus ?thesis by simp
-next
-  case True
-  thus ?thesis
-  proof (induct set: finite)
-    case empty
-    show ?case by simp
-  next
-    case (insert x F)
-    hence "sqrt ((f x + g x)\<twosuperior> + (setL2 (\<lambda>i. f i + g i) F)\<twosuperior>) \<le>
-           sqrt ((f x + g x)\<twosuperior> + (setL2 f F + setL2 g F)\<twosuperior>)"
-      by (intro real_sqrt_le_mono add_left_mono power_mono insert
-                setL2_nonneg add_increasing zero_le_power2)
-    also have
-      "\<dots> \<le> sqrt ((f x)\<twosuperior> + (setL2 f F)\<twosuperior>) + sqrt ((g x)\<twosuperior> + (setL2 g F)\<twosuperior>)"
-      by (rule real_sqrt_sum_squares_triangle_ineq)
-    finally show ?case
-      using insert by simp
-  qed
-qed
-
-lemma sqrt_sum_squares_le_sum:
-  "\<lbrakk>0 \<le> x; 0 \<le> y\<rbrakk> \<Longrightarrow> sqrt (x\<twosuperior> + y\<twosuperior>) \<le> x + y"
-  apply (rule power2_le_imp_le)
-  apply (simp add: power2_sum)
-  apply (simp add: mult_nonneg_nonneg)
-  apply (simp add: add_nonneg_nonneg)
-  done
-
-lemma setL2_le_setsum [rule_format]:
-  "(\<forall>i\<in>A. 0 \<le> f i) \<longrightarrow> setL2 f A \<le> setsum f A"
-  apply (cases "finite A")
-  apply (induct set: finite)
-  apply simp
-  apply clarsimp
-  apply (erule order_trans [OF sqrt_sum_squares_le_sum])
-  apply simp
-  apply simp
-  apply simp
-  done
-
-lemma sqrt_sum_squares_le_sum_abs: "sqrt (x\<twosuperior> + y\<twosuperior>) \<le> \<bar>x\<bar> + \<bar>y\<bar>"
-  apply (rule power2_le_imp_le)
-  apply (simp add: power2_sum)
-  apply (simp add: mult_nonneg_nonneg)
-  apply (simp add: add_nonneg_nonneg)
-  done
-
-lemma setL2_le_setsum_abs: "setL2 f A \<le> (\<Sum>i\<in>A. \<bar>f i\<bar>)"
-  apply (cases "finite A")
-  apply (induct set: finite)
-  apply simp
-  apply simp
-  apply (rule order_trans [OF sqrt_sum_squares_le_sum_abs])
-  apply simp
-  apply simp
-  done
-
-lemma setL2_mult_ineq_lemma:
-  fixes a b c d :: real
-  shows "2 * (a * c) * (b * d) \<le> a\<twosuperior> * d\<twosuperior> + b\<twosuperior> * c\<twosuperior>"
-proof -
-  have "0 \<le> (a * d - b * c)\<twosuperior>" by simp
-  also have "\<dots> = a\<twosuperior> * d\<twosuperior> + b\<twosuperior> * c\<twosuperior> - 2 * (a * d) * (b * c)"
-    by (simp only: power2_diff power_mult_distrib)
-  also have "\<dots> = a\<twosuperior> * d\<twosuperior> + b\<twosuperior> * c\<twosuperior> - 2 * (a * c) * (b * d)"
-    by simp
-  finally show "2 * (a * c) * (b * d) \<le> a\<twosuperior> * d\<twosuperior> + b\<twosuperior> * c\<twosuperior>"
-    by simp
-qed
-
-lemma setL2_mult_ineq: "(\<Sum>i\<in>A. \<bar>f i\<bar> * \<bar>g i\<bar>) \<le> setL2 f A * setL2 g A"
-  apply (cases "finite A")
-  apply (induct set: finite)
-  apply simp
-  apply (rule power2_le_imp_le, simp)
-  apply (rule order_trans)
-  apply (rule power_mono)
-  apply (erule add_left_mono)
-  apply (simp add: add_nonneg_nonneg mult_nonneg_nonneg setsum_nonneg)
-  apply (simp add: power2_sum)
-  apply (simp add: power_mult_distrib)
-  apply (simp add: right_distrib left_distrib)
-  apply (rule ord_le_eq_trans)
-  apply (rule setL2_mult_ineq_lemma)
-  apply simp
-  apply (intro mult_nonneg_nonneg setL2_nonneg)
-  apply simp
-  done
-
-lemma member_le_setL2: "\<lbrakk>finite A; i \<in> A\<rbrakk> \<Longrightarrow> f i \<le> setL2 f A"
-  apply (rule_tac s="insert i (A - {i})" and t="A" in subst)
-  apply fast
-  apply (subst setL2_insert)
-  apply simp
-  apply simp
-  apply simp
-  done
-
-subsection {* Metric *}
-
-(* TODO: move somewhere else *)
-lemma finite_choice: "finite A \<Longrightarrow> \<forall>x\<in>A. \<exists>y. P x y \<Longrightarrow> \<exists>f. \<forall>x\<in>A. P x (f x)"
-apply (induct set: finite, simp_all)
-apply (clarify, rename_tac y)
-apply (rule_tac x="f(x:=y)" in exI, simp)
-done
-
-instantiation "^" :: (metric_space, finite) metric_space
-begin
-
-definition dist_vector_def:
-  "dist (x::'a^'b) (y::'a^'b) = setL2 (\<lambda>i. dist (x$i) (y$i)) UNIV"
-
-lemma dist_nth_le: "dist (x $ i) (y $ i) \<le> dist x y"
-unfolding dist_vector_def
-by (rule member_le_setL2) simp_all
-
-instance proof
-  fix x y :: "'a ^ 'b"
-  show "dist x y = 0 \<longleftrightarrow> x = y"
-    unfolding dist_vector_def
-    by (simp add: setL2_eq_0_iff Cart_eq)
-next
-  fix x y z :: "'a ^ 'b"
-  show "dist x y \<le> dist x z + dist y z"
-    unfolding dist_vector_def
-    apply (rule order_trans [OF _ setL2_triangle_ineq])
-    apply (simp add: setL2_mono dist_triangle2)
-    done
-next
-  (* FIXME: long proof! *)
-  fix S :: "('a ^ 'b) set"
-  show "open S \<longleftrightarrow> (\<forall>x\<in>S. \<exists>e>0. \<forall>y. dist y x < e \<longrightarrow> y \<in> S)"
-    unfolding open_vector_def open_dist
-    apply safe
-     apply (drule (1) bspec)
-     apply clarify
-     apply (subgoal_tac "\<exists>e>0. \<forall>i y. dist y (x$i) < e \<longrightarrow> y \<in> A i")
-      apply clarify
-      apply (rule_tac x=e in exI, clarify)
-      apply (drule spec, erule mp, clarify)
-      apply (drule spec, drule spec, erule mp)
-      apply (erule le_less_trans [OF dist_nth_le])
-     apply (subgoal_tac "\<forall>i\<in>UNIV. \<exists>e>0. \<forall>y. dist y (x$i) < e \<longrightarrow> y \<in> A i")
-      apply (drule finite_choice [OF finite], clarify)
-      apply (rule_tac x="Min (range f)" in exI, simp)
-     apply clarify
-     apply (drule_tac x=i in spec, clarify)
-     apply (erule (1) bspec)
-    apply (drule (1) bspec, clarify)
-    apply (subgoal_tac "\<exists>r. (\<forall>i::'b. 0 < r i) \<and> e = setL2 r UNIV")
-     apply clarify
-     apply (rule_tac x="\<lambda>i. {y. dist y (x$i) < r i}" in exI)
-     apply (rule conjI)
-      apply clarify
-      apply (rule conjI)
-       apply (clarify, rename_tac y)
-       apply (rule_tac x="r i - dist y (x$i)" in exI, rule conjI, simp)
-       apply clarify
-       apply (simp only: less_diff_eq)
-       apply (erule le_less_trans [OF dist_triangle])
-      apply simp
-     apply clarify
-     apply (drule spec, erule mp)
-     apply (simp add: dist_vector_def setL2_strict_mono)
-    apply (rule_tac x="\<lambda>i. e / sqrt (of_nat CARD('b))" in exI)
-    apply (simp add: divide_pos_pos setL2_constant)
-    done
-qed
-
-end
-
-lemma LIMSEQ_Cart_nth:
-  "(X ----> a) \<Longrightarrow> (\<lambda>n. X n $ i) ----> a $ i"
-unfolding LIMSEQ_conv_tendsto by (rule tendsto_Cart_nth)
-
-lemma LIM_Cart_nth:
-  "(f -- x --> y) \<Longrightarrow> (\<lambda>x. f x $ i) -- x --> y $ i"
-unfolding LIM_conv_tendsto by (rule tendsto_Cart_nth)
-
-lemma Cauchy_Cart_nth:
-  "Cauchy (\<lambda>n. X n) \<Longrightarrow> Cauchy (\<lambda>n. X n $ i)"
-unfolding Cauchy_def by (fast intro: le_less_trans [OF dist_nth_le])
-
-lemma LIMSEQ_vector:
-  fixes X :: "nat \<Rightarrow> 'a::metric_space ^ 'n::finite"
-  assumes X: "\<And>i. (\<lambda>n. X n $ i) ----> (a $ i)"
-  shows "X ----> a"
-proof (rule metric_LIMSEQ_I)
-  fix r :: real assume "0 < r"
-  then have "0 < r / of_nat CARD('n)" (is "0 < ?s")
-    by (simp add: divide_pos_pos)
-  def N \<equiv> "\<lambda>i. LEAST N. \<forall>n\<ge>N. dist (X n $ i) (a $ i) < ?s"
-  def M \<equiv> "Max (range N)"
-  have "\<And>i. \<exists>N. \<forall>n\<ge>N. dist (X n $ i) (a $ i) < ?s"
-    using X `0 < ?s` by (rule metric_LIMSEQ_D)
-  hence "\<And>i. \<forall>n\<ge>N i. dist (X n $ i) (a $ i) < ?s"
-    unfolding N_def by (rule LeastI_ex)
-  hence M: "\<And>i. \<forall>n\<ge>M. dist (X n $ i) (a $ i) < ?s"
-    unfolding M_def by simp
-  {
-    fix n :: nat assume "M \<le> n"
-    have "dist (X n) a = setL2 (\<lambda>i. dist (X n $ i) (a $ i)) UNIV"
-      unfolding dist_vector_def ..
-    also have "\<dots> \<le> setsum (\<lambda>i. dist (X n $ i) (a $ i)) UNIV"
-      by (rule setL2_le_setsum [OF zero_le_dist])
-    also have "\<dots> < setsum (\<lambda>i::'n. ?s) UNIV"
-      by (rule setsum_strict_mono, simp_all add: M `M \<le> n`)
-    also have "\<dots> = r"
-      by simp
-    finally have "dist (X n) a < r" .
-  }
-  hence "\<forall>n\<ge>M. dist (X n) a < r"
-    by simp
-  then show "\<exists>M. \<forall>n\<ge>M. dist (X n) a < r" ..
-qed
-
-lemma Cauchy_vector:
-  fixes X :: "nat \<Rightarrow> 'a::metric_space ^ 'n::finite"
-  assumes X: "\<And>i. Cauchy (\<lambda>n. X n $ i)"
-  shows "Cauchy (\<lambda>n. X n)"
-proof (rule metric_CauchyI)
-  fix r :: real assume "0 < r"
-  then have "0 < r / of_nat CARD('n)" (is "0 < ?s")
-    by (simp add: divide_pos_pos)
-  def N \<equiv> "\<lambda>i. LEAST N. \<forall>m\<ge>N. \<forall>n\<ge>N. dist (X m $ i) (X n $ i) < ?s"
-  def M \<equiv> "Max (range N)"
-  have "\<And>i. \<exists>N. \<forall>m\<ge>N. \<forall>n\<ge>N. dist (X m $ i) (X n $ i) < ?s"
-    using X `0 < ?s` by (rule metric_CauchyD)
-  hence "\<And>i. \<forall>m\<ge>N i. \<forall>n\<ge>N i. dist (X m $ i) (X n $ i) < ?s"
-    unfolding N_def by (rule LeastI_ex)
-  hence M: "\<And>i. \<forall>m\<ge>M. \<forall>n\<ge>M. dist (X m $ i) (X n $ i) < ?s"
-    unfolding M_def by simp
-  {
-    fix m n :: nat
-    assume "M \<le> m" "M \<le> n"
-    have "dist (X m) (X n) = setL2 (\<lambda>i. dist (X m $ i) (X n $ i)) UNIV"
-      unfolding dist_vector_def ..
-    also have "\<dots> \<le> setsum (\<lambda>i. dist (X m $ i) (X n $ i)) UNIV"
-      by (rule setL2_le_setsum [OF zero_le_dist])
-    also have "\<dots> < setsum (\<lambda>i::'n. ?s) UNIV"
-      by (rule setsum_strict_mono, simp_all add: M `M \<le> m` `M \<le> n`)
-    also have "\<dots> = r"
-      by simp
-    finally have "dist (X m) (X n) < r" .
-  }
-  hence "\<forall>m\<ge>M. \<forall>n\<ge>M. dist (X m) (X n) < r"
-    by simp
-  then show "\<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. dist (X m) (X n) < r" ..
-qed
-
-instance "^" :: (complete_space, finite) complete_space
-proof
-  fix X :: "nat \<Rightarrow> 'a ^ 'b" assume "Cauchy X"
-  have "\<And>i. (\<lambda>n. X n $ i) ----> lim (\<lambda>n. X n $ i)"
-    using Cauchy_Cart_nth [OF `Cauchy X`]
-    by (simp add: Cauchy_convergent_iff convergent_LIMSEQ_iff)
-  hence "X ----> Cart_lambda (\<lambda>i. lim (\<lambda>n. X n $ i))"
-    by (simp add: LIMSEQ_vector)
-  then show "convergent X"
-    by (rule convergentI)
-qed
-
-subsection {* Norms *}
-
-instantiation "^" :: (real_normed_vector, finite) real_normed_vector
-begin
-
-definition norm_vector_def:
-  "norm (x::'a^'b) = setL2 (\<lambda>i. norm (x$i)) UNIV"
-
-definition vector_sgn_def:
-  "sgn (x::'a^'b) = scaleR (inverse (norm x)) x"
-
-instance proof
-  fix a :: real and x y :: "'a ^ 'b"
-  show "0 \<le> norm x"
-    unfolding norm_vector_def
-    by (rule setL2_nonneg)
-  show "norm x = 0 \<longleftrightarrow> x = 0"
-    unfolding norm_vector_def
-    by (simp add: setL2_eq_0_iff Cart_eq)
-  show "norm (x + y) \<le> norm x + norm y"
-    unfolding norm_vector_def
-    apply (rule order_trans [OF _ setL2_triangle_ineq])
-    apply (simp add: setL2_mono norm_triangle_ineq)
-    done
-  show "norm (scaleR a x) = \<bar>a\<bar> * norm x"
-    unfolding norm_vector_def
-    by (simp add: setL2_right_distrib)
-  show "sgn x = scaleR (inverse (norm x)) x"
-    by (rule vector_sgn_def)
-  show "dist x y = norm (x - y)"
-    unfolding dist_vector_def norm_vector_def
-    by (simp add: dist_norm)
-qed
-
-end
-
-lemma norm_nth_le: "norm (x $ i) \<le> norm x"
-unfolding norm_vector_def
-by (rule member_le_setL2) simp_all
-
-interpretation Cart_nth: bounded_linear "\<lambda>x. x $ i"
-apply default
-apply (rule vector_add_component)
-apply (rule vector_scaleR_component)
-apply (rule_tac x="1" in exI, simp add: norm_nth_le)
-done
-
-instance "^" :: (banach, finite) banach ..
-
-subsection {* Inner products *}
-
-instantiation "^" :: (real_inner, finite) real_inner
-begin
-
-definition inner_vector_def:
-  "inner x y = setsum (\<lambda>i. inner (x$i) (y$i)) UNIV"
-
-instance proof
-  fix r :: real and x y z :: "'a ^ 'b"
-  show "inner x y = inner y x"
-    unfolding inner_vector_def
-    by (simp add: inner_commute)
-  show "inner (x + y) z = inner x z + inner y z"
-    unfolding inner_vector_def
-    by (simp add: inner_add_left setsum_addf)
-  show "inner (scaleR r x) y = r * inner x y"
-    unfolding inner_vector_def
-    by (simp add: setsum_right_distrib)
-  show "0 \<le> inner x x"
-    unfolding inner_vector_def
-    by (simp add: setsum_nonneg)
-  show "inner x x = 0 \<longleftrightarrow> x = 0"
-    unfolding inner_vector_def
-    by (simp add: Cart_eq setsum_nonneg_eq_0_iff)
-  show "norm x = sqrt (inner x x)"
-    unfolding inner_vector_def norm_vector_def setL2_def
-    by (simp add: power2_norm_eq_inner)
-qed
-
-end
-
-subsection{* Properties of the dot product.  *}
-
-lemma dot_sym: "(x::'a:: {comm_monoid_add, ab_semigroup_mult} ^ 'n) \<bullet> y = y \<bullet> x"
-  by (vector mult_commute)
-lemma dot_ladd: "((x::'a::ring ^ 'n) + y) \<bullet> z = (x \<bullet> z) + (y \<bullet> z)"
-  by (vector ring_simps)
-lemma dot_radd: "x \<bullet> (y + (z::'a::ring ^ 'n)) = (x \<bullet> y) + (x \<bullet> z)"
-  by (vector ring_simps)
-lemma dot_lsub: "((x::'a::ring ^ 'n) - y) \<bullet> z = (x \<bullet> z) - (y \<bullet> z)"
-  by (vector ring_simps)
-lemma dot_rsub: "(x::'a::ring ^ 'n) \<bullet> (y - z) = (x \<bullet> y) - (x \<bullet> z)"
-  by (vector ring_simps)
-lemma dot_lmult: "(c *s x) \<bullet> y = (c::'a::ring) * (x \<bullet> y)" by (vector ring_simps)
-lemma dot_rmult: "x \<bullet> (c *s y) = (c::'a::comm_ring) * (x \<bullet> y)" by (vector ring_simps)
-lemma dot_lneg: "(-x) \<bullet> (y::'a::ring ^ 'n) = -(x \<bullet> y)" by vector
-lemma dot_rneg: "(x::'a::ring ^ 'n) \<bullet> (-y) = -(x \<bullet> y)" by vector
-lemma dot_lzero[simp]: "0 \<bullet> x = (0::'a::{comm_monoid_add, mult_zero})" by vector
-lemma dot_rzero[simp]: "x \<bullet> 0 = (0::'a::{comm_monoid_add, mult_zero})" by vector
-lemma dot_pos_le[simp]: "(0::'a\<Colon>ordered_ring_strict) <= x \<bullet> x"
-  by (simp add: dot_def setsum_nonneg)
-
-lemma setsum_squares_eq_0_iff: assumes fS: "finite F" and fp: "\<forall>x \<in> F. f x \<ge> (0 ::'a::pordered_ab_group_add)" shows "setsum f F = 0 \<longleftrightarrow> (ALL x:F. f x = 0)"
-using fS fp setsum_nonneg[OF fp]
-proof (induct set: finite)
-  case empty thus ?case by simp
-next
-  case (insert x F)
-  from insert.prems have Fx: "f x \<ge> 0" and Fp: "\<forall> a \<in> F. f a \<ge> 0" by simp_all
-  from insert.hyps Fp setsum_nonneg[OF Fp]
-  have h: "setsum f F = 0 \<longleftrightarrow> (\<forall>a \<in>F. f a = 0)" by metis
-  from add_nonneg_eq_0_iff[OF Fx  setsum_nonneg[OF Fp]] insert.hyps(1,2)
-  show ?case by (simp add: h)
-qed
-
-lemma dot_eq_0: "x \<bullet> x = 0 \<longleftrightarrow> (x::'a::{ordered_ring_strict,ring_no_zero_divisors} ^ 'n::finite) = 0"
-  by (simp add: dot_def setsum_squares_eq_0_iff Cart_eq)
-
-lemma dot_pos_lt[simp]: "(0 < x \<bullet> x) \<longleftrightarrow> (x::'a::{ordered_ring_strict,ring_no_zero_divisors} ^ 'n::finite) \<noteq> 0" using dot_eq_0[of x] dot_pos_le[of x]
-  by (auto simp add: le_less)
-
-subsection{* The collapse of the general concepts to dimension one. *}
-
-lemma vector_one: "(x::'a ^1) = (\<chi> i. (x$1))"
-  by (simp add: Cart_eq forall_1)
-
-lemma forall_one: "(\<forall>(x::'a ^1). P x) \<longleftrightarrow> (\<forall>x. P(\<chi> i. x))"
-  apply auto
-  apply (erule_tac x= "x$1" in allE)
-  apply (simp only: vector_one[symmetric])
-  done
-
-lemma norm_vector_1: "norm (x :: _^1) = norm (x$1)"
-  by (simp add: norm_vector_def UNIV_1)
-
-lemma norm_real: "norm(x::real ^ 1) = abs(x$1)"
-  by (simp add: norm_vector_1)
-
-lemma dist_real: "dist(x::real ^ 1) y = abs((x$1) - (y$1))"
-  by (auto simp add: norm_real dist_norm)
-
-subsection {* A connectedness or intermediate value lemma with several applications. *}
-
-lemma connected_real_lemma:
-  fixes f :: "real \<Rightarrow> 'a::metric_space"
-  assumes ab: "a \<le> b" and fa: "f a \<in> e1" and fb: "f b \<in> e2"
-  and dst: "\<And>e x. a <= x \<Longrightarrow> x <= b \<Longrightarrow> 0 < e ==> \<exists>d > 0. \<forall>y. abs(y - x) < d \<longrightarrow> dist(f y) (f x) < e"
-  and e1: "\<forall>y \<in> e1. \<exists>e > 0. \<forall>y'. dist y' y < e \<longrightarrow> y' \<in> e1"
-  and e2: "\<forall>y \<in> e2. \<exists>e > 0. \<forall>y'. dist y' y < e \<longrightarrow> y' \<in> e2"
-  and e12: "~(\<exists>x \<ge> a. x <= b \<and> f x \<in> e1 \<and> f x \<in> e2)"
-  shows "\<exists>x \<ge> a. x <= b \<and> f x \<notin> e1 \<and> f x \<notin> e2" (is "\<exists> x. ?P x")
-proof-
-  let ?S = "{c. \<forall>x \<ge> a. x <= c \<longrightarrow> f x \<in> e1}"
-  have Se: " \<exists>x. x \<in> ?S" apply (rule exI[where x=a]) by (auto simp add: fa)
-  have Sub: "\<exists>y. isUb UNIV ?S y"
-    apply (rule exI[where x= b])
-    using ab fb e12 by (auto simp add: isUb_def setle_def)
-  from reals_complete[OF Se Sub] obtain l where
-    l: "isLub UNIV ?S l"by blast
-  have alb: "a \<le> l" "l \<le> b" using l ab fa fb e12
-    apply (auto simp add: isLub_def leastP_def isUb_def setle_def setge_def)
-    by (metis linorder_linear)
-  have ale1: "\<forall>z \<ge> a. z < l \<longrightarrow> f z \<in> e1" using l
-    apply (auto simp add: isLub_def leastP_def isUb_def setle_def setge_def)
-    by (metis linorder_linear not_le)
-    have th1: "\<And>z x e d :: real. z <= x + e \<Longrightarrow> e < d ==> z < x \<or> abs(z - x) < d" by arith
-    have th2: "\<And>e x:: real. 0 < e ==> ~(x + e <= x)" by arith
-    have th3: "\<And>d::real. d > 0 \<Longrightarrow> \<exists>e > 0. e < d" by dlo
-    {assume le2: "f l \<in> e2"
-      from le2 fa fb e12 alb have la: "l \<noteq> a" by metis
-      hence lap: "l - a > 0" using alb by arith
-      from e2[rule_format, OF le2] obtain e where
-        e: "e > 0" "\<forall>y. dist y (f l) < e \<longrightarrow> y \<in> e2" by metis
-      from dst[OF alb e(1)] obtain d where
-        d: "d > 0" "\<forall>y. \<bar>y - l\<bar> < d \<longrightarrow> dist (f y) (f l) < e" by metis
-      have "\<exists>d'. d' < d \<and> d' >0 \<and> l - d' > a" using lap d(1)
-        apply ferrack by arith
-      then obtain d' where d': "d' > 0" "d' < d" "l - d' > a" by metis
-      from d e have th0: "\<forall>y. \<bar>y - l\<bar> < d \<longrightarrow> f y \<in> e2" by metis
-      from th0[rule_format, of "l - d'"] d' have "f (l - d') \<in> e2" by auto
-      moreover
-      have "f (l - d') \<in> e1" using ale1[rule_format, of "l -d'"] d' by auto
-      ultimately have False using e12 alb d' by auto}
-    moreover
-    {assume le1: "f l \<in> e1"
-    from le1 fa fb e12 alb have lb: "l \<noteq> b" by metis
-      hence blp: "b - l > 0" using alb by arith
-      from e1[rule_format, OF le1] obtain e where
-        e: "e > 0" "\<forall>y. dist y (f l) < e \<longrightarrow> y \<in> e1" by metis
-      from dst[OF alb e(1)] obtain d where
-        d: "d > 0" "\<forall>y. \<bar>y - l\<bar> < d \<longrightarrow> dist (f y) (f l) < e" by metis
-      have "\<exists>d'. d' < d \<and> d' >0" using d(1) by dlo
-      then obtain d' where d': "d' > 0" "d' < d" by metis
-      from d e have th0: "\<forall>y. \<bar>y - l\<bar> < d \<longrightarrow> f y \<in> e1" by auto
-      hence "\<forall>y. l \<le> y \<and> y \<le> l + d' \<longrightarrow> f y \<in> e1" using d' by auto
-      with ale1 have "\<forall>y. a \<le> y \<and> y \<le> l + d' \<longrightarrow> f y \<in> e1" by auto
-      with l d' have False
-        by (auto simp add: isLub_def isUb_def setle_def setge_def leastP_def) }
-    ultimately show ?thesis using alb by metis
-qed
-
-text{* One immediately useful corollary is the existence of square roots! --- Should help to get rid of all the development of square-root for reals as a special case @{typ "real^1"} *}
-
-lemma square_bound_lemma: "(x::real) < (1 + x) * (1 + x)"
-proof-
-  have "(x + 1/2)^2 + 3/4 > 0" using zero_le_power2[of "x+1/2"] by arith
-  thus ?thesis by (simp add: ring_simps power2_eq_square)
-qed
-
-lemma square_continuous: "0 < (e::real) ==> \<exists>d. 0 < d \<and> (\<forall>y. abs(y - x) < d \<longrightarrow> abs(y * y - x * x) < e)"
-  using isCont_power[OF isCont_ident, of 2, unfolded isCont_def LIM_eq, rule_format, of e x] apply (auto simp add: power2_eq_square)
-  apply (rule_tac x="s" in exI)
-  apply auto
-  apply (erule_tac x=y in allE)
-  apply auto
-  done
-
-lemma real_le_lsqrt: "0 <= x \<Longrightarrow> 0 <= y \<Longrightarrow> x <= y^2 ==> sqrt x <= y"
-  using real_sqrt_le_iff[of x "y^2"] by simp
-
-lemma real_le_rsqrt: "x^2 \<le> y \<Longrightarrow> x \<le> sqrt y"
-  using real_sqrt_le_mono[of "x^2" y] by simp
-
-lemma real_less_rsqrt: "x^2 < y \<Longrightarrow> x < sqrt y"
-  using real_sqrt_less_mono[of "x^2" y] by simp
-
-lemma sqrt_even_pow2: assumes n: "even n"
-  shows "sqrt(2 ^ n) = 2 ^ (n div 2)"
-proof-
-  from n obtain m where m: "n = 2*m" unfolding even_nat_equiv_def2
-    by (auto simp add: nat_number)
-  from m  have "sqrt(2 ^ n) = sqrt ((2 ^ m) ^ 2)"
-    by (simp only: power_mult[symmetric] mult_commute)
-  then show ?thesis  using m by simp
-qed
-
-lemma real_div_sqrt: "0 <= x ==> x / sqrt(x) = sqrt(x)"
-  apply (cases "x = 0", simp_all)
-  using sqrt_divide_self_eq[of x]
-  apply (simp add: inverse_eq_divide real_sqrt_ge_0_iff field_simps)
-  done
-
-text{* Hence derive more interesting properties of the norm. *}
-
-text {*
-  This type-specific version is only here
-  to make @{text normarith.ML} happy.
-*}
-lemma norm_0: "norm (0::real ^ _) = 0"
-  by (rule norm_zero)
-
-lemma norm_mul[simp]: "norm(a *s x) = abs(a) * norm x"
-  by (simp add: norm_vector_def vector_component setL2_right_distrib
-           abs_mult cong: strong_setL2_cong)
-lemma norm_eq_0_dot: "(norm x = 0) \<longleftrightarrow> (x \<bullet> x = (0::real))"
-  by (simp add: norm_vector_def dot_def setL2_def power2_eq_square)
-lemma real_vector_norm_def: "norm x = sqrt (x \<bullet> x)"
-  by (simp add: norm_vector_def setL2_def dot_def power2_eq_square)
-lemma norm_pow_2: "norm x ^ 2 = x \<bullet> x"
-  by (simp add: real_vector_norm_def)
-lemma norm_eq_0_imp: "norm x = 0 ==> x = (0::real ^'n::finite)" by (metis norm_eq_zero)
-lemma vector_mul_eq_0[simp]: "(a *s x = 0) \<longleftrightarrow> a = (0::'a::idom) \<or> x = 0"
-  by vector
-lemma vector_mul_lcancel[simp]: "a *s x = a *s y \<longleftrightarrow> a = (0::real) \<or> x = y"
-  by (metis eq_iff_diff_eq_0 vector_mul_eq_0 vector_ssub_ldistrib)
-lemma vector_mul_rcancel[simp]: "a *s x = b *s x \<longleftrightarrow> (a::real) = b \<or> x = 0"
-  by (metis eq_iff_diff_eq_0 vector_mul_eq_0 vector_sub_rdistrib)
-lemma vector_mul_lcancel_imp: "a \<noteq> (0::real) ==>  a *s x = a *s y ==> (x = y)"
-  by (metis vector_mul_lcancel)
-lemma vector_mul_rcancel_imp: "x \<noteq> 0 \<Longrightarrow> (a::real) *s x = b *s x ==> a = b"
-  by (metis vector_mul_rcancel)
-lemma norm_cauchy_schwarz:
-  fixes x y :: "real ^ 'n::finite"
-  shows "x \<bullet> y <= norm x * norm y"
-proof-
-  {assume "norm x = 0"
-    hence ?thesis by (simp add: dot_lzero dot_rzero)}
-  moreover
-  {assume "norm y = 0"
-    hence ?thesis by (simp add: dot_lzero dot_rzero)}
-  moreover
-  {assume h: "norm x \<noteq> 0" "norm y \<noteq> 0"
-    let ?z = "norm y *s x - norm x *s y"
-    from h have p: "norm x * norm y > 0" by (metis norm_ge_zero le_less zero_compare_simps)
-    from dot_pos_le[of ?z]
-    have "(norm x * norm y) * (x \<bullet> y) \<le> norm x ^2 * norm y ^2"
-      apply (simp add: dot_rsub dot_lsub dot_lmult dot_rmult ring_simps)
-      by (simp add: norm_pow_2[symmetric] power2_eq_square dot_sym)
-    hence "x\<bullet>y \<le> (norm x ^2 * norm y ^2) / (norm x * norm y)" using p
-      by (simp add: field_simps)
-    hence ?thesis using h by (simp add: power2_eq_square)}
-  ultimately show ?thesis by metis
-qed
-
-lemma norm_cauchy_schwarz_abs:
-  fixes x y :: "real ^ 'n::finite"
-  shows "\<bar>x \<bullet> y\<bar> \<le> norm x * norm y"
-  using norm_cauchy_schwarz[of x y] norm_cauchy_schwarz[of x "-y"]
-  by (simp add: real_abs_def dot_rneg)
-
-lemma norm_triangle_sub:
-  fixes x y :: "'a::real_normed_vector"
-  shows "norm x \<le> norm y  + norm (x - y)"
-  using norm_triangle_ineq[of "y" "x - y"] by (simp add: ring_simps)
-
-lemma norm_triangle_le: "norm(x::real ^'n::finite) + norm y <= e ==> norm(x + y) <= e"
-  by (metis order_trans norm_triangle_ineq)
-lemma norm_triangle_lt: "norm(x::real ^'n::finite) + norm(y) < e ==> norm(x + y) < e"
-  by (metis basic_trans_rules(21) norm_triangle_ineq)
-
-lemma component_le_norm: "\<bar>x$i\<bar> <= norm (x::real ^ 'n::finite)"
-  apply (simp add: norm_vector_def)
-  apply (rule member_le_setL2, simp_all)
-  done
-
-lemma norm_bound_component_le: "norm(x::real ^ 'n::finite) <= e
-                ==> \<bar>x$i\<bar> <= e"
-  by (metis component_le_norm order_trans)
-
-lemma norm_bound_component_lt: "norm(x::real ^ 'n::finite) < e
-                ==> \<bar>x$i\<bar> < e"
-  by (metis component_le_norm basic_trans_rules(21))
-
-lemma norm_le_l1: "norm (x:: real ^'n::finite) <= setsum(\<lambda>i. \<bar>x$i\<bar>) UNIV"
-  by (simp add: norm_vector_def setL2_le_setsum)
-
-lemma real_abs_norm: "\<bar>norm x\<bar> = norm (x :: real ^ _)"
-  by (rule abs_norm_cancel)
-lemma real_abs_sub_norm: "\<bar>norm(x::real ^'n::finite) - norm y\<bar> <= norm(x - y)"
-  by (rule norm_triangle_ineq3)
-lemma norm_le: "norm(x::real ^ _) <= norm(y) \<longleftrightarrow> x \<bullet> x <= y \<bullet> y"
-  by (simp add: real_vector_norm_def)
-lemma norm_lt: "norm(x::real ^ _) < norm(y) \<longleftrightarrow> x \<bullet> x < y \<bullet> y"
-  by (simp add: real_vector_norm_def)
-lemma norm_eq: "norm (x::real ^ _) = norm y \<longleftrightarrow> x \<bullet> x = y \<bullet> y"
-  by (simp add: order_eq_iff norm_le)
-lemma norm_eq_1: "norm(x::real ^ _) = 1 \<longleftrightarrow> x \<bullet> x = 1"
-  by (simp add: real_vector_norm_def)
-
-text{* Squaring equations and inequalities involving norms.  *}
-
-lemma dot_square_norm: "x \<bullet> x = norm(x)^2"
-  by (simp add: real_vector_norm_def)
-
-lemma norm_eq_square: "norm(x) = a \<longleftrightarrow> 0 <= a \<and> x \<bullet> x = a^2"
-  by (auto simp add: real_vector_norm_def)
-
-lemma real_abs_le_square_iff: "\<bar>x\<bar> \<le> \<bar>y\<bar> \<longleftrightarrow> (x::real)^2 \<le> y^2"
-proof-
-  have "x^2 \<le> y^2 \<longleftrightarrow> (x -y) * (y + x) \<le> 0" by (simp add: ring_simps power2_eq_square)
-  also have "\<dots> \<longleftrightarrow> \<bar>x\<bar> \<le> \<bar>y\<bar>" apply (simp add: zero_compare_simps real_abs_def not_less) by arith
-finally show ?thesis ..
-qed
-
-lemma norm_le_square: "norm(x) <= a \<longleftrightarrow> 0 <= a \<and> x \<bullet> x <= a^2"
-  apply (simp add: dot_square_norm real_abs_le_square_iff[symmetric])
-  using norm_ge_zero[of x]
-  apply arith
-  done
-
-lemma norm_ge_square: "norm(x) >= a \<longleftrightarrow> a <= 0 \<or> x \<bullet> x >= a ^ 2"
-  apply (simp add: dot_square_norm real_abs_le_square_iff[symmetric])
-  using norm_ge_zero[of x]
-  apply arith
-  done
-
-lemma norm_lt_square: "norm(x) < a \<longleftrightarrow> 0 < a \<and> x \<bullet> x < a^2"
-  by (metis not_le norm_ge_square)
-lemma norm_gt_square: "norm(x) > a \<longleftrightarrow> a < 0 \<or> x \<bullet> x > a^2"
-  by (metis norm_le_square not_less)
-
-text{* Dot product in terms of the norm rather than conversely. *}
-
-lemma dot_norm: "x \<bullet> y = (norm(x + y) ^2 - norm x ^ 2 - norm y ^ 2) / 2"
-  by (simp add: norm_pow_2 dot_ladd dot_radd dot_sym)
-
-lemma dot_norm_neg: "x \<bullet> y = ((norm x ^ 2 + norm y ^ 2) - norm(x - y) ^ 2) / 2"
-  by (simp add: norm_pow_2 dot_ladd dot_radd dot_lsub dot_rsub dot_sym)
-
-
-text{* Equality of vectors in terms of @{term "op \<bullet>"} products.    *}
-
-lemma vector_eq: "(x:: real ^ 'n::finite) = y \<longleftrightarrow> x \<bullet> x = x \<bullet> y\<and> y \<bullet> y = x \<bullet> x" (is "?lhs \<longleftrightarrow> ?rhs")
-proof
-  assume "?lhs" then show ?rhs by simp
-next
-  assume ?rhs
-  then have "x \<bullet> x - x \<bullet> y = 0 \<and> x \<bullet> y - y\<bullet> y = 0" by simp
-  hence "x \<bullet> (x - y) = 0 \<and> y \<bullet> (x - y) = 0"
-    by (simp add: dot_rsub dot_lsub dot_sym)
-  then have "(x - y) \<bullet> (x - y) = 0" by (simp add: ring_simps dot_lsub dot_rsub)
-  then show "x = y" by (simp add: dot_eq_0)
-qed
-
-
-subsection{* General linear decision procedure for normed spaces. *}
-
-lemma norm_cmul_rule_thm:
-  fixes x :: "'a::real_normed_vector"
-  shows "b >= norm(x) ==> \<bar>c\<bar> * b >= norm(scaleR c x)"
-  unfolding norm_scaleR
-  apply (erule mult_mono1)
-  apply simp
-  done
-
-  (* FIXME: Move all these theorems into the ML code using lemma antiquotation *)
-lemma norm_add_rule_thm:
-  fixes x1 x2 :: "'a::real_normed_vector"
-  shows "norm x1 \<le> b1 \<Longrightarrow> norm x2 \<le> b2 \<Longrightarrow> norm (x1 + x2) \<le> b1 + b2"
-  by (rule order_trans [OF norm_triangle_ineq add_mono])
-
-lemma ge_iff_diff_ge_0: "(a::'a::ordered_ring) \<ge> b == a - b \<ge> 0"
-  by (simp add: ring_simps)
-
-lemma pth_1:
-  fixes x :: "'a::real_normed_vector"
-  shows "x == scaleR 1 x" by simp
-
-lemma pth_2:
-  fixes x :: "'a::real_normed_vector"
-  shows "x - y == x + -y" by (atomize (full)) simp
-
-lemma pth_3:
-  fixes x :: "'a::real_normed_vector"
-  shows "- x == scaleR (-1) x" by simp
-
-lemma pth_4:
-  fixes x :: "'a::real_normed_vector"
-  shows "scaleR 0 x == 0" and "scaleR c 0 = (0::'a)" by simp_all
-
-lemma pth_5:
-  fixes x :: "'a::real_normed_vector"
-  shows "scaleR c (scaleR d x) == scaleR (c * d) x" by simp
-
-lemma pth_6:
-  fixes x :: "'a::real_normed_vector"
-  shows "scaleR c (x + y) == scaleR c x + scaleR c y"
-  by (simp add: scaleR_right_distrib)
-
-lemma pth_7:
-  fixes x :: "'a::real_normed_vector"
-  shows "0 + x == x" and "x + 0 == x" by simp_all
-
-lemma pth_8:
-  fixes x :: "'a::real_normed_vector"
-  shows "scaleR c x + scaleR d x == scaleR (c + d) x"
-  by (simp add: scaleR_left_distrib)
-
-lemma pth_9:
-  fixes x :: "'a::real_normed_vector" shows
-  "(scaleR c x + z) + scaleR d x == scaleR (c + d) x + z"
-  "scaleR c x + (scaleR d x + z) == scaleR (c + d) x + z"
-  "(scaleR c x + w) + (scaleR d x + z) == scaleR (c + d) x + (w + z)"
-  by (simp_all add: algebra_simps)
-
-lemma pth_a:
-  fixes x :: "'a::real_normed_vector"
-  shows "scaleR 0 x + y == y" by simp
-
-lemma pth_b:
-  fixes x :: "'a::real_normed_vector" shows
-  "scaleR c x + scaleR d y == scaleR c x + scaleR d y"
-  "(scaleR c x + z) + scaleR d y == scaleR c x + (z + scaleR d y)"
-  "scaleR c x + (scaleR d y + z) == scaleR c x + (scaleR d y + z)"
-  "(scaleR c x + w) + (scaleR d y + z) == scaleR c x + (w + (scaleR d y + z))"
-  by (simp_all add: algebra_simps)
-
-lemma pth_c:
-  fixes x :: "'a::real_normed_vector" shows
-  "scaleR c x + scaleR d y == scaleR d y + scaleR c x"
-  "(scaleR c x + z) + scaleR d y == scaleR d y + (scaleR c x + z)"
-  "scaleR c x + (scaleR d y + z) == scaleR d y + (scaleR c x + z)"
-  "(scaleR c x + w) + (scaleR d y + z) == scaleR d y + ((scaleR c x + w) + z)"
-  by (simp_all add: algebra_simps)
-
-lemma pth_d:
-  fixes x :: "'a::real_normed_vector"
-  shows "x + 0 == x" by simp
-
-lemma norm_imp_pos_and_ge:
-  fixes x :: "'a::real_normed_vector"
-  shows "norm x == n \<Longrightarrow> norm x \<ge> 0 \<and> n \<ge> norm x"
-  by atomize auto
-
-lemma real_eq_0_iff_le_ge_0: "(x::real) = 0 == x \<ge> 0 \<and> -x \<ge> 0" by arith
-
-lemma norm_pths:
-  fixes x :: "'a::real_normed_vector" shows
-  "x = y \<longleftrightarrow> norm (x - y) \<le> 0"
-  "x \<noteq> y \<longleftrightarrow> \<not> (norm (x - y) \<le> 0)"
-  using norm_ge_zero[of "x - y"] by auto
-
-lemma vector_dist_norm:
-  fixes x :: "'a::real_normed_vector"
-  shows "dist x y = norm (x - y)"
-  by (rule dist_norm)
-
-use "normarith.ML"
-
-method_setup norm = {* Scan.succeed (SIMPLE_METHOD' o NormArith.norm_arith_tac)
-*} "Proves simple linear statements about vector norms"
-
-
-text{* Hence more metric properties. *}
-
-lemma dist_triangle_alt:
-  fixes x y z :: "'a::metric_space"
-  shows "dist y z <= dist x y + dist x z"
-using dist_triangle [of y z x] by (simp add: dist_commute)
-
-lemma dist_pos_lt:
-  fixes x y :: "'a::metric_space"
-  shows "x \<noteq> y ==> 0 < dist x y"
-by (simp add: zero_less_dist_iff)
-
-lemma dist_nz:
-  fixes x y :: "'a::metric_space"
-  shows "x \<noteq> y \<longleftrightarrow> 0 < dist x y"
-by (simp add: zero_less_dist_iff)
-
-lemma dist_triangle_le:
-  fixes x y z :: "'a::metric_space"
-  shows "dist x z + dist y z <= e \<Longrightarrow> dist x y <= e"
-by (rule order_trans [OF dist_triangle2])
-
-lemma dist_triangle_lt:
-  fixes x y z :: "'a::metric_space"
-  shows "dist x z + dist y z < e ==> dist x y < e"
-by (rule le_less_trans [OF dist_triangle2])
-
-lemma dist_triangle_half_l:
-  fixes x1 x2 y :: "'a::metric_space"
-  shows "dist x1 y < e / 2 \<Longrightarrow> dist x2 y < e / 2 \<Longrightarrow> dist x1 x2 < e"
-by (rule dist_triangle_lt [where z=y], simp)
-
-lemma dist_triangle_half_r:
-  fixes x1 x2 y :: "'a::metric_space"
-  shows "dist y x1 < e / 2 \<Longrightarrow> dist y x2 < e / 2 \<Longrightarrow> dist x1 x2 < e"
-by (rule dist_triangle_half_l, simp_all add: dist_commute)
-
-lemma dist_triangle_add:
-  fixes x y x' y' :: "'a::real_normed_vector"
-  shows "dist (x + y) (x' + y') <= dist x x' + dist y y'"
-  by norm
-
-lemma dist_mul[simp]: "dist (c *s x) (c *s y) = \<bar>c\<bar> * dist x y"
-  unfolding dist_norm vector_ssub_ldistrib[symmetric] norm_mul ..
-
-lemma dist_triangle_add_half:
-  fixes x x' y y' :: "'a::real_normed_vector"
-  shows "dist x x' < e / 2 \<Longrightarrow> dist y y' < e / 2 \<Longrightarrow> dist(x + y) (x' + y') < e"
-  by norm
-
-lemma setsum_component [simp]:
-  fixes f:: " 'a \<Rightarrow> ('b::comm_monoid_add) ^'n"
-  shows "(setsum f S)$i = setsum (\<lambda>x. (f x)$i) S"
-  by (cases "finite S", induct S set: finite, simp_all)
-
-lemma setsum_eq: "setsum f S = (\<chi> i. setsum (\<lambda>x. (f x)$i ) S)"
-  by (simp add: Cart_eq)
-
-lemma setsum_clauses:
-  shows "setsum f {} = 0"
-  and "finite S \<Longrightarrow> setsum f (insert x S) =
-                 (if x \<in> S then setsum f S else f x + setsum f S)"
-  by (auto simp add: insert_absorb)
-
-lemma setsum_cmul:
-  fixes f:: "'c \<Rightarrow> ('a::semiring_1)^'n"
-  shows "setsum (\<lambda>x. c *s f x) S = c *s setsum f S"
-  by (simp add: Cart_eq setsum_right_distrib)
-
-lemma setsum_norm:
-  fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"
-  assumes fS: "finite S"
-  shows "norm (setsum f S) <= setsum (\<lambda>x. norm(f x)) S"
-proof(induct rule: finite_induct[OF fS])
-  case 1 thus ?case by simp
-next
-  case (2 x S)
-  from "2.hyps" have "norm (setsum f (insert x S)) \<le> norm (f x) + norm (setsum f S)" by (simp add: norm_triangle_ineq)
-  also have "\<dots> \<le> norm (f x) + setsum (\<lambda>x. norm(f x)) S"
-    using "2.hyps" by simp
-  finally  show ?case  using "2.hyps" by simp
-qed
-
-lemma real_setsum_norm:
-  fixes f :: "'a \<Rightarrow> real ^'n::finite"
-  assumes fS: "finite S"
-  shows "norm (setsum f S) <= setsum (\<lambda>x. norm(f x)) S"
-proof(induct rule: finite_induct[OF fS])
-  case 1 thus ?case by simp
-next
-  case (2 x S)
-  from "2.hyps" have "norm (setsum f (insert x S)) \<le> norm (f x) + norm (setsum f S)" by (simp add: norm_triangle_ineq)
-  also have "\<dots> \<le> norm (f x) + setsum (\<lambda>x. norm(f x)) S"
-    using "2.hyps" by simp
-  finally  show ?case  using "2.hyps" by simp
-qed
-
-lemma setsum_norm_le:
-  fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"
-  assumes fS: "finite S"
-  and fg: "\<forall>x \<in> S. norm (f x) \<le> g x"
-  shows "norm (setsum f S) \<le> setsum g S"
-proof-
-  from fg have "setsum (\<lambda>x. norm(f x)) S <= setsum g S"
-    by - (rule setsum_mono, simp)
-  then show ?thesis using setsum_norm[OF fS, of f] fg
-    by arith
-qed
-
-lemma real_setsum_norm_le:
-  fixes f :: "'a \<Rightarrow> real ^ 'n::finite"
-  assumes fS: "finite S"
-  and fg: "\<forall>x \<in> S. norm (f x) \<le> g x"
-  shows "norm (setsum f S) \<le> setsum g S"
-proof-
-  from fg have "setsum (\<lambda>x. norm(f x)) S <= setsum g S"
-    by - (rule setsum_mono, simp)
-  then show ?thesis using real_setsum_norm[OF fS, of f] fg
-    by arith
-qed
-
-lemma setsum_norm_bound:
-  fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"
-  assumes fS: "finite S"
-  and K: "\<forall>x \<in> S. norm (f x) \<le> K"
-  shows "norm (setsum f S) \<le> of_nat (card S) * K"
-  using setsum_norm_le[OF fS K] setsum_constant[symmetric]
-  by simp
-
-lemma real_setsum_norm_bound:
-  fixes f :: "'a \<Rightarrow> real ^ 'n::finite"
-  assumes fS: "finite S"
-  and K: "\<forall>x \<in> S. norm (f x) \<le> K"
-  shows "norm (setsum f S) \<le> of_nat (card S) * K"
-  using real_setsum_norm_le[OF fS K] setsum_constant[symmetric]
-  by simp
-
-lemma setsum_vmul:
-  fixes f :: "'a \<Rightarrow> 'b::{real_normed_vector,semiring, mult_zero}"
-  assumes fS: "finite S"
-  shows "setsum f S *s v = setsum (\<lambda>x. f x *s v) S"
-proof(induct rule: finite_induct[OF fS])
-  case 1 then show ?case by (simp add: vector_smult_lzero)
-next
-  case (2 x F)
-  from "2.hyps" have "setsum f (insert x F) *s v = (f x + setsum f F) *s v"
-    by simp
-  also have "\<dots> = f x *s v + setsum f F *s v"
-    by (simp add: vector_sadd_rdistrib)
-  also have "\<dots> = setsum (\<lambda>x. f x *s v) (insert x F)" using "2.hyps" by simp
-  finally show ?case .
-qed
-
-(* FIXME : Problem thm setsum_vmul[of _ "f:: 'a \<Rightarrow> real ^'n"]  ---
- Get rid of *s and use real_vector instead! Also prove that ^ creates a real_vector !! *)
-
-    (* FIXME: Here too need stupid finiteness assumption on T!!! *)
-lemma setsum_group:
-  assumes fS: "finite S" and fT: "finite T" and fST: "f ` S \<subseteq> T"
-  shows "setsum (\<lambda>y. setsum g {x. x\<in> S \<and> f x = y}) T = setsum g S"
-
-apply (subst setsum_image_gen[OF fS, of g f])
-apply (rule setsum_mono_zero_right[OF fT fST])
-by (auto intro: setsum_0')
-
-lemma vsum_norm_allsubsets_bound:
-  fixes f:: "'a \<Rightarrow> real ^'n::finite"
-  assumes fP: "finite P" and fPs: "\<And>Q. Q \<subseteq> P \<Longrightarrow> norm (setsum f Q) \<le> e"
-  shows "setsum (\<lambda>x. norm (f x)) P \<le> 2 * real CARD('n) *  e"
-proof-
-  let ?d = "real CARD('n)"
-  let ?nf = "\<lambda>x. norm (f x)"
-  let ?U = "UNIV :: 'n set"
-  have th0: "setsum (\<lambda>x. setsum (\<lambda>i. \<bar>f x $ i\<bar>) ?U) P = setsum (\<lambda>i. setsum (\<lambda>x. \<bar>f x $ i\<bar>) P) ?U"
-    by (rule setsum_commute)
-  have th1: "2 * ?d * e = of_nat (card ?U) * (2 * e)" by (simp add: real_of_nat_def)
-  have "setsum ?nf P \<le> setsum (\<lambda>x. setsum (\<lambda>i. \<bar>f x $ i\<bar>) ?U) P"
-    apply (rule setsum_mono)
-    by (rule norm_le_l1)
-  also have "\<dots> \<le> 2 * ?d * e"
-    unfolding th0 th1
-  proof(rule setsum_bounded)
-    fix i assume i: "i \<in> ?U"
-    let ?Pp = "{x. x\<in> P \<and> f x $ i \<ge> 0}"
-    let ?Pn = "{x. x \<in> P \<and> f x $ i < 0}"
-    have thp: "P = ?Pp \<union> ?Pn" by auto
-    have thp0: "?Pp \<inter> ?Pn ={}" by auto
-    have PpP: "?Pp \<subseteq> P" and PnP: "?Pn \<subseteq> P" by blast+
-    have Ppe:"setsum (\<lambda>x. \<bar>f x $ i\<bar>) ?Pp \<le> e"
-      using component_le_norm[of "setsum (\<lambda>x. f x) ?Pp" i]  fPs[OF PpP]
-      by (auto intro: abs_le_D1)
-    have Pne: "setsum (\<lambda>x. \<bar>f x $ i\<bar>) ?Pn \<le> e"
-      using component_le_norm[of "setsum (\<lambda>x. - f x) ?Pn" i]  fPs[OF PnP]
-      by (auto simp add: setsum_negf intro: abs_le_D1)
-    have "setsum (\<lambda>x. \<bar>f x $ i\<bar>) P = setsum (\<lambda>x. \<bar>f x $ i\<bar>) ?Pp + setsum (\<lambda>x. \<bar>f x $ i\<bar>) ?Pn"
-      apply (subst thp)
-      apply (rule setsum_Un_zero)
-      using fP thp0 by auto
-    also have "\<dots> \<le> 2*e" using Pne Ppe by arith
-    finally show "setsum (\<lambda>x. \<bar>f x $ i\<bar>) P \<le> 2*e" .
-  qed
-  finally show ?thesis .
-qed
-
-lemma dot_lsum: "finite S \<Longrightarrow> setsum f S \<bullet> (y::'a::{comm_ring}^'n) = setsum (\<lambda>x. f x \<bullet> y) S "
-  by (induct rule: finite_induct, auto simp add: dot_lzero dot_ladd dot_radd)
-
-lemma dot_rsum: "finite S \<Longrightarrow> (y::'a::{comm_ring}^'n) \<bullet> setsum f S = setsum (\<lambda>x. y \<bullet> f x) S "
-  by (induct rule: finite_induct, auto simp add: dot_rzero dot_radd)
-
-subsection{* Basis vectors in coordinate directions. *}
-
-
-definition "basis k = (\<chi> i. if i = k then 1 else 0)"
-
-lemma basis_component [simp]: "basis k $ i = (if k=i then 1 else 0)"
-  unfolding basis_def by simp
-
-lemma delta_mult_idempotent:
-  "(if k=a then 1 else (0::'a::semiring_1)) * (if k=a then 1 else 0) = (if k=a then 1 else 0)" by (cases "k=a", auto)
-
-lemma norm_basis:
-  shows "norm (basis k :: real ^'n::finite) = 1"
-  apply (simp add: basis_def real_vector_norm_def dot_def)
-  apply (vector delta_mult_idempotent)
-  using setsum_delta[of "UNIV :: 'n set" "k" "\<lambda>k. 1::real"]
-  apply auto
-  done
-
-lemma norm_basis_1: "norm(basis 1 :: real ^'n::{finite,one}) = 1"
-  by (rule norm_basis)
-
-lemma vector_choose_size: "0 <= c ==> \<exists>(x::real^'n::finite). norm x = c"
-  apply (rule exI[where x="c *s basis arbitrary"])
-  by (simp only: norm_mul norm_basis)
-
-lemma vector_choose_dist: assumes e: "0 <= e"
-  shows "\<exists>(y::real^'n::finite). dist x y = e"
-proof-
-  from vector_choose_size[OF e] obtain c:: "real ^'n"  where "norm c = e"
-    by blast
-  then have "dist x (x - c) = e" by (simp add: dist_norm)
-  then show ?thesis by blast
-qed
-
-lemma basis_inj: "inj (basis :: 'n \<Rightarrow> real ^'n::finite)"
-  by (simp add: inj_on_def Cart_eq)
-
-lemma cond_value_iff: "f (if b then x else y) = (if b then f x else f y)"
-  by auto
-
-lemma basis_expansion:
-  "setsum (\<lambda>i. (x$i) *s basis i) UNIV = (x::('a::ring_1) ^'n::finite)" (is "?lhs = ?rhs" is "setsum ?f ?S = _")
-  by (auto simp add: Cart_eq cond_value_iff setsum_delta[of "?S", where ?'b = "'a", simplified] cong del: if_weak_cong)
-
-lemma basis_expansion_unique:
-  "setsum (\<lambda>i. f i *s basis i) UNIV = (x::('a::comm_ring_1) ^'n::finite) \<longleftrightarrow> (\<forall>i. f i = x$i)"
-  by (simp add: Cart_eq setsum_delta cond_value_iff cong del: if_weak_cong)
-
-lemma cond_application_beta: "(if b then f else g) x = (if b then f x else g x)"
-  by auto
-
-lemma dot_basis:
-  shows "basis i \<bullet> x = x$i" "x \<bullet> (basis i :: 'a^'n::finite) = (x$i :: 'a::semiring_1)"
-  by (auto simp add: dot_def basis_def cond_application_beta  cond_value_iff setsum_delta cong del: if_weak_cong)
-
-lemma inner_basis:
-  fixes x :: "'a::{real_inner, real_algebra_1} ^ 'n::finite"
-  shows "inner (basis i) x = inner 1 (x $ i)"
-    and "inner x (basis i) = inner (x $ i) 1"
-  unfolding inner_vector_def basis_def
-  by (auto simp add: cond_application_beta  cond_value_iff setsum_delta cong del: if_weak_cong)
-
-lemma basis_eq_0: "basis i = (0::'a::semiring_1^'n) \<longleftrightarrow> False"
-  by (auto simp add: Cart_eq)
-
-lemma basis_nonzero:
-  shows "basis k \<noteq> (0:: 'a::semiring_1 ^'n)"
-  by (simp add: basis_eq_0)
-
-lemma vector_eq_ldot: "(\<forall>x. x \<bullet> y = x \<bullet> z) \<longleftrightarrow> y = (z::'a::semiring_1^'n::finite)"
-  apply (auto simp add: Cart_eq dot_basis)
-  apply (erule_tac x="basis i" in allE)
-  apply (simp add: dot_basis)
-  apply (subgoal_tac "y = z")
-  apply simp
-  apply (simp add: Cart_eq)
-  done
-
-lemma vector_eq_rdot: "(\<forall>z. x \<bullet> z = y \<bullet> z) \<longleftrightarrow> x = (y::'a::semiring_1^'n::finite)"
-  apply (auto simp add: Cart_eq dot_basis)
-  apply (erule_tac x="basis i" in allE)
-  apply (simp add: dot_basis)
-  apply (subgoal_tac "x = y")
-  apply simp
-  apply (simp add: Cart_eq)
-  done
-
-subsection{* Orthogonality. *}
-
-definition "orthogonal x y \<longleftrightarrow> (x \<bullet> y = 0)"
-
-lemma orthogonal_basis:
-  shows "orthogonal (basis i :: 'a^'n::finite) x \<longleftrightarrow> x$i = (0::'a::ring_1)"
-  by (auto simp add: orthogonal_def dot_def basis_def cond_value_iff cond_application_beta setsum_delta cong del: if_weak_cong)
-
-lemma orthogonal_basis_basis:
-  shows "orthogonal (basis i :: 'a::ring_1^'n::finite) (basis j) \<longleftrightarrow> i \<noteq> j"
-  unfolding orthogonal_basis[of i] basis_component[of j] by simp
-
-  (* FIXME : Maybe some of these require less than comm_ring, but not all*)
-lemma orthogonal_clauses:
-  "orthogonal a (0::'a::comm_ring ^'n)"
-  "orthogonal a x ==> orthogonal a (c *s x)"
-  "orthogonal a x ==> orthogonal a (-x)"
-  "orthogonal a x \<Longrightarrow> orthogonal a y ==> orthogonal a (x + y)"
-  "orthogonal a x \<Longrightarrow> orthogonal a y ==> orthogonal a (x - y)"
-  "orthogonal 0 a"
-  "orthogonal x a ==> orthogonal (c *s x) a"
-  "orthogonal x a ==> orthogonal (-x) a"
-  "orthogonal x a \<Longrightarrow> orthogonal y a ==> orthogonal (x + y) a"
-  "orthogonal x a \<Longrightarrow> orthogonal y a ==> orthogonal (x - y) a"
-  unfolding orthogonal_def dot_rneg dot_rmult dot_radd dot_rsub
-  dot_lzero dot_rzero dot_lneg dot_lmult dot_ladd dot_lsub
-  by simp_all
-
-lemma orthogonal_commute: "orthogonal (x::'a::{ab_semigroup_mult,comm_monoid_add} ^'n)y \<longleftrightarrow> orthogonal y x"
-  by (simp add: orthogonal_def dot_sym)
-
-subsection{* Explicit vector construction from lists. *}
-
-primrec from_nat :: "nat \<Rightarrow> 'a::{monoid_add,one}"
-where "from_nat 0 = 0" | "from_nat (Suc n) = 1 + from_nat n"
-
-lemma from_nat [simp]: "from_nat = of_nat"
-by (rule ext, induct_tac x, simp_all)
-
-primrec
-  list_fun :: "nat \<Rightarrow> _ list \<Rightarrow> _ \<Rightarrow> _"
-where
-  "list_fun n [] = (\<lambda>x. 0)"
-| "list_fun n (x # xs) = fun_upd (list_fun (Suc n) xs) (from_nat n) x"
-
-definition "vector l = (\<chi> i. list_fun 1 l i)"
-(*definition "vector l = (\<chi> i. if i <= length l then l ! (i - 1) else 0)"*)
-
-lemma vector_1: "(vector[x]) $1 = x"
-  unfolding vector_def by simp
-
-lemma vector_2:
- "(vector[x,y]) $1 = x"
- "(vector[x,y] :: 'a^2)$2 = (y::'a::zero)"
-  unfolding vector_def by simp_all
-
-lemma vector_3:
- "(vector [x,y,z] ::('a::zero)^3)$1 = x"
- "(vector [x,y,z] ::('a::zero)^3)$2 = y"
- "(vector [x,y,z] ::('a::zero)^3)$3 = z"
-  unfolding vector_def by simp_all
-
-lemma forall_vector_1: "(\<forall>v::'a::zero^1. P v) \<longleftrightarrow> (\<forall>x. P(vector[x]))"
-  apply auto
-  apply (erule_tac x="v$1" in allE)
-  apply (subgoal_tac "vector [v$1] = v")
-  apply simp
-  apply (vector vector_def)
-  apply (simp add: forall_1)
-  done
-
-lemma forall_vector_2: "(\<forall>v::'a::zero^2. P v) \<longleftrightarrow> (\<forall>x y. P(vector[x, y]))"
-  apply auto
-  apply (erule_tac x="v$1" in allE)
-  apply (erule_tac x="v$2" in allE)
-  apply (subgoal_tac "vector [v$1, v$2] = v")
-  apply simp
-  apply (vector vector_def)
-  apply (simp add: forall_2)
-  done
-
-lemma forall_vector_3: "(\<forall>v::'a::zero^3. P v) \<longleftrightarrow> (\<forall>x y z. P(vector[x, y, z]))"
-  apply auto
-  apply (erule_tac x="v$1" in allE)
-  apply (erule_tac x="v$2" in allE)
-  apply (erule_tac x="v$3" in allE)
-  apply (subgoal_tac "vector [v$1, v$2, v$3] = v")
-  apply simp
-  apply (vector vector_def)
-  apply (simp add: forall_3)
-  done
-
-subsection{* Linear functions. *}
-
-definition "linear f \<longleftrightarrow> (\<forall>x y. f(x + y) = f x + f y) \<and> (\<forall>c x. f(c *s x) = c *s f x)"
-
-lemma linear_compose_cmul: "linear f ==> linear (\<lambda>x. (c::'a::comm_semiring) *s f x)"
-  by (vector linear_def Cart_eq ring_simps)
-
-lemma linear_compose_neg: "linear (f :: 'a ^'n \<Rightarrow> 'a::comm_ring ^'m) ==> linear (\<lambda>x. -(f(x)))" by (vector linear_def Cart_eq)
-
-lemma linear_compose_add: "linear (f :: 'a ^'n \<Rightarrow> 'a::semiring_1 ^'m) \<Longrightarrow> linear g ==> linear (\<lambda>x. f(x) + g(x))"
-  by (vector linear_def Cart_eq ring_simps)
-
-lemma linear_compose_sub: "linear (f :: 'a ^'n \<Rightarrow> 'a::ring_1 ^'m) \<Longrightarrow> linear g ==> linear (\<lambda>x. f x - g x)"
-  by (vector linear_def Cart_eq ring_simps)
-
-lemma linear_compose: "linear f \<Longrightarrow> linear g ==> linear (g o f)"
-  by (simp add: linear_def)
-
-lemma linear_id: "linear id" by (simp add: linear_def id_def)
-
-lemma linear_zero: "linear (\<lambda>x. 0::'a::semiring_1 ^ 'n)" by (simp add: linear_def)
-
-lemma linear_compose_setsum:
-  assumes fS: "finite S" and lS: "\<forall>a \<in> S. linear (f a :: 'a::semiring_1 ^ 'n \<Rightarrow> 'a ^ 'm)"
-  shows "linear(\<lambda>x. setsum (\<lambda>a. f a x :: 'a::semiring_1 ^'m) S)"
-  using lS
-  apply (induct rule: finite_induct[OF fS])
-  by (auto simp add: linear_zero intro: linear_compose_add)
-
-lemma linear_vmul_component:
-  fixes f:: "'a::semiring_1^'m \<Rightarrow> 'a^'n"
-  assumes lf: "linear f"
-  shows "linear (\<lambda>x. f x $ k *s v)"
-  using lf
-  apply (auto simp add: linear_def )
-  by (vector ring_simps)+
-
-lemma linear_0: "linear f ==> f 0 = (0::'a::semiring_1 ^'n)"
-  unfolding linear_def
-  apply clarsimp
-  apply (erule allE[where x="0::'a"])
-  apply simp
-  done
-
-lemma linear_cmul: "linear f ==> f(c*s x) = c *s f x" by (simp add: linear_def)
-
-lemma linear_neg: "linear (f :: 'a::ring_1 ^'n \<Rightarrow> _) ==> f (-x) = - f x"
-  unfolding vector_sneg_minus1
-  using linear_cmul[of f] by auto
-
-lemma linear_add: "linear f ==> f(x + y) = f x + f y" by (metis linear_def)
-
-lemma linear_sub: "linear (f::'a::ring_1 ^'n \<Rightarrow> _) ==> f(x - y) = f x - f y"
-  by (simp add: diff_def linear_add linear_neg)
-
-lemma linear_setsum:
-  fixes f:: "'a::semiring_1^'n \<Rightarrow> _"
-  assumes lf: "linear f" and fS: "finite S"
-  shows "f (setsum g S) = setsum (f o g) S"
-proof (induct rule: finite_induct[OF fS])
-  case 1 thus ?case by (simp add: linear_0[OF lf])
-next
-  case (2 x F)
-  have "f (setsum g (insert x F)) = f (g x + setsum g F)" using "2.hyps"
-    by simp
-  also have "\<dots> = f (g x) + f (setsum g F)" using linear_add[OF lf] by simp
-  also have "\<dots> = setsum (f o g) (insert x F)" using "2.hyps" by simp
-  finally show ?case .
-qed
-
-lemma linear_setsum_mul:
-  fixes f:: "'a ^'n \<Rightarrow> 'a::semiring_1^'m"
-  assumes lf: "linear f" and fS: "finite S"
-  shows "f (setsum (\<lambda>i. c i *s v i) S) = setsum (\<lambda>i. c i *s f (v i)) S"
-  using linear_setsum[OF lf fS, of "\<lambda>i. c i *s v i" , unfolded o_def]
-  linear_cmul[OF lf] by simp
-
-lemma linear_injective_0:
-  assumes lf: "linear (f:: 'a::ring_1 ^ 'n \<Rightarrow> _)"
-  shows "inj f \<longleftrightarrow> (\<forall>x. f x = 0 \<longrightarrow> x = 0)"
-proof-
-  have "inj f \<longleftrightarrow> (\<forall> x y. f x = f y \<longrightarrow> x = y)" by (simp add: inj_on_def)
-  also have "\<dots> \<longleftrightarrow> (\<forall> x y. f x - f y = 0 \<longrightarrow> x - y = 0)" by simp
-  also have "\<dots> \<longleftrightarrow> (\<forall> x y. f (x - y) = 0 \<longrightarrow> x - y = 0)"
-    by (simp add: linear_sub[OF lf])
-  also have "\<dots> \<longleftrightarrow> (\<forall> x. f x = 0 \<longrightarrow> x = 0)" by auto
-  finally show ?thesis .
-qed
-
-lemma linear_bounded:
-  fixes f:: "real ^'m::finite \<Rightarrow> real ^'n::finite"
-  assumes lf: "linear f"
-  shows "\<exists>B. \<forall>x. norm (f x) \<le> B * norm x"
-proof-
-  let ?S = "UNIV:: 'm set"
-  let ?B = "setsum (\<lambda>i. norm(f(basis i))) ?S"
-  have fS: "finite ?S" by simp
-  {fix x:: "real ^ 'm"
-    let ?g = "(\<lambda>i. (x$i) *s (basis i) :: real ^ 'm)"
-    have "norm (f x) = norm (f (setsum (\<lambda>i. (x$i) *s (basis i)) ?S))"
-      by (simp only:  basis_expansion)
-    also have "\<dots> = norm (setsum (\<lambda>i. (x$i) *s f (basis i))?S)"
-      using linear_setsum[OF lf fS, of ?g, unfolded o_def] linear_cmul[OF lf]
-      by auto
-    finally have th0: "norm (f x) = norm (setsum (\<lambda>i. (x$i) *s f (basis i))?S)" .
-    {fix i assume i: "i \<in> ?S"
-      from component_le_norm[of x i]
-      have "norm ((x$i) *s f (basis i :: real ^'m)) \<le> norm (f (basis i)) * norm x"
-      unfolding norm_mul
-      apply (simp only: mult_commute)
-      apply (rule mult_mono)
-      by (auto simp add: ring_simps norm_ge_zero) }
-    then have th: "\<forall>i\<in> ?S. norm ((x$i) *s f (basis i :: real ^'m)) \<le> norm (f (basis i)) * norm x" by metis
-    from real_setsum_norm_le[OF fS, of "\<lambda>i. (x$i) *s (f (basis i))", OF th]
-    have "norm (f x) \<le> ?B * norm x" unfolding th0 setsum_left_distrib by metis}
-  then show ?thesis by blast
-qed
-
-lemma linear_bounded_pos:
-  fixes f:: "real ^'n::finite \<Rightarrow> real ^ 'm::finite"
-  assumes lf: "linear f"
-  shows "\<exists>B > 0. \<forall>x. norm (f x) \<le> B * norm x"
-proof-
-  from linear_bounded[OF lf] obtain B where
-    B: "\<forall>x. norm (f x) \<le> B * norm x" by blast
-  let ?K = "\<bar>B\<bar> + 1"
-  have Kp: "?K > 0" by arith
-    {assume C: "B < 0"
-      have "norm (1::real ^ 'n) > 0" by (simp add: zero_less_norm_iff)
-      with C have "B * norm (1:: real ^ 'n) < 0"
-        by (simp add: zero_compare_simps)
-      with B[rule_format, of 1] norm_ge_zero[of "f 1"] have False by simp
-    }
-    then have Bp: "B \<ge> 0" by ferrack
-    {fix x::"real ^ 'n"
-      have "norm (f x) \<le> ?K *  norm x"
-      using B[rule_format, of x] norm_ge_zero[of x] norm_ge_zero[of "f x"] Bp
-      apply (auto simp add: ring_simps split add: abs_split)
-      apply (erule order_trans, simp)
-      done
-  }
-  then show ?thesis using Kp by blast
-qed
-
-lemma smult_conv_scaleR: "c *s x = scaleR c x"
-  unfolding vector_scalar_mult_def vector_scaleR_def by simp
-
-lemma linear_conv_bounded_linear:
-  fixes f :: "real ^ _ \<Rightarrow> real ^ _"
-  shows "linear f \<longleftrightarrow> bounded_linear f"
-proof
-  assume "linear f"
-  show "bounded_linear f"
-  proof
-    fix x y show "f (x + y) = f x + f y"
-      using `linear f` unfolding linear_def by simp
-  next
-    fix r x show "f (scaleR r x) = scaleR r (f x)"
-      using `linear f` unfolding linear_def
-      by (simp add: smult_conv_scaleR)
-  next
-    have "\<exists>B. \<forall>x. norm (f x) \<le> B * norm x"
-      using `linear f` by (rule linear_bounded)
-    thus "\<exists>K. \<forall>x. norm (f x) \<le> norm x * K"
-      by (simp add: mult_commute)
-  qed
-next
-  assume "bounded_linear f"
-  then interpret f: bounded_linear f .
-  show "linear f"
-    unfolding linear_def smult_conv_scaleR
-    by (simp add: f.add f.scaleR)
-qed
-
-subsection{* Bilinear functions. *}
-
-definition "bilinear f \<longleftrightarrow> (\<forall>x. linear(\<lambda>y. f x y)) \<and> (\<forall>y. linear(\<lambda>x. f x y))"
-
-lemma bilinear_ladd: "bilinear h ==> h (x + y) z = (h x z) + (h y z)"
-  by (simp add: bilinear_def linear_def)
-lemma bilinear_radd: "bilinear h ==> h x (y + z) = (h x y) + (h x z)"
-  by (simp add: bilinear_def linear_def)
-
-lemma bilinear_lmul: "bilinear h ==> h (c *s x) y = c *s (h x y)"
-  by (simp add: bilinear_def linear_def)
-
-lemma bilinear_rmul: "bilinear h ==> h x (c *s y) = c *s (h x y)"
-  by (simp add: bilinear_def linear_def)
-
-lemma bilinear_lneg: "bilinear h ==> h (- (x:: 'a::ring_1 ^ 'n)) y = -(h x y)"
-  by (simp only: vector_sneg_minus1 bilinear_lmul)
-
-lemma bilinear_rneg: "bilinear h ==> h x (- (y:: 'a::ring_1 ^ 'n)) = - h x y"
-  by (simp only: vector_sneg_minus1 bilinear_rmul)
-
-lemma  (in ab_group_add) eq_add_iff: "x = x + y \<longleftrightarrow> y = 0"
-  using add_imp_eq[of x y 0] by auto
-
-lemma bilinear_lzero:
-  fixes h :: "'a::ring^'n \<Rightarrow> _" assumes bh: "bilinear h" shows "h 0 x = 0"
-  using bilinear_ladd[OF bh, of 0 0 x]
-    by (simp add: eq_add_iff ring_simps)
-
-lemma bilinear_rzero:
-  fixes h :: "'a::ring^'n \<Rightarrow> _" assumes bh: "bilinear h" shows "h x 0 = 0"
-  using bilinear_radd[OF bh, of x 0 0 ]
-    by (simp add: eq_add_iff ring_simps)
-
-lemma bilinear_lsub: "bilinear h ==> h (x - (y:: 'a::ring_1 ^ 'n)) z = h x z - h y z"
-  by (simp  add: diff_def bilinear_ladd bilinear_lneg)
-
-lemma bilinear_rsub: "bilinear h ==> h z (x - (y:: 'a::ring_1 ^ 'n)) = h z x - h z y"
-  by (simp  add: diff_def bilinear_radd bilinear_rneg)
-
-lemma bilinear_setsum:
-  fixes h:: "'a ^'n \<Rightarrow> 'a::semiring_1^'m \<Rightarrow> 'a ^ 'k"
-  assumes bh: "bilinear h" and fS: "finite S" and fT: "finite T"
-  shows "h (setsum f S) (setsum g T) = setsum (\<lambda>(i,j). h (f i) (g j)) (S \<times> T) "
-proof-
-  have "h (setsum f S) (setsum g T) = setsum (\<lambda>x. h (f x) (setsum g T)) S"
-    apply (rule linear_setsum[unfolded o_def])
-    using bh fS by (auto simp add: bilinear_def)
-  also have "\<dots> = setsum (\<lambda>x. setsum (\<lambda>y. h (f x) (g y)) T) S"
-    apply (rule setsum_cong, simp)
-    apply (rule linear_setsum[unfolded o_def])
-    using bh fT by (auto simp add: bilinear_def)
-  finally show ?thesis unfolding setsum_cartesian_product .
-qed
-
-lemma bilinear_bounded:
-  fixes h:: "real ^'m::finite \<Rightarrow> real^'n::finite \<Rightarrow> real ^ 'k::finite"
-  assumes bh: "bilinear h"
-  shows "\<exists>B. \<forall>x y. norm (h x y) \<le> B * norm x * norm y"
-proof-
-  let ?M = "UNIV :: 'm set"
-  let ?N = "UNIV :: 'n set"
-  let ?B = "setsum (\<lambda>(i,j). norm (h (basis i) (basis j))) (?M \<times> ?N)"
-  have fM: "finite ?M" and fN: "finite ?N" by simp_all
-  {fix x:: "real ^ 'm" and  y :: "real^'n"
-    have "norm (h x y) = norm (h (setsum (\<lambda>i. (x$i) *s basis i) ?M) (setsum (\<lambda>i. (y$i) *s basis i) ?N))" unfolding basis_expansion ..
-    also have "\<dots> = norm (setsum (\<lambda> (i,j). h ((x$i) *s basis i) ((y$j) *s basis j)) (?M \<times> ?N))"  unfolding bilinear_setsum[OF bh fM fN] ..
-    finally have th: "norm (h x y) = \<dots>" .
-    have "norm (h x y) \<le> ?B * norm x * norm y"
-      apply (simp add: setsum_left_distrib th)
-      apply (rule real_setsum_norm_le)
-      using fN fM
-      apply simp
-      apply (auto simp add: bilinear_rmul[OF bh] bilinear_lmul[OF bh] norm_mul ring_simps)
-      apply (rule mult_mono)
-      apply (auto simp add: norm_ge_zero zero_le_mult_iff component_le_norm)
-      apply (rule mult_mono)
-      apply (auto simp add: norm_ge_zero zero_le_mult_iff component_le_norm)
-      done}
-  then show ?thesis by metis
-qed
-
-lemma bilinear_bounded_pos:
-  fixes h:: "real ^'m::finite \<Rightarrow> real^'n::finite \<Rightarrow> real ^ 'k::finite"
-  assumes bh: "bilinear h"
-  shows "\<exists>B > 0. \<forall>x y. norm (h x y) \<le> B * norm x * norm y"
-proof-
-  from bilinear_bounded[OF bh] obtain B where
-    B: "\<forall>x y. norm (h x y) \<le> B * norm x * norm y" by blast
-  let ?K = "\<bar>B\<bar> + 1"
-  have Kp: "?K > 0" by arith
-  have KB: "B < ?K" by arith
-  {fix x::"real ^'m" and y :: "real ^'n"
-    from KB Kp
-    have "B * norm x * norm y \<le> ?K * norm x * norm y"
-      apply -
-      apply (rule mult_right_mono, rule mult_right_mono)
-      by (auto simp add: norm_ge_zero)
-    then have "norm (h x y) \<le> ?K * norm x * norm y"
-      using B[rule_format, of x y] by simp}
-  with Kp show ?thesis by blast
-qed
-
-lemma bilinear_conv_bounded_bilinear:
-  fixes h :: "real ^ _ \<Rightarrow> real ^ _ \<Rightarrow> real ^ _"
-  shows "bilinear h \<longleftrightarrow> bounded_bilinear h"
-proof
-  assume "bilinear h"
-  show "bounded_bilinear h"
-  proof
-    fix x y z show "h (x + y) z = h x z + h y z"
-      using `bilinear h` unfolding bilinear_def linear_def by simp
-  next
-    fix x y z show "h x (y + z) = h x y + h x z"
-      using `bilinear h` unfolding bilinear_def linear_def by simp
-  next
-    fix r x y show "h (scaleR r x) y = scaleR r (h x y)"
-      using `bilinear h` unfolding bilinear_def linear_def
-      by (simp add: smult_conv_scaleR)
-  next
-    fix r x y show "h x (scaleR r y) = scaleR r (h x y)"
-      using `bilinear h` unfolding bilinear_def linear_def
-      by (simp add: smult_conv_scaleR)
-  next
-    have "\<exists>B. \<forall>x y. norm (h x y) \<le> B * norm x * norm y"
-      using `bilinear h` by (rule bilinear_bounded)
-    thus "\<exists>K. \<forall>x y. norm (h x y) \<le> norm x * norm y * K"
-      by (simp add: mult_ac)
-  qed
-next
-  assume "bounded_bilinear h"
-  then interpret h: bounded_bilinear h .
-  show "bilinear h"
-    unfolding bilinear_def linear_conv_bounded_linear
-    using h.bounded_linear_left h.bounded_linear_right
-    by simp
-qed
-
-subsection{* Adjoints. *}
-
-definition "adjoint f = (SOME f'. \<forall>x y. f x \<bullet> y = x \<bullet> f' y)"
-
-lemma choice_iff: "(\<forall>x. \<exists>y. P x y) \<longleftrightarrow> (\<exists>f. \<forall>x. P x (f x))" by metis
-
-lemma adjoint_works_lemma:
-  fixes f:: "'a::ring_1 ^'n::finite \<Rightarrow> 'a ^ 'm::finite"
-  assumes lf: "linear f"
-  shows "\<forall>x y. f x \<bullet> y = x \<bullet> adjoint f y"
-proof-
-  let ?N = "UNIV :: 'n set"
-  let ?M = "UNIV :: 'm set"
-  have fN: "finite ?N" by simp
-  have fM: "finite ?M" by simp
-  {fix y:: "'a ^ 'm"
-    let ?w = "(\<chi> i. (f (basis i) \<bullet> y)) :: 'a ^ 'n"
-    {fix x
-      have "f x \<bullet> y = f (setsum (\<lambda>i. (x$i) *s basis i) ?N) \<bullet> y"
-        by (simp only: basis_expansion)
-      also have "\<dots> = (setsum (\<lambda>i. (x$i) *s f (basis i)) ?N) \<bullet> y"
-        unfolding linear_setsum[OF lf fN]
-        by (simp add: linear_cmul[OF lf])
-      finally have "f x \<bullet> y = x \<bullet> ?w"
-        apply (simp only: )
-        apply (simp add: dot_def setsum_left_distrib setsum_right_distrib setsum_commute[of _ ?M ?N] ring_simps)
-        done}
-  }
-  then show ?thesis unfolding adjoint_def
-    some_eq_ex[of "\<lambda>f'. \<forall>x y. f x \<bullet> y = x \<bullet> f' y"]
-    using choice_iff[of "\<lambda>a b. \<forall>x. f x \<bullet> a = x \<bullet> b "]
-    by metis
-qed
-
-lemma adjoint_works:
-  fixes f:: "'a::ring_1 ^'n::finite \<Rightarrow> 'a ^ 'm::finite"
-  assumes lf: "linear f"
-  shows "x \<bullet> adjoint f y = f x \<bullet> y"
-  using adjoint_works_lemma[OF lf] by metis
-
-
-lemma adjoint_linear:
-  fixes f :: "'a::comm_ring_1 ^'n::finite \<Rightarrow> 'a ^ 'm::finite"
-  assumes lf: "linear f"
-  shows "linear (adjoint f)"
-  by (simp add: linear_def vector_eq_ldot[symmetric] dot_radd dot_rmult adjoint_works[OF lf])
-
-lemma adjoint_clauses:
-  fixes f:: "'a::comm_ring_1 ^'n::finite \<Rightarrow> 'a ^ 'm::finite"
-  assumes lf: "linear f"
-  shows "x \<bullet> adjoint f y = f x \<bullet> y"
-  and "adjoint f y \<bullet> x = y \<bullet> f x"
-  by (simp_all add: adjoint_works[OF lf] dot_sym )
-
-lemma adjoint_adjoint:
-  fixes f:: "'a::comm_ring_1 ^ 'n::finite \<Rightarrow> 'a ^ 'm::finite"
-  assumes lf: "linear f"
-  shows "adjoint (adjoint f) = f"
-  apply (rule ext)
-  by (simp add: vector_eq_ldot[symmetric] adjoint_clauses[OF adjoint_linear[OF lf]] adjoint_clauses[OF lf])
-
-lemma adjoint_unique:
-  fixes f:: "'a::comm_ring_1 ^ 'n::finite \<Rightarrow> 'a ^ 'm::finite"
-  assumes lf: "linear f" and u: "\<forall>x y. f' x \<bullet> y = x \<bullet> f y"
-  shows "f' = adjoint f"
-  apply (rule ext)
-  using u
-  by (simp add: vector_eq_rdot[symmetric] adjoint_clauses[OF lf])
-
-text{* Matrix notation. NB: an MxN matrix is of type @{typ "'a^'n^'m"}, not @{typ "'a^'m^'n"} *}
-
-consts generic_mult :: "'a \<Rightarrow> 'b \<Rightarrow> 'c" (infixr "\<star>" 75)
-
-defs (overloaded)
-matrix_matrix_mult_def: "(m:: ('a::semiring_1) ^'n^'m) \<star> (m' :: 'a ^'p^'n) \<equiv> (\<chi> i j. setsum (\<lambda>k. ((m$i)$k) * ((m'$k)$j)) (UNIV :: 'n set)) ::'a ^ 'p ^'m"
-
-abbreviation
-  matrix_matrix_mult' :: "('a::semiring_1) ^'n^'m \<Rightarrow> 'a ^'p^'n \<Rightarrow> 'a ^ 'p ^'m"  (infixl "**" 70)
-  where "m ** m' == m\<star> m'"
-
-defs (overloaded)
-  matrix_vector_mult_def: "(m::('a::semiring_1) ^'n^'m) \<star> (x::'a ^'n) \<equiv> (\<chi> i. setsum (\<lambda>j. ((m$i)$j) * (x$j)) (UNIV ::'n set)) :: 'a^'m"
-
-abbreviation
-  matrix_vector_mult' :: "('a::semiring_1) ^'n^'m \<Rightarrow> 'a ^'n \<Rightarrow> 'a ^ 'm"  (infixl "*v" 70)
-  where
-  "m *v v == m \<star> v"
-
-defs (overloaded)
-  vector_matrix_mult_def: "(x::'a^'m) \<star> (m::('a::semiring_1) ^'n^'m) \<equiv> (\<chi> j. setsum (\<lambda>i. ((m$i)$j) * (x$i)) (UNIV :: 'm set)) :: 'a^'n"
-
-abbreviation
-  vactor_matrix_mult' :: "'a ^ 'm \<Rightarrow> ('a::semiring_1) ^'n^'m \<Rightarrow> 'a ^'n "  (infixl "v*" 70)
-  where
-  "v v* m == v \<star> m"
-
-definition "(mat::'a::zero => 'a ^'n^'n) k = (\<chi> i j. if i = j then k else 0)"
-definition "(transp::'a^'n^'m \<Rightarrow> 'a^'m^'n) A = (\<chi> i j. ((A$j)$i))"
-definition "(row::'m => 'a ^'n^'m \<Rightarrow> 'a ^'n) i A = (\<chi> j. ((A$i)$j))"
-definition "(column::'n =>'a^'n^'m =>'a^'m) j A = (\<chi> i. ((A$i)$j))"
-definition "rows(A::'a^'n^'m) = { row i A | i. i \<in> (UNIV :: 'm set)}"
-definition "columns(A::'a^'n^'m) = { column i A | i. i \<in> (UNIV :: 'n set)}"
-
-lemma mat_0[simp]: "mat 0 = 0" by (vector mat_def)
-lemma matrix_add_ldistrib: "(A ** (B + C)) = (A \<star> B) + (A \<star> C)"
-  by (vector matrix_matrix_mult_def setsum_addf[symmetric] ring_simps)
-
-lemma matrix_mul_lid:
-  fixes A :: "'a::semiring_1 ^ 'm ^ 'n::finite"
-  shows "mat 1 ** A = A"
-  apply (simp add: matrix_matrix_mult_def mat_def)
-  apply vector
-  by (auto simp only: cond_value_iff cond_application_beta setsum_delta'[OF finite]  mult_1_left mult_zero_left if_True UNIV_I)
-
-
-lemma matrix_mul_rid:
-  fixes A :: "'a::semiring_1 ^ 'm::finite ^ 'n"
-  shows "A ** mat 1 = A"
-  apply (simp add: matrix_matrix_mult_def mat_def)
-  apply vector
-  by (auto simp only: cond_value_iff cond_application_beta setsum_delta[OF finite]  mult_1_right mult_zero_right if_True UNIV_I cong: if_cong)
-
-lemma matrix_mul_assoc: "A ** (B ** C) = (A ** B) ** C"
-  apply (vector matrix_matrix_mult_def setsum_right_distrib setsum_left_distrib mult_assoc)
-  apply (subst setsum_commute)
-  apply simp
-  done
-
-lemma matrix_vector_mul_assoc: "A *v (B *v x) = (A ** B) *v x"
-  apply (vector matrix_matrix_mult_def matrix_vector_mult_def setsum_right_distrib setsum_left_distrib mult_assoc)
-  apply (subst setsum_commute)
-  apply simp
-  done
-
-lemma matrix_vector_mul_lid: "mat 1 *v x = (x::'a::semiring_1 ^ 'n::finite)"
-  apply (vector matrix_vector_mult_def mat_def)
-  by (simp add: cond_value_iff cond_application_beta
-    setsum_delta' cong del: if_weak_cong)
-
-lemma matrix_transp_mul: "transp(A ** B) = transp B ** transp (A::'a::comm_semiring_1^'m^'n)"
-  by (simp add: matrix_matrix_mult_def transp_def Cart_eq mult_commute)
-
-lemma matrix_eq:
-  fixes A B :: "'a::semiring_1 ^ 'n::finite ^ 'm"
-  shows "A = B \<longleftrightarrow>  (\<forall>x. A *v x = B *v x)" (is "?lhs \<longleftrightarrow> ?rhs")
-  apply auto
-  apply (subst Cart_eq)
-  apply clarify
-  apply (clarsimp simp add: matrix_vector_mult_def basis_def cond_value_iff cond_application_beta Cart_eq cong del: if_weak_cong)
-  apply (erule_tac x="basis ia" in allE)
-  apply (erule_tac x="i" in allE)
-  by (auto simp add: basis_def cond_value_iff cond_application_beta setsum_delta[OF finite] cong del: if_weak_cong)
-
-lemma matrix_vector_mul_component:
-  shows "((A::'a::semiring_1^'n'^'m) *v x)$k = (A$k) \<bullet> x"
-  by (simp add: matrix_vector_mult_def dot_def)
-
-lemma dot_lmul_matrix: "((x::'a::comm_semiring_1 ^'n) v* A) \<bullet> y = x \<bullet> (A *v y)"
-  apply (simp add: dot_def matrix_vector_mult_def vector_matrix_mult_def setsum_left_distrib setsum_right_distrib mult_ac)
-  apply (subst setsum_commute)
-  by simp
-
-lemma transp_mat: "transp (mat n) = mat n"
-  by (vector transp_def mat_def)
-
-lemma transp_transp: "transp(transp A) = A"
-  by (vector transp_def)
-
-lemma row_transp:
-  fixes A:: "'a::semiring_1^'n^'m"
-  shows "row i (transp A) = column i A"
-  by (simp add: row_def column_def transp_def Cart_eq)
-
-lemma column_transp:
-  fixes A:: "'a::semiring_1^'n^'m"
-  shows "column i (transp A) = row i A"
-  by (simp add: row_def column_def transp_def Cart_eq)
-
-lemma rows_transp: "rows(transp (A::'a::semiring_1^'n^'m)) = columns A"
-by (auto simp add: rows_def columns_def row_transp intro: set_ext)
-
-lemma columns_transp: "columns(transp (A::'a::semiring_1^'n^'m)) = rows A" by (metis transp_transp rows_transp)
-
-text{* Two sometimes fruitful ways of looking at matrix-vector multiplication. *}
-
-lemma matrix_mult_dot: "A *v x = (\<chi> i. A$i \<bullet> x)"
-  by (simp add: matrix_vector_mult_def dot_def)
-
-lemma matrix_mult_vsum: "(A::'a::comm_semiring_1^'n^'m) *v x = setsum (\<lambda>i. (x$i) *s column i A) (UNIV:: 'n set)"
-  by (simp add: matrix_vector_mult_def Cart_eq column_def mult_commute)
-
-lemma vector_componentwise:
-  "(x::'a::ring_1^'n::finite) = (\<chi> j. setsum (\<lambda>i. (x$i) * (basis i :: 'a^'n)$j) (UNIV :: 'n set))"
-  apply (subst basis_expansion[symmetric])
-  by (vector Cart_eq setsum_component)
-
-lemma linear_componentwise:
-  fixes f:: "'a::ring_1 ^ 'm::finite \<Rightarrow> 'a ^ 'n"
-  assumes lf: "linear f"
-  shows "(f x)$j = setsum (\<lambda>i. (x$i) * (f (basis i)$j)) (UNIV :: 'm set)" (is "?lhs = ?rhs")
-proof-
-  let ?M = "(UNIV :: 'm set)"
-  let ?N = "(UNIV :: 'n set)"
-  have fM: "finite ?M" by simp
-  have "?rhs = (setsum (\<lambda>i.(x$i) *s f (basis i) ) ?M)$j"
-    unfolding vector_smult_component[symmetric]
-    unfolding setsum_component[of "(\<lambda>i.(x$i) *s f (basis i :: 'a^'m))" ?M]
-    ..
-  then show ?thesis unfolding linear_setsum_mul[OF lf fM, symmetric] basis_expansion ..
-qed
-
-text{* Inverse matrices  (not necessarily square) *}
-
-definition "invertible(A::'a::semiring_1^'n^'m) \<longleftrightarrow> (\<exists>A'::'a^'m^'n. A ** A' = mat 1 \<and> A' ** A = mat 1)"
-
-definition "matrix_inv(A:: 'a::semiring_1^'n^'m) =
-        (SOME A'::'a^'m^'n. A ** A' = mat 1 \<and> A' ** A = mat 1)"
-
-text{* Correspondence between matrices and linear operators. *}
-
-definition matrix:: "('a::{plus,times, one, zero}^'m \<Rightarrow> 'a ^ 'n) \<Rightarrow> 'a^'m^'n"
-where "matrix f = (\<chi> i j. (f(basis j))$i)"
-
-lemma matrix_vector_mul_linear: "linear(\<lambda>x. A *v (x::'a::comm_semiring_1 ^ 'n))"
-  by (simp add: linear_def matrix_vector_mult_def Cart_eq ring_simps setsum_right_distrib setsum_addf)
-
-lemma matrix_works: assumes lf: "linear f" shows "matrix f *v x = f (x::'a::comm_ring_1 ^ 'n::finite)"
-apply (simp add: matrix_def matrix_vector_mult_def Cart_eq mult_commute)
-apply clarify
-apply (rule linear_componentwise[OF lf, symmetric])
-done
-
-lemma matrix_vector_mul: "linear f ==> f = (\<lambda>x. matrix f *v (x::'a::comm_ring_1 ^ 'n::finite))" by (simp add: ext matrix_works)
-
-lemma matrix_of_matrix_vector_mul: "matrix(\<lambda>x. A *v (x :: 'a:: comm_ring_1 ^ 'n::finite)) = A"
-  by (simp add: matrix_eq matrix_vector_mul_linear matrix_works)
-
-lemma matrix_compose:
-  assumes lf: "linear (f::'a::comm_ring_1^'n::finite \<Rightarrow> 'a^'m::finite)"
-  and lg: "linear (g::'a::comm_ring_1^'m::finite \<Rightarrow> 'a^'k)"
-  shows "matrix (g o f) = matrix g ** matrix f"
-  using lf lg linear_compose[OF lf lg] matrix_works[OF linear_compose[OF lf lg]]
-  by (simp  add: matrix_eq matrix_works matrix_vector_mul_assoc[symmetric] o_def)
-
-lemma matrix_vector_column:"(A::'a::comm_semiring_1^'n^'m) *v x = setsum (\<lambda>i. (x$i) *s ((transp A)$i)) (UNIV:: 'n set)"
-  by (simp add: matrix_vector_mult_def transp_def Cart_eq mult_commute)
-
-lemma adjoint_matrix: "adjoint(\<lambda>x. (A::'a::comm_ring_1^'n::finite^'m::finite) *v x) = (\<lambda>x. transp A *v x)"
-  apply (rule adjoint_unique[symmetric])
-  apply (rule matrix_vector_mul_linear)
-  apply (simp add: transp_def dot_def matrix_vector_mult_def setsum_left_distrib setsum_right_distrib)
-  apply (subst setsum_commute)
-  apply (auto simp add: mult_ac)
-  done
-
-lemma matrix_adjoint: assumes lf: "linear (f :: 'a::comm_ring_1^'n::finite \<Rightarrow> 'a ^ 'm::finite)"
-  shows "matrix(adjoint f) = transp(matrix f)"
-  apply (subst matrix_vector_mul[OF lf])
-  unfolding adjoint_matrix matrix_of_matrix_vector_mul ..
-
-subsection{* Interlude: Some properties of real sets *}
-
-lemma seq_mono_lemma: assumes "\<forall>(n::nat) \<ge> m. (d n :: real) < e n" and "\<forall>n \<ge> m. e n <= e m"
-  shows "\<forall>n \<ge> m. d n < e m"
-  using prems apply auto
-  apply (erule_tac x="n" in allE)
-  apply (erule_tac x="n" in allE)
-  apply auto
-  done
-
-
-lemma real_convex_bound_lt:
-  assumes xa: "(x::real) < a" and ya: "y < a" and u: "0 <= u" and v: "0 <= v"
-  and uv: "u + v = 1"
-  shows "u * x + v * y < a"
-proof-
-  have uv': "u = 0 \<longrightarrow> v \<noteq> 0" using u v uv by arith
-  have "a = a * (u + v)" unfolding uv  by simp
-  hence th: "u * a + v * a = a" by (simp add: ring_simps)
-  from xa u have "u \<noteq> 0 \<Longrightarrow> u*x < u*a" by (simp add: mult_compare_simps)
-  from ya v have "v \<noteq> 0 \<Longrightarrow> v * y < v * a" by (simp add: mult_compare_simps)
-  from xa ya u v have "u * x + v * y < u * a + v * a"
-    apply (cases "u = 0", simp_all add: uv')
-    apply(rule mult_strict_left_mono)
-    using uv' apply simp_all
-
-    apply (rule add_less_le_mono)
-    apply(rule mult_strict_left_mono)
-    apply simp_all
-    apply (rule mult_left_mono)
-    apply simp_all
-    done
-  thus ?thesis unfolding th .
-qed
-
-lemma real_convex_bound_le:
-  assumes xa: "(x::real) \<le> a" and ya: "y \<le> a" and u: "0 <= u" and v: "0 <= v"
-  and uv: "u + v = 1"
-  shows "u * x + v * y \<le> a"
-proof-
-  from xa ya u v have "u * x + v * y \<le> u * a + v * a" by (simp add: add_mono mult_left_mono)
-  also have "\<dots> \<le> (u + v) * a" by (simp add: ring_simps)
-  finally show ?thesis unfolding uv by simp
-qed
-
-lemma infinite_enumerate: assumes fS: "infinite S"
-  shows "\<exists>r. subseq r \<and> (\<forall>n. r n \<in> S)"
-unfolding subseq_def
-using enumerate_in_set[OF fS] enumerate_mono[of _ _ S] fS by auto
-
-lemma approachable_lt_le: "(\<exists>(d::real)>0. \<forall>x. f x < d \<longrightarrow> P x) \<longleftrightarrow> (\<exists>d>0. \<forall>x. f x \<le> d \<longrightarrow> P x)"
-apply auto
-apply (rule_tac x="d/2" in exI)
-apply auto
-done
-
-
-lemma triangle_lemma:
-  assumes x: "0 <= (x::real)" and y:"0 <= y" and z: "0 <= z" and xy: "x^2 <= y^2 + z^2"
-  shows "x <= y + z"
-proof-
-  have "y^2 + z^2 \<le> y^2 + 2*y*z + z^2" using z y  by (simp add: zero_compare_simps)
-  with xy have th: "x ^2 \<le> (y+z)^2" by (simp add: power2_eq_square ring_simps)
-  from y z have yz: "y + z \<ge> 0" by arith
-  from power2_le_imp_le[OF th yz] show ?thesis .
-qed
-
-
-lemma lambda_skolem: "(\<forall>i. \<exists>x. P i x) \<longleftrightarrow>
-   (\<exists>x::'a ^ 'n. \<forall>i. P i (x$i))" (is "?lhs \<longleftrightarrow> ?rhs")
-proof-
-  let ?S = "(UNIV :: 'n set)"
-  {assume H: "?rhs"
-    then have ?lhs by auto}
-  moreover
-  {assume H: "?lhs"
-    then obtain f where f:"\<forall>i. P i (f i)" unfolding choice_iff by metis
-    let ?x = "(\<chi> i. (f i)) :: 'a ^ 'n"
-    {fix i
-      from f have "P i (f i)" by metis
-      then have "P i (?x$i)" by auto
-    }
-    hence "\<forall>i. P i (?x$i)" by metis
-    hence ?rhs by metis }
-  ultimately show ?thesis by metis
-qed
-
-(* Supremum and infimum of real sets *)
-
-
-definition rsup:: "real set \<Rightarrow> real" where
-  "rsup S = (SOME a. isLub UNIV S a)"
-
-lemma rsup_alt: "rsup S = (SOME a. (\<forall>x \<in> S. x \<le> a) \<and> (\<forall>b. (\<forall>x \<in> S. x \<le> b) \<longrightarrow> a \<le> b))"  by (auto simp  add: isLub_def rsup_def leastP_def isUb_def setle_def setge_def)
-
-lemma rsup: assumes Se: "S \<noteq> {}" and b: "\<exists>b. S *<= b"
-  shows "isLub UNIV S (rsup S)"
-using Se b
-unfolding rsup_def
-apply clarify
-apply (rule someI_ex)
-apply (rule reals_complete)
-by (auto simp add: isUb_def setle_def)
-
-lemma rsup_le: assumes Se: "S \<noteq> {}" and Sb: "S *<= b" shows "rsup S \<le> b"
-proof-
-  from Sb have bu: "isUb UNIV S b" by (simp add: isUb_def setle_def)
-  from rsup[OF Se] Sb have "isLub UNIV S (rsup S)"  by blast
-  then show ?thesis using bu by (auto simp add: isLub_def leastP_def setle_def setge_def)
-qed
-
-lemma rsup_finite_Max: assumes fS: "finite S" and Se: "S \<noteq> {}"
-  shows "rsup S = Max S"
-using fS Se
-proof-
-  let ?m = "Max S"
-  from Max_ge[OF fS] have Sm: "\<forall> x\<in> S. x \<le> ?m" by metis
-  with rsup[OF Se] have lub: "isLub UNIV S (rsup S)" by (metis setle_def)
-  from Max_in[OF fS Se] lub have mrS: "?m \<le> rsup S"
-    by (auto simp add: isLub_def leastP_def setle_def setge_def isUb_def)
-  moreover
-  have "rsup S \<le> ?m" using Sm lub
-    by (auto simp add: isLub_def leastP_def isUb_def setle_def setge_def)
-  ultimately  show ?thesis by arith
-qed
-
-lemma rsup_finite_in: assumes fS: "finite S" and Se: "S \<noteq> {}"
-  shows "rsup S \<in> S"
-  using rsup_finite_Max[OF fS Se] Max_in[OF fS Se] by metis
-
-lemma rsup_finite_Ub: assumes fS: "finite S" and Se: "S \<noteq> {}"
-  shows "isUb S S (rsup S)"
-  using rsup_finite_Max[OF fS Se] rsup_finite_in[OF fS Se] Max_ge[OF fS]
-  unfolding isUb_def setle_def by metis
-
-lemma rsup_finite_ge_iff: assumes fS: "finite S" and Se: "S \<noteq> {}"
-  shows "a \<le> rsup S \<longleftrightarrow> (\<exists> x \<in> S. a \<le> x)"
-using rsup_finite_Ub[OF fS Se] by (auto simp add: isUb_def setle_def)
-
-lemma rsup_finite_le_iff: assumes fS: "finite S" and Se: "S \<noteq> {}"
-  shows "a \<ge> rsup S \<longleftrightarrow> (\<forall> x \<in> S. a \<ge> x)"
-using rsup_finite_Ub[OF fS Se] by (auto simp add: isUb_def setle_def)
-
-lemma rsup_finite_gt_iff: assumes fS: "finite S" and Se: "S \<noteq> {}"
-  shows "a < rsup S \<longleftrightarrow> (\<exists> x \<in> S. a < x)"
-using rsup_finite_Ub[OF fS Se] by (auto simp add: isUb_def setle_def)
-
-lemma rsup_finite_lt_iff: assumes fS: "finite S" and Se: "S \<noteq> {}"
-  shows "a > rsup S \<longleftrightarrow> (\<forall> x \<in> S. a > x)"
-using rsup_finite_Ub[OF fS Se] by (auto simp add: isUb_def setle_def)
-
-lemma rsup_unique: assumes b: "S *<= b" and S: "\<forall>b' < b. \<exists>x \<in> S. b' < x"
-  shows "rsup S = b"
-using b S
-unfolding setle_def rsup_alt
-apply -
-apply (rule some_equality)
-apply (metis  linorder_not_le order_eq_iff[symmetric])+
-done
-
-lemma rsup_le_subset: "S\<noteq>{} \<Longrightarrow> S \<subseteq> T \<Longrightarrow> (\<exists>b. T *<= b) \<Longrightarrow> rsup S \<le> rsup T"
-  apply (rule rsup_le)
-  apply simp
-  using rsup[of T] by (auto simp add: isLub_def leastP_def setge_def setle_def isUb_def)
-
-lemma isUb_def': "isUb R S = (\<lambda>x. S *<= x \<and> x \<in> R)"
-  apply (rule ext)
-  by (metis isUb_def)
-
-lemma UNIV_trivial: "UNIV x" using UNIV_I[of x] by (metis mem_def)
-lemma rsup_bounds: assumes Se: "S \<noteq> {}" and l: "a <=* S" and u: "S *<= b"
-  shows "a \<le> rsup S \<and> rsup S \<le> b"
-proof-
-  from rsup[OF Se] u have lub: "isLub UNIV S (rsup S)" by blast
-  hence b: "rsup S \<le> b" using u by (auto simp add: isLub_def leastP_def setle_def setge_def isUb_def')
-  from Se obtain y where y: "y \<in> S" by blast
-  from lub l have "a \<le> rsup S" apply (auto simp add: isLub_def leastP_def setle_def setge_def isUb_def')
-    apply (erule ballE[where x=y])
-    apply (erule ballE[where x=y])
-    apply arith
-    using y apply auto
-    done
-  with b show ?thesis by blast
-qed
-
-lemma rsup_abs_le: "S \<noteq> {} \<Longrightarrow> (\<forall>x\<in>S. \<bar>x\<bar> \<le> a) \<Longrightarrow> \<bar>rsup S\<bar> \<le> a"
-  unfolding abs_le_interval_iff  using rsup_bounds[of S "-a" a]
-  by (auto simp add: setge_def setle_def)
-
-lemma rsup_asclose: assumes S:"S \<noteq> {}" and b: "\<forall>x\<in>S. \<bar>x - l\<bar> \<le> e" shows "\<bar>rsup S - l\<bar> \<le> e"
-proof-
-  have th: "\<And>(x::real) l e. \<bar>x - l\<bar> \<le> e \<longleftrightarrow> l - e \<le> x \<and> x \<le> l + e" by arith
-  show ?thesis using S b rsup_bounds[of S "l - e" "l+e"] unfolding th
-    by  (auto simp add: setge_def setle_def)
-qed
-
-definition rinf:: "real set \<Rightarrow> real" where
-  "rinf S = (SOME a. isGlb UNIV S a)"
-
-lemma rinf_alt: "rinf S = (SOME a. (\<forall>x \<in> S. x \<ge> a) \<and> (\<forall>b. (\<forall>x \<in> S. x \<ge> b) \<longrightarrow> a \<ge> b))"  by (auto simp  add: isGlb_def rinf_def greatestP_def isLb_def setle_def setge_def)
-
-lemma reals_complete_Glb: assumes Se: "\<exists>x. x \<in> S" and lb: "\<exists> y. isLb UNIV S y"
-  shows "\<exists>(t::real). isGlb UNIV S t"
-proof-
-  let ?M = "uminus ` S"
-  from lb have th: "\<exists>y. isUb UNIV ?M y" apply (auto simp add: isUb_def isLb_def setle_def setge_def)
-    by (rule_tac x="-y" in exI, auto)
-  from Se have Me: "\<exists>x. x \<in> ?M" by blast
-  from reals_complete[OF Me th] obtain t where t: "isLub UNIV ?M t" by blast
-  have "isGlb UNIV S (- t)" using t
-    apply (auto simp add: isLub_def isGlb_def leastP_def greatestP_def setle_def setge_def isUb_def isLb_def)
-    apply (erule_tac x="-y" in allE)
-    apply auto
-    done
-  then show ?thesis by metis
-qed
-
-lemma rinf: assumes Se: "S \<noteq> {}" and b: "\<exists>b. b <=* S"
-  shows "isGlb UNIV S (rinf S)"
-using Se b
-unfolding rinf_def
-apply clarify
-apply (rule someI_ex)
-apply (rule reals_complete_Glb)
-apply (auto simp add: isLb_def setle_def setge_def)
-done
-
-lemma rinf_ge: assumes Se: "S \<noteq> {}" and Sb: "b <=* S" shows "rinf S \<ge> b"
-proof-
-  from Sb have bu: "isLb UNIV S b" by (simp add: isLb_def setge_def)
-  from rinf[OF Se] Sb have "isGlb UNIV S (rinf S)"  by blast
-  then show ?thesis using bu by (auto simp add: isGlb_def greatestP_def setle_def setge_def)
-qed
-
-lemma rinf_finite_Min: assumes fS: "finite S" and Se: "S \<noteq> {}"
-  shows "rinf S = Min S"
-using fS Se
-proof-
-  let ?m = "Min S"
-  from Min_le[OF fS] have Sm: "\<forall> x\<in> S. x \<ge> ?m" by metis
-  with rinf[OF Se] have glb: "isGlb UNIV S (rinf S)" by (metis setge_def)
-  from Min_in[OF fS Se] glb have mrS: "?m \<ge> rinf S"
-    by (auto simp add: isGlb_def greatestP_def setle_def setge_def isLb_def)
-  moreover
-  have "rinf S \<ge> ?m" using Sm glb
-    by (auto simp add: isGlb_def greatestP_def isLb_def setle_def setge_def)
-  ultimately  show ?thesis by arith
-qed
-
-lemma rinf_finite_in: assumes fS: "finite S" and Se: "S \<noteq> {}"
-  shows "rinf S \<in> S"
-  using rinf_finite_Min[OF fS Se] Min_in[OF fS Se] by metis
-
-lemma rinf_finite_Lb: assumes fS: "finite S" and Se: "S \<noteq> {}"
-  shows "isLb S S (rinf S)"
-  using rinf_finite_Min[OF fS Se] rinf_finite_in[OF fS Se] Min_le[OF fS]
-  unfolding isLb_def setge_def by metis
-
-lemma rinf_finite_ge_iff: assumes fS: "finite S" and Se: "S \<noteq> {}"
-  shows "a \<le> rinf S \<longleftrightarrow> (\<forall> x \<in> S. a \<le> x)"
-using rinf_finite_Lb[OF fS Se] by (auto simp add: isLb_def setge_def)
-
-lemma rinf_finite_le_iff: assumes fS: "finite S" and Se: "S \<noteq> {}"
-  shows "a \<ge> rinf S \<longleftrightarrow> (\<exists> x \<in> S. a \<ge> x)"
-using rinf_finite_Lb[OF fS Se] by (auto simp add: isLb_def setge_def)
-
-lemma rinf_finite_gt_iff: assumes fS: "finite S" and Se: "S \<noteq> {}"
-  shows "a < rinf S \<longleftrightarrow> (\<forall> x \<in> S. a < x)"
-using rinf_finite_Lb[OF fS Se] by (auto simp add: isLb_def setge_def)
-
-lemma rinf_finite_lt_iff: assumes fS: "finite S" and Se: "S \<noteq> {}"
-  shows "a > rinf S \<longleftrightarrow> (\<exists> x \<in> S. a > x)"
-using rinf_finite_Lb[OF fS Se] by (auto simp add: isLb_def setge_def)
-
-lemma rinf_unique: assumes b: "b <=* S" and S: "\<forall>b' > b. \<exists>x \<in> S. b' > x"
-  shows "rinf S = b"
-using b S
-unfolding setge_def rinf_alt
-apply -
-apply (rule some_equality)
-apply (metis  linorder_not_le order_eq_iff[symmetric])+
-done
-
-lemma rinf_ge_subset: "S\<noteq>{} \<Longrightarrow> S \<subseteq> T \<Longrightarrow> (\<exists>b. b <=* T) \<Longrightarrow> rinf S >= rinf T"
-  apply (rule rinf_ge)
-  apply simp
-  using rinf[of T] by (auto simp add: isGlb_def greatestP_def setge_def setle_def isLb_def)
-
-lemma isLb_def': "isLb R S = (\<lambda>x. x <=* S \<and> x \<in> R)"
-  apply (rule ext)
-  by (metis isLb_def)
-
-lemma rinf_bounds: assumes Se: "S \<noteq> {}" and l: "a <=* S" and u: "S *<= b"
-  shows "a \<le> rinf S \<and> rinf S \<le> b"
-proof-
-  from rinf[OF Se] l have lub: "isGlb UNIV S (rinf S)" by blast
-  hence b: "a \<le> rinf S" using l by (auto simp add: isGlb_def greatestP_def setle_def setge_def isLb_def')
-  from Se obtain y where y: "y \<in> S" by blast
-  from lub u have "b \<ge> rinf S" apply (auto simp add: isGlb_def greatestP_def setle_def setge_def isLb_def')
-    apply (erule ballE[where x=y])
-    apply (erule ballE[where x=y])
-    apply arith
-    using y apply auto
-    done
-  with b show ?thesis by blast
-qed
-
-lemma rinf_abs_ge: "S \<noteq> {} \<Longrightarrow> (\<forall>x\<in>S. \<bar>x\<bar> \<le> a) \<Longrightarrow> \<bar>rinf S\<bar> \<le> a"
-  unfolding abs_le_interval_iff  using rinf_bounds[of S "-a" a]
-  by (auto simp add: setge_def setle_def)
-
-lemma rinf_asclose: assumes S:"S \<noteq> {}" and b: "\<forall>x\<in>S. \<bar>x - l\<bar> \<le> e" shows "\<bar>rinf S - l\<bar> \<le> e"
-proof-
-  have th: "\<And>(x::real) l e. \<bar>x - l\<bar> \<le> e \<longleftrightarrow> l - e \<le> x \<and> x \<le> l + e" by arith
-  show ?thesis using S b rinf_bounds[of S "l - e" "l+e"] unfolding th
-    by  (auto simp add: setge_def setle_def)
-qed
-
-
-
-subsection{* Operator norm. *}
-
-definition "onorm f = rsup {norm (f x)| x. norm x = 1}"
-
-lemma norm_bound_generalize:
-  fixes f:: "real ^'n::finite \<Rightarrow> real^'m::finite"
-  assumes lf: "linear f"
-  shows "(\<forall>x. norm x = 1 \<longrightarrow> norm (f x) \<le> b) \<longleftrightarrow> (\<forall>x. norm (f x) \<le> b * norm x)" (is "?lhs \<longleftrightarrow> ?rhs")
-proof-
-  {assume H: ?rhs
-    {fix x :: "real^'n" assume x: "norm x = 1"
-      from H[rule_format, of x] x have "norm (f x) \<le> b" by simp}
-    then have ?lhs by blast }
-
-  moreover
-  {assume H: ?lhs
-    from H[rule_format, of "basis arbitrary"]
-    have bp: "b \<ge> 0" using norm_ge_zero[of "f (basis arbitrary)"]
-      by (auto simp add: norm_basis elim: order_trans [OF norm_ge_zero])
-    {fix x :: "real ^'n"
-      {assume "x = 0"
-        then have "norm (f x) \<le> b * norm x" by (simp add: linear_0[OF lf] bp)}
-      moreover
-      {assume x0: "x \<noteq> 0"
-        hence n0: "norm x \<noteq> 0" by (metis norm_eq_zero)
-        let ?c = "1/ norm x"
-        have "norm (?c*s x) = 1" using x0 by (simp add: n0 norm_mul)
-        with H have "norm (f(?c*s x)) \<le> b" by blast
-        hence "?c * norm (f x) \<le> b"
-          by (simp add: linear_cmul[OF lf] norm_mul)
-        hence "norm (f x) \<le> b * norm x"
-          using n0 norm_ge_zero[of x] by (auto simp add: field_simps)}
-      ultimately have "norm (f x) \<le> b * norm x" by blast}
-    then have ?rhs by blast}
-  ultimately show ?thesis by blast
-qed
-
-lemma onorm:
-  fixes f:: "real ^'n::finite \<Rightarrow> real ^'m::finite"
-  assumes lf: "linear f"
-  shows "norm (f x) <= onorm f * norm x"
-  and "\<forall>x. norm (f x) <= b * norm x \<Longrightarrow> onorm f <= b"
-proof-
-  {
-    let ?S = "{norm (f x) |x. norm x = 1}"
-    have Se: "?S \<noteq> {}" using  norm_basis by auto
-    from linear_bounded[OF lf] have b: "\<exists> b. ?S *<= b"
-      unfolding norm_bound_generalize[OF lf, symmetric] by (auto simp add: setle_def)
-    {from rsup[OF Se b, unfolded onorm_def[symmetric]]
-      show "norm (f x) <= onorm f * norm x"
-        apply -
-        apply (rule spec[where x = x])
-        unfolding norm_bound_generalize[OF lf, symmetric]
-        by (auto simp add: isLub_def isUb_def leastP_def setge_def setle_def)}
-    {
-      show "\<forall>x. norm (f x) <= b * norm x \<Longrightarrow> onorm f <= b"
-        using rsup[OF Se b, unfolded onorm_def[symmetric]]
-        unfolding norm_bound_generalize[OF lf, symmetric]
-        by (auto simp add: isLub_def isUb_def leastP_def setge_def setle_def)}
-  }
-qed
-
-lemma onorm_pos_le: assumes lf: "linear (f::real ^'n::finite \<Rightarrow> real ^'m::finite)" shows "0 <= onorm f"
-  using order_trans[OF norm_ge_zero onorm(1)[OF lf, of "basis arbitrary"], unfolded norm_basis] by simp
-
-lemma onorm_eq_0: assumes lf: "linear (f::real ^'n::finite \<Rightarrow> real ^'m::finite)"
-  shows "onorm f = 0 \<longleftrightarrow> (\<forall>x. f x = 0)"
-  using onorm[OF lf]
-  apply (auto simp add: onorm_pos_le)
-  apply atomize
-  apply (erule allE[where x="0::real"])
-  using onorm_pos_le[OF lf]
-  apply arith
-  done
-
-lemma onorm_const: "onorm(\<lambda>x::real^'n::finite. (y::real ^ 'm::finite)) = norm y"
-proof-
-  let ?f = "\<lambda>x::real^'n. (y::real ^ 'm)"
-  have th: "{norm (?f x)| x. norm x = 1} = {norm y}"
-    by(auto intro: vector_choose_size set_ext)
-  show ?thesis
-    unfolding onorm_def th
-    apply (rule rsup_unique) by (simp_all  add: setle_def)
-qed
-
-lemma onorm_pos_lt: assumes lf: "linear (f::real ^ 'n::finite \<Rightarrow> real ^'m::finite)"
-  shows "0 < onorm f \<longleftrightarrow> ~(\<forall>x. f x = 0)"
-  unfolding onorm_eq_0[OF lf, symmetric]
-  using onorm_pos_le[OF lf] by arith
-
-lemma onorm_compose:
-  assumes lf: "linear (f::real ^'n::finite \<Rightarrow> real ^'m::finite)"
-  and lg: "linear (g::real^'k::finite \<Rightarrow> real^'n::finite)"
-  shows "onorm (f o g) <= onorm f * onorm g"
-  apply (rule onorm(2)[OF linear_compose[OF lg lf], rule_format])
-  unfolding o_def
-  apply (subst mult_assoc)
-  apply (rule order_trans)
-  apply (rule onorm(1)[OF lf])
-  apply (rule mult_mono1)
-  apply (rule onorm(1)[OF lg])
-  apply (rule onorm_pos_le[OF lf])
-  done
-
-lemma onorm_neg_lemma: assumes lf: "linear (f::real ^'n::finite \<Rightarrow> real^'m::finite)"
-  shows "onorm (\<lambda>x. - f x) \<le> onorm f"
-  using onorm[OF linear_compose_neg[OF lf]] onorm[OF lf]
-  unfolding norm_minus_cancel by metis
-
-lemma onorm_neg: assumes lf: "linear (f::real ^'n::finite \<Rightarrow> real^'m::finite)"
-  shows "onorm (\<lambda>x. - f x) = onorm f"
-  using onorm_neg_lemma[OF lf] onorm_neg_lemma[OF linear_compose_neg[OF lf]]
-  by simp
-
-lemma onorm_triangle:
-  assumes lf: "linear (f::real ^'n::finite \<Rightarrow> real ^'m::finite)" and lg: "linear g"
-  shows "onorm (\<lambda>x. f x + g x) <= onorm f + onorm g"
-  apply(rule onorm(2)[OF linear_compose_add[OF lf lg], rule_format])
-  apply (rule order_trans)
-  apply (rule norm_triangle_ineq)
-  apply (simp add: distrib)
-  apply (rule add_mono)
-  apply (rule onorm(1)[OF lf])
-  apply (rule onorm(1)[OF lg])
-  done
-
-lemma onorm_triangle_le: "linear (f::real ^'n::finite \<Rightarrow> real ^'m::finite) \<Longrightarrow> linear g \<Longrightarrow> onorm(f) + onorm(g) <= e
-  \<Longrightarrow> onorm(\<lambda>x. f x + g x) <= e"
-  apply (rule order_trans)
-  apply (rule onorm_triangle)
-  apply assumption+
-  done
-
-lemma onorm_triangle_lt: "linear (f::real ^'n::finite \<Rightarrow> real ^'m::finite) \<Longrightarrow> linear g \<Longrightarrow> onorm(f) + onorm(g) < e
-  ==> onorm(\<lambda>x. f x + g x) < e"
-  apply (rule order_le_less_trans)
-  apply (rule onorm_triangle)
-  by assumption+
-
-(* "lift" from 'a to 'a^1 and "drop" from 'a^1 to 'a -- FIXME: potential use of transfer *)
-
-definition vec1:: "'a \<Rightarrow> 'a ^ 1" where "vec1 x = (\<chi> i. x)"
-
-definition dest_vec1:: "'a ^1 \<Rightarrow> 'a"
-  where "dest_vec1 x = (x$1)"
-
-lemma vec1_component[simp]: "(vec1 x)$1 = x"
-  by (simp add: vec1_def)
-
-lemma vec1_dest_vec1[simp]: "vec1(dest_vec1 x) = x" "dest_vec1(vec1 y) = y"
-  by (simp_all add: vec1_def dest_vec1_def Cart_eq forall_1)
-
-lemma forall_vec1: "(\<forall>x. P x) \<longleftrightarrow> (\<forall>x. P (vec1 x))" by (metis vec1_dest_vec1)
-
-lemma exists_vec1: "(\<exists>x. P x) \<longleftrightarrow> (\<exists>x. P(vec1 x))" by (metis vec1_dest_vec1)
-
-lemma forall_dest_vec1: "(\<forall>x. P x) \<longleftrightarrow> (\<forall>x. P(dest_vec1 x))"  by (metis vec1_dest_vec1)
-
-lemma exists_dest_vec1: "(\<exists>x. P x) \<longleftrightarrow> (\<exists>x. P(dest_vec1 x))"by (metis vec1_dest_vec1)
-
-lemma vec1_eq[simp]:  "vec1 x = vec1 y \<longleftrightarrow> x = y" by (metis vec1_dest_vec1)
-
-lemma dest_vec1_eq[simp]: "dest_vec1 x = dest_vec1 y \<longleftrightarrow> x = y" by (metis vec1_dest_vec1)
-
-lemma vec1_in_image_vec1: "vec1 x \<in> (vec1 ` S) \<longleftrightarrow> x \<in> S" by auto
-
-lemma vec1_vec: "vec1 x = vec x" by (vector vec1_def)
-
-lemma vec1_add: "vec1(x + y) = vec1 x + vec1 y" by (vector vec1_def)
-lemma vec1_sub: "vec1(x - y) = vec1 x - vec1 y" by (vector vec1_def)
-lemma vec1_cmul: "vec1(c* x) = c *s vec1 x " by (vector vec1_def)
-lemma vec1_neg: "vec1(- x) = - vec1 x " by (vector vec1_def)
-
-lemma vec1_setsum: assumes fS: "finite S"
-  shows "vec1(setsum f S) = setsum (vec1 o f) S"
-  apply (induct rule: finite_induct[OF fS])
-  apply (simp add: vec1_vec)
-  apply (auto simp add: vec1_add)
-  done
-
-lemma dest_vec1_lambda: "dest_vec1(\<chi> i. x i) = x 1"
-  by (simp add: dest_vec1_def)
-
-lemma dest_vec1_vec: "dest_vec1(vec x) = x"
-  by (simp add: vec1_vec[symmetric])
-
-lemma dest_vec1_add: "dest_vec1(x + y) = dest_vec1 x + dest_vec1 y"
- by (metis vec1_dest_vec1 vec1_add)
-
-lemma dest_vec1_sub: "dest_vec1(x - y) = dest_vec1 x - dest_vec1 y"
- by (metis vec1_dest_vec1 vec1_sub)
-
-lemma dest_vec1_cmul: "dest_vec1(c*sx) = c * dest_vec1 x"
- by (metis vec1_dest_vec1 vec1_cmul)
-
-lemma dest_vec1_neg: "dest_vec1(- x) = - dest_vec1 x"
- by (metis vec1_dest_vec1 vec1_neg)
-
-lemma dest_vec1_0[simp]: "dest_vec1 0 = 0" by (metis vec_0 dest_vec1_vec)
-
-lemma dest_vec1_sum: assumes fS: "finite S"
-  shows "dest_vec1(setsum f S) = setsum (dest_vec1 o f) S"
-  apply (induct rule: finite_induct[OF fS])
-  apply (simp add: dest_vec1_vec)
-  apply (auto simp add: dest_vec1_add)
-  done
-
-lemma norm_vec1: "norm(vec1 x) = abs(x)"
-  by (simp add: vec1_def norm_real)
-
-lemma dist_vec1: "dist(vec1 x) (vec1 y) = abs(x - y)"
-  by (simp only: dist_real vec1_component)
-lemma abs_dest_vec1: "norm x = \<bar>dest_vec1 x\<bar>"
-  by (metis vec1_dest_vec1 norm_vec1)
-
-lemma linear_vmul_dest_vec1:
-  fixes f:: "'a::semiring_1^'n \<Rightarrow> 'a^1"
-  shows "linear f \<Longrightarrow> linear (\<lambda>x. dest_vec1(f x) *s v)"
-  unfolding dest_vec1_def
-  apply (rule linear_vmul_component)
-  by auto
-
-lemma linear_from_scalars:
-  assumes lf: "linear (f::'a::comm_ring_1 ^1 \<Rightarrow> 'a^'n)"
-  shows "f = (\<lambda>x. dest_vec1 x *s column 1 (matrix f))"
-  apply (rule ext)
-  apply (subst matrix_works[OF lf, symmetric])
-  apply (auto simp add: Cart_eq matrix_vector_mult_def dest_vec1_def column_def  mult_commute UNIV_1)
-  done
-
-lemma linear_to_scalars: assumes lf: "linear (f::'a::comm_ring_1 ^'n::finite \<Rightarrow> 'a^1)"
-  shows "f = (\<lambda>x. vec1(row 1 (matrix f) \<bullet> x))"
-  apply (rule ext)
-  apply (subst matrix_works[OF lf, symmetric])
-  apply (simp add: Cart_eq matrix_vector_mult_def vec1_def row_def dot_def mult_commute forall_1)
-  done
-
-lemma dest_vec1_eq_0: "dest_vec1 x = 0 \<longleftrightarrow> x = 0"
-  by (simp add: dest_vec1_eq[symmetric])
-
-lemma setsum_scalars: assumes fS: "finite S"
-  shows "setsum f S = vec1 (setsum (dest_vec1 o f) S)"
-  unfolding vec1_setsum[OF fS] by simp
-
-lemma dest_vec1_wlog_le: "(\<And>(x::'a::linorder ^ 1) y. P x y \<longleftrightarrow> P y x)  \<Longrightarrow> (\<And>x y. dest_vec1 x <= dest_vec1 y ==> P x y) \<Longrightarrow> P x y"
-  apply (cases "dest_vec1 x \<le> dest_vec1 y")
-  apply simp
-  apply (subgoal_tac "dest_vec1 y \<le> dest_vec1 x")
-  apply (auto)
-  done
-
-text{* Pasting vectors. *}
-
-lemma linear_fstcart: "linear fstcart"
-  by (auto simp add: linear_def Cart_eq)
-
-lemma linear_sndcart: "linear sndcart"
-  by (auto simp add: linear_def Cart_eq)
-
-lemma fstcart_vec[simp]: "fstcart(vec x) = vec x"
-  by (simp add: Cart_eq)
-
-lemma fstcart_add[simp]:"fstcart(x + y) = fstcart (x::'a::{plus,times}^('b + 'c)) + fstcart y"
-  by (simp add: Cart_eq)
-
-lemma fstcart_cmul[simp]:"fstcart(c*s x) = c*s fstcart (x::'a::{plus,times}^('b + 'c))"
-  by (simp add: Cart_eq)
-
-lemma fstcart_neg[simp]:"fstcart(- x) = - fstcart (x::'a::ring_1^('b + 'c))"
-  by (simp add: Cart_eq)
-
-lemma fstcart_sub[simp]:"fstcart(x - y) = fstcart (x::'a::ring_1^('b + 'c)) - fstcart y"
-  by (simp add: Cart_eq)
-
-lemma fstcart_setsum:
-  fixes f:: "'d \<Rightarrow> 'a::semiring_1^_"
-  assumes fS: "finite S"
-  shows "fstcart (setsum f S) = setsum (\<lambda>i. fstcart (f i)) S"
-  by (induct rule: finite_induct[OF fS], simp_all add: vec_0[symmetric] del: vec_0)
-
-lemma sndcart_vec[simp]: "sndcart(vec x) = vec x"
-  by (simp add: Cart_eq)
-
-lemma sndcart_add[simp]:"sndcart(x + y) = sndcart (x::'a::{plus,times}^('b + 'c)) + sndcart y"
-  by (simp add: Cart_eq)
-
-lemma sndcart_cmul[simp]:"sndcart(c*s x) = c*s sndcart (x::'a::{plus,times}^('b + 'c))"
-  by (simp add: Cart_eq)
-
-lemma sndcart_neg[simp]:"sndcart(- x) = - sndcart (x::'a::ring_1^('b + 'c))"
-  by (simp add: Cart_eq)
-
-lemma sndcart_sub[simp]:"sndcart(x - y) = sndcart (x::'a::ring_1^('b + 'c)) - sndcart y"
-  by (simp add: Cart_eq)
-
-lemma sndcart_setsum:
-  fixes f:: "'d \<Rightarrow> 'a::semiring_1^_"
-  assumes fS: "finite S"
-  shows "sndcart (setsum f S) = setsum (\<lambda>i. sndcart (f i)) S"
-  by (induct rule: finite_induct[OF fS], simp_all add: vec_0[symmetric] del: vec_0)
-
-lemma pastecart_vec[simp]: "pastecart (vec x) (vec x) = vec x"
-  by (simp add: pastecart_eq fstcart_pastecart sndcart_pastecart)
-
-lemma pastecart_add[simp]:"pastecart (x1::'a::{plus,times}^_) y1 + pastecart x2 y2 = pastecart (x1 + x2) (y1 + y2)"
-  by (simp add: pastecart_eq fstcart_pastecart sndcart_pastecart)
-
-lemma pastecart_cmul[simp]: "pastecart (c *s (x1::'a::{plus,times}^_)) (c *s y1) = c *s pastecart x1 y1"
-  by (simp add: pastecart_eq fstcart_pastecart sndcart_pastecart)
-
-lemma pastecart_neg[simp]: "pastecart (- (x::'a::ring_1^_)) (- y) = - pastecart x y"
-  unfolding vector_sneg_minus1 pastecart_cmul ..
-
-lemma pastecart_sub: "pastecart (x1::'a::ring_1^_) y1 - pastecart x2 y2 = pastecart (x1 - x2) (y1 - y2)"
-  by (simp add: diff_def pastecart_neg[symmetric] del: pastecart_neg)
-
-lemma pastecart_setsum:
-  fixes f:: "'d \<Rightarrow> 'a::semiring_1^_"
-  assumes fS: "finite S"
-  shows "pastecart (setsum f S) (setsum g S) = setsum (\<lambda>i. pastecart (f i) (g i)) S"
-  by (simp  add: pastecart_eq fstcart_setsum[OF fS] sndcart_setsum[OF fS] fstcart_pastecart sndcart_pastecart)
-
-lemma setsum_Plus:
-  "\<lbrakk>finite A; finite B\<rbrakk> \<Longrightarrow>
-    (\<Sum>x\<in>A <+> B. g x) = (\<Sum>x\<in>A. g (Inl x)) + (\<Sum>x\<in>B. g (Inr x))"
-  unfolding Plus_def
-  by (subst setsum_Un_disjoint, auto simp add: setsum_reindex)
-
-lemma setsum_UNIV_sum:
-  fixes g :: "'a::finite + 'b::finite \<Rightarrow> _"
-  shows "(\<Sum>x\<in>UNIV. g x) = (\<Sum>x\<in>UNIV. g (Inl x)) + (\<Sum>x\<in>UNIV. g (Inr x))"
-  apply (subst UNIV_Plus_UNIV [symmetric])
-  apply (rule setsum_Plus [OF finite finite])
-  done
-
-lemma norm_fstcart: "norm(fstcart x) <= norm (x::real ^('n::finite + 'm::finite))"
-proof-
-  have th0: "norm x = norm (pastecart (fstcart x) (sndcart x))"
-    by (simp add: pastecart_fst_snd)
-  have th1: "fstcart x \<bullet> fstcart x \<le> pastecart (fstcart x) (sndcart x) \<bullet> pastecart (fstcart x) (sndcart x)"
-    by (simp add: dot_def setsum_UNIV_sum pastecart_def setsum_nonneg)
-  then show ?thesis
-    unfolding th0
-    unfolding real_vector_norm_def real_sqrt_le_iff id_def
-    by (simp add: dot_def)
-qed
-
-lemma dist_fstcart: "dist(fstcart (x::real^_)) (fstcart y) <= dist x y"
-  unfolding dist_norm by (metis fstcart_sub[symmetric] norm_fstcart)
-
-lemma norm_sndcart: "norm(sndcart x) <= norm (x::real ^('n::finite + 'm::finite))"
-proof-
-  have th0: "norm x = norm (pastecart (fstcart x) (sndcart x))"
-    by (simp add: pastecart_fst_snd)
-  have th1: "sndcart x \<bullet> sndcart x \<le> pastecart (fstcart x) (sndcart x) \<bullet> pastecart (fstcart x) (sndcart x)"
-    by (simp add: dot_def setsum_UNIV_sum pastecart_def setsum_nonneg)
-  then show ?thesis
-    unfolding th0
-    unfolding real_vector_norm_def real_sqrt_le_iff id_def
-    by (simp add: dot_def)
-qed
-
-lemma dist_sndcart: "dist(sndcart (x::real^_)) (sndcart y) <= dist x y"
-  unfolding dist_norm by (metis sndcart_sub[symmetric] norm_sndcart)
-
-lemma dot_pastecart: "(pastecart (x1::'a::{times,comm_monoid_add}^'n::finite) (x2::'a::{times,comm_monoid_add}^'m::finite)) \<bullet> (pastecart y1 y2) =  x1 \<bullet> y1 + x2 \<bullet> y2"
-  by (simp add: dot_def setsum_UNIV_sum pastecart_def)
-
-text {* TODO: move to NthRoot *}
-lemma sqrt_add_le_add_sqrt:
-  assumes x: "0 \<le> x" and y: "0 \<le> y"
-  shows "sqrt (x + y) \<le> sqrt x + sqrt y"
-apply (rule power2_le_imp_le)
-apply (simp add: real_sum_squared_expand add_nonneg_nonneg x y)
-apply (simp add: mult_nonneg_nonneg x y)
-apply (simp add: add_nonneg_nonneg x y)
-done
-
-lemma norm_pastecart: "norm (pastecart x y) <= norm x + norm y"
-  unfolding norm_vector_def setL2_def setsum_UNIV_sum
-  by (simp add: sqrt_add_le_add_sqrt setsum_nonneg)
-
-subsection {* A generic notion of "hull" (convex, affine, conic hull and closure). *}
-
-definition hull :: "'a set set \<Rightarrow> 'a set \<Rightarrow> 'a set" (infixl "hull" 75) where
-  "S hull s = Inter {t. t \<in> S \<and> s \<subseteq> t}"
-
-lemma hull_same: "s \<in> S \<Longrightarrow> S hull s = s"
-  unfolding hull_def by auto
-
-lemma hull_in: "(\<And>T. T \<subseteq> S ==> Inter T \<in> S) ==> (S hull s) \<in> S"
-unfolding hull_def subset_iff by auto
-
-lemma hull_eq: "(\<And>T. T \<subseteq> S ==> Inter T \<in> S) ==> (S hull s) = s \<longleftrightarrow> s \<in> S"
-using hull_same[of s S] hull_in[of S s] by metis
-
-
-lemma hull_hull: "S hull (S hull s) = S hull s"
-  unfolding hull_def by blast
-
-lemma hull_subset: "s \<subseteq> (S hull s)"
-  unfolding hull_def by blast
-
-lemma hull_mono: " s \<subseteq> t ==> (S hull s) \<subseteq> (S hull t)"
-  unfolding hull_def by blast
-
-lemma hull_antimono: "S \<subseteq> T ==> (T hull s) \<subseteq> (S hull s)"
-  unfolding hull_def by blast
-
-lemma hull_minimal: "s \<subseteq> t \<Longrightarrow> t \<in> S ==> (S hull s) \<subseteq> t"
-  unfolding hull_def by blast
-
-lemma subset_hull: "t \<in> S ==> S hull s \<subseteq> t \<longleftrightarrow>  s \<subseteq> t"
-  unfolding hull_def by blast
-
-lemma hull_unique: "s \<subseteq> t \<Longrightarrow> t \<in> S \<Longrightarrow> (\<And>t'. s \<subseteq> t' \<Longrightarrow> t' \<in> S ==> t \<subseteq> t')
-           ==> (S hull s = t)"
-unfolding hull_def by auto
-
-lemma hull_induct: "(\<And>x. x\<in> S \<Longrightarrow> P x) \<Longrightarrow> Q {x. P x} \<Longrightarrow> \<forall>x\<in> Q hull S. P x"
-  using hull_minimal[of S "{x. P x}" Q]
-  by (auto simp add: subset_eq Collect_def mem_def)
-
-lemma hull_inc: "x \<in> S \<Longrightarrow> x \<in> P hull S" by (metis hull_subset subset_eq)
-
-lemma hull_union_subset: "(S hull s) \<union> (S hull t) \<subseteq> (S hull (s \<union> t))"
-unfolding Un_subset_iff by (metis hull_mono Un_upper1 Un_upper2)
-
-lemma hull_union: assumes T: "\<And>T. T \<subseteq> S ==> Inter T \<in> S"
-  shows "S hull (s \<union> t) = S hull (S hull s \<union> S hull t)"
-apply rule
-apply (rule hull_mono)
-unfolding Un_subset_iff
-apply (metis hull_subset Un_upper1 Un_upper2 subset_trans)
-apply (rule hull_minimal)
-apply (metis hull_union_subset)
-apply (metis hull_in T)
-done
-
-lemma hull_redundant_eq: "a \<in> (S hull s) \<longleftrightarrow> (S hull (insert a s) = S hull s)"
-  unfolding hull_def by blast
-
-lemma hull_redundant: "a \<in> (S hull s) ==> (S hull (insert a s) = S hull s)"
-by (metis hull_redundant_eq)
-
-text{* Archimedian properties and useful consequences. *}
-
-lemma real_arch_simple: "\<exists>n. x <= real (n::nat)"
-  using reals_Archimedean2[of x] apply auto by (rule_tac x="Suc n" in exI, auto)
-lemmas real_arch_lt = reals_Archimedean2
-
-lemmas real_arch = reals_Archimedean3
-
-lemma real_arch_inv: "0 < e \<longleftrightarrow> (\<exists>n::nat. n \<noteq> 0 \<and> 0 < inverse (real n) \<and> inverse (real n) < e)"
-  using reals_Archimedean
-  apply (auto simp add: field_simps inverse_positive_iff_positive)
-  apply (subgoal_tac "inverse (real n) > 0")
-  apply arith
-  apply simp
-  done
-
-lemma real_pow_lbound: "0 <= x ==> 1 + real n * x <= (1 + x) ^ n"
-proof(induct n)
-  case 0 thus ?case by simp
-next
-  case (Suc n)
-  hence h: "1 + real n * x \<le> (1 + x) ^ n" by simp
-  from h have p: "1 \<le> (1 + x) ^ n" using Suc.prems by simp
-  from h have "1 + real n * x + x \<le> (1 + x) ^ n + x" by simp
-  also have "\<dots> \<le> (1 + x) ^ Suc n" apply (subst diff_le_0_iff_le[symmetric])
-    apply (simp add: ring_simps)
-    using mult_left_mono[OF p Suc.prems] by simp
-  finally show ?case  by (simp add: real_of_nat_Suc ring_simps)
-qed
-
-lemma real_arch_pow: assumes x: "1 < (x::real)" shows "\<exists>n. y < x^n"
-proof-
-  from x have x0: "x - 1 > 0" by arith
-  from real_arch[OF x0, rule_format, of y]
-  obtain n::nat where n:"y < real n * (x - 1)" by metis
-  from x0 have x00: "x- 1 \<ge> 0" by arith
-  from real_pow_lbound[OF x00, of n] n
-  have "y < x^n" by auto
-  then show ?thesis by metis
-qed
-
-lemma real_arch_pow2: "\<exists>n. (x::real) < 2^ n"
-  using real_arch_pow[of 2 x] by simp
-
-lemma real_arch_pow_inv: assumes y: "(y::real) > 0" and x1: "x < 1"
-  shows "\<exists>n. x^n < y"
-proof-
-  {assume x0: "x > 0"
-    from x0 x1 have ix: "1 < 1/x" by (simp add: field_simps)
-    from real_arch_pow[OF ix, of "1/y"]
-    obtain n where n: "1/y < (1/x)^n" by blast
-    then
-    have ?thesis using y x0 by (auto simp add: field_simps power_divide) }
-  moreover
-  {assume "\<not> x > 0" with y x1 have ?thesis apply auto by (rule exI[where x=1], auto)}
-  ultimately show ?thesis by metis
-qed
-
-lemma forall_pos_mono: "(\<And>d e::real. d < e \<Longrightarrow> P d ==> P e) \<Longrightarrow> (\<And>n::nat. n \<noteq> 0 ==> P(inverse(real n))) \<Longrightarrow> (\<And>e. 0 < e ==> P e)"
-  by (metis real_arch_inv)
-
-lemma forall_pos_mono_1: "(\<And>d e::real. d < e \<Longrightarrow> P d ==> P e) \<Longrightarrow> (\<And>n. P(inverse(real (Suc n)))) ==> 0 < e ==> P e"
-  apply (rule forall_pos_mono)
-  apply auto
-  apply (atomize)
-  apply (erule_tac x="n - 1" in allE)
-  apply auto
-  done
-
-lemma real_archimedian_rdiv_eq_0: assumes x0: "x \<ge> 0" and c: "c \<ge> 0" and xc: "\<forall>(m::nat)>0. real m * x \<le> c"
-  shows "x = 0"
-proof-
-  {assume "x \<noteq> 0" with x0 have xp: "x > 0" by arith
-    from real_arch[OF xp, rule_format, of c] obtain n::nat where n: "c < real n * x"  by blast
-    with xc[rule_format, of n] have "n = 0" by arith
-    with n c have False by simp}
-  then show ?thesis by blast
-qed
-
-(* ------------------------------------------------------------------------- *)
-(* Relate max and min to sup and inf.                                        *)
-(* ------------------------------------------------------------------------- *)
-
-lemma real_max_rsup: "max x y = rsup {x,y}"
-proof-
-  have f: "finite {x, y}" "{x,y} \<noteq> {}"  by simp_all
-  from rsup_finite_le_iff[OF f, of "max x y"] have "rsup {x,y} \<le> max x y" by simp
-  moreover
-  have "max x y \<le> rsup {x,y}" using rsup_finite_ge_iff[OF f, of "max x y"]
-    by (simp add: linorder_linear)
-  ultimately show ?thesis by arith
-qed
-
-lemma real_min_rinf: "min x y = rinf {x,y}"
-proof-
-  have f: "finite {x, y}" "{x,y} \<noteq> {}"  by simp_all
-  from rinf_finite_le_iff[OF f, of "min x y"] have "rinf {x,y} \<le> min x y"
-    by (simp add: linorder_linear)
-  moreover
-  have "min x y \<le> rinf {x,y}" using rinf_finite_ge_iff[OF f, of "min x y"]
-    by simp
-  ultimately show ?thesis by arith
-qed
-
-(* ------------------------------------------------------------------------- *)
-(* Geometric progression.                                                    *)
-(* ------------------------------------------------------------------------- *)
-
-lemma sum_gp_basic: "((1::'a::{field}) - x) * setsum (\<lambda>i. x^i) {0 .. n} = (1 - x^(Suc n))"
-  (is "?lhs = ?rhs")
-proof-
-  {assume x1: "x = 1" hence ?thesis by simp}
-  moreover
-  {assume x1: "x\<noteq>1"
-    hence x1': "x - 1 \<noteq> 0" "1 - x \<noteq> 0" "x - 1 = - (1 - x)" "- (1 - x) \<noteq> 0" by auto
-    from geometric_sum[OF x1, of "Suc n", unfolded x1']
-    have "(- (1 - x)) * setsum (\<lambda>i. x^i) {0 .. n} = - (1 - x^(Suc n))"
-      unfolding atLeastLessThanSuc_atLeastAtMost
-      using x1' apply (auto simp only: field_simps)
-      apply (simp add: ring_simps)
-      done
-    then have ?thesis by (simp add: ring_simps) }
-  ultimately show ?thesis by metis
-qed
-
-lemma sum_gp_multiplied: assumes mn: "m <= n"
-  shows "((1::'a::{field}) - x) * setsum (op ^ x) {m..n} = x^m - x^ Suc n"
-  (is "?lhs = ?rhs")
-proof-
-  let ?S = "{0..(n - m)}"
-  from mn have mn': "n - m \<ge> 0" by arith
-  let ?f = "op + m"
-  have i: "inj_on ?f ?S" unfolding inj_on_def by auto
-  have f: "?f ` ?S = {m..n}"
-    using mn apply (auto simp add: image_iff Bex_def) by arith
-  have th: "op ^ x o op + m = (\<lambda>i. x^m * x^i)"
-    by (rule ext, simp add: power_add power_mult)
-  from setsum_reindex[OF i, of "op ^ x", unfolded f th setsum_right_distrib[symmetric]]
-  have "?lhs = x^m * ((1 - x) * setsum (op ^ x) {0..n - m})" by simp
-  then show ?thesis unfolding sum_gp_basic using mn
-    by (simp add: ring_simps power_add[symmetric])
-qed
-
-lemma sum_gp: "setsum (op ^ (x::'a::{field})) {m .. n} =
-   (if n < m then 0 else if x = 1 then of_nat ((n + 1) - m)
-                    else (x^ m - x^ (Suc n)) / (1 - x))"
-proof-
-  {assume nm: "n < m" hence ?thesis by simp}
-  moreover
-  {assume "\<not> n < m" hence nm: "m \<le> n" by arith
-    {assume x: "x = 1"  hence ?thesis by simp}
-    moreover
-    {assume x: "x \<noteq> 1" hence nz: "1 - x \<noteq> 0" by simp
-      from sum_gp_multiplied[OF nm, of x] nz have ?thesis by (simp add: field_simps)}
-    ultimately have ?thesis by metis
-  }
-  ultimately show ?thesis by metis
-qed
-
-lemma sum_gp_offset: "setsum (op ^ (x::'a::{field})) {m .. m+n} =
-  (if x = 1 then of_nat n + 1 else x^m * (1 - x^Suc n) / (1 - x))"
-  unfolding sum_gp[of x m "m + n"] power_Suc
-  by (simp add: ring_simps power_add)
-
-
-subsection{* A bit of linear algebra. *}
-
-definition "subspace S \<longleftrightarrow> 0 \<in> S \<and> (\<forall>x\<in> S. \<forall>y \<in>S. x + y \<in> S) \<and> (\<forall>c. \<forall>x \<in>S. c *s x \<in>S )"
-definition "span S = (subspace hull S)"
-definition "dependent S \<longleftrightarrow> (\<exists>a \<in> S. a \<in> span(S - {a}))"
-abbreviation "independent s == ~(dependent s)"
-
-(* Closure properties of subspaces.                                          *)
-
-lemma subspace_UNIV[simp]: "subspace(UNIV)" by (simp add: subspace_def)
-
-lemma subspace_0: "subspace S ==> 0 \<in> S" by (metis subspace_def)
-
-lemma subspace_add: "subspace S \<Longrightarrow> x \<in> S \<Longrightarrow> y \<in> S ==> x + y \<in> S"
-  by (metis subspace_def)
-
-lemma subspace_mul: "subspace S \<Longrightarrow> x \<in> S \<Longrightarrow> c *s x \<in> S"
-  by (metis subspace_def)
-
-lemma subspace_neg: "subspace S \<Longrightarrow> (x::'a::ring_1^'n) \<in> S \<Longrightarrow> - x \<in> S"
-  by (metis vector_sneg_minus1 subspace_mul)
-
-lemma subspace_sub: "subspace S \<Longrightarrow> (x::'a::ring_1^'n) \<in> S \<Longrightarrow> y \<in> S \<Longrightarrow> x - y \<in> S"
-  by (metis diff_def subspace_add subspace_neg)
-
-lemma subspace_setsum:
-  assumes sA: "subspace A" and fB: "finite B"
-  and f: "\<forall>x\<in> B. f x \<in> A"
-  shows "setsum f B \<in> A"
-  using  fB f sA
-  apply(induct rule: finite_induct[OF fB])
-  by (simp add: subspace_def sA, auto simp add: sA subspace_add)
-
-lemma subspace_linear_image:
-  assumes lf: "linear (f::'a::semiring_1^'n \<Rightarrow> _)" and sS: "subspace S"
-  shows "subspace(f ` S)"
-  using lf sS linear_0[OF lf]
-  unfolding linear_def subspace_def
-  apply (auto simp add: image_iff)
-  apply (rule_tac x="x + y" in bexI, auto)
-  apply (rule_tac x="c*s x" in bexI, auto)
-  done
-
-lemma subspace_linear_preimage: "linear (f::'a::semiring_1^'n \<Rightarrow> _) ==> subspace S ==> subspace {x. f x \<in> S}"
-  by (auto simp add: subspace_def linear_def linear_0[of f])
-
-lemma subspace_trivial: "subspace {0::'a::semiring_1 ^_}"
-  by (simp add: subspace_def)
-
-lemma subspace_inter: "subspace A \<Longrightarrow> subspace B ==> subspace (A \<inter> B)"
-  by (simp add: subspace_def)
-
-
-lemma span_mono: "A \<subseteq> B ==> span A \<subseteq> span B"
-  by (metis span_def hull_mono)
-
-lemma subspace_span: "subspace(span S)"
-  unfolding span_def
-  apply (rule hull_in[unfolded mem_def])
-  apply (simp only: subspace_def Inter_iff Int_iff subset_eq)
-  apply auto
-  apply (erule_tac x="X" in ballE)
-  apply (simp add: mem_def)
-  apply blast
-  apply (erule_tac x="X" in ballE)
-  apply (erule_tac x="X" in ballE)
-  apply (erule_tac x="X" in ballE)
-  apply (clarsimp simp add: mem_def)
-  apply simp
-  apply simp
-  apply simp
-  apply (erule_tac x="X" in ballE)
-  apply (erule_tac x="X" in ballE)
-  apply (simp add: mem_def)
-  apply simp
-  apply simp
-  done
-
-lemma span_clauses:
-  "a \<in> S ==> a \<in> span S"
-  "0 \<in> span S"
-  "x\<in> span S \<Longrightarrow> y \<in> span S ==> x + y \<in> span S"
-  "x \<in> span S \<Longrightarrow> c *s x \<in> span S"
-  by (metis span_def hull_subset subset_eq subspace_span subspace_def)+
-
-lemma span_induct: assumes SP: "\<And>x. x \<in> S ==> P x"
-  and P: "subspace P" and x: "x \<in> span S" shows "P x"
-proof-
-  from SP have SP': "S \<subseteq> P" by (simp add: mem_def subset_eq)
-  from P have P': "P \<in> subspace" by (simp add: mem_def)
-  from x hull_minimal[OF SP' P', unfolded span_def[symmetric]]
-  show "P x" by (metis mem_def subset_eq)
-qed
-
-lemma span_empty: "span {} = {(0::'a::semiring_0 ^ 'n)}"
-  apply (simp add: span_def)
-  apply (rule hull_unique)
-  apply (auto simp add: mem_def subspace_def)
-  unfolding mem_def[of "0::'a^'n", symmetric]
-  apply simp
-  done
-
-lemma independent_empty: "independent {}"
-  by (simp add: dependent_def)
-
-lemma independent_mono: "independent A \<Longrightarrow> B \<subseteq> A ==> independent B"
-  apply (clarsimp simp add: dependent_def span_mono)
-  apply (subgoal_tac "span (B - {a}) \<le> span (A - {a})")
-  apply force
-  apply (rule span_mono)
-  apply auto
-  done
-
-lemma span_subspace: "A \<subseteq> B \<Longrightarrow> B \<le> span A \<Longrightarrow>  subspace B \<Longrightarrow> span A = B"
-  by (metis order_antisym span_def hull_minimal mem_def)
-
-lemma span_induct': assumes SP: "\<forall>x \<in> S. P x"
-  and P: "subspace P" shows "\<forall>x \<in> span S. P x"
-  using span_induct SP P by blast
-
-inductive span_induct_alt_help for S:: "'a::semiring_1^'n \<Rightarrow> bool"
-  where
-  span_induct_alt_help_0: "span_induct_alt_help S 0"
-  | span_induct_alt_help_S: "x \<in> S \<Longrightarrow> span_induct_alt_help S z \<Longrightarrow> span_induct_alt_help S (c *s x + z)"
-
-lemma span_induct_alt':
-  assumes h0: "h (0::'a::semiring_1^'n)" and hS: "\<And>c x y. x \<in> S \<Longrightarrow> h y \<Longrightarrow> h (c*s x + y)" shows "\<forall>x \<in> span S. h x"
-proof-
-  {fix x:: "'a^'n" assume x: "span_induct_alt_help S x"
-    have "h x"
-      apply (rule span_induct_alt_help.induct[OF x])
-      apply (rule h0)
-      apply (rule hS, assumption, assumption)
-      done}
-  note th0 = this
-  {fix x assume x: "x \<in> span S"
-
-    have "span_induct_alt_help S x"
-      proof(rule span_induct[where x=x and S=S])
-        show "x \<in> span S" using x .
-      next
-        fix x assume xS : "x \<in> S"
-          from span_induct_alt_help_S[OF xS span_induct_alt_help_0, of 1]
-          show "span_induct_alt_help S x" by simp
-        next
-        have "span_induct_alt_help S 0" by (rule span_induct_alt_help_0)
-        moreover
-        {fix x y assume h: "span_induct_alt_help S x" "span_induct_alt_help S y"
-          from h
-          have "span_induct_alt_help S (x + y)"
-            apply (induct rule: span_induct_alt_help.induct)
-            apply simp
-            unfolding add_assoc
-            apply (rule span_induct_alt_help_S)
-            apply assumption
-            apply simp
-            done}
-        moreover
-        {fix c x assume xt: "span_induct_alt_help S x"
-          then have "span_induct_alt_help S (c*s x)"
-            apply (induct rule: span_induct_alt_help.induct)
-            apply (simp add: span_induct_alt_help_0)
-            apply (simp add: vector_smult_assoc vector_add_ldistrib)
-            apply (rule span_induct_alt_help_S)
-            apply assumption
-            apply simp
-            done
-        }
-        ultimately show "subspace (span_induct_alt_help S)"
-          unfolding subspace_def mem_def Ball_def by blast
-      qed}
-  with th0 show ?thesis by blast
-qed
-
-lemma span_induct_alt:
-  assumes h0: "h (0::'a::semiring_1^'n)" and hS: "\<And>c x y. x \<in> S \<Longrightarrow> h y \<Longrightarrow> h (c*s x + y)" and x: "x \<in> span S"
-  shows "h x"
-using span_induct_alt'[of h S] h0 hS x by blast
-
-(* Individual closure properties. *)
-
-lemma span_superset: "x \<in> S ==> x \<in> span S" by (metis span_clauses)
-
-lemma span_0: "0 \<in> span S" by (metis subspace_span subspace_0)
-
-lemma span_add: "x \<in> span S \<Longrightarrow> y \<in> span S ==> x + y \<in> span S"
-  by (metis subspace_add subspace_span)
-
-lemma span_mul: "x \<in> span S ==> (c *s x) \<in> span S"
-  by (metis subspace_span subspace_mul)
-
-lemma span_neg: "x \<in> span S ==> - (x::'a::ring_1^'n) \<in> span S"
-  by (metis subspace_neg subspace_span)
-
-lemma span_sub: "(x::'a::ring_1^'n) \<in> span S \<Longrightarrow> y \<in> span S ==> x - y \<in> span S"
-  by (metis subspace_span subspace_sub)
-
-lemma span_setsum: "finite A \<Longrightarrow> \<forall>x \<in> A. f x \<in> span S ==> setsum f A \<in> span S"
-  apply (rule subspace_setsum)
-  by (metis subspace_span subspace_setsum)+
-
-lemma span_add_eq: "(x::'a::ring_1^'n) \<in> span S \<Longrightarrow> x + y \<in> span S \<longleftrightarrow> y \<in> span S"
-  apply (auto simp only: span_add span_sub)
-  apply (subgoal_tac "(x + y) - x \<in> span S", simp)
-  by (simp only: span_add span_sub)
-
-(* Mapping under linear image. *)
-
-lemma span_linear_image: assumes lf: "linear (f::'a::semiring_1 ^ 'n => _)"
-  shows "span (f ` S) = f ` (span S)"
-proof-
-  {fix x
-    assume x: "x \<in> span (f ` S)"
-    have "x \<in> f ` span S"
-      apply (rule span_induct[where x=x and S = "f ` S"])
-      apply (clarsimp simp add: image_iff)
-      apply (frule span_superset)
-      apply blast
-      apply (simp only: mem_def)
-      apply (rule subspace_linear_image[OF lf])
-      apply (rule subspace_span)
-      apply (rule x)
-      done}
-  moreover
-  {fix x assume x: "x \<in> span S"
-    have th0:"(\<lambda>a. f a \<in> span (f ` S)) = {x. f x \<in> span (f ` S)}" apply (rule set_ext)
-      unfolding mem_def Collect_def ..
-    have "f x \<in> span (f ` S)"
-      apply (rule span_induct[where S=S])
-      apply (rule span_superset)
-      apply simp
-      apply (subst th0)
-      apply (rule subspace_linear_preimage[OF lf subspace_span, of "f ` S"])
-      apply (rule x)
-      done}
-  ultimately show ?thesis by blast
-qed
-
-(* The key breakdown property. *)
-
-lemma span_breakdown:
-  assumes bS: "(b::'a::ring_1 ^ 'n) \<in> S" and aS: "a \<in> span S"
-  shows "\<exists>k. a - k*s b \<in> span (S - {b})" (is "?P a")
-proof-
-  {fix x assume xS: "x \<in> S"
-    {assume ab: "x = b"
-      then have "?P x"
-        apply simp
-        apply (rule exI[where x="1"], simp)
-        by (rule span_0)}
-    moreover
-    {assume ab: "x \<noteq> b"
-      then have "?P x"  using xS
-        apply -
-        apply (rule exI[where x=0])
-        apply (rule span_superset)
-        by simp}
-    ultimately have "?P x" by blast}
-  moreover have "subspace ?P"
-    unfolding subspace_def
-    apply auto
-    apply (simp add: mem_def)
-    apply (rule exI[where x=0])
-    using span_0[of "S - {b}"]
-    apply (simp add: mem_def)
-    apply (clarsimp simp add: mem_def)
-    apply (rule_tac x="k + ka" in exI)
-    apply (subgoal_tac "x + y - (k + ka) *s b = (x - k*s b) + (y - ka *s b)")
-    apply (simp only: )
-    apply (rule span_add[unfolded mem_def])
-    apply assumption+
-    apply (vector ring_simps)
-    apply (clarsimp simp add: mem_def)
-    apply (rule_tac x= "c*k" in exI)
-    apply (subgoal_tac "c *s x - (c * k) *s b = c*s (x - k*s b)")
-    apply (simp only: )
-    apply (rule span_mul[unfolded mem_def])
-    apply assumption
-    by (vector ring_simps)
-  ultimately show "?P a" using aS span_induct[where S=S and P= "?P"] by metis
-qed
-
-lemma span_breakdown_eq:
-  "(x::'a::ring_1^'n) \<in> span (insert a S) \<longleftrightarrow> (\<exists>k. (x - k *s a) \<in> span S)" (is "?lhs \<longleftrightarrow> ?rhs")
-proof-
-  {assume x: "x \<in> span (insert a S)"
-    from x span_breakdown[of "a" "insert a S" "x"]
-    have ?rhs apply clarsimp
-      apply (rule_tac x= "k" in exI)
-      apply (rule set_rev_mp[of _ "span (S - {a})" _])
-      apply assumption
-      apply (rule span_mono)
-      apply blast
-      done}
-  moreover
-  { fix k assume k: "x - k *s a \<in> span S"
-    have eq: "x = (x - k *s a) + k *s a" by vector
-    have "(x - k *s a) + k *s a \<in> span (insert a S)"
-      apply (rule span_add)
-      apply (rule set_rev_mp[of _ "span S" _])
-      apply (rule k)
-      apply (rule span_mono)
-      apply blast
-      apply (rule span_mul)
-      apply (rule span_superset)
-      apply blast
-      done
-    then have ?lhs using eq by metis}
-  ultimately show ?thesis by blast
-qed
-
-(* Hence some "reversal" results.*)
-
-lemma in_span_insert:
-  assumes a: "(a::'a::field^'n) \<in> span (insert b S)" and na: "a \<notin> span S"
-  shows "b \<in> span (insert a S)"
-proof-
-  from span_breakdown[of b "insert b S" a, OF insertI1 a]
-  obtain k where k: "a - k*s b \<in> span (S - {b})" by auto
-  {assume k0: "k = 0"
-    with k have "a \<in> span S"
-      apply (simp)
-      apply (rule set_rev_mp)
-      apply assumption
-      apply (rule span_mono)
-      apply blast
-      done
-    with na  have ?thesis by blast}
-  moreover
-  {assume k0: "k \<noteq> 0"
-    have eq: "b = (1/k) *s a - ((1/k) *s a - b)" by vector
-    from k0 have eq': "(1/k) *s (a - k*s b) = (1/k) *s a - b"
-      by (vector field_simps)
-    from k have "(1/k) *s (a - k*s b) \<in> span (S - {b})"
-      by (rule span_mul)
-    hence th: "(1/k) *s a - b \<in> span (S - {b})"
-      unfolding eq' .
-
-    from k
-    have ?thesis
-      apply (subst eq)
-      apply (rule span_sub)
-      apply (rule span_mul)
-      apply (rule span_superset)
-      apply blast
-      apply (rule set_rev_mp)
-      apply (rule th)
-      apply (rule span_mono)
-      using na by blast}
-  ultimately show ?thesis by blast
-qed
-
-lemma in_span_delete:
-  assumes a: "(a::'a::field^'n) \<in> span S"
-  and na: "a \<notin> span (S-{b})"
-  shows "b \<in> span (insert a (S - {b}))"
-  apply (rule in_span_insert)
-  apply (rule set_rev_mp)
-  apply (rule a)
-  apply (rule span_mono)
-  apply blast
-  apply (rule na)
-  done
-
-(* Transitivity property. *)
-
-lemma span_trans:
-  assumes x: "(x::'a::ring_1^'n) \<in> span S" and y: "y \<in> span (insert x S)"
-  shows "y \<in> span S"
-proof-
-  from span_breakdown[of x "insert x S" y, OF insertI1 y]
-  obtain k where k: "y -k*s x \<in> span (S - {x})" by auto
-  have eq: "y = (y - k *s x) + k *s x" by vector
-  show ?thesis
-    apply (subst eq)
-    apply (rule span_add)
-    apply (rule set_rev_mp)
-    apply (rule k)
-    apply (rule span_mono)
-    apply blast
-    apply (rule span_mul)
-    by (rule x)
-qed
-
-(* ------------------------------------------------------------------------- *)
-(* An explicit expansion is sometimes needed.                                *)
-(* ------------------------------------------------------------------------- *)
-
-lemma span_explicit:
-  "span P = {y::'a::semiring_1^'n. \<exists>S u. finite S \<and> S \<subseteq> P \<and> setsum (\<lambda>v. u v *s v) S = y}"
-  (is "_ = ?E" is "_ = {y. ?h y}" is "_ = {y. \<exists>S u. ?Q S u y}")
-proof-
-  {fix x assume x: "x \<in> ?E"
-    then obtain S u where fS: "finite S" and SP: "S\<subseteq>P" and u: "setsum (\<lambda>v. u v *s v) S = x"
-      by blast
-    have "x \<in> span P"
-      unfolding u[symmetric]
-      apply (rule span_setsum[OF fS])
-      using span_mono[OF SP]
-      by (auto intro: span_superset span_mul)}
-  moreover
-  have "\<forall>x \<in> span P. x \<in> ?E"
-    unfolding mem_def Collect_def
-  proof(rule span_induct_alt')
-    show "?h 0"
-      apply (rule exI[where x="{}"]) by simp
-  next
-    fix c x y
-    assume x: "x \<in> P" and hy: "?h y"
-    from hy obtain S u where fS: "finite S" and SP: "S\<subseteq>P"
-      and u: "setsum (\<lambda>v. u v *s v) S = y" by blast
-    let ?S = "insert x S"
-    let ?u = "\<lambda>y. if y = x then (if x \<in> S then u y + c else c)
-                  else u y"
-    from fS SP x have th0: "finite (insert x S)" "insert x S \<subseteq> P" by blast+
-    {assume xS: "x \<in> S"
-      have S1: "S = (S - {x}) \<union> {x}"
-        and Sss:"finite (S - {x})" "finite {x}" "(S -{x}) \<inter> {x} = {}" using xS fS by auto
-      have "setsum (\<lambda>v. ?u v *s v) ?S =(\<Sum>v\<in>S - {x}. u v *s v) + (u x + c) *s x"
-        using xS
-        by (simp add: setsum_Un_disjoint[OF Sss, unfolded S1[symmetric]]
-          setsum_clauses(2)[OF fS] cong del: if_weak_cong)
-      also have "\<dots> = (\<Sum>v\<in>S. u v *s v) + c *s x"
-        apply (simp add: setsum_Un_disjoint[OF Sss, unfolded S1[symmetric]])
-        by (vector ring_simps)
-      also have "\<dots> = c*s x + y"
-        by (simp add: add_commute u)
-      finally have "setsum (\<lambda>v. ?u v *s v) ?S = c*s x + y" .
-    then have "?Q ?S ?u (c*s x + y)" using th0 by blast}
-  moreover
-  {assume xS: "x \<notin> S"
-    have th00: "(\<Sum>v\<in>S. (if v = x then c else u v) *s v) = y"
-      unfolding u[symmetric]
-      apply (rule setsum_cong2)
-      using xS by auto
-    have "?Q ?S ?u (c*s x + y)" using fS xS th0
-      by (simp add: th00 setsum_clauses add_commute cong del: if_weak_cong)}
-  ultimately have "?Q ?S ?u (c*s x + y)"
-    by (cases "x \<in> S", simp, simp)
-    then show "?h (c*s x + y)"
-      apply -
-      apply (rule exI[where x="?S"])
-      apply (rule exI[where x="?u"]) by metis
-  qed
-  ultimately show ?thesis by blast
-qed
-
-lemma dependent_explicit:
-  "dependent P \<longleftrightarrow> (\<exists>S u. finite S \<and> S \<subseteq> P \<and> (\<exists>(v::'a::{idom,field}^'n) \<in>S. u v \<noteq> 0 \<and> setsum (\<lambda>v. u v *s v) S = 0))" (is "?lhs = ?rhs")
-proof-
-  {assume dP: "dependent P"
-    then obtain a S u where aP: "a \<in> P" and fS: "finite S"
-      and SP: "S \<subseteq> P - {a}" and ua: "setsum (\<lambda>v. u v *s v) S = a"
-      unfolding dependent_def span_explicit by blast
-    let ?S = "insert a S"
-    let ?u = "\<lambda>y. if y = a then - 1 else u y"
-    let ?v = a
-    from aP SP have aS: "a \<notin> S" by blast
-    from fS SP aP have th0: "finite ?S" "?S \<subseteq> P" "?v \<in> ?S" "?u ?v \<noteq> 0" by auto
-    have s0: "setsum (\<lambda>v. ?u v *s v) ?S = 0"
-      using fS aS
-      apply (simp add: vector_smult_lneg vector_smult_lid setsum_clauses ring_simps )
-      apply (subst (2) ua[symmetric])
-      apply (rule setsum_cong2)
-      by auto
-    with th0 have ?rhs
-      apply -
-      apply (rule exI[where x= "?S"])
-      apply (rule exI[where x= "?u"])
-      by clarsimp}
-  moreover
-  {fix S u v assume fS: "finite S"
-      and SP: "S \<subseteq> P" and vS: "v \<in> S" and uv: "u v \<noteq> 0"
-    and u: "setsum (\<lambda>v. u v *s v) S = 0"
-    let ?a = v
-    let ?S = "S - {v}"
-    let ?u = "\<lambda>i. (- u i) / u v"
-    have th0: "?a \<in> P" "finite ?S" "?S \<subseteq> P"       using fS SP vS by auto
-    have "setsum (\<lambda>v. ?u v *s v) ?S = setsum (\<lambda>v. (- (inverse (u ?a))) *s (u v *s v)) S - ?u v *s v"
-      using fS vS uv
-      by (simp add: setsum_diff1 vector_smult_lneg divide_inverse
-        vector_smult_assoc field_simps)
-    also have "\<dots> = ?a"
-      unfolding setsum_cmul u
-      using uv by (simp add: vector_smult_lneg)
-    finally  have "setsum (\<lambda>v. ?u v *s v) ?S = ?a" .
-    with th0 have ?lhs
-      unfolding dependent_def span_explicit
-      apply -
-      apply (rule bexI[where x= "?a"])
-      apply simp_all
-      apply (rule exI[where x= "?S"])
-      by auto}
-  ultimately show ?thesis by blast
-qed
-
-
-lemma span_finite:
-  assumes fS: "finite S"
-  shows "span S = {(y::'a::semiring_1^'n). \<exists>u. setsum (\<lambda>v. u v *s v) S = y}"
-  (is "_ = ?rhs")
-proof-
-  {fix y assume y: "y \<in> span S"
-    from y obtain S' u where fS': "finite S'" and SS': "S' \<subseteq> S" and
-      u: "setsum (\<lambda>v. u v *s v) S' = y" unfolding span_explicit by blast
-    let ?u = "\<lambda>x. if x \<in> S' then u x else 0"
-    from setsum_restrict_set[OF fS, of "\<lambda>v. u v *s v" S', symmetric] SS'
-    have "setsum (\<lambda>v. ?u v *s v) S = setsum (\<lambda>v. u v *s v) S'"
-      unfolding cond_value_iff cond_application_beta
-      by (simp add: cond_value_iff inf_absorb2 cong del: if_weak_cong)
-    hence "setsum (\<lambda>v. ?u v *s v) S = y" by (metis u)
-    hence "y \<in> ?rhs" by auto}
-  moreover
-  {fix y u assume u: "setsum (\<lambda>v. u v *s v) S = y"
-    then have "y \<in> span S" using fS unfolding span_explicit by auto}
-  ultimately show ?thesis by blast
-qed
-
-
-(* Standard bases are a spanning set, and obviously finite.                  *)
-
-lemma span_stdbasis:"span {basis i :: 'a::ring_1^'n::finite | i. i \<in> (UNIV :: 'n set)} = UNIV"
-apply (rule set_ext)
-apply auto
-apply (subst basis_expansion[symmetric])
-apply (rule span_setsum)
-apply simp
-apply auto
-apply (rule span_mul)
-apply (rule span_superset)
-apply (auto simp add: Collect_def mem_def)
-done
-
-lemma has_size_stdbasis: "{basis i ::real ^'n::finite | i. i \<in> (UNIV :: 'n set)} hassize CARD('n)" (is "?S hassize ?n")
-proof-
-  have eq: "?S = basis ` UNIV" by blast
-  show ?thesis unfolding eq
-    apply (rule hassize_image_inj[OF basis_inj])
-    by (simp add: hassize_def)
-qed
-
-lemma finite_stdbasis: "finite {basis i ::real^'n::finite |i. i\<in> (UNIV:: 'n set)}"
-  using has_size_stdbasis[unfolded hassize_def]
-  ..
-
-lemma card_stdbasis: "card {basis i ::real^'n::finite |i. i\<in> (UNIV :: 'n set)} = CARD('n)"
-  using has_size_stdbasis[unfolded hassize_def]
-  ..
-
-lemma independent_stdbasis_lemma:
-  assumes x: "(x::'a::semiring_1 ^ 'n) \<in> span (basis ` S)"
-  and iS: "i \<notin> S"
-  shows "(x$i) = 0"
-proof-
-  let ?U = "UNIV :: 'n set"
-  let ?B = "basis ` S"
-  let ?P = "\<lambda>(x::'a^'n). \<forall>i\<in> ?U. i \<notin> S \<longrightarrow> x$i =0"
- {fix x::"'a^'n" assume xS: "x\<in> ?B"
-   from xS have "?P x" by auto}
- moreover
- have "subspace ?P"
-   by (auto simp add: subspace_def Collect_def mem_def)
- ultimately show ?thesis
-   using x span_induct[of ?B ?P x] iS by blast
-qed
-
-lemma independent_stdbasis: "independent {basis i ::real^'n::finite |i. i\<in> (UNIV :: 'n set)}"
-proof-
-  let ?I = "UNIV :: 'n set"
-  let ?b = "basis :: _ \<Rightarrow> real ^'n"
-  let ?B = "?b ` ?I"
-  have eq: "{?b i|i. i \<in> ?I} = ?B"
-    by auto
-  {assume d: "dependent ?B"
-    then obtain k where k: "k \<in> ?I" "?b k \<in> span (?B - {?b k})"
-      unfolding dependent_def by auto
-    have eq1: "?B - {?b k} = ?B - ?b ` {k}"  by simp
-    have eq2: "?B - {?b k} = ?b ` (?I - {k})"
-      unfolding eq1
-      apply (rule inj_on_image_set_diff[symmetric])
-      apply (rule basis_inj) using k(1) by auto
-    from k(2) have th0: "?b k \<in> span (?b ` (?I - {k}))" unfolding eq2 .
-    from independent_stdbasis_lemma[OF th0, of k, simplified]
-    have False by simp}
-  then show ?thesis unfolding eq dependent_def ..
-qed
-
-(* This is useful for building a basis step-by-step.                         *)
-
-lemma independent_insert:
-  "independent(insert (a::'a::field ^'n) S) \<longleftrightarrow>
-      (if a \<in> S then independent S
-                else independent S \<and> a \<notin> span S)" (is "?lhs \<longleftrightarrow> ?rhs")
-proof-
-  {assume aS: "a \<in> S"
-    hence ?thesis using insert_absorb[OF aS] by simp}
-  moreover
-  {assume aS: "a \<notin> S"
-    {assume i: ?lhs
-      then have ?rhs using aS
-        apply simp
-        apply (rule conjI)
-        apply (rule independent_mono)
-        apply assumption
-        apply blast
-        by (simp add: dependent_def)}
-    moreover
-    {assume i: ?rhs
-      have ?lhs using i aS
-        apply simp
-        apply (auto simp add: dependent_def)
-        apply (case_tac "aa = a", auto)
-        apply (subgoal_tac "insert a S - {aa} = insert a (S - {aa})")
-        apply simp
-        apply (subgoal_tac "a \<in> span (insert aa (S - {aa}))")
-        apply (subgoal_tac "insert aa (S - {aa}) = S")
-        apply simp
-        apply blast
-        apply (rule in_span_insert)
-        apply assumption
-        apply blast
-        apply blast
-        done}
-    ultimately have ?thesis by blast}
-  ultimately show ?thesis by blast
-qed
-
-(* The degenerate case of the Exchange Lemma.  *)
-
-lemma mem_delete: "x \<in> (A - {a}) \<longleftrightarrow> x \<noteq> a \<and> x \<in> A"
-  by blast
-
-lemma span_span: "span (span A) = span A"
-  unfolding span_def hull_hull ..
-
-lemma span_inc: "S \<subseteq> span S"
-  by (metis subset_eq span_superset)
-
-lemma spanning_subset_independent:
-  assumes BA: "B \<subseteq> A" and iA: "independent (A::('a::field ^'n) set)"
-  and AsB: "A \<subseteq> span B"
-  shows "A = B"
-proof
-  from BA show "B \<subseteq> A" .
-next
-  from span_mono[OF BA] span_mono[OF AsB]
-  have sAB: "span A = span B" unfolding span_span by blast
-
-  {fix x assume x: "x \<in> A"
-    from iA have th0: "x \<notin> span (A - {x})"
-      unfolding dependent_def using x by blast
-    from x have xsA: "x \<in> span A" by (blast intro: span_superset)
-    have "A - {x} \<subseteq> A" by blast
-    hence th1:"span (A - {x}) \<subseteq> span A" by (metis span_mono)
-    {assume xB: "x \<notin> B"
-      from xB BA have "B \<subseteq> A -{x}" by blast
-      hence "span B \<subseteq> span (A - {x})" by (metis span_mono)
-      with th1 th0 sAB have "x \<notin> span A" by blast
-      with x have False by (metis span_superset)}
-    then have "x \<in> B" by blast}
-  then show "A \<subseteq> B" by blast
-qed
-
-(* The general case of the Exchange Lemma, the key to what follows.  *)
-
-lemma exchange_lemma:
-  assumes f:"finite (t:: ('a::field^'n) set)" and i: "independent s"
-  and sp:"s \<subseteq> span t"
-  shows "\<exists>t'. (t' hassize card t) \<and> s \<subseteq> t' \<and> t' \<subseteq> s \<union> t \<and> s \<subseteq> span t'"
-using f i sp
-proof(induct c\<equiv>"card(t - s)" arbitrary: s t rule: nat_less_induct)
-  fix n:: nat and s t :: "('a ^'n) set"
-  assume H: " \<forall>m<n. \<forall>(x:: ('a ^'n) set) xa.
-                finite xa \<longrightarrow>
-                independent x \<longrightarrow>
-                x \<subseteq> span xa \<longrightarrow>
-                m = card (xa - x) \<longrightarrow>
-                (\<exists>t'. (t' hassize card xa) \<and>
-                      x \<subseteq> t' \<and> t' \<subseteq> x \<union> xa \<and> x \<subseteq> span t')"
-    and ft: "finite t" and s: "independent s" and sp: "s \<subseteq> span t"
-    and n: "n = card (t - s)"
-  let ?P = "\<lambda>t'. (t' hassize card t) \<and> s \<subseteq> t' \<and> t' \<subseteq> s \<union> t \<and> s \<subseteq> span t'"
-  let ?ths = "\<exists>t'. ?P t'"
-  {assume st: "s \<subseteq> t"
-    from st ft span_mono[OF st] have ?ths apply - apply (rule exI[where x=t])
-      by (auto simp add: hassize_def intro: span_superset)}
-  moreover
-  {assume st: "t \<subseteq> s"
-
-    from spanning_subset_independent[OF st s sp]
-      st ft span_mono[OF st] have ?ths apply - apply (rule exI[where x=t])
-      by (auto simp add: hassize_def intro: span_superset)}
-  moreover
-  {assume st: "\<not> s \<subseteq> t" "\<not> t \<subseteq> s"
-    from st(2) obtain b where b: "b \<in> t" "b \<notin> s" by blast
-      from b have "t - {b} - s \<subset> t - s" by blast
-      then have cardlt: "card (t - {b} - s) < n" using n ft
-        by (auto intro: psubset_card_mono)
-      from b ft have ct0: "card t \<noteq> 0" by auto
-    {assume stb: "s \<subseteq> span(t -{b})"
-      from ft have ftb: "finite (t -{b})" by auto
-      from H[rule_format, OF cardlt ftb s stb]
-      obtain u where u: "u hassize card (t-{b})" "s \<subseteq> u" "u \<subseteq> s \<union> (t - {b})" "s \<subseteq> span u" by blast
-      let ?w = "insert b u"
-      have th0: "s \<subseteq> insert b u" using u by blast
-      from u(3) b have "u \<subseteq> s \<union> t" by blast
-      then have th1: "insert b u \<subseteq> s \<union> t" using u b by blast
-      have bu: "b \<notin> u" using b u by blast
-      from u(1) have fu: "finite u" by (simp add: hassize_def)
-      from u(1) ft b have "u hassize (card t - 1)" by auto
-      then
-      have th2: "insert b u hassize card t"
-        using  card_insert_disjoint[OF fu bu] ct0 by (auto simp add: hassize_def)
-      from u(4) have "s \<subseteq> span u" .
-      also have "\<dots> \<subseteq> span (insert b u)" apply (rule span_mono) by blast
-      finally have th3: "s \<subseteq> span (insert b u)" .      from th0 th1 th2 th3 have th: "?P ?w"  by blast
-      from th have ?ths by blast}
-    moreover
-    {assume stb: "\<not> s \<subseteq> span(t -{b})"
-      from stb obtain a where a: "a \<in> s" "a \<notin> span (t - {b})" by blast
-      have ab: "a \<noteq> b" using a b by blast
-      have at: "a \<notin> t" using a ab span_superset[of a "t- {b}"] by auto
-      have mlt: "card ((insert a (t - {b})) - s) < n"
-        using cardlt ft n  a b by auto
-      have ft': "finite (insert a (t - {b}))" using ft by auto
-      {fix x assume xs: "x \<in> s"
-        have t: "t \<subseteq> (insert b (insert a (t -{b})))" using b by auto
-        from b(1) have "b \<in> span t" by (simp add: span_superset)
-        have bs: "b \<in> span (insert a (t - {b}))"
-          by (metis in_span_delete a sp mem_def subset_eq)
-        from xs sp have "x \<in> span t" by blast
-        with span_mono[OF t]
-        have x: "x \<in> span (insert b (insert a (t - {b})))" ..
-        from span_trans[OF bs x] have "x \<in> span (insert a (t - {b}))"  .}
-      then have sp': "s \<subseteq> span (insert a (t - {b}))" by blast
-
-      from H[rule_format, OF mlt ft' s sp' refl] obtain u where
-        u: "u hassize card (insert a (t -{b}))" "s \<subseteq> u" "u \<subseteq> s \<union> insert a (t -{b})"
-        "s \<subseteq> span u" by blast
-      from u a b ft at ct0 have "?P u" by (auto simp add: hassize_def)
-      then have ?ths by blast }
-    ultimately have ?ths by blast
-  }
-  ultimately
-  show ?ths  by blast
-qed
-
-(* This implies corresponding size bounds.                                   *)
-
-lemma independent_span_bound:
-  assumes f: "finite t" and i: "independent (s::('a::field^'n) set)" and sp:"s \<subseteq> span t"
-  shows "finite s \<and> card s \<le> card t"
-  by (metis exchange_lemma[OF f i sp] hassize_def finite_subset card_mono)
-
-
-lemma finite_Atleast_Atmost_nat[simp]: "finite {f x |x. x\<in> (UNIV::'a::finite set)}"
-proof-
-  have eq: "{f x |x. x\<in> UNIV} = f ` UNIV" by auto
-  show ?thesis unfolding eq
-    apply (rule finite_imageI)
-    apply (rule finite)
-    done
-qed
-
-
-lemma independent_bound:
-  fixes S:: "(real^'n::finite) set"
-  shows "independent S \<Longrightarrow> finite S \<and> card S <= CARD('n)"
-  apply (subst card_stdbasis[symmetric])
-  apply (rule independent_span_bound)
-  apply (rule finite_Atleast_Atmost_nat)
-  apply assumption
-  unfolding span_stdbasis
-  apply (rule subset_UNIV)
-  done
-
-lemma dependent_biggerset: "(finite (S::(real ^'n::finite) set) ==> card S > CARD('n)) ==> dependent S"
-  by (metis independent_bound not_less)
-
-(* Hence we can create a maximal independent subset.                         *)
-
-lemma maximal_independent_subset_extend:
-  assumes sv: "(S::(real^'n::finite) set) \<subseteq> V" and iS: "independent S"
-  shows "\<exists>B. S \<subseteq> B \<and> B \<subseteq> V \<and> independent B \<and> V \<subseteq> span B"
-  using sv iS
-proof(induct d\<equiv> "CARD('n) - card S" arbitrary: S rule: nat_less_induct)
-  fix n and S:: "(real^'n) set"
-  assume H: "\<forall>m<n. \<forall>S \<subseteq> V. independent S \<longrightarrow> m = CARD('n) - card S \<longrightarrow>
-              (\<exists>B. S \<subseteq> B \<and> B \<subseteq> V \<and> independent B \<and> V \<subseteq> span B)"
-    and sv: "S \<subseteq> V" and i: "independent S" and n: "n = CARD('n) - card S"
-  let ?P = "\<lambda>B. S \<subseteq> B \<and> B \<subseteq> V \<and> independent B \<and> V \<subseteq> span B"
-  let ?ths = "\<exists>x. ?P x"
-  let ?d = "CARD('n)"
-  {assume "V \<subseteq> span S"
-    then have ?ths  using sv i by blast }
-  moreover
-  {assume VS: "\<not> V \<subseteq> span S"
-    from VS obtain a where a: "a \<in> V" "a \<notin> span S" by blast
-    from a have aS: "a \<notin> S" by (auto simp add: span_superset)
-    have th0: "insert a S \<subseteq> V" using a sv by blast
-    from independent_insert[of a S]  i a
-    have th1: "independent (insert a S)" by auto
-    have mlt: "?d - card (insert a S) < n"
-      using aS a n independent_bound[OF th1]
-      by auto
-
-    from H[rule_format, OF mlt th0 th1 refl]
-    obtain B where B: "insert a S \<subseteq> B" "B \<subseteq> V" "independent B" " V \<subseteq> span B"
-      by blast
-    from B have "?P B" by auto
-    then have ?ths by blast}
-  ultimately show ?ths by blast
-qed
-
-lemma maximal_independent_subset:
-  "\<exists>(B:: (real ^'n::finite) set). B\<subseteq> V \<and> independent B \<and> V \<subseteq> span B"
-  by (metis maximal_independent_subset_extend[of "{}:: (real ^'n) set"] empty_subsetI independent_empty)
-
-(* Notion of dimension.                                                      *)
-
-definition "dim V = (SOME n. \<exists>B. B \<subseteq> V \<and> independent B \<and> V \<subseteq> span B \<and> (B hassize n))"
-
-lemma basis_exists:  "\<exists>B. (B :: (real ^'n::finite) set) \<subseteq> V \<and> independent B \<and> V \<subseteq> span B \<and> (B hassize dim V)"
-unfolding dim_def some_eq_ex[of "\<lambda>n. \<exists>B. B \<subseteq> V \<and> independent B \<and> V \<subseteq> span B \<and> (B hassize n)"]
-unfolding hassize_def
-using maximal_independent_subset[of V] independent_bound
-by auto
-
-(* Consequences of independence or spanning for cardinality.                 *)
-
-lemma independent_card_le_dim: "(B::(real ^'n::finite) set) \<subseteq> V \<Longrightarrow> independent B \<Longrightarrow> finite B \<and> card B \<le> dim V"
-by (metis basis_exists[of V] independent_span_bound[where ?'a=real] hassize_def subset_trans)
-
-lemma span_card_ge_dim:  "(B::(real ^'n::finite) set) \<subseteq> V \<Longrightarrow> V \<subseteq> span B \<Longrightarrow> finite B \<Longrightarrow> dim V \<le> card B"
-  by (metis basis_exists[of V] independent_span_bound hassize_def subset_trans)
-
-lemma basis_card_eq_dim:
-  "B \<subseteq> (V:: (real ^'n::finite) set) \<Longrightarrow> V \<subseteq> span B \<Longrightarrow> independent B \<Longrightarrow> finite B \<and> card B = dim V"
-  by (metis order_eq_iff independent_card_le_dim span_card_ge_dim independent_mono)
-
-lemma dim_unique: "(B::(real ^'n::finite) set) \<subseteq> V \<Longrightarrow> V \<subseteq> span B \<Longrightarrow> independent B \<Longrightarrow> B hassize n \<Longrightarrow> dim V = n"
-  by (metis basis_card_eq_dim hassize_def)
-
-(* More lemmas about dimension.                                              *)
-
-lemma dim_univ: "dim (UNIV :: (real^'n::finite) set) = CARD('n)"
-  apply (rule dim_unique[of "{basis i |i. i\<in> (UNIV :: 'n set)}"])
-  by (auto simp only: span_stdbasis has_size_stdbasis independent_stdbasis)
-
-lemma dim_subset:
-  "(S:: (real ^'n::finite) set) \<subseteq> T \<Longrightarrow> dim S \<le> dim T"
-  using basis_exists[of T] basis_exists[of S]
-  by (metis independent_span_bound[where ?'a = real and ?'n = 'n] subset_eq hassize_def)
-
-lemma dim_subset_univ: "dim (S:: (real^'n::finite) set) \<le> CARD('n)"
-  by (metis dim_subset subset_UNIV dim_univ)
-
-(* Converses to those.                                                       *)
-
-lemma card_ge_dim_independent:
-  assumes BV:"(B::(real ^'n::finite) set) \<subseteq> V" and iB:"independent B" and dVB:"dim V \<le> card B"
-  shows "V \<subseteq> span B"
-proof-
-  {fix a assume aV: "a \<in> V"
-    {assume aB: "a \<notin> span B"
-      then have iaB: "independent (insert a B)" using iB aV  BV by (simp add: independent_insert)
-      from aV BV have th0: "insert a B \<subseteq> V" by blast
-      from aB have "a \<notin>B" by (auto simp add: span_superset)
-      with independent_card_le_dim[OF th0 iaB] dVB  have False by auto}
-    then have "a \<in> span B"  by blast}
-  then show ?thesis by blast
-qed
-
-lemma card_le_dim_spanning:
-  assumes BV: "(B:: (real ^'n::finite) set) \<subseteq> V" and VB: "V \<subseteq> span B"
-  and fB: "finite B" and dVB: "dim V \<ge> card B"
-  shows "independent B"
-proof-
-  {fix a assume a: "a \<in> B" "a \<in> span (B -{a})"
-    from a fB have c0: "card B \<noteq> 0" by auto
-    from a fB have cb: "card (B -{a}) = card B - 1" by auto
-    from BV a have th0: "B -{a} \<subseteq> V" by blast
-    {fix x assume x: "x \<in> V"
-      from a have eq: "insert a (B -{a}) = B" by blast
-      from x VB have x': "x \<in> span B" by blast
-      from span_trans[OF a(2), unfolded eq, OF x']
-      have "x \<in> span (B -{a})" . }
-    then have th1: "V \<subseteq> span (B -{a})" by blast
-    have th2: "finite (B -{a})" using fB by auto
-    from span_card_ge_dim[OF th0 th1 th2]
-    have c: "dim V \<le> card (B -{a})" .
-    from c c0 dVB cb have False by simp}
-  then show ?thesis unfolding dependent_def by blast
-qed
-
-lemma card_eq_dim: "(B:: (real ^'n::finite) set) \<subseteq> V \<Longrightarrow> B hassize dim V \<Longrightarrow> independent B \<longleftrightarrow> V \<subseteq> span B"
-  by (metis hassize_def order_eq_iff card_le_dim_spanning
-    card_ge_dim_independent)
-
-(* ------------------------------------------------------------------------- *)
-(* More general size bound lemmas.                                           *)
-(* ------------------------------------------------------------------------- *)
-
-lemma independent_bound_general:
-  "independent (S:: (real^'n::finite) set) \<Longrightarrow> finite S \<and> card S \<le> dim S"
-  by (metis independent_card_le_dim independent_bound subset_refl)
-
-lemma dependent_biggerset_general: "(finite (S:: (real^'n::finite) set) \<Longrightarrow> card S > dim S) \<Longrightarrow> dependent S"
-  using independent_bound_general[of S] by (metis linorder_not_le)
-
-lemma dim_span: "dim (span (S:: (real ^'n::finite) set)) = dim S"
-proof-
-  have th0: "dim S \<le> dim (span S)"
-    by (auto simp add: subset_eq intro: dim_subset span_superset)
-  from basis_exists[of S]
-  obtain B where B: "B \<subseteq> S" "independent B" "S \<subseteq> span B" "B hassize dim S" by blast
-  from B have fB: "finite B" "card B = dim S" unfolding hassize_def by blast+
-  have bSS: "B \<subseteq> span S" using B(1) by (metis subset_eq span_inc)
-  have sssB: "span S \<subseteq> span B" using span_mono[OF B(3)] by (simp add: span_span)
-  from span_card_ge_dim[OF bSS sssB fB(1)] th0 show ?thesis
-    using fB(2)  by arith
-qed
-
-lemma subset_le_dim: "(S:: (real ^'n::finite) set) \<subseteq> span T \<Longrightarrow> dim S \<le> dim T"
-  by (metis dim_span dim_subset)
-
-lemma span_eq_dim: "span (S:: (real ^'n::finite) set) = span T ==> dim S = dim T"
-  by (metis dim_span)
-
-lemma spans_image:
-  assumes lf: "linear (f::'a::semiring_1^'n \<Rightarrow> _)" and VB: "V \<subseteq> span B"
-  shows "f ` V \<subseteq> span (f ` B)"
-  unfolding span_linear_image[OF lf]
-  by (metis VB image_mono)
-
-lemma dim_image_le:
-  fixes f :: "real^'n::finite \<Rightarrow> real^'m::finite"
-  assumes lf: "linear f" shows "dim (f ` S) \<le> dim (S:: (real ^'n::finite) set)"
-proof-
-  from basis_exists[of S] obtain B where
-    B: "B \<subseteq> S" "independent B" "S \<subseteq> span B" "B hassize dim S" by blast
-  from B have fB: "finite B" "card B = dim S" unfolding hassize_def by blast+
-  have "dim (f ` S) \<le> card (f ` B)"
-    apply (rule span_card_ge_dim)
-    using lf B fB by (auto simp add: span_linear_image spans_image subset_image_iff)
-  also have "\<dots> \<le> dim S" using card_image_le[OF fB(1)] fB by simp
-  finally show ?thesis .
-qed
-
-(* Relation between bases and injectivity/surjectivity of map.               *)
-
-lemma spanning_surjective_image:
-  assumes us: "UNIV \<subseteq> span (S:: ('a::semiring_1 ^'n) set)"
-  and lf: "linear f" and sf: "surj f"
-  shows "UNIV \<subseteq> span (f ` S)"
-proof-
-  have "UNIV \<subseteq> f ` UNIV" using sf by (auto simp add: surj_def)
-  also have " \<dots> \<subseteq> span (f ` S)" using spans_image[OF lf us] .
-finally show ?thesis .
-qed
-
-lemma independent_injective_image:
-  assumes iS: "independent (S::('a::semiring_1^'n) set)" and lf: "linear f" and fi: "inj f"
-  shows "independent (f ` S)"
-proof-
-  {fix a assume a: "a \<in> S" "f a \<in> span (f ` S - {f a})"
-    have eq: "f ` S - {f a} = f ` (S - {a})" using fi
-      by (auto simp add: inj_on_def)
-    from a have "f a \<in> f ` span (S -{a})"
-      unfolding eq span_linear_image[OF lf, of "S - {a}"]  by blast
-    hence "a \<in> span (S -{a})" using fi by (auto simp add: inj_on_def)
-    with a(1) iS  have False by (simp add: dependent_def) }
-  then show ?thesis unfolding dependent_def by blast
-qed
-
-(* ------------------------------------------------------------------------- *)
-(* Picking an orthogonal replacement for a spanning set.                     *)
-(* ------------------------------------------------------------------------- *)
-    (* FIXME : Move to some general theory ?*)
-definition "pairwise R S \<longleftrightarrow> (\<forall>x \<in> S. \<forall>y\<in> S. x\<noteq>y \<longrightarrow> R x y)"
-
-lemma vector_sub_project_orthogonal: "(b::'a::ordered_field^'n::finite) \<bullet> (x - ((b \<bullet> x) / (b\<bullet>b)) *s b) = 0"
-  apply (cases "b = 0", simp)
-  apply (simp add: dot_rsub dot_rmult)
-  unfolding times_divide_eq_right[symmetric]
-  by (simp add: field_simps dot_eq_0)
-
-lemma basis_orthogonal:
-  fixes B :: "(real ^'n::finite) set"
-  assumes fB: "finite B"
-  shows "\<exists>C. finite C \<and> card C \<le> card B \<and> span C = span B \<and> pairwise orthogonal C"
-  (is " \<exists>C. ?P B C")
-proof(induct rule: finite_induct[OF fB])
-  case 1 thus ?case apply (rule exI[where x="{}"]) by (auto simp add: pairwise_def)
-next
-  case (2 a B)
-  note fB = `finite B` and aB = `a \<notin> B`
-  from `\<exists>C. finite C \<and> card C \<le> card B \<and> span C = span B \<and> pairwise orthogonal C`
-  obtain C where C: "finite C" "card C \<le> card B"
-    "span C = span B" "pairwise orthogonal C" by blast
-  let ?a = "a - setsum (\<lambda>x. (x\<bullet>a / (x\<bullet>x)) *s x) C"
-  let ?C = "insert ?a C"
-  from C(1) have fC: "finite ?C" by simp
-  from fB aB C(1,2) have cC: "card ?C \<le> card (insert a B)" by (simp add: card_insert_if)
-  {fix x k
-    have th0: "\<And>(a::'b::comm_ring) b c. a - (b - c) = c + (a - b)" by (simp add: ring_simps)
-    have "x - k *s (a - (\<Sum>x\<in>C. (x \<bullet> a / (x \<bullet> x)) *s x)) \<in> span C \<longleftrightarrow> x - k *s a \<in> span C"
-      apply (simp only: vector_ssub_ldistrib th0)
-      apply (rule span_add_eq)
-      apply (rule span_mul)
-      apply (rule span_setsum[OF C(1)])
-      apply clarify
-      apply (rule span_mul)
-      by (rule span_superset)}
-  then have SC: "span ?C = span (insert a B)"
-    unfolding expand_set_eq span_breakdown_eq C(3)[symmetric] by auto
-  thm pairwise_def
-  {fix x y assume xC: "x \<in> ?C" and yC: "y \<in> ?C" and xy: "x \<noteq> y"
-    {assume xa: "x = ?a" and ya: "y = ?a"
-      have "orthogonal x y" using xa ya xy by blast}
-    moreover
-    {assume xa: "x = ?a" and ya: "y \<noteq> ?a" "y \<in> C"
-      from ya have Cy: "C = insert y (C - {y})" by blast
-      have fth: "finite (C - {y})" using C by simp
-      have "orthogonal x y"
-        using xa ya
-        unfolding orthogonal_def xa dot_lsub dot_rsub diff_eq_0_iff_eq
-        apply simp
-        apply (subst Cy)
-        using C(1) fth
-        apply (simp only: setsum_clauses)
-        thm dot_ladd
-        apply (auto simp add: dot_ladd dot_radd dot_lmult dot_rmult dot_eq_0 dot_sym[of y a] dot_lsum[OF fth])
-        apply (rule setsum_0')
-        apply clarsimp
-        apply (rule C(4)[unfolded pairwise_def orthogonal_def, rule_format])
-        by auto}
-    moreover
-    {assume xa: "x \<noteq> ?a" "x \<in> C" and ya: "y = ?a"
-      from xa have Cx: "C = insert x (C - {x})" by blast
-      have fth: "finite (C - {x})" using C by simp
-      have "orthogonal x y"
-        using xa ya
-        unfolding orthogonal_def ya dot_rsub dot_lsub diff_eq_0_iff_eq
-        apply simp
-        apply (subst Cx)
-        using C(1) fth
-        apply (simp only: setsum_clauses)
-        apply (subst dot_sym[of x])
-        apply (auto simp add: dot_radd dot_rmult dot_eq_0 dot_sym[of x a] dot_rsum[OF fth])
-        apply (rule setsum_0')
-        apply clarsimp
-        apply (rule C(4)[unfolded pairwise_def orthogonal_def, rule_format])
-        by auto}
-    moreover
-    {assume xa: "x \<in> C" and ya: "y \<in> C"
-      have "orthogonal x y" using xa ya xy C(4) unfolding pairwise_def by blast}
-    ultimately have "orthogonal x y" using xC yC by blast}
-  then have CPO: "pairwise orthogonal ?C" unfolding pairwise_def by blast
-  from fC cC SC CPO have "?P (insert a B) ?C" by blast
-  then show ?case by blast
-qed
-
-lemma orthogonal_basis_exists:
-  fixes V :: "(real ^'n::finite) set"
-  shows "\<exists>B. independent B \<and> B \<subseteq> span V \<and> V \<subseteq> span B \<and> (B hassize dim V) \<and> pairwise orthogonal B"
-proof-
-  from basis_exists[of V] obtain B where B: "B \<subseteq> V" "independent B" "V \<subseteq> span B" "B hassize dim V" by blast
-  from B have fB: "finite B" "card B = dim V" by (simp_all add: hassize_def)
-  from basis_orthogonal[OF fB(1)] obtain C where
-    C: "finite C" "card C \<le> card B" "span C = span B" "pairwise orthogonal C" by blast
-  from C B
-  have CSV: "C \<subseteq> span V" by (metis span_inc span_mono subset_trans)
-  from span_mono[OF B(3)]  C have SVC: "span V \<subseteq> span C" by (simp add: span_span)
-  from card_le_dim_spanning[OF CSV SVC C(1)] C(2,3) fB
-  have iC: "independent C" by (simp add: dim_span)
-  from C fB have "card C \<le> dim V" by simp
-  moreover have "dim V \<le> card C" using span_card_ge_dim[OF CSV SVC C(1)]
-    by (simp add: dim_span)
-  ultimately have CdV: "C hassize dim V" unfolding hassize_def using C(1) by simp
-  from C B CSV CdV iC show ?thesis by auto
-qed
-
-lemma span_eq: "span S = span T \<longleftrightarrow> S \<subseteq> span T \<and> T \<subseteq> span S"
-  by (metis set_eq_subset span_mono span_span span_inc) (* FIXME: slow *)
-
-(* ------------------------------------------------------------------------- *)
-(* Low-dimensional subset is in a hyperplane (weak orthogonal complement).   *)
-(* ------------------------------------------------------------------------- *)
-
-lemma span_not_univ_orthogonal:
-  assumes sU: "span S \<noteq> UNIV"
-  shows "\<exists>(a:: real ^'n::finite). a \<noteq>0 \<and> (\<forall>x \<in> span S. a \<bullet> x = 0)"
-proof-
-  from sU obtain a where a: "a \<notin> span S" by blast
-  from orthogonal_basis_exists obtain B where
-    B: "independent B" "B \<subseteq> span S" "S \<subseteq> span B" "B hassize dim S" "pairwise orthogonal B"
-    by blast
-  from B have fB: "finite B" "card B = dim S" by (simp_all add: hassize_def)
-  from span_mono[OF B(2)] span_mono[OF B(3)]
-  have sSB: "span S = span B" by (simp add: span_span)
-  let ?a = "a - setsum (\<lambda>b. (a\<bullet>b / (b\<bullet>b)) *s b) B"
-  have "setsum (\<lambda>b. (a\<bullet>b / (b\<bullet>b)) *s b) B \<in> span S"
-    unfolding sSB
-    apply (rule span_setsum[OF fB(1)])
-    apply clarsimp
-    apply (rule span_mul)
-    by (rule span_superset)
-  with a have a0:"?a  \<noteq> 0" by auto
-  have "\<forall>x\<in>span B. ?a \<bullet> x = 0"
-  proof(rule span_induct')
-    show "subspace (\<lambda>x. ?a \<bullet> x = 0)"
-      by (auto simp add: subspace_def mem_def dot_radd dot_rmult)
-  next
-    {fix x assume x: "x \<in> B"
-      from x have B': "B = insert x (B - {x})" by blast
-      have fth: "finite (B - {x})" using fB by simp
-      have "?a \<bullet> x = 0"
-        apply (subst B') using fB fth
-        unfolding setsum_clauses(2)[OF fth]
-        apply simp
-        apply (clarsimp simp add: dot_lsub dot_ladd dot_lmult dot_lsum dot_eq_0)
-        apply (rule setsum_0', rule ballI)
-        unfolding dot_sym
-        by (auto simp add: x field_simps dot_eq_0 intro: B(5)[unfolded pairwise_def orthogonal_def, rule_format])}
-    then show "\<forall>x \<in> B. ?a \<bullet> x = 0" by blast
-  qed
-  with a0 show ?thesis unfolding sSB by (auto intro: exI[where x="?a"])
-qed
-
-lemma span_not_univ_subset_hyperplane:
-  assumes SU: "span S \<noteq> (UNIV ::(real^'n::finite) set)"
-  shows "\<exists> a. a \<noteq>0 \<and> span S \<subseteq> {x. a \<bullet> x = 0}"
-  using span_not_univ_orthogonal[OF SU] by auto
-
-lemma lowdim_subset_hyperplane:
-  assumes d: "dim S < CARD('n::finite)"
-  shows "\<exists>(a::real ^'n::finite). a  \<noteq> 0 \<and> span S \<subseteq> {x. a \<bullet> x = 0}"
-proof-
-  {assume "span S = UNIV"
-    hence "dim (span S) = dim (UNIV :: (real ^'n) set)" by simp
-    hence "dim S = CARD('n)" by (simp add: dim_span dim_univ)
-    with d have False by arith}
-  hence th: "span S \<noteq> UNIV" by blast
-  from span_not_univ_subset_hyperplane[OF th] show ?thesis .
-qed
-
-(* We can extend a linear basis-basis injection to the whole set.            *)
-
-lemma linear_indep_image_lemma:
-  assumes lf: "linear f" and fB: "finite B"
-  and ifB: "independent (f ` B)"
-  and fi: "inj_on f B" and xsB: "x \<in> span B"
-  and fx: "f (x::'a::field^'n) = 0"
-  shows "x = 0"
-  using fB ifB fi xsB fx
-proof(induct arbitrary: x rule: finite_induct[OF fB])
-  case 1 thus ?case by (auto simp add:  span_empty)
-next
-  case (2 a b x)
-  have fb: "finite b" using "2.prems" by simp
-  have th0: "f ` b \<subseteq> f ` (insert a b)"
-    apply (rule image_mono) by blast
-  from independent_mono[ OF "2.prems"(2) th0]
-  have ifb: "independent (f ` b)"  .
-  have fib: "inj_on f b"
-    apply (rule subset_inj_on [OF "2.prems"(3)])
-    by blast
-  from span_breakdown[of a "insert a b", simplified, OF "2.prems"(4)]
-  obtain k where k: "x - k*s a \<in> span (b -{a})" by blast
-  have "f (x - k*s a) \<in> span (f ` b)"
-    unfolding span_linear_image[OF lf]
-    apply (rule imageI)
-    using k span_mono[of "b-{a}" b] by blast
-  hence "f x - k*s f a \<in> span (f ` b)"
-    by (simp add: linear_sub[OF lf] linear_cmul[OF lf])
-  hence th: "-k *s f a \<in> span (f ` b)"
-    using "2.prems"(5) by (simp add: vector_smult_lneg)
-  {assume k0: "k = 0"
-    from k0 k have "x \<in> span (b -{a})" by simp
-    then have "x \<in> span b" using span_mono[of "b-{a}" b]
-      by blast}
-  moreover
-  {assume k0: "k \<noteq> 0"
-    from span_mul[OF th, of "- 1/ k"] k0
-    have th1: "f a \<in> span (f ` b)"
-      by (auto simp add: vector_smult_assoc)
-    from inj_on_image_set_diff[OF "2.prems"(3), of "insert a b " "{a}", symmetric]
-    have tha: "f ` insert a b - f ` {a} = f ` (insert a b - {a})" by blast
-    from "2.prems"(2)[unfolded dependent_def bex_simps(10), rule_format, of "f a"]
-    have "f a \<notin> span (f ` b)" using tha
-      using "2.hyps"(2)
-      "2.prems"(3) by auto
-    with th1 have False by blast
-    then have "x \<in> span b" by blast}
-  ultimately have xsb: "x \<in> span b" by blast
-  from "2.hyps"(3)[OF fb ifb fib xsb "2.prems"(5)]
-  show "x = 0" .
-qed
-
-(* We can extend a linear mapping from basis.                                *)
-
-lemma linear_independent_extend_lemma:
-  assumes fi: "finite B" and ib: "independent B"
-  shows "\<exists>g. (\<forall>x\<in> span B. \<forall>y\<in> span B. g ((x::'a::field^'n) + y) = g x + g y)
-           \<and> (\<forall>x\<in> span B. \<forall>c. g (c*s x) = c *s g x)
-           \<and> (\<forall>x\<in> B. g x = f x)"
-using ib fi
-proof(induct rule: finite_induct[OF fi])
-  case 1 thus ?case by (auto simp add: span_empty)
-next
-  case (2 a b)
-  from "2.prems" "2.hyps" have ibf: "independent b" "finite b"
-    by (simp_all add: independent_insert)
-  from "2.hyps"(3)[OF ibf] obtain g where
-    g: "\<forall>x\<in>span b. \<forall>y\<in>span b. g (x + y) = g x + g y"
-    "\<forall>x\<in>span b. \<forall>c. g (c *s x) = c *s g x" "\<forall>x\<in>b. g x = f x" by blast
-  let ?h = "\<lambda>z. SOME k. (z - k *s a) \<in> span b"
-  {fix z assume z: "z \<in> span (insert a b)"
-    have th0: "z - ?h z *s a \<in> span b"
-      apply (rule someI_ex)
-      unfolding span_breakdown_eq[symmetric]
-      using z .
-    {fix k assume k: "z - k *s a \<in> span b"
-      have eq: "z - ?h z *s a - (z - k*s a) = (k - ?h z) *s a"
-        by (simp add: ring_simps vector_sadd_rdistrib[symmetric])
-      from span_sub[OF th0 k]
-      have khz: "(k - ?h z) *s a \<in> span b" by (simp add: eq)
-      {assume "k \<noteq> ?h z" hence k0: "k - ?h z \<noteq> 0" by simp
-        from k0 span_mul[OF khz, of "1 /(k - ?h z)"]
-        have "a \<in> span b" by (simp add: vector_smult_assoc)
-        with "2.prems"(1) "2.hyps"(2) have False
-          by (auto simp add: dependent_def)}
-      then have "k = ?h z" by blast}
-    with th0 have "z - ?h z *s a \<in> span b \<and> (\<forall>k. z - k *s a \<in> span b \<longrightarrow> k = ?h z)" by blast}
-  note h = this
-  let ?g = "\<lambda>z. ?h z *s f a + g (z - ?h z *s a)"
-  {fix x y assume x: "x \<in> span (insert a b)" and y: "y \<in> span (insert a b)"
-    have tha: "\<And>(x::'a^'n) y a k l. (x + y) - (k + l) *s a = (x - k *s a) + (y - l *s a)"
-      by (vector ring_simps)
-    have addh: "?h (x + y) = ?h x + ?h y"
-      apply (rule conjunct2[OF h, rule_format, symmetric])
-      apply (rule span_add[OF x y])
-      unfolding tha
-      by (metis span_add x y conjunct1[OF h, rule_format])
-    have "?g (x + y) = ?g x + ?g y"
-      unfolding addh tha
-      g(1)[rule_format,OF conjunct1[OF h, OF x] conjunct1[OF h, OF y]]
-      by (simp add: vector_sadd_rdistrib)}
-  moreover
-  {fix x:: "'a^'n" and c:: 'a  assume x: "x \<in> span (insert a b)"
-    have tha: "\<And>(x::'a^'n) c k a. c *s x - (c * k) *s a = c *s (x - k *s a)"
-      by (vector ring_simps)
-    have hc: "?h (c *s x) = c * ?h x"
-      apply (rule conjunct2[OF h, rule_format, symmetric])
-      apply (metis span_mul x)
-      by (metis tha span_mul x conjunct1[OF h])
-    have "?g (c *s x) = c*s ?g x"
-      unfolding hc tha g(2)[rule_format, OF conjunct1[OF h, OF x]]
-      by (vector ring_simps)}
-  moreover
-  {fix x assume x: "x \<in> (insert a b)"
-    {assume xa: "x = a"
-      have ha1: "1 = ?h a"
-        apply (rule conjunct2[OF h, rule_format])
-        apply (metis span_superset insertI1)
-        using conjunct1[OF h, OF span_superset, OF insertI1]
-        by (auto simp add: span_0)
-
-      from xa ha1[symmetric] have "?g x = f x"
-        apply simp
-        using g(2)[rule_format, OF span_0, of 0]
-        by simp}
-    moreover
-    {assume xb: "x \<in> b"
-      have h0: "0 = ?h x"
-        apply (rule conjunct2[OF h, rule_format])
-        apply (metis  span_superset insertI1 xb x)
-        apply simp
-        apply (metis span_superset xb)
-        done
-      have "?g x = f x"
-        by (simp add: h0[symmetric] g(3)[rule_format, OF xb])}
-    ultimately have "?g x = f x" using x by blast }
-  ultimately show ?case apply - apply (rule exI[where x="?g"]) by blast
-qed
-
-lemma linear_independent_extend:
-  assumes iB: "independent (B:: (real ^'n::finite) set)"
-  shows "\<exists>g. linear g \<and> (\<forall>x\<in>B. g x = f x)"
-proof-
-  from maximal_independent_subset_extend[of B UNIV] iB
-  obtain C where C: "B \<subseteq> C" "independent C" "\<And>x. x \<in> span C" by auto
-
-  from C(2) independent_bound[of C] linear_independent_extend_lemma[of C f]
-  obtain g where g: "(\<forall>x\<in> span C. \<forall>y\<in> span C. g (x + y) = g x + g y)
-           \<and> (\<forall>x\<in> span C. \<forall>c. g (c*s x) = c *s g x)
-           \<and> (\<forall>x\<in> C. g x = f x)" by blast
-  from g show ?thesis unfolding linear_def using C
-    apply clarsimp by blast
-qed
-
-(* Can construct an isomorphism between spaces of same dimension.            *)
-
-lemma card_le_inj: assumes fA: "finite A" and fB: "finite B"
-  and c: "card A \<le> card B" shows "(\<exists>f. f ` A \<subseteq> B \<and> inj_on f A)"
-using fB c
-proof(induct arbitrary: B rule: finite_induct[OF fA])
-  case 1 thus ?case by simp
-next
-  case (2 x s t)
-  thus ?case
-  proof(induct rule: finite_induct[OF "2.prems"(1)])
-    case 1    then show ?case by simp
-  next
-    case (2 y t)
-    from "2.prems"(1,2,5) "2.hyps"(1,2) have cst:"card s \<le> card t" by simp
-    from "2.prems"(3) [OF "2.hyps"(1) cst] obtain f where
-      f: "f ` s \<subseteq> t \<and> inj_on f s" by blast
-    from f "2.prems"(2) "2.hyps"(2) show ?case
-      apply -
-      apply (rule exI[where x = "\<lambda>z. if z = x then y else f z"])
-      by (auto simp add: inj_on_def)
-  qed
-qed
-
-lemma card_subset_eq: assumes fB: "finite B" and AB: "A \<subseteq> B" and
-  c: "card A = card B"
-  shows "A = B"
-proof-
-  from fB AB have fA: "finite A" by (auto intro: finite_subset)
-  from fA fB have fBA: "finite (B - A)" by auto
-  have e: "A \<inter> (B - A) = {}" by blast
-  have eq: "A \<union> (B - A) = B" using AB by blast
-  from card_Un_disjoint[OF fA fBA e, unfolded eq c]
-  have "card (B - A) = 0" by arith
-  hence "B - A = {}" unfolding card_eq_0_iff using fA fB by simp
-  with AB show "A = B" by blast
-qed
-
-lemma subspace_isomorphism:
-  assumes s: "subspace (S:: (real ^'n::finite) set)"
-  and t: "subspace (T :: (real ^ 'm::finite) set)"
-  and d: "dim S = dim T"
-  shows "\<exists>f. linear f \<and> f ` S = T \<and> inj_on f S"
-proof-
-  from basis_exists[of S] obtain B where
-    B: "B \<subseteq> S" "independent B" "S \<subseteq> span B" "B hassize dim S" by blast
-  from basis_exists[of T] obtain C where
-    C: "C \<subseteq> T" "independent C" "T \<subseteq> span C" "C hassize dim T" by blast
-  from B(4) C(4) card_le_inj[of B C] d obtain f where
-    f: "f ` B \<subseteq> C" "inj_on f B" unfolding hassize_def by auto
-  from linear_independent_extend[OF B(2)] obtain g where
-    g: "linear g" "\<forall>x\<in> B. g x = f x" by blast
-  from B(4) have fB: "finite B" by (simp add: hassize_def)
-  from C(4) have fC: "finite C" by (simp add: hassize_def)
-  from inj_on_iff_eq_card[OF fB, of f] f(2)
-  have "card (f ` B) = card B" by simp
-  with B(4) C(4) have ceq: "card (f ` B) = card C" using d
-    by (simp add: hassize_def)
-  have "g ` B = f ` B" using g(2)
-    by (auto simp add: image_iff)
-  also have "\<dots> = C" using card_subset_eq[OF fC f(1) ceq] .
-  finally have gBC: "g ` B = C" .
-  have gi: "inj_on g B" using f(2) g(2)
-    by (auto simp add: inj_on_def)
-  note g0 = linear_indep_image_lemma[OF g(1) fB, unfolded gBC, OF C(2) gi]
-  {fix x y assume x: "x \<in> S" and y: "y \<in> S" and gxy:"g x = g y"
-    from B(3) x y have x': "x \<in> span B" and y': "y \<in> span B" by blast+
-    from gxy have th0: "g (x - y) = 0" by (simp add: linear_sub[OF g(1)])
-    have th1: "x - y \<in> span B" using x' y' by (metis span_sub)
-    have "x=y" using g0[OF th1 th0] by simp }
-  then have giS: "inj_on g S"
-    unfolding inj_on_def by blast
-  from span_subspace[OF B(1,3) s]
-  have "g ` S = span (g ` B)" by (simp add: span_linear_image[OF g(1)])
-  also have "\<dots> = span C" unfolding gBC ..
-  also have "\<dots> = T" using span_subspace[OF C(1,3) t] .
-  finally have gS: "g ` S = T" .
-  from g(1) gS giS show ?thesis by blast
-qed
-
-(* linear functions are equal on a subspace if they are on a spanning set.   *)
-
-lemma subspace_kernel:
-  assumes lf: "linear (f::'a::semiring_1 ^'n \<Rightarrow> _)"
-  shows "subspace {x. f x = 0}"
-apply (simp add: subspace_def)
-by (simp add: linear_add[OF lf] linear_cmul[OF lf] linear_0[OF lf])
-
-lemma linear_eq_0_span:
-  assumes lf: "linear f" and f0: "\<forall>x\<in>B. f x = 0"
-  shows "\<forall>x \<in> span B. f x = (0::'a::semiring_1 ^'n)"
-proof
-  fix x assume x: "x \<in> span B"
-  let ?P = "\<lambda>x. f x = 0"
-  from subspace_kernel[OF lf] have "subspace ?P" unfolding Collect_def .
-  with x f0 span_induct[of B "?P" x] show "f x = 0" by blast
-qed
-
-lemma linear_eq_0:
-  assumes lf: "linear f" and SB: "S \<subseteq> span B" and f0: "\<forall>x\<in>B. f x = 0"
-  shows "\<forall>x \<in> S. f x = (0::'a::semiring_1^'n)"
-  by (metis linear_eq_0_span[OF lf] subset_eq SB f0)
-
-lemma linear_eq:
-  assumes lf: "linear (f::'a::ring_1^'n \<Rightarrow> _)" and lg: "linear g" and S: "S \<subseteq> span B"
-  and fg: "\<forall> x\<in> B. f x = g x"
-  shows "\<forall>x\<in> S. f x = g x"
-proof-
-  let ?h = "\<lambda>x. f x - g x"
-  from fg have fg': "\<forall>x\<in> B. ?h x = 0" by simp
-  from linear_eq_0[OF linear_compose_sub[OF lf lg] S fg']
-  show ?thesis by simp
-qed
-
-lemma linear_eq_stdbasis:
-  assumes lf: "linear (f::'a::ring_1^'m::finite \<Rightarrow> 'a^'n::finite)" and lg: "linear g"
-  and fg: "\<forall>i. f (basis i) = g(basis i)"
-  shows "f = g"
-proof-
-  let ?U = "UNIV :: 'm set"
-  let ?I = "{basis i:: 'a^'m|i. i \<in> ?U}"
-  {fix x assume x: "x \<in> (UNIV :: ('a^'m) set)"
-    from equalityD2[OF span_stdbasis]
-    have IU: " (UNIV :: ('a^'m) set) \<subseteq> span ?I" by blast
-    from linear_eq[OF lf lg IU] fg x
-    have "f x = g x" unfolding Collect_def  Ball_def mem_def by metis}
-  then show ?thesis by (auto intro: ext)
-qed
-
-(* Similar results for bilinear functions.                                   *)
-
-lemma bilinear_eq:
-  assumes bf: "bilinear (f:: 'a::ring^'m \<Rightarrow> 'a^'n \<Rightarrow> 'a^'p)"
-  and bg: "bilinear g"
-  and SB: "S \<subseteq> span B" and TC: "T \<subseteq> span C"
-  and fg: "\<forall>x\<in> B. \<forall>y\<in> C. f x y = g x y"
-  shows "\<forall>x\<in>S. \<forall>y\<in>T. f x y = g x y "
-proof-
-  let ?P = "\<lambda>x. \<forall>y\<in> span C. f x y = g x y"
-  from bf bg have sp: "subspace ?P"
-    unfolding bilinear_def linear_def subspace_def bf bg
-    by(auto simp add: span_0 mem_def bilinear_lzero[OF bf] bilinear_lzero[OF bg] span_add Ball_def intro:  bilinear_ladd[OF bf])
-
-  have "\<forall>x \<in> span B. \<forall>y\<in> span C. f x y = g x y"
-    apply -
-    apply (rule ballI)
-    apply (rule span_induct[of B ?P])
-    defer
-    apply (rule sp)
-    apply assumption
-    apply (clarsimp simp add: Ball_def)
-    apply (rule_tac P="\<lambda>y. f xa y = g xa y" and S=C in span_induct)
-    using fg
-    apply (auto simp add: subspace_def)
-    using bf bg unfolding bilinear_def linear_def
-    by(auto simp add: span_0 mem_def bilinear_rzero[OF bf] bilinear_rzero[OF bg] span_add Ball_def intro:  bilinear_ladd[OF bf])
-  then show ?thesis using SB TC by (auto intro: ext)
-qed
-
-lemma bilinear_eq_stdbasis:
-  assumes bf: "bilinear (f:: 'a::ring_1^'m::finite \<Rightarrow> 'a^'n::finite \<Rightarrow> 'a^'p)"
-  and bg: "bilinear g"
-  and fg: "\<forall>i j. f (basis i) (basis j) = g (basis i) (basis j)"
-  shows "f = g"
-proof-
-  from fg have th: "\<forall>x \<in> {basis i| i. i\<in> (UNIV :: 'm set)}. \<forall>y\<in>  {basis j |j. j \<in> (UNIV :: 'n set)}. f x y = g x y" by blast
-  from bilinear_eq[OF bf bg equalityD2[OF span_stdbasis] equalityD2[OF span_stdbasis] th] show ?thesis by (blast intro: ext)
-qed
-
-(* Detailed theorems about left and right invertibility in general case.     *)
-
-lemma left_invertible_transp:
-  "(\<exists>(B::'a^'n^'m). B ** transp (A::'a^'n^'m) = mat (1::'a::comm_semiring_1)) \<longleftrightarrow> (\<exists>(B::'a^'m^'n). A ** B = mat 1)"
-  by (metis matrix_transp_mul transp_mat transp_transp)
-
-lemma right_invertible_transp:
-  "(\<exists>(B::'a^'n^'m). transp (A::'a^'n^'m) ** B = mat (1::'a::comm_semiring_1)) \<longleftrightarrow> (\<exists>(B::'a^'m^'n). B ** A = mat 1)"
-  by (metis matrix_transp_mul transp_mat transp_transp)
-
-lemma linear_injective_left_inverse:
-  assumes lf: "linear (f::real ^'n::finite \<Rightarrow> real ^'m::finite)" and fi: "inj f"
-  shows "\<exists>g. linear g \<and> g o f = id"
-proof-
-  from linear_independent_extend[OF independent_injective_image, OF independent_stdbasis, OF lf fi]
-  obtain h:: "real ^'m \<Rightarrow> real ^'n" where h: "linear h" " \<forall>x \<in> f ` {basis i|i. i \<in> (UNIV::'n set)}. h x = inv f x" by blast
-  from h(2)
-  have th: "\<forall>i. (h \<circ> f) (basis i) = id (basis i)"
-    using inv_o_cancel[OF fi, unfolded stupid_ext[symmetric] id_def o_def]
-    by auto
-
-  from linear_eq_stdbasis[OF linear_compose[OF lf h(1)] linear_id th]
-  have "h o f = id" .
-  then show ?thesis using h(1) by blast
-qed
-
-lemma linear_surjective_right_inverse:
-  assumes lf: "linear (f:: real ^'m::finite \<Rightarrow> real ^'n::finite)" and sf: "surj f"
-  shows "\<exists>g. linear g \<and> f o g = id"
-proof-
-  from linear_independent_extend[OF independent_stdbasis]
-  obtain h:: "real ^'n \<Rightarrow> real ^'m" where
-    h: "linear h" "\<forall> x\<in> {basis i| i. i\<in> (UNIV :: 'n set)}. h x = inv f x" by blast
-  from h(2)
-  have th: "\<forall>i. (f o h) (basis i) = id (basis i)"
-    using sf
-    apply (auto simp add: surj_iff o_def stupid_ext[symmetric])
-    apply (erule_tac x="basis i" in allE)
-    by auto
-
-  from linear_eq_stdbasis[OF linear_compose[OF h(1) lf] linear_id th]
-  have "f o h = id" .
-  then show ?thesis using h(1) by blast
-qed
-
-lemma matrix_left_invertible_injective:
-"(\<exists>B. (B::real^'m^'n) ** (A::real^'n::finite^'m::finite) = mat 1) \<longleftrightarrow> (\<forall>x y. A *v x = A *v y \<longrightarrow> x = y)"
-proof-
-  {fix B:: "real^'m^'n" and x y assume B: "B ** A = mat 1" and xy: "A *v x = A*v y"
-    from xy have "B*v (A *v x) = B *v (A*v y)" by simp
-    hence "x = y"
-      unfolding matrix_vector_mul_assoc B matrix_vector_mul_lid .}
-  moreover
-  {assume A: "\<forall>x y. A *v x = A *v y \<longrightarrow> x = y"
-    hence i: "inj (op *v A)" unfolding inj_on_def by auto
-    from linear_injective_left_inverse[OF matrix_vector_mul_linear i]
-    obtain g where g: "linear g" "g o op *v A = id" by blast
-    have "matrix g ** A = mat 1"
-      unfolding matrix_eq matrix_vector_mul_lid matrix_vector_mul_assoc[symmetric] matrix_works[OF g(1)]
-      using g(2) by (simp add: o_def id_def stupid_ext)
-    then have "\<exists>B. (B::real ^'m^'n) ** A = mat 1" by blast}
-  ultimately show ?thesis by blast
-qed
-
-lemma matrix_left_invertible_ker:
-  "(\<exists>B. (B::real ^'m::finite^'n::finite) ** (A::real^'n^'m) = mat 1) \<longleftrightarrow> (\<forall>x. A *v x = 0 \<longrightarrow> x = 0)"
-  unfolding matrix_left_invertible_injective
-  using linear_injective_0[OF matrix_vector_mul_linear, of A]
-  by (simp add: inj_on_def)
-
-lemma matrix_right_invertible_surjective:
-"(\<exists>B. (A::real^'n::finite^'m::finite) ** (B::real^'m^'n) = mat 1) \<longleftrightarrow> surj (\<lambda>x. A *v x)"
-proof-
-  {fix B :: "real ^'m^'n"  assume AB: "A ** B = mat 1"
-    {fix x :: "real ^ 'm"
-      have "A *v (B *v x) = x"
-        by (simp add: matrix_vector_mul_lid matrix_vector_mul_assoc AB)}
-    hence "surj (op *v A)" unfolding surj_def by metis }
-  moreover
-  {assume sf: "surj (op *v A)"
-    from linear_surjective_right_inverse[OF matrix_vector_mul_linear sf]
-    obtain g:: "real ^'m \<Rightarrow> real ^'n" where g: "linear g" "op *v A o g = id"
-      by blast
-
-    have "A ** (matrix g) = mat 1"
-      unfolding matrix_eq  matrix_vector_mul_lid
-        matrix_vector_mul_assoc[symmetric] matrix_works[OF g(1)]
-      using g(2) unfolding o_def stupid_ext[symmetric] id_def
-      .
-    hence "\<exists>B. A ** (B::real^'m^'n) = mat 1" by blast
-  }
-  ultimately show ?thesis unfolding surj_def by blast
-qed
-
-lemma matrix_left_invertible_independent_columns:
-  fixes A :: "real^'n::finite^'m::finite"
-  shows "(\<exists>(B::real ^'m^'n). B ** A = mat 1) \<longleftrightarrow> (\<forall>c. setsum (\<lambda>i. c i *s column i A) (UNIV :: 'n set) = 0 \<longrightarrow> (\<forall>i. c i = 0))"
-   (is "?lhs \<longleftrightarrow> ?rhs")
-proof-
-  let ?U = "UNIV :: 'n set"
-  {assume k: "\<forall>x. A *v x = 0 \<longrightarrow> x = 0"
-    {fix c i assume c: "setsum (\<lambda>i. c i *s column i A) ?U = 0"
-      and i: "i \<in> ?U"
-      let ?x = "\<chi> i. c i"
-      have th0:"A *v ?x = 0"
-        using c
-        unfolding matrix_mult_vsum Cart_eq
-        by auto
-      from k[rule_format, OF th0] i
-      have "c i = 0" by (vector Cart_eq)}
-    hence ?rhs by blast}
-  moreover
-  {assume H: ?rhs
-    {fix x assume x: "A *v x = 0"
-      let ?c = "\<lambda>i. ((x$i ):: real)"
-      from H[rule_format, of ?c, unfolded matrix_mult_vsum[symmetric], OF x]
-      have "x = 0" by vector}}
-  ultimately show ?thesis unfolding matrix_left_invertible_ker by blast
-qed
-
-lemma matrix_right_invertible_independent_rows:
-  fixes A :: "real^'n::finite^'m::finite"
-  shows "(\<exists>(B::real^'m^'n). A ** B = mat 1) \<longleftrightarrow> (\<forall>c. setsum (\<lambda>i. c i *s row i A) (UNIV :: 'm set) = 0 \<longrightarrow> (\<forall>i. c i = 0))"
-  unfolding left_invertible_transp[symmetric]
-    matrix_left_invertible_independent_columns
-  by (simp add: column_transp)
-
-lemma matrix_right_invertible_span_columns:
-  "(\<exists>(B::real ^'n::finite^'m::finite). (A::real ^'m^'n) ** B = mat 1) \<longleftrightarrow> span (columns A) = UNIV" (is "?lhs = ?rhs")
-proof-
-  let ?U = "UNIV :: 'm set"
-  have fU: "finite ?U" by simp
-  have lhseq: "?lhs \<longleftrightarrow> (\<forall>y. \<exists>(x::real^'m). setsum (\<lambda>i. (x$i) *s column i A) ?U = y)"
-    unfolding matrix_right_invertible_surjective matrix_mult_vsum surj_def
-    apply (subst eq_commute) ..
-  have rhseq: "?rhs \<longleftrightarrow> (\<forall>x. x \<in> span (columns A))" by blast
-  {assume h: ?lhs
-    {fix x:: "real ^'n"
-        from h[unfolded lhseq, rule_format, of x] obtain y:: "real ^'m"
-          where y: "setsum (\<lambda>i. (y$i) *s column i A) ?U = x" by blast
-        have "x \<in> span (columns A)"
-          unfolding y[symmetric]
-          apply (rule span_setsum[OF fU])
-          apply clarify
-          apply (rule span_mul)
-          apply (rule span_superset)
-          unfolding columns_def
-          by blast}
-    then have ?rhs unfolding rhseq by blast}
-  moreover
-  {assume h:?rhs
-    let ?P = "\<lambda>(y::real ^'n). \<exists>(x::real^'m). setsum (\<lambda>i. (x$i) *s column i A) ?U = y"
-    {fix y have "?P y"
-      proof(rule span_induct_alt[of ?P "columns A"])
-        show "\<exists>x\<Colon>real ^ 'm. setsum (\<lambda>i. (x$i) *s column i A) ?U = 0"
-          apply (rule exI[where x=0])
-          by (simp add: zero_index vector_smult_lzero)
-      next
-        fix c y1 y2 assume y1: "y1 \<in> columns A" and y2: "?P y2"
-        from y1 obtain i where i: "i \<in> ?U" "y1 = column i A"
-          unfolding columns_def by blast
-        from y2 obtain x:: "real ^'m" where
-          x: "setsum (\<lambda>i. (x$i) *s column i A) ?U = y2" by blast
-        let ?x = "(\<chi> j. if j = i then c + (x$i) else (x$j))::real^'m"
-        show "?P (c*s y1 + y2)"
-          proof(rule exI[where x= "?x"], vector, auto simp add: i x[symmetric] cond_value_iff right_distrib cond_application_beta cong del: if_weak_cong)
-            fix j
-            have th: "\<forall>xa \<in> ?U. (if xa = i then (c + (x$i)) * ((column xa A)$j)
-           else (x$xa) * ((column xa A$j))) = (if xa = i then c * ((column i A)$j) else 0) + ((x$xa) * ((column xa A)$j))" using i(1)
-              by (simp add: ring_simps)
-            have "setsum (\<lambda>xa. if xa = i then (c + (x$i)) * ((column xa A)$j)
-           else (x$xa) * ((column xa A$j))) ?U = setsum (\<lambda>xa. (if xa = i then c * ((column i A)$j) else 0) + ((x$xa) * ((column xa A)$j))) ?U"
-              apply (rule setsum_cong[OF refl])
-              using th by blast
-            also have "\<dots> = setsum (\<lambda>xa. if xa = i then c * ((column i A)$j) else 0) ?U + setsum (\<lambda>xa. ((x$xa) * ((column xa A)$j))) ?U"
-              by (simp add: setsum_addf)
-            also have "\<dots> = c * ((column i A)$j) + setsum (\<lambda>xa. ((x$xa) * ((column xa A)$j))) ?U"
-              unfolding setsum_delta[OF fU]
-              using i(1) by simp
-            finally show "setsum (\<lambda>xa. if xa = i then (c + (x$i)) * ((column xa A)$j)
-           else (x$xa) * ((column xa A$j))) ?U = c * ((column i A)$j) + setsum (\<lambda>xa. ((x$xa) * ((column xa A)$j))) ?U" .
-          qed
-        next
-          show "y \<in> span (columns A)" unfolding h by blast
-        qed}
-    then have ?lhs unfolding lhseq ..}
-  ultimately show ?thesis by blast
-qed
-
-lemma matrix_left_invertible_span_rows:
-  "(\<exists>(B::real^'m::finite^'n::finite). B ** (A::real^'n^'m) = mat 1) \<longleftrightarrow> span (rows A) = UNIV"
-  unfolding right_invertible_transp[symmetric]
-  unfolding columns_transp[symmetric]
-  unfolding matrix_right_invertible_span_columns
- ..
-
-(* An injective map real^'n->real^'n is also surjective.                       *)
-
-lemma linear_injective_imp_surjective:
-  assumes lf: "linear (f:: real ^'n::finite \<Rightarrow> real ^'n)" and fi: "inj f"
-  shows "surj f"
-proof-
-  let ?U = "UNIV :: (real ^'n) set"
-  from basis_exists[of ?U] obtain B
-    where B: "B \<subseteq> ?U" "independent B" "?U \<subseteq> span B" "B hassize dim ?U"
-    by blast
-  from B(4) have d: "dim ?U = card B" by (simp add: hassize_def)
-  have th: "?U \<subseteq> span (f ` B)"
-    apply (rule card_ge_dim_independent)
-    apply blast
-    apply (rule independent_injective_image[OF B(2) lf fi])
-    apply (rule order_eq_refl)
-    apply (rule sym)
-    unfolding d
-    apply (rule card_image)
-    apply (rule subset_inj_on[OF fi])
-    by blast
-  from th show ?thesis
-    unfolding span_linear_image[OF lf] surj_def
-    using B(3) by blast
-qed
-
-(* And vice versa.                                                           *)
-
-lemma surjective_iff_injective_gen:
-  assumes fS: "finite S" and fT: "finite T" and c: "card S = card T"
-  and ST: "f ` S \<subseteq> T"
-  shows "(\<forall>y \<in> T. \<exists>x \<in> S. f x = y) \<longleftrightarrow> inj_on f S" (is "?lhs \<longleftrightarrow> ?rhs")
-proof-
-  {assume h: "?lhs"
-    {fix x y assume x: "x \<in> S" and y: "y \<in> S" and f: "f x = f y"
-      from x fS have S0: "card S \<noteq> 0" by auto
-      {assume xy: "x \<noteq> y"
-        have th: "card S \<le> card (f ` (S - {y}))"
-          unfolding c
-          apply (rule card_mono)
-          apply (rule finite_imageI)
-          using fS apply simp
-          using h xy x y f unfolding subset_eq image_iff
-          apply auto
-          apply (case_tac "xa = f x")
-          apply (rule bexI[where x=x])
-          apply auto
-          done
-        also have " \<dots> \<le> card (S -{y})"
-          apply (rule card_image_le)
-          using fS by simp
-        also have "\<dots> \<le> card S - 1" using y fS by simp
-        finally have False  using S0 by arith }
-      then have "x = y" by blast}
-    then have ?rhs unfolding inj_on_def by blast}
-  moreover
-  {assume h: ?rhs
-    have "f ` S = T"
-      apply (rule card_subset_eq[OF fT ST])
-      unfolding card_image[OF h] using c .
-    then have ?lhs by blast}
-  ultimately show ?thesis by blast
-qed
-
-lemma linear_surjective_imp_injective:
-  assumes lf: "linear (f::real ^'n::finite => real ^'n)" and sf: "surj f"
-  shows "inj f"
-proof-
-  let ?U = "UNIV :: (real ^'n) set"
-  from basis_exists[of ?U] obtain B
-    where B: "B \<subseteq> ?U" "independent B" "?U \<subseteq> span B" "B hassize dim ?U"
-    by blast
-  {fix x assume x: "x \<in> span B" and fx: "f x = 0"
-    from B(4) have fB: "finite B" by (simp add: hassize_def)
-    from B(4) have d: "dim ?U = card B" by (simp add: hassize_def)
-    have fBi: "independent (f ` B)"
-      apply (rule card_le_dim_spanning[of "f ` B" ?U])
-      apply blast
-      using sf B(3)
-      unfolding span_linear_image[OF lf] surj_def subset_eq image_iff
-      apply blast
-      using fB apply (blast intro: finite_imageI)
-      unfolding d
-      apply (rule card_image_le)
-      apply (rule fB)
-      done
-    have th0: "dim ?U \<le> card (f ` B)"
-      apply (rule span_card_ge_dim)
-      apply blast
-      unfolding span_linear_image[OF lf]
-      apply (rule subset_trans[where B = "f ` UNIV"])
-      using sf unfolding surj_def apply blast
-      apply (rule image_mono)
-      apply (rule B(3))
-      apply (metis finite_imageI fB)
-      done
-
-    moreover have "card (f ` B) \<le> card B"
-      by (rule card_image_le, rule fB)
-    ultimately have th1: "card B = card (f ` B)" unfolding d by arith
-    have fiB: "inj_on f B"
-      unfolding surjective_iff_injective_gen[OF fB finite_imageI[OF fB] th1 subset_refl, symmetric] by blast
-    from linear_indep_image_lemma[OF lf fB fBi fiB x] fx
-    have "x = 0" by blast}
-  note th = this
-  from th show ?thesis unfolding linear_injective_0[OF lf]
-    using B(3) by blast
-qed
-
-(* Hence either is enough for isomorphism.                                   *)
-
-lemma left_right_inverse_eq:
-  assumes fg: "f o g = id" and gh: "g o h = id"
-  shows "f = h"
-proof-
-  have "f = f o (g o h)" unfolding gh by simp
-  also have "\<dots> = (f o g) o h" by (simp add: o_assoc)
-  finally show "f = h" unfolding fg by simp
-qed
-
-lemma isomorphism_expand:
-  "f o g = id \<and> g o f = id \<longleftrightarrow> (\<forall>x. f(g x) = x) \<and> (\<forall>x. g(f x) = x)"
-  by (simp add: expand_fun_eq o_def id_def)
-
-lemma linear_injective_isomorphism:
-  assumes lf: "linear (f :: real^'n::finite \<Rightarrow> real ^'n)" and fi: "inj f"
-  shows "\<exists>f'. linear f' \<and> (\<forall>x. f' (f x) = x) \<and> (\<forall>x. f (f' x) = x)"
-unfolding isomorphism_expand[symmetric]
-using linear_surjective_right_inverse[OF lf linear_injective_imp_surjective[OF lf fi]] linear_injective_left_inverse[OF lf fi]
-by (metis left_right_inverse_eq)
-
-lemma linear_surjective_isomorphism:
-  assumes lf: "linear (f::real ^'n::finite \<Rightarrow> real ^'n)" and sf: "surj f"
-  shows "\<exists>f'. linear f' \<and> (\<forall>x. f' (f x) = x) \<and> (\<forall>x. f (f' x) = x)"
-unfolding isomorphism_expand[symmetric]
-using linear_surjective_right_inverse[OF lf sf] linear_injective_left_inverse[OF lf linear_surjective_imp_injective[OF lf sf]]
-by (metis left_right_inverse_eq)
-
-(* Left and right inverses are the same for R^N->R^N.                        *)
-
-lemma linear_inverse_left:
-  assumes lf: "linear (f::real ^'n::finite \<Rightarrow> real ^'n)" and lf': "linear f'"
-  shows "f o f' = id \<longleftrightarrow> f' o f = id"
-proof-
-  {fix f f':: "real ^'n \<Rightarrow> real ^'n"
-    assume lf: "linear f" "linear f'" and f: "f o f' = id"
-    from f have sf: "surj f"
-
-      apply (auto simp add: o_def stupid_ext[symmetric] id_def surj_def)
-      by metis
-    from linear_surjective_isomorphism[OF lf(1) sf] lf f
-    have "f' o f = id" unfolding stupid_ext[symmetric] o_def id_def
-      by metis}
-  then show ?thesis using lf lf' by metis
-qed
-
-(* Moreover, a one-sided inverse is automatically linear.                    *)
-
-lemma left_inverse_linear:
-  assumes lf: "linear (f::real ^'n::finite \<Rightarrow> real ^'n)" and gf: "g o f = id"
-  shows "linear g"
-proof-
-  from gf have fi: "inj f" apply (auto simp add: inj_on_def o_def id_def stupid_ext[symmetric])
-    by metis
-  from linear_injective_isomorphism[OF lf fi]
-  obtain h:: "real ^'n \<Rightarrow> real ^'n" where
-    h: "linear h" "\<forall>x. h (f x) = x" "\<forall>x. f (h x) = x" by blast
-  have "h = g" apply (rule ext) using gf h(2,3)
-    apply (simp add: o_def id_def stupid_ext[symmetric])
-    by metis
-  with h(1) show ?thesis by blast
-qed
-
-lemma right_inverse_linear:
-  assumes lf: "linear (f:: real ^'n::finite \<Rightarrow> real ^'n)" and gf: "f o g = id"
-  shows "linear g"
-proof-
-  from gf have fi: "surj f" apply (auto simp add: surj_def o_def id_def stupid_ext[symmetric])
-    by metis
-  from linear_surjective_isomorphism[OF lf fi]
-  obtain h:: "real ^'n \<Rightarrow> real ^'n" where
-    h: "linear h" "\<forall>x. h (f x) = x" "\<forall>x. f (h x) = x" by blast
-  have "h = g" apply (rule ext) using gf h(2,3)
-    apply (simp add: o_def id_def stupid_ext[symmetric])
-    by metis
-  with h(1) show ?thesis by blast
-qed
-
-(* The same result in terms of square matrices.                              *)
-
-lemma matrix_left_right_inverse:
-  fixes A A' :: "real ^'n::finite^'n"
-  shows "A ** A' = mat 1 \<longleftrightarrow> A' ** A = mat 1"
-proof-
-  {fix A A' :: "real ^'n^'n" assume AA': "A ** A' = mat 1"
-    have sA: "surj (op *v A)"
-      unfolding surj_def
-      apply clarify
-      apply (rule_tac x="(A' *v y)" in exI)
-      by (simp add: matrix_vector_mul_assoc AA' matrix_vector_mul_lid)
-    from linear_surjective_isomorphism[OF matrix_vector_mul_linear sA]
-    obtain f' :: "real ^'n \<Rightarrow> real ^'n"
-      where f': "linear f'" "\<forall>x. f' (A *v x) = x" "\<forall>x. A *v f' x = x" by blast
-    have th: "matrix f' ** A = mat 1"
-      by (simp add: matrix_eq matrix_works[OF f'(1)] matrix_vector_mul_assoc[symmetric] matrix_vector_mul_lid f'(2)[rule_format])
-    hence "(matrix f' ** A) ** A' = mat 1 ** A'" by simp
-    hence "matrix f' = A'" by (simp add: matrix_mul_assoc[symmetric] AA' matrix_mul_rid matrix_mul_lid)
-    hence "matrix f' ** A = A' ** A" by simp
-    hence "A' ** A = mat 1" by (simp add: th)}
-  then show ?thesis by blast
-qed
-
-(* Considering an n-element vector as an n-by-1 or 1-by-n matrix.            *)
-
-definition "rowvector v = (\<chi> i j. (v$j))"
-
-definition "columnvector v = (\<chi> i j. (v$i))"
-
-lemma transp_columnvector:
- "transp(columnvector v) = rowvector v"
-  by (simp add: transp_def rowvector_def columnvector_def Cart_eq)
-
-lemma transp_rowvector: "transp(rowvector v) = columnvector v"
-  by (simp add: transp_def columnvector_def rowvector_def Cart_eq)
-
-lemma dot_rowvector_columnvector:
-  "columnvector (A *v v) = A ** columnvector v"
-  by (vector columnvector_def matrix_matrix_mult_def matrix_vector_mult_def)
-
-lemma dot_matrix_product: "(x::'a::semiring_1^'n::finite) \<bullet> y = (((rowvector x ::'a^'n^1) ** (columnvector y :: 'a^1^'n))$1)$1"
-  by (vector matrix_matrix_mult_def rowvector_def columnvector_def dot_def)
-
-lemma dot_matrix_vector_mul:
-  fixes A B :: "real ^'n::finite ^'n" and x y :: "real ^'n"
-  shows "(A *v x) \<bullet> (B *v y) =
-      (((rowvector x :: real^'n^1) ** ((transp A ** B) ** (columnvector y :: real ^1^'n)))$1)$1"
-unfolding dot_matrix_product transp_columnvector[symmetric]
-  dot_rowvector_columnvector matrix_transp_mul matrix_mul_assoc ..
-
-(* Infinity norm.                                                            *)
-
-definition "infnorm (x::real^'n::finite) = rsup {abs(x$i) |i. i\<in> (UNIV :: 'n set)}"
-
-lemma numseg_dimindex_nonempty: "\<exists>i. i \<in> (UNIV :: 'n set)"
-  by auto
-
-lemma infnorm_set_image:
-  "{abs(x$i) |i. i\<in> (UNIV :: 'n set)} =
-  (\<lambda>i. abs(x$i)) ` (UNIV :: 'n set)" by blast
-
-lemma infnorm_set_lemma:
-  shows "finite {abs((x::'a::abs ^'n::finite)$i) |i. i\<in> (UNIV :: 'n set)}"
-  and "{abs(x$i) |i. i\<in> (UNIV :: 'n::finite set)} \<noteq> {}"
-  unfolding infnorm_set_image
-  by (auto intro: finite_imageI)
-
-lemma infnorm_pos_le: "0 \<le> infnorm (x::real^'n::finite)"
-  unfolding infnorm_def
-  unfolding rsup_finite_ge_iff[ OF infnorm_set_lemma]
-  unfolding infnorm_set_image
-  by auto
-
-lemma infnorm_triangle: "infnorm ((x::real^'n::finite) + y) \<le> infnorm x + infnorm y"
-proof-
-  have th: "\<And>x y (z::real). x - y <= z \<longleftrightarrow> x - z <= y" by arith
-  have th1: "\<And>S f. f ` S = { f i| i. i \<in> S}" by blast
-  have th2: "\<And>x (y::real). abs(x + y) - abs(x) <= abs(y)" by arith
-  show ?thesis
-  unfolding infnorm_def
-  unfolding rsup_finite_le_iff[ OF infnorm_set_lemma]
-  apply (subst diff_le_eq[symmetric])
-  unfolding rsup_finite_ge_iff[ OF infnorm_set_lemma]
-  unfolding infnorm_set_image bex_simps
-  apply (subst th)
-  unfolding th1
-  unfolding rsup_finite_ge_iff[ OF infnorm_set_lemma]
-
-  unfolding infnorm_set_image ball_simps bex_simps
-  apply simp
-  apply (metis th2)
-  done
-qed
-
-lemma infnorm_eq_0: "infnorm x = 0 \<longleftrightarrow> (x::real ^'n::finite) = 0"
-proof-
-  have "infnorm x <= 0 \<longleftrightarrow> x = 0"
-    unfolding infnorm_def
-    unfolding rsup_finite_le_iff[OF infnorm_set_lemma]
-    unfolding infnorm_set_image ball_simps
-    by vector
-  then show ?thesis using infnorm_pos_le[of x] by simp
-qed
-
-lemma infnorm_0: "infnorm 0 = 0"
-  by (simp add: infnorm_eq_0)
-
-lemma infnorm_neg: "infnorm (- x) = infnorm x"
-  unfolding infnorm_def
-  apply (rule cong[of "rsup" "rsup"])
-  apply blast
-  apply (rule set_ext)
-  apply auto
-  done
-
-lemma infnorm_sub: "infnorm (x - y) = infnorm (y - x)"
-proof-
-  have "y - x = - (x - y)" by simp
-  then show ?thesis  by (metis infnorm_neg)
-qed
-
-lemma real_abs_sub_infnorm: "\<bar> infnorm x - infnorm y\<bar> \<le> infnorm (x - y)"
-proof-
-  have th: "\<And>(nx::real) n ny. nx <= n + ny \<Longrightarrow> ny <= n + nx ==> \<bar>nx - ny\<bar> <= n"
-    by arith
-  from infnorm_triangle[of "x - y" " y"] infnorm_triangle[of "x - y" "-x"]
-  have ths: "infnorm x \<le> infnorm (x - y) + infnorm y"
-    "infnorm y \<le> infnorm (x - y) + infnorm x"
-    by (simp_all add: ring_simps infnorm_neg diff_def[symmetric])
-  from th[OF ths]  show ?thesis .
-qed
-
-lemma real_abs_infnorm: " \<bar>infnorm x\<bar> = infnorm x"
-  using infnorm_pos_le[of x] by arith
-
-lemma component_le_infnorm:
-  shows "\<bar>x$i\<bar> \<le> infnorm (x::real^'n::finite)"
-proof-
-  let ?U = "UNIV :: 'n set"
-  let ?S = "{\<bar>x$i\<bar> |i. i\<in> ?U}"
-  have fS: "finite ?S" unfolding image_Collect[symmetric]
-    apply (rule finite_imageI) unfolding Collect_def mem_def by simp
-  have S0: "?S \<noteq> {}" by blast
-  have th1: "\<And>S f. f ` S = { f i| i. i \<in> S}" by blast
-  from rsup_finite_in[OF fS S0] rsup_finite_Ub[OF fS S0]
-  show ?thesis unfolding infnorm_def isUb_def setle_def
-    unfolding infnorm_set_image ball_simps by auto
-qed
-
-lemma infnorm_mul_lemma: "infnorm(a *s x) <= \<bar>a\<bar> * infnorm x"
-  apply (subst infnorm_def)
-  unfolding rsup_finite_le_iff[OF infnorm_set_lemma]
-  unfolding infnorm_set_image ball_simps
-  apply (simp add: abs_mult)
-  apply (rule allI)
-  apply (cut_tac component_le_infnorm[of x])
-  apply (rule mult_mono)
-  apply auto
-  done
-
-lemma infnorm_mul: "infnorm(a *s x) = abs a * infnorm x"
-proof-
-  {assume a0: "a = 0" hence ?thesis by (simp add: infnorm_0) }
-  moreover
-  {assume a0: "a \<noteq> 0"
-    from a0 have th: "(1/a) *s (a *s x) = x"
-      by (simp add: vector_smult_assoc)
-    from a0 have ap: "\<bar>a\<bar> > 0" by arith
-    from infnorm_mul_lemma[of "1/a" "a *s x"]
-    have "infnorm x \<le> 1/\<bar>a\<bar> * infnorm (a*s x)"
-      unfolding th by simp
-    with ap have "\<bar>a\<bar> * infnorm x \<le> \<bar>a\<bar> * (1/\<bar>a\<bar> * infnorm (a *s x))" by (simp add: field_simps)
-    then have "\<bar>a\<bar> * infnorm x \<le> infnorm (a*s x)"
-      using ap by (simp add: field_simps)
-    with infnorm_mul_lemma[of a x] have ?thesis by arith }
-  ultimately show ?thesis by blast
-qed
-
-lemma infnorm_pos_lt: "infnorm x > 0 \<longleftrightarrow> x \<noteq> 0"
-  using infnorm_pos_le[of x] infnorm_eq_0[of x] by arith
-
-(* Prove that it differs only up to a bound from Euclidean norm.             *)
-
-lemma infnorm_le_norm: "infnorm x \<le> norm x"
-  unfolding infnorm_def rsup_finite_le_iff[OF infnorm_set_lemma]
-  unfolding infnorm_set_image  ball_simps
-  by (metis component_le_norm)
-lemma card_enum: "card {1 .. n} = n" by auto
-lemma norm_le_infnorm: "norm(x) <= sqrt(real CARD('n)) * infnorm(x::real ^'n::finite)"
-proof-
-  let ?d = "CARD('n)"
-  have "real ?d \<ge> 0" by simp
-  hence d2: "(sqrt (real ?d))^2 = real ?d"
-    by (auto intro: real_sqrt_pow2)
-  have th: "sqrt (real ?d) * infnorm x \<ge> 0"
-    by (simp add: zero_le_mult_iff real_sqrt_ge_0_iff infnorm_pos_le)
-  have th1: "x\<bullet>x \<le> (sqrt (real ?d) * infnorm x)^2"
-    unfolding power_mult_distrib d2
-    apply (subst power2_abs[symmetric])
-    unfolding real_of_nat_def dot_def power2_eq_square[symmetric]
-    apply (subst power2_abs[symmetric])
-    apply (rule setsum_bounded)
-    apply (rule power_mono)
-    unfolding abs_of_nonneg[OF infnorm_pos_le]
-    unfolding infnorm_def  rsup_finite_ge_iff[OF infnorm_set_lemma]
-    unfolding infnorm_set_image bex_simps
-    apply blast
-    by (rule abs_ge_zero)
-  from real_le_lsqrt[OF dot_pos_le th th1]
-  show ?thesis unfolding real_vector_norm_def id_def .
-qed
-
-(* Equality in Cauchy-Schwarz and triangle inequalities.                     *)
-
-lemma norm_cauchy_schwarz_eq: "(x::real ^'n::finite) \<bullet> y = norm x * norm y \<longleftrightarrow> norm x *s y = norm y *s x" (is "?lhs \<longleftrightarrow> ?rhs")
-proof-
-  {assume h: "x = 0"
-    hence ?thesis by simp}
-  moreover
-  {assume h: "y = 0"
-    hence ?thesis by simp}
-  moreover
-  {assume x: "x \<noteq> 0" and y: "y \<noteq> 0"
-    from dot_eq_0[of "norm y *s x - norm x *s y"]
-    have "?rhs \<longleftrightarrow> (norm y * (norm y * norm x * norm x - norm x * (x \<bullet> y)) - norm x * (norm y * (y \<bullet> x) - norm x * norm y * norm y) =  0)"
-      using x y
-      unfolding dot_rsub dot_lsub dot_lmult dot_rmult
-      unfolding norm_pow_2[symmetric] power2_eq_square diff_eq_0_iff_eq apply (simp add: dot_sym)
-      apply (simp add: ring_simps)
-      apply metis
-      done
-    also have "\<dots> \<longleftrightarrow> (2 * norm x * norm y * (norm x * norm y - x \<bullet> y) = 0)" using x y
-      by (simp add: ring_simps dot_sym)
-    also have "\<dots> \<longleftrightarrow> ?lhs" using x y
-      apply simp
-      by metis
-    finally have ?thesis by blast}
-  ultimately show ?thesis by blast
-qed
-
-lemma norm_cauchy_schwarz_abs_eq:
-  fixes x y :: "real ^ 'n::finite"
-  shows "abs(x \<bullet> y) = norm x * norm y \<longleftrightarrow>
-                norm x *s y = norm y *s x \<or> norm(x) *s y = - norm y *s x" (is "?lhs \<longleftrightarrow> ?rhs")
-proof-
-  have th: "\<And>(x::real) a. a \<ge> 0 \<Longrightarrow> abs x = a \<longleftrightarrow> x = a \<or> x = - a" by arith
-  have "?rhs \<longleftrightarrow> norm x *s y = norm y *s x \<or> norm (- x) *s y = norm y *s (- x)"
-    apply simp by vector
-  also have "\<dots> \<longleftrightarrow>(x \<bullet> y = norm x * norm y \<or>
-     (-x) \<bullet> y = norm x * norm y)"
-    unfolding norm_cauchy_schwarz_eq[symmetric]
-    unfolding norm_minus_cancel
-      norm_mul by blast
-  also have "\<dots> \<longleftrightarrow> ?lhs"
-    unfolding th[OF mult_nonneg_nonneg, OF norm_ge_zero[of x] norm_ge_zero[of y]] dot_lneg
-    by arith
-  finally show ?thesis ..
-qed
-
-lemma norm_triangle_eq:
-  fixes x y :: "real ^ 'n::finite"
-  shows "norm(x + y) = norm x + norm y \<longleftrightarrow> norm x *s y = norm y *s x"
-proof-
-  {assume x: "x =0 \<or> y =0"
-    hence ?thesis by (cases "x=0", simp_all)}
-  moreover
-  {assume x: "x \<noteq> 0" and y: "y \<noteq> 0"
-    hence "norm x \<noteq> 0" "norm y \<noteq> 0"
-      by simp_all
-    hence n: "norm x > 0" "norm y > 0"
-      using norm_ge_zero[of x] norm_ge_zero[of y]
-      by arith+
-    have th: "\<And>(a::real) b c. a + b + c \<noteq> 0 ==> (a = b + c \<longleftrightarrow> a^2 = (b + c)^2)" by algebra
-    have "norm(x + y) = norm x + norm y \<longleftrightarrow> norm(x + y)^ 2 = (norm x + norm y) ^2"
-      apply (rule th) using n norm_ge_zero[of "x + y"]
-      by arith
-    also have "\<dots> \<longleftrightarrow> norm x *s y = norm y *s x"
-      unfolding norm_cauchy_schwarz_eq[symmetric]
-      unfolding norm_pow_2 dot_ladd dot_radd
-      by (simp add: norm_pow_2[symmetric] power2_eq_square dot_sym ring_simps)
-    finally have ?thesis .}
-  ultimately show ?thesis by blast
-qed
-
-(* Collinearity.*)
-
-definition "collinear S \<longleftrightarrow> (\<exists>u. \<forall>x \<in> S. \<forall> y \<in> S. \<exists>c. x - y = c *s u)"
-
-lemma collinear_empty:  "collinear {}" by (simp add: collinear_def)
-
-lemma collinear_sing: "collinear {(x::'a::ring_1^'n)}"
-  apply (simp add: collinear_def)
-  apply (rule exI[where x=0])
-  by simp
-
-lemma collinear_2: "collinear {(x::'a::ring_1^'n),y}"
-  apply (simp add: collinear_def)
-  apply (rule exI[where x="x - y"])
-  apply auto
-  apply (rule exI[where x=0], simp)
-  apply (rule exI[where x=1], simp)
-  apply (rule exI[where x="- 1"], simp add: vector_sneg_minus1[symmetric])
-  apply (rule exI[where x=0], simp)
-  done
-
-lemma collinear_lemma: "collinear {(0::real^'n),x,y} \<longleftrightarrow> x = 0 \<or> y = 0 \<or> (\<exists>c. y = c *s x)" (is "?lhs \<longleftrightarrow> ?rhs")
-proof-
-  {assume "x=0 \<or> y = 0" hence ?thesis
-      by (cases "x = 0", simp_all add: collinear_2 insert_commute)}
-  moreover
-  {assume x: "x \<noteq> 0" and y: "y \<noteq> 0"
-    {assume h: "?lhs"
-      then obtain u where u: "\<forall> x\<in> {0,x,y}. \<forall>y\<in> {0,x,y}. \<exists>c. x - y = c *s u" unfolding collinear_def by blast
-      from u[rule_format, of x 0] u[rule_format, of y 0]
-      obtain cx and cy where
-        cx: "x = cx*s u" and cy: "y = cy*s u"
-        by auto
-      from cx x have cx0: "cx \<noteq> 0" by auto
-      from cy y have cy0: "cy \<noteq> 0" by auto
-      let ?d = "cy / cx"
-      from cx cy cx0 have "y = ?d *s x"
-        by (simp add: vector_smult_assoc)
-      hence ?rhs using x y by blast}
-    moreover
-    {assume h: "?rhs"
-      then obtain c where c: "y = c*s x" using x y by blast
-      have ?lhs unfolding collinear_def c
-        apply (rule exI[where x=x])
-        apply auto
-        apply (rule exI[where x="- 1"], simp only: vector_smult_lneg vector_smult_lid)
-        apply (rule exI[where x= "-c"], simp only: vector_smult_lneg)
-        apply (rule exI[where x=1], simp)
-        apply (rule exI[where x="1 - c"], simp add: vector_smult_lneg vector_sub_rdistrib)
-        apply (rule exI[where x="c - 1"], simp add: vector_smult_lneg vector_sub_rdistrib)
-        done}
-    ultimately have ?thesis by blast}
-  ultimately show ?thesis by blast
-qed
-
-lemma norm_cauchy_schwarz_equal:
-  fixes x y :: "real ^ 'n::finite"
-  shows "abs(x \<bullet> y) = norm x * norm y \<longleftrightarrow> collinear {(0::real^'n),x,y}"
-unfolding norm_cauchy_schwarz_abs_eq
-apply (cases "x=0", simp_all add: collinear_2)
-apply (cases "y=0", simp_all add: collinear_2 insert_commute)
-unfolding collinear_lemma
-apply simp
-apply (subgoal_tac "norm x \<noteq> 0")
-apply (subgoal_tac "norm y \<noteq> 0")
-apply (rule iffI)
-apply (cases "norm x *s y = norm y *s x")
-apply (rule exI[where x="(1/norm x) * norm y"])
-apply (drule sym)
-unfolding vector_smult_assoc[symmetric]
-apply (simp add: vector_smult_assoc field_simps)
-apply (rule exI[where x="(1/norm x) * - norm y"])
-apply clarify
-apply (drule sym)
-unfolding vector_smult_assoc[symmetric]
-apply (simp add: vector_smult_assoc field_simps)
-apply (erule exE)
-apply (erule ssubst)
-unfolding vector_smult_assoc
-unfolding norm_mul
-apply (subgoal_tac "norm x * c = \<bar>c\<bar> * norm x \<or> norm x * c = - \<bar>c\<bar> * norm x")
-apply (case_tac "c <= 0", simp add: ring_simps)
-apply (simp add: ring_simps)
-apply (case_tac "c <= 0", simp add: ring_simps)
-apply (simp add: ring_simps)
-apply simp
-apply simp
-done
-
-end
--- a/src/HOL/Library/Finite_Cartesian_Product.thy	Fri Oct 23 14:33:07 2009 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,95 +0,0 @@
-(* Title:      HOL/Library/Finite_Cartesian_Product
-   Author:     Amine Chaieb, University of Cambridge
-*)
-
-header {* Definition of finite Cartesian product types. *}
-
-theory Finite_Cartesian_Product
-imports Main (*FIXME: ATP_Linkup is only needed for metis at a few places. We could dispense of that by changing the proofs.*)
-begin
-
-definition hassize (infixr "hassize" 12) where
-  "(S hassize n) = (finite S \<and> card S = n)"
-
-lemma hassize_image_inj: assumes f: "inj_on f S" and S: "S hassize n"
-  shows "f ` S hassize n"
-  using f S card_image[OF f]
-    by (simp add: hassize_def inj_on_def)
-
-
-subsection {* Finite Cartesian products, with indexing and lambdas. *}
-
-typedef (open Cart)
-  ('a, 'b) "^" (infixl "^" 15)
-    = "UNIV :: ('b \<Rightarrow> 'a) set"
-  morphisms Cart_nth Cart_lambda ..
-
-notation Cart_nth (infixl "$" 90)
-
-notation (xsymbols) Cart_lambda (binder "\<chi>" 10)
-
-lemma stupid_ext: "(\<forall>x. f x = g x) \<longleftrightarrow> (f = g)"
-  apply auto
-  apply (rule ext)
-  apply auto
-  done
-
-lemma Cart_eq: "((x:: 'a ^ 'b) = y) \<longleftrightarrow> (\<forall>i. x$i = y$i)"
-  by (simp add: Cart_nth_inject [symmetric] expand_fun_eq)
-
-lemma Cart_lambda_beta [simp]: "Cart_lambda g $ i = g i"
-  by (simp add: Cart_lambda_inverse)
-
-lemma Cart_lambda_unique:
-  fixes f :: "'a ^ 'b"
-  shows "(\<forall>i. f$i = g i) \<longleftrightarrow> Cart_lambda g = f"
-  by (auto simp add: Cart_eq)
-
-lemma Cart_lambda_eta: "(\<chi> i. (g$i)) = g"
-  by (simp add: Cart_eq)
-
-text{* A non-standard sum to "paste" Cartesian products. *}
-
-definition pastecart :: "'a ^ 'm \<Rightarrow> 'a ^ 'n \<Rightarrow> 'a ^ ('m + 'n)" where
-  "pastecart f g = (\<chi> i. case i of Inl a \<Rightarrow> f$a | Inr b \<Rightarrow> g$b)"
-
-definition fstcart:: "'a ^('m + 'n) \<Rightarrow> 'a ^ 'm" where
-  "fstcart f = (\<chi> i. (f$(Inl i)))"
-
-definition sndcart:: "'a ^('m + 'n) \<Rightarrow> 'a ^ 'n" where
-  "sndcart f = (\<chi> i. (f$(Inr i)))"
-
-lemma nth_pastecart_Inl [simp]: "pastecart f g $ Inl a = f$a"
-  unfolding pastecart_def by simp
-
-lemma nth_pastecart_Inr [simp]: "pastecart f g $ Inr b = g$b"
-  unfolding pastecart_def by simp
-
-lemma nth_fstcart [simp]: "fstcart f $ i = f $ Inl i"
-  unfolding fstcart_def by simp
-
-lemma nth_sndtcart [simp]: "sndcart f $ i = f $ Inr i"
-  unfolding sndcart_def by simp
-
-lemma finite_sum_image: "(UNIV::('a + 'b) set) = range Inl \<union> range Inr"
-by (auto, case_tac x, auto)
-
-lemma fstcart_pastecart: "fstcart (pastecart (x::'a ^'m ) (y:: 'a ^ 'n)) = x"
-  by (simp add: Cart_eq)
-
-lemma sndcart_pastecart: "sndcart (pastecart (x::'a ^'m ) (y:: 'a ^ 'n)) = y"
-  by (simp add: Cart_eq)
-
-lemma pastecart_fst_snd: "pastecart (fstcart z) (sndcart z) = z"
-  by (simp add: Cart_eq pastecart_def fstcart_def sndcart_def split: sum.split)
-
-lemma pastecart_eq: "(x = y) \<longleftrightarrow> (fstcart x = fstcart y) \<and> (sndcart x = sndcart y)"
-  using pastecart_fst_snd[of x] pastecart_fst_snd[of y] by metis
-
-lemma forall_pastecart: "(\<forall>p. P p) \<longleftrightarrow> (\<forall>x y. P (pastecart x y))"
-  by (metis pastecart_fst_snd fstcart_pastecart sndcart_pastecart)
-
-lemma exists_pastecart: "(\<exists>p. P p)  \<longleftrightarrow> (\<exists>x y. P (pastecart x y))"
-  by (metis pastecart_fst_snd fstcart_pastecart sndcart_pastecart)
-
-end
--- a/src/HOL/Library/Library.thy	Fri Oct 23 14:33:07 2009 +0200
+++ b/src/HOL/Library/Library.thy	Fri Oct 23 13:23:18 2009 +0200
@@ -14,9 +14,7 @@
   Commutative_Ring
   Continuity
   ContNotDenum
-  Convex_Euclidean_Space
   Countable
-  Determinants
   Diagonalize
   Efficient_Nat
   Enum
@@ -54,7 +52,6 @@
   RBT
   State_Monad
   Sum_Of_Squares
-  Topology_Euclidean_Space
   Univ_Poly
   While_Combinator
   Word
--- a/src/HOL/Library/Topology_Euclidean_Space.thy	Fri Oct 23 14:33:07 2009 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,6027 +0,0 @@
-(*  Title:      HOL/Library/Topology_Euclidian_Space.thy
-    Author:     Amine Chaieb, University of Cambridge
-    Author:     Robert Himmelmann, TU Muenchen
-*)
-
-header {* Elementary topology in Euclidean space. *}
-
-theory Topology_Euclidean_Space
-imports SEQ Euclidean_Space Product_Vector
-begin
-
-declare fstcart_pastecart[simp] sndcart_pastecart[simp]
-
-subsection{* General notion of a topology *}
-
-definition "istopology L \<longleftrightarrow> {} \<in> L \<and> (\<forall>S \<in>L. \<forall>T \<in>L. S \<inter> T \<in> L) \<and> (\<forall>K. K \<subseteq>L \<longrightarrow> \<Union> K \<in> L)"
-typedef (open) 'a topology = "{L::('a set) set. istopology L}"
-  morphisms "openin" "topology"
-  unfolding istopology_def by blast
-
-lemma istopology_open_in[intro]: "istopology(openin U)"
-  using openin[of U] by blast
-
-lemma topology_inverse': "istopology U \<Longrightarrow> openin (topology U) = U"
-  using topology_inverse[unfolded mem_def Collect_def] .
-
-lemma topology_inverse_iff: "istopology U \<longleftrightarrow> openin (topology U) = U"
-  using topology_inverse[of U] istopology_open_in[of "topology U"] by auto
-
-lemma topology_eq: "T1 = T2 \<longleftrightarrow> (\<forall>S. openin T1 S \<longleftrightarrow> openin T2 S)"
-proof-
-  {assume "T1=T2" hence "\<forall>S. openin T1 S \<longleftrightarrow> openin T2 S" by simp}
-  moreover
-  {assume H: "\<forall>S. openin T1 S \<longleftrightarrow> openin T2 S"
-    hence "openin T1 = openin T2" by (metis mem_def set_ext)
-    hence "topology (openin T1) = topology (openin T2)" by simp
-    hence "T1 = T2" unfolding openin_inverse .}
-  ultimately show ?thesis by blast
-qed
-
-text{* Infer the "universe" from union of all sets in the topology. *}
-
-definition "topspace T =  \<Union>{S. openin T S}"
-
-subsection{* Main properties of open sets *}
-
-lemma openin_clauses:
-  fixes U :: "'a topology"
-  shows "openin U {}"
-  "\<And>S T. openin U S \<Longrightarrow> openin U T \<Longrightarrow> openin U (S\<inter>T)"
-  "\<And>K. (\<forall>S \<in> K. openin U S) \<Longrightarrow> openin U (\<Union>K)"
-  using openin[of U] unfolding istopology_def Collect_def mem_def
-  by (metis mem_def subset_eq)+
-
-lemma openin_subset[intro]: "openin U S \<Longrightarrow> S \<subseteq> topspace U"
-  unfolding topspace_def by blast
-lemma openin_empty[simp]: "openin U {}" by (simp add: openin_clauses)
-
-lemma openin_Int[intro]: "openin U S \<Longrightarrow> openin U T \<Longrightarrow> openin U (S \<inter> T)"
-  by (simp add: openin_clauses)
-
-lemma openin_Union[intro]: "(\<forall>S \<in>K. openin U S) \<Longrightarrow> openin U (\<Union> K)" by (simp add: openin_clauses)
-
-lemma openin_Un[intro]: "openin U S \<Longrightarrow> openin U T \<Longrightarrow> openin U (S \<union> T)"
-  using openin_Union[of "{S,T}" U] by auto
-
-lemma openin_topspace[intro, simp]: "openin U (topspace U)" by (simp add: openin_Union topspace_def)
-
-lemma openin_subopen: "openin U S \<longleftrightarrow> (\<forall>x \<in> S. \<exists>T. openin U T \<and> x \<in> T \<and> T \<subseteq> S)" (is "?lhs \<longleftrightarrow> ?rhs")
-proof-
-  {assume ?lhs then have ?rhs by auto }
-  moreover
-  {assume H: ?rhs
-    then obtain t where t: "\<forall>x\<in>S. openin U (t x) \<and> x \<in> t x \<and> t x \<subseteq> S"
-      unfolding Ball_def ex_simps(6)[symmetric] choice_iff by blast
-    from t have th0: "\<forall>x\<in> t`S. openin U x" by auto
-    have "\<Union> t`S = S" using t by auto
-    with openin_Union[OF th0] have "openin U S" by simp }
-  ultimately show ?thesis by blast
-qed
-
-subsection{* Closed sets *}
-
-definition "closedin U S \<longleftrightarrow> S \<subseteq> topspace U \<and> openin U (topspace U - S)"
-
-lemma closedin_subset: "closedin U S \<Longrightarrow> S \<subseteq> topspace U" by (metis closedin_def)
-lemma closedin_empty[simp]: "closedin U {}" by (simp add: closedin_def)
-lemma closedin_topspace[intro,simp]:
-  "closedin U (topspace U)" by (simp add: closedin_def)
-lemma closedin_Un[intro]: "closedin U S \<Longrightarrow> closedin U T \<Longrightarrow> closedin U (S \<union> T)"
-  by (auto simp add: Diff_Un closedin_def)
-
-lemma Diff_Inter[intro]: "A - \<Inter>S = \<Union> {A - s|s. s\<in>S}" by auto
-lemma closedin_Inter[intro]: assumes Ke: "K \<noteq> {}" and Kc: "\<forall>S \<in>K. closedin U S"
-  shows "closedin U (\<Inter> K)"  using Ke Kc unfolding closedin_def Diff_Inter by auto
-
-lemma closedin_Int[intro]: "closedin U S \<Longrightarrow> closedin U T \<Longrightarrow> closedin U (S \<inter> T)"
-  using closedin_Inter[of "{S,T}" U] by auto
-
-lemma Diff_Diff_Int: "A - (A - B) = A \<inter> B" by blast
-lemma openin_closedin_eq: "openin U S \<longleftrightarrow> S \<subseteq> topspace U \<and> closedin U (topspace U - S)"
-  apply (auto simp add: closedin_def Diff_Diff_Int inf_absorb2)
-  apply (metis openin_subset subset_eq)
-  done
-
-lemma openin_closedin:  "S \<subseteq> topspace U \<Longrightarrow> (openin U S \<longleftrightarrow> closedin U (topspace U - S))"
-  by (simp add: openin_closedin_eq)
-
-lemma openin_diff[intro]: assumes oS: "openin U S" and cT: "closedin U T" shows "openin U (S - T)"
-proof-
-  have "S - T = S \<inter> (topspace U - T)" using openin_subset[of U S]  oS cT
-    by (auto simp add: topspace_def openin_subset)
-  then show ?thesis using oS cT by (auto simp add: closedin_def)
-qed
-
-lemma closedin_diff[intro]: assumes oS: "closedin U S" and cT: "openin U T" shows "closedin U (S - T)"
-proof-
-  have "S - T = S \<inter> (topspace U - T)" using closedin_subset[of U S]  oS cT
-    by (auto simp add: topspace_def )
-  then show ?thesis using oS cT by (auto simp add: openin_closedin_eq)
-qed
-
-subsection{* Subspace topology. *}
-
-definition "subtopology U V = topology {S \<inter> V |S. openin U S}"
-
-lemma istopology_subtopology: "istopology {S \<inter> V |S. openin U S}" (is "istopology ?L")
-proof-
-  have "{} \<in> ?L" by blast
-  {fix A B assume A: "A \<in> ?L" and B: "B \<in> ?L"
-    from A B obtain Sa and Sb where Sa: "openin U Sa" "A = Sa \<inter> V" and Sb: "openin U Sb" "B = Sb \<inter> V" by blast
-    have "A\<inter>B = (Sa \<inter> Sb) \<inter> V" "openin U (Sa \<inter> Sb)"  using Sa Sb by blast+
-    then have "A \<inter> B \<in> ?L" by blast}
-  moreover
-  {fix K assume K: "K \<subseteq> ?L"
-    have th0: "?L = (\<lambda>S. S \<inter> V) ` openin U "
-      apply (rule set_ext)
-      apply (simp add: Ball_def image_iff)
-      by (metis mem_def)
-    from K[unfolded th0 subset_image_iff]
-    obtain Sk where Sk: "Sk \<subseteq> openin U" "K = (\<lambda>S. S \<inter> V) ` Sk" by blast
-    have "\<Union>K = (\<Union>Sk) \<inter> V" using Sk by auto
-    moreover have "openin U (\<Union> Sk)" using Sk by (auto simp add: subset_eq mem_def)
-    ultimately have "\<Union>K \<in> ?L" by blast}
-  ultimately show ?thesis unfolding istopology_def by blast
-qed
-
-lemma openin_subtopology:
-  "openin (subtopology U V) S \<longleftrightarrow> (\<exists> T. (openin U T) \<and> (S = T \<inter> V))"
-  unfolding subtopology_def topology_inverse'[OF istopology_subtopology]
-  by (auto simp add: Collect_def)
-
-lemma topspace_subtopology: "topspace(subtopology U V) = topspace U \<inter> V"
-  by (auto simp add: topspace_def openin_subtopology)
-
-lemma closedin_subtopology:
-  "closedin (subtopology U V) S \<longleftrightarrow> (\<exists>T. closedin U T \<and> S = T \<inter> V)"
-  unfolding closedin_def topspace_subtopology
-  apply (simp add: openin_subtopology)
-  apply (rule iffI)
-  apply clarify
-  apply (rule_tac x="topspace U - T" in exI)
-  by auto
-
-lemma openin_subtopology_refl: "openin (subtopology U V) V \<longleftrightarrow> V \<subseteq> topspace U"
-  unfolding openin_subtopology
-  apply (rule iffI, clarify)
-  apply (frule openin_subset[of U])  apply blast
-  apply (rule exI[where x="topspace U"])
-  by auto
-
-lemma subtopology_superset: assumes UV: "topspace U \<subseteq> V"
-  shows "subtopology U V = U"
-proof-
-  {fix S
-    {fix T assume T: "openin U T" "S = T \<inter> V"
-      from T openin_subset[OF T(1)] UV have eq: "S = T" by blast
-      have "openin U S" unfolding eq using T by blast}
-    moreover
-    {assume S: "openin U S"
-      hence "\<exists>T. openin U T \<and> S = T \<inter> V"
-        using openin_subset[OF S] UV by auto}
-    ultimately have "(\<exists>T. openin U T \<and> S = T \<inter> V) \<longleftrightarrow> openin U S" by blast}
-  then show ?thesis unfolding topology_eq openin_subtopology by blast
-qed
-
-
-lemma subtopology_topspace[simp]: "subtopology U (topspace U) = U"
-  by (simp add: subtopology_superset)
-
-lemma subtopology_UNIV[simp]: "subtopology U UNIV = U"
-  by (simp add: subtopology_superset)
-
-subsection{* The universal Euclidean versions are what we use most of the time *}
-
-definition
-  euclidean :: "'a::topological_space topology" where
-  "euclidean = topology open"
-
-lemma open_openin: "open S \<longleftrightarrow> openin euclidean S"
-  unfolding euclidean_def
-  apply (rule cong[where x=S and y=S])
-  apply (rule topology_inverse[symmetric])
-  apply (auto simp add: istopology_def)
-  by (auto simp add: mem_def subset_eq)
-
-lemma topspace_euclidean: "topspace euclidean = UNIV"
-  apply (simp add: topspace_def)
-  apply (rule set_ext)
-  by (auto simp add: open_openin[symmetric])
-
-lemma topspace_euclidean_subtopology[simp]: "topspace (subtopology euclidean S) = S"
-  by (simp add: topspace_euclidean topspace_subtopology)
-
-lemma closed_closedin: "closed S \<longleftrightarrow> closedin euclidean S"
-  by (simp add: closed_def closedin_def topspace_euclidean open_openin Compl_eq_Diff_UNIV)
-
-lemma open_subopen: "open S \<longleftrightarrow> (\<forall>x\<in>S. \<exists>T. open T \<and> x \<in> T \<and> T \<subseteq> S)"
-  by (simp add: open_openin openin_subopen[symmetric])
-
-subsection{* Open and closed balls. *}
-
-definition
-  ball :: "'a::metric_space \<Rightarrow> real \<Rightarrow> 'a set" where
-  "ball x e = {y. dist x y < e}"
-
-definition
-  cball :: "'a::metric_space \<Rightarrow> real \<Rightarrow> 'a set" where
-  "cball x e = {y. dist x y \<le> e}"
-
-lemma mem_ball[simp]: "y \<in> ball x e \<longleftrightarrow> dist x y < e" by (simp add: ball_def)
-lemma mem_cball[simp]: "y \<in> cball x e \<longleftrightarrow> dist x y \<le> e" by (simp add: cball_def)
-
-lemma mem_ball_0 [simp]:
-  fixes x :: "'a::real_normed_vector"
-  shows "x \<in> ball 0 e \<longleftrightarrow> norm x < e"
-  by (simp add: dist_norm)
-
-lemma mem_cball_0 [simp]:
-  fixes x :: "'a::real_normed_vector"
-  shows "x \<in> cball 0 e \<longleftrightarrow> norm x \<le> e"
-  by (simp add: dist_norm)
-
-lemma centre_in_cball[simp]: "x \<in> cball x e \<longleftrightarrow> 0\<le> e"  by simp
-lemma ball_subset_cball[simp,intro]: "ball x e \<subseteq> cball x e" by (simp add: subset_eq)
-lemma subset_ball[intro]: "d <= e ==> ball x d \<subseteq> ball x e" by (simp add: subset_eq)
-lemma subset_cball[intro]: "d <= e ==> cball x d \<subseteq> cball x e" by (simp add: subset_eq)
-lemma ball_max_Un: "ball a (max r s) = ball a r \<union> ball a s"
-  by (simp add: expand_set_eq) arith
-
-lemma ball_min_Int: "ball a (min r s) = ball a r \<inter> ball a s"
-  by (simp add: expand_set_eq)
-
-subsection{* Topological properties of open balls *}
-
-lemma diff_less_iff: "(a::real) - b > 0 \<longleftrightarrow> a > b"
-  "(a::real) - b < 0 \<longleftrightarrow> a < b"
-  "a - b < c \<longleftrightarrow> a < c +b" "a - b > c \<longleftrightarrow> a > c +b" by arith+
-lemma diff_le_iff: "(a::real) - b \<ge> 0 \<longleftrightarrow> a \<ge> b" "(a::real) - b \<le> 0 \<longleftrightarrow> a \<le> b"
-  "a - b \<le> c \<longleftrightarrow> a \<le> c +b" "a - b \<ge> c \<longleftrightarrow> a \<ge> c +b"  by arith+
-
-lemma open_ball[intro, simp]: "open (ball x e)"
-  unfolding open_dist ball_def Collect_def Ball_def mem_def
-  unfolding dist_commute
-  apply clarify
-  apply (rule_tac x="e - dist xa x" in exI)
-  using dist_triangle_alt[where z=x]
-  apply (clarsimp simp add: diff_less_iff)
-  apply atomize
-  apply (erule_tac x="y" in allE)
-  apply (erule_tac x="xa" in allE)
-  by arith
-
-lemma centre_in_ball[simp]: "x \<in> ball x e \<longleftrightarrow> e > 0" by (metis mem_ball dist_self)
-lemma open_contains_ball: "open S \<longleftrightarrow> (\<forall>x\<in>S. \<exists>e>0. ball x e \<subseteq> S)"
-  unfolding open_dist subset_eq mem_ball Ball_def dist_commute ..
-
-lemma open_contains_ball_eq: "open S \<Longrightarrow> \<forall>x. x\<in>S \<longleftrightarrow> (\<exists>e>0. ball x e \<subseteq> S)"
-  by (metis open_contains_ball subset_eq centre_in_ball)
-
-lemma ball_eq_empty[simp]: "ball x e = {} \<longleftrightarrow> e \<le> 0"
-  unfolding mem_ball expand_set_eq
-  apply (simp add: not_less)
-  by (metis zero_le_dist order_trans dist_self)
-
-lemma ball_empty[intro]: "e \<le> 0 ==> ball x e = {}" by simp
-
-subsection{* Basic "localization" results are handy for connectedness. *}
-
-lemma openin_open: "openin (subtopology euclidean U) S \<longleftrightarrow> (\<exists>T. open T \<and> (S = U \<inter> T))"
-  by (auto simp add: openin_subtopology open_openin[symmetric])
-
-lemma openin_open_Int[intro]: "open S \<Longrightarrow> openin (subtopology euclidean U) (U \<inter> S)"
-  by (auto simp add: openin_open)
-
-lemma open_openin_trans[trans]:
- "open S \<Longrightarrow> open T \<Longrightarrow> T \<subseteq> S \<Longrightarrow> openin (subtopology euclidean S) T"
-  by (metis Int_absorb1  openin_open_Int)
-
-lemma open_subset:  "S \<subseteq> T \<Longrightarrow> open S \<Longrightarrow> openin (subtopology euclidean T) S"
-  by (auto simp add: openin_open)
-
-lemma closedin_closed: "closedin (subtopology euclidean U) S \<longleftrightarrow> (\<exists>T. closed T \<and> S = U \<inter> T)"
-  by (simp add: closedin_subtopology closed_closedin Int_ac)
-
-lemma closedin_closed_Int: "closed S ==> closedin (subtopology euclidean U) (U \<inter> S)"
-  by (metis closedin_closed)
-
-lemma closed_closedin_trans: "closed S \<Longrightarrow> closed T \<Longrightarrow> T \<subseteq> S \<Longrightarrow> closedin (subtopology euclidean S) T"
-  apply (subgoal_tac "S \<inter> T = T" )
-  apply auto
-  apply (frule closedin_closed_Int[of T S])
-  by simp
-
-lemma closed_subset: "S \<subseteq> T \<Longrightarrow> closed S \<Longrightarrow> closedin (subtopology euclidean T) S"
-  by (auto simp add: closedin_closed)
-
-lemma openin_euclidean_subtopology_iff:
-  fixes S U :: "'a::metric_space set"
-  shows "openin (subtopology euclidean U) S
-  \<longleftrightarrow> S \<subseteq> U \<and> (\<forall>x\<in>S. \<exists>e>0. \<forall>x'\<in>U. dist x' x < e \<longrightarrow> x'\<in> S)" (is "?lhs \<longleftrightarrow> ?rhs")
-proof-
-  {assume ?lhs hence ?rhs unfolding openin_subtopology open_openin[symmetric]
-      by (simp add: open_dist) blast}
-  moreover
-  {assume SU: "S \<subseteq> U" and H: "\<And>x. x \<in> S \<Longrightarrow> \<exists>e>0. \<forall>x'\<in>U. dist x' x < e \<longrightarrow> x' \<in> S"
-    from H obtain d where d: "\<And>x . x\<in> S \<Longrightarrow> d x > 0 \<and> (\<forall>x' \<in> U. dist x' x < d x \<longrightarrow> x' \<in> S)"
-      by metis
-    let ?T = "\<Union>{B. \<exists>x\<in>S. B = ball x (d x)}"
-    have oT: "open ?T" by auto
-    { fix x assume "x\<in>S"
-      hence "x \<in> \<Union>{B. \<exists>x\<in>S. B = ball x (d x)}"
-        apply simp apply(rule_tac x="ball x(d x)" in exI) apply auto
-        by (rule d [THEN conjunct1])
-      hence "x\<in> ?T \<inter> U" using SU and `x\<in>S` by auto  }
-    moreover
-    { fix y assume "y\<in>?T"
-      then obtain B where "y\<in>B" "B\<in>{B. \<exists>x\<in>S. B = ball x (d x)}" by auto
-      then obtain x where "x\<in>S" and x:"y \<in> ball x (d x)" by auto
-      assume "y\<in>U"
-      hence "y\<in>S" using d[OF `x\<in>S`] and x by(auto simp add: dist_commute) }
-    ultimately have "S = ?T \<inter> U" by blast
-    with oT have ?lhs unfolding openin_subtopology open_openin[symmetric] by blast}
-  ultimately show ?thesis by blast
-qed
-
-text{* These "transitivity" results are handy too. *}
-
-lemma openin_trans[trans]: "openin (subtopology euclidean T) S \<Longrightarrow> openin (subtopology euclidean U) T
-  \<Longrightarrow> openin (subtopology euclidean U) S"
-  unfolding open_openin openin_open by blast
-
-lemma openin_open_trans: "openin (subtopology euclidean T) S \<Longrightarrow> open T \<Longrightarrow> open S"
-  by (auto simp add: openin_open intro: openin_trans)
-
-lemma closedin_trans[trans]:
- "closedin (subtopology euclidean T) S \<Longrightarrow>
-           closedin (subtopology euclidean U) T
-           ==> closedin (subtopology euclidean U) S"
-  by (auto simp add: closedin_closed closed_closedin closed_Inter Int_assoc)
-
-lemma closedin_closed_trans: "closedin (subtopology euclidean T) S \<Longrightarrow> closed T \<Longrightarrow> closed S"
-  by (auto simp add: closedin_closed intro: closedin_trans)
-
-subsection{* Connectedness *}
-
-definition "connected S \<longleftrightarrow>
-  ~(\<exists>e1 e2. open e1 \<and> open e2 \<and> S \<subseteq> (e1 \<union> e2) \<and> (e1 \<inter> e2 \<inter> S = {})
-  \<and> ~(e1 \<inter> S = {}) \<and> ~(e2 \<inter> S = {}))"
-
-lemma connected_local:
- "connected S \<longleftrightarrow> ~(\<exists>e1 e2.
-                 openin (subtopology euclidean S) e1 \<and>
-                 openin (subtopology euclidean S) e2 \<and>
-                 S \<subseteq> e1 \<union> e2 \<and>
-                 e1 \<inter> e2 = {} \<and>
-                 ~(e1 = {}) \<and>
-                 ~(e2 = {}))"
-unfolding connected_def openin_open by (safe, blast+)
-
-lemma exists_diff: "(\<exists>S. P(UNIV - S)) \<longleftrightarrow> (\<exists>S. P S)" (is "?lhs \<longleftrightarrow> ?rhs")
-proof-
-
-  {assume "?lhs" hence ?rhs by blast }
-  moreover
-  {fix S assume H: "P S"
-    have "S = UNIV - (UNIV - S)" by auto
-    with H have "P (UNIV - (UNIV - S))" by metis }
-  ultimately show ?thesis by metis
-qed
-
-lemma connected_clopen: "connected S \<longleftrightarrow>
-        (\<forall>T. openin (subtopology euclidean S) T \<and>
-            closedin (subtopology euclidean S) T \<longrightarrow> T = {} \<or> T = S)" (is "?lhs \<longleftrightarrow> ?rhs")
-proof-
-  have " \<not> connected S \<longleftrightarrow> (\<exists>e1 e2. open e1 \<and> open (UNIV - e2) \<and> S \<subseteq> e1 \<union> (UNIV - e2) \<and> e1 \<inter> (UNIV - e2) \<inter> S = {} \<and> e1 \<inter> S \<noteq> {} \<and> (UNIV - e2) \<inter> S \<noteq> {})"
-    unfolding connected_def openin_open closedin_closed
-    apply (subst exists_diff) by blast
-  hence th0: "connected S \<longleftrightarrow> \<not> (\<exists>e2 e1. closed e2 \<and> open e1 \<and> S \<subseteq> e1 \<union> (UNIV - e2) \<and> e1 \<inter> (UNIV - e2) \<inter> S = {} \<and> e1 \<inter> S \<noteq> {} \<and> (UNIV - e2) \<inter> S \<noteq> {})"
-    (is " _ \<longleftrightarrow> \<not> (\<exists>e2 e1. ?P e2 e1)") apply (simp add: closed_def Compl_eq_Diff_UNIV) by metis
-
-  have th1: "?rhs \<longleftrightarrow> \<not> (\<exists>t' t. closed t'\<and>t = S\<inter>t' \<and> t\<noteq>{} \<and> t\<noteq>S \<and> (\<exists>t'. open t' \<and> t = S \<inter> t'))"
-    (is "_ \<longleftrightarrow> \<not> (\<exists>t' t. ?Q t' t)")
-    unfolding connected_def openin_open closedin_closed by auto
-  {fix e2
-    {fix e1 have "?P e2 e1 \<longleftrightarrow> (\<exists>t.  closed e2 \<and> t = S\<inter>e2 \<and> open e1 \<and> t = S\<inter>e1 \<and> t\<noteq>{} \<and> t\<noteq>S)"
-        by auto}
-    then have "(\<exists>e1. ?P e2 e1) \<longleftrightarrow> (\<exists>t. ?Q e2 t)" by metis}
-  then have "\<forall>e2. (\<exists>e1. ?P e2 e1) \<longleftrightarrow> (\<exists>t. ?Q e2 t)" by blast
-  then show ?thesis unfolding th0 th1 by simp
-qed
-
-lemma connected_empty[simp, intro]: "connected {}"
-  by (simp add: connected_def)
-
-subsection{* Hausdorff and other separation properties *}
-
-class t0_space =
-  assumes t0_space: "x \<noteq> y \<Longrightarrow> \<exists>U. open U \<and> \<not> (x \<in> U \<longleftrightarrow> y \<in> U)"
-
-class t1_space =
-  assumes t1_space: "x \<noteq> y \<Longrightarrow> \<exists>U V. open U \<and> open V \<and> x \<in> U \<and> y \<notin> U \<and> x \<notin> V \<and> y \<in> V"
-begin
-
-subclass t0_space
-proof
-qed (fast dest: t1_space)
-
-end
-
-text {* T2 spaces are also known as Hausdorff spaces. *}
-
-class t2_space =
-  assumes hausdorff: "x \<noteq> y \<Longrightarrow> \<exists>U V. open U \<and> open V \<and> x \<in> U \<and> y \<in> V \<and> U \<inter> V = {}"
-begin
-
-subclass t1_space
-proof
-qed (fast dest: hausdorff)
-
-end
-
-instance metric_space \<subseteq> t2_space
-proof
-  fix x y :: "'a::metric_space"
-  assume xy: "x \<noteq> y"
-  let ?U = "ball x (dist x y / 2)"
-  let ?V = "ball y (dist x y / 2)"
-  have th0: "\<And>d x y z. (d x z :: real) <= d x y + d y z \<Longrightarrow> d y z = d z y
-               ==> ~(d x y * 2 < d x z \<and> d z y * 2 < d x z)" by arith
-  have "open ?U \<and> open ?V \<and> x \<in> ?U \<and> y \<in> ?V \<and> ?U \<inter> ?V = {}"
-    using dist_pos_lt[OF xy] th0[of dist,OF dist_triangle dist_commute]
-    by (auto simp add: expand_set_eq)
-  then show "\<exists>U V. open U \<and> open V \<and> x \<in> U \<and> y \<in> V \<and> U \<inter> V = {}"
-    by blast
-qed
-
-lemma separation_t2:
-  fixes x y :: "'a::t2_space"
-  shows "x \<noteq> y \<longleftrightarrow> (\<exists>U V. open U \<and> open V \<and> x \<in> U \<and> y \<in> V \<and> U \<inter> V = {})"
-  using hausdorff[of x y] by blast
-
-lemma separation_t1:
-  fixes x y :: "'a::t1_space"
-  shows "x \<noteq> y \<longleftrightarrow> (\<exists>U V. open U \<and> open V \<and> x \<in>U \<and> y\<notin> U \<and> x\<notin>V \<and> y\<in>V)"
-  using t1_space[of x y] by blast
-
-lemma separation_t0:
-  fixes x y :: "'a::t0_space"
-  shows "x \<noteq> y \<longleftrightarrow> (\<exists>U. open U \<and> ~(x\<in>U \<longleftrightarrow> y\<in>U))"
-  using t0_space[of x y] by blast
-
-subsection{* Limit points *}
-
-definition
-  islimpt:: "'a::topological_space \<Rightarrow> 'a set \<Rightarrow> bool"
-    (infixr "islimpt" 60) where
-  "x islimpt S \<longleftrightarrow> (\<forall>T. x\<in>T \<longrightarrow> open T \<longrightarrow> (\<exists>y\<in>S. y\<in>T \<and> y\<noteq>x))"
-
-lemma islimptI:
-  assumes "\<And>T. x \<in> T \<Longrightarrow> open T \<Longrightarrow> \<exists>y\<in>S. y \<in> T \<and> y \<noteq> x"
-  shows "x islimpt S"
-  using assms unfolding islimpt_def by auto
-
-lemma islimptE:
-  assumes "x islimpt S" and "x \<in> T" and "open T"
-  obtains y where "y \<in> S" and "y \<in> T" and "y \<noteq> x"
-  using assms unfolding islimpt_def by auto
-
-lemma islimpt_subset: "x islimpt S \<Longrightarrow> S \<subseteq> T ==> x islimpt T" by (auto simp add: islimpt_def)
-
-lemma islimpt_approachable:
-  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)"
-  unfolding islimpt_def
-  apply auto
-  apply(erule_tac x="ball x e" in allE)
-  apply auto
-  apply(rule_tac x=y in bexI)
-  apply (auto simp add: dist_commute)
-  apply (simp add: open_dist, drule (1) bspec)
-  apply (clarify, drule spec, drule (1) mp, auto)
-  done
-
-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)"
-  unfolding islimpt_approachable
-  using approachable_lt_le[where f="\<lambda>x'. dist x' x" and P="\<lambda>x'. \<not> (x'\<in>S \<and> x'\<noteq>x)"]
-  by metis (* FIXME: VERY slow! *)
-
-class perfect_space =
-  (* FIXME: perfect_space should inherit from topological_space *)
-  assumes islimpt_UNIV [simp, intro]: "(x::'a::metric_space) islimpt UNIV"
-
-lemma perfect_choose_dist:
-  fixes x :: "'a::perfect_space"
-  shows "0 < r \<Longrightarrow> \<exists>a. a \<noteq> x \<and> dist a x < r"
-using islimpt_UNIV [of x]
-by (simp add: islimpt_approachable)
-
-instance real :: perfect_space
-apply default
-apply (rule islimpt_approachable [THEN iffD2])
-apply (clarify, rule_tac x="x + e/2" in bexI)
-apply (auto simp add: dist_norm)
-done
-
-instance "^" :: (perfect_space, finite) perfect_space
-proof
-  fix x :: "'a ^ 'b"
-  {
-    fix e :: real assume "0 < e"
-    def a \<equiv> "x $ undefined"
-    have "a islimpt UNIV" by (rule islimpt_UNIV)
-    with `0 < e` obtain b where "b \<noteq> a" and "dist b a < e"
-      unfolding islimpt_approachable by auto
-    def y \<equiv> "Cart_lambda ((Cart_nth x)(undefined := b))"
-    from `b \<noteq> a` have "y \<noteq> x"
-      unfolding a_def y_def by (simp add: Cart_eq)
-    from `dist b a < e` have "dist y x < e"
-      unfolding dist_vector_def a_def y_def
-      apply simp
-      apply (rule le_less_trans [OF setL2_le_setsum [OF zero_le_dist]])
-      apply (subst setsum_diff1' [where a=undefined], simp, simp, simp)
-      done
-    from `y \<noteq> x` and `dist y x < e`
-    have "\<exists>y\<in>UNIV. y \<noteq> x \<and> dist y x < e" by auto
-  }
-  then show "x islimpt UNIV" unfolding islimpt_approachable by blast
-qed
-
-lemma closed_limpt: "closed S \<longleftrightarrow> (\<forall>x. x islimpt S \<longrightarrow> x \<in> S)"
-  unfolding closed_def
-  apply (subst open_subopen)
-  apply (simp add: islimpt_def subset_eq Compl_eq_Diff_UNIV)
-  by (metis DiffE DiffI UNIV_I insertCI insert_absorb mem_def)
-
-lemma islimpt_EMPTY[simp]: "\<not> x islimpt {}"
-  unfolding islimpt_def by auto
-
-lemma closed_positive_orthant: "closed {x::real^'n::finite. \<forall>i. 0 \<le>x$i}"
-proof-
-  let ?U = "UNIV :: 'n set"
-  let ?O = "{x::real^'n. \<forall>i. x$i\<ge>0}"
-  {fix x:: "real^'n" and i::'n assume H: "\<forall>e>0. \<exists>x'\<in>?O. x' \<noteq> x \<and> dist x' x < e"
-    and xi: "x$i < 0"
-    from xi have th0: "-x$i > 0" by arith
-    from H[rule_format, OF th0] obtain x' where x': "x' \<in>?O" "x' \<noteq> x" "dist x' x < -x $ i" by blast
-      have th:" \<And>b a (x::real). abs x <= b \<Longrightarrow> b <= a ==> ~(a + x < 0)" by arith
-      have th': "\<And>x (y::real). x < 0 \<Longrightarrow> 0 <= y ==> abs x <= abs (y - x)" by arith
-      have th1: "\<bar>x$i\<bar> \<le> \<bar>(x' - x)$i\<bar>" using x'(1) xi
-        apply (simp only: vector_component)
-        by (rule th') auto
-      have th2: "\<bar>dist x x'\<bar> \<ge> \<bar>(x' - x)$i\<bar>" using  component_le_norm[of "x'-x" i]
-        apply (simp add: dist_norm) by norm
-      from th[OF th1 th2] x'(3) have False by (simp add: dist_commute) }
-  then show ?thesis unfolding closed_limpt islimpt_approachable
-    unfolding not_le[symmetric] by blast
-qed
-
-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"
-proof(induct rule: finite_induct[OF fS])
-  case 1 thus ?case apply auto by ferrack
-next
-  case (2 x F)
-  from 2 obtain d where d: "d >0" "\<forall>x\<in>F. x\<noteq>a \<longrightarrow> d \<le> dist a x" by blast
-  {assume "x = a" hence ?case using d by auto  }
-  moreover
-  {assume xa: "x\<noteq>a"
-    let ?d = "min d (dist a x)"
-    have dp: "?d > 0" using xa d(1) using dist_nz by auto
-    from d have d': "\<forall>x\<in>F. x\<noteq>a \<longrightarrow> ?d \<le> dist a x" by auto
-    with dp xa have ?case by(auto intro!: exI[where x="?d"]) }
-  ultimately show ?case by blast
-qed
-
-lemma islimpt_finite:
-  fixes S :: "'a::metric_space set"
-  assumes fS: "finite S" shows "\<not> a islimpt S"
-  unfolding islimpt_approachable
-  using finite_set_avoid[OF fS, of a] by (metis dist_commute  not_le)
-
-lemma islimpt_Un: "x islimpt (S \<union> T) \<longleftrightarrow> x islimpt S \<or> x islimpt T"
-  apply (rule iffI)
-  defer
-  apply (metis Un_upper1 Un_upper2 islimpt_subset)
-  unfolding islimpt_def
-  apply (rule ccontr, clarsimp, rename_tac A B)
-  apply (drule_tac x="A \<inter> B" in spec)
-  apply (auto simp add: open_Int)
-  done
-
-lemma discrete_imp_closed:
-  fixes S :: "'a::metric_space set"
-  assumes e: "0 < e" and d: "\<forall>x \<in> S. \<forall>y \<in> S. dist y x < e \<longrightarrow> y = x"
-  shows "closed S"
-proof-
-  {fix x assume C: "\<forall>e>0. \<exists>x'\<in>S. x' \<noteq> x \<and> dist x' x < e"
-    from e have e2: "e/2 > 0" by arith
-    from C[rule_format, OF e2] obtain y where y: "y \<in> S" "y\<noteq>x" "dist y x < e/2" by blast
-    let ?m = "min (e/2) (dist x y) "
-    from e2 y(2) have mp: "?m > 0" by (simp add: dist_nz[THEN sym])
-    from C[rule_format, OF mp] obtain z where z: "z \<in> S" "z\<noteq>x" "dist z x < ?m" by blast
-    have th: "dist z y < e" using z y
-      by (intro dist_triangle_lt [where z=x], simp)
-    from d[rule_format, OF y(1) z(1) th] y z
-    have False by (auto simp add: dist_commute)}
-  then show ?thesis by (metis islimpt_approachable closed_limpt [where 'a='a])
-qed
-
-subsection{* Interior of a Set *}
-definition "interior S = {x. \<exists>T. open T \<and> x \<in> T \<and> T \<subseteq> S}"
-
-lemma interior_eq: "interior S = S \<longleftrightarrow> open S"
-  apply (simp add: expand_set_eq interior_def)
-  apply (subst (2) open_subopen) by (safe, blast+)
-
-lemma interior_open: "open S ==> (interior S = S)" by (metis interior_eq)
-
-lemma interior_empty[simp]: "interior {} = {}" by (simp add: interior_def)
-
-lemma open_interior[simp, intro]: "open(interior S)"
-  apply (simp add: interior_def)
-  apply (subst open_subopen) by blast
-
-lemma interior_interior[simp]: "interior(interior S) = interior S" by (metis interior_eq open_interior)
-lemma interior_subset: "interior S \<subseteq> S" by (auto simp add: interior_def)
-lemma subset_interior: "S \<subseteq> T ==> (interior S) \<subseteq> (interior T)" by (auto simp add: interior_def)
-lemma interior_maximal: "T \<subseteq> S \<Longrightarrow> open T ==> T \<subseteq> (interior S)" by (auto simp add: interior_def)
-lemma interior_unique: "T \<subseteq> S \<Longrightarrow> open T  \<Longrightarrow> (\<forall>T'. T' \<subseteq> S \<and> open T' \<longrightarrow> T' \<subseteq> T) \<Longrightarrow> interior S = T"
-  by (metis equalityI interior_maximal interior_subset open_interior)
-lemma mem_interior: "x \<in> interior S \<longleftrightarrow> (\<exists>e. 0 < e \<and> ball x e \<subseteq> S)"
-  apply (simp add: interior_def)
-  by (metis open_contains_ball centre_in_ball open_ball subset_trans)
-
-lemma open_subset_interior: "open S ==> S \<subseteq> interior T \<longleftrightarrow> S \<subseteq> T"
-  by (metis interior_maximal interior_subset subset_trans)
-
-lemma interior_inter[simp]: "interior(S \<inter> T) = interior S \<inter> interior T"
-  apply (rule equalityI, simp)
-  apply (metis Int_lower1 Int_lower2 subset_interior)
-  by (metis Int_mono interior_subset open_Int open_interior open_subset_interior)
-
-lemma interior_limit_point [intro]:
-  fixes x :: "'a::perfect_space"
-  assumes x: "x \<in> interior S" shows "x islimpt S"
-proof-
-  from x obtain e where e: "e>0" "\<forall>x'. dist x x' < e \<longrightarrow> x' \<in> S"
-    unfolding mem_interior subset_eq Ball_def mem_ball by blast
-  {
-    fix d::real assume d: "d>0"
-    let ?m = "min d e"
-    have mde2: "0 < ?m" using e(1) d(1) by simp
-    from perfect_choose_dist [OF mde2, of x]
-    obtain y where "y \<noteq> x" and "dist y x < ?m" by blast
-    then have "dist y x < e" "dist y x < d" by simp_all
-    from `dist y x < e` e(2) have "y \<in> S" by (simp add: dist_commute)
-    have "\<exists>x'\<in>S. x'\<noteq> x \<and> dist x' x < d"
-      using `y \<in> S` `y \<noteq> x` `dist y x < d` by fast
-  }
-  then show ?thesis unfolding islimpt_approachable by blast
-qed
-
-lemma interior_closed_Un_empty_interior:
-  assumes cS: "closed S" and iT: "interior T = {}"
-  shows "interior(S \<union> T) = interior S"
-proof
-  show "interior S \<subseteq> interior (S\<union>T)"
-    by (rule subset_interior, blast)
-next
-  show "interior (S \<union> T) \<subseteq> interior S"
-  proof
-    fix x assume "x \<in> interior (S \<union> T)"
-    then obtain R where "open R" "x \<in> R" "R \<subseteq> S \<union> T"
-      unfolding interior_def by fast
-    show "x \<in> interior S"
-    proof (rule ccontr)
-      assume "x \<notin> interior S"
-      with `x \<in> R` `open R` obtain y where "y \<in> R - S"
-        unfolding interior_def expand_set_eq by fast
-      from `open R` `closed S` have "open (R - S)" by (rule open_Diff)
-      from `R \<subseteq> S \<union> T` have "R - S \<subseteq> T" by fast
-      from `y \<in> R - S` `open (R - S)` `R - S \<subseteq> T` `interior T = {}`
-      show "False" unfolding interior_def by fast
-    qed
-  qed
-qed
-
-
-subsection{* Closure of a Set *}
-
-definition "closure S = S \<union> {x | x. x islimpt S}"
-
-lemma closure_interior: "closure S = UNIV - interior (UNIV - S)"
-proof-
-  { fix x
-    have "x\<in>UNIV - interior (UNIV - S) \<longleftrightarrow> x \<in> closure S"  (is "?lhs = ?rhs")
-    proof
-      let ?exT = "\<lambda> y. (\<exists>T. open T \<and> y \<in> T \<and> T \<subseteq> UNIV - S)"
-      assume "?lhs"
-      hence *:"\<not> ?exT x"
-        unfolding interior_def
-        by simp
-      { assume "\<not> ?rhs"
-        hence False using *
-          unfolding closure_def islimpt_def
-          by blast
-      }
-      thus "?rhs"
-        by blast
-    next
-      assume "?rhs" thus "?lhs"
-        unfolding closure_def interior_def islimpt_def
-        by blast
-    qed
-  }
-  thus ?thesis
-    by blast
-qed
-
-lemma interior_closure: "interior S = UNIV - (closure (UNIV - S))"
-proof-
-  { fix x
-    have "x \<in> interior S \<longleftrightarrow> x \<in> UNIV - (closure (UNIV - S))"
-      unfolding interior_def closure_def islimpt_def
-      by blast (* FIXME: VERY slow! *)
-  }
-  thus ?thesis
-    by blast
-qed
-
-lemma closed_closure[simp, intro]: "closed (closure S)"
-proof-
-  have "closed (UNIV - interior (UNIV -S))" by blast
-  thus ?thesis using closure_interior[of S] by simp
-qed
-
-lemma closure_hull: "closure S = closed hull S"
-proof-
-  have "S \<subseteq> closure S"
-    unfolding closure_def
-    by blast
-  moreover
-  have "closed (closure S)"
-    using closed_closure[of S]
-    by assumption
-  moreover
-  { fix t
-    assume *:"S \<subseteq> t" "closed t"
-    { fix x
-      assume "x islimpt S"
-      hence "x islimpt t" using *(1)
-        using islimpt_subset[of x, of S, of t]
-        by blast
-    }
-    with * have "closure S \<subseteq> t"
-      unfolding closure_def
-      using closed_limpt[of t]
-      by auto
-  }
-  ultimately show ?thesis
-    using hull_unique[of S, of "closure S", of closed]
-    unfolding mem_def
-    by simp
-qed
-
-lemma closure_eq: "closure S = S \<longleftrightarrow> closed S"
-  unfolding closure_hull
-  using hull_eq[of closed, unfolded mem_def, OF  closed_Inter, of S]
-  by (metis mem_def subset_eq)
-
-lemma closure_closed[simp]: "closed S \<Longrightarrow> closure S = S"
-  using closure_eq[of S]
-  by simp
-
-lemma closure_closure[simp]: "closure (closure S) = closure S"
-  unfolding closure_hull
-  using hull_hull[of closed S]
-  by assumption
-
-lemma closure_subset: "S \<subseteq> closure S"
-  unfolding closure_hull
-  using hull_subset[of S closed]
-  by assumption
-
-lemma subset_closure: "S \<subseteq> T \<Longrightarrow> closure S \<subseteq> closure T"
-  unfolding closure_hull
-  using hull_mono[of S T closed]
-  by assumption
-
-lemma closure_minimal: "S \<subseteq> T \<Longrightarrow>  closed T \<Longrightarrow> closure S \<subseteq> T"
-  using hull_minimal[of S T closed]
-  unfolding closure_hull mem_def
-  by simp
-
-lemma closure_unique: "S \<subseteq> T \<and> closed T \<and> (\<forall> T'. S \<subseteq> T' \<and> closed T' \<longrightarrow> T \<subseteq> T') \<Longrightarrow> closure S = T"
-  using hull_unique[of S T closed]
-  unfolding closure_hull mem_def
-  by simp
-
-lemma closure_empty[simp]: "closure {} = {}"
-  using closed_empty closure_closed[of "{}"]
-  by simp
-
-lemma closure_univ[simp]: "closure UNIV = UNIV"
-  using closure_closed[of UNIV]
-  by simp
-
-lemma closure_eq_empty: "closure S = {} \<longleftrightarrow> S = {}"
-  using closure_empty closure_subset[of S]
-  by blast
-
-lemma closure_subset_eq: "closure S \<subseteq> S \<longleftrightarrow> closed S"
-  using closure_eq[of S] closure_subset[of S]
-  by simp
-
-lemma open_inter_closure_eq_empty:
-  "open S \<Longrightarrow> (S \<inter> closure T) = {} \<longleftrightarrow> S \<inter> T = {}"
-  using open_subset_interior[of S "UNIV - T"]
-  using interior_subset[of "UNIV - T"]
-  unfolding closure_interior
-  by auto
-
-lemma open_inter_closure_subset:
-  "open S \<Longrightarrow> (S \<inter> (closure T)) \<subseteq> closure(S \<inter> T)"
-proof
-  fix x
-  assume as: "open S" "x \<in> S \<inter> closure T"
-  { assume *:"x islimpt T"
-    have "x islimpt (S \<inter> T)"
-    proof (rule islimptI)
-      fix A
-      assume "x \<in> A" "open A"
-      with as have "x \<in> A \<inter> S" "open (A \<inter> S)"
-        by (simp_all add: open_Int)
-      with * obtain y where "y \<in> T" "y \<in> A \<inter> S" "y \<noteq> x"
-        by (rule islimptE)
-      hence "y \<in> S \<inter> T" "y \<in> A \<and> y \<noteq> x"
-        by simp_all
-      thus "\<exists>y\<in>(S \<inter> T). y \<in> A \<and> y \<noteq> x" ..
-    qed
-  }
-  then show "x \<in> closure (S \<inter> T)" using as
-    unfolding closure_def
-    by blast
-qed
-
-lemma closure_complement: "closure(UNIV - S) = UNIV - interior(S)"
-proof-
-  have "S = UNIV - (UNIV - S)"
-    by auto
-  thus ?thesis
-    unfolding closure_interior
-    by auto
-qed
-
-lemma interior_complement: "interior(UNIV - S) = UNIV - closure(S)"
-  unfolding closure_interior
-  by blast
-
-subsection{* Frontier (aka boundary) *}
-
-definition "frontier S = closure S - interior S"
-
-lemma frontier_closed: "closed(frontier S)"
-  by (simp add: frontier_def closed_Diff)
-
-lemma frontier_closures: "frontier S = (closure S) \<inter> (closure(UNIV - S))"
-  by (auto simp add: frontier_def interior_closure)
-
-lemma frontier_straddle:
-  fixes a :: "'a::metric_space"
-  shows "a \<in> frontier S \<longleftrightarrow> (\<forall>e>0. (\<exists>x\<in>S. dist a x < e) \<and> (\<exists>x. x \<notin> S \<and> dist a x < e))" (is "?lhs \<longleftrightarrow> ?rhs")
-proof
-  assume "?lhs"
-  { fix e::real
-    assume "e > 0"
-    let ?rhse = "(\<exists>x\<in>S. dist a x < e) \<and> (\<exists>x. x \<notin> S \<and> dist a x < e)"
-    { assume "a\<in>S"
-      have "\<exists>x\<in>S. dist a x < e" using `e>0` `a\<in>S` by(rule_tac x=a in bexI) auto
-      moreover have "\<exists>x. x \<notin> S \<and> dist a x < e" using `?lhs` `a\<in>S`
-        unfolding frontier_closures closure_def islimpt_def using `e>0`
-        by (auto, erule_tac x="ball a e" in allE, auto)
-      ultimately have ?rhse by auto
-    }
-    moreover
-    { assume "a\<notin>S"
-      hence ?rhse using `?lhs`
-        unfolding frontier_closures closure_def islimpt_def
-        using open_ball[of a e] `e > 0`
-        by (auto, erule_tac x = "ball a e" in allE, auto) (* FIXME: VERY slow! *)
-    }
-    ultimately have ?rhse by auto
-  }
-  thus ?rhs by auto
-next
-  assume ?rhs
-  moreover
-  { fix T assume "a\<notin>S" and
-    as:"\<forall>e>0. (\<exists>x\<in>S. dist a x < e) \<and> (\<exists>x. x \<notin> S \<and> dist a x < e)" "a \<notin> S" "a \<in> T" "open T"
-    from `open T` `a \<in> T` have "\<exists>e>0. ball a e \<subseteq> T" unfolding open_contains_ball[of T] by auto
-    then obtain e where "e>0" "ball a e \<subseteq> T" by auto
-    then obtain y where y:"y\<in>S" "dist a y < e"  using as(1) by auto
-    have "\<exists>y\<in>S. y \<in> T \<and> y \<noteq> a"
-      using `dist a y < e` `ball a e \<subseteq> T` unfolding ball_def using `y\<in>S` `a\<notin>S` by auto
-  }
-  hence "a \<in> closure S" unfolding closure_def islimpt_def using `?rhs` by auto
-  moreover
-  { fix T assume "a \<in> T"  "open T" "a\<in>S"
-    then obtain e where "e>0" and balle: "ball a e \<subseteq> T" unfolding open_contains_ball using `?rhs` by auto
-    obtain x where "x \<notin> S" "dist a x < e" using `?rhs` using `e>0` by auto
-    hence "\<exists>y\<in>UNIV - S. y \<in> T \<and> y \<noteq> a" using balle `a\<in>S` unfolding ball_def by (rule_tac x=x in bexI)auto
-  }
-  hence "a islimpt (UNIV - S) \<or> a\<notin>S" unfolding islimpt_def by auto
-  ultimately show ?lhs unfolding frontier_closures using closure_def[of "UNIV - S"] by auto
-qed
-
-lemma frontier_subset_closed: "closed S \<Longrightarrow> frontier S \<subseteq> S"
-  by (metis frontier_def closure_closed Diff_subset)
-
-lemma frontier_empty: "frontier {} = {}"
-  by (simp add: frontier_def closure_empty)
-
-lemma frontier_subset_eq: "frontier S \<subseteq> S \<longleftrightarrow> closed S"
-proof-
-  { assume "frontier S \<subseteq> S"
-    hence "closure S \<subseteq> S" using interior_subset unfolding frontier_def by auto
-    hence "closed S" using closure_subset_eq by auto
-  }
-  thus ?thesis using frontier_subset_closed[of S] by auto
-qed
-
-lemma frontier_complement: "frontier(UNIV - S) = frontier S"
-  by (auto simp add: frontier_def closure_complement interior_complement)
-
-lemma frontier_disjoint_eq: "frontier S \<inter> S = {} \<longleftrightarrow> open S"
-  using frontier_complement frontier_subset_eq[of "UNIV - S"]
-  unfolding open_closed Compl_eq_Diff_UNIV by auto
-
-subsection{* Common nets and The "within" modifier for nets. *}
-
-definition
-  at_infinity :: "'a::real_normed_vector net" where
-  "at_infinity = Abs_net (range (\<lambda>r. {x. r \<le> norm x}))"
-
-definition
-  indirection :: "'a::real_normed_vector \<Rightarrow> 'a \<Rightarrow> 'a net" (infixr "indirection" 70) where
-  "a indirection v = (at a) within {b. \<exists>c\<ge>0. b - a = scaleR c v}"
-
-text{* Prove That They are all nets. *}
-
-lemma Rep_net_at_infinity:
-  "Rep_net at_infinity = range (\<lambda>r. {x. r \<le> norm x})"
-unfolding at_infinity_def
-apply (rule Abs_net_inverse')
-apply (rule image_nonempty, simp)
-apply (clarsimp, rename_tac r s)
-apply (rule_tac x="max r s" in exI, auto)
-done
-
-lemma within_UNIV: "net within UNIV = net"
-  by (simp add: Rep_net_inject [symmetric] Rep_net_within)
-
-subsection{* Identify Trivial limits, where we can't approach arbitrarily closely. *}
-
-definition
-  trivial_limit :: "'a net \<Rightarrow> bool" where
-  "trivial_limit net \<longleftrightarrow> {} \<in> Rep_net net"
-
-lemma trivial_limit_within:
-  shows "trivial_limit (at a within S) \<longleftrightarrow> \<not> a islimpt S"
-proof
-  assume "trivial_limit (at a within S)"
-  thus "\<not> a islimpt S"
-    unfolding trivial_limit_def
-    unfolding Rep_net_within Rep_net_at
-    unfolding islimpt_def
-    apply (clarsimp simp add: expand_set_eq)
-    apply (rename_tac T, rule_tac x=T in exI)
-    apply (clarsimp, drule_tac x=y in spec, simp)
-    done
-next
-  assume "\<not> a islimpt S"
-  thus "trivial_limit (at a within S)"
-    unfolding trivial_limit_def
-    unfolding Rep_net_within Rep_net_at
-    unfolding islimpt_def
-    apply (clarsimp simp add: image_image)
-    apply (rule_tac x=T in image_eqI)
-    apply (auto simp add: expand_set_eq)
-    done
-qed
-
-lemma trivial_limit_at_iff: "trivial_limit (at a) \<longleftrightarrow> \<not> a islimpt UNIV"
-  using trivial_limit_within [of a UNIV]
-  by (simp add: within_UNIV)
-
-lemma trivial_limit_at:
-  fixes a :: "'a::perfect_space"
-  shows "\<not> trivial_limit (at a)"
-  by (simp add: trivial_limit_at_iff)
-
-lemma trivial_limit_at_infinity:
-  "\<not> trivial_limit (at_infinity :: ('a::{real_normed_vector,zero_neq_one}) net)"
-  (* FIXME: find a more appropriate type class *)
-  unfolding trivial_limit_def Rep_net_at_infinity
-  apply (clarsimp simp add: expand_set_eq)
-  apply (drule_tac x="scaleR r (sgn 1)" in spec)
-  apply (simp add: norm_sgn)
-  done
-
-lemma trivial_limit_sequentially: "\<not> trivial_limit sequentially"
-  by (auto simp add: trivial_limit_def Rep_net_sequentially)
-
-subsection{* Some property holds "sufficiently close" to the limit point. *}
-
-lemma eventually_at: (* FIXME: this replaces Limits.eventually_at *)
-  "eventually P (at a) \<longleftrightarrow> (\<exists>d>0. \<forall>x. 0 < dist x a \<and> dist x a < d \<longrightarrow> P x)"
-unfolding eventually_at dist_nz by auto
-
-lemma eventually_at_infinity:
-  "eventually P at_infinity \<longleftrightarrow> (\<exists>b. \<forall>x. norm x >= b \<longrightarrow> P x)"
-unfolding eventually_def Rep_net_at_infinity by auto
-
-lemma eventually_within: "eventually P (at a within S) \<longleftrightarrow>
-        (\<exists>d>0. \<forall>x\<in>S. 0 < dist x a \<and> dist x a < d \<longrightarrow> P x)"
-unfolding eventually_within eventually_at dist_nz by auto
-
-lemma eventually_within_le: "eventually P (at a within S) \<longleftrightarrow>
-        (\<exists>d>0. \<forall>x\<in>S. 0 < dist x a \<and> dist x a <= d \<longrightarrow> P x)" (is "?lhs = ?rhs")
-unfolding eventually_within
-apply safe
-apply (rule_tac x="d/2" in exI, simp)
-apply (rule_tac x="d" in exI, simp)
-done
-
-lemma eventually_happens: "eventually P net ==> trivial_limit net \<or> (\<exists>x. P x)"
-  unfolding eventually_def trivial_limit_def
-  using Rep_net_nonempty [of net] by auto
-
-lemma always_eventually: "(\<forall>x. P x) ==> eventually P net"
-  unfolding eventually_def trivial_limit_def
-  using Rep_net_nonempty [of net] by auto
-
-lemma trivial_limit_eventually: "trivial_limit net \<Longrightarrow> eventually P net"
-  unfolding trivial_limit_def eventually_def by auto
-
-lemma eventually_False: "eventually (\<lambda>x. False) net \<longleftrightarrow> trivial_limit net"
-  unfolding trivial_limit_def eventually_def by auto
-
-lemma trivial_limit_eq: "trivial_limit net \<longleftrightarrow> (\<forall>P. eventually P net)"
-  apply (safe elim!: trivial_limit_eventually)
-  apply (simp add: eventually_False [symmetric])
-  done
-
-text{* Combining theorems for "eventually" *}
-
-lemma eventually_conjI:
-  "\<lbrakk>eventually (\<lambda>x. P x) net; eventually (\<lambda>x. Q x) net\<rbrakk>
-    \<Longrightarrow> eventually (\<lambda>x. P x \<and> Q x) net"
-by (rule eventually_conj)
-
-lemma eventually_rev_mono:
-  "eventually P net \<Longrightarrow> (\<forall>x. P x \<longrightarrow> Q x) \<Longrightarrow> eventually Q net"
-using eventually_mono [of P Q] by fast
-
-lemma eventually_and: " eventually (\<lambda>x. P x \<and> Q x) net \<longleftrightarrow> eventually P net \<and> eventually Q net"
-  by (auto intro!: eventually_conjI elim: eventually_rev_mono)
-
-lemma eventually_false: "eventually (\<lambda>x. False) net \<longleftrightarrow> trivial_limit net"
-  by (auto simp add: eventually_False)
-
-lemma not_eventually: "(\<forall>x. \<not> P x ) \<Longrightarrow> ~(trivial_limit net) ==> ~(eventually (\<lambda>x. P x) net)"
-  by (simp add: eventually_False)
-
-subsection{* Limits, defined as vacuously true when the limit is trivial. *}
-
-  text{* Notation Lim to avoid collition with lim defined in analysis *}
-definition
-  Lim :: "'a net \<Rightarrow> ('a \<Rightarrow> 'b::t2_space) \<Rightarrow> 'b" where
-  "Lim net f = (THE l. (f ---> l) net)"
-
-lemma Lim:
- "(f ---> l) net \<longleftrightarrow>
-        trivial_limit net \<or>
-        (\<forall>e>0. eventually (\<lambda>x. dist (f x) l < e) net)"
-  unfolding tendsto_iff trivial_limit_eq by auto
-
-
-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)"
-  by (auto simp add: tendsto_iff eventually_within_le)
-
-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)"
-  by (auto simp add: tendsto_iff eventually_within)
-
-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)"
-  by (auto simp add: tendsto_iff eventually_at)
-
-lemma Lim_at_iff_LIM: "(f ---> l) (at a) \<longleftrightarrow> f -- a --> l"
-  unfolding Lim_at LIM_def by (simp only: zero_less_dist_iff)
-
-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)"
-  by (auto simp add: tendsto_iff eventually_at_infinity)
-
-lemma Lim_sequentially:
- "(S ---> l) sequentially \<longleftrightarrow>
-          (\<forall>e>0. \<exists>N. \<forall>n\<ge>N. dist (S n) l < e)"
-  by (auto simp add: tendsto_iff eventually_sequentially)
-
-lemma Lim_sequentially_iff_LIMSEQ: "(S ---> l) sequentially \<longleftrightarrow> S ----> l"
-  unfolding Lim_sequentially LIMSEQ_def ..
-
-lemma Lim_eventually: "eventually (\<lambda>x. f x = l) net \<Longrightarrow> (f ---> l) net"
-  by (rule topological_tendstoI, auto elim: eventually_rev_mono)
-
-text{* The expected monotonicity property. *}
-
-lemma Lim_within_empty: "(f ---> l) (net within {})"
-  unfolding tendsto_def Limits.eventually_within by simp
-
-lemma Lim_within_subset: "(f ---> l) (net within S) \<Longrightarrow> T \<subseteq> S \<Longrightarrow> (f ---> l) (net within T)"
-  unfolding tendsto_def Limits.eventually_within
-  by (auto elim!: eventually_elim1)
-
-lemma Lim_Un: assumes "(f ---> l) (net within S)" "(f ---> l) (net within T)"
-  shows "(f ---> l) (net within (S \<union> T))"
-  using assms unfolding tendsto_def Limits.eventually_within
-  apply clarify
-  apply (drule spec, drule (1) mp, drule (1) mp)
-  apply (drule spec, drule (1) mp, drule (1) mp)
-  apply (auto elim: eventually_elim2)
-  done
-
-lemma Lim_Un_univ:
- "(f ---> l) (net within S) \<Longrightarrow> (f ---> l) (net within T) \<Longrightarrow>  S \<union> T = UNIV
-        ==> (f ---> l) net"
-  by (metis Lim_Un within_UNIV)
-
-text{* Interrelations between restricted and unrestricted limits. *}
-
-lemma Lim_at_within: "(f ---> l) net ==> (f ---> l)(net within S)"
-  (* FIXME: rename *)
-  unfolding tendsto_def Limits.eventually_within
-  apply (clarify, drule spec, drule (1) mp, drule (1) mp)
-  by (auto elim!: eventually_elim1)
-
-lemma Lim_within_open:
-  fixes f :: "'a::topological_space \<Rightarrow> 'b::topological_space"
-  assumes"a \<in> S" "open S"
-  shows "(f ---> l)(at a within S) \<longleftrightarrow> (f ---> l)(at a)" (is "?lhs \<longleftrightarrow> ?rhs")
-proof
-  assume ?lhs
-  { fix A assume "open A" "l \<in> A"
-    with `?lhs` have "eventually (\<lambda>x. f x \<in> A) (at a within S)"
-      by (rule topological_tendstoD)
-    hence "eventually (\<lambda>x. x \<in> S \<longrightarrow> f x \<in> A) (at a)"
-      unfolding Limits.eventually_within .
-    then obtain T where "open T" "a \<in> T" "\<forall>x\<in>T. x \<noteq> a \<longrightarrow> x \<in> S \<longrightarrow> f x \<in> A"
-      unfolding eventually_at_topological by fast
-    hence "open (T \<inter> S)" "a \<in> T \<inter> S" "\<forall>x\<in>(T \<inter> S). x \<noteq> a \<longrightarrow> f x \<in> A"
-      using assms by auto
-    hence "\<exists>T. open T \<and> a \<in> T \<and> (\<forall>x\<in>T. x \<noteq> a \<longrightarrow> f x \<in> A)"
-      by fast
-    hence "eventually (\<lambda>x. f x \<in> A) (at a)"
-      unfolding eventually_at_topological .
-  }
-  thus ?rhs by (rule topological_tendstoI)
-next
-  assume ?rhs
-  thus ?lhs by (rule Lim_at_within)
-qed
-
-text{* Another limit point characterization. *}
-
-lemma islimpt_sequential:
-  fixes x :: "'a::metric_space" (* FIXME: generalize to topological_space *)
-  shows "x islimpt S \<longleftrightarrow> (\<exists>f. (\<forall>n::nat. f n \<in> S -{x}) \<and> (f ---> x) sequentially)"
-    (is "?lhs = ?rhs")
-proof
-  assume ?lhs
-  then obtain f where f:"\<forall>y. y>0 \<longrightarrow> f y \<in> S \<and> f y \<noteq> x \<and> dist (f y) x < y"
-    unfolding islimpt_approachable using choice[of "\<lambda>e y. e>0 \<longrightarrow> y\<in>S \<and> y\<noteq>x \<and> dist y x < e"] by auto
-  { fix n::nat
-    have "f (inverse (real n + 1)) \<in> S - {x}" using f by auto
-  }
-  moreover
-  { fix e::real assume "e>0"
-    hence "\<exists>N::nat. inverse (real (N + 1)) < e" using real_arch_inv[of e] apply (auto simp add: Suc_pred') apply(rule_tac x="n - 1" in exI) by auto
-    then obtain N::nat where "inverse (real (N + 1)) < e" by auto
-    hence "\<forall>n\<ge>N. inverse (real n + 1) < e" by (auto, metis Suc_le_mono le_SucE less_imp_inverse_less nat_le_real_less order_less_trans real_of_nat_Suc real_of_nat_Suc_gt_zero)
-    moreover have "\<forall>n\<ge>N. dist (f (inverse (real n + 1))) x < (inverse (real n + 1))" using f `e>0` by auto
-    ultimately have "\<exists>N::nat. \<forall>n\<ge>N. dist (f (inverse (real n + 1))) x < e" apply(rule_tac x=N in exI) apply auto apply(erule_tac x=n in allE)+ by auto
-  }
-  hence " ((\<lambda>n. f (inverse (real n + 1))) ---> x) sequentially"
-    unfolding Lim_sequentially using f by auto
-  ultimately show ?rhs apply (rule_tac x="(\<lambda>n::nat. f (inverse (real n + 1)))" in exI) by auto
-next
-  assume ?rhs
-  then obtain f::"nat\<Rightarrow>'a"  where f:"(\<forall>n. f n \<in> S - {x})" "(\<forall>e>0. \<exists>N. \<forall>n\<ge>N. dist (f n) x < e)" unfolding Lim_sequentially by auto
-  { fix e::real assume "e>0"
-    then obtain N where "dist (f N) x < e" using f(2) by auto
-    moreover have "f N\<in>S" "f N \<noteq> x" using f(1) by auto
-    ultimately have "\<exists>x'\<in>S. x' \<noteq> x \<and> dist x' x < e" by auto
-  }
-  thus ?lhs unfolding islimpt_approachable by auto
-qed
-
-text{* Basic arithmetical combining theorems for limits. *}
-
-lemma Lim_linear:
-  assumes "(f ---> l) net" "bounded_linear h"
-  shows "((\<lambda>x. h (f x)) ---> h l) net"
-using `bounded_linear h` `(f ---> l) net`
-by (rule bounded_linear.tendsto)
-
-lemma Lim_ident_at: "((\<lambda>x. x) ---> a) (at a)"
-  unfolding tendsto_def Limits.eventually_at_topological by fast
-
-lemma Lim_const: "((\<lambda>x. a) ---> a) net"
-  by (rule tendsto_const)
-
-lemma Lim_cmul:
-  fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"
-  shows "(f ---> l) net ==> ((\<lambda>x. c *\<^sub>R f x) ---> c *\<^sub>R l) net"
-  by (intro tendsto_intros)
-
-lemma Lim_neg:
-  fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"
-  shows "(f ---> l) net ==> ((\<lambda>x. -(f x)) ---> -l) net"
-  by (rule tendsto_minus)
-
-lemma Lim_add: fixes f :: "'a \<Rightarrow> 'b::real_normed_vector" shows
- "(f ---> l) net \<Longrightarrow> (g ---> m) net \<Longrightarrow> ((\<lambda>x. f(x) + g(x)) ---> l + m) net"
-  by (rule tendsto_add)
-
-lemma Lim_sub:
-  fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"
-  shows "(f ---> l) net \<Longrightarrow> (g ---> m) net \<Longrightarrow> ((\<lambda>x. f(x) - g(x)) ---> l - m) net"
-  by (rule tendsto_diff)
-
-lemma Lim_null:
-  fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"
-  shows "(f ---> l) net \<longleftrightarrow> ((\<lambda>x. f(x) - l) ---> 0) net" by (simp add: Lim dist_norm)
-
-lemma Lim_null_norm:
-  fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"
-  shows "(f ---> 0) net \<longleftrightarrow> ((\<lambda>x. norm(f x)) ---> 0) net"
-  by (simp add: Lim dist_norm)
-
-lemma Lim_null_comparison:
-  fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"
-  assumes "eventually (\<lambda>x. norm (f x) \<le> g x) net" "(g ---> 0) net"
-  shows "(f ---> 0) net"
-proof(simp add: tendsto_iff, rule+)
-  fix e::real assume "0<e"
-  { fix x
-    assume "norm (f x) \<le> g x" "dist (g x) 0 < e"
-    hence "dist (f x) 0 < e" by (simp add: dist_norm)
-  }
-  thus "eventually (\<lambda>x. dist (f x) 0 < e) net"
-    using eventually_and[of "\<lambda>x. norm(f x) <= g x" "\<lambda>x. dist (g x) 0 < e" net]
-    using eventually_mono[of "(\<lambda>x. norm (f x) \<le> g x \<and> dist (g x) 0 < e)" "(\<lambda>x. dist (f x) 0 < e)" net]
-    using assms `e>0` unfolding tendsto_iff by auto
-qed
-
-lemma Lim_component:
-  fixes f :: "'a \<Rightarrow> 'b::metric_space ^ 'n::finite"
-  shows "(f ---> l) net \<Longrightarrow> ((\<lambda>a. f a $i) ---> l$i) net"
-  unfolding tendsto_iff
-  apply (clarify)
-  apply (drule spec, drule (1) mp)
-  apply (erule eventually_elim1)
-  apply (erule le_less_trans [OF dist_nth_le])
-  done
-
-lemma Lim_transform_bound:
-  fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"
-  fixes g :: "'a \<Rightarrow> 'c::real_normed_vector"
-  assumes "eventually (\<lambda>n. norm(f n) <= norm(g n)) net"  "(g ---> 0) net"
-  shows "(f ---> 0) net"
-proof (rule tendstoI)
-  fix e::real assume "e>0"
-  { fix x
-    assume "norm (f x) \<le> norm (g x)" "dist (g x) 0 < e"
-    hence "dist (f x) 0 < e" by (simp add: dist_norm)}
-  thus "eventually (\<lambda>x. dist (f x) 0 < e) net"
-    using eventually_and[of "\<lambda>x. norm (f x) \<le> norm (g x)" "\<lambda>x. dist (g x) 0 < e" net]
-    using eventually_mono[of "\<lambda>x. norm (f x) \<le> norm (g x) \<and> dist (g x) 0 < e" "\<lambda>x. dist (f x) 0 < e" net]
-    using assms `e>0` unfolding tendsto_iff by blast
-qed
-
-text{* Deducing things about the limit from the elements. *}
-
-lemma Lim_in_closed_set:
-  assumes "closed S" "eventually (\<lambda>x. f(x) \<in> S) net" "\<not>(trivial_limit net)" "(f ---> l) net"
-  shows "l \<in> S"
-proof (rule ccontr)
-  assume "l \<notin> S"
-  with `closed S` have "open (- S)" "l \<in> - S"
-    by (simp_all add: open_Compl)
-  with assms(4) have "eventually (\<lambda>x. f x \<in> - S) net"
-    by (rule topological_tendstoD)
-  with assms(2) have "eventually (\<lambda>x. False) net"
-    by (rule eventually_elim2) simp
-  with assms(3) show "False"
-    by (simp add: eventually_False)
-qed
-
-text{* Need to prove closed(cball(x,e)) before deducing this as a corollary. *}
-
-lemma Lim_dist_ubound:
-  assumes "\<not>(trivial_limit net)" "(f ---> l) net" "eventually (\<lambda>x. dist a (f x) <= e) net"
-  shows "dist a l <= e"
-proof (rule ccontr)
-  assume "\<not> dist a l \<le> e"
-  then have "0 < dist a l - e" by simp
-  with assms(2) have "eventually (\<lambda>x. dist (f x) l < dist a l - e) net"
-    by (rule tendstoD)
-  with assms(3) have "eventually (\<lambda>x. dist a (f x) \<le> e \<and> dist (f x) l < dist a l - e) net"
-    by (rule eventually_conjI)
-  then obtain w where "dist a (f w) \<le> 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)"
-    by (rule add_le_less_mono)
-  hence "dist a (f w) + dist (f w) l < dist a l"
-    by simp
-  also have "\<dots> \<le> dist a (f w) + dist (f w) l"
-    by (rule dist_triangle)
-  finally show False by simp
-qed
-
-lemma Lim_norm_ubound:
-  fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"
-  assumes "\<not>(trivial_limit net)" "(f ---> l) net" "eventually (\<lambda>x. norm(f x) <= e) net"
-  shows "norm(l) <= e"
-proof (rule ccontr)
-  assume "\<not> norm l \<le> e"
-  then have "0 < norm l - e" by simp
-  with assms(2) have "eventually (\<lambda>x. dist (f x) l < norm l - e) net"
-    by (rule tendstoD)
-  with assms(3) have "eventually (\<lambda>x. norm (f x) \<le> e \<and> dist (f x) l < norm l - e) net"
-    by (rule eventually_conjI)
-  then obtain w where "norm (f w) \<le> 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) \<le> e" by (simp_all add: dist_norm)
-  hence "norm (f w - l) + norm (f w) < norm l" by simp
-  hence "norm (f w - l - f w) < norm l" by (rule le_less_trans [OF norm_triangle_ineq4])
-  thus False using `\<not> norm l \<le> e` by simp
-qed
-
-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"
-  shows "e \<le> norm l"
-proof (rule ccontr)
-  assume "\<not> e \<le> norm l"
-  then have "0 < e - norm l" by simp
-  with assms(2) have "eventually (\<lambda>x. dist (f x) l < e - norm l) net"
-    by (rule tendstoD)
-  with assms(3) have "eventually (\<lambda>x. e \<le> norm (f x) \<and> dist (f x) l < e - norm l) net"
-    by (rule eventually_conjI)
-  then obtain w where "e \<le> 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 \<le> norm (f w)" by (simp_all add: dist_norm)
-  hence "norm (f w - l) + norm l < norm (f w)" by (rule less_le_trans)
-  hence "norm (f w - l + l) < norm (f w)" by (rule le_less_trans [OF norm_triangle_ineq])
-  thus False by simp
-qed
-
-text{* Uniqueness of the limit, when nontrivial. *}
-
-lemma Lim_unique:
-  fixes f :: "'a \<Rightarrow> 'b::t2_space"
-  assumes "\<not> trivial_limit net"  "(f ---> l) net"  "(f ---> l') net"
-  shows "l = l'"
-proof (rule ccontr)
-  assume "l \<noteq> l'"
-  obtain U V where "open U" "open V" "l \<in> U" "l' \<in> V" "U \<inter> V = {}"
-    using hausdorff [OF `l \<noteq> l'`] by fast
-  have "eventually (\<lambda>x. f x \<in> U) net"
-    using `(f ---> l) net` `open U` `l \<in> U` by (rule topological_tendstoD)
-  moreover
-  have "eventually (\<lambda>x. f x \<in> V) net"
-    using `(f ---> l') net` `open V` `l' \<in> V` by (rule topological_tendstoD)
-  ultimately
-  have "eventually (\<lambda>x. False) net"
-  proof (rule eventually_elim2)
-    fix x
-    assume "f x \<in> U" "f x \<in> V"
-    hence "f x \<in> U \<inter> V" by simp
-    with `U \<inter> V = {}` show "False" by simp
-  qed
-  with `\<not> trivial_limit net` show "False"
-    by (simp add: eventually_False)
-qed
-
-lemma tendsto_Lim:
-  fixes f :: "'a \<Rightarrow> 'b::t2_space"
-  shows "~(trivial_limit net) \<Longrightarrow> (f ---> l) net ==> Lim net f = l"
-  unfolding Lim_def using Lim_unique[of net f] by auto
-
-text{* Limit under bilinear function *}
-
-lemma Lim_bilinear:
-  assumes "(f ---> l) net" and "(g ---> m) net" and "bounded_bilinear h"
-  shows "((\<lambda>x. h (f x) (g x)) ---> (h l m)) net"
-using `bounded_bilinear h` `(f ---> l) net` `(g ---> m) net`
-by (rule bounded_bilinear.tendsto)
-
-text{* These are special for limits out of the same vector space. *}
-
-lemma Lim_within_id: "(id ---> a) (at a within s)"
-  unfolding tendsto_def Limits.eventually_within eventually_at_topological
-  by auto
-
-lemma Lim_at_id: "(id ---> a) (at a)"
-apply (subst within_UNIV[symmetric]) by (simp add: Lim_within_id)
-
-lemma Lim_at_zero:
-  fixes a :: "'a::real_normed_vector"
-  fixes l :: "'b::topological_space"
-  shows "(f ---> l) (at a) \<longleftrightarrow> ((\<lambda>x. f(a + x)) ---> l) (at 0)" (is "?lhs = ?rhs")
-proof
-  assume "?lhs"
-  { fix S assume "open S" "l \<in> S"
-    with `?lhs` have "eventually (\<lambda>x. f x \<in> S) (at a)"
-      by (rule topological_tendstoD)
-    then obtain d where d: "d>0" "\<forall>x. x \<noteq> a \<and> dist x a < d \<longrightarrow> f x \<in> S"
-      unfolding Limits.eventually_at by fast
-    { fix x::"'a" assume "x \<noteq> 0 \<and> dist x 0 < d"
-      hence "f (a + x) \<in> S" using d
-      apply(erule_tac x="x+a" in allE)
-      by(auto simp add: comm_monoid_add.mult_commute dist_norm dist_commute)
-    }
-    hence "\<exists>d>0. \<forall>x. x \<noteq> 0 \<and> dist x 0 < d \<longrightarrow> f (a + x) \<in> S"
-      using d(1) by auto
-    hence "eventually (\<lambda>x. f (a + x) \<in> S) (at 0)"
-      unfolding Limits.eventually_at .
-  }
-  thus "?rhs" by (rule topological_tendstoI)
-next
-  assume "?rhs"
-  { fix S assume "open S" "l \<in> S"
-    with `?rhs` have "eventually (\<lambda>x. f (a + x) \<in> S) (at 0)"
-      by (rule topological_tendstoD)
-    then obtain d where d: "d>0" "\<forall>x. x \<noteq> 0 \<and> dist x 0 < d \<longrightarrow> f (a + x) \<in> S"
-      unfolding Limits.eventually_at by fast
-    { fix x::"'a" assume "x \<noteq> a \<and> dist x a < d"
-      hence "f x \<in> S" using d apply(erule_tac x="x-a" in allE)
-        by(auto simp add: comm_monoid_add.mult_commute dist_norm dist_commute)
-    }
-    hence "\<exists>d>0. \<forall>x. x \<noteq> a \<and> dist x a < d \<longrightarrow> f x \<in> S" using d(1) by auto
-    hence "eventually (\<lambda>x. f x \<in> S) (at a)" unfolding Limits.eventually_at .
-  }
-  thus "?lhs" by (rule topological_tendstoI)
-qed
-
-text{* It's also sometimes useful to extract the limit point from the net.  *}
-
-definition
-  netlimit :: "'a::t2_space net \<Rightarrow> 'a" where
-  "netlimit net = (SOME a. ((\<lambda>x. x) ---> a) net)"
-
-lemma netlimit_within:
-  assumes "\<not> trivial_limit (at a within S)"
-  shows "netlimit (at a within S) = a"
-unfolding netlimit_def
-apply (rule some_equality)
-apply (rule Lim_at_within)
-apply (rule Lim_ident_at)
-apply (erule Lim_unique [OF assms])
-apply (rule Lim_at_within)
-apply (rule Lim_ident_at)
-done
-
-lemma netlimit_at:
-  fixes a :: "'a::perfect_space"
-  shows "netlimit (at a) = a"
-  apply (subst within_UNIV[symmetric])
-  using netlimit_within[of a UNIV]
-  by (simp add: trivial_limit_at within_UNIV)
-
-text{* Transformation of limit. *}
-
-lemma Lim_transform:
-  fixes f g :: "'a::type \<Rightarrow> 'b::real_normed_vector"
-  assumes "((\<lambda>x. f x - g x) ---> 0) net" "(f ---> l) net"
-  shows "(g ---> l) net"
-proof-
-  from assms have "((\<lambda>x. f x - g x - f x) ---> 0 - l) net" using Lim_sub[of "\<lambda>x. f x - g x" 0 net f l] by auto
-  thus "?thesis" using Lim_neg [of "\<lambda> x. - g x" "-l" net] by auto
-qed
-
-lemma Lim_transform_eventually:
-  "eventually (\<lambda>x. f x = g x) net \<Longrightarrow> (f ---> l) net ==> (g ---> l) net"
-  apply (rule topological_tendstoI)
-  apply (drule (2) topological_tendstoD)
-  apply (erule (1) eventually_elim2, simp)
-  done
-
-lemma Lim_transform_within:
-  fixes l :: "'b::metric_space" (* TODO: generalize *)
-  assumes "0 < d" "(\<forall>x'\<in>S. 0 < dist x' x \<and> dist x' x < d \<longrightarrow> f x' = g x')"
-          "(f ---> l) (at x within S)"
-  shows   "(g ---> l) (at x within S)"
-  using assms(1,3) unfolding Lim_within
-  apply -
-  apply (clarify, rename_tac e)
-  apply (drule_tac x=e in spec, clarsimp, rename_tac r)
-  apply (rule_tac x="min d r" in exI, clarsimp, rename_tac y)
-  apply (drule_tac x=y in bspec, assumption, clarsimp)
-  apply (simp add: assms(2))
-  done
-
-lemma Lim_transform_at:
-  fixes l :: "'b::metric_space" (* TODO: generalize *)
-  shows "0 < d \<Longrightarrow> (\<forall>x'. 0 < dist x' x \<and> dist x' x < d \<longrightarrow> f x' = g x') \<Longrightarrow>
-  (f ---> l) (at x) ==> (g ---> l) (at x)"
-  apply (subst within_UNIV[symmetric])
-  using Lim_transform_within[of d UNIV x f g l]
-  by (auto simp add: within_UNIV)
-
-text{* Common case assuming being away from some crucial point like 0. *}
-
-lemma Lim_transform_away_within:
-  fixes a b :: "'a::metric_space"
-  fixes l :: "'b::metric_space" (* TODO: generalize *)
-  assumes "a\<noteq>b" "\<forall>x\<in> S. x \<noteq> a \<and> x \<noteq> b \<longrightarrow> f x = g x"
-  and "(f ---> l) (at a within S)"
-  shows "(g ---> l) (at a within S)"
-proof-
-  have "\<forall>x'\<in>S. 0 < dist x' a \<and> dist x' a < dist a b \<longrightarrow> f x' = g x'" using assms(2)
-    apply auto apply(erule_tac x=x' in ballE) by (auto simp add: dist_commute)
-  thus ?thesis using Lim_transform_within[of "dist a b" S a f g l] using assms(1,3) unfolding dist_nz by auto
-qed
-
-lemma Lim_transform_away_at:
-  fixes a b :: "'a::metric_space"
-  fixes l :: "'b::metric_space" (* TODO: generalize *)
-  assumes ab: "a\<noteq>b" and fg: "\<forall>x. x \<noteq> a \<and> x \<noteq> b \<longrightarrow> f x = g x"
-  and fl: "(f ---> l) (at a)"
-  shows "(g ---> l) (at a)"
-  using Lim_transform_away_within[OF ab, of UNIV f g l] fg fl
-  by (auto simp add: within_UNIV)
-
-text{* Alternatively, within an open set. *}
-
-lemma Lim_transform_within_open:
-  fixes a :: "'a::metric_space"
-  fixes l :: "'b::metric_space" (* TODO: generalize *)
-  assumes "open S"  "a \<in> S"  "\<forall>x\<in>S. x \<noteq> a \<longrightarrow> f x = g x"  "(f ---> l) (at a)"
-  shows "(g ---> l) (at a)"
-proof-
-  from assms(1,2) obtain e::real where "e>0" and e:"ball a e \<subseteq> S" unfolding open_contains_ball by auto
-  hence "\<forall>x'. 0 < dist x' a \<and> dist x' a < e \<longrightarrow> f x' = g x'" using assms(3)
-    unfolding ball_def subset_eq apply auto apply(erule_tac x=x' in allE) apply(erule_tac x=x' in ballE) by(auto simp add: dist_commute)
-  thus ?thesis using Lim_transform_at[of e a f g l] `e>0` assms(4) by auto
-qed
-
-text{* A congruence rule allowing us to transform limits assuming not at point. *}
-
-(* FIXME: Only one congruence rule for tendsto can be used at a time! *)
-
-lemma Lim_cong_within[cong add]:
-  fixes a :: "'a::metric_space"
-  fixes l :: "'b::metric_space" (* TODO: generalize *)
-  shows "(\<And>x. x \<noteq> a \<Longrightarrow> f x = g x) ==> ((\<lambda>x. f x) ---> l) (at a within S) \<longleftrightarrow> ((g ---> l) (at a within S))"
-  by (simp add: Lim_within dist_nz[symmetric])
-
-lemma Lim_cong_at[cong add]:
-  fixes a :: "'a::metric_space"
-  fixes l :: "'b::metric_space" (* TODO: generalize *)
-  shows "(\<And>x. x \<noteq> a ==> f x = g x) ==> (((\<lambda>x. f x) ---> l) (at a) \<longleftrightarrow> ((g ---> l) (at a)))"
-  by (simp add: Lim_at dist_nz[symmetric])
-
-text{* Useful lemmas on closure and set of possible sequential limits.*}
-
-lemma closure_sequential:
-  fixes l :: "'a::metric_space" (* TODO: generalize *)
-  shows "l \<in> closure S \<longleftrightarrow> (\<exists>x. (\<forall>n. x n \<in> S) \<and> (x ---> l) sequentially)" (is "?lhs = ?rhs")
-proof
-  assume "?lhs" moreover
-  { assume "l \<in> S"
-    hence "?rhs" using Lim_const[of l sequentially] by auto
-  } moreover
-  { assume "l islimpt S"
-    hence "?rhs" unfolding islimpt_sequential by auto
-  } ultimately
-  show "?rhs" unfolding closure_def by auto
-next
-  assume "?rhs"
-  thus "?lhs" unfolding closure_def unfolding islimpt_sequential by auto
-qed
-
-lemma closed_sequential_limits:
-  fixes S :: "'a::metric_space set"
-  shows "closed S \<longleftrightarrow> (\<forall>x l. (\<forall>n. x n \<in> S) \<and> (x ---> l) sequentially \<longrightarrow> l \<in> S)"
-  unfolding closed_limpt
-  using closure_sequential [where 'a='a] closure_closed [where 'a='a] closed_limpt [where 'a='a] islimpt_sequential [where 'a='a] mem_delete [where 'a='a]
-  by metis
-
-lemma closure_approachable:
-  fixes S :: "'a::metric_space set"
-  shows "x \<in> closure S \<longleftrightarrow> (\<forall>e>0. \<exists>y\<in>S. dist y x < e)"
-  apply (auto simp add: closure_def islimpt_approachable)
-  by (metis dist_self)
-
-lemma closed_approachable:
-  fixes S :: "'a::metric_space set"
-  shows "closed S ==> (\<forall>e>0. \<exists>y\<in>S. dist y x < e) \<longleftrightarrow> x \<in> S"
-  by (metis closure_closed closure_approachable)
-
-text{* Some other lemmas about sequences. *}
-
-lemma seq_offset:
-  fixes l :: "'a::metric_space" (* TODO: generalize *)
-  shows "(f ---> l) sequentially ==> ((\<lambda>i. f( i + k)) ---> l) sequentially"
-  apply (auto simp add: Lim_sequentially)
-  by (metis trans_le_add1 )
-
-lemma seq_offset_neg:
-  "(f ---> l) sequentially ==> ((\<lambda>i. f(i - k)) ---> l) sequentially"
-  apply (rule topological_tendstoI)
-  apply (drule (2) topological_tendstoD)
-  apply (simp only: eventually_sequentially)
-  apply (subgoal_tac "\<And>N k (n::nat). N + k <= n ==> N <= n - k")
-  apply metis
-  by arith
-
-lemma seq_offset_rev:
-  "((\<lambda>i. f(i + k)) ---> l) sequentially ==> (f ---> l) sequentially"
-  apply (rule topological_tendstoI)
-  apply (drule (2) topological_tendstoD)
-  apply (simp only: eventually_sequentially)
-  apply (subgoal_tac "\<And>N k (n::nat). N + k <= n ==> N <= n - k \<and> (n - k) + k = n")
-  by metis arith
-
-lemma seq_harmonic: "((\<lambda>n. inverse (real n)) ---> 0) sequentially"
-proof-
-  { fix e::real assume "e>0"
-    hence "\<exists>N::nat. \<forall>n::nat\<ge>N. inverse (real n) < e"
-      using real_arch_inv[of e] apply auto apply(rule_tac x=n in exI)
-      by (metis not_le le_imp_inverse_le not_less real_of_nat_gt_zero_cancel_iff real_of_nat_less_iff xt1(7))
-  }
-  thus ?thesis unfolding Lim_sequentially dist_norm by simp
-qed
-
-text{* More properties of closed balls. *}
-
-lemma closed_cball: "closed (cball x e)"
-unfolding cball_def closed_def
-unfolding Collect_neg_eq [symmetric] not_le
-apply (clarsimp simp add: open_dist, rename_tac y)
-apply (rule_tac x="dist x y - e" in exI, clarsimp)
-apply (rename_tac x')
-apply (cut_tac x=x and y=x' and z=y in dist_triangle)
-apply simp
-done
-
-lemma open_contains_cball: "open S \<longleftrightarrow> (\<forall>x\<in>S. \<exists>e>0.  cball x e \<subseteq> S)"
-proof-
-  { fix x and e::real assume "x\<in>S" "e>0" "ball x e \<subseteq> S"
-    hence "\<exists>d>0. cball x d \<subseteq> S" unfolding subset_eq by (rule_tac x="e/2" in exI, auto)
-  } moreover
-  { fix x and e::real assume "x\<in>S" "e>0" "cball x e \<subseteq> S"
-    hence "\<exists>d>0. ball x d \<subseteq> S" unfolding subset_eq apply(rule_tac x="e/2" in exI) by auto
-  } ultimately
-  show ?thesis unfolding open_contains_ball by auto
-qed
-
-lemma open_contains_cball_eq: "open S ==> (\<forall>x. x \<in> S \<longleftrightarrow> (\<exists>e>0. cball x e \<subseteq> S))"
-  by (metis open_contains_cball subset_eq order_less_imp_le centre_in_cball mem_def)
-
-lemma mem_interior_cball: "x \<in> interior S \<longleftrightarrow> (\<exists>e>0. cball x e \<subseteq> S)"
-  apply (simp add: interior_def, safe)
-  apply (force simp add: open_contains_cball)
-  apply (rule_tac x="ball x e" in exI)
-  apply (simp add: open_ball centre_in_ball subset_trans [OF ball_subset_cball])
-  done
-
-lemma islimpt_ball:
-  fixes x y :: "'a::{real_normed_vector,perfect_space}"
-  shows "y islimpt ball x e \<longleftrightarrow> 0 < e \<and> y \<in> cball x e" (is "?lhs = ?rhs")
-proof
-  assume "?lhs"
-  { assume "e \<le> 0"
-    hence *:"ball x e = {}" using ball_eq_empty[of x e] by auto
-    have False using `?lhs` unfolding * using islimpt_EMPTY[of y] by auto
-  }
-  hence "e > 0" by (metis not_less)
-  moreover
-  have "y \<in> cball x e" using closed_cball[of x e] islimpt_subset[of y "ball x e" "cball x e"] ball_subset_cball[of x e] `?lhs` unfolding closed_limpt by auto
-  ultimately show "?rhs" by auto
-next
-  assume "?rhs" hence "e>0"  by auto
-  { fix d::real assume "d>0"
-    have "\<exists>x'\<in>ball x e. x' \<noteq> y \<and> dist x' y < d"
-    proof(cases "d \<le> dist x y")
-      case True thus "\<exists>x'\<in>ball x e. x' \<noteq> y \<and> dist x' y < d"
-      proof(cases "x=y")
-        case True hence False using `d \<le> dist x y` `d>0` by auto
-        thus "\<exists>x'\<in>ball x e. x' \<noteq> y \<and> dist x' y < d" by auto
-      next
-        case False
-
-        have "dist x (y - (d / (2 * dist y x)) *\<^sub>R (y - x))
-              = norm (x - y + (d / (2 * norm (y - x))) *\<^sub>R (y - x))"
-          unfolding mem_cball mem_ball dist_norm diff_diff_eq2 diff_add_eq[THEN sym] by auto
-        also have "\<dots> = \<bar>- 1 + d / (2 * norm (x - y))\<bar> * norm (x - y)"
-          using scaleR_left_distrib[of "- 1" "d / (2 * norm (y - x))", THEN sym, of "y - x"]
-          unfolding scaleR_minus_left scaleR_one
-          by (auto simp add: norm_minus_commute)
-        also have "\<dots> = \<bar>- norm (x - y) + d / 2\<bar>"
-          unfolding abs_mult_pos[of "norm (x - y)", OF norm_ge_zero[of "x - y"]]
-          unfolding real_add_mult_distrib using `x\<noteq>y`[unfolded dist_nz, unfolded dist_norm] by auto
-        also have "\<dots> \<le> e - d/2" using `d \<le> dist x y` and `d>0` and `?rhs` by(auto simp add: dist_norm)
-        finally have "y - (d / (2 * dist y x)) *\<^sub>R (y - x) \<in> ball x e" using `d>0` by auto
-
-        moreover
-
-        have "(d / (2*dist y x)) *\<^sub>R (y - x) \<noteq> 0"
-          using `x\<noteq>y`[unfolded dist_nz] `d>0` unfolding scaleR_eq_0_iff by (auto simp add: dist_commute)
-        moreover
-        have "dist (y - (d / (2 * dist y x)) *\<^sub>R (y - x)) y < d" unfolding dist_norm apply simp unfolding norm_minus_cancel
-          using `d>0` `x\<noteq>y`[unfolded dist_nz] dist_commute[of x y]
-          unfolding dist_norm by auto
-        ultimately show "\<exists>x'\<in>ball x e. x' \<noteq> y \<and> dist x' y < d" by (rule_tac  x="y - (d / (2*dist y x)) *\<^sub>R (y - x)" in bexI) auto
-      qed
-    next
-      case False hence "d > dist x y" by auto
-      show "\<exists>x'\<in>ball x e. x' \<noteq> y \<and> dist x' y < d"
-      proof(cases "x=y")
-        case True
-        obtain z where **: "z \<noteq> y" "dist z y < min e d"
-          using perfect_choose_dist[of "min e d" y]
-          using `d > 0` `e>0` by auto
-        show "\<exists>x'\<in>ball x e. x' \<noteq> y \<and> dist x' y < d"
-          unfolding `x = y`
-          using `z \<noteq> y` **
-          by (rule_tac x=z in bexI, auto simp add: dist_commute)
-      next
-        case False thus "\<exists>x'\<in>ball x e. x' \<noteq> y \<and> dist x' y < d"
-          using `d>0` `d > dist x y` `?rhs` by(rule_tac x=x in bexI, auto)
-      qed
-    qed  }
-  thus "?lhs" unfolding mem_cball islimpt_approachable mem_ball by auto
-qed
-
-lemma closure_ball_lemma:
-  fixes x y :: "'a::real_normed_vector"
-  assumes "x \<noteq> y" shows "y islimpt ball x (dist x y)"
-proof (rule islimptI)
-  fix T assume "y \<in> T" "open T"
-  then obtain r where "0 < r" "\<forall>z. dist z y < r \<longrightarrow> z \<in> T"
-    unfolding open_dist by fast
-  (* choose point between x and y, within distance r of y. *)
-  def k \<equiv> "min 1 (r / (2 * dist x y))"
-  def z \<equiv> "y + scaleR k (x - y)"
-  have z_def2: "z = x + scaleR (1 - k) (y - x)"
-    unfolding z_def by (simp add: algebra_simps)
-  have "dist z y < r"
-    unfolding z_def k_def using `0 < r`
-    by (simp add: dist_norm min_def)
-  hence "z \<in> T" using `\<forall>z. dist z y < r \<longrightarrow> z \<in> T` by simp
-  have "dist x z < dist x y"
-    unfolding z_def2 dist_norm
-    apply (simp add: norm_minus_commute)
-    apply (simp only: dist_norm [symmetric])
-    apply (subgoal_tac "\<bar>1 - k\<bar> * dist x y < 1 * dist x y", simp)
-    apply (rule mult_strict_right_mono)
-    apply (simp add: k_def divide_pos_pos zero_less_dist_iff `0 < r` `x \<noteq> y`)
-    apply (simp add: zero_less_dist_iff `x \<noteq> y`)
-    done
-  hence "z \<in> ball x (dist x y)" by simp
-  have "z \<noteq> y"
-    unfolding z_def k_def using `x \<noteq> y` `0 < r`
-    by (simp add: min_def)
-  show "\<exists>z\<in>ball x (dist x y). z \<in> T \<and> z \<noteq> y"
-    using `z \<in> ball x (dist x y)` `z \<in> T` `z \<noteq> y`
-    by fast
-qed
-
-lemma closure_ball:
-  fixes x :: "'a::real_normed_vector"
-  shows "0 < e \<Longrightarrow> closure (ball x e) = cball x e"
-apply (rule equalityI)
-apply (rule closure_minimal)
-apply (rule ball_subset_cball)
-apply (rule closed_cball)
-apply (rule subsetI, rename_tac y)
-apply (simp add: le_less [where 'a=real])
-apply (erule disjE)
-apply (rule subsetD [OF closure_subset], simp)
-apply (simp add: closure_def)
-apply clarify
-apply (rule closure_ball_lemma)
-apply (simp add: zero_less_dist_iff)
-done
-
-(* In a trivial vector space, this fails for e = 0. *)
-lemma interior_cball:
-  fixes x :: "'a::{real_normed_vector, perfect_space}"
-  shows "interior (cball x e) = ball x e"
-proof(cases "e\<ge>0")
-  case False note cs = this
-  from cs have "ball x e = {}" using ball_empty[of e x] by auto moreover
-  { fix y assume "y \<in> cball x e"
-    hence False unfolding mem_cball using dist_nz[of x y] cs by auto  }
-  hence "cball x e = {}" by auto
-  hence "interior (cball x e) = {}" using interior_empty by auto
-  ultimately show ?thesis by blast
-next
-  case True note cs = this
-  have "ball x e \<subseteq> cball x e" using ball_subset_cball by auto moreover
-  { fix S y assume as: "S \<subseteq> cball x e" "open S" "y\<in>S"
-    then obtain d where "d>0" and d:"\<forall>x'. dist x' y < d \<longrightarrow> x' \<in> S" unfolding open_dist by blast
-
-    then obtain xa where xa_y: "xa \<noteq> y" and xa: "dist xa y < d"
-      using perfect_choose_dist [of d] by auto
-    have "xa\<in>S" using d[THEN spec[where x=xa]] using xa by(auto simp add: dist_commute)
-    hence xa_cball:"xa \<in> cball x e" using as(1) by auto
-
-    hence "y \<in> ball x e" proof(cases "x = y")
-      case True
-      hence "e>0" using xa_y[unfolded dist_nz] xa_cball[unfolded mem_cball] by (auto simp add: dist_commute)
-      thus "y \<in> ball x e" using `x = y ` by simp
-    next
-      case False
-      have "dist (y + (d / 2 / dist y x) *\<^sub>R (y - x)) y < d" unfolding dist_norm
-        using `d>0` norm_ge_zero[of "y - x"] `x \<noteq> y` by auto
-      hence *:"y + (d / 2 / dist y x) *\<^sub>R (y - x) \<in> cball x e" using d as(1)[unfolded subset_eq] by blast
-      have "y - x \<noteq> 0" using `x \<noteq> y` by auto
-      hence **:"d / (2 * norm (y - x)) > 0" unfolding zero_less_norm_iff[THEN sym]
-        using `d>0` divide_pos_pos[of d "2*norm (y - x)"] by auto
-
-      have "dist (y + (d / 2 / dist y x) *\<^sub>R (y - x)) x = norm (y + (d / (2 * norm (y - x))) *\<^sub>R y - (d / (2 * norm (y - x))) *\<^sub>R x - x)"
-        by (auto simp add: dist_norm algebra_simps)
-      also have "\<dots> = norm ((1 + d / (2 * norm (y - x))) *\<^sub>R (y - x))"
-        by (auto simp add: algebra_simps)
-      also have "\<dots> = \<bar>1 + d / (2 * norm (y - x))\<bar> * norm (y - x)"
-        using ** by auto
-      also have "\<dots> = (dist y x) + d/2"using ** by (auto simp add: left_distrib dist_norm)
-      finally have "e \<ge> dist x y +d/2" using *[unfolded mem_cball] by (auto simp add: dist_commute)
-      thus "y \<in> ball x e" unfolding mem_ball using `d>0` by auto
-    qed  }
-  hence "\<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
-qed
-
-lemma frontier_ball:
-  fixes a :: "'a::real_normed_vector"
-  shows "0 < e ==> frontier(ball a e) = {x. dist a x = e}"
-  apply (simp add: frontier_def closure_ball interior_open open_ball order_less_imp_le)
-  apply (simp add: expand_set_eq)
-  by arith
-
-lemma frontier_cball:
-  fixes a :: "'a::{real_normed_vector, perfect_space}"
-  shows "frontier(cball a e) = {x. dist a x = e}"
-  apply (simp add: frontier_def interior_cball closed_cball closure_closed order_less_imp_le)
-  apply (simp add: expand_set_eq)
-  by arith
-
-lemma cball_eq_empty: "(cball x e = {}) \<longleftrightarrow> e < 0"
-  apply (simp add: expand_set_eq not_le)
-  by (metis zero_le_dist dist_self order_less_le_trans)
-lemma cball_empty: "e < 0 ==> cball x e = {}" by (simp add: cball_eq_empty)
-
-lemma cball_eq_sing:
-  fixes x :: "'a::perfect_space"
-  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"
-    using perfect_choose_dist [OF e] by auto
-  hence "a \<noteq> x" "dist x a \<le> e" by (auto simp add: dist_commute)
-  with e show ?thesis by (auto simp add: expand_set_eq)
-qed auto
-
-lemma cball_sing:
-  fixes x :: "'a::metric_space"
-  shows "e = 0 ==> cball x e = {x}"
-  by (auto simp add: expand_set_eq)
-
-text{* For points in the interior, localization of limits makes no difference.   *}
-
-lemma eventually_within_interior:
-  assumes "x \<in> interior S"
-  shows "eventually P (at x within S) \<longleftrightarrow> eventually P (at x)" (is "?lhs = ?rhs")
-proof-
-  from assms obtain T where T: "open T" "x \<in> T" "T \<subseteq> S"
-    unfolding interior_def by fast
-  { 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"
-      unfolding Limits.eventually_within Limits.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"
-      by auto
-    then have "?rhs"
-      unfolding Limits.eventually_at_topological by auto
-  } moreover
-  { assume "?rhs" hence "?lhs"
-      unfolding Limits.eventually_within
-      by (auto elim: eventually_elim1)
-  } ultimately
-  show "?thesis" ..
-qed
-
-lemma lim_within_interior:
-  "x \<in> interior S \<Longrightarrow> (f ---> l) (at x within S) \<longleftrightarrow> (f ---> l) (at x)"
-  unfolding tendsto_def by (simp add: eventually_within_interior)
-
-lemma netlimit_within_interior:
-  fixes x :: "'a::{perfect_space, real_normed_vector}"
-    (* FIXME: generalize to perfect_space *)
-  assumes "x \<in> interior S"
-  shows "netlimit(at x within S) = x" (is "?lhs = ?rhs")
-proof-
-  from assms obtain e::real where e:"e>0" "ball x e \<subseteq> S" using open_interior[of S] unfolding open_contains_ball using interior_subset[of S] by auto
-  hence "\<not> trivial_limit (at x within S)" using islimpt_subset[of x "ball x e" S] unfolding trivial_limit_within islimpt_ball centre_in_cball by auto
-  thus ?thesis using netlimit_within by auto
-qed
-
-subsection{* Boundedness. *}
-
-  (* FIXME: This has to be unified with BSEQ!! *)
-definition
-  bounded :: "'a::metric_space set \<Rightarrow> bool" where
-  "bounded S \<longleftrightarrow> (\<exists>x e. \<forall>y\<in>S. dist x y \<le> e)"
-
-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 (drule (1) bspec)
-apply (erule order_trans [OF dist_triangle add_left_mono])
-apply auto
-done
-
-lemma bounded_iff: "bounded S \<longleftrightarrow> (\<exists>a. \<forall>x\<in>S. norm x \<le> a)"
-unfolding bounded_any_center [where a=0]
-by (simp add: dist_norm)
-
-lemma bounded_empty[simp]: "bounded {}" by (simp add: bounded_def)
-lemma bounded_subset: "bounded T \<Longrightarrow> S \<subseteq> T ==> bounded S"
-  by (metis bounded_def subset_eq)
-
-lemma bounded_interior[intro]: "bounded S ==> bounded(interior S)"
-  by (metis bounded_subset interior_subset)
-
-lemma bounded_closure[intro]: assumes "bounded S" shows "bounded(closure S)"
-proof-
-  from assms obtain x and a where a: "\<forall>y\<in>S. dist x y \<le> a" unfolding bounded_def by auto
-  { fix y assume "y \<in> closure S"
-    then obtain f where f: "\<forall>n. f n \<in> S"  "(f ---> y) sequentially"
-      unfolding closure_sequential by auto
-    have "\<forall>n. f n \<in> S \<longrightarrow> dist x (f n) \<le> a" using a by simp
-    hence "eventually (\<lambda>n. dist x (f n) \<le> a) sequentially"
-      by (rule eventually_mono, simp add: f(1))
-    have "dist x y \<le> a"
-      apply (rule Lim_dist_ubound [of sequentially f])
-      apply (rule trivial_limit_sequentially)
-      apply (rule f(2))
-      apply fact
-      done
-  }
-  thus ?thesis unfolding bounded_def by auto
-qed
-
-lemma bounded_cball[simp,intro]: "bounded (cball x e)"
-  apply (simp add: bounded_def)
-  apply (rule_tac x=x in exI)
-  apply (rule_tac x=e in exI)
-  apply auto
-  done
-
-lemma bounded_ball[simp,intro]: "bounded(ball x e)"
-  by (metis ball_subset_cball bounded_cball bounded_subset)
-
-lemma finite_imp_bounded[intro]: assumes "finite S" shows "bounded S"
-proof-
-  { fix a F assume as:"bounded F"
-    then obtain x e where "\<forall>y\<in>F. dist x y \<le> e" unfolding bounded_def by auto
-    hence "\<forall>y\<in>(insert a F). dist x y \<le> max e (dist x a)" by auto
-    hence "bounded (insert a F)" unfolding bounded_def by (intro exI)
-  }
-  thus ?thesis using finite_induct[of S bounded]  using bounded_empty assms by auto
-qed
-
-lemma bounded_Un[simp]: "bounded (S \<union> T) \<longleftrightarrow> bounded S \<and> bounded T"
-  apply (auto simp add: bounded_def)
-  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 (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)"
-  by (induct rule: finite_induct[of F], auto)
-
-lemma bounded_pos: "bounded S \<longleftrightarrow> (\<exists>b>0. \<forall>x\<in> S. norm x <= 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)")
-  by metis arith
-
-lemma bounded_Int[intro]: "bounded S \<or> bounded T \<Longrightarrow> bounded (S \<inter> T)"
-  by (metis Int_lower1 Int_lower2 bounded_subset)
-
-lemma bounded_diff[intro]: "bounded S ==> bounded (S - T)"
-apply (metis Diff_subset bounded_subset)
-done
-
-lemma bounded_insert[intro]:"bounded(insert x S) \<longleftrightarrow> bounded S"
-  by (metis Diff_cancel Un_empty_right Un_insert_right bounded_Un bounded_subset finite.emptyI finite_imp_bounded infinite_remove subset_insertI)
-
-lemma not_bounded_UNIV[simp, intro]:
-  "\<not> bounded (UNIV :: 'a::{real_normed_vector, perfect_space} set)"
-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
-  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" ..
-qed
-
-lemma bounded_linear_image:
-  assumes "bounded S" "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" unfolding bounded_pos by auto
-  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"
-    hence "norm x \<le> b" using b by auto
-    hence "norm (f x) \<le> B * b" using B(2) apply(erule_tac x=x in allE)
-      by (metis B(1) B(2) real_le_trans real_mult_le_cancel_iff2)
-  }
-  thus ?thesis unfolding bounded_pos apply(rule_tac x="b*B" in exI)
-    using b B real_mult_order[of b B] by (auto simp add: real_mult_commute)
-qed
-
-lemma bounded_scaling:
-  fixes S :: "'a::real_normed_vector set"
-  shows "bounded S \<Longrightarrow> bounded ((\<lambda>x. c *\<^sub>R x) ` S)"
-  apply (rule bounded_linear_image, assumption)
-  apply (rule scaleR.bounded_linear_right)
-  done
-
-lemma bounded_translation:
-  fixes S :: "'a::real_normed_vector set"
-  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" unfolding bounded_pos by auto
-  { fix x assume "x\<in>S"
-    hence "norm (a + x) \<le> b + norm a" using norm_triangle_ineq[of a x] b by auto
-  }
-  thus ?thesis unfolding bounded_pos using norm_ge_zero[of a] b(1) using add_strict_increasing[of b 0 "norm a"]
-    by (auto intro!: add exI[of _ "b + norm a"])
-qed
-
-
-text{* Some theorems on sups and infs using the notion "bounded". *}
-
-lemma bounded_real:
-  fixes S :: "real set"
-  shows "bounded S \<longleftrightarrow>  (\<exists>a. \<forall>x\<in>S. abs x <= a)"
-  by (simp add: bounded_iff)
-
-lemma bounded_has_rsup: assumes "bounded S" "S \<noteq> {}"
-  shows "\<forall>x\<in>S. x <= rsup S" and "\<forall>b. (\<forall>x\<in>S. x <= b) \<longrightarrow> rsup S <= b"
-proof
-  fix x 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
-  hence *:"S *<= a" using setleI[of S a] by (metis abs_le_interval_iff mem_def)
-  thus "x \<le> rsup S" using rsup[OF `S\<noteq>{}`] using assms(1)[unfolded bounded_real] using isLubD2[of UNIV S "rsup S" x] using `x\<in>S` by auto
-next
-  show "\<forall>b. (\<forall>x\<in>S. x \<le> b) \<longrightarrow> rsup S \<le> b" using assms
-  using rsup[of S, unfolded isLub_def isUb_def leastP_def setle_def setge_def]
-  apply (auto simp add: bounded_real)
-  by (auto simp add: isLub_def isUb_def leastP_def setle_def setge_def)
-qed
-
-lemma rsup_insert: assumes "bounded S"
-  shows "rsup(insert x S) = (if S = {} then x else max x (rsup S))"
-proof(cases "S={}")
-  case True thus ?thesis using rsup_finite_in[of "{x}"] by auto
-next
-  let ?S = "insert x S"
-  case False
-  hence *:"\<forall>x\<in>S. x \<le> rsup S" using bounded_has_rsup(1)[of S] using assms by auto
-  hence "insert x S *<= max x (rsup S)" unfolding setle_def by auto
-  hence "isLub UNIV ?S (rsup ?S)" using rsup[of ?S] by auto
-  moreover
-  have **:"isUb UNIV ?S (max x (rsup S))" unfolding isUb_def setle_def using * by auto
-  { fix y assume as:"isUb UNIV (insert x S) y"
-    hence "max x (rsup S) \<le> y" unfolding isUb_def using rsup_le[OF `S\<noteq>{}`]
-      unfolding setle_def by auto  }
-  hence "max x (rsup S) <=* isUb UNIV (insert x S)" unfolding setge_def Ball_def mem_def by auto
-  hence "isLub UNIV ?S (max x (rsup S))" using ** isLubI2[of UNIV ?S "max x (rsup S)"] unfolding Collect_def by auto
-  ultimately show ?thesis using real_isLub_unique[of UNIV ?S] using `S\<noteq>{}` by auto
-qed
-
-lemma sup_insert_finite: "finite S \<Longrightarrow> rsup(insert x S) = (if S = {} then x else max x (rsup S))"
-  apply (rule rsup_insert)
-  apply (rule finite_imp_bounded)
-  by simp
-
-lemma bounded_has_rinf:
-  assumes "bounded S"  "S \<noteq> {}"
-  shows "\<forall>x\<in>S. x >= rinf S" and "\<forall>b. (\<forall>x\<in>S. x >= b) \<longrightarrow> rinf S >= b"
-proof
-  fix x 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
-  hence *:"- a <=* S" using setgeI[of S "-a"] unfolding abs_le_interval_iff by auto
-  thus "x \<ge> rinf S" using rinf[OF `S\<noteq>{}`] using isGlbD2[of UNIV S "rinf S" x] using `x\<in>S` by auto
-next
-  show "\<forall>b. (\<forall>x\<in>S. x >= b) \<longrightarrow> rinf S \<ge> b" using assms
-  using rinf[of S, unfolded isGlb_def isLb_def greatestP_def setle_def setge_def]
-  apply (auto simp add: bounded_real)
-  by (auto simp add: isGlb_def isLb_def greatestP_def setle_def setge_def)
-qed
-
-(* TODO: Move this to RComplete.thy -- would need to include Glb into RComplete *)
-lemma real_isGlb_unique: "[| isGlb R S x; isGlb R S y |] ==> x = (y::real)"
-  apply (frule isGlb_isLb)
-  apply (frule_tac x = y in isGlb_isLb)
-  apply (blast intro!: order_antisym dest!: isGlb_le_isLb)
-  done
-
-lemma rinf_insert: assumes "bounded S"
-  shows "rinf(insert x S) = (if S = {} then x else min x (rinf S))" (is "?lhs = ?rhs")
-proof(cases "S={}")
-  case True thus ?thesis using rinf_finite_in[of "{x}"] by auto
-next
-  let ?S = "insert x S"
-  case False
-  hence *:"\<forall>x\<in>S. x \<ge> rinf S" using bounded_has_rinf(1)[of S] using assms by auto
-  hence "min x (rinf S) <=* insert x S" unfolding setge_def by auto
-  hence "isGlb UNIV ?S (rinf ?S)" using rinf[of ?S] by auto
-  moreover
-  have **:"isLb UNIV ?S (min x (rinf S))" unfolding isLb_def setge_def using * by auto
-  { fix y assume as:"isLb UNIV (insert x S) y"
-    hence "min x (rinf S) \<ge> y" unfolding isLb_def using rinf_ge[OF `S\<noteq>{}`]
-      unfolding setge_def by auto  }
-  hence "isLb UNIV (insert x S) *<= min x (rinf S)" unfolding setle_def Ball_def mem_def by auto
-  hence "isGlb UNIV ?S (min x (rinf S))" using ** isGlbI2[of UNIV ?S "min x (rinf S)"] unfolding Collect_def by auto
-  ultimately show ?thesis using real_isGlb_unique[of UNIV ?S] using `S\<noteq>{}` by auto
-qed
-
-lemma inf_insert_finite: "finite S ==> rinf(insert x S) = (if S = {} then x else min x (rinf S))"
-  by (rule rinf_insert, rule finite_imp_bounded, simp)
-
-subsection{* Compactness (the definition is the one based on convegent subsequences). *}
-
-definition
-  compact :: "'a::metric_space set \<Rightarrow> bool" where (* TODO: generalize *)
-  "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))"
-
-text {*
-  A metric space (or topological vector space) is said to have the
-  Heine-Borel property if every closed and bounded subset is compact.
-*}
-
-class heine_borel =
-  assumes bounded_imp_convergent_subsequence:
-    "bounded s \<Longrightarrow> \<forall>n. f n \<in> s
-      \<Longrightarrow> \<exists>l r. subseq r \<and> ((f \<circ> r) ---> l) sequentially"
-
-lemma bounded_closed_imp_compact:
-  fixes s::"'a::heine_borel set"
-  assumes "bounded s" and "closed s" shows "compact s"
-proof (unfold compact_def, clarify)
-  fix f :: "nat \<Rightarrow> 'a" assume f: "\<forall>n. f n \<in> s"
-  obtain l r where r: "subseq r" and l: "((f \<circ> r) ---> l) sequentially"
-    using bounded_imp_convergent_subsequence [OF `bounded s` `\<forall>n. f n \<in> s`] by auto
-  from f have fr: "\<forall>n. (f \<circ> r) n \<in> s" by simp
-  have "l \<in> s" using `closed s` fr l
-    unfolding closed_sequential_limits by blast
-  show "\<exists>l\<in>s. \<exists>r. subseq r \<and> ((f \<circ> r) ---> l) sequentially"
-    using `l \<in> s` r l by blast
-qed
-
-lemma subseq_bigger: assumes "subseq r" shows "n \<le> r n"
-proof(induct n)
-  show "0 \<le> r 0" by auto
-next
-  fix n assume "n \<le> r n"
-  moreover have "r n < r (Suc n)"
-    using assms [unfolded subseq_def] by auto
-  ultimately show "Suc n \<le> r (Suc n)" by auto
-qed
-
-lemma eventually_subseq:
-  assumes r: "subseq r"
-  shows "eventually P sequentially \<Longrightarrow> eventually (\<lambda>n. P (r n)) sequentially"
-unfolding eventually_sequentially
-by (metis subseq_bigger [OF r] le_trans)
-
-lemma lim_subseq:
-  "subseq r \<Longrightarrow> (s ---> l) sequentially \<Longrightarrow> ((s o r) ---> l) sequentially"
-unfolding tendsto_def eventually_sequentially o_def
-by (metis subseq_bigger le_trans)
-
-lemma num_Axiom: "EX! g. g 0 = e \<and> (\<forall>n. g (Suc n) = f n (g n))"
-  unfolding Ex1_def
-  apply (rule_tac x="nat_rec e f" in exI)
-  apply (rule conjI)+
-apply (rule def_nat_rec_0, simp)
-apply (rule allI, rule def_nat_rec_Suc, simp)
-apply (rule allI, rule impI, rule ext)
-apply (erule conjE)
-apply (induct_tac x)
-apply (simp add: nat_rec_0)
-apply (erule_tac x="n" in allE)
-apply (simp)
-done
-
-lemma convergent_bounded_increasing: fixes s ::"nat\<Rightarrow>real"
-  assumes "incseq s" and "\<forall>n. abs(s n) \<le> b"
-  shows "\<exists> l. \<forall>e::real>0. \<exists> N. \<forall>n \<ge> N.  abs(s n - l) < e"
-proof-
-  have "isUb UNIV (range s) b" using assms(2) and abs_le_D1 unfolding isUb_def and setle_def by auto
-  then obtain t where t:"isLub UNIV (range s) t" using reals_complete[of "range s" ] by auto
-  { fix e::real assume "e>0" and as:"\<forall>N. \<exists>n\<ge>N. \<not> \<bar>s n - t\<bar> < e"
-    { fix n::nat
-      obtain N where "N\<ge>n" and n:"\<bar>s N - t\<bar> \<ge> e" using as[THEN spec[where x=n]] by auto
-      have "t \<ge> s N" using isLub_isUb[OF t, unfolded isUb_def setle_def] by auto
-      with n have "s N \<le> t - e" using `e>0` by auto
-      hence "s n \<le> t - e" using assms(1)[unfolded incseq_def, THEN spec[where x=n], THEN spec[where x=N]] using `n\<le>N` by auto  }
-    hence "isUb UNIV (range s) (t - e)" unfolding isUb_def and setle_def by auto
-    hence False using isLub_le_isUb[OF t, of "t - e"] and `e>0` by auto  }
-  thus ?thesis by blast
-qed
-
-lemma convergent_bounded_monotone: fixes s::"nat \<Rightarrow> real"
-  assumes "\<forall>n. abs(s n) \<le> b" and "monoseq s"
-  shows "\<exists>l. \<forall>e::real>0. \<exists>N. \<forall>n\<ge>N. abs(s n - l) < e"
-  using convergent_bounded_increasing[of s b] assms using convergent_bounded_increasing[of "\<lambda>n. - s n" b]
-  unfolding monoseq_def incseq_def
-  apply auto unfolding minus_add_distrib[THEN sym, unfolded diff_minus[THEN sym]]
-  unfolding abs_minus_cancel by(rule_tac x="-l" in exI)auto
-
-lemma compact_real_lemma:
-  assumes "\<forall>n::nat. abs(s n) \<le> b"
-  shows "\<exists>(l::real) r. subseq r \<and> ((s \<circ> r) ---> l) sequentially"
-proof-
-  obtain r where r:"subseq r" "monoseq (\<lambda>n. s (r n))"
-    using seq_monosub[of s] by auto
-  thus ?thesis using convergent_bounded_monotone[of "\<lambda>n. s (r n)" b] and assms
-    unfolding tendsto_iff dist_norm eventually_sequentially by auto
-qed
-
-instance real :: heine_borel
-proof
-  fix s :: "real set" and f :: "nat \<Rightarrow> real"
-  assume s: "bounded s" and f: "\<forall>n. f n \<in> s"
-  then obtain b where b: "\<forall>n. abs (f n) \<le> b"
-    unfolding bounded_iff by auto
-  obtain l :: real and r :: "nat \<Rightarrow> nat" where
-    r: "subseq r" and l: "((f \<circ> r) ---> l) sequentially"
-    using compact_real_lemma [OF b] by auto
-  thus "\<exists>l r. subseq r \<and> ((f \<circ> r) ---> l) sequentially"
-    by auto
-qed
-
-lemma bounded_component: "bounded s \<Longrightarrow> bounded ((\<lambda>x. x $ i) ` s)"
-unfolding bounded_def
-apply clarify
-apply (rule_tac x="x $ i" in exI)
-apply (rule_tac x="e" in exI)
-apply clarify
-apply (rule order_trans [OF dist_nth_le], simp)
-done
-
-lemma compact_lemma:
-  fixes f :: "nat \<Rightarrow> 'a::heine_borel ^ 'n::finite"
-  assumes "bounded s" and "\<forall>n. f n \<in> s"
-  shows "\<forall>d.
-        \<exists>l r. subseq r \<and>
-        (\<forall>e>0. eventually (\<lambda>n. \<forall>i\<in>d. dist (f (r n) $ i) (l $ i) < e) sequentially)"
-proof
-  fix d::"'n set" have "finite d" by simp
-  thus "\<exists>l::'a ^ 'n. \<exists>r. subseq r \<and>
-      (\<forall>e>0. eventually (\<lambda>n. \<forall>i\<in>d. dist (f (r n) $ i) (l $ i) < e) sequentially)"
-  proof(induct d) case empty thus ?case unfolding subseq_def by auto
-  next case (insert k d)
-    have s': "bounded ((\<lambda>x. x $ k) ` s)" using `bounded s` by (rule bounded_component)
-    obtain l1::"'a^'n" and r1 where r1:"subseq r1" and lr1:"\<forall>e>0. eventually (\<lambda>n. \<forall>i\<in>d. dist (f (r1 n) $ i) (l1 $ i) < e) sequentially"
-      using insert(3) by auto
-    have f': "\<forall>n. f (r1 n) $ k \<in> (\<lambda>x. x $ k) ` s" using `\<forall>n. f n \<in> s` by simp
-    obtain l2 r2 where r2:"subseq r2" and lr2:"((\<lambda>i. f (r1 (r2 i)) $ k) ---> l2) sequentially"
-      using bounded_imp_convergent_subsequence[OF s' f'] unfolding o_def by auto
-    def r \<equiv> "r1 \<circ> r2" have r:"subseq r"
-      using r1 and r2 unfolding r_def o_def subseq_def by auto
-    moreover
-    def l \<equiv> "(\<chi> i. if i = k then l2 else l1$i)::'a^'n"
-    { fix e::real assume "e>0"
-      from lr1 `e>0` have N1:"eventually (\<lambda>n. \<forall>i\<in>d. dist (f (r1 n) $ i) (l1 $ i) < e) sequentially" by blast
-      from lr2 `e>0` have N2:"eventually (\<lambda>n. dist (f (r1 (r2 n)) $ k) l2 < e) sequentially" by (rule tendstoD)
-      from r2 N1 have N1': "eventually (\<lambda>n. \<forall>i\<in>d. dist (f (r1 (r2 n)) $ i) (l1 $ i) < e) sequentially"
-        by (rule eventually_subseq)
-      have "eventually (\<lambda>n. \<forall>i\<in>(insert k d). dist (f (r n) $ i) (l $ i) < e) sequentially"
-        using N1' N2 by (rule eventually_elim2, simp add: l_def r_def)
-    }
-    ultimately show ?case by auto
-  qed
-qed
-
-instance "^" :: (heine_borel, finite) heine_borel
-proof
-  fix s :: "('a ^ 'b) set" and f :: "nat \<Rightarrow> 'a ^ 'b"
-  assume s: "bounded s" and f: "\<forall>n. f n \<in> s"
-  then obtain l r where r: "subseq r"
-    and l: "\<forall>e>0. eventually (\<lambda>n. \<forall>i\<in>UNIV. dist (f (r n) $ i) (l $ i) < e) sequentially"
-    using compact_lemma [OF s f] by blast
-  let ?d = "UNIV::'b set"
-  { fix e::real assume "e>0"
-    hence "0 < e / (real_of_nat (card ?d))"
-      using zero_less_card_finite using divide_pos_pos[of e, of "real_of_nat (card ?d)"] by auto
-    with l have "eventually (\<lambda>n. \<forall>i. dist (f (r n) $ i) (l $ i) < e / (real_of_nat (card ?d))) sequentially"
-      by simp
-    moreover
-    { fix n assume n: "\<forall>i. dist (f (r n) $ i) (l $ i) < e / (real_of_nat (card ?d))"
-      have "dist (f (r n)) l \<le> (\<Sum>i\<in>?d. dist (f (r n) $ i) (l $ i))"
-        unfolding dist_vector_def using zero_le_dist by (rule setL2_le_setsum)
-      also have "\<dots> < (\<Sum>i\<in>?d. e / (real_of_nat (card ?d)))"
-        by (rule setsum_strict_mono) (simp_all add: n)
-      finally have "dist (f (r n)) l < e" by simp
-    }
-    ultimately have "eventually (\<lambda>n. dist (f (r n)) l < e) sequentially"
-      by (rule eventually_elim1)
-  }
-  hence *:"((f \<circ> r) ---> l) sequentially" unfolding o_def tendsto_iff by simp
-  with r show "\<exists>l r. subseq r \<and> ((f \<circ> r) ---> l) sequentially" by auto
-qed
-
-lemma bounded_fst: "bounded s \<Longrightarrow> bounded (fst ` s)"
-unfolding bounded_def
-apply clarify
-apply (rule_tac x="a" in exI)
-apply (rule_tac x="e" in exI)
-apply clarsimp
-apply (drule (1) bspec)
-apply (simp add: dist_Pair_Pair)
-apply (erule order_trans [OF real_sqrt_sum_squares_ge1])
-done
-
-lemma bounded_snd: "bounded s \<Longrightarrow> bounded (snd ` s)"
-unfolding bounded_def
-apply clarify
-apply (rule_tac x="b" in exI)
-apply (rule_tac x="e" in exI)
-apply clarsimp
-apply (drule (1) bspec)
-apply (simp add: dist_Pair_Pair)
-apply (erule order_trans [OF real_sqrt_sum_squares_ge2])
-done
-
-instance "*" :: (heine_borel, heine_borel) heine_borel
-proof
-  fix s :: "('a * 'b) set" and f :: "nat \<Rightarrow> 'a * 'b"
-  assume s: "bounded s" and f: "\<forall>n. f n \<in> s"
-  from s have s1: "bounded (fst ` s)" by (rule bounded_fst)
-  from f have f1: "\<forall>n. fst (f n) \<in> fst ` s" by simp
-  obtain l1 r1 where r1: "subseq r1"
-    and l1: "((\<lambda>n. fst (f (r1 n))) ---> l1) sequentially"
-    using bounded_imp_convergent_subsequence [OF s1 f1]
-    unfolding o_def by fast
-  from s have s2: "bounded (snd ` s)" by (rule bounded_snd)
-  from f have f2: "\<forall>n. snd (f (r1 n)) \<in> snd ` s" by simp
-  obtain l2 r2 where r2: "subseq r2"
-    and l2: "((\<lambda>n. snd (f (r1 (r2 n)))) ---> l2) sequentially"
-    using bounded_imp_convergent_subsequence [OF s2 f2]
-    unfolding o_def by fast
-  have l1': "((\<lambda>n. fst (f (r1 (r2 n)))) ---> l1) sequentially"
-    using lim_subseq [OF r2 l1] unfolding o_def .
-  have l: "((f \<circ> (r1 \<circ> r2)) ---> (l1, l2)) sequentially"
-    using tendsto_Pair [OF l1' l2] unfolding o_def by simp
-  have r: "subseq (r1 \<circ> r2)"
-    using r1 r2 unfolding subseq_def by simp
-  show "\<exists>l r. subseq r \<and> ((f \<circ> r) ---> l) sequentially"
-    using l r by fast
-qed
-
-subsection{* Completeness. *}
-
-lemma cauchy_def:
-  "Cauchy s \<longleftrightarrow> (\<forall>e>0. \<exists>N. \<forall>m n. m \<ge> N \<and> n \<ge> N --> dist(s m)(s n) < e)"
-unfolding Cauchy_def by blast
-
-definition
-  complete :: "'a::metric_space set \<Rightarrow> bool" where
-  "complete s \<longleftrightarrow> (\<forall>f. (\<forall>n. f n \<in> s) \<and> Cauchy f
-                      --> (\<exists>l \<in> s. (f ---> l) sequentially))"
-
-lemma cauchy: "Cauchy s \<longleftrightarrow> (\<forall>e>0.\<exists> N::nat. \<forall>n\<ge>N. dist(s n)(s N) < e)" (is "?lhs = ?rhs")
-proof-
-  { assume ?rhs
-    { fix e::real
-      assume "e>0"
-      with `?rhs` obtain N where N:"\<forall>n\<ge>N. dist (s n) (s N) < e/2"
-        by (erule_tac x="e/2" in allE) auto
-      { fix n m
-        assume nm:"N \<le> m \<and> N \<le> n"
-        hence "dist (s m) (s n) < e" using N
-          using dist_triangle_half_l[of "s m" "s N" "e" "s n"]
-          by blast
-      }
-      hence "\<exists>N. \<forall>m n. N \<le> m \<and> N \<le> n \<longrightarrow> dist (s m) (s n) < e"
-        by blast
-    }
-    hence ?lhs
-      unfolding cauchy_def
-      by blast
-  }
-  thus ?thesis
-    unfolding cauchy_def
-    using dist_triangle_half_l
-    by blast
-qed
-
-lemma convergent_imp_cauchy:
- "(s ---> l) sequentially ==> Cauchy s"
-proof(simp only: cauchy_def, rule, rule)
-  fix e::real assume "e>0" "(s ---> l) sequentially"
-  then obtain N::nat where N:"\<forall>n\<ge>N. dist (s n) l < e/2" unfolding Lim_sequentially by(erule_tac x="e/2" in allE) auto
-  thus "\<exists>N. \<forall>m n. N \<le> m \<and> N \<le> n \<longrightarrow> dist (s m) (s n) < e"  using dist_triangle_half_l[of _ l e _] by (rule_tac x=N in exI) auto
-qed
-
-lemma cauchy_imp_bounded: assumes "Cauchy s" shows "bounded {y. (\<exists>n::nat. y = s n)}"
-proof-
-  from assms obtain N::nat where "\<forall>m n. N \<le> m \<and> N \<le> n \<longrightarrow> dist (s m) (s n) < 1" unfolding cauchy_def apply(erule_tac x= 1 in allE) by auto
-  hence N:"\<forall>n. N \<le> n \<longrightarrow> dist (s N) (s n) < 1" by auto
-  moreover
-  have "bounded (s ` {0..N})" using finite_imp_bounded[of "s ` {1..N}"] by auto
-  then obtain a where a:"\<forall>x\<in>s ` {0..N}. dist (s N) x \<le> a"
-    unfolding bounded_any_center [where a="s N"] by auto
-  ultimately show "?thesis"
-    unfolding bounded_any_center [where a="s N"]
-    apply(rule_tac x="max a 1" in exI) apply auto
-    apply(erule_tac x=n in allE) apply(erule_tac x=n in ballE) by auto
-qed
-
-lemma compact_imp_complete: assumes "compact s" shows "complete s"
-proof-
-  { fix f assume as: "(\<forall>n::nat. f n \<in> s)" "Cauchy f"
-    from as(1) obtain l r where lr: "l\<in>s" "subseq r" "((f \<circ> r) ---> l) sequentially" using assms unfolding compact_def by blast
-
-    note lr' = subseq_bigger [OF lr(2)]
-
-    { fix e::real assume "e>0"
-      from as(2) obtain N where N:"\<forall>m n. N \<le> m \<and> N \<le> n \<longrightarrow> dist (f m) (f n) < e/2" unfolding cauchy_def using `e>0` apply (erule_tac x="e/2" in allE) by auto
-      from lr(3)[unfolded Lim_sequentially, THEN spec[where x="e/2"]] obtain M where M:"\<forall>n\<ge>M. dist ((f \<circ> r) n) l < e/2" using `e>0` by auto
-      { fix n::nat assume n:"n \<ge> max N M"
-        have "dist ((f \<circ> r) n) l < e/2" using n M by auto
-        moreover have "r n \<ge> N" using lr'[of n] n by auto
-        hence "dist (f n) ((f \<circ> r) n) < e / 2" using N using n by auto
-        ultimately have "dist (f n) l < e" using dist_triangle_half_r[of "f (r n)" "f n" e l] by (auto simp add: dist_commute)  }
-      hence "\<exists>N. \<forall>n\<ge>N. dist (f n) l < e" by blast  }
-    hence "\<exists>l\<in>s. (f ---> l) sequentially" using `l\<in>s` unfolding Lim_sequentially by auto  }
-  thus ?thesis unfolding complete_def by auto
-qed
-
-instance heine_borel < complete_space
-proof
-  fix f :: "nat \<Rightarrow> 'a" assume "Cauchy f"
-  hence "bounded (range f)" unfolding image_def
-    using cauchy_imp_bounded [of f] by auto
-  hence "compact (closure (range f))"
-    using bounded_closed_imp_compact [of "closure (range f)"] by auto
-  hence "complete (closure (range f))"
-    using compact_imp_complete by auto
-  moreover have "\<forall>n. f n \<in> closure (range f)"
-    using closure_subset [of "range f"] by auto
-  ultimately have "\<exists>l\<in>closure (range f). (f ---> l) sequentially"
-    using `Cauchy f` unfolding complete_def by auto
-  then show "convergent f"
-    unfolding convergent_def LIMSEQ_conv_tendsto [symmetric] by auto
-qed
-
-lemma complete_univ: "complete (UNIV :: 'a::complete_space set)"
-proof(simp add: complete_def, rule, rule)
-  fix f :: "nat \<Rightarrow> 'a" assume "Cauchy f"
-  hence "convergent f" by (rule Cauchy_convergent)
-  hence "\<exists>l. f ----> l" unfolding convergent_def .  
-  thus "\<exists>l. (f ---> l) sequentially" unfolding LIMSEQ_conv_tendsto .
-qed
-
-lemma complete_imp_closed: assumes "complete s" shows "closed s"
-proof -
-  { fix x assume "x islimpt s"
-    then obtain f where f: "\<forall>n. f n \<in> s - {x}" "(f ---> x) sequentially"
-      unfolding islimpt_sequential by auto
-    then obtain l where l: "l\<in>s" "(f ---> l) sequentially"
-      using `complete s`[unfolded complete_def] using convergent_imp_cauchy[of f x] by auto
-    hence "x \<in> s"  using Lim_unique[of sequentially f l x] trivial_limit_sequentially f(2) by auto
-  }
-  thus "closed s" unfolding closed_limpt by auto
-qed
-
-lemma complete_eq_closed:
-  fixes s :: "'a::complete_space set"
-  shows "complete s \<longleftrightarrow> closed s" (is "?lhs = ?rhs")
-proof
-  assume ?lhs thus ?rhs by (rule complete_imp_closed)
-next
-  assume ?rhs
-  { fix f assume as:"\<forall>n::nat. f n \<in> s" "Cauchy f"
-    then obtain l where "(f ---> l) sequentially" using complete_univ[unfolded complete_def, THEN spec[where x=f]] by auto
-    hence "\<exists>l\<in>s. (f ---> l) sequentially" using `?rhs`[unfolded closed_sequential_limits, THEN spec[where x=f], THEN spec[where x=l]] using as(1) by auto  }
-  thus ?lhs unfolding complete_def by auto
-qed
-
-lemma convergent_eq_cauchy:
-  fixes s :: "nat \<Rightarrow> 'a::complete_space"
-  shows "(\<exists>l. (s ---> l) sequentially) \<longleftrightarrow> Cauchy s" (is "?lhs = ?rhs")
-proof
-  assume ?lhs then obtain l where "(s ---> l) sequentially" by auto
-  thus ?rhs using convergent_imp_cauchy by auto
-next
-  assume ?rhs thus ?lhs using complete_univ[unfolded complete_def, THEN spec[where x=s]] by auto
-qed
-
-lemma convergent_imp_bounded:
-  fixes s :: "nat \<Rightarrow> 'a::metric_space"
-  shows "(s ---> l) sequentially ==> bounded (s ` (UNIV::(nat set)))"
-  using convergent_imp_cauchy[of s]
-  using cauchy_imp_bounded[of s]
-  unfolding image_def
-  by auto
-
-subsection{* Total boundedness. *}
-
-fun helper_1::"('a::metric_space set) \<Rightarrow> real \<Rightarrow> nat \<Rightarrow> 'a" where
-  "helper_1 s e n = (SOME y::'a. y \<in> s \<and> (\<forall>m<n. \<not> (dist (helper_1 s e m) y < e)))"
-declare helper_1.simps[simp del]
-
-lemma compact_imp_totally_bounded:
-  assumes "compact s"
-  shows "\<forall>e>0. \<exists>k. finite k \<and> k \<subseteq> s \<and> s \<subseteq> (\<Union>((\<lambda>x. ball x e) ` k))"
-proof(rule, rule, rule ccontr)
-  fix e::real assume "e>0" and assm:"\<not> (\<exists>k. finite k \<and> k \<subseteq> s \<and> s \<subseteq> \<Union>(\<lambda>x. ball x e) ` k)"
-  def x \<equiv> "helper_1 s e"
-  { fix n
-    have "x n \<in> s \<and> (\<forall>m<n. \<not> dist (x m) (x n) < e)"
-    proof(induct_tac rule:nat_less_induct)
-      fix n  def Q \<equiv> "(\<lambda>y. y \<in> s \<and> (\<forall>m<n. \<not> dist (x m) y < e))"
-      assume as:"\<forall>m<n. x m \<in> s \<and> (\<forall>ma<m. \<not> dist (x ma) (x m) < e)"
-      have "\<not> s \<subseteq> (\<Union>x\<in>x ` {0..<n}. ball x e)" using assm apply simp apply(erule_tac x="x ` {0 ..< n}" in allE) using as by auto
-      then obtain z where z:"z\<in>s" "z \<notin> (\<Union>x\<in>x ` {0..<n}. ball x e)" unfolding subset_eq by auto
-      have "Q (x n)" unfolding x_def and helper_1.simps[of s e n]
-        apply(rule someI2[where a=z]) unfolding x_def[symmetric] and Q_def using z by auto
-      thus "x n \<in> s \<and> (\<forall>m<n. \<not> dist (x m) (x n) < e)" unfolding Q_def by auto
-    qed }
-  hence "\<forall>n::nat. x n \<in> s" and x:"\<forall>n. \<forall>m < n. \<not> (dist (x m) (x n) < e)" by blast+
-  then obtain l r where "l\<in>s" and r:"subseq r" and "((x \<circ> r) ---> l) sequentially" using assms(1)[unfolded compact_def, THEN spec[where x=x]] by auto
-  from this(3) have "Cauchy (x \<circ> r)" using convergent_imp_cauchy by auto
-  then obtain N::nat where N:"\<forall>m n. N \<le> m \<and> N \<le> n \<longrightarrow> dist ((x \<circ> r) m) ((x \<circ> r) n) < e" unfolding cauchy_def using `e>0` by auto
-  show False
-    using N[THEN spec[where x=N], THEN spec[where x="N+1"]]
-    using r[unfolded subseq_def, THEN spec[where x=N], THEN spec[where x="N+1"]]
-    using x[THEN spec[where x="r (N+1)"], THEN spec[where x="r (N)"]] by auto
-qed
-
-subsection{* Heine-Borel theorem (following Burkill \& Burkill vol. 2) *}
-
-lemma heine_borel_lemma: fixes s::"'a::metric_space set"
-  assumes "compact s"  "s \<subseteq> (\<Union> t)"  "\<forall>b \<in> t. open b"
-  shows "\<exists>e>0. \<forall>x \<in> s. \<exists>b \<in> t. ball x e \<subseteq> b"
-proof(rule ccontr)
-  assume "\<not> (\<exists>e>0. \<forall>x\<in>s. \<exists>b\<in>t. ball x e \<subseteq> b)"
-  hence cont:"\<forall>e>0. \<exists>x\<in>s. \<forall>xa\<in>t. \<not> (ball x e \<subseteq> xa)" by auto
-  { fix n::nat
-    have "1 / real (n + 1) > 0" by auto
-    hence "\<exists>x. x\<in>s \<and> (\<forall>xa\<in>t. \<not> (ball x (inverse (real (n+1))) \<subseteq> xa))" using cont unfolding Bex_def by auto }
-  hence "\<forall>n::nat. \<exists>x. x \<in> s \<and> (\<forall>xa\<in>t. \<not> ball x (inverse (real (n + 1))) \<subseteq> xa)" by auto
-  then obtain f where f:"\<forall>n::nat. f n \<in> s \<and> (\<forall>xa\<in>t. \<not> ball (f n) (inverse (real (n + 1))) \<subseteq> xa)"
-    using choice[of "\<lambda>n::nat. \<lambda>x. x\<in>s \<and> (\<forall>xa\<in>t. \<not> ball x (inverse (real (n + 1))) \<subseteq> xa)"] by auto
-
-  then obtain l r where l:"l\<in>s" and r:"subseq r" and lr:"((f \<circ> r) ---> l) sequentially"
-    using assms(1)[unfolded compact_def, THEN spec[where x=f]] by auto
-
-  obtain b where "l\<in>b" "b\<in>t" using assms(2) and l by auto
-  then obtain e where "e>0" and e:"\<forall>z. dist z l < e \<longrightarrow> z\<in>b"
-    using assms(3)[THEN bspec[where x=b]] unfolding open_dist by auto
-
-  then obtain N1 where N1:"\<forall>n\<ge>N1. dist ((f \<circ> r) n) l < e / 2"
-    using lr[unfolded Lim_sequentially, THEN spec[where x="e/2"]] by auto
-
-  obtain N2::nat where N2:"N2>0" "inverse (real N2) < e /2" using real_arch_inv[of "e/2"] and `e>0` by auto
-  have N2':"inverse (real (r (N1 + N2) +1 )) < e/2"
-    apply(rule order_less_trans) apply(rule less_imp_inverse_less) using N2
-    using subseq_bigger[OF r, of "N1 + N2"] by auto
-
-  def x \<equiv> "(f (r (N1 + N2)))"
-  have x:"\<not> ball x (inverse (real (r (N1 + N2) + 1))) \<subseteq> b" unfolding x_def
-    using f[THEN spec[where x="r (N1 + N2)"]] using `b\<in>t` by auto
-  have "\<exists>y\<in>ball x (inverse (real (r (N1 + N2) + 1))). y\<notin>b" apply(rule ccontr) using x by auto
-  then obtain y where y:"y \<in> ball x (inverse (real (r (N1 + N2) + 1)))" "y \<notin> b" by auto
-
-  have "dist x l < e/2" using N1 unfolding x_def o_def by auto
-  hence "dist y l < e" using y N2' using dist_triangle[of y l x]by (auto simp add:dist_commute)
-
-  thus False using e and `y\<notin>b` by auto
-qed
-
-lemma compact_imp_heine_borel: "compact s ==> (\<forall>f. (\<forall>t \<in> f. open t) \<and> s \<subseteq> (\<Union> f)
-               \<longrightarrow> (\<exists>f'. f' \<subseteq> f \<and> finite f' \<and> s \<subseteq> (\<Union> f')))"
-proof clarify
-  fix f assume "compact s" " \<forall>t\<in>f. open t" "s \<subseteq> \<Union>f"
-  then obtain e::real where "e>0" and "\<forall>x\<in>s. \<exists>b\<in>f. ball x e \<subseteq> b" using heine_borel_lemma[of s f] by auto
-  hence "\<forall>x\<in>s. \<exists>b. b\<in>f \<and> ball x e \<subseteq> b" by auto
-  hence "\<exists>bb. \<forall>x\<in>s. bb x \<in>f \<and> ball x e \<subseteq> bb x" using bchoice[of s "\<lambda>x b. b\<in>f \<and> ball x e \<subseteq> b"] by auto
-  then obtain  bb where bb:"\<forall>x\<in>s. (bb x) \<in> f \<and> ball x e \<subseteq> (bb x)" by blast
-
-  from `compact s` have  "\<exists> k. finite k \<and> k \<subseteq> s \<and> s \<subseteq> \<Union>(\<lambda>x. ball x e) ` k" using compact_imp_totally_bounded[of s] `e>0` by auto
-  then obtain k where k:"finite k" "k \<subseteq> s" "s \<subseteq> \<Union>(\<lambda>x. ball x e) ` k" by auto
-
-  have "finite (bb ` k)" using k(1) by auto
-  moreover
-  { fix x assume "x\<in>s"
-    hence "x\<in>\<Union>(\<lambda>x. ball x e) ` k" using k(3)  unfolding subset_eq by auto
-    hence "\<exists>X\<in>bb ` k. x \<in> X" using bb k(2) by blast
-    hence "x \<in> \<Union>(bb ` k)" using  Union_iff[of x "bb ` k"] by auto
-  }
-  ultimately show "\<exists>f'\<subseteq>f. finite f' \<and> s \<subseteq> \<Union>f'" using bb k(2) by (rule_tac x="bb ` k" in exI) auto
-qed
-
-subsection{* Bolzano-Weierstrass property. *}
-
-lemma heine_borel_imp_bolzano_weierstrass:
-  assumes "\<forall>f. (\<forall>t \<in> f. open t) \<and> s \<subseteq> (\<Union> f) --> (\<exists>f'. f' \<subseteq> f \<and> finite f' \<and> s \<subseteq> (\<Union> f'))"
-          "infinite t"  "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)" 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"
-    using assms(1)[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
-  { fix x y assume "x\<in>t" "y\<in>t" "f x = f y"
-    hence "x \<in> f x"  "y \<in> f x \<longrightarrow> y = x" using f[THEN bspec[where x=x]] and `t\<subseteq>s` by auto
-    hence "x = y" using `f x = f y` and f[THEN bspec[where x=y]] and `y\<in>t` and `t\<subseteq>s` by auto  }
-  hence "infinite (f ` t)" using assms(2) using finite_imageD[unfolded inj_on_def, of f t] by auto
-  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" using g'[THEN bspec[where x=h]] by auto
-    hence "y = x" using f[THEN bspec[where x=y]] and `x\<in>t` and `x\<in>h`[unfolded `h = f y`] by auto
-    hence False using `f x \<notin> g` `h\<in>g` unfolding `h = f y` by auto  }
-  hence "f ` t \<subseteq> g" by auto
-  ultimately show False using g(2) using finite_subset by auto
-qed
-
-subsection{* Complete the chain of compactness variants. *}
-
-primrec helper_2::"(real \<Rightarrow> 'a::metric_space) \<Rightarrow> nat \<Rightarrow> 'a" where
-  "helper_2 beyond 0 = beyond 0" |
-  "helper_2 beyond (Suc n) = beyond (dist undefined (helper_2 beyond n) + 1 )"
-
-lemma bolzano_weierstrass_imp_bounded: fixes s::"'a::metric_space set"
-  assumes "\<forall>t. infinite t \<and> t \<subseteq> s --> (\<exists>x \<in> s. x islimpt t)"
-  shows "bounded s"
-proof(rule ccontr)
-  assume "\<not> bounded s"
-  then obtain beyond where "\<forall>a. beyond a \<in>s \<and> \<not> dist undefined (beyond a) \<le> a"
-    unfolding bounded_any_center [where a=undefined]
-    apply simp using choice[of "\<lambda>a x. x\<in>s \<and> \<not> dist undefined x \<le> a"] by auto
-  hence beyond:"\<And>a. beyond a \<in>s" "\<And>a. dist undefined (beyond a) > a"
-    unfolding linorder_not_le by auto
-  def x \<equiv> "helper_2 beyond"
-
-  { fix m n ::nat assume "m<n"
-    hence "dist undefined (x m) + 1 < dist undefined (x n)"
-    proof(induct n)
-      case 0 thus ?case by auto
-    next
-      case (Suc n)
-      have *:"dist undefined (x n) + 1 < dist undefined (x (Suc n))"
-        unfolding x_def and helper_2.simps
-        using beyond(2)[of "dist undefined (helper_2 beyond n) + 1"] by auto
-      thus ?case proof(cases "m < n")
-        case True thus ?thesis using Suc and * by auto
-      next
-        case False hence "m = n" using Suc(2) by auto
-        thus ?thesis using * by auto
-      qed
-    qed  } note * = this
-  { fix m n ::nat assume "m\<noteq>n"
-    have "1 < dist (x m) (x n)"
-    proof(cases "m<n")
-      case True
-      hence "1 < dist undefined (x n) - dist undefined (x m)" using *[of m n] by auto
-      thus ?thesis using dist_triangle [of undefined "x n" "x m"] by arith
-    next
-      case False hence "n<m" using `m\<noteq>n` by auto
-      hence "1 < dist undefined (x m) - dist undefined (x n)" using *[of n m] by auto
-      thus ?thesis using dist_triangle2 [of undefined "x m" "x n"] by arith
-    qed  } note ** = this
-  { fix a b assume "x a = x b" "a \<noteq> b"
-    hence False using **[of a b] by auto  }
-  hence "inj x" unfolding inj_on_def by auto
-  moreover
-  { fix n::nat
-    have "x n \<in> s"
-    proof(cases "n = 0")
-      case True thus ?thesis unfolding x_def using beyond by auto
-    next
-      case False then obtain z where "n = Suc z" using not0_implies_Suc by auto
-      thus ?thesis unfolding x_def using beyond by auto
-    qed  }
-  ultimately have "infinite (range x) \<and> range x \<subseteq> s" unfolding x_def using range_inj_infinite[of "helper_2 beyond"] using beyond(1) by auto
-
-  then obtain l where "l\<in>s" and l:"l islimpt range x" using assms[THEN spec[where x="range x"]] by auto
-  then obtain y where "x y \<noteq> l" and y:"dist (x y) l < 1/2" unfolding islimpt_approachable apply(erule_tac x="1/2" in allE) by auto
-  then obtain z where "x z \<noteq> l" and z:"dist (x z) l < dist (x y) l" using l[unfolded islimpt_approachable, THEN spec[where x="dist (x y) l"]]
-    unfolding dist_nz by auto
-  show False using y and z and dist_triangle_half_l[of "x y" l 1 "x z"] and **[of y z] by auto
-qed
-
-lemma sequence_infinite_lemma:
-  fixes l :: "'a::metric_space" (* TODO: generalize *)
-  assumes "\<forall>n::nat. (f n  \<noteq> l)"  "(f ---> l) sequentially"
-  shows "infinite {y. (\<exists> n. y = f n)}"
-proof(rule ccontr)
-  let ?A = "(\<lambda>x. dist x l) ` {y. \<exists>n. y = f n}"
-  assume "\<not> infinite {y. \<exists>n. y = f n}"
-  hence **:"finite ?A" "?A \<noteq> {}" by auto
-  obtain k where k:"dist (f k) l = Min ?A" using Min_in[OF **] by auto
-  have "0 < Min ?A" using assms(1) unfolding dist_nz unfolding Min_gr_iff[OF **] by auto
-  then obtain N where "dist (f N) l < Min ?A" using assms(2)[unfolded Lim_sequentially, THEN spec[where x="Min ?A"]] by auto
-  moreover have "dist (f N) l \<in> ?A" by auto
-  ultimately show False using Min_le[OF **(1), of "dist (f N) l"] by auto
-qed
-
-lemma sequence_unique_limpt:
-  fixes l :: "'a::metric_space" (* TODO: generalize *)
-  assumes "\<forall>n::nat. (f n \<noteq> l)"  "(f ---> l) sequentially"  "l' islimpt {y.  (\<exists>n. y = f n)}"
-  shows "l' = l"
-proof(rule ccontr)
-  def e \<equiv> "dist l' l"
-  assume "l' \<noteq> l" hence "e>0" unfolding dist_nz e_def by auto
-  then obtain N::nat where N:"\<forall>n\<ge>N. dist (f n) l < e / 2"
-    using assms(2)[unfolded Lim_sequentially, THEN spec[where x="e/2"]] by auto
-  def d \<equiv> "Min (insert (e/2) ((\<lambda>n. if dist (f n) l' = 0 then e/2 else dist (f n) l') ` {0 .. N}))"
-  have "d>0" using `e>0` unfolding d_def e_def using zero_le_dist[of _ l', unfolded order_le_less] by auto
-  obtain k where k:"f k \<noteq> l'"  "dist (f k) l' < d" using `d>0` and assms(3)[unfolded islimpt_approachable, THEN spec[where x="d"]] by auto
-  have "k\<ge>N" using k(1)[unfolded dist_nz] using k(2)[unfolded d_def]
-    by force
-  hence "dist l' l < e" using N[THEN spec[where x=k]] using k(2)[unfolded d_def] and dist_triangle_half_r[of "f k" l' e l] by auto
-  thus False unfolding e_def by auto
-qed
-
-lemma bolzano_weierstrass_imp_closed:
-  fixes s :: "'a::metric_space set" (* TODO: can this be generalized? *)
-  assumes "\<forall>t. infinite t \<and> t \<subseteq> s --> (\<exists>x \<in> s. x islimpt t)"
-  shows "closed s"
-proof-
-  { fix x l assume as: "\<forall>n::nat. x n \<in> s" "(x ---> l) sequentially"
-    hence "l \<in> s"
-    proof(cases "\<forall>n. x n \<noteq> l")
-      case False thus "l\<in>s" using as(1) by auto
-    next
-      case True note cas = this
-      with as(2) have "infinite {y. \<exists>n. y = x n}" using sequence_infinite_lemma[of x l] by auto
-      then obtain l' where "l'\<in>s" "l' islimpt {y. \<exists>n. y = x n}" using assms[THEN spec[where x="{y. \<exists>n. y = x n}"]] as(1) by auto
-      thus "l\<in>s" using sequence_unique_limpt[of x l l'] using as cas by auto
-    qed  }
-  thus ?thesis unfolding closed_sequential_limits by fast
-qed
-
-text{* Hence express everything as an equivalence.   *}
-
-lemma compact_eq_heine_borel:
-  fixes s :: "'a::heine_borel set"
-  shows "compact s \<longleftrightarrow>
-           (\<forall>f. (\<forall>t \<in> f. open t) \<and> s \<subseteq> (\<Union> f)
-               --> (\<exists>f'. f' \<subseteq> f \<and> finite f' \<and> s \<subseteq> (\<Union> f')))" (is "?lhs = ?rhs")
-proof
-  assume ?lhs thus ?rhs using compact_imp_heine_borel[of s] by blast
-next
-  assume ?rhs
-  hence "\<forall>t. infinite t \<and> t \<subseteq> s \<longrightarrow> (\<exists>x\<in>s. x islimpt t)"
-    by (blast intro: heine_borel_imp_bolzano_weierstrass[of s])
-  thus ?lhs using bolzano_weierstrass_imp_bounded[of s] bolzano_weierstrass_imp_closed[of s] bounded_closed_imp_compact[of s] by blast
-qed
-
-lemma compact_eq_bolzano_weierstrass:
-  fixes s :: "'a::heine_borel set"
-  shows "compact s \<longleftrightarrow> (\<forall>t. infinite t \<and> t \<subseteq> s --> (\<exists>x \<in> s. x islimpt t))" (is "?lhs = ?rhs")
-proof
-  assume ?lhs thus ?rhs unfolding compact_eq_heine_borel using heine_borel_imp_bolzano_weierstrass[of s] by auto
-next
-  assume ?rhs thus ?lhs using bolzano_weierstrass_imp_bounded bolzano_weierstrass_imp_closed bounded_closed_imp_compact by auto
-qed
-
-lemma compact_eq_bounded_closed:
-  fixes s :: "'a::heine_borel set"
-  shows "compact s \<longleftrightarrow> bounded s \<and> closed s"  (is "?lhs = ?rhs")
-proof
-  assume ?lhs thus ?rhs unfolding compact_eq_bolzano_weierstrass using bolzano_weierstrass_imp_bounded bolzano_weierstrass_imp_closed by auto
-next
-  assume ?rhs thus ?lhs using bounded_closed_imp_compact by auto
-qed
-
-lemma compact_imp_bounded:
-  fixes s :: "'a::metric_space set"
-  shows "compact s ==> bounded s"
-proof -
-  assume "compact s"
-  hence "\<forall>f. (\<forall>t\<in>f. open t) \<and> s \<subseteq> \<Union>f \<longrightarrow> (\<exists>f'\<subseteq>f. finite f' \<and> s \<subseteq> \<Union>f')"
-    by (rule compact_imp_heine_borel)
-  hence "\<forall>t. infinite t \<and> t \<subseteq> s \<longrightarrow> (\<exists>x \<in> s. x islimpt t)"
-    using heine_borel_imp_bolzano_weierstrass[of s] by auto
-  thus "bounded s"
-    by (rule bolzano_weierstrass_imp_bounded)
-qed
-
-lemma compact_imp_closed:
-  fixes s :: "'a::metric_space set"
-  shows "compact s ==> closed s"
-proof -
-  assume "compact s"
-  hence "\<forall>f. (\<forall>t\<in>f. open t) \<and> s \<subseteq> \<Union>f \<longrightarrow> (\<exists>f'\<subseteq>f. finite f' \<and> s \<subseteq> \<Union>f')"
-    by (rule compact_imp_heine_borel)
-  hence "\<forall>t. infinite t \<and> t \<subseteq> s \<longrightarrow> (\<exists>x \<in> s. x islimpt t)"
-    using heine_borel_imp_bolzano_weierstrass[of s] by auto
-  thus "closed s"
-    by (rule bolzano_weierstrass_imp_closed)
-qed
-
-text{* In particular, some common special cases. *}
-
-lemma compact_empty[simp]:
- "compact {}"
-  unfolding compact_def
-  by simp
-
-(* TODO: can any of the next 3 lemmas be generalized to metric spaces? *)
-
-  (* FIXME : Rename *)
-lemma compact_union[intro]:
-  fixes s t :: "'a::heine_borel set"
-  shows "compact s \<Longrightarrow> compact t ==> compact (s \<union> t)"
-  unfolding compact_eq_bounded_closed
-  using bounded_Un[of s t]
-  using closed_Un[of s t]
-  by simp
-
-lemma compact_inter[intro]:
-  fixes s t :: "'a::heine_borel set"
-  shows "compact s \<Longrightarrow> compact t ==> compact (s \<inter> t)"
-  unfolding compact_eq_bounded_closed
-  using bounded_Int[of s t]
-  using closed_Int[of s t]
-  by simp
-
-lemma compact_inter_closed[intro]:
-  fixes s t :: "'a::heine_borel set"
-  shows "compact s \<Longrightarrow> closed t ==> compact (s \<inter> t)"
-  unfolding compact_eq_bounded_closed
-  using closed_Int[of s t]
-  using bounded_subset[of "s \<inter> t" s]
-  by blast
-
-lemma closed_inter_compact[intro]:
-  fixes s t :: "'a::heine_borel set"
-  shows "closed s \<Longrightarrow> compact t ==> compact (s \<inter> t)"
-proof-
-  assume "closed s" "compact t"
-  moreover
-  have "s \<inter> t = t \<inter> s" by auto ultimately
-  show ?thesis
-    using compact_inter_closed[of t s]
-    by auto
-qed
-
-lemma closed_sing [simp]:
-  fixes a :: "'a::metric_space"
-  shows "closed {a}"
-  apply (clarsimp simp add: closed_def open_dist)
-  apply (rule ccontr)
-  apply (drule_tac x="dist x a" in spec)
-  apply (simp add: dist_nz dist_commute)
-  done
-
-lemma finite_imp_closed:
-  fixes s :: "'a::metric_space set"
-  shows "finite s ==> closed s"
-proof (induct set: finite)
-  case empty show "closed {}" by simp
-next
-  case (insert x F)
-  hence "closed ({x} \<union> F)" by (simp only: closed_Un closed_sing)
-  thus "closed (insert x F)" by simp
-qed
-
-lemma finite_imp_compact:
-  fixes s :: "'a::heine_borel set"
-  shows "finite s ==> compact s"
-  unfolding compact_eq_bounded_closed
-  using finite_imp_closed finite_imp_bounded
-  by blast
-
-lemma compact_sing [simp]: "compact {a}"
-  unfolding compact_def o_def subseq_def
-  by (auto simp add: tendsto_const)
-
-lemma compact_cball[simp]:
-  fixes x :: "'a::heine_borel"
-  shows "compact(cball x e)"
-  using compact_eq_bounded_closed bounded_cball closed_cball
-  by blast
-
-lemma compact_frontier_bounded[intro]:
-  fixes s :: "'a::heine_borel set"
-  shows "bounded s ==> compact(frontier s)"
-  unfolding frontier_def
-  using compact_eq_bounded_closed
-  by blast
-
-lemma compact_frontier[intro]:
-  fixes s :: "'a::heine_borel set"
-  shows "compact s ==> compact (frontier s)"
-  using compact_eq_bounded_closed compact_frontier_bounded
-  by blast
-
-lemma frontier_subset_compact:
-  fixes s :: "'a::heine_borel set"
-  shows "compact s ==> frontier s \<subseteq> s"
-  using frontier_subset_closed compact_eq_bounded_closed
-  by blast
-
-lemma open_delete:
-  fixes s :: "'a::metric_space set"
-  shows "open s ==> open(s - {x})"
-  using open_Diff[of s "{x}"] closed_sing
-  by blast
-
-text{* Finite intersection property. I could make it an equivalence in fact. *}
-
-lemma compact_imp_fip:
-  fixes s :: "'a::heine_borel set"
-  assumes "compact s"  "\<forall>t \<in> f. closed t"
-        "\<forall>f'. finite f' \<and> f' \<subseteq> f --> (s \<inter> (\<Inter> f') \<noteq> {})"
-  shows "s \<inter> (\<Inter> f) \<noteq> {}"
-proof
-  assume as:"s \<inter> (\<Inter> f) = {}"
-  hence "s \<subseteq> \<Union>op - UNIV ` f" by auto
-  moreover have "Ball (op - UNIV ` f) open" using open_Diff closed_Diff using assms(2) by auto
-  ultimately obtain f' where f':"f' \<subseteq> op - UNIV ` f"  "finite f'"  "s \<subseteq> \<Union>f'" using assms(1)[unfolded compact_eq_heine_borel, THEN spec[where x="(\<lambda>t. UNIV - t) ` f"]] by auto
-  hence "finite (op - UNIV ` f') \<and> op - UNIV ` f' \<subseteq> f" by(auto simp add: Diff_Diff_Int)
-  hence "s \<inter> \<Inter>op - UNIV ` f' \<noteq> {}" using assms(3)[THEN spec[where x="op - UNIV ` f'"]] by auto
-  thus False using f'(3) unfolding subset_eq and Union_iff by blast
-qed
-
-subsection{* Bounded closed nest property (proof does not use Heine-Borel).            *}
-
-lemma bounded_closed_nest:
-  assumes "\<forall>n. closed(s n)" "\<forall>n. (s n \<noteq> {})"
-  "(\<forall>m n. m \<le> n --> s n \<subseteq> s m)"  "bounded(s 0)"
-  shows "\<exists>a::'a::heine_borel. \<forall>n::nat. a \<in> s(n)"
-proof-
-  from assms(2) obtain x where x:"\<forall>n::nat. x n \<in> s n" using choice[of "\<lambda>n x. x\<in> s n"] by auto
-  from assms(4,1) have *:"compact (s 0)" using bounded_closed_imp_compact[of "s 0"] by auto
-
-  then obtain l r where lr:"l\<in>s 0" "subseq r" "((x \<circ> r) ---> l) sequentially"
-    unfolding compact_def apply(erule_tac x=x in allE)  using x using assms(3) by blast
-
-  { fix n::nat
-    { fix e::real assume "e>0"
-      with lr(3) obtain N where N:"\<forall>m\<ge>N. dist ((x \<circ> r) m) l < e" unfolding Lim_sequentially by auto
-      hence "dist ((x \<circ> r) (max N n)) l < e" by auto
-      moreover
-      have "r (max N n) \<ge> n" using lr(2) using subseq_bigger[of r "max N n"] by auto
-      hence "(x \<circ> r) (max N n) \<in> s n"
-        using x apply(erule_tac x=n in allE)
-        using x apply(erule_tac x="r (max N n)" in allE)
-        using assms(3) apply(erule_tac x=n in allE)apply( erule_tac x="r (max N n)" in allE) by auto
-      ultimately have "\<exists>y\<in>s n. dist y l < e" by auto
-    }
-    hence "l \<in> s n" using closed_approachable[of "s n" l] assms(1) by blast
-  }
-  thus ?thesis by auto
-qed
-
-text{* Decreasing case does not even need compactness, just completeness.        *}
-
-lemma decreasing_closed_nest:
-  assumes "\<forall>n. closed(s n)"
-          "\<forall>n. (s n \<noteq> {})"
-          "\<forall>m n. m \<le> n --> s n \<subseteq> s m"
-          "\<forall>e>0. \<exists>n. \<forall>x \<in> (s n). \<forall> y \<in> (s n). dist x y < e"
-  shows "\<exists>a::'a::heine_borel. \<forall>n::nat. a \<in> s n"
-proof-
-  have "\<forall>n. \<exists> x. x\<in>s n" using assms(2) by auto
-  hence "\<exists>t. \<forall>n. t n \<in> s n" using choice[of "\<lambda> n x. x \<in> s n"] by auto
-  then obtain t where t: "\<forall>n. t n \<in> s n" by auto
-  { fix e::real assume "e>0"
-    then obtain N where N:"\<forall>x\<in>s N. \<forall>y\<in>s N. dist x y < e" using assms(4) by auto
-    { fix m n ::nat assume "N \<le> m \<and> N \<le> n"
-      hence "t m \<in> s N" "t n \<in> s N" using assms(3) t unfolding  subset_eq t by blast+
-      hence "dist (t m) (t n) < e" using N by auto
-    }
-    hence "\<exists>N. \<forall>m n. N \<le> m \<and> N \<le> n \<longrightarrow> dist (t m) (t n) < e" by auto
-  }
-  hence  "Cauchy t" unfolding cauchy_def by auto
-  then obtain l where l:"(t ---> l) sequentially" using complete_univ unfolding complete_def by auto
-  { fix n::nat
-    { fix e::real assume "e>0"
-      then obtain N::nat where N:"\<forall>n\<ge>N. dist (t n) l < e" using l[unfolded Lim_sequentially] by auto
-      have "t (max n N) \<in> s n" using assms(3) unfolding subset_eq apply(erule_tac x=n in allE) apply (erule_tac x="max n N" in allE) using t by auto
-      hence "\<exists>y\<in>s n. dist y l < e" apply(rule_tac x="t (max n N)" in bexI) using N by auto
-    }
-    hence "l \<in> s n" using closed_approachable[of "s n" l] assms(1) by auto
-  }
-  then show ?thesis by auto
-qed
-
-text{* Strengthen it to the intersection actually being a singleton.             *}
-
-lemma decreasing_closed_nest_sing:
-  assumes "\<forall>n. closed(s n)"
-          "\<forall>n. s n \<noteq> {}"
-          "\<forall>m n. m \<le> n --> s n \<subseteq> s m"
-          "\<forall>e>0. \<exists>n. \<forall>x \<in> (s n). \<forall> y\<in>(s n). dist x y < e"
-  shows "\<exists>a::'a::heine_borel. \<Inter> {t. (\<exists>n::nat. t = s n)} = {a}"
-proof-
-  obtain a where a:"\<forall>n. a \<in> s n" using decreasing_closed_nest[of s] using assms by auto
-  { fix b assume b:"b \<in> \<Inter>{t. \<exists>n. t = s n}"
-    { fix e::real assume "e>0"
-      hence "dist a b < e" using assms(4 )using b using a by blast
-    }
-    hence "dist a b = 0" by (metis dist_eq_0_iff dist_nz real_less_def)
-  }
-  with a have "\<Inter>{t. \<exists>n. t = s n} = {a}"  by auto
-  thus ?thesis by auto
-qed
-
-text{* Cauchy-type criteria for uniform convergence. *}
-
-lemma uniformly_convergent_eq_cauchy: fixes s::"nat \<Rightarrow> 'b \<Rightarrow> 'a::heine_borel" shows
- "(\<exists>l. \<forall>e>0. \<exists>N. \<forall>n x. N \<le> n \<and> P x --> dist(s n x)(l x) < e) \<longleftrightarrow>
-  (\<forall>e>0. \<exists>N. \<forall>m n x. N \<le> m \<and> N \<le> n \<and> P x  --> dist (s m x) (s n x) < e)" (is "?lhs = ?rhs")
-proof(rule)
-  assume ?lhs
-  then obtain l where l:"\<forall>e>0. \<exists>N. \<forall>n x. N \<le> n \<and> P x \<longrightarrow> dist (s n x) (l x) < e" by auto
-  { fix e::real assume "e>0"
-    then obtain N::nat where N:"\<forall>n x. N \<le> n \<and> P x \<longrightarrow> dist (s n x) (l x) < e / 2" using l[THEN spec[where x="e/2"]] by auto
-    { fix n m::nat and x::"'b" assume "N \<le> m \<and> N \<le> n \<and> P x"
-      hence "dist (s m x) (s n x) < e"
-        using N[THEN spec[where x=m], THEN spec[where x=x]]
-        using N[THEN spec[where x=n], THEN spec[where x=x]]
-        using dist_triangle_half_l[of "s m x" "l x" e "s n x"] by auto  }
-    hence "\<exists>N. \<forall>m n x. N \<le> m \<and> N \<le> n \<and> P x  --> dist (s m x) (s n x) < e"  by auto  }
-  thus ?rhs by auto
-next
-  assume ?rhs
-  hence "\<forall>x. P x \<longrightarrow> Cauchy (\<lambda>n. s n x)" unfolding cauchy_def apply auto by (erule_tac x=e in allE)auto
-  then obtain l where l:"\<forall>x. P x \<longrightarrow> ((\<lambda>n. s n x) ---> l x) sequentially" unfolding convergent_eq_cauchy[THEN sym]
-    using choice[of "\<lambda>x l. P x \<longrightarrow> ((\<lambda>n. s n x) ---> l) sequentially"] by auto
-  { fix e::real assume "e>0"
-    then obtain N where N:"\<forall>m n x. N \<le> m \<and> N \<le> n \<and> P x \<longrightarrow> dist (s m x) (s n x) < e/2"
-      using `?rhs`[THEN spec[where x="e/2"]] by auto
-    { fix x assume "P x"
-      then obtain M where M:"\<forall>n\<ge>M. dist (s n x) (l x) < e/2"
-        using l[THEN spec[where x=x], unfolded Lim_sequentially] using `e>0` by(auto elim!: allE[where x="e/2"])
-      fix n::nat assume "n\<ge>N"
-      hence "dist(s n x)(l x) < e"  using `P x`and N[THEN spec[where x=n], THEN spec[where x="N+M"], THEN spec[where x=x]]
-        using M[THEN spec[where x="N+M"]] and dist_triangle_half_l[of "s n x" "s (N+M) x" e "l x"] by (auto simp add: dist_commute)  }
-    hence "\<exists>N. \<forall>n x. N \<le> n \<and> P x \<longrightarrow> dist(s n x)(l x) < e" by auto }
-  thus ?lhs by auto
-qed
-
-lemma uniformly_cauchy_imp_uniformly_convergent:
-  fixes s :: "nat \<Rightarrow> 'a \<Rightarrow> 'b::heine_borel"
-  assumes "\<forall>e>0.\<exists>N. \<forall>m (n::nat) x. N \<le> m \<and> N \<le> n \<and> P x --> dist(s m x)(s n x) < e"
-          "\<forall>x. P x --> (\<forall>e>0. \<exists>N. \<forall>n. N \<le> n --> dist(s n x)(l x) < e)"
-  shows "\<forall>e>0. \<exists>N. \<forall>n x. N \<le> n \<and> P x --> dist(s n x)(l x) < e"
-proof-
-  obtain l' where l:"\<forall>e>0. \<exists>N. \<forall>n x. N \<le> n \<and> P x \<longrightarrow> dist (s n x) (l' x) < e"
-    using assms(1) unfolding uniformly_convergent_eq_cauchy[THEN sym] by auto
-  moreover
-  { fix x assume "P x"
-    hence "l x = l' x" using Lim_unique[OF trivial_limit_sequentially, of "\<lambda>n. s n x" "l x" "l' x"]
-      using l and assms(2) unfolding Lim_sequentially by blast  }
-  ultimately show ?thesis by auto
-qed
-
-subsection{* Define continuity over a net to take in restrictions of the set. *}
-
-definition
-  continuous :: "'a::t2_space net \<Rightarrow> ('a \<Rightarrow> 'b::topological_space) \<Rightarrow> bool" where
-  "continuous net f \<longleftrightarrow> (f ---> f(netlimit net)) net"
-
-lemma continuous_trivial_limit:
- "trivial_limit net ==> continuous net f"
-  unfolding continuous_def tendsto_def trivial_limit_eq by auto
-
-lemma continuous_within: "continuous (at x within s) f \<longleftrightarrow> (f ---> f(x)) (at x within s)"
-  unfolding continuous_def
-  unfolding tendsto_def
-  using netlimit_within[of x s]
-  by (cases "trivial_limit (at x within s)") (auto simp add: trivial_limit_eventually)
-
-lemma continuous_at: "continuous (at x) f \<longleftrightarrow> (f ---> f(x)) (at x)"
-  using continuous_within [of x UNIV f] by (simp add: within_UNIV)
-
-lemma continuous_at_within:
-  assumes "continuous (at x) f"  shows "continuous (at x within s) f"
-  using assms unfolding continuous_at continuous_within
-  by (rule Lim_at_within)
-
-text{* Derive the epsilon-delta forms, which we often use as "definitions" *}
-
-lemma continuous_within_eps_delta:
-  "continuous (at x within s) f \<longleftrightarrow> (\<forall>e>0. \<exists>d>0. \<forall>x'\<in> s.  dist x' x < d --> dist (f x') (f x) < e)"
-  unfolding continuous_within and Lim_within
-  apply auto unfolding dist_nz[THEN sym] apply(auto elim!:allE) apply(rule_tac x=d in exI) by auto
-
-lemma continuous_at_eps_delta: "continuous (at x) f \<longleftrightarrow>  (\<forall>e>0. \<exists>d>0.
-                           \<forall>x'. dist x' x < d --> dist(f x')(f x) < e)"
-  using continuous_within_eps_delta[of x UNIV f]
-  unfolding within_UNIV by blast
-
-text{* Versions in terms of open balls. *}
-
-lemma continuous_within_ball:
- "continuous (at x within s) f \<longleftrightarrow> (\<forall>e>0. \<exists>d>0.
-                            f ` (ball x d \<inter> s) \<subseteq> ball (f x) e)" (is "?lhs = ?rhs")
-proof
-  assume ?lhs
-  { fix e::real assume "e>0"
-    then obtain d where d: "d>0" "\<forall>xa\<in>s. 0 < dist xa x \<and> dist xa x < d \<longrightarrow> dist (f xa) (f x) < e"
-      using `?lhs`[unfolded continuous_within Lim_within] by auto
-    { fix y assume "y\<in>f ` (ball x d \<inter> s)"
-      hence "y \<in> ball (f x) e" using d(2) unfolding dist_nz[THEN sym]
-        apply (auto simp add: dist_commute mem_ball) apply(erule_tac x=xa in ballE) apply auto using `e>0` by auto
-    }
-    hence "\<exists>d>0. f ` (ball x d \<inter> s) \<subseteq> ball (f x) e" using `d>0` unfolding subset_eq ball_def by (auto simp add: dist_commute)  }
-  thus ?rhs by auto
-next
-  assume ?rhs thus ?lhs unfolding continuous_within Lim_within ball_def subset_eq
-    apply (auto simp add: dist_commute) apply(erule_tac x=e in allE) by auto
-qed
-
-lemma continuous_at_ball:
-  "continuous (at x) f \<longleftrightarrow> (\<forall>e>0. \<exists>d>0. f ` (ball x d) \<subseteq> ball (f x) e)" (is "?lhs = ?rhs")
-proof
-  assume ?lhs thus ?rhs unfolding continuous_at Lim_at subset_eq Ball_def Bex_def image_iff mem_ball
-    apply auto apply(erule_tac x=e in allE) apply auto apply(rule_tac x=d in exI) apply auto apply(erule_tac x=xa in allE) apply (auto simp add: dist_commute dist_nz)
-    unfolding dist_nz[THEN sym] by auto
-next
-  assume ?rhs thus ?lhs unfolding continuous_at Lim_at subset_eq Ball_def Bex_def image_iff mem_ball
-    apply auto apply(erule_tac x=e in allE) apply auto apply(rule_tac x=d in exI) apply auto apply(erule_tac x="f xa" in allE) by (auto simp add: dist_commute dist_nz)
-qed
-
-text{* For setwise continuity, just start from the epsilon-delta definitions. *}
-
-definition
-  continuous_on :: "'a::metric_space set \<Rightarrow> ('a \<Rightarrow> 'b::metric_space) \<Rightarrow> bool" where
-  "continuous_on s f \<longleftrightarrow> (\<forall>x \<in> s. \<forall>e>0. \<exists>d::real>0. \<forall>x' \<in> s. dist x' x < d --> dist (f x') (f x) < e)"
-
-
-definition
-  uniformly_continuous_on ::
-    "'a::metric_space set \<Rightarrow> ('a \<Rightarrow> 'b::metric_space) \<Rightarrow> bool" where
-  "uniformly_continuous_on s f \<longleftrightarrow>
-        (\<forall>e>0. \<exists>d>0. \<forall>x\<in>s. \<forall> x'\<in>s. dist x' x < d
-                           --> dist (f x') (f x) < e)"
-
-text{* Some simple consequential lemmas. *}
-
-lemma uniformly_continuous_imp_continuous:
- " uniformly_continuous_on s f ==> continuous_on s f"
-  unfolding uniformly_continuous_on_def continuous_on_def by blast
-
-lemma continuous_at_imp_continuous_within:
- "continuous (at x) f ==> continuous (at x within s) f"
-  unfolding continuous_within continuous_at using Lim_at_within by auto
-
-lemma continuous_at_imp_continuous_on: assumes "(\<forall>x \<in> s. continuous (at x) f)"
-  shows "continuous_on s f"
-proof(simp add: continuous_at continuous_on_def, rule, rule, rule)
-  fix x and e::real assume "x\<in>s" "e>0"
-  hence "eventually (\<lambda>xa. dist (f xa) (f x) < e) (at x)" using assms unfolding continuous_at tendsto_iff by auto
-  then obtain d where d:"d>0" "\<forall>xa. 0 < dist xa x \<and> dist xa x < d \<longrightarrow> dist (f xa) (f x) < e" unfolding eventually_at by auto
-  { fix x' assume "\<not> 0 < dist x' x"
-    hence "x=x'"
-      using dist_nz[of x' x] by auto
-    hence "dist (f x') (f x) < e" using `e>0` by auto
-  }
-  thus "\<exists>d>0. \<forall>x'\<in>s. dist x' x < d \<longrightarrow> dist (f x') (f x) < e" using d by auto
-qed
-
-lemma continuous_on_eq_continuous_within:
- "continuous_on s f \<longleftrightarrow> (\<forall>x \<in> s. continuous (at x within s) f)" (is "?lhs = ?rhs")
-proof
-  assume ?rhs
-  { fix x assume "x\<in>s"
-    fix e::real assume "e>0"
-    assume "\<exists>d>0. \<forall>xa\<in>s. 0 < dist xa x \<and> dist xa x < d \<longrightarrow> dist (f xa) (f x) < e"
-    then obtain d where "d>0" and d:"\<forall>xa\<in>s. 0 < dist xa x \<and> dist xa x < d \<longrightarrow> dist (f xa) (f x) < e" by auto
-    { fix x' assume as:"x'\<in>s" "dist x' x < d"
-      hence "dist (f x') (f x) < e" using `e>0` d `x'\<in>s` dist_eq_0_iff[of x' x] zero_le_dist[of x' x] as(2) by (metis dist_eq_0_iff dist_nz) }
-    hence "\<exists>d>0. \<forall>x'\<in>s. dist x' x < d \<longrightarrow> dist (f x') (f x) < e" using `d>0` by auto
-  }
-  thus ?lhs using `?rhs` unfolding continuous_on_def continuous_within Lim_within by auto
-next
-  assume ?lhs
-  thus ?rhs unfolding continuous_on_def continuous_within Lim_within by blast
-qed
-
-lemma continuous_on:
- "continuous_on s f \<longleftrightarrow> (\<forall>x \<in> s. (f ---> f(x)) (at x within s))"
-  by (auto simp add: continuous_on_eq_continuous_within continuous_within)
-
-lemma continuous_on_eq_continuous_at:
- "open s ==> (continuous_on s f \<longleftrightarrow> (\<forall>x \<in> s. continuous (at x) f))"
-  by (auto simp add: continuous_on continuous_at Lim_within_open)
-
-lemma continuous_within_subset:
- "continuous (at x within s) f \<Longrightarrow> t \<subseteq> s
-             ==> continuous (at x within t) f"
-  unfolding continuous_within by(metis Lim_within_subset)
-
-lemma continuous_on_subset:
- "continuous_on s f \<Longrightarrow> t \<subseteq> s ==> continuous_on t f"
-  unfolding continuous_on by (metis subset_eq Lim_within_subset)
-
-lemma continuous_on_interior:
- "continuous_on s f \<Longrightarrow> x \<in> interior s ==> continuous (at x) f"
-unfolding interior_def
-apply simp
-by (meson continuous_on_eq_continuous_at continuous_on_subset)
-
-lemma continuous_on_eq:
- "(\<forall>x \<in> s. f x = g x) \<Longrightarrow> continuous_on s f
-           ==> continuous_on s g"
-  by (simp add: continuous_on_def)
-
-text{* Characterization of various kinds of continuity in terms of sequences.  *}
-
-(* \<longrightarrow> could be generalized, but \<longleftarrow> requires metric space *)
-lemma continuous_within_sequentially:
-  fixes f :: "'a::metric_space \<Rightarrow> 'b::metric_space"
-  shows "continuous (at a within s) f \<longleftrightarrow>
-                (\<forall>x. (\<forall>n::nat. x n \<in> s) \<and> (x ---> a) sequentially
-                     --> ((f o x) ---> f a) sequentially)" (is "?lhs = ?rhs")
-proof
-  assume ?lhs
-  { fix x::"nat \<Rightarrow> 'a" assume x:"\<forall>n. x n \<in> s" "\<forall>e>0. \<exists>N. \<forall>n\<ge>N. dist (x n) a < e"
-    fix e::real assume "e>0"
-    from `?lhs` obtain d where "d>0" and d:"\<forall>x\<in>s. 0 < dist x a \<and> dist x a < d \<longrightarrow> dist (f x) (f a) < e" unfolding continuous_within Lim_within using `e>0` by auto
-    from x(2) `d>0` obtain N where N:"\<forall>n\<ge>N. dist (x n) a < d" by auto
-    hence "\<exists>N. \<forall>n\<ge>N. dist ((f \<circ> x) n) (f a) < e"
-      apply(rule_tac  x=N in exI) using N d  apply auto using x(1)
-      apply(erule_tac x=n in allE) apply(erule_tac x=n in allE)
-      apply(erule_tac x="x n" in ballE)  apply auto unfolding dist_nz[THEN sym] apply auto using `e>0` by auto
-  }
-  thus ?rhs unfolding continuous_within unfolding Lim_sequentially by simp
-next
-  assume ?rhs
-  { fix e::real assume "e>0"
-    assume "\<not> (\<exists>d>0. \<forall>x\<in>s. 0 < dist x a \<and> dist x a < d \<longrightarrow> dist (f x) (f a) < e)"
-    hence "\<forall>d. \<exists>x. d>0 \<longrightarrow> x\<in>s \<and> (0 < dist x a \<and> dist x a < d \<and> \<not> dist (f x) (f a) < e)" by blast
-    then obtain x where x:"\<forall>d>0. x d \<in> s \<and> (0 < dist (x d) a \<and> dist (x d) a < d \<and> \<not> dist (f (x d)) (f a) < e)"
-      using choice[of "\<lambda>d x.0<d \<longrightarrow> x\<in>s \<and> (0 < dist x a \<and> dist x a < d \<and> \<not> dist (f x) (f a) < e)"] by auto
-    { fix d::real assume "d>0"
-      hence "\<exists>N::nat. inverse (real (N + 1)) < d" using real_arch_inv[of d] by (auto, rule_tac x="n - 1" in exI)auto
-      then obtain N::nat where N:"inverse (real (N + 1)) < d" by auto
-      { fix n::nat assume n:"n\<ge>N"
-        hence "dist (x (inverse (real (n + 1)))) a < inverse (real (n + 1))" using x[THEN spec[where x="inverse (real (n + 1))"]] by auto
-        moreover have "inverse (real (n + 1)) < d" using N n by (auto, metis Suc_le_mono le_SucE less_imp_inverse_less nat_le_real_less order_less_trans real_of_nat_Suc real_of_nat_Suc_gt_zero)
-        ultimately have "dist (x (inverse (real (n + 1)))) a < d" by auto
-      }
-      hence "\<exists>N::nat. \<forall>n\<ge>N. dist (x (inverse (real (n + 1)))) a < d" by auto
-    }
-    hence "(\<forall>n::nat. x (inverse (real (n + 1))) \<in> s) \<and> (\<forall>e>0. \<exists>N::nat. \<forall>n\<ge>N. dist (x (inverse (real (n + 1)))) a < e)" using x by auto
-    hence "\<forall>e>0. \<exists>N::nat. \<forall>n\<ge>N. dist (f (x (inverse (real (n + 1))))) (f a) < e"  using `?rhs`[THEN spec[where x="\<lambda>n::nat. x (inverse (real (n+1)))"], unfolded Lim_sequentially] by auto
-    hence "False" apply(erule_tac x=e in allE) using `e>0` using x by auto
-  }
-  thus ?lhs  unfolding continuous_within unfolding Lim_within unfolding Lim_sequentially by blast
-qed
-
-lemma continuous_at_sequentially:
-  fixes f :: "'a::metric_space \<Rightarrow> 'b::metric_space"
-  shows "continuous (at a) f \<longleftrightarrow> (\<forall>x. (x ---> a) sequentially
-                  --> ((f o x) ---> f a) sequentially)"
-  using continuous_within_sequentially[of a UNIV f] unfolding within_UNIV by auto
-
-lemma continuous_on_sequentially:
- "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)" (is "?lhs = ?rhs")
-proof
-  assume ?rhs thus ?lhs using continuous_within_sequentially[of _ s f] unfolding continuous_on_eq_continuous_within by auto
-next
-  assume ?lhs thus ?rhs unfolding continuous_on_eq_continuous_within using continuous_within_sequentially[of _ s f] by auto
-qed
-
-lemma uniformly_continuous_on_sequentially:
-  fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
-  shows "uniformly_continuous_on s f \<longleftrightarrow> (\<forall>x y. (\<forall>n. x n \<in> s) \<and> (\<forall>n. y n \<in> s) \<and>
-                    ((\<lambda>n. x n - y n) ---> 0) sequentially
-                    \<longrightarrow> ((\<lambda>n. f(x n) - f(y n)) ---> 0) sequentially)" (is "?lhs = ?rhs")
-proof
-  assume ?lhs
-  { fix x y assume x:"\<forall>n. x n \<in> s" and y:"\<forall>n. y n \<in> s" and xy:"((\<lambda>n. x n - y n) ---> 0) sequentially"
-    { fix e::real assume "e>0"
-      then obtain d where "d>0" and d:"\<forall>x\<in>s. \<forall>x'\<in>s. dist x' x < d \<longrightarrow> dist (f x') (f x) < e"
-        using `?lhs`[unfolded uniformly_continuous_on_def, THEN spec[where x=e]] by auto
-      obtain N where N:"\<forall>n\<ge>N. norm (x n - y n - 0) < d" using xy[unfolded Lim_sequentially dist_norm] and `d>0` by auto
-      { fix n assume "n\<ge>N"
-        hence "norm (f (x n) - f (y n) - 0) < e"
-          using N[THEN spec[where x=n]] using d[THEN bspec[where x="x n"], THEN bspec[where x="y n"]] using x and y
-          unfolding dist_commute and dist_norm by simp  }
-      hence "\<exists>N. \<forall>n\<ge>N. norm (f (x n) - f (y n) - 0) < e"  by auto  }
-    hence "((\<lambda>n. f(x n) - f(y n)) ---> 0) sequentially" unfolding Lim_sequentially and dist_norm by auto  }
-  thus ?rhs by auto
-next
-  assume ?rhs
-  { assume "\<not> ?lhs"
-    then obtain e where "e>0" "\<forall>d>0. \<exists>x\<in>s. \<exists>x'\<in>s. dist x' x < d \<and> \<not> dist (f x') (f x) < e" unfolding uniformly_continuous_on_def by auto
-    then obtain fa where fa:"\<forall>x.  0 < x \<longrightarrow> fst (fa x) \<in> s \<and> snd (fa x) \<in> s \<and> dist (fst (fa x)) (snd (fa x)) < x \<and> \<not> dist (f (fst (fa x))) (f (snd (fa x))) < e"
-      using choice[of "\<lambda>d x. d>0 \<longrightarrow> fst x \<in> s \<and> snd x \<in> s \<and> dist (snd x) (fst x) < d \<and> \<not> dist (f (snd x)) (f (fst x)) < e"] unfolding Bex_def
-      by (auto simp add: dist_commute)
-    def x \<equiv> "\<lambda>n::nat. fst (fa (inverse (real n + 1)))"
-    def y \<equiv> "\<lambda>n::nat. snd (fa (inverse (real n + 1)))"
-    have xyn:"\<forall>n. x n \<in> s \<and> y n \<in> s" and xy0:"\<forall>n. dist (x n) (y n) < inverse (real n + 1)" and fxy:"\<forall>n. \<not> dist (f (x n)) (f (y n)) < e"
-      unfolding x_def and y_def using fa by auto
-    have 1:"\<And>(x::'a) y. dist (x - y) 0 = dist x y" unfolding dist_norm by auto
-    have 2:"\<And>(x::'b) y. dist (x - y) 0 = dist x y" unfolding dist_norm by auto
-    { fix e::real assume "e>0"
-      then obtain N::nat where "N \<noteq> 0" and N:"0 < inverse (real N) \<and> inverse (real N) < e" unfolding real_arch_inv[of e]   by auto
-      { fix n::nat assume "n\<ge>N"
-        hence "inverse (real n + 1) < inverse (real N)" using real_of_nat_ge_zero and `N\<noteq>0` by auto
-        also have "\<dots> < e" using N by auto
-        finally have "inverse (real n + 1) < e" by auto
-        hence "dist (x n - y n) 0 < e" unfolding 1 using xy0[THEN spec[where x=n]] by auto  }
-      hence "\<exists>N. \<forall>n\<ge>N. dist (x n - y n) 0 < e" by auto  }
-    hence "\<forall>e>0. \<exists>N. \<forall>n\<ge>N. dist (f (x n) - f (y n)) 0 < e" using `?rhs`[THEN spec[where x=x], THEN spec[where x=y]] and xyn unfolding Lim_sequentially by auto
-    hence False unfolding 2 using fxy and `e>0` by auto  }
-  thus ?lhs unfolding uniformly_continuous_on_def by blast
-qed
-
-text{* The usual transformation theorems. *}
-
-lemma continuous_transform_within:
-  fixes f g :: "'a::metric_space \<Rightarrow> 'b::metric_space"
-  assumes "0 < d" "x \<in> s" "\<forall>x' \<in> s. dist x' x < d --> f x' = g x'"
-          "continuous (at x within s) f"
-  shows "continuous (at x within s) g"
-proof-
-  { fix e::real assume "e>0"
-    then obtain d' where d':"d'>0" "\<forall>xa\<in>s. 0 < dist xa x \<and> dist xa x < d' \<longrightarrow> dist (f xa) (f x) < e" using assms(4) unfolding continuous_within Lim_within by auto
-    { fix x' assume "x'\<in>s" "0 < dist x' x" "dist x' x < (min d d')"
-      hence "dist (f x') (g x) < e" using assms(2,3) apply(erule_tac x=x in ballE) using d' by auto  }
-    hence "\<forall>xa\<in>s. 0 < dist xa x \<and> dist xa x < (min d d') \<longrightarrow> dist (f xa) (g x) < e" by blast
-    hence "\<exists>d>0. \<forall>xa\<in>s. 0 < dist xa x \<and> dist xa x < d \<longrightarrow> dist (f xa) (g x) < e" using `d>0` `d'>0` by(rule_tac x="min d d'" in exI)auto  }
-  hence "(f ---> g x) (at x within s)" unfolding Lim_within using assms(1) by auto
-  thus ?thesis unfolding continuous_within using Lim_transform_within[of d s x f g "g x"] using assms by blast
-qed
-
-lemma continuous_transform_at:
-  fixes f g :: "'a::metric_space \<Rightarrow> 'b::metric_space"
-  assumes "0 < d" "\<forall>x'. dist x' x < d --> f x' = g x'"
-          "continuous (at x) f"
-  shows "continuous (at x) g"
-proof-
-  { fix e::real assume "e>0"
-    then obtain d' where d':"d'>0" "\<forall>xa. 0 < dist xa x \<and> dist xa x < d' \<longrightarrow> dist (f xa) (f x) < e" using assms(3) unfolding continuous_at Lim_at by auto
-    { fix x' assume "0 < dist x' x" "dist x' x < (min d d')"
-      hence "dist (f x') (g x) < e" using assms(2) apply(erule_tac x=x in allE) using d' by auto
-    }
-    hence "\<forall>xa. 0 < dist xa x \<and> dist xa x < (min d d') \<longrightarrow> dist (f xa) (g x) < e" by blast
-    hence "\<exists>d>0. \<forall>xa. 0 < dist xa x \<and> dist xa x < d \<longrightarrow> dist (f xa) (g x) < e" using `d>0` `d'>0` by(rule_tac x="min d d'" in exI)auto
-  }
-  hence "(f ---> g x) (at x)" unfolding Lim_at using assms(1) by auto
-  thus ?thesis unfolding continuous_at using Lim_transform_at[of d x f g "g x"] using assms by blast
-qed
-
-text{* Combination results for pointwise continuity. *}
-
-lemma continuous_const: "continuous net (\<lambda>x. c)"
-  by (auto simp add: continuous_def Lim_const)
-
-lemma continuous_cmul:
-  fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector"
-  shows "continuous net f ==> continuous net (\<lambda>x. c *\<^sub>R f x)"
-  by (auto simp add: continuous_def Lim_cmul)
-
-lemma continuous_neg:
-  fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector"
-  shows "continuous net f ==> continuous net (\<lambda>x. -(f x))"
-  by (auto simp add: continuous_def Lim_neg)
-
-lemma continuous_add:
-  fixes f g :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector"
-  shows "continuous net f \<Longrightarrow> continuous net g \<Longrightarrow> continuous net (\<lambda>x. f x + g x)"
-  by (auto simp add: continuous_def Lim_add)
-
-lemma continuous_sub:
-  fixes f g :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector"
-  shows "continuous net f \<Longrightarrow> continuous net g \<Longrightarrow> continuous net (\<lambda>x. f x - g x)"
-  by (auto simp add: continuous_def Lim_sub)
-
-text{* Same thing for setwise continuity. *}
-
-lemma continuous_on_const:
- "continuous_on s (\<lambda>x. c)"
-  unfolding continuous_on_eq_continuous_within using continuous_const by blast
-
-lemma continuous_on_cmul:
-  fixes f :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector"
-  shows "continuous_on s f ==>  continuous_on s (\<lambda>x. c *\<^sub>R (f x))"
-  unfolding continuous_on_eq_continuous_within using continuous_cmul by blast
-
-lemma continuous_on_neg:
-  fixes f :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector"
-  shows "continuous_on s f \<Longrightarrow> continuous_on s (\<lambda>x. - f x)"
-  unfolding continuous_on_eq_continuous_within using continuous_neg by blast
-
-lemma continuous_on_add:
-  fixes f g :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector"
-  shows "continuous_on s f \<Longrightarrow> continuous_on s g
-           \<Longrightarrow> continuous_on s (\<lambda>x. f x + g x)"
-  unfolding continuous_on_eq_continuous_within using continuous_add by blast
-
-lemma continuous_on_sub:
-  fixes f g :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector"
-  shows "continuous_on s f \<Longrightarrow> continuous_on s g
-           \<Longrightarrow> continuous_on s (\<lambda>x. f x - g x)"
-  unfolding continuous_on_eq_continuous_within using continuous_sub by blast
-
-text{* Same thing for uniform continuity, using sequential formulations. *}
-
-lemma uniformly_continuous_on_const:
- "uniformly_continuous_on s (\<lambda>x. c)"
-  unfolding uniformly_continuous_on_def by simp
-
-lemma uniformly_continuous_on_cmul:
-  fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
-    (* FIXME: generalize 'a to metric_space *)
-  assumes "uniformly_continuous_on s f"
-  shows "uniformly_continuous_on s (\<lambda>x. c *\<^sub>R f(x))"
-proof-
-  { fix x y assume "((\<lambda>n. f (x n) - f (y n)) ---> 0) sequentially"
-    hence "((\<lambda>n. c *\<^sub>R f (x n) - c *\<^sub>R f (y n)) ---> 0) sequentially"
-      using Lim_cmul[of "(\<lambda>n. f (x n) - f (y n))" 0 sequentially c]
-      unfolding scaleR_zero_right scaleR_right_diff_distrib by auto
-  }
-  thus ?thesis using assms unfolding uniformly_continuous_on_sequentially by auto
-qed
-
-lemma dist_minus:
-  fixes x y :: "'a::real_normed_vector"
-  shows "dist (- x) (- y) = dist x y"
-  unfolding dist_norm minus_diff_minus norm_minus_cancel ..
-
-lemma uniformly_continuous_on_neg:
-  fixes f :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector"
-  shows "uniformly_continuous_on s f
-         ==> uniformly_continuous_on s (\<lambda>x. -(f x))"
-  unfolding uniformly_continuous_on_def dist_minus .
-
-lemma uniformly_continuous_on_add:
-  fixes f g :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector" (* FIXME: generalize 'a *)
-  assumes "uniformly_continuous_on s f" "uniformly_continuous_on s g"
-  shows "uniformly_continuous_on s (\<lambda>x. f x + g x)"
-proof-
-  {  fix x y assume "((\<lambda>n. f (x n) - f (y n)) ---> 0) sequentially"
-                    "((\<lambda>n. g (x n) - g (y n)) ---> 0) sequentially"
-    hence "((\<lambda>xa. f (x xa) - f (y xa) + (g (x xa) - g (y xa))) ---> 0 + 0) sequentially"
-      using Lim_add[of "\<lambda> n. f (x n) - f (y n)" 0  sequentially "\<lambda> n. g (x n) - g (y n)" 0] by auto
-    hence "((\<lambda>n. f (x n) + g (x n) - (f (y n) + g (y n))) ---> 0) sequentially" unfolding Lim_sequentially and add_diff_add [symmetric] by auto  }
-  thus ?thesis using assms unfolding uniformly_continuous_on_sequentially by auto
-qed
-
-lemma uniformly_continuous_on_sub:
-  fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector" (* FIXME: generalize 'a *)
-  shows "uniformly_continuous_on s f \<Longrightarrow> uniformly_continuous_on s g
-           ==> uniformly_continuous_on s  (\<lambda>x. f x - g x)"
-  unfolding ab_diff_minus
-  using uniformly_continuous_on_add[of s f "\<lambda>x. - g x"]
-  using uniformly_continuous_on_neg[of s g] by auto
-
-text{* Identity function is continuous in every sense. *}
-
-lemma continuous_within_id:
- "continuous (at a within s) (\<lambda>x. x)"
-  unfolding continuous_within by (rule Lim_at_within [OF Lim_ident_at])
-
-lemma continuous_at_id:
- "continuous (at a) (\<lambda>x. x)"
-  unfolding continuous_at by (rule Lim_ident_at)
-
-lemma continuous_on_id:
- "continuous_on s (\<lambda>x. x)"
-  unfolding continuous_on Lim_within by auto
-
-lemma uniformly_continuous_on_id:
- "uniformly_continuous_on s (\<lambda>x. x)"
-  unfolding uniformly_continuous_on_def by auto
-
-text{* Continuity of all kinds is preserved under composition. *}
-
-lemma continuous_within_compose:
-  fixes f :: "'a::metric_space \<Rightarrow> 'b::metric_space" (* FIXME: generalize *)
-  fixes g :: "'b::metric_space \<Rightarrow> 'c::metric_space"
-  assumes "continuous (at x within s) f"   "continuous (at (f x) within f ` s) g"
-  shows "continuous (at x within s) (g o f)"
-proof-
-  { fix e::real assume "e>0"
-    with assms(2)[unfolded continuous_within Lim_within] obtain d  where "d>0" and d:"\<forall>xa\<in>f ` s. 0 < dist xa (f x) \<and> dist xa (f x) < d \<longrightarrow> dist (g xa) (g (f x)) < e" by auto
-    from assms(1)[unfolded continuous_within Lim_within] obtain d' where "d'>0" and d':"\<forall>xa\<in>s. 0 < dist xa x \<and> dist xa x < d' \<longrightarrow> dist (f xa) (f x) < d" using `d>0` by auto
-    { fix y assume as:"y\<in>s"  "0 < dist y x"  "dist y x < d'"
-      hence "dist (f y) (f x) < d" using d'[THEN bspec[where x=y]] by (auto simp add:dist_commute)
-      hence "dist (g (f y)) (g (f x)) < e" using as(1) d[THEN bspec[where x="f y"]] unfolding dist_nz[THEN sym] using `e>0` by auto   }
-    hence "\<exists>d>0. \<forall>xa\<in>s. 0 < dist xa x \<and> dist xa x < d \<longrightarrow> dist (g (f xa)) (g (f x)) < e" using `d'>0` by auto  }
-  thus ?thesis unfolding continuous_within Lim_within by auto
-qed
-
-lemma continuous_at_compose:
-  fixes f :: "'a::metric_space \<Rightarrow> 'b::metric_space" (* FIXME: generalize *)
-  fixes g :: "'b::metric_space \<Rightarrow> 'c::metric_space"
-  assumes "continuous (at x) f"  "continuous (at (f x)) g"
-  shows "continuous (at x) (g o f)"
-proof-
-  have " continuous (at (f x) within range f) g" using assms(2) using continuous_within_subset[of "f x" UNIV g "range f", unfolded within_UNIV] by auto
-  thus ?thesis using assms(1) using continuous_within_compose[of x UNIV f g, unfolded within_UNIV] by auto
-qed
-
-lemma continuous_on_compose:
- "continuous_on s f \<Longrightarrow> continuous_on (f ` s) g \<Longrightarrow> continuous_on s (g o f)"
-  unfolding continuous_on_eq_continuous_within using continuous_within_compose[of _ s f g] by auto
-
-lemma uniformly_continuous_on_compose:
-  assumes "uniformly_continuous_on s f"  "uniformly_continuous_on (f ` s) g"
-  shows "uniformly_continuous_on s (g o f)"
-proof-
-  { fix e::real assume "e>0"
-    then obtain d where "d>0" and d:"\<forall>x\<in>f ` s. \<forall>x'\<in>f ` s. dist x' x < d \<longrightarrow> dist (g x') (g x) < e" using assms(2) unfolding uniformly_continuous_on_def by auto
-    obtain d' where "d'>0" "\<forall>x\<in>s. \<forall>x'\<in>s. dist x' x < d' \<longrightarrow> dist (f x') (f x) < d" using `d>0` using assms(1) unfolding uniformly_continuous_on_def by auto
-    hence "\<exists>d>0. \<forall>x\<in>s. \<forall>x'\<in>s. dist x' x < d \<longrightarrow> dist ((g \<circ> f) x') ((g \<circ> f) x) < e" using `d>0` using d by auto  }
-  thus ?thesis using assms unfolding uniformly_continuous_on_def by auto
-qed
-
-text{* Continuity in terms of open preimages. *}
-
-lemma continuous_at_open:
-  fixes f :: "'a::metric_space \<Rightarrow> 'b::metric_space" (* FIXME: generalize *)
-  shows "continuous (at x) f \<longleftrightarrow> (\<forall>t. open t \<and> f x \<in> t --> (\<exists>s. open s \<and> x \<in> s \<and> (\<forall>x' \<in> s. (f x') \<in> t)))" (is "?lhs = ?rhs")
-proof
-  assume ?lhs
-  { fix t assume as: "open t" "f x \<in> t"
-    then obtain e where "e>0" and e:"ball (f x) e \<subseteq> t" unfolding open_contains_ball by auto
-
-    obtain d where "d>0" and d:"\<forall>y. 0 < dist y x \<and> dist y x < d \<longrightarrow> dist (f y) (f x) < e" using `e>0` using `?lhs`[unfolded continuous_at Lim_at open_dist] by auto
-
-    have "open (ball x d)" using open_ball by auto
-    moreover have "x \<in> ball x d" unfolding centre_in_ball using `d>0` by simp
-    moreover
-    { fix x' assume "x'\<in>ball x d" hence "f x' \<in> t"
-        using e[unfolded subset_eq Ball_def mem_ball, THEN spec[where x="f x'"]]    d[THEN spec[where x=x']]
-        unfolding mem_ball apply (auto simp add: dist_commute)
-        unfolding dist_nz[THEN sym] using as(2) by auto  }
-    hence "\<forall>x'\<in>ball x d. f x' \<in> t" by auto
-    ultimately have "\<exists>s. open s \<and> x \<in> s \<and> (\<forall>x'\<in>s. f x' \<in> t)"
-      apply(rule_tac x="ball x d" in exI) by simp  }
-  thus ?rhs by auto
-next
-  assume ?rhs
-  { fix e::real assume "e>0"
-    then obtain s where s: "open s"  "x \<in> s"  "\<forall>x'\<in>s. f x' \<in> ball (f x) e" using `?rhs`[unfolded continuous_at Lim_at, THEN spec[where x="ball (f x) e"]]
-      unfolding centre_in_ball[of "f x" e, THEN sym] by auto
-    then obtain d where "d>0" and d:"ball x d \<subseteq> s" unfolding open_contains_ball by auto
-    { fix y assume "0 < dist y x \<and> dist y x < d"
-      hence "dist (f y) (f x) < e" using d[unfolded subset_eq Ball_def mem_ball, THEN spec[where x=y]]
-        using s(3)[THEN bspec[where x=y], unfolded mem_ball] by (auto simp add: dist_commute)  }
-    hence "\<exists>d>0. \<forall>xa. 0 < dist xa x \<and> dist xa x < d \<longrightarrow> dist (f xa) (f x) < e" using `d>0` by auto  }
-  thus ?lhs unfolding continuous_at Lim_at by auto
-qed
-
-lemma continuous_on_open:
- "continuous_on s f \<longleftrightarrow>
-        (\<forall>t. openin (subtopology euclidean (f ` s)) t
-            --> openin (subtopology euclidean s) {x \<in> s. f x \<in> t})" (is "?lhs = ?rhs")
-proof
-  assume ?lhs
-  { fix t assume as:"openin (subtopology euclidean (f ` s)) t"
-    have "{x \<in> s. f x \<in> t} \<subseteq> s" using as[unfolded openin_euclidean_subtopology_iff] by auto
-    moreover
-    { fix x assume as':"x\<in>{x \<in> s. f x \<in> t}"
-      then obtain e where e: "e>0" "\<forall>x'\<in>f ` s. dist x' (f x) < e \<longrightarrow> x' \<in> t" using as[unfolded openin_euclidean_subtopology_iff, THEN conjunct2, THEN bspec[where x="f x"]] by auto
-      from this(1) obtain d where d: "d>0" "\<forall>xa\<in>s. 0 < dist xa x \<and> dist xa x < d \<longrightarrow> dist (f xa) (f x) < e" using `?lhs`[unfolded continuous_on Lim_within, THEN bspec[where x=x]] using as' by auto
-      have "\<exists>e>0. \<forall>x'\<in>s. dist x' x < e \<longrightarrow> x' \<in> {x \<in> s. f x \<in> t}" using d e unfolding dist_nz[THEN sym] by (rule_tac x=d in exI, auto)  }
-    ultimately have "openin (subtopology euclidean s) {x \<in> s. f x \<in> t}" unfolding openin_euclidean_subtopology_iff by auto  }
-  thus ?rhs unfolding continuous_on Lim_within using openin by auto
-next
-  assume ?rhs
-  { fix e::real and x assume "x\<in>s" "e>0"
-    { fix xa x' assume "dist (f xa) (f x) < e" "xa \<in> s" "x' \<in> s" "dist (f xa) (f x') < e - dist (f xa) (f x)"
-      hence "dist (f x') (f x) < e" using dist_triangle[of "f x'" "f x" "f xa"]
-        by (auto simp add: dist_commute)  }
-    hence "ball (f x) e \<inter> f ` s \<subseteq> f ` s \<and> (\<forall>xa\<in>ball (f x) e \<inter> f ` s. \<exists>ea>0. \<forall>x'\<in>f ` s. dist x' xa < ea \<longrightarrow> x' \<in> ball (f x) e \<inter> f ` s)" apply auto
-      apply(rule_tac x="e - dist (f xa) (f x)" in exI) using `e>0` by (auto simp add: dist_commute)
-    hence "\<forall>xa\<in>{xa \<in> s. f xa \<in> ball (f x) e \<inter> f ` s}. \<exists>ea>0. \<forall>x'\<in>s. dist x' xa < ea \<longrightarrow> x' \<in> {xa \<in> s. f xa \<in> ball (f x) e \<inter> f ` s}"
-      using `?rhs`[unfolded openin_euclidean_subtopology_iff, THEN spec[where x="ball (f x) e \<inter> f ` s"]] by auto
-    hence "\<exists>d>0. \<forall>xa\<in>s. 0 < dist xa x \<and> dist xa x < d \<longrightarrow> dist (f xa) (f x) < e" apply(erule_tac x=x in ballE) apply auto using `e>0` `x\<in>s` by (auto simp add: dist_commute)  }
-  thus ?lhs unfolding continuous_on Lim_within by auto
-qed
-
-(* ------------------------------------------------------------------------- *)
-(* Similarly in terms of closed sets.                                        *)
-(* ------------------------------------------------------------------------- *)
-
-lemma continuous_on_closed:
- "continuous_on s f \<longleftrightarrow>  (\<forall>t. closedin (subtopology euclidean (f ` s)) t  --> closedin (subtopology euclidean s) {x \<in> s. f x \<in> t})" (is "?lhs = ?rhs")
-proof
-  assume ?lhs
-  { fix t
-    have *:"s - {x \<in> s. f x \<in> f ` s - t} = {x \<in> s. f x \<in> t}" by auto
-    have **:"f ` s - (f ` s - (f ` s - t)) = f ` s - t" by auto
-    assume as:"closedin (subtopology euclidean (f ` s)) t"
-    hence "closedin (subtopology euclidean (f ` s)) (f ` s - (f ` s - t))" unfolding closedin_def topspace_euclidean_subtopology unfolding ** by auto
-    hence "closedin (subtopology euclidean s) {x \<in> s. f x \<in> t}" using `?lhs`[unfolded continuous_on_open, THEN spec[where x="(f ` s) - t"]]
-      unfolding openin_closedin_eq topspace_euclidean_subtopology unfolding * by auto  }
-  thus ?rhs by auto
-next
-  assume ?rhs
-  { fix t
-    have *:"s - {x \<in> s. f x \<in> f ` s - t} = {x \<in> s. f x \<in> t}" by auto
-    assume as:"openin (subtopology euclidean (f ` s)) t"
-    hence "openin (subtopology euclidean s) {x \<in> s. f x \<in> t}" using `?rhs`[THEN spec[where x="(f ` s) - t"]]
-      unfolding openin_closedin_eq topspace_euclidean_subtopology *[THEN sym] closedin_subtopology by auto }
-  thus ?lhs unfolding continuous_on_open by auto
-qed
-
-text{* Half-global and completely global cases.                                  *}
-
-lemma continuous_open_in_preimage:
-  assumes "continuous_on s f"  "open t"
-  shows "openin (subtopology euclidean s) {x \<in> s. f x \<in> t}"
-proof-
-  have *:"\<forall>x. x \<in> s \<and> f x \<in> t \<longleftrightarrow> x \<in> s \<and> f x \<in> (t \<inter> f ` s)" by auto
-  have "openin (subtopology euclidean (f ` s)) (t \<inter> f ` s)"
-    using openin_open_Int[of t "f ` s", OF assms(2)] unfolding openin_open by auto
-  thus ?thesis using assms(1)[unfolded continuous_on_open, THEN spec[where x="t \<inter> f ` s"]] using * by auto
-qed
-
-lemma continuous_closed_in_preimage:
-  assumes "continuous_on s f"  "closed t"
-  shows "closedin (subtopology euclidean s) {x \<in> s. f x \<in> t}"
-proof-
-  have *:"\<forall>x. x \<in> s \<and> f x \<in> t \<longleftrightarrow> x \<in> s \<and> f x \<in> (t \<inter> f ` s)" by auto
-  have "closedin (subtopology euclidean (f ` s)) (t \<inter> f ` s)"
-    using closedin_closed_Int[of t "f ` s", OF assms(2)] unfolding Int_commute by auto
-  thus ?thesis
-    using assms(1)[unfolded continuous_on_closed, THEN spec[where x="t \<inter> f ` s"]] using * by auto
-qed
-
-lemma continuous_open_preimage:
-  assumes "continuous_on s f" "open s" "open t"
-  shows "open {x \<in> s. f x \<in> t}"
-proof-
-  obtain T where T: "open T" "{x \<in> s. f x \<in> t} = s \<inter> T"
-    using continuous_open_in_preimage[OF assms(1,3)] unfolding openin_open by auto
-  thus ?thesis using open_Int[of s T, OF assms(2)] by auto
-qed
-
-lemma continuous_closed_preimage:
-  assumes "continuous_on s f" "closed s" "closed t"
-  shows "closed {x \<in> s. f x \<in> t}"
-proof-
-  obtain T where T: "closed T" "{x \<in> s. f x \<in> t} = s \<inter> T"
-    using continuous_closed_in_preimage[OF assms(1,3)] unfolding closedin_closed by auto
-  thus ?thesis using closed_Int[of s T, OF assms(2)] by auto
-qed
-
-lemma continuous_open_preimage_univ:
-  fixes f :: "'a::metric_space \<Rightarrow> 'b::metric_space" (* FIXME: generalize *)
-  shows "\<forall>x. continuous (at x) f \<Longrightarrow> open s \<Longrightarrow> open {x. f x \<in> s}"
-  using continuous_open_preimage[of UNIV f s] open_UNIV continuous_at_imp_continuous_on by auto
-
-lemma continuous_closed_preimage_univ:
-  fixes f :: "'a::metric_space \<Rightarrow> 'b::metric_space" (* FIXME: generalize *)
-  shows "(\<forall>x. continuous (at x) f) \<Longrightarrow> closed s ==> closed {x. f x \<in> s}"
-  using continuous_closed_preimage[of UNIV f s] closed_UNIV continuous_at_imp_continuous_on by auto
-
-lemma continuous_open_vimage:
-  fixes f :: "'a::metric_space \<Rightarrow> 'b::metric_space" (* FIXME: generalize *)
-  shows "\<forall>x. continuous (at x) f \<Longrightarrow> open s \<Longrightarrow> open (f -` s)"
-  unfolding vimage_def by (rule continuous_open_preimage_univ)
-
-lemma continuous_closed_vimage:
-  fixes f :: "'a::metric_space \<Rightarrow> 'b::metric_space" (* FIXME: generalize *)
-  shows "\<forall>x. continuous (at x) f \<Longrightarrow> closed s \<Longrightarrow> closed (f -` s)"
-  unfolding vimage_def by (rule continuous_closed_preimage_univ)
-
-text{* Equality of continuous functions on closure and related results.          *}
-
-lemma continuous_closed_in_preimage_constant:
- "continuous_on s f ==> closedin (subtopology euclidean s) {x \<in> s. f x = a}"
-  using continuous_closed_in_preimage[of s f "{a}"] closed_sing by auto
-
-lemma continuous_closed_preimage_constant:
- "continuous_on s f \<Longrightarrow> closed s ==> closed {x \<in> s. f x = a}"
-  using continuous_closed_preimage[of s f "{a}"] closed_sing by auto
-
-lemma continuous_constant_on_closure:
-  assumes "continuous_on (closure s) f"
-          "\<forall>x \<in> s. f x = a"
-  shows "\<forall>x \<in> (closure s). f x = a"
-    using continuous_closed_preimage_constant[of "closure s" f a]
-    assms closure_minimal[of s "{x \<in> closure s. f x = a}"] closure_subset unfolding subset_eq by auto
-
-lemma image_closure_subset:
-  assumes "continuous_on (closure s) f"  "closed t"  "(f ` s) \<subseteq> t"
-  shows "f ` (closure s) \<subseteq> t"
-proof-
-  have "s \<subseteq> {x \<in> closure s. f x \<in> t}" using assms(3) closure_subset by auto
-  moreover have "closed {x \<in> closure s. f x \<in> t}"
-    using continuous_closed_preimage[OF assms(1)] and assms(2) by auto
-  ultimately have "closure s = {x \<in> closure s . f x \<in> t}"
-    using closure_minimal[of s "{x \<in> closure s. f x \<in> t}"] by auto
-  thus ?thesis by auto
-qed
-
-lemma continuous_on_closure_norm_le:
-  fixes f :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector"
-  assumes "continuous_on (closure s) f"  "\<forall>y \<in> s. norm(f y) \<le> b"  "x \<in> (closure s)"
-  shows "norm(f x) \<le> b"
-proof-
-  have *:"f ` s \<subseteq> cball 0 b" using assms(2)[unfolded mem_cball_0[THEN sym]] by auto
-  show ?thesis
-    using image_closure_subset[OF assms(1) closed_cball[of 0 b] *] assms(3)
-    unfolding subset_eq apply(erule_tac x="f x" in ballE) by (auto simp add: dist_norm)
-qed
-
-text{* Making a continuous function avoid some value in a neighbourhood.         *}
-
-lemma continuous_within_avoid:
-  fixes f :: "'a::metric_space \<Rightarrow> 'b::metric_space" (* FIXME: generalize *)
-  assumes "continuous (at x within s) f"  "x \<in> s"  "f x \<noteq> a"
-  shows "\<exists>e>0. \<forall>y \<in> s. dist x y < e --> f y \<noteq> a"
-proof-
-  obtain d where "d>0" and d:"\<forall>xa\<in>s. 0 < dist xa x \<and> dist xa x < d \<longrightarrow> dist (f xa) (f x) < dist (f x) a"
-    using assms(1)[unfolded continuous_within Lim_within, THEN spec[where x="dist (f x) a"]] assms(3)[unfolded dist_nz] by auto
-  { fix y assume " y\<in>s"  "dist x y < d"
-    hence "f y \<noteq> a" using d[THEN bspec[where x=y]] assms(3)[unfolded dist_nz]
-      apply auto unfolding dist_nz[THEN sym] by (auto simp add: dist_commute) }
-  thus ?thesis using `d>0` by auto
-qed
-
-lemma continuous_at_avoid:
-  fixes f :: "'a::metric_space \<Rightarrow> 'b::metric_space" (* FIXME: generalize *)
-  assumes "continuous (at x) f"  "f x \<noteq> a"
-  shows "\<exists>e>0. \<forall>y. dist x y < e \<longrightarrow> f y \<noteq> a"
-using assms using continuous_within_avoid[of x UNIV f a, unfolded within_UNIV] by auto
-
-lemma continuous_on_avoid:
-  assumes "continuous_on s f"  "x \<in> s"  "f x \<noteq> a"
-  shows "\<exists>e>0. \<forall>y \<in> s. dist x y < e \<longrightarrow> f y \<noteq> a"
-using assms(1)[unfolded continuous_on_eq_continuous_within, THEN bspec[where x=x], OF assms(2)]  continuous_within_avoid[of x s f a]  assms(2,3) by auto
-
-lemma continuous_on_open_avoid:
-  assumes "continuous_on s f"  "open s"  "x \<in> s"  "f x \<noteq> a"
-  shows "\<exists>e>0. \<forall>y. dist x y < e \<longrightarrow> f y \<noteq> a"
-using assms(1)[unfolded continuous_on_eq_continuous_at[OF assms(2)], THEN bspec[where x=x], OF assms(3)]  continuous_at_avoid[of x f a]  assms(3,4) by auto
-
-text{* Proving a function is constant by proving open-ness of level set.         *}
-
-lemma continuous_levelset_open_in_cases:
- "connected s \<Longrightarrow> continuous_on s f \<Longrightarrow>
-        openin (subtopology euclidean s) {x \<in> s. f x = a}
-        ==> (\<forall>x \<in> s. f x \<noteq> a) \<or> (\<forall>x \<in> s. f x = a)"
-unfolding connected_clopen using continuous_closed_in_preimage_constant by auto
-
-lemma continuous_levelset_open_in:
- "connected s \<Longrightarrow> continuous_on s f \<Longrightarrow>
-        openin (subtopology euclidean s) {x \<in> s. f x = a} \<Longrightarrow>
-        (\<exists>x \<in> s. f x = a)  ==> (\<forall>x \<in> s. f x = a)"
-using continuous_levelset_open_in_cases[of s f ]
-by meson
-
-lemma continuous_levelset_open:
-  assumes "connected s"  "continuous_on s f"  "open {x \<in> s. f x = a}"  "\<exists>x \<in> s.  f x = a"
-  shows "\<forall>x \<in> s. f x = a"
-using continuous_levelset_open_in[OF assms(1,2), of a, unfolded openin_open] using assms (3,4) by auto
-
-text{* Some arithmetical combinations (more to prove).                           *}
-
-lemma open_scaling[intro]:
-  fixes s :: "'a::real_normed_vector set"
-  assumes "c \<noteq> 0"  "open s"
-  shows "open((\<lambda>x. c *\<^sub>R x) ` s)"
-proof-
-  { fix x assume "x \<in> s"
-    then obtain e where "e>0" and e:"\<forall>x'. dist x' x < e \<longrightarrow> x' \<in> s" using assms(2)[unfolded open_dist, THEN bspec[where x=x]] by auto
-    have "e * abs c > 0" using assms(1)[unfolded zero_less_abs_iff[THEN sym]] using real_mult_order[OF `e>0`] by auto
-    moreover
-    { fix y assume "dist y (c *\<^sub>R x) < e * \<bar>c\<bar>"
-      hence "norm ((1 / c) *\<^sub>R y - x) < e" unfolding dist_norm
-        using norm_scaleR[of c "(1 / c) *\<^sub>R y - x", unfolded scaleR_right_diff_distrib, unfolded scaleR_scaleR] assms(1)
-          assms(1)[unfolded zero_less_abs_iff[THEN sym]] by (simp del:zero_less_abs_iff)
-      hence "y \<in> op *\<^sub>R c ` s" using rev_image_eqI[of "(1 / c) *\<^sub>R y" s y "op *\<^sub>R c"]  e[THEN spec[where x="(1 / c) *\<^sub>R y"]]  assms(1) unfolding dist_norm scaleR_scaleR by auto  }
-    ultimately have "\<exists>e>0. \<forall>x'. dist x' (c *\<^sub>R x) < e \<longrightarrow> x' \<in> op *\<^sub>R c ` s" apply(rule_tac x="e * abs c" in exI) by auto  }
-  thus ?thesis unfolding open_dist by auto
-qed
-
-lemma minus_image_eq_vimage:
-  fixes A :: "'a::ab_group_add set"
-  shows "(\<lambda>x. - x) ` A = (\<lambda>x. - x) -` A"
-  by (auto intro!: image_eqI [where f="\<lambda>x. - x"])
-
-lemma open_negations:
-  fixes s :: "'a::real_normed_vector set"
-  shows "open s ==> open ((\<lambda> x. -x) ` s)"
-  unfolding scaleR_minus1_left [symmetric]
-  by (rule open_scaling, auto)
-
-lemma open_translation:
-  fixes s :: "'a::real_normed_vector set"
-  assumes "open s"  shows "open((\<lambda>x. a + x) ` s)"
-proof-
-  { fix x have "continuous (at x) (\<lambda>x. x - a)" using continuous_sub[of "at x" "\<lambda>x. x" "\<lambda>x. a"] continuous_at_id[of x] continuous_const[of "at x" a] by auto  }
-  moreover have "{x. x - a \<in> s}  = op + a ` s" apply auto unfolding image_iff apply(rule_tac x="x - a" in bexI) by auto
-  ultimately show ?thesis using continuous_open_preimage_univ[of "\<lambda>x. x - a" s] using assms by auto
-qed
-
-lemma open_affinity:
-  fixes s :: "'a::real_normed_vector set"
-  assumes "open s"  "c \<noteq> 0"
-  shows "open ((\<lambda>x. a + c *\<^sub>R x) ` s)"
-proof-
-  have *:"(\<lambda>x. a + c *\<^sub>R x) = (\<lambda>x. a + x) \<circ> (\<lambda>x. c *\<^sub>R x)" unfolding o_def ..
-  have "op + a ` op *\<^sub>R c ` s = (op + a \<circ> op *\<^sub>R c) ` s" by auto
-  thus ?thesis using assms open_translation[of "op *\<^sub>R c ` s" a] unfolding * by auto
-qed
-
-lemma interior_translation:
-  fixes s :: "'a::real_normed_vector set"
-  shows "interior ((\<lambda>x. a + x) ` s) = (\<lambda>x. a + x) ` (interior s)"
-proof (rule set_ext, rule)
-  fix x assume "x \<in> interior (op + a ` s)"
-  then obtain e where "e>0" and e:"ball x e \<subseteq> op + a ` s" unfolding mem_interior by auto
-  hence "ball (x - a) e \<subseteq> s" unfolding subset_eq Ball_def mem_ball dist_norm apply auto apply(erule_tac x="a + xa" in allE) unfolding ab_group_add_class.diff_diff_eq[THEN sym] by auto
-  thus "x \<in> op + a ` interior s" unfolding image_iff apply(rule_tac x="x - a" in bexI) unfolding mem_interior using `e > 0` by auto
-next
-  fix x assume "x \<in> op + a ` interior s"
-  then obtain y e where "e>0" and e:"ball y e \<subseteq> s" and y:"x = a + y" unfolding image_iff Bex_def mem_interior by auto
-  { fix z have *:"a + y - z = y + a - z" by auto
-    assume "z\<in>ball x e"
-    hence "z - a \<in> s" using e[unfolded subset_eq, THEN bspec[where x="z - a"]] unfolding mem_ball dist_norm y ab_group_add_class.diff_diff_eq2 * by auto
-    hence "z \<in> op + a ` s" unfolding image_iff by(auto intro!: bexI[where x="z - a"])  }
-  hence "ball x e \<subseteq> op + a ` s" unfolding subset_eq by auto
-  thus "x \<in> interior (op + a ` s)" unfolding mem_interior using `e>0` by auto
-qed
-
-subsection {* Preservation of compactness and connectedness under continuous function.  *}
-
-lemma compact_continuous_image:
-  assumes "continuous_on s f"  "compact s"
-  shows "compact(f ` s)"
-proof-
-  { fix x assume x:"\<forall>n::nat. x n \<in> f ` s"
-    then obtain y where y:"\<forall>n. y n \<in> s \<and> x n = f (y n)" unfolding image_iff Bex_def using choice[of "\<lambda>n xa. xa \<in> s \<and> x n = f xa"] by auto
-    then obtain l r where "l\<in>s" and r:"subseq r" and lr:"((y \<circ> r) ---> l) sequentially" using assms(2)[unfolded compact_def, THEN spec[where x=y]] by auto
-    { fix e::real assume "e>0"
-      then obtain d where "d>0" and d:"\<forall>x'\<in>s. dist x' l < d \<longrightarrow> dist (f x') (f l) < e" using assms(1)[unfolded continuous_on_def, THEN bspec[where x=l], OF `l\<in>s`] by auto
-      then obtain N::nat where N:"\<forall>n\<ge>N. dist ((y \<circ> r) n) l < d" using lr[unfolded Lim_sequentially, THEN spec[where x=d]] by auto
-      { fix n::nat assume "n\<ge>N" hence "dist ((x \<circ> r) n) (f l) < e" using N[THEN spec[where x=n]] d[THEN bspec[where x="y (r n)"]] y[THEN spec[where x="r n"]] by auto  }
-      hence "\<exists>N. \<forall>n\<ge>N. dist ((x \<circ> r) n) (f l) < e" by auto  }
-    hence "\<exists>l\<in>f ` s. \<exists>r. subseq r \<and> ((x \<circ> r) ---> l) sequentially" unfolding Lim_sequentially using r lr `l\<in>s` by auto  }
-  thus ?thesis unfolding compact_def by auto
-qed
-
-lemma connected_continuous_image:
-  assumes "continuous_on s f"  "connected s"
-  shows "connected(f ` s)"
-proof-
-  { fix T assume as: "T \<noteq> {}"  "T \<noteq> f ` s"  "openin (subtopology euclidean (f ` s)) T"  "closedin (subtopology euclidean (f ` s)) T"
-    have "{x \<in> s. f x \<in> T} = {} \<or> {x \<in> s. f x \<in> T} = s"
-      using assms(1)[unfolded continuous_on_open, THEN spec[where x=T]]
-      using assms(1)[unfolded continuous_on_closed, THEN spec[where x=T]]
-      using assms(2)[unfolded connected_clopen, THEN spec[where x="{x \<in> s. f x \<in> T}"]] as(3,4) by auto
-    hence False using as(1,2)
-      using as(4)[unfolded closedin_def topspace_euclidean_subtopology] by auto }
-  thus ?thesis unfolding connected_clopen by auto
-qed
-
-text{* Continuity implies uniform continuity on a compact domain.                *}
-
-lemma compact_uniformly_continuous:
-  assumes "continuous_on s f"  "compact s"
-  shows "uniformly_continuous_on s f"
-proof-
-    { fix x assume x:"x\<in>s"
-      hence "\<forall>xa. \<exists>y. 0 < xa \<longrightarrow> (y > 0 \<and> (\<forall>x'\<in>s. dist x' x < y \<longrightarrow> dist (f x') (f x) < xa))" using assms(1)[unfolded continuous_on_def, THEN bspec[where x=x]] by auto
-      hence "\<exists>fa. \<forall>xa>0. \<forall>x'\<in>s. fa xa > 0 \<and> (dist x' x < fa xa \<longrightarrow> dist (f x') (f x) < xa)" using choice[of "\<lambda>e d. e>0 \<longrightarrow> d>0 \<and>(\<forall>x'\<in>s. (dist x' x < d \<longrightarrow> dist (f x') (f x) < e))"] by auto  }
-    then have "\<forall>x\<in>s. \<exists>y. \<forall>xa. 0 < xa \<longrightarrow> (\<forall>x'\<in>s. y xa > 0 \<and> (dist x' x < y xa \<longrightarrow> dist (f x') (f x) < xa))" by auto
-    then obtain d where d:"\<forall>e>0. \<forall>x\<in>s. \<forall>x'\<in>s. d x e > 0 \<and> (dist x' x < d x e \<longrightarrow> dist (f x') (f x) < e)"
-      using bchoice[of s "\<lambda>x fa. \<forall>xa>0. \<forall>x'\<in>s. fa xa > 0 \<and> (dist x' x < fa xa \<longrightarrow> dist (f x') (f x) < xa)"] by blast
-
-  { fix e::real assume "e>0"
-
-    { fix x assume "x\<in>s" hence "x \<in> ball x (d x (e / 2))" unfolding centre_in_ball using d[THEN spec[where x="e/2"]] using `e>0` by auto  }
-    hence "s \<subseteq> \<Union>{ball x (d x (e / 2)) |x. x \<in> s}" unfolding subset_eq by auto
-    moreover
-    { fix b assume "b\<in>{ball x (d x (e / 2)) |x. x \<in> s}" hence "open b" by auto  }
-    ultimately obtain ea where "ea>0" and ea:"\<forall>x\<in>s. \<exists>b\<in>{ball x (d x (e / 2)) |x. x \<in> s}. ball x ea \<subseteq> b" using heine_borel_lemma[OF assms(2), of "{ball x (d x (e / 2)) | x. x\<in>s }"] by auto
-
-    { fix x y assume "x\<in>s" "y\<in>s" and as:"dist y x < ea"
-      obtain z where "z\<in>s" and z:"ball x ea \<subseteq> ball z (d z (e / 2))" using ea[THEN bspec[where x=x]] and `x\<in>s` by auto
-      hence "x\<in>ball z (d z (e / 2))" using `ea>0` unfolding subset_eq by auto
-      hence "dist (f z) (f x) < e / 2" using d[THEN spec[where x="e/2"]] and `e>0` and `x\<in>s` and `z\<in>s`
-        by (auto  simp add: dist_commute)
-      moreover have "y\<in>ball z (d z (e / 2))" using as and `ea>0` and z[unfolded subset_eq]
-        by (auto simp add: dist_commute)
-      hence "dist (f z) (f y) < e / 2" using d[THEN spec[where x="e/2"]] and `e>0` and `y\<in>s` and `z\<in>s`
-        by (auto  simp add: dist_commute)
-      ultimately have "dist (f y) (f x) < e" using dist_triangle_half_r[of "f z" "f x" e "f y"]
-        by (auto simp add: dist_commute)  }
-    then have "\<exists>d>0. \<forall>x\<in>s. \<forall>x'\<in>s. dist x' x < d \<longrightarrow> dist (f x') (f x) < e" using `ea>0` by auto  }
-  thus ?thesis unfolding uniformly_continuous_on_def by auto
-qed
-
-text{* Continuity of inverse function on compact domain. *}
-
-lemma continuous_on_inverse:
-  fixes f :: "'a::heine_borel \<Rightarrow> 'b::heine_borel"
-    (* TODO: can this be generalized more? *)
-  assumes "continuous_on s f"  "compact s"  "\<forall>x \<in> s. g (f x) = x"
-  shows "continuous_on (f ` s) g"
-proof-
-  have *:"g ` f ` s = s" using assms(3) by (auto simp add: image_iff)
-  { fix t assume t:"closedin (subtopology euclidean (g ` f ` s)) t"
-    then obtain T where T: "closed T" "t = s \<inter> T" unfolding closedin_closed unfolding * by auto
-    have "continuous_on (s \<inter> T) f" using continuous_on_subset[OF assms(1), of "s \<inter> t"]
-      unfolding T(2) and Int_left_absorb by auto
-    moreover have "compact (s \<inter> T)"
-      using assms(2) unfolding compact_eq_bounded_closed
-      using bounded_subset[of s "s \<inter> T"] and T(1) by auto
-    ultimately have "closed (f ` t)" using T(1) unfolding T(2)
-      using compact_continuous_image [of "s \<inter> T" f] unfolding compact_eq_bounded_closed by auto
-    moreover have "{x \<in> f ` s. g x \<in> t} = f ` s \<inter> f ` t" using assms(3) unfolding T(2) by auto
-    ultimately have "closedin (subtopology euclidean (f ` s)) {x \<in> f ` s. g x \<in> t}"
-      unfolding closedin_closed by auto  }
-  thus ?thesis unfolding continuous_on_closed by auto
-qed
-
-subsection{* A uniformly convergent limit of continuous functions is continuous.       *}
-
-lemma norm_triangle_lt:
-  fixes x y :: "'a::real_normed_vector"
-  shows "norm x + norm y < e \<Longrightarrow> norm (x + y) < e"
-by (rule le_less_trans [OF norm_triangle_ineq])
-
-lemma continuous_uniform_limit:
-  fixes f :: "'a \<Rightarrow> 'b::metric_space \<Rightarrow> 'c::real_normed_vector"
-  assumes "\<not> (trivial_limit net)"  "eventually (\<lambda>n. continuous_on s (f n)) net"
-  "\<forall>e>0. eventually (\<lambda>n. \<forall>x \<in> s. norm(f n x - g x) < e) net"
-  shows "continuous_on s g"
-proof-
-  { fix x and e::real assume "x\<in>s" "e>0"
-    have "eventually (\<lambda>n. \<forall>x\<in>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:"\<forall>xa\<in>s. norm (f n xa - g xa) < e / 3"  "continuous_on s (f n)"
-      using eventually_and[of "(\<lambda>n. \<forall>x\<in>s. norm (f n x - g x) < e / 3)" "(\<lambda>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:"\<forall>x'\<in>s. dist x' x < d \<longrightarrow> dist (f n x') (f n x) < e / 3"
-      using n(2)[unfolded continuous_on_def, THEN bspec[where x=x], OF `x\<in>s`, THEN spec[where x="e/3"]] by blast
-    { fix y assume "y\<in>s" "dist y x < d"
-      hence "dist (f n y) (f n x) < e / 3" using d[THEN bspec[where x=y]] by auto
-      hence "norm (f n y - g x) < 2 * e / 3" using norm_triangle_lt[of "f n y - f n x" "f n x - g x" "2*e/3"]
-        using n(1)[THEN bspec[where x=x], OF `x\<in>s`] unfolding dist_norm unfolding ab_group_add_class.ab_diff_minus by auto
-      hence "dist (g y) (g x) < e" unfolding dist_norm using n(1)[THEN bspec[where x=y], OF `y\<in>s`]
-        unfolding norm_minus_cancel[of "f n y - g y", THEN sym] using norm_triangle_lt[of "f n y - g x" "g y - f n y" e] by (auto simp add: uminus_add_conv_diff)  }
-    hence "\<exists>d>0. \<forall>x'\<in>s. dist x' x < d \<longrightarrow> dist (g x') (g x) < e" using `d>0` by auto  }
-  thus ?thesis unfolding continuous_on_def by auto
-qed
-
-subsection{* Topological properties of linear functions.                               *}
-
-lemma linear_lim_0:
-  assumes "bounded_linear f" shows "(f ---> 0) (at (0))"
-proof-
-  interpret f: bounded_linear f by fact
-  have "(f ---> f 0) (at 0)"
-    using tendsto_ident_at by (rule f.tendsto)
-  thus ?thesis unfolding f.zero .
-qed
-
-lemma linear_continuous_at:
-  assumes "bounded_linear f"  shows "continuous (at a) f"
-  unfolding continuous_at using assms
-  apply (rule bounded_linear.tendsto)
-  apply (rule tendsto_ident_at)
-  done
-
-lemma linear_continuous_within:
-  shows "bounded_linear f ==> continuous (at x within s) f"
-  using continuous_at_imp_continuous_within[of x f s] using linear_continuous_at[of f] by auto
-
-lemma linear_continuous_on:
-  shows "bounded_linear f ==> continuous_on s f"
-  using continuous_at_imp_continuous_on[of s f] using linear_continuous_at[of f] by auto
-
-text{* Also bilinear functions, in composition form.                             *}
-
-lemma bilinear_continuous_at_compose:
-  shows "continuous (at x) f \<Longrightarrow> continuous (at x) g \<Longrightarrow> bounded_bilinear h
-        ==> continuous (at x) (\<lambda>x. h (f x) (g x))"
-  unfolding continuous_at using Lim_bilinear[of f "f x" "(at x)" g "g x" h] by auto
-
-lemma bilinear_continuous_within_compose:
-  shows "continuous (at x within s) f \<Longrightarrow> continuous (at x within s) g \<Longrightarrow> bounded_bilinear h
-        ==> continuous (at x within s) (\<lambda>x. h (f x) (g x))"
-  unfolding continuous_within using Lim_bilinear[of f "f x"] by auto
-
-lemma bilinear_continuous_on_compose:
-  shows "continuous_on s f \<Longrightarrow> continuous_on s g \<Longrightarrow> bounded_bilinear h
-             ==> continuous_on s (\<lambda>x. h (f x) (g x))"
-  unfolding continuous_on_eq_continuous_within apply auto apply(erule_tac x=x in ballE) apply auto apply(erule_tac x=x in ballE) apply auto
-  using bilinear_continuous_within_compose[of _ s f g h] by auto
-
-subsection{* Topological stuff lifted from and dropped to R                            *}
-
-
-lemma open_real:
-  fixes s :: "real set" shows
- "open s \<longleftrightarrow>
-        (\<forall>x \<in> s. \<exists>e>0. \<forall>x'. abs(x' - x) < e --> x' \<in> s)" (is "?lhs = ?rhs")
-  unfolding open_dist dist_norm by simp
-
-lemma islimpt_approachable_real:
-  fixes s :: "real set"
-  shows "x islimpt s \<longleftrightarrow> (\<forall>e>0.  \<exists>x'\<in> s. x' \<noteq> x \<and> abs(x' - x) < e)"
-  unfolding islimpt_approachable dist_norm by simp
-
-lemma closed_real:
-  fixes s :: "real set"
-  shows "closed s \<longleftrightarrow>
-        (\<forall>x. (\<forall>e>0.  \<exists>x' \<in> s. x' \<noteq> x \<and> abs(x' - x) < e)
-            --> x \<in> s)"
-  unfolding closed_limpt islimpt_approachable dist_norm by simp
-
-lemma continuous_at_real_range:
-  fixes f :: "'a::real_normed_vector \<Rightarrow> real"
-  shows "continuous (at x) f \<longleftrightarrow> (\<forall>e>0. \<exists>d>0.
-        \<forall>x'. norm(x' - x) < d --> abs(f x' - f x) < e)"
-  unfolding continuous_at unfolding Lim_at
-  unfolding dist_nz[THEN sym] unfolding dist_norm apply auto
-  apply(erule_tac x=e in allE) apply auto apply (rule_tac x=d in exI) apply auto apply (erule_tac x=x' in allE) apply auto
-  apply(erule_tac x=e in allE) by auto
-
-lemma continuous_on_real_range:
-  fixes f :: "'a::real_normed_vector \<Rightarrow> real"
-  shows "continuous_on s f \<longleftrightarrow> (\<forall>x \<in> s. \<forall>e>0. \<exists>d>0. (\<forall>x' \<in> s. norm(x' - x) < d --> abs(f x' - f x) < e))"
-  unfolding continuous_on_def dist_norm by simp
-
-lemma continuous_at_norm: "continuous (at x) norm"
-  unfolding continuous_at by (intro tendsto_intros)
-
-lemma continuous_on_norm: "continuous_on s norm"
-unfolding continuous_on by (intro ballI tendsto_intros)
-
-lemma continuous_at_component: "continuous (at a) (\<lambda>x. x $ i)"
-unfolding continuous_at by (intro tendsto_intros)
-
-lemma continuous_on_component: "continuous_on s (\<lambda>x. x $ i)"
-unfolding continuous_on by (intro ballI tendsto_intros)
-
-lemma continuous_at_infnorm: "continuous (at x) infnorm"
-  unfolding continuous_at Lim_at o_def unfolding dist_norm
-  apply auto apply (rule_tac x=e in exI) apply auto
-  using order_trans[OF real_abs_sub_infnorm infnorm_le_norm, of _ x] by (metis xt1(7))
-
-text{* Hence some handy theorems on distance, diameter etc. of/from a set.       *}
-
-lemma compact_attains_sup:
-  fixes s :: "real set"
-  assumes "compact s"  "s \<noteq> {}"
-  shows "\<exists>x \<in> s. \<forall>y \<in> s. y \<le> x"
-proof-
-  from assms(1) have a:"bounded s" "closed s" unfolding compact_eq_bounded_closed by auto
-  { fix e::real assume as: "\<forall>x\<in>s. x \<le> rsup s" "rsup s \<notin> s"  "0 < e" "\<forall>x'\<in>s. x' = rsup s \<or> \<not> rsup s - x' < e"
-    have "isLub UNIV s (rsup s)" using rsup[OF assms(2)] unfolding setle_def using as(1) by auto
-    moreover have "isUb UNIV s (rsup s - e)" unfolding isUb_def unfolding setle_def using as(4,2) by auto
-    ultimately have False using isLub_le_isUb[of UNIV s "rsup s" "rsup s - e"] using `e>0` by auto  }
-  thus ?thesis using bounded_has_rsup(1)[OF a(1) assms(2)] using a(2)[unfolded closed_real, THEN spec[where x="rsup s"]]
-    apply(rule_tac x="rsup s" in bexI) by auto
-qed
-
-lemma compact_attains_inf:
-  fixes s :: "real set"
-  assumes "compact s" "s \<noteq> {}"  shows "\<exists>x \<in> s. \<forall>y \<in> s. x \<le> y"
-proof-
-  from assms(1) have a:"bounded s" "closed s" unfolding compact_eq_bounded_closed by auto
-  { fix e::real assume as: "\<forall>x\<in>s. x \<ge> rinf s"  "rinf s \<notin> s"  "0 < e"
-      "\<forall>x'\<in>s. x' = rinf s \<or> \<not> abs (x' - rinf s) < e"
-    have "isGlb UNIV s (rinf s)" using rinf[OF assms(2)] unfolding setge_def using as(1) by auto
-    moreover
-    { fix x assume "x \<in> s"
-      hence *:"abs (x - rinf s) = x - rinf s" using as(1)[THEN bspec[where x=x]] by auto
-      have "rinf s + e \<le> x" using as(4)[THEN bspec[where x=x]] using as(2) `x\<in>s` unfolding * by auto }
-    hence "isLb UNIV s (rinf s + e)" unfolding isLb_def and setge_def by auto
-    ultimately have False using isGlb_le_isLb[of UNIV s "rinf s" "rinf s + e"] using `e>0` by auto  }
-  thus ?thesis using bounded_has_rinf(1)[OF a(1) assms(2)] using a(2)[unfolded closed_real, THEN spec[where x="rinf s"]]
-    apply(rule_tac x="rinf s" in bexI) by auto
-qed
-
-lemma continuous_attains_sup:
-  fixes f :: "'a::metric_space \<Rightarrow> real"
-  shows "compact s \<Longrightarrow> s \<noteq> {} \<Longrightarrow> continuous_on s f
-        ==> (\<exists>x \<in> s. \<forall>y \<in> s.  f y \<le> f x)"
-  using compact_attains_sup[of "f ` s"]
-  using compact_continuous_image[of s f] by auto
-
-lemma continuous_attains_inf:
-  fixes f :: "'a::metric_space \<Rightarrow> real"
-  shows "compact s \<Longrightarrow> s \<noteq> {} \<Longrightarrow> continuous_on s f
-        \<Longrightarrow> (\<exists>x \<in> s. \<forall>y \<in> s. f x \<le> f y)"
-  using compact_attains_inf[of "f ` s"]
-  using compact_continuous_image[of s f] by auto
-
-lemma distance_attains_sup:
-  assumes "compact s" "s \<noteq> {}"
-  shows "\<exists>x \<in> s. \<forall>y \<in> s. dist a y \<le> dist a x"
-proof (rule continuous_attains_sup [OF assms])
-  { fix x assume "x\<in>s"
-    have "(dist a ---> dist a x) (at x within s)"
-      by (intro tendsto_dist tendsto_const Lim_at_within Lim_ident_at)
-  }
-  thus "continuous_on s (dist a)"
-    unfolding continuous_on ..
-qed
-
-text{* For *minimal* distance, we only need closure, not compactness.            *}
-
-lemma distance_attains_inf:
-  fixes a :: "'a::heine_borel"
-  assumes "closed s"  "s \<noteq> {}"
-  shows "\<exists>x \<in> s. \<forall>y \<in> s. dist a x \<le> dist a y"
-proof-
-  from assms(2) obtain b where "b\<in>s" by auto
-  let ?B = "cball a (dist b a) \<inter> s"
-  have "b \<in> ?B" using `b\<in>s` by (simp add: dist_commute)
-  hence "?B \<noteq> {}" by auto
-  moreover
-  { fix x assume "x\<in>?B"
-    fix e::real assume "e>0"
-    { fix x' assume "x'\<in>?B" and as:"dist x' x < e"
-      from as have "\<bar>dist a x' - dist a x\<bar> < e"
-        unfolding abs_less_iff minus_diff_eq
-        using dist_triangle2 [of a x' x]
-        using dist_triangle [of a x x']
-        by arith
-    }
-    hence "\<exists>d>0. \<forall>x'\<in>?B. dist x' x < d \<longrightarrow> \<bar>dist a x' - dist a x\<bar> < e"
-      using `e>0` by auto
-  }
-  hence "continuous_on (cball a (dist b a) \<inter> s) (dist a)"
-    unfolding continuous_on Lim_within dist_norm real_norm_def
-    by fast
-  moreover have "compact ?B"
-    using compact_cball[of a "dist b a"]
-    unfolding compact_eq_bounded_closed
-    using bounded_Int and closed_Int and assms(1) by auto
-  ultimately obtain x where "x\<in>cball a (dist b a) \<inter> s" "\<forall>y\<in>cball a (dist b a) \<inter> s. dist a x \<le> dist a y"
-    using continuous_attains_inf[of ?B "dist a"] by fastsimp
-  thus ?thesis by fastsimp
-qed
-
-subsection{* We can now extend limit compositions to consider the scalar multiplier.   *}
-
-lemma Lim_mul:
-  fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"
-  assumes "(c ---> d) net"  "(f ---> l) net"
-  shows "((\<lambda>x. c(x) *\<^sub>R f x) ---> (d *\<^sub>R l)) net"
-  using assms by (rule scaleR.tendsto)
-
-lemma Lim_vmul:
-  fixes c :: "'a \<Rightarrow> real" and v :: "'b::real_normed_vector"
-  shows "(c ---> d) net ==> ((\<lambda>x. c(x) *\<^sub>R v) ---> d *\<^sub>R v) net"
-  by (intro tendsto_intros)
-
-lemma continuous_vmul:
-  fixes c :: "'a::metric_space \<Rightarrow> real" and v :: "'b::real_normed_vector"
-  shows "continuous net c ==> continuous net (\<lambda>x. c(x) *\<^sub>R v)"
-  unfolding continuous_def using Lim_vmul[of c] by auto
-
-lemma continuous_mul:
-  fixes c :: "'a::metric_space \<Rightarrow> real"
-  fixes f :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector"
-  shows "continuous net c \<Longrightarrow> continuous net f
-             ==> continuous net (\<lambda>x. c(x) *\<^sub>R f x) "
-  unfolding continuous_def by (intro tendsto_intros)
-
-lemma continuous_on_vmul:
-  fixes c :: "'a::metric_space \<Rightarrow> real" and v :: "'b::real_normed_vector"
-  shows "continuous_on s c ==> continuous_on s (\<lambda>x. c(x) *\<^sub>R v)"
-  unfolding continuous_on_eq_continuous_within using continuous_vmul[of _ c] by auto
-
-lemma continuous_on_mul:
-  fixes c :: "'a::metric_space \<Rightarrow> real"
-  fixes f :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector"
-  shows "continuous_on s c \<Longrightarrow> continuous_on s f
-             ==> continuous_on s (\<lambda>x. c(x) *\<^sub>R f x)"
-  unfolding continuous_on_eq_continuous_within using continuous_mul[of _ c] by auto
-
-text{* And so we have continuity of inverse.                                     *}
-
-lemma Lim_inv:
-  fixes f :: "'a \<Rightarrow> real"
-  assumes "(f ---> l) (net::'a net)"  "l \<noteq> 0"
-  shows "((inverse o f) ---> inverse l) net"
-  unfolding o_def using assms by (rule tendsto_inverse)
-
-lemma continuous_inv:
-  fixes f :: "'a::metric_space \<Rightarrow> real"
-  shows "continuous net f \<Longrightarrow> f(netlimit net) \<noteq> 0
-           ==> continuous net (inverse o f)"
-  unfolding continuous_def using Lim_inv by auto
-
-lemma continuous_at_within_inv:
-  fixes f :: "'a::metric_space \<Rightarrow> 'b::real_normed_field"
-  assumes "continuous (at a within s) f" "f a \<noteq> 0"
-  shows "continuous (at a within s) (inverse o f)"
-  using assms unfolding continuous_within o_def
-  by (intro tendsto_intros)
-
-lemma continuous_at_inv:
-  fixes f :: "'a::metric_space \<Rightarrow> 'b::real_normed_field"
-  shows "continuous (at a) f \<Longrightarrow> f a \<noteq> 0
-         ==> continuous (at a) (inverse o f) "
-  using within_UNIV[THEN sym, of "at a"] using continuous_at_within_inv[of a UNIV] by auto
-
-subsection{* Preservation properties for pasted sets.                                  *}
-
-lemma bounded_pastecart:
-  fixes s :: "('a::real_normed_vector ^ _) set" (* FIXME: generalize to metric_space *)
-  assumes "bounded s" "bounded t"
-  shows "bounded { pastecart x y | x y . (x \<in> s \<and> y \<in> t)}"
-proof-
-  obtain a b where ab:"\<forall>x\<in>s. norm x \<le> a" "\<forall>x\<in>t. norm x \<le> b" using assms[unfolded bounded_iff] by auto
-  { fix x y assume "x\<in>s" "y\<in>t"
-    hence "norm x \<le> a" "norm y \<le> b" using ab by auto
-    hence "norm (pastecart x y) \<le> a + b" using norm_pastecart[of x y] by auto }
-  thus ?thesis unfolding bounded_iff by auto
-qed
-
-lemma bounded_Times:
-  assumes "bounded s" "bounded t" shows "bounded (s \<times> t)"
-proof-
-  obtain x y a b where "\<forall>z\<in>s. dist x z \<le> a" "\<forall>z\<in>t. dist y z \<le> b"
-    using assms [unfolded bounded_def] by auto
-  then have "\<forall>z\<in>s \<times> t. dist (x, y) z \<le> sqrt (a\<twosuperior> + b\<twosuperior>)"
-    by (auto simp add: dist_Pair_Pair real_sqrt_le_mono add_mono power_mono)
-  thus ?thesis unfolding bounded_any_center [where a="(x, y)"] by auto
-qed
-
-lemma closed_pastecart:
-  fixes s :: "(real ^ 'a::finite) set" (* FIXME: generalize *)
-  assumes "closed s"  "closed t"
-  shows "closed {pastecart x y | x y . x \<in> s \<and> y \<in> t}"
-proof-
-  { fix x l assume as:"\<forall>n::nat. x n \<in> {pastecart x y |x y. x \<in> s \<and> y \<in> t}"  "(x ---> l) sequentially"
-    { fix n::nat have "fstcart (x n) \<in> s" "sndcart (x n) \<in> t" using as(1)[THEN spec[where x=n]] by auto } note * = this
-    moreover
-    { fix e::real assume "e>0"
-      then obtain N::nat where N:"\<forall>n\<ge>N. dist (x n) l < e" using as(2)[unfolded Lim_sequentially, THEN spec[where x=e]] by auto
-      { fix n::nat assume "n\<ge>N"
-        hence "dist (fstcart (x n)) (fstcart l) < e" "dist (sndcart (x n)) (sndcart l) < e"
-          using N[THEN spec[where x=n]] dist_fstcart[of "x n" l] dist_sndcart[of "x n" l] by auto   }
-      hence "\<exists>N. \<forall>n\<ge>N. dist (fstcart (x n)) (fstcart l) < e" "\<exists>N. \<forall>n\<ge>N. dist (sndcart (x n)) (sndcart l) < e" by auto  }
-    ultimately have "fstcart l \<in> s" "sndcart l \<in> t"
-      using assms(1)[unfolded closed_sequential_limits, THEN spec[where x="\<lambda>n. fstcart (x n)"], THEN spec[where x="fstcart l"]]
-      using assms(2)[unfolded closed_sequential_limits, THEN spec[where x="\<lambda>n. sndcart (x n)"], THEN spec[where x="sndcart l"]]
-      unfolding Lim_sequentially by auto
-    hence "l \<in> {pastecart x y |x y. x \<in> s \<and> y \<in> t}" using pastecart_fst_snd[THEN sym, of l] by auto  }
-  thus ?thesis unfolding closed_sequential_limits by auto
-qed
-
-lemma compact_pastecart:
-  fixes s t :: "(real ^ _) set"
-  shows "compact s \<Longrightarrow> compact t ==> compact {pastecart x y | x y . x \<in> s \<and> y \<in> t}"
-  unfolding compact_eq_bounded_closed using bounded_pastecart[of s t] closed_pastecart[of s t] by auto
-
-lemma mem_Times_iff: "x \<in> A \<times> B \<longleftrightarrow> fst x \<in> A \<and> snd x \<in> B"
-by (induct x) simp
-
-lemma compact_Times: "compact s \<Longrightarrow> compact t \<Longrightarrow> compact (s \<times> t)"
-unfolding compact_def
-apply clarify
-apply (drule_tac x="fst \<circ> f" in spec)
-apply (drule mp, simp add: mem_Times_iff)
-apply (clarify, rename_tac l1 r1)
-apply (drule_tac x="snd \<circ> f \<circ> r1" in spec)
-apply (drule mp, simp add: mem_Times_iff)
-apply (clarify, rename_tac l2 r2)
-apply (rule_tac x="(l1, l2)" in rev_bexI, simp)
-apply (rule_tac x="r1 \<circ> r2" in exI)
-apply (rule conjI, simp add: subseq_def)
-apply (drule_tac r=r2 in lim_subseq [COMP swap_prems_rl], assumption)
-apply (drule (1) tendsto_Pair) back
-apply (simp add: o_def)
-done
-
-text{* Hence some useful properties follow quite easily.                         *}
-
-lemma compact_scaling:
-  fixes s :: "'a::real_normed_vector set"
-  assumes "compact s"  shows "compact ((\<lambda>x. c *\<^sub>R x) ` s)"
-proof-
-  let ?f = "\<lambda>x. scaleR c x"
-  have *:"bounded_linear ?f" by (rule scaleR.bounded_linear_right)
-  show ?thesis using compact_continuous_image[of s ?f] continuous_at_imp_continuous_on[of s ?f]
-    using linear_continuous_at[OF *] assms by auto
-qed
-
-lemma compact_negations:
-  fixes s :: "'a::real_normed_vector set"
-  assumes "compact s"  shows "compact ((\<lambda>x. -x) ` s)"
-  using compact_scaling [OF assms, of "- 1"] by auto
-
-lemma compact_sums:
-  fixes s t :: "'a::real_normed_vector set"
-  assumes "compact s"  "compact t"  shows "compact {x + y | x y. x \<in> s \<and> y \<in> t}"
-proof-
-  have *:"{x + y | x y. x \<in> s \<and> y \<in> t} = (\<lambda>z. fst z + snd z) ` (s \<times> t)"
-    apply auto unfolding image_iff apply(rule_tac x="(xa, y)" in bexI) by auto
-  have "continuous_on (s \<times> t) (\<lambda>z. fst z + snd z)"
-    unfolding continuous_on by (rule ballI) (intro tendsto_intros)
-  thus ?thesis unfolding * using compact_continuous_image compact_Times [OF assms] by auto
-qed
-
-lemma compact_differences:
-  fixes s t :: "'a::real_normed_vector set"
-  assumes "compact s" "compact t"  shows "compact {x - y | x y. x \<in> s \<and> y \<in> t}"
-proof-
-  have "{x - y | x y. x\<in>s \<and> y \<in> t} =  {x + y | x y. x \<in> s \<and> y \<in> (uminus ` t)}"
-    apply auto apply(rule_tac x= xa in exI) apply auto apply(rule_tac x=xa in exI) by auto
-  thus ?thesis using compact_sums[OF assms(1) compact_negations[OF assms(2)]] by auto
-qed
-
-lemma compact_translation:
-  fixes s :: "'a::real_normed_vector set"
-  assumes "compact s"  shows "compact ((\<lambda>x. a + x) ` s)"
-proof-
-  have "{x + y |x y. x \<in> s \<and> y \<in> {a}} = (\<lambda>x. a + x) ` s" by auto
-  thus ?thesis using compact_sums[OF assms compact_sing[of a]] by auto
-qed
-
-lemma compact_affinity:
-  fixes s :: "'a::real_normed_vector set"
-  assumes "compact s"  shows "compact ((\<lambda>x. a + c *\<^sub>R x) ` s)"
-proof-
-  have "op + a ` op *\<^sub>R c ` s = (\<lambda>x. a + c *\<^sub>R x) ` s" by auto
-  thus ?thesis using compact_translation[OF compact_scaling[OF assms], of a c] by auto
-qed
-
-text{* Hence we get the following.                                               *}
-
-lemma compact_sup_maxdistance:
-  fixes s :: "'a::real_normed_vector set"
-  assumes "compact s"  "s \<noteq> {}"
-  shows "\<exists>x\<in>s. \<exists>y\<in>s. \<forall>u\<in>s. \<forall>v\<in>s. norm(u - v) \<le> norm(x - y)"
-proof-
-  have "{x - y | x y . x\<in>s \<and> y\<in>s} \<noteq> {}" using `s \<noteq> {}` by auto
-  then obtain x where x:"x\<in>{x - y |x y. x \<in> s \<and> y \<in> s}"  "\<forall>y\<in>{x - y |x y. x \<in> s \<and> y \<in> s}. norm y \<le> norm x"
-    using compact_differences[OF assms(1) assms(1)]
-    using distance_attains_sup[where 'a="'a", unfolded dist_norm, of "{x - y | x y . x\<in>s \<and> y\<in>s}" 0] by(auto simp add: norm_minus_cancel)
-  from x(1) obtain a b where "a\<in>s" "b\<in>s" "x = a - b" by auto
-  thus ?thesis using x(2)[unfolded `x = a - b`] by blast
-qed
-
-text{* We can state this in terms of diameter of a set.                          *}
-
-definition "diameter s = (if s = {} then 0::real else rsup {norm(x - y) | x y. x \<in> s \<and> y \<in> s})"
-  (* TODO: generalize to class metric_space *)
-
-lemma diameter_bounded:
-  assumes "bounded s"
-  shows "\<forall>x\<in>s. \<forall>y\<in>s. norm(x - y) \<le> diameter s"
-        "\<forall>d>0. d < diameter s --> (\<exists>x\<in>s. \<exists>y\<in>s. norm(x - y) > d)"
-proof-
-  let ?D = "{norm (x - y) |x y. x \<in> s \<and> y \<in> s}"
-  obtain a where a:"\<forall>x\<in>s. norm x \<le> a" using assms[unfolded bounded_iff] by auto
-  { fix x y assume "x \<in> s" "y \<in> s"
-    hence "norm (x - y) \<le> 2 * a" using norm_triangle_ineq[of x "-y", unfolded norm_minus_cancel] a[THEN bspec[where x=x]] a[THEN bspec[where x=y]] by (auto simp add: ring_simps)  }
-  note * = this
-  { fix x y assume "x\<in>s" "y\<in>s"  hence "s \<noteq> {}" by auto
-    have lub:"isLub UNIV ?D (rsup ?D)" using * rsup[of ?D] using `s\<noteq>{}` unfolding setle_def by auto
-    have "norm(x - y) \<le> diameter s" unfolding diameter_def using `s\<noteq>{}` *[OF `x\<in>s` `y\<in>s`] `x\<in>s` `y\<in>s` isLubD1[OF lub] unfolding setle_def by auto  }
-  moreover
-  { fix d::real assume "d>0" "d < diameter s"
-    hence "s\<noteq>{}" unfolding diameter_def by auto
-    hence lub:"isLub UNIV ?D (rsup ?D)" using * rsup[of ?D] unfolding setle_def by auto
-    have "\<exists>d' \<in> ?D. d' > d"
-    proof(rule ccontr)
-      assume "\<not> (\<exists>d'\<in>{norm (x - y) |x y. x \<in> s \<and> y \<in> s}. d < d')"
-      hence as:"\<forall>d'\<in>?D. d' \<le> d" apply auto apply(erule_tac x="norm (x - y)" in allE) by auto
-      hence "isUb UNIV ?D d" unfolding isUb_def unfolding setle_def by auto
-      thus False using `d < diameter s` `s\<noteq>{}` isLub_le_isUb[OF lub, of d] unfolding diameter_def  by auto
-    qed
-    hence "\<exists>x\<in>s. \<exists>y\<in>s. norm(x - y) > d" by auto  }
-  ultimately show "\<forall>x\<in>s. \<forall>y\<in>s. norm(x - y) \<le> diameter s"
-        "\<forall>d>0. d < diameter s --> (\<exists>x\<in>s. \<exists>y\<in>s. norm(x - y) > d)" by auto
-qed
-
-lemma diameter_bounded_bound:
- "bounded s \<Longrightarrow> x \<in> s \<Longrightarrow> y \<in> s ==> norm(x - y) \<le> diameter s"
-  using diameter_bounded by blast
-
-lemma diameter_compact_attained:
-  fixes s :: "'a::real_normed_vector set"
-  assumes "compact s"  "s \<noteq> {}"
-  shows "\<exists>x\<in>s. \<exists>y\<in>s. (norm(x - y) = diameter s)"
-proof-
-  have b:"bounded s" using assms(1) by (rule compact_imp_bounded)
-  then obtain x y where xys:"x\<in>s" "y\<in>s" and xy:"\<forall>u\<in>s. \<forall>v\<in>s. norm (u - v) \<le> norm (x - y)" using compact_sup_maxdistance[OF assms] by auto
-  hence "diameter s \<le> norm (x - y)" using rsup_le[of "{norm (x - y) |x y. x \<in> s \<and> y \<in> s}" "norm (x - y)"]
-    unfolding setle_def and diameter_def by auto
-  thus ?thesis using diameter_bounded(1)[OF b, THEN bspec[where x=x], THEN bspec[where x=y], OF xys] and xys by auto
-qed
-
-text{* Related results with closure as the conclusion.                           *}
-
-lemma closed_scaling:
-  fixes s :: "'a::real_normed_vector set"
-  assumes "closed s" shows "closed ((\<lambda>x. c *\<^sub>R x) ` s)"
-proof(cases "s={}")
-  case True thus ?thesis by auto
-next
-  case False
-  show ?thesis
-  proof(cases "c=0")
-    have *:"(\<lambda>x. 0) ` s = {0}" using `s\<noteq>{}` by auto
-    case True thus ?thesis apply auto unfolding * using closed_sing by auto
-  next
-    case False
-    { fix x l assume as:"\<forall>n::nat. x n \<in> scaleR c ` s"  "(x ---> l) sequentially"
-      { fix n::nat have "scaleR (1 / c) (x n) \<in> s"
-          using as(1)[THEN spec[where x=n]]
-          using `c\<noteq>0` by (auto simp add: vector_smult_assoc)
-      }
-      moreover
-      { fix e::real assume "e>0"
-        hence "0 < e *\<bar>c\<bar>"  using `c\<noteq>0` mult_pos_pos[of e "abs c"] by auto
-        then obtain N where "\<forall>n\<ge>N. dist (x n) l < e * \<bar>c\<bar>"
-          using as(2)[unfolded Lim_sequentially, THEN spec[where x="e * abs c"]] by auto
-        hence "\<exists>N. \<forall>n\<ge>N. dist (scaleR (1 / c) (x n)) (scaleR (1 / c) l) < e"
-          unfolding dist_norm unfolding scaleR_right_diff_distrib[THEN sym]
-          using mult_imp_div_pos_less[of "abs c" _ e] `c\<noteq>0` by auto  }
-      hence "((\<lambda>n. scaleR (1 / c) (x n)) ---> scaleR (1 / c) l) sequentially" unfolding Lim_sequentially by auto
-      ultimately have "l \<in> scaleR c ` s"
-        using assms[unfolded closed_sequential_limits, THEN spec[where x="\<lambda>n. scaleR (1/c) (x n)"], THEN spec[where x="scaleR (1/c) l"]]
-        unfolding image_iff using `c\<noteq>0` apply(rule_tac x="scaleR (1 / c) l" in bexI) by auto  }
-    thus ?thesis unfolding closed_sequential_limits by fast
-  qed
-qed
-
-lemma closed_negations:
-  fixes s :: "'a::real_normed_vector set"
-  assumes "closed s"  shows "closed ((\<lambda>x. -x) ` s)"
-  using closed_scaling[OF assms, of "- 1"] by simp
-
-lemma compact_closed_sums:
-  fixes s :: "'a::real_normed_vector set"
-  assumes "compact s"  "closed t"  shows "closed {x + y | x y. x \<in> s \<and> y \<in> t}"
-proof-
-  let ?S = "{x + y |x y. x \<in> s \<and> y \<in> t}"
-  { fix x l assume as:"\<forall>n. x n \<in> ?S"  "(x ---> l) sequentially"
-    from as(1) obtain f where f:"\<forall>n. x n = fst (f n) + snd (f n)"  "\<forall>n. fst (f n) \<in> s"  "\<forall>n. snd (f n) \<in> t"
-      using choice[of "\<lambda>n y. x n = (fst y) + (snd y) \<and> fst y \<in> s \<and> snd y \<in> t"] by auto
-    obtain l' r where "l'\<in>s" and r:"subseq r" and lr:"(((\<lambda>n. fst (f n)) \<circ> r) ---> l') sequentially"
-      using assms(1)[unfolded compact_def, THEN spec[where x="\<lambda> n. fst (f n)"]] using f(2) by auto
-    have "((\<lambda>n. snd (f (r n))) ---> l - l') sequentially"
-      using Lim_sub[OF lim_subseq[OF r as(2)] lr] and f(1) unfolding o_def by auto
-    hence "l - l' \<in> t"
-      using assms(2)[unfolded closed_sequential_limits, THEN spec[where x="\<lambda> n. snd (f (r n))"], THEN spec[where x="l - l'"]]
-      using f(3) by auto
-    hence "l \<in> ?S" using `l' \<in> s` apply auto apply(rule_tac x=l' in exI) apply(rule_tac x="l - l'" in exI) by auto
-  }
-  thus ?thesis unfolding closed_sequential_limits by fast
-qed
-
-lemma closed_compact_sums:
-  fixes s t :: "'a::real_normed_vector set"
-  assumes "closed s"  "compact t"
-  shows "closed {x + y | x y. x \<in> s \<and> y \<in> t}"
-proof-
-  have "{x + y |x y. x \<in> t \<and> y \<in> s} = {x + y |x y. x \<in> s \<and> y \<in> t}" apply auto
-    apply(rule_tac x=y in exI) apply auto apply(rule_tac x=y in exI) by auto
-  thus ?thesis using compact_closed_sums[OF assms(2,1)] by simp
-qed
-
-lemma compact_closed_differences:
-  fixes s t :: "'a::real_normed_vector set"
-  assumes "compact s"  "closed t"
-  shows "closed {x - y | x y. x \<in> s \<and> y \<in> t}"
-proof-
-  have "{x + y |x y. x \<in> s \<and> y \<in> uminus ` t} =  {x - y |x y. x \<in> s \<and> y \<in> t}"
-    apply auto apply(rule_tac x=xa in exI) apply auto apply(rule_tac x=xa in exI) by auto
-  thus ?thesis using compact_closed_sums[OF assms(1) closed_negations[OF assms(2)]] by auto
-qed
-
-lemma closed_compact_differences:
-  fixes s t :: "'a::real_normed_vector set"
-  assumes "closed s" "compact t"
-  shows "closed {x - y | x y. x \<in> s \<and> y \<in> t}"
-proof-
-  have "{x + y |x y. x \<in> s \<and> y \<in> uminus ` t} = {x - y |x y. x \<in> s \<and> y \<in> t}"
-    apply auto apply(rule_tac x=xa in exI) apply auto apply(rule_tac x=xa in exI) by auto
- thus ?thesis using closed_compact_sums[OF assms(1) compact_negations[OF assms(2)]] by simp
-qed
-
-lemma closed_translation:
-  fixes a :: "'a::real_normed_vector"
-  assumes "closed s"  shows "closed ((\<lambda>x. a + x) ` s)"
-proof-
-  have "{a + y |y. y \<in> s} = (op + a ` s)" by auto
-  thus ?thesis using compact_closed_sums[OF compact_sing[of a] assms] by auto
-qed
-
-lemma translation_UNIV:
-  fixes a :: "'a::ab_group_add" shows "range (\<lambda>x. a + x) = UNIV"
-  apply (auto simp add: image_iff) apply(rule_tac x="x - a" in exI) by auto
-
-lemma translation_diff:
-  fixes a :: "'a::ab_group_add"
-  shows "(\<lambda>x. a + x) ` (s - t) = ((\<lambda>x. a + x) ` s) - ((\<lambda>x. a + x) ` t)"
-  by auto
-
-lemma closure_translation:
-  fixes a :: "'a::real_normed_vector"
-  shows "closure ((\<lambda>x. a + x) ` s) = (\<lambda>x. a + x) ` (closure s)"
-proof-
-  have *:"op + a ` (UNIV - s) = UNIV - op + a ` s"
-    apply auto unfolding image_iff apply(rule_tac x="x - a" in bexI) by auto
-  show ?thesis unfolding closure_interior translation_diff translation_UNIV
-    using interior_translation[of a "UNIV - s"] unfolding * by auto
-qed
-
-lemma frontier_translation:
-  fixes a :: "'a::real_normed_vector"
-  shows "frontier((\<lambda>x. a + x) ` s) = (\<lambda>x. a + x) ` (frontier s)"
-  unfolding frontier_def translation_diff interior_translation closure_translation by auto
-
-subsection{* Separation between points and sets.                                       *}
-
-lemma separate_point_closed:
-  fixes s :: "'a::heine_borel set"
-  shows "closed s \<Longrightarrow> a \<notin> s  ==> (\<exists>d>0. \<forall>x\<in>s. d \<le> dist a x)"
-proof(cases "s = {}")
-  case True
-  thus ?thesis by(auto intro!: exI[where x=1])
-next
-  case False
-  assume "closed s" "a \<notin> s"
-  then obtain x where "x\<in>s" "\<forall>y\<in>s. dist a x \<le> dist a y" using `s \<noteq> {}` distance_attains_inf [of s a] by blast
-  with `x\<in>s` show ?thesis using dist_pos_lt[of a x] and`a \<notin> s` by blast
-qed
-
-lemma separate_compact_closed:
-  fixes s t :: "'a::{heine_borel, real_normed_vector} set"
-    (* TODO: does this generalize to heine_borel? *)
-  assumes "compact s" and "closed t" and "s \<inter> t = {}"
-  shows "\<exists>d>0. \<forall>x\<in>s. \<forall>y\<in>t. d \<le> dist x y"
-proof-
-  have "0 \<notin> {x - y |x y. x \<in> s \<and> y \<in> t}" using assms(3) by auto
-  then obtain d where "d>0" and d:"\<forall>x\<in>{x - y |x y. x \<in> s \<and> y \<in> t}. d \<le> dist 0 x"
-    using separate_point_closed[OF compact_closed_differences[OF assms(1,2)], of 0] by auto
-  { fix x y assume "x\<in>s" "y\<in>t"
-    hence "x - y \<in> {x - y |x y. x \<in> s \<and> y \<in> t}" by auto
-    hence "d \<le> dist (x - y) 0" using d[THEN bspec[where x="x - y"]] using dist_commute
-      by (auto  simp add: dist_commute)
-    hence "d \<le> dist x y" unfolding dist_norm by auto  }
-  thus ?thesis using `d>0` by auto
-qed
-
-lemma separate_closed_compact:
-  fixes s t :: "'a::{heine_borel, real_normed_vector} set"
-  assumes "closed s" and "compact t" and "s \<inter> t = {}"
-  shows "\<exists>d>0. \<forall>x\<in>s. \<forall>y\<in>t. d \<le> dist x y"
-proof-
-  have *:"t \<inter> s = {}" using assms(3) by auto
-  show ?thesis using separate_compact_closed[OF assms(2,1) *]
-    apply auto apply(rule_tac x=d in exI) apply auto apply (erule_tac x=y in ballE)
-    by (auto simp add: dist_commute)
-qed
-
-(* A cute way of denoting open and closed intervals using overloading.       *)
-
-lemma interval: fixes a :: "'a::ord^'n::finite" shows
-  "{a <..< b} = {x::'a^'n. \<forall>i. a$i < x$i \<and> x$i < b$i}" and
-  "{a .. b} = {x::'a^'n. \<forall>i. a$i \<le> x$i \<and> x$i \<le> b$i}"
-  by (auto simp add: expand_set_eq vector_less_def vector_less_eq_def)
-
-lemma mem_interval: fixes a :: "'a::ord^'n::finite" shows
-  "x \<in> {a<..<b} \<longleftrightarrow> (\<forall>i. a$i < x$i \<and> x$i < b$i)"
-  "x \<in> {a .. b} \<longleftrightarrow> (\<forall>i. a$i \<le> x$i \<and> x$i \<le> b$i)"
-  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 \<in> {a .. b} \<longleftrightarrow> dest_vec1 a \<le> dest_vec1 x \<and> dest_vec1 x \<le> dest_vec1 b)"
- "(x \<in> {a<..<b} \<longleftrightarrow> dest_vec1 a < dest_vec1 x \<and> 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} = {} \<longleftrightarrow> (\<exists>i. b$i \<le> a$i))" (is ?th1) and
- "({a  ..  b} = {} \<longleftrightarrow> (\<exists>i. b$i < a$i))" (is ?th2)
-proof-
-  { fix i x assume as:"b$i \<le> a$i" and x:"x\<in>{a <..< b}"
-    hence "a $ i < x $ i \<and> x $ i < b $ i" unfolding mem_interval by auto
-    hence "a$i < b$i" by auto
-    hence False using as by auto  }
-  moreover
-  { assume as:"\<forall>i. \<not> (b$i \<le> a$i)"
-    let ?x = "(1/2) *\<^sub>R (a + b)"
-    { fix i
-      have "a$i < b$i" using as[THEN spec[where x=i]] by auto
-      hence "a$i < ((1/2) *\<^sub>R (a+b)) $ i" "((1/2) *\<^sub>R (a+b)) $ i < b$i"
-        unfolding vector_smult_component and vector_add_component
-        by (auto simp add: less_divide_eq_number_of1)  }
-    hence "{a <..< b} \<noteq> {}" using mem_interval(1)[of "?x" a b] by auto  }
-  ultimately show ?th1 by blast
-
-  { fix i x assume as:"b$i < a$i" and x:"x\<in>{a .. b}"
-    hence "a $ i \<le> x $ i \<and> x $ i \<le> b $ i" unfolding mem_interval by auto
-    hence "a$i \<le> b$i" by auto
-    hence False using as by auto  }
-  moreover
-  { assume as:"\<forall>i. \<not> (b$i < a$i)"
-    let ?x = "(1/2) *\<^sub>R (a + b)"
-    { fix i
-      have "a$i \<le> b$i" using as[THEN spec[where x=i]] by auto
-      hence "a$i \<le> ((1/2) *\<^sub>R (a+b)) $ i" "((1/2) *\<^sub>R (a+b)) $ i \<le> b$i"
-        unfolding vector_smult_component and vector_add_component
-        by (auto simp add: less_divide_eq_number_of1)  }
-    hence "{a .. b} \<noteq> {}" using mem_interval(2)[of "?x" a b] by auto  }
-  ultimately show ?th2 by blast
-qed
-
-lemma interval_ne_empty: fixes a :: "real^'n::finite" shows
-  "{a  ..  b} \<noteq> {} \<longleftrightarrow> (\<forall>i. a$i \<le> b$i)" and
-  "{a <..< b} \<noteq> {} \<longleftrightarrow> (\<forall>i. a$i < b$i)"
-  unfolding interval_eq_empty[of a b] by (auto simp add: not_less not_le) (* BH: Why doesn't just "auto" work here? *)
-
-lemma subset_interval_imp: fixes a :: "real^'n::finite" shows
- "(\<forall>i. a$i \<le> c$i \<and> d$i \<le> b$i) \<Longrightarrow> {c .. d} \<subseteq> {a .. b}" and
- "(\<forall>i. a$i < c$i \<and> d$i < b$i) \<Longrightarrow> {c .. d} \<subseteq> {a<..<b}" and
- "(\<forall>i. a$i \<le> c$i \<and> d$i \<le> b$i) \<Longrightarrow> {c<..<d} \<subseteq> {a .. b}" and
- "(\<forall>i. a$i \<le> c$i \<and> d$i \<le> b$i) \<Longrightarrow> {c<..<d} \<subseteq> {a<..<b}"
-  unfolding subset_eq[unfolded Ball_def] unfolding mem_interval
-  by (auto intro: order_trans less_le_trans le_less_trans less_imp_le) (* BH: Why doesn't just "auto" work here? *)
-
-lemma interval_sing: fixes a :: "'a::linorder^'n::finite" shows
- "{a .. a} = {a} \<and> {a<..<a} = {}"
-apply(auto simp add: expand_set_eq vector_less_def vector_less_eq_def Cart_eq)
-apply (simp add: order_eq_iff)
-apply (auto simp add: not_less less_imp_le)
-done
-
-lemma interval_open_subset_closed:  fixes a :: "'a::preorder^'n::finite" shows
- "{a<..<b} \<subseteq> {a .. b}"
-proof(simp add: subset_eq, rule)
-  fix x
-  assume x:"x \<in>{a<..<b}"
-  { fix i
-    have "a $ i \<le> x $ i"
-      using x order_less_imp_le[of "a$i" "x$i"]
-      by(simp add: expand_set_eq vector_less_def vector_less_eq_def Cart_eq)
-  }
-  moreover
-  { fix i
-    have "x $ i \<le> b $ i"
-      using x order_less_imp_le[of "x$i" "b$i"]
-      by(simp add: expand_set_eq vector_less_def vector_less_eq_def Cart_eq)
-  }
-  ultimately
-  show "a \<le> x \<and> x \<le> b"
-    by(simp add: expand_set_eq vector_less_def vector_less_eq_def Cart_eq)
-qed
-
-lemma subset_interval: fixes a :: "real^'n::finite" shows
- "{c .. d} \<subseteq> {a .. b} \<longleftrightarrow> (\<forall>i. c$i \<le> d$i) --> (\<forall>i. a$i \<le> c$i \<and> d$i \<le> b$i)" (is ?th1) and
- "{c .. d} \<subseteq> {a<..<b} \<longleftrightarrow> (\<forall>i. c$i \<le> d$i) --> (\<forall>i. a$i < c$i \<and> d$i < b$i)" (is ?th2) and
- "{c<..<d} \<subseteq> {a .. b} \<longleftrightarrow> (\<forall>i. c$i < d$i) --> (\<forall>i. a$i \<le> c$i \<and> d$i \<le> b$i)" (is ?th3) and
- "{c<..<d} \<subseteq> {a<..<b} \<longleftrightarrow> (\<forall>i. c$i < d$i) --> (\<forall>i. a$i \<le> c$i \<and> d$i \<le> b$i)" (is ?th4)
-proof-
-  show ?th1 unfolding subset_eq and Ball_def and mem_interval by (auto intro: order_trans)
-  show ?th2 unfolding subset_eq and Ball_def and mem_interval by (auto intro: le_less_trans less_le_trans order_trans less_imp_le)
-  { assume as: "{c<..<d} \<subseteq> {a .. b}" "\<forall>i. c$i < d$i"
-    hence "{c<..<d} \<noteq> {}" unfolding interval_eq_empty by (auto, drule_tac x=i in spec, simp) (* BH: Why doesn't just "auto" work? *)
-    fix i
-    (** TODO combine the following two parts as done in the HOL_light version. **)
-    { let ?x = "(\<chi> j. (if j=i then ((min (a$j) (d$j))+c$j)/2 else (c$j+d$j)/2))::real^'n"
-      assume as2: "a$i > c$i"
-      { fix j
-        have "c $ j < ?x $ j \<and> ?x $ j < d $ j" unfolding Cart_lambda_beta
-          apply(cases "j=i") using as(2)[THEN spec[where x=j]]
-          by (auto simp add: less_divide_eq_number_of1 as2)  }
-      hence "?x\<in>{c<..<d}" unfolding mem_interval by auto
-      moreover
-      have "?x\<notin>{a .. b}"
-        unfolding mem_interval apply auto apply(rule_tac x=i in exI)
-        using as(2)[THEN spec[where x=i]] and as2
-        by (auto simp add: less_divide_eq_number_of1)
-      ultimately have False using as by auto  }
-    hence "a$i \<le> c$i" by(rule ccontr)auto
-    moreover
-    { let ?x = "(\<chi> j. (if j=i then ((max (b$j) (c$j))+d$j)/2 else (c$j+d$j)/2))::real^'n"
-      assume as2: "b$i < d$i"
-      { fix j
-        have "d $ j > ?x $ j \<and> ?x $ j > c $ j" unfolding Cart_lambda_beta
-          apply(cases "j=i") using as(2)[THEN spec[where x=j]]
-          by (auto simp add: less_divide_eq_number_of1 as2)  }
-      hence "?x\<in>{c<..<d}" unfolding mem_interval by auto
-      moreover
-      have "?x\<notin>{a .. b}"
-        unfolding mem_interval apply auto apply(rule_tac x=i in exI)
-        using as(2)[THEN spec[where x=i]] and as2
-        by (auto simp add: less_divide_eq_number_of1)
-      ultimately have False using as by auto  }
-    hence "b$i \<ge> d$i" by(rule ccontr)auto
-    ultimately
-    have "a$i \<le> c$i \<and> d$i \<le> b$i" by auto
-  } note part1 = this
-  thus ?th3 unfolding subset_eq and Ball_def and mem_interval apply auto apply (erule_tac x=ia in allE, simp)+ by (erule_tac x=i in allE, erule_tac x=i in allE, simp)+
-  { assume as:"{c<..<d} \<subseteq> {a<..<b}" "\<forall>i. c$i < d$i"
-    fix i
-    from as(1) have "{c<..<d} \<subseteq> {a..b}" using interval_open_subset_closed[of a b] by auto
-    hence "a$i \<le> c$i \<and> d$i \<le> b$i" using part1 and as(2) by auto  } note * = this
-  thus ?th4 unfolding subset_eq and Ball_def and mem_interval apply auto apply (erule_tac x=ia in allE, simp)+ by (erule_tac x=i in allE, erule_tac x=i in allE, simp)+
-qed
-
-lemma disjoint_interval: fixes a::"real^'n::finite" shows
-  "{a .. b} \<inter> {c .. d} = {} \<longleftrightarrow> (\<exists>i. (b$i < a$i \<or> d$i < c$i \<or> b$i < c$i \<or> d$i < a$i))" (is ?th1) and
-  "{a .. b} \<inter> {c<..<d} = {} \<longleftrightarrow> (\<exists>i. (b$i < a$i \<or> d$i \<le> c$i \<or> b$i \<le> c$i \<or> d$i \<le> a$i))" (is ?th2) and
-  "{a<..<b} \<inter> {c .. d} = {} \<longleftrightarrow> (\<exists>i. (b$i \<le> a$i \<or> d$i < c$i \<or> b$i \<le> c$i \<or> d$i \<le> a$i))" (is ?th3) and
-  "{a<..<b} \<inter> {c<..<d} = {} \<longleftrightarrow> (\<exists>i. (b$i \<le> a$i \<or> d$i \<le> c$i \<or> b$i \<le> c$i \<or> d$i \<le> a$i))" (is ?th4)
-proof-
-  let ?z = "(\<chi> i. ((max (a$i) (c$i)) + (min (b$i) (d$i))) / 2)::real^'n"
-  show ?th1 ?th2 ?th3 ?th4
-  unfolding expand_set_eq and Int_iff and empty_iff and mem_interval and all_conj_distrib[THEN sym] and eq_False
-  apply (auto elim!: allE[where x="?z"])
-  apply ((rule_tac x=x in exI, force) | (rule_tac x=i in exI, force))+
-  done
-qed
-
-lemma inter_interval: fixes a :: "'a::linorder^'n::finite" shows
- "{a .. b} \<inter> {c .. d} =  {(\<chi> i. max (a$i) (c$i)) .. (\<chi> i. min (b$i) (d$i))}"
-  unfolding expand_set_eq and Int_iff and mem_interval
-  by (auto simp add: less_divide_eq_number_of1 intro!: bexI)
-
-(* Moved interval_open_subset_closed a bit upwards *)
-
-lemma open_interval_lemma: fixes x :: "real" shows
- "a < x \<Longrightarrow> x < b ==> (\<exists>d>0. \<forall>x'. abs(x' - x) < d --> a < x' \<and> x' < b)"
-  by(rule_tac x="min (x - a) (b - x)" in exI, auto)
-
-lemma open_interval: fixes a :: "real^'n::finite" shows "open {a<..<b}"
-proof-
-  { fix x assume x:"x\<in>{a<..<b}"
-    { fix i
-      have "\<exists>d>0. \<forall>x'. abs (x' - (x$i)) < d \<longrightarrow> a$i < x' \<and> x' < b$i"
-        using x[unfolded mem_interval, THEN spec[where x=i]]
-        using open_interval_lemma[of "a$i" "x$i" "b$i"] by auto  }
-
-    hence "\<forall>i. \<exists>d>0. \<forall>x'. abs (x' - (x$i)) < d \<longrightarrow> a$i < x' \<and> x' < b$i" by auto
-    then obtain d where d:"\<forall>i. 0 < d i \<and> (\<forall>x'. \<bar>x' - x $ i\<bar> < d i \<longrightarrow> a $ i < x' \<and> x' < b $ i)"
-      using bchoice[of "UNIV" "\<lambda>i d. d>0 \<and> (\<forall>x'. \<bar>x' - x $ i\<bar> < d \<longrightarrow> a $ i < x' \<and> x' < b $ i)"] by auto
-
-    let ?d = "Min (range d)"
-    have **:"finite (range d)" "range d \<noteq> {}" by auto
-    have "?d>0" unfolding Min_gr_iff[OF **] using d by auto
-    moreover
-    { fix x' assume as:"dist x' x < ?d"
-      { fix i
-        have "\<bar>x'$i - x $ i\<bar> < d i"
-          using norm_bound_component_lt[OF as[unfolded dist_norm], of i]
-          unfolding vector_minus_component and Min_gr_iff[OF **] by auto
-        hence "a $ i < x' $ i" "x' $ i < b $ i" using d[THEN spec[where x=i]] by auto  }
-      hence "a < x' \<and> x' < b" unfolding vector_less_def by auto  }
-    ultimately have "\<exists>e>0. \<forall>x'. dist x' x < e \<longrightarrow> x' \<in> {a<..<b}" by (auto, rule_tac x="?d" in exI, simp)
-  }
-  thus ?thesis unfolding open_dist using open_interval_lemma by auto
-qed
-
-lemma closed_interval: fixes a :: "real^'n::finite" shows "closed {a .. b}"
-proof-
-  { fix x i assume as:"\<forall>e>0. \<exists>x'\<in>{a..b}. x' \<noteq> x \<and> dist x' x < e"(* and xab:"a$i > x$i \<or> b$i < x$i"*)
-    { assume xa:"a$i > x$i"
-      with as obtain y where y:"y\<in>{a..b}" "y \<noteq> x" "dist y x < a$i - x$i" by(erule_tac x="a$i - x$i" in allE)auto
-      hence False unfolding mem_interval and dist_norm
-        using component_le_norm[of "y-x" i, unfolded vector_minus_component] and xa by(auto elim!: allE[where x=i])
-    } hence "a$i \<le> x$i" by(rule ccontr)auto
-    moreover
-    { assume xb:"b$i < x$i"
-      with as obtain y where y:"y\<in>{a..b}" "y \<noteq> x" "dist y x < x$i - b$i" by(erule_tac x="x$i - b$i" in allE)auto
-      hence False unfolding mem_interval and dist_norm
-        using component_le_norm[of "y-x" i, unfolded vector_minus_component] and xb by(auto elim!: allE[where x=i])
-    } hence "x$i \<le> b$i" by(rule ccontr)auto
-    ultimately
-    have "a $ i \<le> x $ i \<and> x $ i \<le> b $ i" by auto }
-  thus ?thesis unfolding closed_limpt islimpt_approachable mem_interval by auto
-qed
-
-lemma interior_closed_interval: fixes a :: "real^'n::finite" shows
- "interior {a .. b} = {a<..<b}" (is "?L = ?R")
-proof(rule subset_antisym)
-  show "?R \<subseteq> ?L" using interior_maximal[OF interval_open_subset_closed open_interval] by auto
-next
-  { fix x assume "\<exists>T. open T \<and> x \<in> T \<and> T \<subseteq> {a..b}"
-    then obtain s where s:"open s" "x \<in> s" "s \<subseteq> {a..b}" by auto
-    then obtain e where "e>0" and e:"\<forall>x'. dist x' x < e \<longrightarrow> x' \<in> {a..b}" unfolding open_dist and subset_eq by auto
-    { fix i
-      have "dist (x - (e / 2) *\<^sub>R basis i) x < e"
-           "dist (x + (e / 2) *\<^sub>R basis i) x < e"
-        unfolding dist_norm apply auto
-        unfolding norm_minus_cancel using norm_basis[of i] and `e>0` by auto
-      hence "a $ i \<le> (x - (e / 2) *\<^sub>R basis i) $ i"
-                    "(x + (e / 2) *\<^sub>R basis i) $ i \<le> b $ i"
-        using e[THEN spec[where x="x - (e/2) *\<^sub>R basis i"]]
-        and   e[THEN spec[where x="x + (e/2) *\<^sub>R basis i"]]
-        unfolding mem_interval by (auto elim!: allE[where x=i])
-      hence "a $ i < x $ i" and "x $ i < b $ i"
-        unfolding vector_minus_component and vector_add_component
-        unfolding vector_smult_component and basis_component using `e>0` by auto   }
-    hence "x \<in> {a<..<b}" unfolding mem_interval by auto  }
-  thus "?L \<subseteq> ?R" unfolding interior_def and subset_eq by auto
-qed
-
-lemma bounded_closed_interval: fixes a :: "real^'n::finite" shows
- "bounded {a .. b}"
-proof-
-  let ?b = "\<Sum>i\<in>UNIV. \<bar>a$i\<bar> + \<bar>b$i\<bar>"
-  { fix x::"real^'n" assume x:"\<forall>i. a $ i \<le> x $ i \<and> x $ i \<le> b $ i"
-    { fix i
-      have "\<bar>x$i\<bar> \<le> \<bar>a$i\<bar> + \<bar>b$i\<bar>" using x[THEN spec[where x=i]] by auto  }
-    hence "(\<Sum>i\<in>UNIV. \<bar>x $ i\<bar>) \<le> ?b" by(rule setsum_mono)
-    hence "norm x \<le> ?b" using norm_le_l1[of x] by auto  }
-  thus ?thesis unfolding interval and bounded_iff by auto
-qed
-
-lemma bounded_interval: fixes a :: "real^'n::finite" shows
- "bounded {a .. b} \<and> bounded {a<..<b}"
-  using bounded_closed_interval[of a b]
-  using interval_open_subset_closed[of a b]
-  using bounded_subset[of "{a..b}" "{a<..<b}"]
-  by simp
-
-lemma not_interval_univ: fixes a :: "real^'n::finite" shows
- "({a .. b} \<noteq> UNIV) \<and> ({a<..<b} \<noteq> UNIV)"
-  using bounded_interval[of a b]
-  by auto
-
-lemma compact_interval: fixes a :: "real^'n::finite" shows
- "compact {a .. b}"
-  using bounded_closed_imp_compact using bounded_interval[of a b] using closed_interval[of a b] by auto
-
-lemma open_interval_midpoint: fixes a :: "real^'n::finite"
-  assumes "{a<..<b} \<noteq> {}" shows "((1/2) *\<^sub>R (a + b)) \<in> {a<..<b}"
-proof-
-  { fix i
-    have "a $ i < ((1 / 2) *\<^sub>R (a + b)) $ i \<and> ((1 / 2) *\<^sub>R (a + b)) $ i < b $ i"
-      using assms[unfolded interval_ne_empty, THEN spec[where x=i]]
-      unfolding vector_smult_component and vector_add_component
-      by(auto simp add: less_divide_eq_number_of1)  }
-  thus ?thesis unfolding mem_interval by auto
-qed
-
-lemma open_closed_interval_convex: fixes x :: "real^'n::finite"
-  assumes x:"x \<in> {a<..<b}" and y:"y \<in> {a .. b}" and e:"0 < e" "e \<le> 1"
-  shows "(e *\<^sub>R x + (1 - e) *\<^sub>R y) \<in> {a<..<b}"
-proof-
-  { fix i
-    have "a $ i = e * a$i + (1 - e) * a$i" unfolding left_diff_distrib by simp
-    also have "\<dots> < e * x $ i + (1 - e) * y $ i" apply(rule add_less_le_mono)
-      using e unfolding mult_less_cancel_left and mult_le_cancel_left apply simp_all
-      using x unfolding mem_interval  apply simp
-      using y unfolding mem_interval  apply simp
-      done
-    finally have "a $ i < (e *\<^sub>R x + (1 - e) *\<^sub>R y) $ i" by auto
-    moreover {
-    have "b $ i = e * b$i + (1 - e) * b$i" unfolding left_diff_distrib by simp
-    also have "\<dots> > e * x $ i + (1 - e) * y $ i" apply(rule add_less_le_mono)
-      using e unfolding mult_less_cancel_left and mult_le_cancel_left apply simp_all
-      using x unfolding mem_interval  apply simp
-      using y unfolding mem_interval  apply simp
-      done
-    finally have "(e *\<^sub>R x + (1 - e) *\<^sub>R y) $ i < b $ i" by auto
-    } ultimately have "a $ i < (e *\<^sub>R x + (1 - e) *\<^sub>R y) $ i \<and> (e *\<^sub>R x + (1 - e) *\<^sub>R y) $ i < b $ i" by auto }
-  thus ?thesis unfolding mem_interval by auto
-qed
-
-lemma closure_open_interval: fixes a :: "real^'n::finite"
-  assumes "{a<..<b} \<noteq> {}"
-  shows "closure {a<..<b} = {a .. b}"
-proof-
-  have ab:"a < b" using assms[unfolded interval_ne_empty] unfolding vector_less_def by auto
-  let ?c = "(1 / 2) *\<^sub>R (a + b)"
-  { fix x assume as:"x \<in> {a .. b}"
-    def f == "\<lambda>n::nat. x + (inverse (real n + 1)) *\<^sub>R (?c - x)"
-    { fix n assume fn:"f n < b \<longrightarrow> a < f n \<longrightarrow> f n = x" and xc:"x \<noteq> ?c"
-      have *:"0 < inverse (real n + 1)" "inverse (real n + 1) \<le> 1" unfolding inverse_le_1_iff by auto
-      have "(inverse (real n + 1)) *\<^sub>R ((1 / 2) *\<^sub>R (a + b)) + (1 - inverse (real n + 1)) *\<^sub>R x =
-        x + (inverse (real n + 1)) *\<^sub>R (((1 / 2) *\<^sub>R (a + b)) - x)"
-        by (auto simp add: algebra_simps)
-      hence "f n < b" and "a < f n" using open_closed_interval_convex[OF open_interval_midpoint[OF assms] as *] unfolding f_def by auto
-      hence False using fn unfolding f_def using xc by(auto simp add: vector_mul_lcancel vector_ssub_ldistrib)  }
-    moreover
-    { assume "\<not> (f ---> x) sequentially"
-      { fix e::real assume "e>0"
-        hence "\<exists>N::nat. inverse (real (N + 1)) < e" using real_arch_inv[of e] apply (auto simp add: Suc_pred') apply(rule_tac x="n - 1" in exI) by auto
-        then obtain N::nat where "inverse (real (N + 1)) < e" by auto
-        hence "\<forall>n\<ge>N. inverse (real n + 1) < e" by (auto, metis Suc_le_mono le_SucE less_imp_inverse_less nat_le_real_less order_less_trans real_of_nat_Suc real_of_nat_Suc_gt_zero)
-        hence "\<exists>N::nat. \<forall>n\<ge>N. inverse (real n + 1) < e" by auto  }
-      hence "((\<lambda>n. inverse (real n + 1)) ---> 0) sequentially"
-        unfolding Lim_sequentially by(auto simp add: dist_norm)
-      hence "(f ---> x) sequentially" unfolding f_def
-        using Lim_add[OF Lim_const, of "\<lambda>n::nat. (inverse (real n + 1)) *\<^sub>R ((1 / 2) *\<^sub>R (a + b) - x)" 0 sequentially x]
-        using Lim_vmul[of "\<lambda>n::nat. inverse (real n + 1)" 0 sequentially "((1 / 2) *\<^sub>R (a + b) - x)"] by auto  }
-    ultimately have "x \<in> closure {a<..<b}"
-      using as and open_interval_midpoint[OF assms] unfolding closure_def unfolding islimpt_sequential by(cases "x=?c")auto  }
-  thus ?thesis using closure_minimal[OF interval_open_subset_closed closed_interval, of a b] by blast
-qed
-
-lemma bounded_subset_open_interval_symmetric: fixes s::"(real^'n::finite) set"
-  assumes "bounded s"  shows "\<exists>a. s \<subseteq> {-a<..<a}"
-proof-
-  obtain b where "b>0" and b:"\<forall>x\<in>s. norm x \<le> b" using assms[unfolded bounded_pos] by auto
-  def a \<equiv> "(\<chi> i. b+1)::real^'n"
-  { fix x assume "x\<in>s"
-    fix i
-    have "(-a)$i < x$i" and "x$i < a$i" using b[THEN bspec[where x=x], OF `x\<in>s`] and component_le_norm[of x i]
-      unfolding vector_uminus_component and a_def and Cart_lambda_beta by auto
-  }
-  thus ?thesis by(auto intro: exI[where x=a] simp add: vector_less_def)
-qed
-
-lemma bounded_subset_open_interval:
-  fixes s :: "(real ^ 'n::finite) set"
-  shows "bounded s ==> (\<exists>a b. s \<subseteq> {a<..<b})"
-  by (auto dest!: bounded_subset_open_interval_symmetric)
-
-lemma bounded_subset_closed_interval_symmetric:
-  fixes s :: "(real ^ 'n::finite) set"
-  assumes "bounded s" shows "\<exists>a. s \<subseteq> {-a .. a}"
-proof-
-  obtain a where "s \<subseteq> {- a<..<a}" using bounded_subset_open_interval_symmetric[OF assms] by auto
-  thus ?thesis using interval_open_subset_closed[of "-a" a] by auto
-qed
-
-lemma bounded_subset_closed_interval:
-  fixes s :: "(real ^ 'n::finite) set"
-  shows "bounded s ==> (\<exists>a b. s \<subseteq> {a .. b})"
-  using bounded_subset_closed_interval_symmetric[of s] by auto
-
-lemma frontier_closed_interval:
-  fixes a b :: "real ^ _"
-  shows "frontier {a .. b} = {a .. b} - {a<..<b}"
-  unfolding frontier_def unfolding interior_closed_interval and closure_closed[OF closed_interval] ..
-
-lemma frontier_open_interval:
-  fixes a b :: "real ^ _"
-  shows "frontier {a<..<b} = (if {a<..<b} = {} then {} else {a .. b} - {a<..<b})"
-proof(cases "{a<..<b} = {}")
-  case True thus ?thesis using frontier_empty by auto
-next
-  case False thus ?thesis unfolding frontier_def and closure_open_interval[OF False] and interior_open[OF open_interval] by auto
-qed
-
-lemma inter_interval_mixed_eq_empty: fixes a :: "real^'n::finite"
-  assumes "{c<..<d} \<noteq> {}"  shows "{a<..<b} \<inter> {c .. d} = {} \<longleftrightarrow> {a<..<b} \<inter> {c<..<d} = {}"
-  unfolding closure_open_interval[OF assms, THEN sym] unfolding open_inter_closure_eq_empty[OF open_interval] ..
-
-
-(* Some special cases for intervals in R^1.                                  *)
-
-lemma all_1: "(\<forall>x::1. P x) \<longleftrightarrow> P 1"
-  by (metis num1_eq_iff)
-
-lemma ex_1: "(\<exists>x::1. P x) \<longleftrightarrow> P 1"
-  by auto (metis num1_eq_iff)
-
-lemma interval_cases_1: fixes x :: "real^1" shows
- "x \<in> {a .. b} ==> x \<in> {a<..<b} \<or> (x = a) \<or> (x = b)"
-  by(simp add:  Cart_eq vector_less_def vector_less_eq_def all_1, auto)
-
-lemma in_interval_1: fixes x :: "real^1" shows
- "(x \<in> {a .. b} \<longleftrightarrow> dest_vec1 a \<le> dest_vec1 x \<and> dest_vec1 x \<le> dest_vec1 b) \<and>
-  (x \<in> {a<..<b} \<longleftrightarrow> dest_vec1 a < dest_vec1 x \<and> dest_vec1 x < dest_vec1 b)"
-by(simp add: Cart_eq vector_less_def vector_less_eq_def all_1 dest_vec1_def)
-
-lemma interval_eq_empty_1: fixes a :: "real^1" shows
-  "{a .. b} = {} \<longleftrightarrow> dest_vec1 b < dest_vec1 a"
-  "{a<..<b} = {} \<longleftrightarrow> dest_vec1 b \<le> dest_vec1 a"
-  unfolding interval_eq_empty and ex_1 and dest_vec1_def by auto
-
-lemma subset_interval_1: fixes a :: "real^1" shows
- "({a .. b} \<subseteq> {c .. d} \<longleftrightarrow>  dest_vec1 b < dest_vec1 a \<or>
-                dest_vec1 c \<le> dest_vec1 a \<and> dest_vec1 a \<le> dest_vec1 b \<and> dest_vec1 b \<le> dest_vec1 d)"
- "({a .. b} \<subseteq> {c<..<d} \<longleftrightarrow>  dest_vec1 b < dest_vec1 a \<or>
-                dest_vec1 c < dest_vec1 a \<and> dest_vec1 a \<le> dest_vec1 b \<and> dest_vec1 b < dest_vec1 d)"
- "({a<..<b} \<subseteq> {c .. d} \<longleftrightarrow>  dest_vec1 b \<le> dest_vec1 a \<or>
-                dest_vec1 c \<le> dest_vec1 a \<and> dest_vec1 a < dest_vec1 b \<and> dest_vec1 b \<le> dest_vec1 d)"
- "({a<..<b} \<subseteq> {c<..<d} \<longleftrightarrow> dest_vec1 b \<le> dest_vec1 a \<or>
-                dest_vec1 c \<le> dest_vec1 a \<and> dest_vec1 a < dest_vec1 b \<and> dest_vec1 b \<le> dest_vec1 d)"
-  unfolding subset_interval[of a b c d] unfolding all_1 and dest_vec1_def by auto
-
-lemma eq_interval_1: fixes a :: "real^1" shows
- "{a .. b} = {c .. d} \<longleftrightarrow>
-          dest_vec1 b < dest_vec1 a \<and> dest_vec1 d < dest_vec1 c \<or>
-          dest_vec1 a = dest_vec1 c \<and> dest_vec1 b = dest_vec1 d"
-using set_eq_subset[of "{a .. b}" "{c .. d}"]
-using subset_interval_1(1)[of a b c d]
-using subset_interval_1(1)[of c d a b]
-by auto (* FIXME: slow *)
-
-lemma disjoint_interval_1: fixes a :: "real^1" shows
-  "{a .. b} \<inter> {c .. d} = {} \<longleftrightarrow> dest_vec1 b < dest_vec1 a \<or> dest_vec1 d < dest_vec1 c  \<or>  dest_vec1 b < dest_vec1 c \<or> dest_vec1 d < dest_vec1 a"
-  "{a .. b} \<inter> {c<..<d} = {} \<longleftrightarrow> dest_vec1 b < dest_vec1 a \<or> dest_vec1 d \<le> dest_vec1 c  \<or>  dest_vec1 b \<le> dest_vec1 c \<or> dest_vec1 d \<le> dest_vec1 a"
-  "{a<..<b} \<inter> {c .. d} = {} \<longleftrightarrow> dest_vec1 b \<le> dest_vec1 a \<or> dest_vec1 d < dest_vec1 c  \<or>  dest_vec1 b \<le> dest_vec1 c \<or> dest_vec1 d \<le> dest_vec1 a"
-  "{a<..<b} \<inter> {c<..<d} = {} \<longleftrightarrow> dest_vec1 b \<le> dest_vec1 a \<or> dest_vec1 d \<le> dest_vec1 c  \<or>  dest_vec1 b \<le> dest_vec1 c \<or> dest_vec1 d \<le> dest_vec1 a"
-  unfolding disjoint_interval and dest_vec1_def ex_1 by auto
-
-lemma open_closed_interval_1: fixes a :: "real^1" shows
- "{a<..<b} = {a .. b} - {a, b}"
-  unfolding expand_set_eq apply simp unfolding vector_less_def and vector_less_eq_def and all_1 and dest_vec1_eq[THEN sym] and dest_vec1_def by auto
-
-lemma closed_open_interval_1: "dest_vec1 (a::real^1) \<le> dest_vec1 b ==> {a .. b} = {a<..<b} \<union> {a,b}"
-  unfolding expand_set_eq apply simp unfolding vector_less_def and vector_less_eq_def and all_1 and dest_vec1_eq[THEN sym] and dest_vec1_def by auto
-
-(* Some stuff for half-infinite intervals too; FIXME: notation?  *)
-
-lemma closed_interval_left: fixes b::"real^'n::finite"
-  shows "closed {x::real^'n. \<forall>i. x$i \<le> b$i}"
-proof-
-  { fix i
-    fix x::"real^'n" assume x:"\<forall>e>0. \<exists>x'\<in>{x. \<forall>i. x $ i \<le> b $ i}. x' \<noteq> x \<and> dist x' x < e"
-    { assume "x$i > b$i"
-      then obtain y where "y $ i \<le> b $ i"  "y \<noteq> x"  "dist y x < x$i - b$i" using x[THEN spec[where x="x$i - b$i"]] by auto
-      hence False using component_le_norm[of "y - x" i] unfolding dist_norm and vector_minus_component by auto   }
-    hence "x$i \<le> b$i" by(rule ccontr)auto  }
-  thus ?thesis unfolding closed_limpt unfolding islimpt_approachable by blast
-qed
-
-lemma closed_interval_right: fixes a::"real^'n::finite"
-  shows "closed {x::real^'n. \<forall>i. a$i \<le> x$i}"
-proof-
-  { fix i
-    fix x::"real^'n" assume x:"\<forall>e>0. \<exists>x'\<in>{x. \<forall>i. a $ i \<le> x $ i}. x' \<noteq> x \<and> dist x' x < e"
-    { assume "a$i > x$i"
-      then obtain y where "a $ i \<le> y $ i"  "y \<noteq> x"  "dist y x < a$i - x$i" using x[THEN spec[where x="a$i - x$i"]] by auto
-      hence False using component_le_norm[of "y - x" i] unfolding dist_norm and vector_minus_component by auto   }
-    hence "a$i \<le> x$i" by(rule ccontr)auto  }
-  thus ?thesis unfolding closed_limpt unfolding islimpt_approachable by blast
-qed
-
-subsection{* Intervals in general, including infinite and mixtures of open and closed. *}
-
-definition "is_interval s \<longleftrightarrow> (\<forall>a\<in>s. \<forall>b\<in>s. \<forall>x. (\<forall>i. ((a$i \<le> x$i \<and> x$i \<le> b$i) \<or> (b$i \<le> x$i \<and> x$i \<le> a$i)))  \<longrightarrow> x \<in> s)"
-
-lemma is_interval_interval: "is_interval {a .. b::real^'n::finite}" (is ?th1) "is_interval {a<..<b}" (is ?th2) proof - 
-  have *:"\<And>x y z::real. x < y \<Longrightarrow> y < z \<Longrightarrow> 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 {}"
-  unfolding is_interval_def
-  by simp
-
-lemma is_interval_univ:
- "is_interval UNIV"
-  unfolding is_interval_def
-  by simp
-
-subsection{* Closure of halfspaces and hyperplanes.                                    *}
-
-lemma Lim_inner:
-  assumes "(f ---> l) net"  shows "((\<lambda>y. inner a (f y)) ---> inner a l) net"
-  by (intro tendsto_intros assms)
-
-lemma continuous_at_inner: "continuous (at x) (inner a)"
-  unfolding continuous_at by (intro tendsto_intros)
-
-lemma continuous_on_inner:
-  fixes s :: "'a::real_inner set"
-  shows "continuous_on s (inner a)"
-  unfolding continuous_on by (rule ballI) (intro tendsto_intros)
-
-lemma closed_halfspace_le: "closed {x. inner a x \<le> b}"
-proof-
-  have "\<forall>x. continuous (at x) (inner a)"
-    unfolding continuous_at by (rule allI) (intro tendsto_intros)
-  hence "closed (inner a -` {..b})"
-    using closed_real_atMost by (rule continuous_closed_vimage)
-  moreover have "{x. inner a x \<le> b} = inner a -` {..b}" by auto
-  ultimately show ?thesis by simp
-qed
-
-lemma closed_halfspace_ge: "closed {x. inner a x \<ge> b}"
-  using closed_halfspace_le[of "-a" "-b"] unfolding inner_minus_left by auto
-
-lemma closed_hyperplane: "closed {x. inner a x = b}"
-proof-
-  have "{x. inner a x = b} = {x. inner a x \<ge> b} \<inter> {x. inner a x \<le> b}" by auto
-  thus ?thesis using closed_halfspace_le[of a b] and closed_halfspace_ge[of b a] using closed_Int by auto
-qed
-
-lemma closed_halfspace_component_le:
-  shows "closed {x::real^'n::finite. x$i \<le> a}"
-  using closed_halfspace_le[of "(basis i)::real^'n" a] unfolding inner_basis[OF assms] by auto
-
-lemma closed_halfspace_component_ge:
-  shows "closed {x::real^'n::finite. x$i \<ge> a}"
-  using closed_halfspace_ge[of a "(basis i)::real^'n"] unfolding inner_basis[OF assms] by auto
-
-text{* Openness of halfspaces.                                                   *}
-
-lemma open_halfspace_lt: "open {x. inner a x < b}"
-proof-
-  have "UNIV - {x. b \<le> inner a x} = {x. inner a x < b}" by auto
-  thus ?thesis using closed_halfspace_ge[unfolded closed_def Compl_eq_Diff_UNIV, of b a] by auto
-qed
-
-lemma open_halfspace_gt: "open {x. inner a x > b}"
-proof-
-  have "UNIV - {x. b \<ge> inner a x} = {x. inner a x > b}" by auto
-  thus ?thesis using closed_halfspace_le[unfolded closed_def Compl_eq_Diff_UNIV, of a b] by auto
-qed
-
-lemma open_halfspace_component_lt:
-  shows "open {x::real^'n::finite. x$i < a}"
-  using open_halfspace_lt[of "(basis i)::real^'n" a] unfolding inner_basis[OF assms] by auto
-
-lemma open_halfspace_component_gt:
-  shows "open {x::real^'n::finite. x$i  > a}"
-  using open_halfspace_gt[of a "(basis i)::real^'n"] unfolding inner_basis[OF assms] by auto
-
-text{* This gives a simple derivation of limit component bounds.                 *}
-
-lemma Lim_component_le: fixes f :: "'a \<Rightarrow> real^'n::finite"
-  assumes "(f ---> l) net" "\<not> (trivial_limit net)"  "eventually (\<lambda>x. f(x)$i \<le> b) net"
-  shows "l$i \<le> b"
-proof-
-  { fix x have "x \<in> {x::real^'n. inner (basis i) x \<le> b} \<longleftrightarrow> x$i \<le> b" unfolding inner_basis by auto } note * = this
-  show ?thesis using Lim_in_closed_set[of "{x. inner (basis i) x \<le> b}" f net l] unfolding *
-    using closed_halfspace_le[of "(basis i)::real^'n" b] and assms(1,2,3) by auto
-qed
-
-lemma Lim_component_ge: fixes f :: "'a \<Rightarrow> real^'n::finite"
-  assumes "(f ---> l) net"  "\<not> (trivial_limit net)"  "eventually (\<lambda>x. b \<le> (f x)$i) net"
-  shows "b \<le> l$i"
-proof-
-  { fix x have "x \<in> {x::real^'n. inner (basis i) x \<ge> b} \<longleftrightarrow> x$i \<ge> b" unfolding inner_basis by auto } note * = this
-  show ?thesis using Lim_in_closed_set[of "{x. inner (basis i) x \<ge> b}" f net l] unfolding *
-    using closed_halfspace_ge[of b "(basis i)::real^'n"] and assms(1,2,3) by auto
-qed
-
-lemma Lim_component_eq: fixes f :: "'a \<Rightarrow> real^'n::finite"
-  assumes net:"(f ---> l) net" "~(trivial_limit net)" and ev:"eventually (\<lambda>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
-
-lemma Lim_drop_le: fixes f :: "'a \<Rightarrow> real^1" shows
-  "(f ---> l) net \<Longrightarrow> ~(trivial_limit net) \<Longrightarrow> eventually (\<lambda>x. dest_vec1 (f x) \<le> b) net ==> dest_vec1 l \<le> b"
-  using Lim_component_le[of f l net 1 b] unfolding dest_vec1_def by auto
-
-lemma Lim_drop_ge: fixes f :: "'a \<Rightarrow> real^1" shows
- "(f ---> l) net \<Longrightarrow> ~(trivial_limit net) \<Longrightarrow> eventually (\<lambda>x. b \<le> dest_vec1 (f x)) net ==> b \<le> dest_vec1 l"
-  using Lim_component_ge[of f l net b 1] unfolding dest_vec1_def by auto
-
-text{* Limits relative to a union.                                               *}
-
-lemma eventually_within_Un:
-  "eventually P (net within (s \<union> t)) \<longleftrightarrow>
-    eventually P (net within s) \<and> eventually P (net within t)"
-  unfolding Limits.eventually_within
-  by (auto elim!: eventually_rev_mp)
-
-lemma Lim_within_union:
- "(f ---> l) (net within (s \<union> t)) \<longleftrightarrow>
-  (f ---> l) (net within s) \<and> (f ---> l) (net within t)"
-  unfolding tendsto_def
-  by (auto simp add: eventually_within_Un)
-
-lemma continuous_on_union:
-  assumes "closed s" "closed t" "continuous_on s f" "continuous_on t f"
-  shows "continuous_on (s \<union> t) f"
-  using assms unfolding continuous_on unfolding Lim_within_union
-  unfolding Lim unfolding trivial_limit_within unfolding closed_limpt by auto
-
-lemma continuous_on_cases:
-  assumes "closed s" "closed t" "continuous_on s f" "continuous_on t g"
-          "\<forall>x. (x\<in>s \<and> \<not> P x) \<or> (x \<in> t \<and> P x) \<longrightarrow> f x = g x"
-  shows "continuous_on (s \<union> t) (\<lambda>x. if P x then f x else g x)"
-proof-
-  let ?h = "(\<lambda>x. if P x then f x else g x)"
-  have "\<forall>x\<in>s. f x = (if P x then f x else g x)" using assms(5) by auto
-  hence "continuous_on s ?h" using continuous_on_eq[of s f ?h] using assms(3) by auto
-  moreover
-  have "\<forall>x\<in>t. g x = (if P x then f x else g x)" using assms(5) by auto
-  hence "continuous_on t ?h" using continuous_on_eq[of t g ?h] using assms(4) by auto
-  ultimately show ?thesis using continuous_on_union[OF assms(1,2), of ?h] by auto
-qed
-
-
-text{* Some more convenient intermediate-value theorem formulations.             *}
-
-lemma connected_ivt_hyperplane:
-  assumes "connected s" "x \<in> s" "y \<in> s" "inner a x \<le> b" "b \<le> inner a y"
-  shows "\<exists>z \<in> s. inner a z = b"
-proof(rule ccontr)
-  assume as:"\<not> (\<exists>z\<in>s. inner a z = b)"
-  let ?A = "{x. inner a x < b}"
-  let ?B = "{x. inner a x > b}"
-  have "open ?A" "open ?B" using open_halfspace_lt and open_halfspace_gt by auto
-  moreover have "?A \<inter> ?B = {}" by auto
-  moreover have "s \<subseteq> ?A \<union> ?B" using as by auto
-  ultimately show False using assms(1)[unfolded connected_def not_ex, THEN spec[where x="?A"], THEN spec[where x="?B"]] and assms(2-5) by auto
-qed
-
-lemma connected_ivt_component: fixes x::"real^'n::finite" shows
- "connected s \<Longrightarrow> x \<in> s \<Longrightarrow> y \<in> s \<Longrightarrow> x$k \<le> a \<Longrightarrow> a \<le> y$k \<Longrightarrow> (\<exists>z\<in>s.  z$k = a)"
-  using connected_ivt_hyperplane[of s x y "(basis k)::real^'n" a] by (auto simp add: inner_basis)
-
-text{* Also more convenient formulations of monotone convergence.                *}
-
-lemma bounded_increasing_convergent: fixes s::"nat \<Rightarrow> real^1"
-  assumes "bounded {s n| n::nat. True}"  "\<forall>n. dest_vec1(s n) \<le> dest_vec1(s(Suc n))"
-  shows "\<exists>l. (s ---> l) sequentially"
-proof-
-  obtain a where a:"\<forall>n. \<bar>dest_vec1 (s n)\<bar> \<le>  a" using assms(1)[unfolded bounded_iff abs_dest_vec1] by auto
-  { fix m::nat
-    have "\<And> n. n\<ge>m \<longrightarrow> dest_vec1 (s m) \<le> dest_vec1 (s n)"
-      apply(induct_tac n) apply simp using assms(2) apply(erule_tac x="na" in allE) by(auto simp add: not_less_eq_eq)  }
-  hence "\<forall>m n. m \<le> n \<longrightarrow> dest_vec1 (s m) \<le> dest_vec1 (s n)" by auto
-  then obtain l where "\<forall>e>0. \<exists>N. \<forall>n\<ge>N. \<bar>dest_vec1 (s n) - l\<bar> < e" using convergent_bounded_monotone[OF a] unfolding monoseq_def by auto
-  thus ?thesis unfolding Lim_sequentially apply(rule_tac x="vec1 l" in exI)
-    unfolding dist_norm unfolding abs_dest_vec1 and dest_vec1_sub by auto
-qed
-
-subsection{* Basic homeomorphism definitions.                                          *}
-
-definition "homeomorphism s t f g \<equiv>
-     (\<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::metric_space set \<Rightarrow> 'b::metric_space set \<Rightarrow> bool"
-    (infixr "homeomorphic" 60) where
-  homeomorphic_def: "s homeomorphic t \<equiv> (\<exists>f g. homeomorphism s t f g)"
-
-lemma homeomorphic_refl: "s homeomorphic s"
-  unfolding homeomorphic_def
-  unfolding homeomorphism_def
-  using continuous_on_id
-  apply(rule_tac x = "(\<lambda>x. x)" in exI)
-  apply(rule_tac x = "(\<lambda>x. x)" in exI)
-  by blast
-
-lemma homeomorphic_sym:
- "s homeomorphic t \<longleftrightarrow> t homeomorphic s"
-unfolding homeomorphic_def
-unfolding homeomorphism_def
-by blast (* FIXME: slow *)
-
-lemma homeomorphic_trans:
-  assumes "s homeomorphic t" "t homeomorphic u" shows "s homeomorphic u"
-proof-
-  obtain f1 g1 where fg1:"\<forall>x\<in>s. g1 (f1 x) = x"  "f1 ` s = t" "continuous_on s f1" "\<forall>y\<in>t. f1 (g1 y) = y" "g1 ` t = s" "continuous_on t g1"
-    using assms(1) unfolding homeomorphic_def homeomorphism_def by auto
-  obtain f2 g2 where fg2:"\<forall>x\<in>t. g2 (f2 x) = x"  "f2 ` t = u" "continuous_on t f2" "\<forall>y\<in>u. f2 (g2 y) = y" "g2 ` u = t" "continuous_on u g2"
-    using assms(2) unfolding homeomorphic_def homeomorphism_def by auto
-
-  { fix x assume "x\<in>s" hence "(g1 \<circ> g2) ((f2 \<circ> f1) x) = x" using fg1(1)[THEN bspec[where x=x]] and fg2(1)[THEN bspec[where x="f1 x"]] and fg1(2) by auto }
-  moreover have "(f2 \<circ> f1) ` s = u" using fg1(2) fg2(2) by auto
-  moreover have "continuous_on s (f2 \<circ> f1)" using continuous_on_compose[OF fg1(3)] and fg2(3) unfolding fg1(2) by auto
-  moreover { fix y assume "y\<in>u" hence "(f2 \<circ> f1) ((g1 \<circ> g2) y) = y" using fg2(4)[THEN bspec[where x=y]] and fg1(4)[THEN bspec[where x="g2 y"]] and fg2(5) by auto }
-  moreover have "(g1 \<circ> g2) ` u = s" using fg1(5) fg2(5) by auto
-  moreover have "continuous_on u (g1 \<circ> g2)" using continuous_on_compose[OF fg2(6)] and fg1(6)  unfolding fg2(5) by auto
-  ultimately show ?thesis unfolding homeomorphic_def homeomorphism_def apply(rule_tac x="f2 \<circ> f1" in exI) apply(rule_tac x="g1 \<circ> g2" in exI) by auto
-qed
-
-lemma homeomorphic_minimal:
- "s homeomorphic t \<longleftrightarrow>
-    (\<exists>f g. (\<forall>x\<in>s. f(x) \<in> t \<and> (g(f(x)) = x)) \<and>
-           (\<forall>y\<in>t. g(y) \<in> s \<and> (f(g(y)) = y)) \<and>
-           continuous_on s f \<and> continuous_on t g)"
-unfolding homeomorphic_def homeomorphism_def
-apply auto apply (rule_tac x=f in exI) apply (rule_tac x=g in exI)
-apply auto apply (rule_tac x=f in exI) apply (rule_tac x=g in exI) apply auto
-unfolding image_iff
-apply(erule_tac x="g x" in ballE) apply(erule_tac x="x" in ballE)
-apply auto apply(rule_tac x="g x" in bexI) apply auto
-apply(erule_tac x="f x" in ballE) apply(erule_tac x="x" in ballE)
-apply auto apply(rule_tac x="f x" in bexI) by auto
-
-subsection{* Relatively weak hypotheses if a set is compact.                           *}
-
-definition "inv_on f s = (\<lambda>x. SOME y. y\<in>s \<and> f y = x)"
-
-lemma assumes "inj_on f s" "x\<in>s"
-  shows "inv_on f s (f x) = x"
- using assms unfolding inj_on_def inv_on_def by auto
-
-lemma homeomorphism_compact:
-  fixes f :: "'a::heine_borel \<Rightarrow> 'b::heine_borel"
-    (* class constraint due to continuous_on_inverse *)
-  assumes "compact s" "continuous_on s f"  "f ` s = t"  "inj_on f s"
-  shows "\<exists>g. homeomorphism s t f g"
-proof-
-  def g \<equiv> "\<lambda>x. SOME y. y\<in>s \<and> f y = x"
-  have g:"\<forall>x\<in>s. g (f x) = x" using assms(3) assms(4)[unfolded inj_on_def] unfolding g_def by auto
-  { fix y assume "y\<in>t"
-    then obtain x where x:"f x = y" "x\<in>s" using assms(3) by auto
-    hence "g (f x) = x" using g by auto
-    hence "f (g y) = y" unfolding x(1)[THEN sym] by auto  }
-  hence g':"\<forall>x\<in>t. f (g x) = x" by auto
-  moreover
-  { fix x
-    have "x\<in>s \<Longrightarrow> x \<in> g ` t" using g[THEN bspec[where x=x]] unfolding image_iff using assms(3) by(auto intro!: bexI[where x="f x"])
-    moreover
-    { assume "x\<in>g ` t"
-      then obtain y where y:"y\<in>t" "g y = x" by auto
-      then obtain x' where x':"x'\<in>s" "f x' = y" using assms(3) by auto
-      hence "x \<in> s" unfolding g_def using someI2[of "\<lambda>b. b\<in>s \<and> f b = y" x' "\<lambda>x. x\<in>s"] unfolding y(2)[THEN sym] and g_def by auto }
-    ultimately have "x\<in>s \<longleftrightarrow> x \<in> g ` t" by auto  }
-  hence "g ` t = s" by auto
-  ultimately
-  show ?thesis unfolding homeomorphism_def homeomorphic_def
-    apply(rule_tac x=g in exI) using g and assms(3) and continuous_on_inverse[OF assms(2,1), of g, unfolded assms(3)] and assms(2) by auto
-qed
-
-lemma homeomorphic_compact:
-  fixes f :: "'a::heine_borel \<Rightarrow> 'b::heine_borel"
-    (* class constraint due to continuous_on_inverse *)
-  shows "compact s \<Longrightarrow> continuous_on s f \<Longrightarrow> (f ` s = t) \<Longrightarrow> inj_on f s
-          \<Longrightarrow> s homeomorphic t"
-  unfolding homeomorphic_def by(metis homeomorphism_compact)
-
-text{* Preservation of topological properties.                                   *}
-
-lemma homeomorphic_compactness:
- "s homeomorphic t ==> (compact s \<longleftrightarrow> compact t)"
-unfolding homeomorphic_def homeomorphism_def
-by (metis compact_continuous_image)
-
-text{* Results on translation, scaling etc.                                      *}
-
-lemma homeomorphic_scaling:
-  fixes s :: "'a::real_normed_vector set"
-  assumes "c \<noteq> 0"  shows "s homeomorphic ((\<lambda>x. c *\<^sub>R x) ` s)"
-  unfolding homeomorphic_minimal
-  apply(rule_tac x="\<lambda>x. c *\<^sub>R x" in exI)
-  apply(rule_tac x="\<lambda>x. (1 / c) *\<^sub>R x" in exI)
-  using assms apply auto
-  using continuous_on_cmul[OF continuous_on_id] by auto
-
-lemma homeomorphic_translation:
-  fixes s :: "'a::real_normed_vector set"
-  shows "s homeomorphic ((\<lambda>x. a + x) ` s)"
-  unfolding homeomorphic_minimal
-  apply(rule_tac x="\<lambda>x. a + x" in exI)
-  apply(rule_tac x="\<lambda>x. -a + x" in exI)
-  using continuous_on_add[OF continuous_on_const continuous_on_id] by auto
-
-lemma homeomorphic_affinity:
-  fixes s :: "'a::real_normed_vector set"
-  assumes "c \<noteq> 0"  shows "s homeomorphic ((\<lambda>x. a + c *\<^sub>R x) ` s)"
-proof-
-  have *:"op + a ` op *\<^sub>R c ` s = (\<lambda>x. a + c *\<^sub>R x) ` s" by auto
-  show ?thesis
-    using homeomorphic_trans
-    using homeomorphic_scaling[OF assms, of s]
-    using homeomorphic_translation[of "(\<lambda>x. c *\<^sub>R x) ` s" a] unfolding * by auto
-qed
-
-lemma homeomorphic_balls:
-  fixes a b ::"'a::real_normed_vector" (* FIXME: generalize to metric_space *)
-  assumes "0 < d"  "0 < e"
-  shows "(ball a d) homeomorphic  (ball b e)" (is ?th)
-        "(cball a d) homeomorphic (cball b e)" (is ?cth)
-proof-
-  have *:"\<bar>e / d\<bar> > 0" "\<bar>d / e\<bar> >0" using assms using divide_pos_pos by auto
-  show ?th unfolding homeomorphic_minimal
-    apply(rule_tac x="\<lambda>x. b + (e/d) *\<^sub>R (x - a)" in exI)
-    apply(rule_tac x="\<lambda>x. a + (d/e) *\<^sub>R (x - b)" in exI)
-    using assms apply (auto simp add: dist_commute)
-    unfolding dist_norm
-    apply (auto simp add: pos_divide_less_eq mult_strict_left_mono)
-    unfolding continuous_on
-    by (intro ballI tendsto_intros, simp, assumption)+
-next
-  have *:"\<bar>e / d\<bar> > 0" "\<bar>d / e\<bar> >0" using assms using divide_pos_pos by auto
-  show ?cth unfolding homeomorphic_minimal
-    apply(rule_tac x="\<lambda>x. b + (e/d) *\<^sub>R (x - a)" in exI)
-    apply(rule_tac x="\<lambda>x. a + (d/e) *\<^sub>R (x - b)" in exI)
-    using assms apply (auto simp add: dist_commute)
-    unfolding dist_norm
-    apply (auto simp add: pos_divide_le_eq)
-    unfolding continuous_on
-    by (intro ballI tendsto_intros, simp, assumption)+
-qed
-
-text{* "Isometry" (up to constant bounds) of injective linear map etc.           *}
-
-lemma cauchy_isometric:
-  fixes x :: "nat \<Rightarrow> real ^ 'n::finite"
-  assumes e:"0 < e" 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)"
-  shows "Cauchy x"
-proof-
-  interpret f: bounded_linear f by fact
-  { fix d::real assume "d>0"
-    then obtain N where N:"\<forall>n\<ge>N. norm (f (x n) - f (x N)) < e * d"
-      using cf[unfolded cauchy o_def dist_norm, THEN spec[where x="e*d"]] and e and mult_pos_pos[of e d] by auto
-    { fix n assume "n\<ge>N"
-      hence "norm (f (x n - x N)) < e * d" using N[THEN spec[where x=n]] unfolding f.diff[THEN sym] by auto
-      moreover have "e * norm (x n - x N) \<le> norm (f (x n - x N))"
-        using subspace_sub[OF s, of "x n" "x N"] using xs[THEN spec[where x=N]] and xs[THEN spec[where x=n]]
-        using normf[THEN bspec[where x="x n - x N"]] by auto
-      ultimately have "norm (x n - x N) < d" using `e>0`
-        using mult_left_less_imp_less[of e "norm (x n - x N)" d] by auto   }
-    hence "\<exists>N. \<forall>n\<ge>N. norm (x n - x N) < d" by auto }
-  thus ?thesis unfolding cauchy and dist_norm by auto
-qed
-
-lemma complete_isometric_image:
-  fixes f :: "real ^ _ \<Rightarrow> real ^ _"
-  assumes "0 < e" and s:"subspace s" and f:"bounded_linear f" and normf:"\<forall>x\<in>s. norm(f x) \<ge> e * norm(x)" and cs:"complete s"
-  shows "complete(f ` s)"
-proof-
-  { 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)" unfolding image_iff and Bex_def
-      using choice[of "\<lambda> n xa. xa \<in> s \<and> g n = f xa"] by auto
-    hence x:"\<forall>n. x n \<in> s"  "\<forall>n. g n = f (x n)" by auto
-    hence "f \<circ> x = g" unfolding expand_fun_eq 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
-    hence "\<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  }
-  thus ?thesis unfolding complete_def by auto
-qed
-
-lemma dist_0_norm:
-  fixes x :: "'a::real_normed_vector"
-  shows "dist 0 x = norm x"
-unfolding dist_norm by simp
-
-lemma injective_imp_isometric: fixes f::"real^'m::finite \<Rightarrow> real^'n::finite"
-  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)"
-proof(cases "s \<subseteq> {0::real^'m}")
-  case True
-  { fix x assume "x \<in> s"
-    hence "x = 0" using True by auto
-    hence "norm x \<le> norm (f x)" by auto  }
-  thus ?thesis by(auto intro!: exI[where x=1])
-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
-  let ?S = "{f x| x. (x \<in> s \<and> norm x = norm a)}"
-  let ?S' = "{x::real^'m. x\<in>s \<and> norm x = norm a}"
-  let ?S'' = "{x::real^'m. norm x = norm a}"
-
-  have "?S'' = frontier(cball 0 (norm a))" unfolding frontier_cball and dist_norm by (auto simp add: norm_minus_cancel)
-  hence "compact ?S''" using compact_frontier[OF compact_cball, of 0 "norm a"] by auto
-  moreover have "?S' = s \<inter> ?S''" by auto
-  ultimately have "compact ?S'" using closed_inter_compact[of s ?S''] using s(1) by auto
-  moreover have *:"f ` ?S' = ?S" by auto
-  ultimately have "compact ?S" using compact_continuous_image[OF linear_continuous_on[OF f(1)], of ?S'] by auto
-  hence "closed ?S" using compact_imp_closed by auto
-  moreover have "?S \<noteq> {}" using a by auto
-  ultimately obtain b' where "b'\<in>?S" "\<forall>y\<in>?S. norm b' \<le> norm y" using distance_attains_inf[of ?S 0] unfolding dist_0_norm by auto
-  then obtain b where "b\<in>s" and ba:"norm b = norm a" and b:"\<forall>x\<in>{x \<in> s. norm x = norm a}. norm (f b) \<le> norm (f x)" unfolding *[THEN sym] unfolding image_iff by auto
-
-  let ?e = "norm (f b) / norm b"
-  have "norm b > 0" using ba and a and norm_ge_zero by auto
-  moreover have "norm (f b) > 0" using f(2)[THEN bspec[where x=b], OF `b\<in>s`] using `norm b >0` unfolding zero_less_norm_iff by auto
-  ultimately have "0 < norm (f b) / norm b" by(simp only: divide_pos_pos)
-  moreover
-  { fix x assume "x\<in>s"
-    hence "norm (f b) / norm b * norm x \<le> norm (f x)"
-    proof(cases "x=0")
-      case True thus "norm (f b) / norm b * norm x \<le> norm (f x)" by auto
-    next
-      case False
-      hence *:"0 < norm a / norm x" using `a\<noteq>0` unfolding zero_less_norm_iff[THEN sym] by(simp only: divide_pos_pos)
-      have "\<forall>c. \<forall>x\<in>s. c *\<^sub>R x \<in> s" using s[unfolded subspace_def smult_conv_scaleR] by auto
-      hence "(norm a / norm x) *\<^sub>R x \<in> {x \<in> s. norm x = norm a}" using `x\<in>s` and `x\<noteq>0` by auto
-      thus "norm (f b) / norm b * norm x \<le> norm (f x)" using b[THEN bspec[where x="(norm a / norm x) *\<^sub>R x"]]
-        unfolding f.scaleR and ba using `x\<noteq>0` `a\<noteq>0`
-        by (auto simp add: real_mult_commute pos_le_divide_eq pos_divide_le_eq)
-    qed }
-  ultimately
-  show ?thesis by auto
-qed
-
-lemma closed_injective_image_subspace:
-  fixes f :: "real ^ _ \<Rightarrow> real ^ _"
-  assumes "subspace s" "bounded_linear f" "\<forall>x\<in>s. f x = 0 --> x = 0" "closed s"
-  shows "closed(f ` s)"
-proof-
-  obtain e where "e>0" and e:"\<forall>x\<in>s. e * norm x \<le> norm (f x)" using injective_imp_isometric[OF assms(4,1,2,3)] by auto
-  show ?thesis using complete_isometric_image[OF `e>0` assms(1,2) e] and assms(4)
-    unfolding complete_eq_closed[THEN sym] by auto
-qed
-
-subsection{* Some properties of a canonical subspace.                                  *}
-
-lemma subspace_substandard:
- "subspace {x::real^'n. (\<forall>i. P i \<longrightarrow> x$i = 0)}"
-  unfolding subspace_def by(auto simp add: vector_add_component vector_smult_component elim!: ballE)
-
-lemma closed_substandard:
- "closed {x::real^'n::finite. \<forall>i. P i --> x$i = 0}" (is "closed ?A")
-proof-
-  let ?D = "{i. P i}"
-  let ?Bs = "{{x::real^'n. inner (basis i) x = 0}| i. i \<in> ?D}"
-  { fix x
-    { assume "x\<in>?A"
-      hence x:"\<forall>i\<in>?D. x $ i = 0" by auto
-      hence "x\<in> \<Inter> ?Bs" by(auto simp add: inner_basis x) }
-    moreover
-    { assume x:"x\<in>\<Inter>?Bs"
-      { fix i assume i:"i \<in> ?D"
-        then obtain B where BB:"B \<in> ?Bs" and B:"B = {x::real^'n. inner (basis i) x = 0}" by auto
-        hence "x $ i = 0" unfolding B using x unfolding inner_basis by auto  }
-      hence "x\<in>?A" by auto }
-    ultimately have "x\<in>?A \<longleftrightarrow> x\<in> \<Inter>?Bs" by auto }
-  hence "?A = \<Inter> ?Bs" by auto
-  thus ?thesis by(auto simp add: closed_Inter closed_hyperplane)
-qed
-
-lemma dim_substandard:
-  shows "dim {x::real^'n::finite. \<forall>i. i \<notin> d \<longrightarrow> x$i = 0} = card d" (is "dim ?A = _")
-proof-
-  let ?D = "UNIV::'n set"
-  let ?B = "(basis::'n\<Rightarrow>real^'n) ` d"
-
-    let ?bas = "basis::'n \<Rightarrow> real^'n"
-
-  have "?B \<subseteq> ?A" by auto
-
-  moreover
-  { fix x::"real^'n" assume "x\<in>?A"
-    with finite[of d]
-    have "x\<in> span ?B"
-    proof(induct d arbitrary: x)
-      case empty hence "x=0" unfolding Cart_eq by auto
-      thus ?case using subspace_0[OF subspace_span[of "{}"]] by auto
-    next
-      case (insert k F)
-      hence *:"\<forall>i. i \<notin> insert k F \<longrightarrow> x $ i = 0" by auto
-      have **:"F \<subseteq> insert k F" by auto
-      def y \<equiv> "x - x$k *\<^sub>R basis k"
-      have y:"x = y + (x$k) *\<^sub>R basis k" unfolding y_def by auto
-      { fix i assume i':"i \<notin> F"
-        hence "y $ i = 0" unfolding y_def unfolding vector_minus_component
-          and vector_smult_component and basis_component
-          using *[THEN spec[where x=i]] by auto }
-      hence "y \<in> span (basis ` (insert k F))" using insert(3)
-        using span_mono[of "?bas ` F" "?bas ` (insert k F)"]
-        using image_mono[OF **, of basis] by auto
-      moreover
-      have "basis k \<in> span (?bas ` (insert k F))" by(rule span_superset, auto)
-      hence "x$k *\<^sub>R basis k \<in> span (?bas ` (insert k F))"
-        using span_mul [where 'a=real, unfolded smult_conv_scaleR] by auto
-      ultimately
-      have "y + x$k *\<^sub>R basis k \<in> span (?bas ` (insert k F))"
-        using span_add by auto
-      thus ?case using y by auto
-    qed
-  }
-  hence "?A \<subseteq> span ?B" by auto
-
-  moreover
-  { fix x assume "x \<in> ?B"
-    hence "x\<in>{(basis i)::real^'n |i. i \<in> ?D}" using assms by auto  }
-  hence "independent ?B" using independent_mono[OF independent_stdbasis, of ?B] and assms by auto
-
-  moreover
-  have "d \<subseteq> ?D" unfolding subset_eq using assms by auto
-  hence *:"inj_on (basis::'n\<Rightarrow>real^'n) d" using subset_inj_on[OF basis_inj, of "d"] by auto
-  have "?B hassize (card d)" unfolding hassize_def and card_image[OF *] by auto
-
-  ultimately show ?thesis using dim_unique[of "basis ` d" ?A] by auto
-qed
-
-text{* Hence closure and completeness of all subspaces.                          *}
-
-lemma closed_subspace_lemma: "n \<le> card (UNIV::'n::finite set) \<Longrightarrow> \<exists>A::'n set. card A = n"
-apply (induct n)
-apply (rule_tac x="{}" in exI, simp)
-apply clarsimp
-apply (subgoal_tac "\<exists>x. x \<notin> A")
-apply (erule exE)
-apply (rule_tac x="insert x A" in exI, simp)
-apply (subgoal_tac "A \<noteq> UNIV", auto)
-done
-
-lemma closed_subspace: fixes s::"(real^'n::finite) set"
-  assumes "subspace s" shows "closed s"
-proof-
-  have "dim s \<le> card (UNIV :: 'n set)" using dim_subset_univ by auto
-  then obtain d::"'n set" where t: "card d = dim s"
-    using closed_subspace_lemma by auto
-  let ?t = "{x::real^'n. \<forall>i. i \<notin> d \<longrightarrow> x$i = 0}"
-  obtain f where f:"bounded_linear f"  "f ` ?t = s" "inj_on f ?t"
-    using subspace_isomorphism[unfolded linear_conv_bounded_linear, OF subspace_substandard[of "\<lambda>i. i \<notin> d"] assms]
-    using dim_substandard[of d] and t by auto
-  interpret f: bounded_linear f by fact
-  have "\<forall>x\<in>?t. f x = 0 \<longrightarrow> x = 0" using f.zero using f(3)[unfolded inj_on_def]
-    by(erule_tac x=0 in ballE) auto
-  moreover have "closed ?t" using closed_substandard .
-  moreover have "subspace ?t" using subspace_substandard .
-  ultimately show ?thesis using closed_injective_image_subspace[of ?t f]
-    unfolding f(2) using f(1) by auto
-qed
-
-lemma complete_subspace:
-  fixes s :: "(real ^ _) set" shows "subspace s ==> complete s"
-  using complete_eq_closed closed_subspace
-  by auto
-
-lemma dim_closure:
-  fixes s :: "(real ^ _) set"
-  shows "dim(closure s) = dim s" (is "?dc = ?d")
-proof-
-  have "?dc \<le> ?d" using closure_minimal[OF span_inc, of s]
-    using closed_subspace[OF subspace_span, of s]
-    using dim_subset[of "closure s" "span s"] unfolding dim_span by auto
-  thus ?thesis using dim_subset[OF closure_subset, of s] by auto
-qed
-
-text{* Affine transformations of intervals.                                      *}
-
-lemma affinity_inverses:
-  assumes m0: "m \<noteq> (0::'a::field)"
-  shows "(\<lambda>x. m *s x + c) o (\<lambda>x. inverse(m) *s x + (-(inverse(m) *s c))) = id"
-  "(\<lambda>x. inverse(m) *s x + (-(inverse(m) *s c))) o (\<lambda>x. m *s x + c) = id"
-  using m0
-apply (auto simp add: expand_fun_eq vector_add_ldistrib vector_smult_assoc)
-by (simp add: vector_smult_lneg[symmetric] vector_smult_assoc vector_sneg_minus1[symmetric])
-
-lemma real_affinity_le:
- "0 < (m::'a::ordered_field) ==> (m * x + c \<le> y \<longleftrightarrow> x \<le> inverse(m) * y + -(c / m))"
-  by (simp add: field_simps inverse_eq_divide)
-
-lemma real_le_affinity:
- "0 < (m::'a::ordered_field) ==> (y \<le> m * x + c \<longleftrightarrow> inverse(m) * y + -(c / m) \<le> x)"
-  by (simp add: field_simps inverse_eq_divide)
-
-lemma real_affinity_lt:
- "0 < (m::'a::ordered_field) ==> (m * x + c < y \<longleftrightarrow> x < inverse(m) * y + -(c / m))"
-  by (simp add: field_simps inverse_eq_divide)
-
-lemma real_lt_affinity:
- "0 < (m::'a::ordered_field) ==> (y < m * x + c \<longleftrightarrow> inverse(m) * y + -(c / m) < x)"
-  by (simp add: field_simps inverse_eq_divide)
-
-lemma real_affinity_eq:
- "(m::'a::ordered_field) \<noteq> 0 ==> (m * x + c = y \<longleftrightarrow> x = inverse(m) * y + -(c / m))"
-  by (simp add: field_simps inverse_eq_divide)
-
-lemma real_eq_affinity:
- "(m::'a::ordered_field) \<noteq> 0 ==> (y = m * x + c  \<longleftrightarrow> inverse(m) * y + -(c / m) = x)"
-  by (simp add: field_simps inverse_eq_divide)
-
-lemma vector_affinity_eq:
-  assumes m0: "(m::'a::field) \<noteq> 0"
-  shows "m *s x + c = y \<longleftrightarrow> x = inverse m *s y + -(inverse m *s c)"
-proof
-  assume h: "m *s x + c = y"
-  hence "m *s x = y - c" by (simp add: ring_simps)
-  hence "inverse m *s (m *s x) = inverse m *s (y - c)" by simp
-  then show "x = inverse m *s y + - (inverse m *s c)"
-    using m0 by (simp add: vector_smult_assoc vector_ssub_ldistrib)
-next
-  assume h: "x = inverse m *s y + - (inverse m *s c)"
-  show "m *s x + c = y" unfolding h diff_minus[symmetric]
-    using m0 by (simp add: vector_smult_assoc vector_ssub_ldistrib)
-qed
-
-lemma vector_eq_affinity:
- "(m::'a::field) \<noteq> 0 ==> (y = m *s x + c \<longleftrightarrow> inverse(m) *s y + -(inverse(m) *s c) = x)"
-  using vector_affinity_eq[where m=m and x=x and y=y and c=c]
-  by metis
-
-lemma image_affinity_interval: fixes m::real
-  fixes a b c :: "real^'n::finite"
-  shows "(\<lambda>x. m *\<^sub>R x + c) ` {a .. b} =
-            (if {a .. b} = {} then {}
-            else (if 0 \<le> m then {m *\<^sub>R a + c .. m *\<^sub>R b + c}
-            else {m *\<^sub>R b + c .. m *\<^sub>R a + c}))"
-proof(cases "m=0")
-  { fix x assume "x \<le> c" "c \<le> x"
-    hence "x=c" unfolding vector_less_eq_def and Cart_eq by (auto intro: order_antisym) }
-  moreover case True
-  moreover have "c \<in> {m *\<^sub>R a + c..m *\<^sub>R b + c}" unfolding True by(auto simp add: vector_less_eq_def)
-  ultimately show ?thesis by auto
-next
-  case False
-  { fix y assume "a \<le> y" "y \<le> b" "m > 0"
-    hence "m *\<^sub>R a + c \<le> m *\<^sub>R y + c"  "m *\<^sub>R y + c \<le> m *\<^sub>R b + c"
-      unfolding vector_less_eq_def by(auto simp add: vector_smult_component vector_add_component)
-  } moreover
-  { fix y assume "a \<le> y" "y \<le> b" "m < 0"
-    hence "m *\<^sub>R b + c \<le> m *\<^sub>R y + c"  "m *\<^sub>R y + c \<le> m *\<^sub>R a + c"
-      unfolding vector_less_eq_def by(auto simp add: vector_smult_component vector_add_component mult_left_mono_neg elim!:ballE)
-  } moreover
-  { fix y assume "m > 0"  "m *\<^sub>R a + c \<le> y"  "y \<le> m *\<^sub>R b + c"
-    hence "y \<in> (\<lambda>x. m *\<^sub>R x + c) ` {a..b}"
-      unfolding image_iff Bex_def mem_interval vector_less_eq_def
-      apply(auto simp add: vector_smult_component vector_add_component vector_minus_component vector_smult_assoc pth_3[symmetric]
-        intro!: exI[where x="(1 / m) *\<^sub>R (y - c)"])
-      by(auto simp add: pos_le_divide_eq pos_divide_le_eq real_mult_commute diff_le_iff)
-  } moreover
-  { fix y assume "m *\<^sub>R b + c \<le> y" "y \<le> m *\<^sub>R a + c" "m < 0"
-    hence "y \<in> (\<lambda>x. m *\<^sub>R x + c) ` {a..b}"
-      unfolding image_iff Bex_def mem_interval vector_less_eq_def
-      apply(auto simp add: vector_smult_component vector_add_component vector_minus_component vector_smult_assoc pth_3[symmetric]
-        intro!: exI[where x="(1 / m) *\<^sub>R (y - c)"])
-      by(auto simp add: neg_le_divide_eq neg_divide_le_eq real_mult_commute diff_le_iff)
-  }
-  ultimately show ?thesis using False by auto
-qed
-
-lemma image_smult_interval:"(\<lambda>x. m *\<^sub>R (x::real^'n::finite)) ` {a..b} =
-  (if {a..b} = {} then {} else if 0 \<le> m then {m *\<^sub>R a..m *\<^sub>R b} else {m *\<^sub>R b..m *\<^sub>R 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 \<noteq> {}" and c:"0 \<le> c" "c < 1" and f:"(f ` s) \<subseteq> s" and
-          lipschitz:"\<forall>x\<in>s. \<forall>y\<in>s. dist (f x) (f y) \<le> c * dist x y"
-  shows "\<exists>! x\<in>s. (f x = x)"
-proof-
-  have "1 - c > 0" using c by auto
-
-  from s(2) obtain z0 where "z0 \<in> s" by auto
-  def z \<equiv> "\<lambda>n. (f ^^ n) z0"
-  { fix n::nat
-    have "z n \<in> s" unfolding z_def
-    proof(induct n) case 0 thus ?case using `z0 \<in>s` by auto
-    next case Suc thus ?case using f by auto qed }
-  note z_in_s = this
-
-  def d \<equiv> "dist (z 0) (z 1)"
-
-  have fzn:"\<And>n. f (z n) = z (Suc n)" unfolding z_def by auto
-  { fix n::nat
-    have "dist (z n) (z (Suc n)) \<le> (c ^ n) * d"
-    proof(induct n)
-      case 0 thus ?case unfolding d_def by auto
-    next
-      case (Suc m)
-      hence "c * dist (z m) (z (Suc m)) \<le> c ^ Suc m * d"
-        using `0 \<le> c` using mult_mono1_class.mult_mono1[of "dist (z m) (z (Suc m))" "c ^ m * d" c] by auto
-      thus ?case using lipschitz[THEN bspec[where x="z m"], OF z_in_s, THEN bspec[where x="z (Suc m)"], OF z_in_s]
-        unfolding fzn and mult_le_cancel_left by auto
-    qed
-  } note cf_z = this
-
-  { fix n m::nat
-    have "(1 - c) * dist (z m) (z (m+n)) \<le> (c ^ m) * d * (1 - c ^ n)"
-    proof(induct n)
-      case 0 show ?case by auto
-    next
-      case (Suc k)
-      have "(1 - c) * dist (z m) (z (m + Suc k)) \<le> (1 - c) * (dist (z m) (z (m + k)) + dist (z (m + k)) (z (Suc (m + k))))"
-        using dist_triangle and c by(auto simp add: dist_triangle)
-      also have "\<dots> \<le> (1 - c) * (dist (z m) (z (m + k)) + c ^ (m + k) * d)"
-        using cf_z[of "m + k"] and c by auto
-      also have "\<dots> \<le> c ^ m * d * (1 - c ^ k) + (1 - c) * c ^ (m + k) * d"
-        using Suc by (auto simp add: ring_simps)
-      also have "\<dots> = (c ^ m) * (d * (1 - c ^ k) + (1 - c) * c ^ k * d)"
-        unfolding power_add by (auto simp add: ring_simps)
-      also have "\<dots> \<le> (c ^ m) * d * (1 - c ^ Suc k)"
-        using c by (auto simp add: ring_simps)
-      finally show ?case by auto
-    qed
-  } note cf_z2 = this
-  { fix e::real assume "e>0"
-    hence "\<exists>N. \<forall>m n. N \<le> m \<and> N \<le> n \<longrightarrow> dist (z m) (z n) < e"
-    proof(cases "d = 0")
-      case True
-      hence "\<And>n. z n = z0" using cf_z2[of 0] and c unfolding z_def by (auto simp add: pos_prod_le[OF `1 - c > 0`])
-      thus ?thesis using `e>0` by auto
-    next
-      case False hence "d>0" unfolding d_def using zero_le_dist[of "z 0" "z 1"]
-        by (metis False d_def real_less_def)
-      hence "0 < e * (1 - c) / d" using `e>0` and `1-c>0`
-        using divide_pos_pos[of "e * (1 - c)" d] and mult_pos_pos[of e "1 - c"] by auto
-      then obtain N where N:"c ^ N < e * (1 - c) / d" using real_arch_pow_inv[of "e * (1 - c) / d" c] and c by auto
-      { fix m n::nat assume "m>n" and as:"m\<ge>N" "n\<ge>N"
-        have *:"c ^ n \<le> c ^ N" using `n\<ge>N` and c using power_decreasing[OF `n\<ge>N`, of c] by auto
-        have "1 - c ^ (m - n) > 0" using c and power_strict_mono[of c 1 "m - n"] using `m>n` by auto
-        hence **:"d * (1 - c ^ (m - n)) / (1 - c) > 0"
-          using real_mult_order[OF `d>0`, of "1 - c ^ (m - n)"]
-          using divide_pos_pos[of "d * (1 - c ^ (m - n))" "1 - c"]
-          using `0 < 1 - c` by auto
-
-        have "dist (z m) (z n) \<le> c ^ n * d * (1 - c ^ (m - n)) / (1 - c)"
-          using cf_z2[of n "m - n"] and `m>n` unfolding pos_le_divide_eq[OF `1-c>0`]
-          by (auto simp add: real_mult_commute dist_commute)
-        also have "\<dots> \<le> c ^ N * d * (1 - c ^ (m - n)) / (1 - c)"
-          using mult_right_mono[OF * order_less_imp_le[OF **]]
-          unfolding real_mult_assoc by auto
-        also have "\<dots> < (e * (1 - c) / d) * d * (1 - c ^ (m - n)) / (1 - c)"
-          using mult_strict_right_mono[OF N **] unfolding real_mult_assoc by auto
-        also have "\<dots> = e * (1 - c ^ (m - n))" using c and `d>0` and `1 - c > 0` by auto
-        also have "\<dots> \<le> e" using c and `1 - c ^ (m - n) > 0` and `e>0` using mult_right_le_one_le[of e "1 - c ^ (m - n)"] by auto
-        finally have  "dist (z m) (z n) < e" by auto
-      } note * = this
-      { fix m n::nat assume as:"N\<le>m" "N\<le>n"
-        hence "dist (z n) (z m) < e"
-        proof(cases "n = m")
-          case True thus ?thesis using `e>0` by auto
-        next
-          case False thus ?thesis using as and *[of n m] *[of m n] unfolding nat_neq_iff by (auto simp add: dist_commute)
-        qed }
-      thus ?thesis by auto
-    qed
-  }
-  hence "Cauchy z" unfolding cauchy_def by auto
-  then obtain x where "x\<in>s" and x:"(z ---> x) sequentially" using s(1)[unfolded compact_def complete_def, THEN spec[where x=z]] and z_in_s by auto
-
-  def e \<equiv> "dist (f x) x"
-  have "e = 0" proof(rule ccontr)
-    assume "e \<noteq> 0" hence "e>0" unfolding e_def using zero_le_dist[of "f x" x]
-      by (metis dist_eq_0_iff dist_nz e_def)
-    then obtain N where N:"\<forall>n\<ge>N. dist (z n) x < e / 2"
-      using x[unfolded Lim_sequentially, THEN spec[where x="e/2"]] by auto
-    hence N':"dist (z N) x < e / 2" by auto
-
-    have *:"c * dist (z N) x \<le> dist (z N) x" unfolding mult_le_cancel_right2
-      using zero_le_dist[of "z N" x] and c
-      by (metis dist_eq_0_iff dist_nz order_less_asym real_less_def)
-    have "dist (f (z N)) (f x) \<le> c * dist (z N) x" using lipschitz[THEN bspec[where x="z N"], THEN bspec[where x=x]]
-      using z_in_s[of N] `x\<in>s` using c by auto
-    also have "\<dots> < e / 2" using N' and c using * by auto
-    finally show False unfolding fzn
-      using N[THEN spec[where x="Suc N"]] and dist_triangle_half_r[of "z (Suc N)" "f x" e x]
-      unfolding e_def by auto
-  qed
-  hence "f x = x" unfolding e_def by auto
-  moreover
-  { fix y assume "f y = y" "y\<in>s"
-    hence "dist x y \<le> c * dist x y" using lipschitz[THEN bspec[where x=x], THEN bspec[where x=y]]
-      using `x\<in>s` and `f x = x` by auto
-    hence "dist x y = 0" unfolding mult_le_cancel_right1
-      using c and zero_le_dist[of x y] by auto
-    hence "y = x" by auto
-  }
-  ultimately show ?thesis unfolding Bex1_def using `x\<in>s` by blast+
-qed
-
-subsection{* Edelstein fixed point theorem.                                            *}
-
-lemma edelstein_fix:
-  fixes s :: "'a::real_normed_vector set"
-  assumes s:"compact s" "s \<noteq> {}" and gs:"(g ` s) \<subseteq> s"
-      and dist:"\<forall>x\<in>s. \<forall>y\<in>s. x \<noteq> y \<longrightarrow> dist (g x) (g y) < dist x y"
-  shows "\<exists>! x\<in>s. g x = x"
-proof(cases "\<exists>x\<in>s. g x \<noteq> x")
-  obtain x where "x\<in>s" using s(2) by auto
-  case False hence g:"\<forall>x\<in>s. g x = x" by auto
-  { fix y assume "y\<in>s"
-    hence "x = y" using `x\<in>s` and dist[THEN bspec[where x=x], THEN bspec[where x=y]]
-      unfolding g[THEN bspec[where x=x], OF `x\<in>s`]
-      unfolding g[THEN bspec[where x=y], OF `y\<in>s`] by auto  }
-  thus ?thesis unfolding Bex1_def using `x\<in>s` and g by blast+
-next
-  case True
-  then obtain x where [simp]:"x\<in>s" and "g x \<noteq> x" by auto
-  { fix x y assume "x \<in> s" "y \<in> s"
-    hence "dist (g x) (g y) \<le> dist x y"
-      using dist[THEN bspec[where x=x], THEN bspec[where x=y]] by auto } note dist' = this
-  def y \<equiv> "g x"
-  have [simp]:"y\<in>s" unfolding y_def using gs[unfolded image_subset_iff] and `x\<in>s` by blast
-  def f \<equiv> "\<lambda>n. g ^^ n"
-  have [simp]:"\<And>n z. g (f n z) = f (Suc n) z" unfolding f_def by auto
-  have [simp]:"\<And>z. f 0 z = z" unfolding f_def by auto
-  { fix n::nat and z assume "z\<in>s"
-    have "f n z \<in> s" unfolding f_def
-    proof(induct n)
-      case 0 thus ?case using `z\<in>s` by simp
-    next
-      case (Suc n) thus ?case using gs[unfolded image_subset_iff] by auto
-    qed } note fs = this
-  { fix m n ::nat assume "m\<le>n"
-    fix w z assume "w\<in>s" "z\<in>s"
-    have "dist (f n w) (f n z) \<le> dist (f m w) (f m z)" using `m\<le>n`
-    proof(induct n)
-      case 0 thus ?case by auto
-    next
-      case (Suc n)
-      thus ?case proof(cases "m\<le>n")
-        case True thus ?thesis using Suc(1)
-          using dist'[OF fs fs, OF `w\<in>s` `z\<in>s`, of n n] by auto
-      next
-        case False hence mn:"m = Suc n" using Suc(2) by simp
-        show ?thesis unfolding mn  by auto
-      qed
-    qed } note distf = this
-
-  def h \<equiv> "\<lambda>n. (f n x, f n y)"
-  let ?s2 = "s \<times> s"
-  obtain l r where "l\<in>?s2" and r:"subseq r" and lr:"((h \<circ> r) ---> l) sequentially"
-    using compact_Times [OF s(1) s(1), unfolded compact_def, THEN spec[where x=h]] unfolding  h_def
-    using fs[OF `x\<in>s`] and fs[OF `y\<in>s`] by blast
-  def a \<equiv> "fst l" def b \<equiv> "snd l"
-  have lab:"l = (a, b)" unfolding a_def b_def by simp
-  have [simp]:"a\<in>s" "b\<in>s" unfolding a_def b_def using `l\<in>?s2` by auto
-
-  have lima:"((fst \<circ> (h \<circ> r)) ---> a) sequentially"
-   and limb:"((snd \<circ> (h \<circ> r)) ---> b) sequentially"
-    using lr
-    unfolding o_def a_def b_def by (simp_all add: tendsto_intros)
-
-  { fix n::nat
-    have *:"\<And>fx fy (x::'a) y. dist fx fy \<le> dist x y \<Longrightarrow> \<not> (dist (fx - fy) (a - b) < dist a b - dist x y)" unfolding dist_norm by norm
-    { fix x y :: 'a
-      have "dist (-x) (-y) = dist x y" unfolding dist_norm
-        using norm_minus_cancel[of "x - y"] by (auto simp add: uminus_add_conv_diff) } note ** = this
-
-    { assume as:"dist a b > dist (f n x) (f n y)"
-      then obtain Na Nb where "\<forall>m\<ge>Na. dist (f (r m) x) a < (dist a b - dist (f n x) (f n y)) / 2"
-        and "\<forall>m\<ge>Nb. dist (f (r m) y) b < (dist a b - dist (f n x) (f n y)) / 2"
-        using lima limb unfolding h_def Lim_sequentially by (fastsimp simp del: less_divide_eq_number_of1)
-      hence "dist (f (r (Na + Nb + n)) x - f (r (Na + Nb + n)) y) (a - b) < dist a b - dist (f n x) (f n y)"
-        apply(erule_tac x="Na+Nb+n" in allE)
-        apply(erule_tac x="Na+Nb+n" in allE) apply simp
-        using dist_triangle_add_half[of a "f (r (Na + Nb + n)) x" "dist a b - dist (f n x) (f n y)"
-          "-b"  "- f (r (Na + Nb + n)) y"]
-        unfolding ** unfolding group_simps(12) by (auto simp add: dist_commute)
-      moreover
-      have "dist (f (r (Na + Nb + n)) x - f (r (Na + Nb + n)) y) (a - b) \<ge> dist a b - dist (f n x) (f n y)"
-        using distf[of n "r (Na+Nb+n)", OF _ `x\<in>s` `y\<in>s`]
-        using subseq_bigger[OF r, of "Na+Nb+n"]
-        using *[of "f (r (Na + Nb + n)) x" "f (r (Na + Nb + n)) y" "f n x" "f n y"] by auto
-      ultimately have False by simp
-    }
-    hence "dist a b \<le> dist (f n x) (f n y)" by(rule ccontr)auto }
-  note ab_fn = this
-
-  have [simp]:"a = b" proof(rule ccontr)
-    def e \<equiv> "dist a b - dist (g a) (g b)"
-    assume "a\<noteq>b" hence "e > 0" unfolding e_def using dist by fastsimp
-    hence "\<exists>n. dist (f n x) a < e/2 \<and> dist (f n y) b < e/2"
-      using lima limb unfolding Lim_sequentially
-      apply (auto elim!: allE[where x="e/2"]) apply(rule_tac x="r (max N Na)" in exI) unfolding h_def by fastsimp
-    then obtain n where n:"dist (f n x) a < e/2 \<and> dist (f n y) b < e/2" by auto
-    have "dist (f (Suc n) x) (g a) \<le> dist (f n x) a"
-      using dist[THEN bspec[where x="f n x"], THEN bspec[where x="a"]] and fs by auto
-    moreover have "dist (f (Suc n) y) (g b) \<le> dist (f n y) b"
-      using dist[THEN bspec[where x="f n y"], THEN bspec[where x="b"]] and fs by auto
-    ultimately have "dist (f (Suc n) x) (g a) + dist (f (Suc n) y) (g b) < e" using n by auto
-    thus False unfolding e_def using ab_fn[of "Suc n"] by norm
-  qed
-
-  have [simp]:"\<And>n. f (Suc n) x = f n y" unfolding f_def y_def by(induct_tac n)auto
-  { fix x y assume "x\<in>s" "y\<in>s" moreover
-    fix e::real assume "e>0" ultimately
-    have "dist y x < e \<longrightarrow> dist (g y) (g x) < e" using dist by fastsimp }
-  hence "continuous_on s g" unfolding continuous_on_def by auto
-
-  hence "((snd \<circ> h \<circ> r) ---> g a) sequentially" unfolding continuous_on_sequentially
-    apply (rule allE[where x="\<lambda>n. (fst \<circ> h \<circ> r) n"]) apply (erule ballE[where x=a])
-    using lima unfolding h_def o_def using fs[OF `x\<in>s`] by (auto simp add: y_def)
-  hence "g a = a" using Lim_unique[OF trivial_limit_sequentially limb, of "g a"]
-    unfolding `a=b` and o_assoc by auto
-  moreover
-  { fix x assume "x\<in>s" "g x = x" "x\<noteq>a"
-    hence "False" using dist[THEN bspec[where x=a], THEN bspec[where x=x]]
-      using `g a = a` and `a\<in>s` by auto  }
-  ultimately show "\<exists>!x\<in>s. g x = x" unfolding Bex1_def using `a\<in>s` by blast
-qed
-
-end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy	Fri Oct 23 13:23:18 2009 +0200
@@ -0,0 +1,3371 @@
+(*  Title:      HOL/Library/Convex_Euclidean_Space.thy
+    Author:     Robert Himmelmann, TU Muenchen
+*)
+
+header {* Convex sets, functions and related things. *}
+
+theory Convex_Euclidean_Space
+imports Topology_Euclidean_Space
+begin
+
+
+(* ------------------------------------------------------------------------- *)
+(* To be moved elsewhere                                                     *)
+(* ------------------------------------------------------------------------- *)
+
+declare vector_add_ldistrib[simp] vector_ssub_ldistrib[simp] vector_smult_assoc[simp] vector_smult_rneg[simp]
+declare vector_sadd_rdistrib[simp] vector_sub_rdistrib[simp]
+declare dot_ladd[simp] dot_radd[simp] dot_lsub[simp] dot_rsub[simp]
+declare dot_lmult[simp] dot_rmult[simp] dot_lneg[simp] dot_rneg[simp]
+declare UNIV_1[simp]
+
+term "(x::real^'n \<Rightarrow> real) 0"
+
+lemma dim1in[intro]:"Suc 0 \<in> {1::nat .. CARD(1)}" by auto
+
+lemmas vector_component_simps = vector_minus_component vector_smult_component vector_add_component vector_less_eq_def Cart_lambda_beta dest_vec1_def basis_component vector_uminus_component
+
+lemmas continuous_intros = continuous_add continuous_vmul continuous_cmul continuous_const continuous_sub continuous_at_id continuous_within_id
+
+lemmas continuous_on_intros = continuous_on_add continuous_on_const continuous_on_id continuous_on_compose continuous_on_cmul continuous_on_neg continuous_on_sub
+  uniformly_continuous_on_add uniformly_continuous_on_const uniformly_continuous_on_id uniformly_continuous_on_compose uniformly_continuous_on_cmul uniformly_continuous_on_neg uniformly_continuous_on_sub
+
+lemma dest_vec1_simps[simp]: fixes a::"real^1"
+  shows "a$1 = 0 \<longleftrightarrow> a = 0" (*"a \<le> 1 \<longleftrightarrow> dest_vec1 a \<le> 1" "0 \<le> a \<longleftrightarrow> 0 \<le> dest_vec1 a"*)
+  "a \<le> b \<longleftrightarrow> dest_vec1 a \<le> dest_vec1 b" "dest_vec1 (1::real^1) = 1"
+  by(auto simp add:vector_component_simps all_1 Cart_eq)
+
+lemma nequals0I:"x\<in>A \<Longrightarrow> A \<noteq> {}" by auto
+
+lemma norm_not_0:"(x::real^'n::finite)\<noteq>0 \<Longrightarrow> norm x \<noteq> 0" by auto
+
+lemma setsum_delta_notmem: assumes "x\<notin>s"
+  shows "setsum (\<lambda>y. if (y = x) then P x else Q y) s = setsum Q s"
+        "setsum (\<lambda>y. if (x = y) then P x else Q y) s = setsum Q s"
+        "setsum (\<lambda>y. if (y = x) then P y else Q y) s = setsum Q s"
+        "setsum (\<lambda>y. if (x = y) then P y else Q y) s = setsum Q s"
+  apply(rule_tac [!] setsum_cong2) using assms by auto
+
+lemma setsum_delta'':
+  fixes s::"'a::real_vector set" assumes "finite s"
+  shows "(\<Sum>x\<in>s. (if y = x then f x else 0) *\<^sub>R x) = (if y\<in>s then (f y) *\<^sub>R y else 0)"
+proof-
+  have *:"\<And>x y. (if y = x then f x else (0::real)) *\<^sub>R x = (if x=y then (f x) *\<^sub>R x else 0)" by auto
+  show ?thesis unfolding * using setsum_delta[OF assms, of y "\<lambda>x. f x *\<^sub>R x"] by auto
+qed
+
+lemma not_disjointI:"x\<in>A \<Longrightarrow> x\<in>B \<Longrightarrow> A \<inter> B \<noteq> {}" by blast
+
+lemma if_smult:"(if P then x else (y::real)) *\<^sub>R v = (if P then x *\<^sub>R v else y *\<^sub>R v)" by auto
+
+lemma mem_interval_1: fixes x :: "real^1" shows
+ "(x \<in> {a .. b} \<longleftrightarrow> dest_vec1 a \<le> dest_vec1 x \<and> dest_vec1 x \<le> dest_vec1 b)"
+ "(x \<in> {a<..<b} \<longleftrightarrow> dest_vec1 a < dest_vec1 x \<and> dest_vec1 x < dest_vec1 b)"
+by(simp_all add: Cart_eq vector_less_def vector_less_eq_def dest_vec1_def all_1)
+
+lemma image_smult_interval:"(\<lambda>x. m *\<^sub>R (x::real^'n::finite)) ` {a..b} =
+  (if {a..b} = {} then {} else if 0 \<le> m then {m *\<^sub>R a..m *\<^sub>R b} else {m *\<^sub>R b..m *\<^sub>R a})"
+  using image_affinity_interval[of m 0 a b] by auto
+
+lemma dest_vec1_inverval:
+  "dest_vec1 ` {a .. b} = {dest_vec1 a .. dest_vec1 b}"
+  "dest_vec1 ` {a<.. b} = {dest_vec1 a<.. dest_vec1 b}"
+  "dest_vec1 ` {a ..<b} = {dest_vec1 a ..<dest_vec1 b}"
+  "dest_vec1 ` {a<..<b} = {dest_vec1 a<..<dest_vec1 b}"
+  apply(rule_tac [!] equalityI)
+  unfolding subset_eq Ball_def Bex_def mem_interval_1 image_iff
+  apply(rule_tac [!] allI)apply(rule_tac [!] impI)
+  apply(rule_tac[2] x="vec1 x" in exI)apply(rule_tac[4] x="vec1 x" in exI)
+  apply(rule_tac[6] x="vec1 x" in exI)apply(rule_tac[8] x="vec1 x" in exI)
+  by (auto simp add: vector_less_def vector_less_eq_def all_1 dest_vec1_def
+    vec1_dest_vec1[unfolded dest_vec1_def One_nat_def])
+
+lemma dest_vec1_setsum: assumes "finite S"
+  shows " dest_vec1 (setsum f S) = setsum (\<lambda>x. dest_vec1 (f x)) S"
+  using dest_vec1_sum[OF assms] by auto
+
+lemma dist_triangle_eq:
+  fixes x y z :: "real ^ _"
+  shows "dist x z = dist x y + dist y z \<longleftrightarrow> norm (x - y) *\<^sub>R (y - z) = norm (y - z) *\<^sub>R (x - y)"
+proof- have *:"x - y + (y - z) = x - z" by auto
+  show ?thesis unfolding dist_norm norm_triangle_eq[of "x - y" "y - z", unfolded smult_conv_scaleR *]
+    by(auto simp add:norm_minus_commute) qed
+
+lemma norm_eqI:"x = y \<Longrightarrow> norm x = norm y" by auto 
+lemma norm_minus_eqI:"(x::real^'n::finite) = - y \<Longrightarrow> norm x = norm y" by auto
+
+lemma Min_grI: assumes "finite A" "A \<noteq> {}" "\<forall>a\<in>A. x < a" shows "x < Min A"
+  unfolding Min_gr_iff[OF assms(1,2)] using assms(3) by auto
+
+lemma dimindex_ge_1:"CARD(_::finite) \<ge> 1"
+  using one_le_card_finite by auto
+
+lemma real_dimindex_ge_1:"real (CARD('n::finite)) \<ge> 1" 
+  by(metis dimindex_ge_1 linorder_not_less real_eq_of_nat real_le_trans real_of_nat_1 real_of_nat_le_iff) 
+
+lemma real_dimindex_gt_0:"real (CARD('n::finite)) > 0" apply(rule less_le_trans[OF _ real_dimindex_ge_1]) by auto
+
+subsection {* Affine set and affine hull.*}
+
+definition
+  affine :: "'a::real_vector set \<Rightarrow> bool" where
+  "affine s \<longleftrightarrow> (\<forall>x\<in>s. \<forall>y\<in>s. \<forall>u v. u + v = 1 \<longrightarrow> u *\<^sub>R x + v *\<^sub>R y \<in> s)"
+
+lemma affine_alt: "affine s \<longleftrightarrow> (\<forall>x\<in>s. \<forall>y\<in>s. \<forall>u::real. (1 - u) *\<^sub>R x + u *\<^sub>R y \<in> s)"
+proof- have *:"\<And>u v ::real. u + v = 1 \<longleftrightarrow> v = 1 - u" by auto
+  { fix x y assume "x\<in>s" "y\<in>s"
+    hence "(\<forall>u v::real. u + v = 1 \<longrightarrow> u *\<^sub>R x + v *\<^sub>R y \<in> s) \<longleftrightarrow> (\<forall>u::real. (1 - u) *\<^sub>R x + u *\<^sub>R y \<in> s)" apply auto 
+      apply(erule_tac[!] x="1 - u" in allE) unfolding * by auto  }
+  thus ?thesis unfolding affine_def by auto qed
+
+lemma affine_empty[intro]: "affine {}"
+  unfolding affine_def by auto
+
+lemma affine_sing[intro]: "affine {x}"
+  unfolding affine_alt by (auto simp add: scaleR_left_distrib [symmetric])
+
+lemma affine_UNIV[intro]: "affine UNIV"
+  unfolding affine_def by auto
+
+lemma affine_Inter: "(\<forall>s\<in>f. affine s) \<Longrightarrow> affine (\<Inter> f)"
+  unfolding affine_def by auto 
+
+lemma affine_Int: "affine s \<Longrightarrow> affine t \<Longrightarrow> affine (s \<inter> t)"
+  unfolding affine_def by auto
+
+lemma affine_affine_hull: "affine(affine hull s)"
+  unfolding hull_def using affine_Inter[of "{t \<in> affine. s \<subseteq> t}"]
+  unfolding mem_def by auto
+
+lemma affine_hull_eq[simp]: "(affine hull s = s) \<longleftrightarrow> affine s"
+proof-
+  { fix f assume "f \<subseteq> affine"
+    hence "affine (\<Inter>f)" using affine_Inter[of f] unfolding subset_eq mem_def by auto  }
+  thus ?thesis using hull_eq[unfolded mem_def, of affine s] by auto
+qed
+
+lemma setsum_restrict_set'': assumes "finite A"
+  shows "setsum f {x \<in> A. P x} = (\<Sum>x\<in>A. if P x  then f x else 0)"
+  unfolding mem_def[of _ P, symmetric] unfolding setsum_restrict_set'[OF assms] ..
+
+subsection {* Some explicit formulations (from Lars Schewe). *}
+
+lemma affine: fixes V::"'a::real_vector set"
+  shows "affine V \<longleftrightarrow> (\<forall>s u. finite s \<and> s \<noteq> {} \<and> s \<subseteq> V \<and> setsum u s = 1 \<longrightarrow> (setsum (\<lambda>x. (u x) *\<^sub>R x)) s \<in> V)"
+unfolding affine_def apply rule apply(rule, rule, rule) apply(erule conjE)+ 
+defer apply(rule, rule, rule, rule, rule) proof-
+  fix x y u v assume as:"x \<in> V" "y \<in> V" "u + v = (1::real)"
+    "\<forall>s u. finite s \<and> s \<noteq> {} \<and> s \<subseteq> V \<and> setsum u s = 1 \<longrightarrow> (\<Sum>x\<in>s. u x *\<^sub>R x) \<in> V"
+  thus "u *\<^sub>R x + v *\<^sub>R y \<in> V" apply(cases "x=y")
+    using as(4)[THEN spec[where x="{x,y}"], THEN spec[where x="\<lambda>w. if w = x then u else v"]] and as(1-3) 
+    by(auto simp add: scaleR_left_distrib[THEN sym])
+next
+  fix s u assume as:"\<forall>x\<in>V. \<forall>y\<in>V. \<forall>u v. u + v = 1 \<longrightarrow> u *\<^sub>R x + v *\<^sub>R y \<in> V"
+    "finite s" "s \<noteq> {}" "s \<subseteq> V" "setsum u s = (1::real)"
+  def n \<equiv> "card s"
+  have "card s = 0 \<or> card s = 1 \<or> card s = 2 \<or> card s > 2" by auto
+  thus "(\<Sum>x\<in>s. u x *\<^sub>R x) \<in> V" proof(auto simp only: disjE)
+    assume "card s = 2" hence "card s = Suc (Suc 0)" by auto
+    then obtain a b where "s = {a, b}" unfolding card_Suc_eq by auto
+    thus ?thesis using as(1)[THEN bspec[where x=a], THEN bspec[where x=b]] using as(4,5)
+      by(auto simp add: setsum_clauses(2))
+  next assume "card s > 2" thus ?thesis using as and n_def proof(induct n arbitrary: u s)
+      case (Suc n) fix s::"'a set" and u::"'a \<Rightarrow> real"
+      assume IA:"\<And>u s.  \<lbrakk>2 < card s; \<forall>x\<in>V. \<forall>y\<in>V. \<forall>u v. u + v = 1 \<longrightarrow> u *\<^sub>R x + v *\<^sub>R y \<in> V; finite s;
+               s \<noteq> {}; s \<subseteq> V; setsum u s = 1; n \<equiv> card s \<rbrakk> \<Longrightarrow> (\<Sum>x\<in>s. u x *\<^sub>R x) \<in> V" and
+        as:"Suc n \<equiv> card s" "2 < card s" "\<forall>x\<in>V. \<forall>y\<in>V. \<forall>u v. u + v = 1 \<longrightarrow> u *\<^sub>R x + v *\<^sub>R y \<in> V"
+           "finite s" "s \<noteq> {}" "s \<subseteq> V" "setsum u s = 1"
+      have "\<exists>x\<in>s. u x \<noteq> 1" proof(rule_tac ccontr)
+        assume " \<not> (\<exists>x\<in>s. u x \<noteq> 1)" hence "setsum u s = real_of_nat (card s)" unfolding card_eq_setsum by auto
+        thus False using as(7) and `card s > 2` by (metis Numeral1_eq1_nat less_0_number_of less_int_code(15)
+          less_nat_number_of not_less_iff_gr_or_eq of_nat_1 of_nat_eq_iff pos2 rel_simps(4)) qed
+      then obtain x where x:"x\<in>s" "u x \<noteq> 1" by auto
+
+      have c:"card (s - {x}) = card s - 1" apply(rule card_Diff_singleton) using `x\<in>s` as(4) by auto
+      have *:"s = insert x (s - {x})" "finite (s - {x})" using `x\<in>s` and as(4) by auto
+      have **:"setsum u (s - {x}) = 1 - u x"
+        using setsum_clauses(2)[OF *(2), of u x, unfolded *(1)[THEN sym] as(7)] by auto
+      have ***:"inverse (1 - u x) * setsum u (s - {x}) = 1" unfolding ** using `u x \<noteq> 1` by auto
+      have "(\<Sum>xa\<in>s - {x}. (inverse (1 - u x) * u xa) *\<^sub>R xa) \<in> V" proof(cases "card (s - {x}) > 2")
+        case True hence "s - {x} \<noteq> {}" "card (s - {x}) = n" unfolding c and as(1)[symmetric] proof(rule_tac ccontr) 
+          assume "\<not> s - {x} \<noteq> {}" hence "card (s - {x}) = 0" unfolding card_0_eq[OF *(2)] by simp 
+          thus False using True by auto qed auto
+        thus ?thesis apply(rule_tac IA[of "s - {x}" "\<lambda>y. (inverse (1 - u x) * u y)"])
+        unfolding setsum_right_distrib[THEN sym] using as and *** and True by auto
+      next case False hence "card (s - {x}) = Suc (Suc 0)" using as(2) and c by auto
+        then obtain a b where "(s - {x}) = {a, b}" "a\<noteq>b" unfolding card_Suc_eq by auto
+        thus ?thesis using as(3)[THEN bspec[where x=a], THEN bspec[where x=b]]
+          using *** *(2) and `s \<subseteq> V` unfolding setsum_right_distrib by(auto simp add: setsum_clauses(2)) qed
+      thus "(\<Sum>x\<in>s. u x *\<^sub>R x) \<in> V" unfolding scaleR_scaleR[THEN sym] and scaleR_right.setsum [symmetric]
+         apply(subst *) unfolding setsum_clauses(2)[OF *(2)]
+         using as(3)[THEN bspec[where x=x], THEN bspec[where x="(inverse (1 - u x)) *\<^sub>R (\<Sum>xa\<in>s - {x}. u xa *\<^sub>R xa)"], 
+         THEN spec[where x="u x"], THEN spec[where x="1 - u x"]] and rev_subsetD[OF `x\<in>s` `s\<subseteq>V`] and `u x \<noteq> 1` by auto
+    qed auto
+  next assume "card s = 1" then obtain a where "s={a}" by(auto simp add: card_Suc_eq)
+    thus ?thesis using as(4,5) by simp
+  qed(insert `s\<noteq>{}` `finite s`, auto)
+qed
+
+lemma affine_hull_explicit:
+  "affine hull p = {y. \<exists>s u. finite s \<and> s \<noteq> {} \<and> s \<subseteq> p \<and> setsum u s = 1 \<and> setsum (\<lambda>v. (u v) *\<^sub>R v) s = y}"
+  apply(rule hull_unique) apply(subst subset_eq) prefer 3 apply rule unfolding mem_Collect_eq and mem_def[of _ affine]
+  apply (erule exE)+ apply(erule conjE)+ prefer 2 apply rule proof-
+  fix x assume "x\<in>p" thus "\<exists>s u. finite s \<and> s \<noteq> {} \<and> s \<subseteq> p \<and> setsum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = x"
+    apply(rule_tac x="{x}" in exI, rule_tac x="\<lambda>x. 1" in exI) by auto
+next
+  fix t x s u assume as:"p \<subseteq> t" "affine t" "finite s" "s \<noteq> {}" "s \<subseteq> p" "setsum u s = 1" "(\<Sum>v\<in>s. u v *\<^sub>R v) = x" 
+  thus "x \<in> t" using as(2)[unfolded affine, THEN spec[where x=s], THEN spec[where x=u]] by auto
+next
+  show "affine {y. \<exists>s u. finite s \<and> s \<noteq> {} \<and> s \<subseteq> p \<and> setsum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = y}" unfolding affine_def
+    apply(rule,rule,rule,rule,rule) unfolding mem_Collect_eq proof-
+    fix u v ::real assume uv:"u + v = 1"
+    fix x assume "\<exists>s u. finite s \<and> s \<noteq> {} \<and> s \<subseteq> p \<and> setsum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = x"
+    then obtain sx ux where x:"finite sx" "sx \<noteq> {}" "sx \<subseteq> p" "setsum ux sx = 1" "(\<Sum>v\<in>sx. ux v *\<^sub>R v) = x" by auto
+    fix y assume "\<exists>s u. finite s \<and> s \<noteq> {} \<and> s \<subseteq> p \<and> setsum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = y"
+    then obtain sy uy where y:"finite sy" "sy \<noteq> {}" "sy \<subseteq> p" "setsum uy sy = 1" "(\<Sum>v\<in>sy. uy v *\<^sub>R v) = y" by auto
+    have xy:"finite (sx \<union> sy)" using x(1) y(1) by auto
+    have **:"(sx \<union> sy) \<inter> sx = sx" "(sx \<union> sy) \<inter> sy = sy" by auto
+    show "\<exists>s ua. finite s \<and> s \<noteq> {} \<and> s \<subseteq> p \<and> setsum ua s = 1 \<and> (\<Sum>v\<in>s. ua v *\<^sub>R v) = u *\<^sub>R x + v *\<^sub>R y"
+      apply(rule_tac x="sx \<union> sy" in exI)
+      apply(rule_tac x="\<lambda>a. (if a\<in>sx then u * ux a else 0) + (if a\<in>sy then v * uy a else 0)" in exI)
+      unfolding scaleR_left_distrib setsum_addf if_smult scaleR_zero_left  ** setsum_restrict_set[OF xy, THEN sym]
+      unfolding scaleR_scaleR[THEN sym] scaleR_right.setsum [symmetric] and setsum_right_distrib[THEN sym]
+      unfolding x y using x(1-3) y(1-3) uv by simp qed qed
+
+lemma affine_hull_finite:
+  assumes "finite s"
+  shows "affine hull s = {y. \<exists>u. setsum u s = 1 \<and> setsum (\<lambda>v. u v *\<^sub>R v) s = y}"
+  unfolding affine_hull_explicit and expand_set_eq and mem_Collect_eq apply (rule,rule)
+  apply(erule exE)+ apply(erule conjE)+ defer apply(erule exE) apply(erule conjE) proof-
+  fix x u assume "setsum u s = 1" "(\<Sum>v\<in>s. u v *\<^sub>R v) = x"
+  thus "\<exists>sa u. finite sa \<and> \<not> (\<forall>x. (x \<in> sa) = (x \<in> {})) \<and> sa \<subseteq> s \<and> setsum u sa = 1 \<and> (\<Sum>v\<in>sa. u v *\<^sub>R v) = x"
+    apply(rule_tac x=s in exI, rule_tac x=u in exI) using assms by auto
+next
+  fix x t u assume "t \<subseteq> s" hence *:"s \<inter> t = t" by auto
+  assume "finite t" "\<not> (\<forall>x. (x \<in> t) = (x \<in> {}))" "setsum u t = 1" "(\<Sum>v\<in>t. u v *\<^sub>R v) = x"
+  thus "\<exists>u. setsum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = x" apply(rule_tac x="\<lambda>x. if x\<in>t then u x else 0" in exI)
+    unfolding if_smult scaleR_zero_left and setsum_restrict_set[OF assms, THEN sym] and * by auto qed
+
+subsection {* Stepping theorems and hence small special cases. *}
+
+lemma affine_hull_empty[simp]: "affine hull {} = {}"
+  apply(rule hull_unique) unfolding mem_def by auto
+
+lemma affine_hull_finite_step:
+  fixes y :: "'a::real_vector"
+  shows "(\<exists>u. setsum u {} = w \<and> setsum (\<lambda>x. u x *\<^sub>R x) {} = y) \<longleftrightarrow> w = 0 \<and> y = 0" (is ?th1)
+  "finite s \<Longrightarrow> (\<exists>u. setsum u (insert a s) = w \<and> setsum (\<lambda>x. u x *\<^sub>R x) (insert a s) = y) \<longleftrightarrow>
+                (\<exists>v u. setsum u s = w - v \<and> setsum (\<lambda>x. u x *\<^sub>R x) s = y - v *\<^sub>R a)" (is "?as \<Longrightarrow> (?lhs = ?rhs)")
+proof-
+  show ?th1 by simp
+  assume ?as 
+  { assume ?lhs
+    then obtain u where u:"setsum u (insert a s) = w \<and> (\<Sum>x\<in>insert a s. u x *\<^sub>R x) = y" by auto
+    have ?rhs proof(cases "a\<in>s")
+      case True hence *:"insert a s = s" by auto
+      show ?thesis using u[unfolded *] apply(rule_tac x=0 in exI) by auto
+    next
+      case False thus ?thesis apply(rule_tac x="u a" in exI) using u and `?as` by auto 
+    qed  } moreover
+  { assume ?rhs
+    then obtain v u where vu:"setsum u s = w - v"  "(\<Sum>x\<in>s. u x *\<^sub>R x) = y - v *\<^sub>R a" by auto
+    have *:"\<And>x M. (if x = a then v else M) *\<^sub>R x = (if x = a then v *\<^sub>R x else M *\<^sub>R x)" by auto
+    have ?lhs proof(cases "a\<in>s")
+      case True thus ?thesis
+        apply(rule_tac x="\<lambda>x. (if x=a then v else 0) + u x" in exI)
+        unfolding setsum_clauses(2)[OF `?as`]  apply simp
+        unfolding scaleR_left_distrib and setsum_addf 
+        unfolding vu and * and scaleR_zero_left
+        by (auto simp add: setsum_delta[OF `?as`])
+    next
+      case False 
+      hence **:"\<And>x. x \<in> s \<Longrightarrow> u x = (if x = a then v else u x)"
+               "\<And>x. x \<in> s \<Longrightarrow> u x *\<^sub>R x = (if x = a then v *\<^sub>R x else u x *\<^sub>R x)" by auto
+      from False show ?thesis
+        apply(rule_tac x="\<lambda>x. if x=a then v else u x" in exI)
+        unfolding setsum_clauses(2)[OF `?as`] and * using vu
+        using setsum_cong2[of s "\<lambda>x. u x *\<^sub>R x" "\<lambda>x. if x = a then v *\<^sub>R x else u x *\<^sub>R x", OF **(2)]
+        using setsum_cong2[of s u "\<lambda>x. if x = a then v else u x", OF **(1)] by auto  
+    qed }
+  ultimately show "?lhs = ?rhs" by blast
+qed
+
+lemma affine_hull_2:
+  fixes a b :: "'a::real_vector"
+  shows "affine hull {a,b} = {u *\<^sub>R a + v *\<^sub>R b| u v. (u + v = 1)}" (is "?lhs = ?rhs")
+proof-
+  have *:"\<And>x y z. z = x - y \<longleftrightarrow> y + z = (x::real)" 
+         "\<And>x y z. z = x - y \<longleftrightarrow> y + z = (x::'a)" by auto
+  have "?lhs = {y. \<exists>u. setsum u {a, b} = 1 \<and> (\<Sum>v\<in>{a, b}. u v *\<^sub>R v) = y}"
+    using affine_hull_finite[of "{a,b}"] by auto
+  also have "\<dots> = {y. \<exists>v u. u b = 1 - v \<and> u b *\<^sub>R b = y - v *\<^sub>R a}"
+    by(simp add: affine_hull_finite_step(2)[of "{b}" a]) 
+  also have "\<dots> = ?rhs" unfolding * by auto
+  finally show ?thesis by auto
+qed
+
+lemma affine_hull_3:
+  fixes a b c :: "'a::real_vector"
+  shows "affine hull {a,b,c} = { u *\<^sub>R a + v *\<^sub>R b + w *\<^sub>R c| u v w. u + v + w = 1}" (is "?lhs = ?rhs")
+proof-
+  have *:"\<And>x y z. z = x - y \<longleftrightarrow> y + z = (x::real)" 
+         "\<And>x y z. z = x - y \<longleftrightarrow> y + z = (x::'a)" by auto
+  show ?thesis apply(simp add: affine_hull_finite affine_hull_finite_step)
+    unfolding * apply auto
+    apply(rule_tac x=v in exI) apply(rule_tac x=va in exI) apply auto
+    apply(rule_tac x=u in exI) by(auto intro!: exI)
+qed
+
+subsection {* Some relations between affine hull and subspaces. *}
+
+lemma affine_hull_insert_subset_span:
+  fixes a :: "real ^ _"
+  shows "affine hull (insert a s) \<subseteq> {a + v| v . v \<in> span {x - a | x . x \<in> s}}"
+  unfolding subset_eq Ball_def unfolding affine_hull_explicit span_explicit mem_Collect_eq smult_conv_scaleR
+  apply(rule,rule) apply(erule exE)+ apply(erule conjE)+ proof-
+  fix x t u assume as:"finite t" "t \<noteq> {}" "t \<subseteq> insert a s" "setsum u t = 1" "(\<Sum>v\<in>t. u v *\<^sub>R v) = x"
+  have "(\<lambda>x. x - a) ` (t - {a}) \<subseteq> {x - a |x. x \<in> s}" using as(3) by auto
+  thus "\<exists>v. x = a + v \<and> (\<exists>S u. finite S \<and> S \<subseteq> {x - a |x. x \<in> s} \<and> (\<Sum>v\<in>S. u v *\<^sub>R v) = v)"
+    apply(rule_tac x="x - a" in exI)
+    apply (rule conjI, simp)
+    apply(rule_tac x="(\<lambda>x. x - a) ` (t - {a})" in exI)
+    apply(rule_tac x="\<lambda>x. u (x + a)" in exI)
+    apply (rule conjI) using as(1) apply simp
+    apply (erule conjI)
+    using as(1)
+    apply (simp add: setsum_reindex[unfolded inj_on_def] scaleR_right_diff_distrib setsum_subtractf scaleR_left.setsum[THEN sym] setsum_diff1 scaleR_left_diff_distrib)
+    unfolding as by simp qed
+
+lemma affine_hull_insert_span:
+  fixes a :: "real ^ _"
+  assumes "a \<notin> s"
+  shows "affine hull (insert a s) =
+            {a + v | v . v \<in> span {x - a | x.  x \<in> s}}"
+  apply(rule, rule affine_hull_insert_subset_span) unfolding subset_eq Ball_def
+  unfolding affine_hull_explicit and mem_Collect_eq proof(rule,rule,erule exE,erule conjE)
+  fix y v assume "y = a + v" "v \<in> span {x - a |x. x \<in> s}"
+  then obtain t u where obt:"finite t" "t \<subseteq> {x - a |x. x \<in> s}" "a + (\<Sum>v\<in>t. u v *\<^sub>R v) = y" unfolding span_explicit smult_conv_scaleR by auto
+  def f \<equiv> "(\<lambda>x. x + a) ` t"
+  have f:"finite f" "f \<subseteq> s" "(\<Sum>v\<in>f. u (v - a) *\<^sub>R (v - a)) = y - a" unfolding f_def using obt 
+    by(auto simp add: setsum_reindex[unfolded inj_on_def])
+  have *:"f \<inter> {a} = {}" "f \<inter> - {a} = f" using f(2) assms by auto
+  show "\<exists>sa u. finite sa \<and> sa \<noteq> {} \<and> sa \<subseteq> insert a s \<and> setsum u sa = 1 \<and> (\<Sum>v\<in>sa. u v *\<^sub>R v) = y"
+    apply(rule_tac x="insert a f" in exI)
+    apply(rule_tac x="\<lambda>x. if x=a then 1 - setsum (\<lambda>x. u (x - a)) f else u (x - a)" in exI)
+    using assms and f unfolding setsum_clauses(2)[OF f(1)] and if_smult
+    unfolding setsum_cases[OF f(1), of "{a}", unfolded singleton_iff] and *
+    by (auto simp add: setsum_subtractf scaleR_left.setsum algebra_simps) qed
+
+lemma affine_hull_span:
+  fixes a :: "real ^ _"
+  assumes "a \<in> s"
+  shows "affine hull s = {a + v | v. v \<in> span {x - a | x. x \<in> s - {a}}}"
+  using affine_hull_insert_span[of a "s - {a}", unfolded insert_Diff[OF assms]] by auto
+
+subsection {* Convexity. *}
+
+definition
+  convex :: "'a::real_vector set \<Rightarrow> bool" where
+  "convex s \<longleftrightarrow> (\<forall>x\<in>s. \<forall>y\<in>s. \<forall>u\<ge>0. \<forall>v\<ge>0. u + v = 1 \<longrightarrow> u *\<^sub>R x + v *\<^sub>R y \<in> s)"
+
+lemma convex_alt: "convex s \<longleftrightarrow> (\<forall>x\<in>s. \<forall>y\<in>s. \<forall>u. 0 \<le> u \<and> u \<le> 1 \<longrightarrow> ((1 - u) *\<^sub>R x + u *\<^sub>R y) \<in> s)"
+proof- have *:"\<And>u v::real. u + v = 1 \<longleftrightarrow> u = 1 - v" by auto
+  show ?thesis unfolding convex_def apply auto
+    apply(erule_tac x=x in ballE) apply(erule_tac x=y in ballE) apply(erule_tac x="1 - u" in allE)
+    by (auto simp add: *) qed
+
+lemma mem_convex:
+  assumes "convex s" "a \<in> s" "b \<in> s" "0 \<le> u" "u \<le> 1"
+  shows "((1 - u) *\<^sub>R a + u *\<^sub>R b) \<in> s"
+  using assms unfolding convex_alt by auto
+
+lemma convex_empty[intro]: "convex {}"
+  unfolding convex_def by simp
+
+lemma convex_singleton[intro]: "convex {a}"
+  unfolding convex_def by (auto simp add:scaleR_left_distrib[THEN sym])
+
+lemma convex_UNIV[intro]: "convex UNIV"
+  unfolding convex_def by auto
+
+lemma convex_Inter: "(\<forall>s\<in>f. convex s) ==> convex(\<Inter> f)"
+  unfolding convex_def by auto
+
+lemma convex_Int: "convex s \<Longrightarrow> convex t \<Longrightarrow> convex (s \<inter> t)"
+  unfolding convex_def by auto
+
+lemma convex_halfspace_le: "convex {x. inner a x \<le> b}"
+  unfolding convex_def apply auto
+  unfolding inner_add inner_scaleR
+  by (metis real_convex_bound_le)
+
+lemma convex_halfspace_ge: "convex {x. inner a x \<ge> b}"
+proof- have *:"{x. inner a x \<ge> b} = {x. inner (-a) x \<le> -b}" by auto
+  show ?thesis apply(unfold *) using convex_halfspace_le[of "-a" "-b"] by auto qed
+
+lemma convex_hyperplane: "convex {x. inner a x = b}"
+proof-
+  have *:"{x. inner a x = b} = {x. inner a x \<le> b} \<inter> {x. inner a x \<ge> b}" by auto
+  show ?thesis unfolding * apply(rule convex_Int)
+    using convex_halfspace_le convex_halfspace_ge by auto
+qed
+
+lemma convex_halfspace_lt: "convex {x. inner a x < b}"
+  unfolding convex_def
+  by(auto simp add: real_convex_bound_lt inner_add)
+
+lemma convex_halfspace_gt: "convex {x. inner a x > b}"
+   using convex_halfspace_lt[of "-a" "-b"] by auto
+
+lemma convex_positive_orthant: "convex {x::real^'n::finite. (\<forall>i. 0 \<le> x$i)}"
+  unfolding convex_def apply auto apply(erule_tac x=i in allE)+
+  apply(rule add_nonneg_nonneg) by(auto simp add: mult_nonneg_nonneg)
+
+subsection {* Explicit expressions for convexity in terms of arbitrary sums. *}
+
+lemma convex: "convex s \<longleftrightarrow>
+  (\<forall>(k::nat) u x. (\<forall>i. 1\<le>i \<and> i\<le>k \<longrightarrow> 0 \<le> u i \<and> x i \<in>s) \<and> (setsum u {1..k} = 1)
+           \<longrightarrow> setsum (\<lambda>i. u i *\<^sub>R x i) {1..k} \<in> s)"
+  unfolding convex_def apply rule apply(rule allI)+ defer apply(rule ballI)+ apply(rule allI)+ proof(rule,rule,rule,rule)
+  fix x y u v assume as:"\<forall>(k::nat) u x. (\<forall>i. 1 \<le> i \<and> i \<le> k \<longrightarrow> 0 \<le> u i \<and> x i \<in> s) \<and> setsum u {1..k} = 1 \<longrightarrow> (\<Sum>i = 1..k. u i *\<^sub>R x i) \<in> s"
+    "x \<in> s" "y \<in> s" "0 \<le> u" "0 \<le> v" "u + v = (1::real)"
+  show "u *\<^sub>R x + v *\<^sub>R y \<in> s" using as(1)[THEN spec[where x=2], THEN spec[where x="\<lambda>n. if n=1 then u else v"], THEN spec[where x="\<lambda>n. if n=1 then x else y"]] and as(2-)
+    by (auto simp add: setsum_head_Suc) 
+next
+  fix k u x assume as:"\<forall>x\<in>s. \<forall>y\<in>s. \<forall>u\<ge>0. \<forall>v\<ge>0. u + v = 1 \<longrightarrow> u *\<^sub>R x + v *\<^sub>R y \<in> s" 
+  show "(\<forall>i::nat. 1 \<le> i \<and> i \<le> k \<longrightarrow> 0 \<le> u i \<and> x i \<in> s) \<and> setsum u {1..k} = 1 \<longrightarrow> (\<Sum>i = 1..k. u i *\<^sub>R x i) \<in> s" apply(rule,erule conjE) proof(induct k arbitrary: u)
+  case (Suc k) show ?case proof(cases "u (Suc k) = 1")
+    case True hence "(\<Sum>i = Suc 0..k. u i *\<^sub>R x i) = 0" apply(rule_tac setsum_0') apply(rule ccontr) unfolding ball_simps apply(erule bexE) proof-
+      fix i assume i:"i \<in> {Suc 0..k}" "u i *\<^sub>R x i \<noteq> 0"
+      hence ui:"u i \<noteq> 0" by auto
+      hence "setsum (\<lambda>k. if k=i then u i else 0) {1 .. k} \<le> setsum u {1 .. k}" apply(rule_tac setsum_mono) using Suc(2) by auto
+      hence "setsum u {1 .. k} \<ge> u i" using i(1) by(auto simp add: setsum_delta) 
+      hence "setsum u {1 .. k} > 0"  using ui apply(rule_tac less_le_trans[of _ "u i"]) using Suc(2)[THEN spec[where x=i]] and i(1) by auto
+      thus False using Suc(3) unfolding setsum_cl_ivl_Suc and True by simp qed
+    thus ?thesis unfolding setsum_cl_ivl_Suc using True and Suc(2) by auto
+  next
+    have *:"setsum u {1..k} = 1 - u (Suc k)" using Suc(3)[unfolded setsum_cl_ivl_Suc] by auto
+    have **:"u (Suc k) \<le> 1" apply(rule ccontr) unfolding not_le using Suc(3) using setsum_nonneg[of "{1..k}" u] using Suc(2) by auto
+    have ***:"\<And>i k. (u i / (1 - u (Suc k))) *\<^sub>R x i = (inverse (1 - u (Suc k))) *\<^sub>R (u i *\<^sub>R x i)" unfolding real_divide_def by (auto simp add: algebra_simps)
+    case False hence nn:"1 - u (Suc k) \<noteq> 0" by auto
+    have "(\<Sum>i = 1..k. (u i / (1 - u (Suc k))) *\<^sub>R x i) \<in> s" apply(rule Suc(1)) unfolding setsum_divide_distrib[THEN sym] and *
+      apply(rule_tac allI) apply(rule,rule) apply(rule divide_nonneg_pos) using nn Suc(2) ** by auto
+    hence "(1 - u (Suc k)) *\<^sub>R (\<Sum>i = 1..k. (u i / (1 - u (Suc k))) *\<^sub>R x i) + u (Suc k) *\<^sub>R x (Suc k) \<in> s"
+      apply(rule as[THEN bspec, THEN bspec, THEN spec, THEN mp, THEN spec, THEN mp, THEN mp]) using Suc(2)[THEN spec[where x="Suc k"]] and ** by auto
+    thus ?thesis unfolding setsum_cl_ivl_Suc and *** and scaleR_right.setsum [symmetric] using nn by auto qed qed auto qed
+
+
+lemma convex_explicit:
+  fixes s :: "'a::real_vector set"
+  shows "convex s \<longleftrightarrow>
+  (\<forall>t u. finite t \<and> t \<subseteq> s \<and> (\<forall>x\<in>t. 0 \<le> u x) \<and> setsum u t = 1 \<longrightarrow> setsum (\<lambda>x. u x *\<^sub>R x) t \<in> s)"
+  unfolding convex_def apply(rule,rule,rule) apply(subst imp_conjL,rule) defer apply(rule,rule,rule,rule,rule,rule,rule) proof-
+  fix x y u v assume as:"\<forall>t u. finite t \<and> t \<subseteq> s \<and> (\<forall>x\<in>t. 0 \<le> u x) \<and> setsum u t = 1 \<longrightarrow> (\<Sum>x\<in>t. u x *\<^sub>R x) \<in> s" "x \<in> s" "y \<in> s" "0 \<le> u" "0 \<le> v" "u + v = (1::real)"
+  show "u *\<^sub>R x + v *\<^sub>R y \<in> s" proof(cases "x=y")
+    case True show ?thesis unfolding True and scaleR_left_distrib[THEN sym] using as(3,6) by auto next
+    case False thus ?thesis using as(1)[THEN spec[where x="{x,y}"], THEN spec[where x="\<lambda>z. if z=x then u else v"]] and as(2-) by auto qed
+next 
+  fix t u assume asm:"\<forall>x\<in>s. \<forall>y\<in>s. \<forall>u\<ge>0. \<forall>v\<ge>0. u + v = 1 \<longrightarrow> u *\<^sub>R x + v *\<^sub>R y \<in> s" "finite (t::'a set)"
+  (*"finite t" "t \<subseteq> s" "\<forall>x\<in>t. (0::real) \<le> u x" "setsum u t = 1"*)
+  from this(2) have "\<forall>u. t \<subseteq> s \<and> (\<forall>x\<in>t. 0 \<le> u x) \<and> setsum u t = 1 \<longrightarrow> (\<Sum>x\<in>t. u x *\<^sub>R x) \<in> s" apply(induct_tac t rule:finite_induct)
+    prefer 3 apply (rule,rule) apply(erule conjE)+ proof-
+    fix x f u assume ind:"\<forall>u. f \<subseteq> s \<and> (\<forall>x\<in>f. 0 \<le> u x) \<and> setsum u f = 1 \<longrightarrow> (\<Sum>x\<in>f. u x *\<^sub>R x) \<in> s"
+    assume as:"finite f" "x \<notin> f" "insert x f \<subseteq> s" "\<forall>x\<in>insert x f. 0 \<le> u x" "setsum u (insert x f) = (1::real)"
+    show "(\<Sum>x\<in>insert x f. u x *\<^sub>R x) \<in> s" proof(cases "u x = 1")
+      case True hence "setsum (\<lambda>x. u x *\<^sub>R x) f = 0" apply(rule_tac setsum_0') apply(rule ccontr) unfolding ball_simps apply(erule bexE) proof-
+        fix y assume y:"y \<in> f" "u y *\<^sub>R y \<noteq> 0"
+        hence uy:"u y \<noteq> 0" by auto
+        hence "setsum (\<lambda>k. if k=y then u y else 0) f \<le> setsum u f" apply(rule_tac setsum_mono) using as(4) by auto
+        hence "setsum u f \<ge> u y" using y(1) and as(1) by(auto simp add: setsum_delta) 
+        hence "setsum u f > 0" using uy apply(rule_tac less_le_trans[of _ "u y"]) using as(4) and y(1) by auto
+        thus False using as(2,5) unfolding setsum_clauses(2)[OF as(1)] and True by auto qed
+      thus ?thesis unfolding setsum_clauses(2)[OF as(1)] using as(2,3) unfolding True by auto
+    next
+      have *:"setsum u f = setsum u (insert x f) - u x" using as(2) unfolding setsum_clauses(2)[OF as(1)] by auto
+      have **:"u x \<le> 1" apply(rule ccontr) unfolding not_le using as(5)[unfolded setsum_clauses(2)[OF as(1)]] and as(2)
+        using setsum_nonneg[of f u] and as(4) by auto
+      case False hence "inverse (1 - u x) *\<^sub>R (\<Sum>x\<in>f. u x *\<^sub>R x) \<in> s" unfolding scaleR_right.setsum and scaleR_scaleR
+        apply(rule_tac ind[THEN spec, THEN mp]) apply rule defer apply rule apply rule apply(rule mult_nonneg_nonneg)
+        unfolding setsum_right_distrib[THEN sym] and * using as and ** by auto
+      hence "u x *\<^sub>R x + (1 - u x) *\<^sub>R ((inverse (1 - u x)) *\<^sub>R setsum (\<lambda>x. u x *\<^sub>R x) f) \<in>s" 
+        apply(rule_tac asm(1)[THEN bspec, THEN bspec, THEN spec, THEN mp, THEN spec, THEN mp, THEN mp]) using as and ** False by auto 
+      thus ?thesis unfolding setsum_clauses(2)[OF as(1)] using as(2) and False by auto qed
+  qed auto thus "t \<subseteq> s \<and> (\<forall>x\<in>t. 0 \<le> u x) \<and> setsum u t = 1 \<longrightarrow> (\<Sum>x\<in>t. u x *\<^sub>R x) \<in> s" by auto
+qed
+
+lemma convex_finite: assumes "finite s"
+  shows "convex s \<longleftrightarrow> (\<forall>u. (\<forall>x\<in>s. 0 \<le> u x) \<and> setsum u s = 1
+                      \<longrightarrow> setsum (\<lambda>x. u x *\<^sub>R x) s \<in> s)"
+  unfolding convex_explicit apply(rule, rule, rule) defer apply(rule,rule,rule)apply(erule conjE)+ proof-
+  fix t u assume as:"\<forall>u. (\<forall>x\<in>s. 0 \<le> u x) \<and> setsum u s = 1 \<longrightarrow> (\<Sum>x\<in>s. u x *\<^sub>R x) \<in> s" " finite t" "t \<subseteq> s" "\<forall>x\<in>t. 0 \<le> u x" "setsum u t = (1::real)"
+  have *:"s \<inter> t = t" using as(3) by auto
+  show "(\<Sum>x\<in>t. u x *\<^sub>R x) \<in> s" using as(1)[THEN spec[where x="\<lambda>x. if x\<in>t then u x else 0"]]
+    unfolding if_smult and setsum_cases[OF assms] and * using as(2-) by auto
+qed (erule_tac x=s in allE, erule_tac x=u in allE, auto)
+
+subsection {* Cones. *}
+
+definition
+  cone :: "'a::real_vector set \<Rightarrow> bool" where
+  "cone s \<longleftrightarrow> (\<forall>x\<in>s. \<forall>c\<ge>0. (c *\<^sub>R x) \<in> s)"
+
+lemma cone_empty[intro, simp]: "cone {}"
+  unfolding cone_def by auto
+
+lemma cone_univ[intro, simp]: "cone UNIV"
+  unfolding cone_def by auto
+
+lemma cone_Inter[intro]: "(\<forall>s\<in>f. cone s) \<Longrightarrow> cone(\<Inter> f)"
+  unfolding cone_def by auto
+
+subsection {* Conic hull. *}
+
+lemma cone_cone_hull: "cone (cone hull s)"
+  unfolding hull_def using cone_Inter[of "{t \<in> conic. s \<subseteq> t}"] 
+  by (auto simp add: mem_def)
+
+lemma cone_hull_eq: "(cone hull s = s) \<longleftrightarrow> cone s"
+  apply(rule hull_eq[unfolded mem_def])
+  using cone_Inter unfolding subset_eq by (auto simp add: mem_def)
+
+subsection {* Affine dependence and consequential theorems (from Lars Schewe). *}
+
+definition
+  affine_dependent :: "'a::real_vector set \<Rightarrow> bool" where
+  "affine_dependent s \<longleftrightarrow> (\<exists>x\<in>s. x \<in> (affine hull (s - {x})))"
+
+lemma affine_dependent_explicit:
+  "affine_dependent p \<longleftrightarrow>
+    (\<exists>s u. finite s \<and> s \<subseteq> p \<and> setsum u s = 0 \<and>
+    (\<exists>v\<in>s. u v \<noteq> 0) \<and> setsum (\<lambda>v. u v *\<^sub>R v) s = 0)"
+  unfolding affine_dependent_def affine_hull_explicit mem_Collect_eq apply(rule)
+  apply(erule bexE,erule exE,erule exE) apply(erule conjE)+ defer apply(erule exE,erule exE) apply(erule conjE)+ apply(erule bexE)
+proof-
+  fix x s u assume as:"x \<in> p" "finite s" "s \<noteq> {}" "s \<subseteq> p - {x}" "setsum u s = 1" "(\<Sum>v\<in>s. u v *\<^sub>R v) = x"
+  have "x\<notin>s" using as(1,4) by auto
+  show "\<exists>s u. finite s \<and> s \<subseteq> p \<and> setsum u s = 0 \<and> (\<exists>v\<in>s. u v \<noteq> 0) \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = 0"
+    apply(rule_tac x="insert x s" in exI, rule_tac x="\<lambda>v. if v = x then - 1 else u v" in exI)
+    unfolding if_smult and setsum_clauses(2)[OF as(2)] and setsum_delta_notmem[OF `x\<notin>s`] and as using as by auto 
+next
+  fix s u v assume as:"finite s" "s \<subseteq> p" "setsum u s = 0" "(\<Sum>v\<in>s. u v *\<^sub>R v) = 0" "v \<in> s" "u v \<noteq> 0"
+  have "s \<noteq> {v}" using as(3,6) by auto
+  thus "\<exists>x\<in>p. \<exists>s u. finite s \<and> s \<noteq> {} \<and> s \<subseteq> p - {x} \<and> setsum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = x" 
+    apply(rule_tac x=v in bexI, rule_tac x="s - {v}" in exI, rule_tac x="\<lambda>x. - (1 / u v) * u x" in exI)
+    unfolding scaleR_scaleR[THEN sym] and scaleR_right.setsum [symmetric] unfolding setsum_right_distrib[THEN sym] and setsum_diff1[OF as(1)] using as by auto
+qed
+
+lemma affine_dependent_explicit_finite:
+  fixes s :: "'a::real_vector set" assumes "finite s"
+  shows "affine_dependent s \<longleftrightarrow> (\<exists>u. setsum u s = 0 \<and> (\<exists>v\<in>s. u v \<noteq> 0) \<and> setsum (\<lambda>v. u v *\<^sub>R v) s = 0)"
+  (is "?lhs = ?rhs")
+proof
+  have *:"\<And>vt u v. (if vt then u v else 0) *\<^sub>R v = (if vt then (u v) *\<^sub>R v else (0::'a))" by auto
+  assume ?lhs
+  then obtain t u v where "finite t" "t \<subseteq> s" "setsum u t = 0" "v\<in>t" "u v \<noteq> 0"  "(\<Sum>v\<in>t. u v *\<^sub>R v) = 0"
+    unfolding affine_dependent_explicit by auto
+  thus ?rhs apply(rule_tac x="\<lambda>x. if x\<in>t then u x else 0" in exI)
+    apply auto unfolding * and setsum_restrict_set[OF assms, THEN sym]
+    unfolding Int_absorb1[OF `t\<subseteq>s`] by auto
+next
+  assume ?rhs
+  then obtain u v where "setsum u s = 0"  "v\<in>s" "u v \<noteq> 0" "(\<Sum>v\<in>s. u v *\<^sub>R v) = 0" by auto
+  thus ?lhs unfolding affine_dependent_explicit using assms by auto
+qed
+
+subsection {* A general lemma. *}
+
+lemma convex_connected:
+  fixes s :: "'a::real_normed_vector set"
+  assumes "convex s" shows "connected s"
+proof-
+  { fix e1 e2 assume as:"open e1" "open e2" "e1 \<inter> e2 \<inter> s = {}" "s \<subseteq> e1 \<union> e2" 
+    assume "e1 \<inter> s \<noteq> {}" "e2 \<inter> s \<noteq> {}"
+    then obtain x1 x2 where x1:"x1\<in>e1" "x1\<in>s" and x2:"x2\<in>e2" "x2\<in>s" by auto
+    hence n:"norm (x1 - x2) > 0" unfolding zero_less_norm_iff using as(3) by auto
+
+    { fix x e::real assume as:"0 \<le> x" "x \<le> 1" "0 < e"
+      { fix y have *:"(1 - x) *\<^sub>R x1 + x *\<^sub>R x2 - ((1 - y) *\<^sub>R x1 + y *\<^sub>R x2) = (y - x) *\<^sub>R x1 - (y - x) *\<^sub>R x2"
+          by (simp add: algebra_simps)
+        assume "\<bar>y - x\<bar> < e / norm (x1 - x2)"
+        hence "norm ((1 - x) *\<^sub>R x1 + x *\<^sub>R x2 - ((1 - y) *\<^sub>R x1 + y *\<^sub>R x2)) < e"
+          unfolding * and scaleR_right_diff_distrib[THEN sym]
+          unfolding less_divide_eq using n by auto  }
+      hence "\<exists>d>0. \<forall>y. \<bar>y - x\<bar> < d \<longrightarrow> norm ((1 - x) *\<^sub>R x1 + x *\<^sub>R x2 - ((1 - y) *\<^sub>R x1 + y *\<^sub>R x2)) < e"
+        apply(rule_tac x="e / norm (x1 - x2)" in exI) using as
+        apply auto unfolding zero_less_divide_iff using n by simp  }  note * = this
+
+    have "\<exists>x\<ge>0. x \<le> 1 \<and> (1 - x) *\<^sub>R x1 + x *\<^sub>R x2 \<notin> e1 \<and> (1 - x) *\<^sub>R x1 + x *\<^sub>R x2 \<notin> e2"
+      apply(rule connected_real_lemma) apply (simp add: `x1\<in>e1` `x2\<in>e2` dist_commute)+
+      using * apply(simp add: dist_norm)
+      using as(1,2)[unfolded open_dist] apply simp
+      using as(1,2)[unfolded open_dist] apply simp
+      using assms[unfolded convex_alt, THEN bspec[where x=x1], THEN bspec[where x=x2]] using x1 x2
+      using as(3) by auto
+    then obtain x where "x\<ge>0" "x\<le>1" "(1 - x) *\<^sub>R x1 + x *\<^sub>R x2 \<notin> e1"  "(1 - x) *\<^sub>R x1 + x *\<^sub>R x2 \<notin> e2" by auto
+    hence False using as(4) 
+      using assms[unfolded convex_alt, THEN bspec[where x=x1], THEN bspec[where x=x2]]
+      using x1(2) x2(2) by auto  }
+  thus ?thesis unfolding connected_def by auto
+qed
+
+subsection {* One rather trivial consequence. *}
+
+lemma connected_UNIV: "connected (UNIV :: 'a::real_normed_vector set)"
+  by(simp add: convex_connected convex_UNIV)
+
+subsection {* Convex functions into the reals. *}
+
+definition
+  convex_on :: "'a::real_vector set \<Rightarrow> ('a \<Rightarrow> real) \<Rightarrow> bool" where
+  "convex_on s f \<longleftrightarrow>
+  (\<forall>x\<in>s. \<forall>y\<in>s. \<forall>u\<ge>0. \<forall>v\<ge>0. u + v = 1 \<longrightarrow> f (u *\<^sub>R x + v *\<^sub>R y) \<le> u * f x + v * f y)"
+
+lemma convex_on_subset: "convex_on t f \<Longrightarrow> s \<subseteq> t \<Longrightarrow> convex_on s f"
+  unfolding convex_on_def by auto
+
+lemma convex_add:
+  assumes "convex_on s f" "convex_on s g"
+  shows "convex_on s (\<lambda>x. f x + g x)"
+proof-
+  { fix x y assume "x\<in>s" "y\<in>s" moreover
+    fix u v ::real assume "0 \<le> u" "0 \<le> v" "u + v = 1"
+    ultimately have "f (u *\<^sub>R x + v *\<^sub>R y) + g (u *\<^sub>R x + v *\<^sub>R y) \<le> (u * f x + v * f y) + (u * g x + v * g y)"
+      using assms(1)[unfolded convex_on_def, THEN bspec[where x=x], THEN bspec[where x=y], THEN spec[where x=u]]
+      using assms(2)[unfolded convex_on_def, THEN bspec[where x=x], THEN bspec[where x=y], THEN spec[where x=u]]
+      apply - apply(rule add_mono) by auto
+    hence "f (u *\<^sub>R x + v *\<^sub>R y) + g (u *\<^sub>R x + v *\<^sub>R y) \<le> u * (f x + g x) + v * (f y + g y)" by (simp add: ring_simps)  }
+  thus ?thesis unfolding convex_on_def by auto 
+qed
+
+lemma convex_cmul: 
+  assumes "0 \<le> (c::real)" "convex_on s f"
+  shows "convex_on s (\<lambda>x. c * f x)"
+proof-
+  have *:"\<And>u c fx v fy ::real. u * (c * fx) + v * (c * fy) = c * (u * fx + v * fy)" by (simp add: ring_simps)
+  show ?thesis using assms(2) and mult_mono1[OF _ assms(1)] unfolding convex_on_def and * by auto
+qed
+
+lemma convex_lower:
+  assumes "convex_on s f"  "x\<in>s"  "y \<in> s"  "0 \<le> u"  "0 \<le> v"  "u + v = 1"
+  shows "f (u *\<^sub>R x + v *\<^sub>R y) \<le> max (f x) (f y)"
+proof-
+  let ?m = "max (f x) (f y)"
+  have "u * f x + v * f y \<le> u * max (f x) (f y) + v * max (f x) (f y)" apply(rule add_mono) 
+    using assms(4,5) by(auto simp add: mult_mono1)
+  also have "\<dots> = max (f x) (f y)" using assms(6) unfolding distrib[THEN sym] by auto
+  finally show ?thesis using assms(1)[unfolded convex_on_def, THEN bspec[where x=x], THEN bspec[where x=y], THEN spec[where x=u]]
+    using assms(2-6) by auto 
+qed
+
+lemma convex_local_global_minimum:
+  fixes s :: "'a::real_normed_vector set"
+  assumes "0<e" "convex_on s f" "ball x e \<subseteq> s" "\<forall>y\<in>ball x e. f x \<le> f y"
+  shows "\<forall>y\<in>s. f x \<le> f y"
+proof(rule ccontr)
+  have "x\<in>s" using assms(1,3) by auto
+  assume "\<not> (\<forall>y\<in>s. f x \<le> f y)"
+  then obtain y where "y\<in>s" and y:"f x > f y" by auto
+  hence xy:"0 < dist x y" by (auto simp add: dist_nz[THEN sym])
+
+  then obtain u where "0 < u" "u \<le> 1" and u:"u < e / dist x y"
+    using real_lbound_gt_zero[of 1 "e / dist x y"] using xy `e>0` and divide_pos_pos[of e "dist x y"] by auto
+  hence "f ((1-u) *\<^sub>R x + u *\<^sub>R y) \<le> (1-u) * f x + u * f y" using `x\<in>s` `y\<in>s`
+    using assms(2)[unfolded convex_on_def, THEN bspec[where x=x], THEN bspec[where x=y], THEN spec[where x="1-u"]] by auto
+  moreover
+  have *:"x - ((1 - u) *\<^sub>R x + u *\<^sub>R y) = u *\<^sub>R (x - y)" by (simp add: algebra_simps)
+  have "(1 - u) *\<^sub>R x + u *\<^sub>R y \<in> ball x e" unfolding mem_ball dist_norm unfolding * and norm_scaleR and abs_of_pos[OF `0<u`] unfolding dist_norm[THEN sym]
+    using u unfolding pos_less_divide_eq[OF xy] by auto
+  hence "f x \<le> f ((1 - u) *\<^sub>R x + u *\<^sub>R y)" using assms(4) by auto
+  ultimately show False using mult_strict_left_mono[OF y `u>0`] unfolding left_diff_distrib by auto
+qed
+
+lemma convex_distance:
+  fixes s :: "'a::real_normed_vector set"
+  shows "convex_on s (\<lambda>x. dist a x)"
+proof(auto simp add: convex_on_def dist_norm)
+  fix x y assume "x\<in>s" "y\<in>s"
+  fix u v ::real assume "0 \<le> u" "0 \<le> v" "u + v = 1"
+  have "a = u *\<^sub>R a + v *\<^sub>R a" unfolding scaleR_left_distrib[THEN sym] and `u+v=1` by simp
+  hence *:"a - (u *\<^sub>R x + v *\<^sub>R y) = (u *\<^sub>R (a - x)) + (v *\<^sub>R (a - y))"
+    by (auto simp add: algebra_simps)
+  show "norm (a - (u *\<^sub>R x + v *\<^sub>R y)) \<le> u * norm (a - x) + v * norm (a - y)"
+    unfolding * using norm_triangle_ineq[of "u *\<^sub>R (a - x)" "v *\<^sub>R (a - y)"]
+    using `0 \<le> u` `0 \<le> v` by auto
+qed
+
+subsection {* Arithmetic operations on sets preserve convexity. *}
+
+lemma convex_scaling: "convex s \<Longrightarrow> convex ((\<lambda>x. c *\<^sub>R x) ` s)"
+  unfolding convex_def and image_iff apply auto
+  apply (rule_tac x="u *\<^sub>R x+v *\<^sub>R y" in bexI) by (auto simp add: algebra_simps)
+
+lemma convex_negations: "convex s \<Longrightarrow> convex ((\<lambda>x. -x)` s)"
+  unfolding convex_def and image_iff apply auto
+  apply (rule_tac x="u *\<^sub>R x+v *\<^sub>R y" in bexI) by auto
+
+lemma convex_sums:
+  assumes "convex s" "convex t"
+  shows "convex {x + y| x y. x \<in> s \<and> y \<in> t}"
+proof(auto simp add: convex_def image_iff scaleR_right_distrib)
+  fix xa xb ya yb assume xy:"xa\<in>s" "xb\<in>s" "ya\<in>t" "yb\<in>t"
+  fix u v ::real assume uv:"0 \<le> u" "0 \<le> v" "u + v = 1"
+  show "\<exists>x y. u *\<^sub>R xa + u *\<^sub>R ya + (v *\<^sub>R xb + v *\<^sub>R yb) = x + y \<and> x \<in> s \<and> y \<in> t"
+    apply(rule_tac x="u *\<^sub>R xa + v *\<^sub>R xb" in exI) apply(rule_tac x="u *\<^sub>R ya + v *\<^sub>R yb" in exI)
+    using assms(1)[unfolded convex_def, THEN bspec[where x=xa], THEN bspec[where x=xb]]
+    using assms(2)[unfolded convex_def, THEN bspec[where x=ya], THEN bspec[where x=yb]]
+    using uv xy by auto
+qed
+
+lemma convex_differences: 
+  assumes "convex s" "convex t"
+  shows "convex {x - y| x y. x \<in> s \<and> y \<in> t}"
+proof-
+  have "{x - y| x y. x \<in> s \<and> y \<in> t} = {x + y |x y. x \<in> s \<and> y \<in> uminus ` t}" unfolding image_iff apply auto
+    apply(rule_tac x=xa in exI) apply(rule_tac x="-y" in exI) apply simp
+    apply(rule_tac x=xa in exI) apply(rule_tac x=xb in exI) by simp
+  thus ?thesis using convex_sums[OF assms(1)  convex_negations[OF assms(2)]] by auto
+qed
+
+lemma convex_translation: assumes "convex s" shows "convex ((\<lambda>x. a + x) ` s)"
+proof- have "{a + y |y. y \<in> s} = (\<lambda>x. a + x) ` s" by auto
+  thus ?thesis using convex_sums[OF convex_singleton[of a] assms] by auto qed
+
+lemma convex_affinity: assumes "convex s" shows "convex ((\<lambda>x. a + c *\<^sub>R x) ` s)"
+proof- have "(\<lambda>x. a + c *\<^sub>R x) ` s = op + a ` op *\<^sub>R c ` s" by auto
+  thus ?thesis using convex_translation[OF convex_scaling[OF assms], of a c] by auto qed
+
+lemma convex_linear_image:
+  assumes c:"convex s" and l:"bounded_linear f"
+  shows "convex(f ` s)"
+proof(auto simp add: convex_def)
+  interpret f: bounded_linear f by fact
+  fix x y assume xy:"x \<in> s" "y \<in> s"
+  fix u v ::real assume uv:"0 \<le> u" "0 \<le> v" "u + v = 1"
+  show "u *\<^sub>R f x + v *\<^sub>R f y \<in> f ` s" unfolding image_iff
+    apply(rule_tac x="u *\<^sub>R x + v *\<^sub>R y" in bexI)
+    unfolding f.add f.scaleR
+    using c[unfolded convex_def] xy uv by auto
+qed
+
+subsection {* Balls, being convex, are connected. *}
+
+lemma convex_ball:
+  fixes x :: "'a::real_normed_vector"
+  shows "convex (ball x e)" 
+proof(auto simp add: convex_def)
+  fix y z assume yz:"dist x y < e" "dist x z < e"
+  fix u v ::real assume uv:"0 \<le> u" "0 \<le> v" "u + v = 1"
+  have "dist x (u *\<^sub>R y + v *\<^sub>R z) \<le> u * dist x y + v * dist x z" using uv yz
+    using convex_distance[of "ball x e" x, unfolded convex_on_def, THEN bspec[where x=y], THEN bspec[where x=z]] by auto
+  thus "dist x (u *\<^sub>R y + v *\<^sub>R z) < e" using real_convex_bound_lt[OF yz uv] by auto 
+qed
+
+lemma convex_cball:
+  fixes x :: "'a::real_normed_vector"
+  shows "convex(cball x e)"
+proof(auto simp add: convex_def Ball_def mem_cball)
+  fix y z assume yz:"dist x y \<le> e" "dist x z \<le> e"
+  fix u v ::real assume uv:" 0 \<le> u" "0 \<le> v" "u + v = 1"
+  have "dist x (u *\<^sub>R y + v *\<^sub>R z) \<le> u * dist x y + v * dist x z" using uv yz
+    using convex_distance[of "cball x e" x, unfolded convex_on_def, THEN bspec[where x=y], THEN bspec[where x=z]] by auto
+  thus "dist x (u *\<^sub>R y + v *\<^sub>R z) \<le> e" using real_convex_bound_le[OF yz uv] by auto 
+qed
+
+lemma connected_ball:
+  fixes x :: "'a::real_normed_vector"
+  shows "connected (ball x e)"
+  using convex_connected convex_ball by auto
+
+lemma connected_cball:
+  fixes x :: "'a::real_normed_vector"
+  shows "connected(cball x e)"
+  using convex_connected convex_cball by auto
+
+subsection {* Convex hull. *}
+
+lemma convex_convex_hull: "convex(convex hull s)"
+  unfolding hull_def using convex_Inter[of "{t\<in>convex. s\<subseteq>t}"]
+  unfolding mem_def by auto
+
+lemma convex_hull_eq: "(convex hull s = s) \<longleftrightarrow> convex s" apply(rule hull_eq[unfolded mem_def])
+  using convex_Inter[unfolded Ball_def mem_def] by auto
+
+lemma bounded_convex_hull:
+  fixes s :: "'a::real_normed_vector set"
+  assumes "bounded s" shows "bounded(convex hull s)"
+proof- from assms obtain B where B:"\<forall>x\<in>s. norm x \<le> B" unfolding bounded_iff by auto
+  show ?thesis apply(rule bounded_subset[OF bounded_cball, of _ 0 B])
+    unfolding subset_hull[unfolded mem_def, of convex, OF convex_cball]
+    unfolding subset_eq mem_cball dist_norm using B by auto qed
+
+lemma finite_imp_bounded_convex_hull:
+  fixes s :: "'a::real_normed_vector set"
+  shows "finite s \<Longrightarrow> bounded(convex hull s)"
+  using bounded_convex_hull finite_imp_bounded by auto
+
+subsection {* Stepping theorems for convex hulls of finite sets. *}
+
+lemma convex_hull_empty[simp]: "convex hull {} = {}"
+  apply(rule hull_unique) unfolding mem_def by auto
+
+lemma convex_hull_singleton[simp]: "convex hull {a} = {a}"
+  apply(rule hull_unique) unfolding mem_def by auto
+
+lemma convex_hull_insert:
+  fixes s :: "'a::real_vector set"
+  assumes "s \<noteq> {}"
+  shows "convex hull (insert a s) = {x. \<exists>u\<ge>0. \<exists>v\<ge>0. \<exists>b. (u + v = 1) \<and>
+                                    b \<in> (convex hull s) \<and> (x = u *\<^sub>R a + v *\<^sub>R b)}" (is "?xyz = ?hull")
+ apply(rule,rule hull_minimal,rule) unfolding mem_def[of _ convex] and insert_iff prefer 3 apply rule proof-
+ fix x assume x:"x = a \<or> x \<in> s"
+ thus "x\<in>?hull" apply rule unfolding mem_Collect_eq apply(rule_tac x=1 in exI) defer 
+   apply(rule_tac x=0 in exI) using assms hull_subset[of s convex] by auto
+next
+  fix x assume "x\<in>?hull"
+  then obtain u v b where obt:"u\<ge>0" "v\<ge>0" "u + v = 1" "b \<in> convex hull s" "x = u *\<^sub>R a + v *\<^sub>R b" by auto
+  have "a\<in>convex hull insert a s" "b\<in>convex hull insert a s"
+    using hull_mono[of s "insert a s" convex] hull_mono[of "{a}" "insert a s" convex] and obt(4) by auto
+  thus "x\<in> convex hull insert a s" unfolding obt(5) using convex_convex_hull[of "insert a s", unfolded convex_def]
+    apply(erule_tac x=a in ballE) apply(erule_tac x=b in ballE) apply(erule_tac x=u in allE) using obt by auto
+next
+  show "convex ?hull" unfolding convex_def apply(rule,rule,rule,rule,rule,rule,rule) proof-
+    fix x y u v assume as:"(0::real) \<le> u" "0 \<le> v" "u + v = 1" "x\<in>?hull" "y\<in>?hull"
+    from as(4) obtain u1 v1 b1 where obt1:"u1\<ge>0" "v1\<ge>0" "u1 + v1 = 1" "b1 \<in> convex hull s" "x = u1 *\<^sub>R a + v1 *\<^sub>R b1" by auto
+    from as(5) obtain u2 v2 b2 where obt2:"u2\<ge>0" "v2\<ge>0" "u2 + v2 = 1" "b2 \<in> convex hull s" "y = u2 *\<^sub>R a + v2 *\<^sub>R b2" by auto
+    have *:"\<And>(x::'a) s1 s2. x - s1 *\<^sub>R x - s2 *\<^sub>R x = ((1::real) - (s1 + s2)) *\<^sub>R x" by (auto simp add: algebra_simps)
+    have "\<exists>b \<in> convex hull s. u *\<^sub>R x + v *\<^sub>R y = (u * u1) *\<^sub>R a + (v * u2) *\<^sub>R a + (b - (u * u1) *\<^sub>R b - (v * u2) *\<^sub>R b)"
+    proof(cases "u * v1 + v * v2 = 0")
+      have *:"\<And>(x::'a) s1 s2. x - s1 *\<^sub>R x - s2 *\<^sub>R x = ((1::real) - (s1 + s2)) *\<^sub>R x" by (auto simp add: algebra_simps)
+      case True hence **:"u * v1 = 0" "v * v2 = 0" apply- apply(rule_tac [!] ccontr)
+        using mult_nonneg_nonneg[OF `u\<ge>0` `v1\<ge>0`] mult_nonneg_nonneg[OF `v\<ge>0` `v2\<ge>0`] by auto
+      hence "u * u1 + v * u2 = 1" using as(3) obt1(3) obt2(3) by auto
+      thus ?thesis unfolding obt1(5) obt2(5) * using assms hull_subset[of s convex] by(auto simp add: ** scaleR_right_distrib)
+    next
+      have "1 - (u * u1 + v * u2) = (u + v) - (u * u1 + v * u2)" using as(3) obt1(3) obt2(3) by (auto simp add: field_simps)
+      also have "\<dots> = u * (v1 + u1 - u1) + v * (v2 + u2 - u2)" using as(3) obt1(3) obt2(3) by (auto simp add: field_simps) 
+      also have "\<dots> = u * v1 + v * v2" by simp finally have **:"1 - (u * u1 + v * u2) = u * v1 + v * v2" by auto
+      case False have "0 \<le> u * v1 + v * v2" "0 \<le> u * v1" "0 \<le> u * v1 + v * v2" "0 \<le> v * v2" apply -
+        apply(rule add_nonneg_nonneg) prefer 4 apply(rule add_nonneg_nonneg) apply(rule_tac [!] mult_nonneg_nonneg)
+        using as(1,2) obt1(1,2) obt2(1,2) by auto 
+      thus ?thesis unfolding obt1(5) obt2(5) unfolding * and ** using False
+        apply(rule_tac x="((u * v1) / (u * v1 + v * v2)) *\<^sub>R b1 + ((v * v2) / (u * v1 + v * v2)) *\<^sub>R b2" in bexI) defer
+        apply(rule convex_convex_hull[of s, unfolded convex_def, rule_format]) using obt1(4) obt2(4)
+        unfolding add_divide_distrib[THEN sym] and real_0_le_divide_iff
+        by (auto simp add: scaleR_left_distrib scaleR_right_distrib)
+    qed note * = this
+    have u1:"u1 \<le> 1" apply(rule ccontr) unfolding obt1(3)[THEN sym] and not_le using obt1(2) by auto
+    have u2:"u2 \<le> 1" apply(rule ccontr) unfolding obt2(3)[THEN sym] and not_le using obt2(2) by auto
+    have "u1 * u + u2 * v \<le> (max u1 u2) * u + (max u1 u2) * v" apply(rule add_mono)
+      apply(rule_tac [!] mult_right_mono) using as(1,2) obt1(1,2) obt2(1,2) by auto
+    also have "\<dots> \<le> 1" unfolding mult.add_right[THEN sym] and as(3) using u1 u2 by auto
+    finally 
+    show "u *\<^sub>R x + v *\<^sub>R y \<in> ?hull" unfolding mem_Collect_eq apply(rule_tac x="u * u1 + v * u2" in exI)
+      apply(rule conjI) defer apply(rule_tac x="1 - u * u1 - v * u2" in exI) unfolding Bex_def
+      using as(1,2) obt1(1,2) obt2(1,2) * by(auto intro!: mult_nonneg_nonneg add_nonneg_nonneg simp add: algebra_simps)
+  qed
+qed
+
+
+subsection {* Explicit expression for convex hull. *}
+
+lemma convex_hull_indexed:
+  fixes s :: "'a::real_vector set"
+  shows "convex hull s = {y. \<exists>k u x. (\<forall>i\<in>{1::nat .. k}. 0 \<le> u i \<and> x i \<in> s) \<and>
+                            (setsum u {1..k} = 1) \<and>
+                            (setsum (\<lambda>i. u i *\<^sub>R x i) {1..k} = y)}" (is "?xyz = ?hull")
+  apply(rule hull_unique) unfolding mem_def[of _ convex] apply(rule) defer
+  apply(subst convex_def) apply(rule,rule,rule,rule,rule,rule,rule)
+proof-
+  fix x assume "x\<in>s"
+  thus "x \<in> ?hull" unfolding mem_Collect_eq apply(rule_tac x=1 in exI, rule_tac x="\<lambda>x. 1" in exI) by auto
+next
+  fix t assume as:"s \<subseteq> t" "convex t"
+  show "?hull \<subseteq> t" apply(rule) unfolding mem_Collect_eq apply(erule exE | erule conjE)+ proof-
+    fix x k u y assume assm:"\<forall>i\<in>{1::nat..k}. 0 \<le> u i \<and> y i \<in> s" "setsum u {1..k} = 1" "(\<Sum>i = 1..k. u i *\<^sub>R y i) = x"
+    show "x\<in>t" unfolding assm(3)[THEN sym] apply(rule as(2)[unfolded convex, rule_format])
+      using assm(1,2) as(1) by auto qed
+next
+  fix x y u v assume uv:"0\<le>u" "0\<le>v" "u+v=(1::real)" and xy:"x\<in>?hull" "y\<in>?hull"
+  from xy obtain k1 u1 x1 where x:"\<forall>i\<in>{1::nat..k1}. 0\<le>u1 i \<and> x1 i \<in> s" "setsum u1 {Suc 0..k1} = 1" "(\<Sum>i = Suc 0..k1. u1 i *\<^sub>R x1 i) = x" by auto
+  from xy obtain k2 u2 x2 where y:"\<forall>i\<in>{1::nat..k2}. 0\<le>u2 i \<and> x2 i \<in> s" "setsum u2 {Suc 0..k2} = 1" "(\<Sum>i = Suc 0..k2. u2 i *\<^sub>R x2 i) = y" by auto
+  have *:"\<And>P (x1::'a) x2 s1 s2 i.(if P i then s1 else s2) *\<^sub>R (if P i then x1 else x2) = (if P i then s1 *\<^sub>R x1 else s2 *\<^sub>R x2)"
+    "{1..k1 + k2} \<inter> {1..k1} = {1..k1}" "{1..k1 + k2} \<inter> - {1..k1} = (\<lambda>i. i + k1) ` {1..k2}"
+    prefer 3 apply(rule,rule) unfolding image_iff apply(rule_tac x="x - k1" in bexI) by(auto simp add: not_le)
+  have inj:"inj_on (\<lambda>i. i + k1) {1..k2}" unfolding inj_on_def by auto  
+  show "u *\<^sub>R x + v *\<^sub>R y \<in> ?hull" apply(rule)
+    apply(rule_tac x="k1 + k2" in exI, rule_tac x="\<lambda>i. if i \<in> {1..k1} then u * u1 i else v * u2 (i - k1)" in exI)
+    apply(rule_tac x="\<lambda>i. if i \<in> {1..k1} then x1 i else x2 (i - k1)" in exI) apply(rule,rule) defer apply(rule)
+    unfolding * and setsum_cases[OF finite_atLeastAtMost[of 1 "k1 + k2"]] and setsum_reindex[OF inj] and o_def
+    unfolding scaleR_scaleR[THEN sym] scaleR_right.setsum [symmetric] setsum_right_distrib[THEN sym] proof-
+    fix i assume i:"i \<in> {1..k1+k2}"
+    show "0 \<le> (if i \<in> {1..k1} then u * u1 i else v * u2 (i - k1)) \<and> (if i \<in> {1..k1} then x1 i else x2 (i - k1)) \<in> s"
+    proof(cases "i\<in>{1..k1}")
+      case True thus ?thesis using mult_nonneg_nonneg[of u "u1 i"] and uv(1) x(1)[THEN bspec[where x=i]] by auto
+    next def j \<equiv> "i - k1"
+      case False with i have "j \<in> {1..k2}" unfolding j_def by auto
+      thus ?thesis unfolding j_def[symmetric] using False
+        using mult_nonneg_nonneg[of v "u2 j"] and uv(2) y(1)[THEN bspec[where x=j]] by auto qed
+  qed(auto simp add: not_le x(2,3) y(2,3) uv(3))
+qed
+
+lemma convex_hull_finite:
+  fixes s :: "'a::real_vector set"
+  assumes "finite s"
+  shows "convex hull s = {y. \<exists>u. (\<forall>x\<in>s. 0 \<le> u x) \<and>
+         setsum u s = 1 \<and> setsum (\<lambda>x. u x *\<^sub>R x) s = y}" (is "?HULL = ?set")
+proof(rule hull_unique, auto simp add: mem_def[of _ convex] convex_def[of ?set])
+  fix x assume "x\<in>s" thus " \<exists>u. (\<forall>x\<in>s. 0 \<le> u x) \<and> setsum u s = 1 \<and> (\<Sum>x\<in>s. u x *\<^sub>R x) = x" 
+    apply(rule_tac x="\<lambda>y. if x=y then 1 else 0" in exI) apply auto
+    unfolding setsum_delta'[OF assms] and setsum_delta''[OF assms] by auto 
+next
+  fix u v ::real assume uv:"0 \<le> u" "0 \<le> v" "u + v = 1"
+  fix ux assume ux:"\<forall>x\<in>s. 0 \<le> ux x" "setsum ux s = (1::real)"
+  fix uy assume uy:"\<forall>x\<in>s. 0 \<le> uy x" "setsum uy s = (1::real)"
+  { fix x assume "x\<in>s"
+    hence "0 \<le> u * ux x + v * uy x" using ux(1)[THEN bspec[where x=x]] uy(1)[THEN bspec[where x=x]] and uv(1,2)
+      by (auto, metis add_nonneg_nonneg mult_nonneg_nonneg uv(1) uv(2))  }
+  moreover have "(\<Sum>x\<in>s. u * ux x + v * uy x) = 1"
+    unfolding setsum_addf and setsum_right_distrib[THEN sym] and ux(2) uy(2) using uv(3) by auto
+  moreover have "(\<Sum>x\<in>s. (u * ux x + v * uy x) *\<^sub>R x) = u *\<^sub>R (\<Sum>x\<in>s. ux x *\<^sub>R x) + v *\<^sub>R (\<Sum>x\<in>s. uy x *\<^sub>R x)"
+    unfolding scaleR_left_distrib and setsum_addf and scaleR_scaleR[THEN sym] and scaleR_right.setsum [symmetric] by auto
+  ultimately show "\<exists>uc. (\<forall>x\<in>s. 0 \<le> uc x) \<and> setsum uc s = 1 \<and> (\<Sum>x\<in>s. uc x *\<^sub>R x) = u *\<^sub>R (\<Sum>x\<in>s. ux x *\<^sub>R x) + v *\<^sub>R (\<Sum>x\<in>s. uy x *\<^sub>R x)"
+    apply(rule_tac x="\<lambda>x. u * ux x + v * uy x" in exI) by auto 
+next
+  fix t assume t:"s \<subseteq> t" "convex t" 
+  fix u assume u:"\<forall>x\<in>s. 0 \<le> u x" "setsum u s = (1::real)"
+  thus "(\<Sum>x\<in>s. u x *\<^sub>R x) \<in> t" using t(2)[unfolded convex_explicit, THEN spec[where x=s], THEN spec[where x=u]]
+    using assms and t(1) by auto
+qed
+
+subsection {* Another formulation from Lars Schewe. *}
+
+lemma setsum_constant_scaleR:
+  fixes y :: "'a::real_vector"
+  shows "(\<Sum>x\<in>A. y) = of_nat (card A) *\<^sub>R y"
+apply (cases "finite A")
+apply (induct set: finite)
+apply (simp_all add: algebra_simps)
+done
+
+lemma convex_hull_explicit:
+  fixes p :: "'a::real_vector set"
+  shows "convex hull p = {y. \<exists>s u. finite s \<and> s \<subseteq> p \<and>
+             (\<forall>x\<in>s. 0 \<le> u x) \<and> setsum u s = 1 \<and> setsum (\<lambda>v. u v *\<^sub>R v) s = y}" (is "?lhs = ?rhs")
+proof-
+  { fix x assume "x\<in>?lhs"
+    then obtain k u y where obt:"\<forall>i\<in>{1::nat..k}. 0 \<le> u i \<and> y i \<in> p" "setsum u {1..k} = 1" "(\<Sum>i = 1..k. u i *\<^sub>R y i) = x"
+      unfolding convex_hull_indexed by auto
+
+    have fin:"finite {1..k}" by auto
+    have fin':"\<And>v. finite {i \<in> {1..k}. y i = v}" by auto
+    { fix j assume "j\<in>{1..k}"
+      hence "y j \<in> p" "0 \<le> setsum u {i. Suc 0 \<le> i \<and> i \<le> k \<and> y i = y j}"
+        using obt(1)[THEN bspec[where x=j]] and obt(2) apply simp
+        apply(rule setsum_nonneg) using obt(1) by auto } 
+    moreover
+    have "(\<Sum>v\<in>y ` {1..k}. setsum u {i \<in> {1..k}. y i = v}) = 1"  
+      unfolding setsum_image_gen[OF fin, THEN sym] using obt(2) by auto
+    moreover have "(\<Sum>v\<in>y ` {1..k}. setsum u {i \<in> {1..k}. y i = v} *\<^sub>R v) = x"
+      using setsum_image_gen[OF fin, of "\<lambda>i. u i *\<^sub>R y i" y, THEN sym]
+      unfolding scaleR_left.setsum using obt(3) by auto
+    ultimately have "\<exists>s u. finite s \<and> s \<subseteq> p \<and> (\<forall>x\<in>s. 0 \<le> u x) \<and> setsum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = x"
+      apply(rule_tac x="y ` {1..k}" in exI)
+      apply(rule_tac x="\<lambda>v. setsum u {i\<in>{1..k}. y i = v}" in exI) by auto
+    hence "x\<in>?rhs" by auto  }
+  moreover
+  { fix y assume "y\<in>?rhs"
+    then obtain s u where obt:"finite s" "s \<subseteq> p" "\<forall>x\<in>s. 0 \<le> u x" "setsum u s = 1" "(\<Sum>v\<in>s. u v *\<^sub>R v) = y" by auto
+
+    obtain f where f:"inj_on f {1..card s}" "f ` {1..card s} = s" using ex_bij_betw_nat_finite_1[OF obt(1)] unfolding bij_betw_def by auto
+    
+    { fix i::nat assume "i\<in>{1..card s}"
+      hence "f i \<in> s"  apply(subst f(2)[THEN sym]) by auto
+      hence "0 \<le> u (f i)" "f i \<in> p" using obt(2,3) by auto  }
+    moreover have *:"finite {1..card s}" by auto
+    { fix y assume "y\<in>s"
+      then obtain i where "i\<in>{1..card s}" "f i = y" using f using image_iff[of y f "{1..card s}"] by auto
+      hence "{x. Suc 0 \<le> x \<and> x \<le> card s \<and> f x = y} = {i}" apply auto using f(1)[unfolded inj_on_def] apply(erule_tac x=x in ballE) by auto
+      hence "card {x. Suc 0 \<le> x \<and> x \<le> card s \<and> f x = y} = 1" by auto
+      hence "(\<Sum>x\<in>{x \<in> {1..card s}. f x = y}. u (f x)) = u y"
+            "(\<Sum>x\<in>{x \<in> {1..card s}. f x = y}. u (f x) *\<^sub>R f x) = u y *\<^sub>R y"
+        by (auto simp add: setsum_constant_scaleR)   }
+
+    hence "(\<Sum>x = 1..card s. u (f x)) = 1" "(\<Sum>i = 1..card s. u (f i) *\<^sub>R f i) = y"
+      unfolding setsum_image_gen[OF *(1), of "\<lambda>x. u (f x) *\<^sub>R f x" f] and setsum_image_gen[OF *(1), of "\<lambda>x. u (f x)" f] 
+      unfolding f using setsum_cong2[of s "\<lambda>y. (\<Sum>x\<in>{x \<in> {1..card s}. f x = y}. u (f x) *\<^sub>R f x)" "\<lambda>v. u v *\<^sub>R v"]
+      using setsum_cong2 [of s "\<lambda>y. (\<Sum>x\<in>{x \<in> {1..card s}. f x = y}. u (f x))" u] unfolding obt(4,5) by auto
+    
+    ultimately have "\<exists>k u x. (\<forall>i\<in>{1..k}. 0 \<le> u i \<and> x i \<in> p) \<and> setsum u {1..k} = 1 \<and> (\<Sum>i::nat = 1..k. u i *\<^sub>R x i) = y"
+      apply(rule_tac x="card s" in exI) apply(rule_tac x="u \<circ> f" in exI) apply(rule_tac x=f in exI) by fastsimp
+    hence "y \<in> ?lhs" unfolding convex_hull_indexed by auto  }
+  ultimately show ?thesis unfolding expand_set_eq by blast
+qed
+
+subsection {* A stepping theorem for that expansion. *}
+
+lemma convex_hull_finite_step:
+  fixes s :: "'a::real_vector set" assumes "finite s"
+  shows "(\<exists>u. (\<forall>x\<in>insert a s. 0 \<le> u x) \<and> setsum u (insert a s) = w \<and> setsum (\<lambda>x. u x *\<^sub>R x) (insert a s) = y)
+     \<longleftrightarrow> (\<exists>v\<ge>0. \<exists>u. (\<forall>x\<in>s. 0 \<le> u x) \<and> setsum u s = w - v \<and> setsum (\<lambda>x. u x *\<^sub>R x) s = y - v *\<^sub>R a)" (is "?lhs = ?rhs")
+proof(rule, case_tac[!] "a\<in>s")
+  assume "a\<in>s" hence *:"insert a s = s" by auto
+  assume ?lhs thus ?rhs unfolding * apply(rule_tac x=0 in exI) by auto
+next
+  assume ?lhs then obtain u where u:"\<forall>x\<in>insert a s. 0 \<le> u x" "setsum u (insert a s) = w" "(\<Sum>x\<in>insert a s. u x *\<^sub>R x) = y" by auto
+  assume "a\<notin>s" thus ?rhs apply(rule_tac x="u a" in exI) using u(1)[THEN bspec[where x=a]] apply simp
+    apply(rule_tac x=u in exI) using u[unfolded setsum_clauses(2)[OF assms]] and `a\<notin>s` by auto
+next
+  assume "a\<in>s" hence *:"insert a s = s" by auto
+  have fin:"finite (insert a s)" using assms by auto
+  assume ?rhs then obtain v u where uv:"v\<ge>0" "\<forall>x\<in>s. 0 \<le> u x" "setsum u s = w - v" "(\<Sum>x\<in>s. u x *\<^sub>R x) = y - v *\<^sub>R a" by auto
+  show ?lhs apply(rule_tac x="\<lambda>x. (if a = x then v else 0) + u x" in exI) unfolding scaleR_left_distrib and setsum_addf and setsum_delta''[OF fin] and setsum_delta'[OF fin]
+    unfolding setsum_clauses(2)[OF assms] using uv and uv(2)[THEN bspec[where x=a]] and `a\<in>s` by auto
+next
+  assume ?rhs then obtain v u where uv:"v\<ge>0" "\<forall>x\<in>s. 0 \<le> u x" "setsum u s = w - v" "(\<Sum>x\<in>s. u x *\<^sub>R x) = y - v *\<^sub>R a" by auto
+  moreover assume "a\<notin>s" moreover have "(\<Sum>x\<in>s. if a = x then v else u x) = setsum u s" "(\<Sum>x\<in>s. (if a = x then v else u x) *\<^sub>R x) = (\<Sum>x\<in>s. u x *\<^sub>R x)"
+    apply(rule_tac setsum_cong2) defer apply(rule_tac setsum_cong2) using `a\<notin>s` by auto
+  ultimately show ?lhs apply(rule_tac x="\<lambda>x. if a = x then v else u x" in exI)  unfolding setsum_clauses(2)[OF assms] by auto
+qed
+
+subsection {* Hence some special cases. *}
+
+lemma convex_hull_2:
+  "convex hull {a,b} = {u *\<^sub>R a + v *\<^sub>R b | u v. 0 \<le> u \<and> 0 \<le> v \<and> u + v = 1}"
+proof- have *:"\<And>u. (\<forall>x\<in>{a, b}. 0 \<le> u x) \<longleftrightarrow> 0 \<le> u a \<and> 0 \<le> u b" by auto have **:"finite {b}" by auto
+show ?thesis apply(simp add: convex_hull_finite) unfolding convex_hull_finite_step[OF **, of a 1, unfolded * conj_assoc]
+  apply auto apply(rule_tac x=v in exI) apply(rule_tac x="1 - v" in exI) apply simp
+  apply(rule_tac x=u in exI) apply simp apply(rule_tac x="\<lambda>x. v" in exI) by simp qed
+
+lemma convex_hull_2_alt: "convex hull {a,b} = {a + u *\<^sub>R (b - a) | u.  0 \<le> u \<and> u \<le> 1}"
+  unfolding convex_hull_2 unfolding Collect_def 
+proof(rule ext) have *:"\<And>x y ::real. x + y = 1 \<longleftrightarrow> x = 1 - y" by auto
+  fix x show "(\<exists>v u. x = v *\<^sub>R a + u *\<^sub>R b \<and> 0 \<le> v \<and> 0 \<le> u \<and> v + u = 1) = (\<exists>u. x = a + u *\<^sub>R (b - a) \<and> 0 \<le> u \<and> u \<le> 1)"
+    unfolding * apply auto apply(rule_tac[!] x=u in exI) by (auto simp add: algebra_simps) qed
+
+lemma convex_hull_3:
+  "convex hull {a,b,c} = { u *\<^sub>R a + v *\<^sub>R b + w *\<^sub>R c | u v w. 0 \<le> u \<and> 0 \<le> v \<and> 0 \<le> w \<and> u + v + w = 1}"
+proof-
+  have fin:"finite {a,b,c}" "finite {b,c}" "finite {c}" by auto
+  have *:"\<And>x y z ::real. x + y + z = 1 \<longleftrightarrow> x = 1 - y - z"
+         "\<And>x y z ::real^'n. x + y + z = 1 \<longleftrightarrow> x = 1 - y - z" by (auto simp add: ring_simps)
+  show ?thesis unfolding convex_hull_finite[OF fin(1)] and Collect_def and convex_hull_finite_step[OF fin(2)] and *
+    unfolding convex_hull_finite_step[OF fin(3)] apply(rule ext) apply simp apply auto
+    apply(rule_tac x=va in exI) apply (rule_tac x="u c" in exI) apply simp
+    apply(rule_tac x="1 - v - w" in exI) apply simp apply(rule_tac x=v in exI) apply simp apply(rule_tac x="\<lambda>x. w" in exI) by simp qed
+
+lemma convex_hull_3_alt:
+  "convex hull {a,b,c} = {a + u *\<^sub>R (b - a) + v *\<^sub>R (c - a) | u v.  0 \<le> u \<and> 0 \<le> v \<and> u + v \<le> 1}"
+proof- have *:"\<And>x y z ::real. x + y + z = 1 \<longleftrightarrow> x = 1 - y - z" by auto
+  show ?thesis unfolding convex_hull_3 apply (auto simp add: *) apply(rule_tac x=v in exI) apply(rule_tac x=w in exI) apply (simp add: algebra_simps)
+    apply(rule_tac x=u in exI) apply(rule_tac x=v in exI) by (simp add: algebra_simps) qed
+
+subsection {* Relations among closure notions and corresponding hulls. *}
+
+text {* TODO: Generalize linear algebra concepts defined in @{text
+Euclidean_Space.thy} so that we can generalize these lemmas. *}
+
+lemma subspace_imp_affine:
+  fixes s :: "(real ^ _) set" shows "subspace s \<Longrightarrow> affine s"
+  unfolding subspace_def affine_def smult_conv_scaleR by auto
+
+lemma affine_imp_convex: "affine s \<Longrightarrow> convex s"
+  unfolding affine_def convex_def by auto
+
+lemma subspace_imp_convex:
+  fixes s :: "(real ^ _) set" shows "subspace s \<Longrightarrow> convex s"
+  using subspace_imp_affine affine_imp_convex by auto
+
+lemma affine_hull_subset_span:
+  fixes s :: "(real ^ _) set" shows "(affine hull s) \<subseteq> (span s)"
+  unfolding span_def apply(rule hull_antimono) unfolding subset_eq Ball_def mem_def
+  using subspace_imp_affine  by auto
+
+lemma convex_hull_subset_span:
+  fixes s :: "(real ^ _) set" shows "(convex hull s) \<subseteq> (span s)"
+  unfolding span_def apply(rule hull_antimono) unfolding subset_eq Ball_def mem_def
+  using subspace_imp_convex by auto
+
+lemma convex_hull_subset_affine_hull: "(convex hull s) \<subseteq> (affine hull s)"
+  unfolding span_def apply(rule hull_antimono) unfolding subset_eq Ball_def mem_def
+  using affine_imp_convex by auto
+
+lemma affine_dependent_imp_dependent:
+  fixes s :: "(real ^ _) set" shows "affine_dependent s \<Longrightarrow> dependent s"
+  unfolding affine_dependent_def dependent_def 
+  using affine_hull_subset_span by auto
+
+lemma dependent_imp_affine_dependent:
+  fixes s :: "(real ^ _) set"
+  assumes "dependent {x - a| x . x \<in> s}" "a \<notin> s"
+  shows "affine_dependent (insert a s)"
+proof-
+  from assms(1)[unfolded dependent_explicit smult_conv_scaleR] obtain S u v 
+    where obt:"finite S" "S \<subseteq> {x - a |x. x \<in> s}" "v\<in>S" "u v  \<noteq> 0" "(\<Sum>v\<in>S. u v *\<^sub>R v) = 0" by auto
+  def t \<equiv> "(\<lambda>x. x + a) ` S"
+
+  have inj:"inj_on (\<lambda>x. x + a) S" unfolding inj_on_def by auto
+  have "0\<notin>S" using obt(2) assms(2) unfolding subset_eq by auto
+  have fin:"finite t" and  "t\<subseteq>s" unfolding t_def using obt(1,2) by auto 
+
+  hence "finite (insert a t)" and "insert a t \<subseteq> insert a s" by auto 
+  moreover have *:"\<And>P Q. (\<Sum>x\<in>t. (if x = a then P x else Q x)) = (\<Sum>x\<in>t. Q x)"
+    apply(rule setsum_cong2) using `a\<notin>s` `t\<subseteq>s` by auto
+  have "(\<Sum>x\<in>insert a t. if x = a then - (\<Sum>x\<in>t. u (x - a)) else u (x - a)) = 0"
+    unfolding setsum_clauses(2)[OF fin] using `a\<notin>s` `t\<subseteq>s` apply auto unfolding * by auto
+  moreover have "\<exists>v\<in>insert a t. (if v = a then - (\<Sum>x\<in>t. u (x - a)) else u (v - a)) \<noteq> 0"
+    apply(rule_tac x="v + a" in bexI) using obt(3,4) and `0\<notin>S` unfolding t_def by auto
+  moreover have *:"\<And>P Q. (\<Sum>x\<in>t. (if x = a then P x else Q x) *\<^sub>R x) = (\<Sum>x\<in>t. Q x *\<^sub>R x)"
+    apply(rule setsum_cong2) using `a\<notin>s` `t\<subseteq>s` by auto
+  have "(\<Sum>x\<in>t. u (x - a)) *\<^sub>R a = (\<Sum>v\<in>t. u (v - a) *\<^sub>R v)" 
+    unfolding scaleR_left.setsum unfolding t_def and setsum_reindex[OF inj] and o_def
+    using obt(5) by (auto simp add: setsum_addf scaleR_right_distrib)
+  hence "(\<Sum>v\<in>insert a t. (if v = a then - (\<Sum>x\<in>t. u (x - a)) else u (v - a)) *\<^sub>R v) = 0"
+    unfolding setsum_clauses(2)[OF fin] using `a\<notin>s` `t\<subseteq>s` by (auto simp add: *  vector_smult_lneg) 
+  ultimately show ?thesis unfolding affine_dependent_explicit
+    apply(rule_tac x="insert a t" in exI) by auto 
+qed
+
+lemma convex_cone:
+  "convex s \<and> cone s \<longleftrightarrow> (\<forall>x\<in>s. \<forall>y\<in>s. (x + y) \<in> s) \<and> (\<forall>x\<in>s. \<forall>c\<ge>0. (c *\<^sub>R x) \<in> s)" (is "?lhs = ?rhs")
+proof-
+  { fix x y assume "x\<in>s" "y\<in>s" and ?lhs
+    hence "2 *\<^sub>R x \<in>s" "2 *\<^sub>R y \<in> s" unfolding cone_def by auto
+    hence "x + y \<in> s" using `?lhs`[unfolded convex_def, THEN conjunct1]
+      apply(erule_tac x="2*\<^sub>R x" in ballE) apply(erule_tac x="2*\<^sub>R y" in ballE)
+      apply(erule_tac x="1/2" in allE) apply simp apply(erule_tac x="1/2" in allE) by auto  }
+  thus ?thesis unfolding convex_def cone_def by blast
+qed
+
+lemma affine_dependent_biggerset: fixes s::"(real^'n::finite) set"
+  assumes "finite s" "card s \<ge> CARD('n) + 2"
+  shows "affine_dependent s"
+proof-
+  have "s\<noteq>{}" using assms by auto then obtain a where "a\<in>s" by auto
+  have *:"{x - a |x. x \<in> s - {a}} = (\<lambda>x. x - a) ` (s - {a})" by auto
+  have "card {x - a |x. x \<in> s - {a}} = card (s - {a})" unfolding * 
+    apply(rule card_image) unfolding inj_on_def by auto
+  also have "\<dots> > CARD('n)" using assms(2)
+    unfolding card_Diff_singleton[OF assms(1) `a\<in>s`] by auto
+  finally show ?thesis apply(subst insert_Diff[OF `a\<in>s`, THEN sym])
+    apply(rule dependent_imp_affine_dependent)
+    apply(rule dependent_biggerset) by auto qed
+
+lemma affine_dependent_biggerset_general:
+  assumes "finite (s::(real^'n::finite) set)" "card s \<ge> dim s + 2"
+  shows "affine_dependent s"
+proof-
+  from assms(2) have "s \<noteq> {}" by auto
+  then obtain a where "a\<in>s" by auto
+  have *:"{x - a |x. x \<in> s - {a}} = (\<lambda>x. x - a) ` (s - {a})" by auto
+  have **:"card {x - a |x. x \<in> s - {a}} = card (s - {a})" unfolding * 
+    apply(rule card_image) unfolding inj_on_def by auto
+  have "dim {x - a |x. x \<in> s - {a}} \<le> dim s"
+    apply(rule subset_le_dim) unfolding subset_eq
+    using `a\<in>s` by (auto simp add:span_superset span_sub)
+  also have "\<dots> < dim s + 1" by auto
+  also have "\<dots> \<le> card (s - {a})" using assms
+    using card_Diff_singleton[OF assms(1) `a\<in>s`] by auto
+  finally show ?thesis apply(subst insert_Diff[OF `a\<in>s`, THEN sym])
+    apply(rule dependent_imp_affine_dependent) apply(rule dependent_biggerset_general) unfolding ** by auto qed
+
+subsection {* Caratheodory's theorem. *}
+
+lemma convex_hull_caratheodory: fixes p::"(real^'n::finite) set"
+  shows "convex hull p = {y. \<exists>s u. finite s \<and> s \<subseteq> p \<and> card s \<le> CARD('n) + 1 \<and>
+  (\<forall>x\<in>s. 0 \<le> u x) \<and> setsum u s = 1 \<and> setsum (\<lambda>v. u v *\<^sub>R v) s = y}"
+  unfolding convex_hull_explicit expand_set_eq mem_Collect_eq
+proof(rule,rule)
+  fix y let ?P = "\<lambda>n. \<exists>s u. finite s \<and> card s = n \<and> s \<subseteq> p \<and> (\<forall>x\<in>s. 0 \<le> u x) \<and> setsum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = y"
+  assume "\<exists>s u. finite s \<and> s \<subseteq> p \<and> (\<forall>x\<in>s. 0 \<le> u x) \<and> setsum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = y"
+  then obtain N where "?P N" by auto
+  hence "\<exists>n\<le>N. (\<forall>k<n. \<not> ?P k) \<and> ?P n" apply(rule_tac ex_least_nat_le) by auto
+  then obtain n where "?P n" and smallest:"\<forall>k<n. \<not> ?P k" by blast
+  then obtain s u where obt:"finite s" "card s = n" "s\<subseteq>p" "\<forall>x\<in>s. 0 \<le> u x" "setsum u s = 1"  "(\<Sum>v\<in>s. u v *\<^sub>R v) = y" by auto
+
+  have "card s \<le> CARD('n) + 1" proof(rule ccontr, simp only: not_le)
+    assume "CARD('n) + 1 < card s"
+    hence "affine_dependent s" using affine_dependent_biggerset[OF obt(1)] by auto
+    then obtain w v where wv:"setsum w s = 0" "v\<in>s" "w v \<noteq> 0" "(\<Sum>v\<in>s. w v *\<^sub>R v) = 0"
+      using affine_dependent_explicit_finite[OF obt(1)] by auto
+    def i \<equiv> "(\<lambda>v. (u v) / (- w v)) ` {v\<in>s. w v < 0}"  def t \<equiv> "Min i"
+    have "\<exists>x\<in>s. w x < 0" proof(rule ccontr, simp add: not_less)
+      assume as:"\<forall>x\<in>s. 0 \<le> w x"
+      hence "setsum w (s - {v}) \<ge> 0" apply(rule_tac setsum_nonneg) by auto
+      hence "setsum w s > 0" unfolding setsum_diff1'[OF obt(1) `v\<in>s`]
+        using as[THEN bspec[where x=v]] and `v\<in>s` using `w v \<noteq> 0` by auto
+      thus False using wv(1) by auto
+    qed hence "i\<noteq>{}" unfolding i_def by auto
+
+    hence "t \<ge> 0" using Min_ge_iff[of i 0 ] and obt(1) unfolding t_def i_def
+      using obt(4)[unfolded le_less] apply auto unfolding divide_le_0_iff by auto 
+    have t:"\<forall>v\<in>s. u v + t * w v \<ge> 0" proof
+      fix v assume "v\<in>s" hence v:"0\<le>u v" using obt(4)[THEN bspec[where x=v]] by auto
+      show"0 \<le> u v + t * w v" proof(cases "w v < 0")
+        case False thus ?thesis apply(rule_tac add_nonneg_nonneg) 
+          using v apply simp apply(rule mult_nonneg_nonneg) using `t\<ge>0` by auto next
+        case True hence "t \<le> u v / (- w v)" using `v\<in>s`
+          unfolding t_def i_def apply(rule_tac Min_le) using obt(1) by auto 
+        thus ?thesis unfolding real_0_le_add_iff
+          using pos_le_divide_eq[OF True[unfolded neg_0_less_iff_less[THEN sym]]] by auto
+      qed qed
+
+    obtain a where "a\<in>s" and "t = (\<lambda>v. (u v) / (- w v)) a" and "w a < 0"
+      using Min_in[OF _ `i\<noteq>{}`] and obt(1) unfolding i_def t_def by auto
+    hence a:"a\<in>s" "u a + t * w a = 0" by auto
+    have *:"\<And>f. setsum f (s - {a}) = setsum f s - ((f a)::'a::ring)" unfolding setsum_diff1'[OF obt(1) `a\<in>s`] by auto 
+    have "(\<Sum>v\<in>s. u v + t * w v) = 1"
+      unfolding setsum_addf wv(1) setsum_right_distrib[THEN sym] obt(5) by auto
+    moreover have "(\<Sum>v\<in>s. u v *\<^sub>R v + (t * w v) *\<^sub>R v) - (u a *\<^sub>R a + (t * w a) *\<^sub>R a) = y" 
+      unfolding setsum_addf obt(6) scaleR_scaleR[THEN sym] scaleR_right.setsum [symmetric] wv(4)
+      using a(2) [THEN eq_neg_iff_add_eq_0 [THEN iffD2]]
+      by (simp add: vector_smult_lneg)
+    ultimately have "?P (n - 1)" apply(rule_tac x="(s - {a})" in exI)
+      apply(rule_tac x="\<lambda>v. u v + t * w v" in exI) using obt(1-3) and t and a by (auto simp add: * scaleR_left_distrib)
+    thus False using smallest[THEN spec[where x="n - 1"]] by auto qed
+  thus "\<exists>s u. finite s \<and> s \<subseteq> p \<and> card s \<le> CARD('n) + 1
+    \<and> (\<forall>x\<in>s. 0 \<le> u x) \<and> setsum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = y" using obt by auto
+qed auto
+
+lemma caratheodory:
+ "convex hull p = {x::real^'n::finite. \<exists>s. finite s \<and> s \<subseteq> p \<and>
+      card s \<le> CARD('n) + 1 \<and> x \<in> convex hull s}"
+  unfolding expand_set_eq apply(rule, rule) unfolding mem_Collect_eq proof-
+  fix x assume "x \<in> convex hull p"
+  then obtain s u where "finite s" "s \<subseteq> p" "card s \<le> CARD('n) + 1"
+     "\<forall>x\<in>s. 0 \<le> u x" "setsum u s = 1" "(\<Sum>v\<in>s. u v *\<^sub>R v) = x"unfolding convex_hull_caratheodory by auto
+  thus "\<exists>s. finite s \<and> s \<subseteq> p \<and> card s \<le> CARD('n) + 1 \<and> x \<in> convex hull s"
+    apply(rule_tac x=s in exI) using hull_subset[of s convex]
+  using convex_convex_hull[unfolded convex_explicit, of s, THEN spec[where x=s], THEN spec[where x=u]] by auto
+next
+  fix x assume "\<exists>s. finite s \<and> s \<subseteq> p \<and> card s \<le> CARD('n) + 1 \<and> x \<in> convex hull s"
+  then obtain s where "finite s" "s \<subseteq> p" "card s \<le> CARD('n) + 1" "x \<in> convex hull s" by auto
+  thus "x \<in> convex hull p" using hull_mono[OF `s\<subseteq>p`] by auto
+qed
+
+subsection {* Openness and compactness are preserved by convex hull operation. *}
+
+lemma open_convex_hull:
+  fixes s :: "'a::real_normed_vector set"
+  assumes "open s"
+  shows "open(convex hull s)"
+  unfolding open_contains_cball convex_hull_explicit unfolding mem_Collect_eq ball_simps(10) 
+proof(rule, rule) fix a
+  assume "\<exists>sa u. finite sa \<and> sa \<subseteq> s \<and> (\<forall>x\<in>sa. 0 \<le> u x) \<and> setsum u sa = 1 \<and> (\<Sum>v\<in>sa. u v *\<^sub>R v) = a"
+  then obtain t u where obt:"finite t" "t\<subseteq>s" "\<forall>x\<in>t. 0 \<le> u x" "setsum u t = 1" "(\<Sum>v\<in>t. u v *\<^sub>R v) = a" by auto
+
+  from assms[unfolded open_contains_cball] obtain b where b:"\<forall>x\<in>s. 0 < b x \<and> cball x (b x) \<subseteq> s"
+    using bchoice[of s "\<lambda>x e. e>0 \<and> cball x e \<subseteq> s"] by auto
+  have "b ` t\<noteq>{}" unfolding i_def using obt by auto  def i \<equiv> "b ` t"
+
+  show "\<exists>e>0. cball a e \<subseteq> {y. \<exists>sa u. finite sa \<and> sa \<subseteq> s \<and> (\<forall>x\<in>sa. 0 \<le> u x) \<and> setsum u sa = 1 \<and> (\<Sum>v\<in>sa. u v *\<^sub>R v) = y}"
+    apply(rule_tac x="Min i" in exI) unfolding subset_eq apply rule defer apply rule unfolding mem_Collect_eq
+  proof-
+    show "0 < Min i" unfolding i_def and Min_gr_iff[OF finite_imageI[OF obt(1)] `b \` t\<noteq>{}`]
+      using b apply simp apply rule apply(erule_tac x=x in ballE) using `t\<subseteq>s` by auto
+  next  fix y assume "y \<in> cball a (Min i)"
+    hence y:"norm (a - y) \<le> Min i" unfolding dist_norm[THEN sym] by auto
+    { fix x assume "x\<in>t"
+      hence "Min i \<le> b x" unfolding i_def apply(rule_tac Min_le) using obt(1) by auto
+      hence "x + (y - a) \<in> cball x (b x)" using y unfolding mem_cball dist_norm by auto
+      moreover from `x\<in>t` have "x\<in>s" using obt(2) by auto
+      ultimately have "x + (y - a) \<in> s" using y and b[THEN bspec[where x=x]] unfolding subset_eq by auto }
+    moreover
+    have *:"inj_on (\<lambda>v. v + (y - a)) t" unfolding inj_on_def by auto
+    have "(\<Sum>v\<in>(\<lambda>v. v + (y - a)) ` t. u (v - (y - a))) = 1"
+      unfolding setsum_reindex[OF *] o_def using obt(4) by auto
+    moreover have "(\<Sum>v\<in>(\<lambda>v. v + (y - a)) ` t. u (v - (y - a)) *\<^sub>R v) = y"
+      unfolding setsum_reindex[OF *] o_def using obt(4,5)
+      by (simp add: setsum_addf setsum_subtractf scaleR_left.setsum[THEN sym] scaleR_right_distrib)
+    ultimately show "\<exists>sa u. finite sa \<and> (\<forall>x\<in>sa. x \<in> s) \<and> (\<forall>x\<in>sa. 0 \<le> u x) \<and> setsum u sa = 1 \<and> (\<Sum>v\<in>sa. u v *\<^sub>R v) = y"
+      apply(rule_tac x="(\<lambda>v. v + (y - a)) ` t" in exI) apply(rule_tac x="\<lambda>v. u (v - (y - a))" in exI)
+      using obt(1, 3) by auto
+  qed
+qed
+
+lemma open_dest_vec1_vimage: "open S \<Longrightarrow> open (dest_vec1 -` S)"
+unfolding open_vector_def all_1
+by (auto simp add: dest_vec1_def)
+
+lemma tendsto_dest_vec1 [tendsto_intros]:
+  "(f ---> l) net \<Longrightarrow> ((\<lambda>x. dest_vec1 (f x)) ---> dest_vec1 l) net"
+  unfolding tendsto_def
+  apply clarify
+  apply (drule_tac x="dest_vec1 -` S" in spec)
+  apply (simp add: open_dest_vec1_vimage)
+  done
+
+lemma continuous_dest_vec1: "continuous net f \<Longrightarrow> continuous net (\<lambda>x. dest_vec1 (f x))"
+  unfolding continuous_def by (rule tendsto_dest_vec1)
+
+(* TODO: move *)
+lemma compact_real_interval:
+  fixes a b :: real shows "compact {a..b}"
+proof -
+  have "continuous_on {vec1 a .. vec1 b} dest_vec1"
+    unfolding continuous_on
+    by (simp add: tendsto_dest_vec1 Lim_at_within Lim_ident_at)
+  moreover have "compact {vec1 a .. vec1 b}" by (rule compact_interval)
+  ultimately have "compact (dest_vec1 ` {vec1 a .. vec1 b})"
+    by (rule compact_continuous_image)
+  also have "dest_vec1 ` {vec1 a .. vec1 b} = {a..b}"
+    by (auto simp add: image_def Bex_def exists_vec1)
+  finally show ?thesis .
+qed
+
+lemma compact_convex_combinations:
+  fixes s t :: "'a::real_normed_vector set"
+  assumes "compact s" "compact t"
+  shows "compact { (1 - u) *\<^sub>R x + u *\<^sub>R y | x y u. 0 \<le> u \<and> u \<le> 1 \<and> x \<in> s \<and> y \<in> t}"
+proof-
+  let ?X = "{0..1} \<times> s \<times> t"
+  let ?h = "(\<lambda>z. (1 - fst z) *\<^sub>R fst (snd z) + fst z *\<^sub>R snd (snd z))"
+  have *:"{ (1 - u) *\<^sub>R x + u *\<^sub>R y | x y u. 0 \<le> u \<and> u \<le> 1 \<and> x \<in> s \<and> y \<in> t} = ?h ` ?X"
+    apply(rule set_ext) unfolding image_iff mem_Collect_eq
+    apply rule apply auto
+    apply (rule_tac x=u in rev_bexI, simp)
+    apply (erule rev_bexI, erule rev_bexI, simp)
+    by auto
+  have "continuous_on ({0..1} \<times> s \<times> t)
+     (\<lambda>z. (1 - fst z) *\<^sub>R fst (snd z) + fst z *\<^sub>R snd (snd z))"
+    unfolding continuous_on by (rule ballI) (intro tendsto_intros)
+  thus ?thesis unfolding *
+    apply (rule compact_continuous_image)
+    apply (intro compact_Times compact_real_interval assms)
+    done
+qed
+
+lemma compact_convex_hull: fixes s::"(real^'n::finite) set"
+  assumes "compact s"  shows "compact(convex hull s)"
+proof(cases "s={}")
+  case True thus ?thesis using compact_empty by simp
+next
+  case False then obtain w where "w\<in>s" by auto
+  show ?thesis unfolding caratheodory[of s]
+  proof(induct "CARD('n) + 1")
+    have *:"{x.\<exists>sa. finite sa \<and> sa \<subseteq> s \<and> card sa \<le> 0 \<and> x \<in> convex hull sa} = {}" 
+      using compact_empty by (auto simp add: convex_hull_empty)
+    case 0 thus ?case unfolding * by simp
+  next
+    case (Suc n)
+    show ?case proof(cases "n=0")
+      case True have "{x. \<exists>t. finite t \<and> t \<subseteq> s \<and> card t \<le> Suc n \<and> x \<in> convex hull t} = s"
+        unfolding expand_set_eq and mem_Collect_eq proof(rule, rule)
+        fix x assume "\<exists>t. finite t \<and> t \<subseteq> s \<and> card t \<le> Suc n \<and> x \<in> convex hull t"
+        then obtain t where t:"finite t" "t \<subseteq> s" "card t \<le> Suc n" "x \<in> convex hull t" by auto
+        show "x\<in>s" proof(cases "card t = 0")
+          case True thus ?thesis using t(4) unfolding card_0_eq[OF t(1)] by(simp add: convex_hull_empty)
+        next
+          case False hence "card t = Suc 0" using t(3) `n=0` by auto
+          then obtain a where "t = {a}" unfolding card_Suc_eq by auto
+          thus ?thesis using t(2,4) by (simp add: convex_hull_singleton)
+        qed
+      next
+        fix x assume "x\<in>s"
+        thus "\<exists>t. finite t \<and> t \<subseteq> s \<and> card t \<le> Suc n \<and> x \<in> convex hull t"
+          apply(rule_tac x="{x}" in exI) unfolding convex_hull_singleton by auto 
+      qed thus ?thesis using assms by simp
+    next
+      case False have "{x. \<exists>t. finite t \<and> t \<subseteq> s \<and> card t \<le> Suc n \<and> x \<in> convex hull t} =
+        { (1 - u) *\<^sub>R x + u *\<^sub>R y | x y u. 
+        0 \<le> u \<and> u \<le> 1 \<and> x \<in> s \<and> y \<in> {x. \<exists>t. finite t \<and> t \<subseteq> s \<and> card t \<le> n \<and> x \<in> convex hull t}}"
+        unfolding expand_set_eq and mem_Collect_eq proof(rule,rule)
+        fix x assume "\<exists>u v c. x = (1 - c) *\<^sub>R u + c *\<^sub>R v \<and>
+          0 \<le> c \<and> c \<le> 1 \<and> u \<in> s \<and> (\<exists>t. finite t \<and> t \<subseteq> s \<and> card t \<le> n \<and> v \<in> convex hull t)"
+        then obtain u v c t where obt:"x = (1 - c) *\<^sub>R u + c *\<^sub>R v"
+          "0 \<le> c \<and> c \<le> 1" "u \<in> s" "finite t" "t \<subseteq> s" "card t \<le> n"  "v \<in> convex hull t" by auto
+        moreover have "(1 - c) *\<^sub>R u + c *\<^sub>R v \<in> convex hull insert u t"
+          apply(rule mem_convex) using obt(2) and convex_convex_hull and hull_subset[of "insert u t" convex]
+          using obt(7) and hull_mono[of t "insert u t"] by auto
+        ultimately show "\<exists>t. finite t \<and> t \<subseteq> s \<and> card t \<le> Suc n \<and> x \<in> convex hull t"
+          apply(rule_tac x="insert u t" in exI) by (auto simp add: card_insert_if)
+      next
+        fix x assume "\<exists>t. finite t \<and> t \<subseteq> s \<and> card t \<le> Suc n \<and> x \<in> convex hull t"
+        then obtain t where t:"finite t" "t \<subseteq> s" "card t \<le> Suc n" "x \<in> convex hull t" by auto
+        let ?P = "\<exists>u v c. x = (1 - c) *\<^sub>R u + c *\<^sub>R v \<and>
+          0 \<le> c \<and> c \<le> 1 \<and> u \<in> s \<and> (\<exists>t. finite t \<and> t \<subseteq> s \<and> card t \<le> n \<and> v \<in> convex hull t)"
+        show ?P proof(cases "card t = Suc n")
+          case False hence "card t \<le> n" using t(3) by auto
+          thus ?P apply(rule_tac x=w in exI, rule_tac x=x in exI, rule_tac x=1 in exI) using `w\<in>s` and t
+            by(auto intro!: exI[where x=t])
+        next
+          case True then obtain a u where au:"t = insert a u" "a\<notin>u" apply(drule_tac card_eq_SucD) by auto
+          show ?P proof(cases "u={}")
+            case True hence "x=a" using t(4)[unfolded au] by auto
+            show ?P unfolding `x=a` apply(rule_tac x=a in exI, rule_tac x=a in exI, rule_tac x=1 in exI)
+              using t and `n\<noteq>0` unfolding au by(auto intro!: exI[where x="{a}"] simp add: convex_hull_singleton)
+          next
+            case False obtain ux vx b where obt:"ux\<ge>0" "vx\<ge>0" "ux + vx = 1" "b \<in> convex hull u" "x = ux *\<^sub>R a + vx *\<^sub>R b"
+              using t(4)[unfolded au convex_hull_insert[OF False]] by auto
+            have *:"1 - vx = ux" using obt(3) by auto
+            show ?P apply(rule_tac x=a in exI, rule_tac x=b in exI, rule_tac x=vx in exI)
+              using obt and t(1-3) unfolding au and * using card_insert_disjoint[OF _ au(2)]
+              by(auto intro!: exI[where x=u])
+          qed
+        qed
+      qed
+      thus ?thesis using compact_convex_combinations[OF assms Suc] by simp 
+    qed
+  qed 
+qed
+
+lemma finite_imp_compact_convex_hull:
+  fixes s :: "(real ^ _) set"
+  shows "finite s \<Longrightarrow> compact(convex hull s)"
+  apply(drule finite_imp_compact, drule compact_convex_hull) by assumption
+
+subsection {* Extremal points of a simplex are some vertices. *}
+
+lemma dist_increases_online:
+  fixes a b d :: "'a::real_inner"
+  assumes "d \<noteq> 0"
+  shows "dist a (b + d) > dist a b \<or> dist a (b - d) > dist a b"
+proof(cases "inner a d - inner b d > 0")
+  case True hence "0 < inner d d + (inner a d * 2 - inner b d * 2)" 
+    apply(rule_tac add_pos_pos) using assms by auto
+  thus ?thesis apply(rule_tac disjI2) unfolding dist_norm and norm_eq_sqrt_inner and real_sqrt_less_iff
+    by (simp add: algebra_simps inner_commute)
+next
+  case False hence "0 < inner d d + (inner b d * 2 - inner a d * 2)" 
+    apply(rule_tac add_pos_nonneg) using assms by auto
+  thus ?thesis apply(rule_tac disjI1) unfolding dist_norm and norm_eq_sqrt_inner and real_sqrt_less_iff
+    by (simp add: algebra_simps inner_commute)
+qed
+
+lemma norm_increases_online:
+  fixes d :: "'a::real_inner"
+  shows "d \<noteq> 0 \<Longrightarrow> norm(a + d) > norm a \<or> norm(a - d) > norm a"
+  using dist_increases_online[of d a 0] unfolding dist_norm by auto
+
+lemma simplex_furthest_lt:
+  fixes s::"'a::real_inner set" assumes "finite s"
+  shows "\<forall>x \<in> (convex hull s).  x \<notin> s \<longrightarrow> (\<exists>y\<in>(convex hull s). norm(x - a) < norm(y - a))"
+proof(induct_tac rule: finite_induct[of s])
+  fix x s assume as:"finite s" "x\<notin>s" "\<forall>x\<in>convex hull s. x \<notin> s \<longrightarrow> (\<exists>y\<in>convex hull s. norm (x - a) < norm (y - a))"
+  show "\<forall>xa\<in>convex hull insert x s. xa \<notin> insert x s \<longrightarrow> (\<exists>y\<in>convex hull insert x s. norm (xa - a) < norm (y - a))"
+  proof(rule,rule,cases "s = {}")
+    case False fix y assume y:"y \<in> convex hull insert x s" "y \<notin> insert x s"
+    obtain u v b where obt:"u\<ge>0" "v\<ge>0" "u + v = 1" "b \<in> convex hull s" "y = u *\<^sub>R x + v *\<^sub>R b"
+      using y(1)[unfolded convex_hull_insert[OF False]] by auto
+    show "\<exists>z\<in>convex hull insert x s. norm (y - a) < norm (z - a)"
+    proof(cases "y\<in>convex hull s")
+      case True then obtain z where "z\<in>convex hull s" "norm (y - a) < norm (z - a)"
+        using as(3)[THEN bspec[where x=y]] and y(2) by auto
+      thus ?thesis apply(rule_tac x=z in bexI) unfolding convex_hull_insert[OF False] by auto
+    next
+      case False show ?thesis  using obt(3) proof(cases "u=0", case_tac[!] "v=0")
+        assume "u=0" "v\<noteq>0" hence "y = b" using obt by auto
+        thus ?thesis using False and obt(4) by auto
+      next
+        assume "u\<noteq>0" "v=0" hence "y = x" using obt by auto
+        thus ?thesis using y(2) by auto
+      next
+        assume "u\<noteq>0" "v\<noteq>0"
+        then obtain w where w:"w>0" "w<u" "w<v" using real_lbound_gt_zero[of u v] and obt(1,2) by auto
+        have "x\<noteq>b" proof(rule ccontr) 
+          assume "\<not> x\<noteq>b" hence "y=b" unfolding obt(5)
+            using obt(3) by(auto simp add: scaleR_left_distrib[THEN sym])
+          thus False using obt(4) and False by simp qed
+        hence *:"w *\<^sub>R (x - b) \<noteq> 0" using w(1) by auto
+        show ?thesis using dist_increases_online[OF *, of a y]
+        proof(erule_tac disjE)
+          assume "dist a y < dist a (y + w *\<^sub>R (x - b))"
+          hence "norm (y - a) < norm ((u + w) *\<^sub>R x + (v - w) *\<^sub>R b - a)"
+            unfolding dist_commute[of a] unfolding dist_norm obt(5) by (simp add: algebra_simps)
+          moreover have "(u + w) *\<^sub>R x + (v - w) *\<^sub>R b \<in> convex hull insert x s"
+            unfolding convex_hull_insert[OF `s\<noteq>{}`] and mem_Collect_eq
+            apply(rule_tac x="u + w" in exI) apply rule defer 
+            apply(rule_tac x="v - w" in exI) using `u\<ge>0` and w and obt(3,4) by auto
+          ultimately show ?thesis by auto
+        next
+          assume "dist a y < dist a (y - w *\<^sub>R (x - b))"
+          hence "norm (y - a) < norm ((u - w) *\<^sub>R x + (v + w) *\<^sub>R b - a)"
+            unfolding dist_commute[of a] unfolding dist_norm obt(5) by (simp add: algebra_simps)
+          moreover have "(u - w) *\<^sub>R x + (v + w) *\<^sub>R b \<in> convex hull insert x s"
+            unfolding convex_hull_insert[OF `s\<noteq>{}`] and mem_Collect_eq
+            apply(rule_tac x="u - w" in exI) apply rule defer 
+            apply(rule_tac x="v + w" in exI) using `u\<ge>0` and w and obt(3,4) by auto
+          ultimately show ?thesis by auto
+        qed
+      qed auto
+    qed
+  qed auto
+qed (auto simp add: assms)
+
+lemma simplex_furthest_le:
+  fixes s :: "(real ^ _) set"
+  assumes "finite s" "s \<noteq> {}"
+  shows "\<exists>y\<in>s. \<forall>x\<in>(convex hull s). norm(x - a) \<le> norm(y - a)"
+proof-
+  have "convex hull s \<noteq> {}" using hull_subset[of s convex] and assms(2) by auto
+  then obtain x where x:"x\<in>convex hull s" "\<forall>y\<in>convex hull s. norm (y - a) \<le> norm (x - a)"
+    using distance_attains_sup[OF finite_imp_compact_convex_hull[OF assms(1)], of a]
+    unfolding dist_commute[of a] unfolding dist_norm by auto
+  thus ?thesis proof(cases "x\<in>s")
+    case False then obtain y where "y\<in>convex hull s" "norm (x - a) < norm (y - a)"
+      using simplex_furthest_lt[OF assms(1), THEN bspec[where x=x]] and x(1) by auto
+    thus ?thesis using x(2)[THEN bspec[where x=y]] by auto
+  qed auto
+qed
+
+lemma simplex_furthest_le_exists:
+  fixes s :: "(real ^ _) set"
+  shows "finite s \<Longrightarrow> (\<forall>x\<in>(convex hull s). \<exists>y\<in>s. norm(x - a) \<le> norm(y - a))"
+  using simplex_furthest_le[of s] by (cases "s={}")auto
+
+lemma simplex_extremal_le:
+  fixes s :: "(real ^ _) set"
+  assumes "finite s" "s \<noteq> {}"
+  shows "\<exists>u\<in>s. \<exists>v\<in>s. \<forall>x\<in>convex hull s. \<forall>y \<in> convex hull s. norm(x - y) \<le> norm(u - v)"
+proof-
+  have "convex hull s \<noteq> {}" using hull_subset[of s convex] and assms(2) by auto
+  then obtain u v where obt:"u\<in>convex hull s" "v\<in>convex hull s"
+    "\<forall>x\<in>convex hull s. \<forall>y\<in>convex hull s. norm (x - y) \<le> norm (u - v)"
+    using compact_sup_maxdistance[OF finite_imp_compact_convex_hull[OF assms(1)]] by auto
+  thus ?thesis proof(cases "u\<notin>s \<or> v\<notin>s", erule_tac disjE)
+    assume "u\<notin>s" then obtain y where "y\<in>convex hull s" "norm (u - v) < norm (y - v)"
+      using simplex_furthest_lt[OF assms(1), THEN bspec[where x=u]] and obt(1) by auto
+    thus ?thesis using obt(3)[THEN bspec[where x=y], THEN bspec[where x=v]] and obt(2) by auto
+  next
+    assume "v\<notin>s" then obtain y where "y\<in>convex hull s" "norm (v - u) < norm (y - u)"
+      using simplex_furthest_lt[OF assms(1), THEN bspec[where x=v]] and obt(2) by auto
+    thus ?thesis using obt(3)[THEN bspec[where x=u], THEN bspec[where x=y]] and obt(1)
+      by (auto simp add: norm_minus_commute)
+  qed auto
+qed 
+
+lemma simplex_extremal_le_exists:
+  fixes s :: "(real ^ _) set"
+  shows "finite s \<Longrightarrow> x \<in> convex hull s \<Longrightarrow> y \<in> convex hull s
+  \<Longrightarrow> (\<exists>u\<in>s. \<exists>v\<in>s. norm(x - y) \<le> norm(u - v))"
+  using convex_hull_empty simplex_extremal_le[of s] by(cases "s={}")auto
+
+subsection {* Closest point of a convex set is unique, with a continuous projection. *}
+
+definition
+  closest_point :: "(real ^ 'n::finite) set \<Rightarrow> real ^ 'n \<Rightarrow> real ^ 'n" where
+ "closest_point s a = (SOME x. x \<in> s \<and> (\<forall>y\<in>s. dist a x \<le> dist a y))"
+
+lemma closest_point_exists:
+  assumes "closed s" "s \<noteq> {}"
+  shows  "closest_point s a \<in> s" "\<forall>y\<in>s. dist a (closest_point s a) \<le> dist a y"
+  unfolding closest_point_def apply(rule_tac[!] someI2_ex) 
+  using distance_attains_inf[OF assms(1,2), of a] by auto
+
+lemma closest_point_in_set:
+  "closed s \<Longrightarrow> s \<noteq> {} \<Longrightarrow> (closest_point s a) \<in> s"
+  by(meson closest_point_exists)
+
+lemma closest_point_le:
+  "closed s \<Longrightarrow> x \<in> s \<Longrightarrow> dist a (closest_point s a) \<le> dist a x"
+  using closest_point_exists[of s] by auto
+
+lemma closest_point_self:
+  assumes "x \<in> s"  shows "closest_point s x = x"
+  unfolding closest_point_def apply(rule some1_equality, rule ex1I[of _ x]) 
+  using assms by auto
+
+lemma closest_point_refl:
+ "closed s \<Longrightarrow> s \<noteq> {} \<Longrightarrow> (closest_point s x = x \<longleftrightarrow> x \<in> s)"
+  using closest_point_in_set[of s x] closest_point_self[of x s] by auto
+
+(* TODO: move *)
+lemma norm_lt: "norm x < norm y \<longleftrightarrow> inner x x < inner y y"
+  unfolding norm_eq_sqrt_inner by simp
+
+(* TODO: move *)
+lemma norm_le: "norm x \<le> norm y \<longleftrightarrow> inner x x \<le> inner y y"
+  unfolding norm_eq_sqrt_inner by simp
+
+lemma closer_points_lemma: fixes y::"real^'n::finite"
+  assumes "inner y z > 0"
+  shows "\<exists>u>0. \<forall>v>0. v \<le> u \<longrightarrow> norm(v *\<^sub>R z - y) < norm y"
+proof- have z:"inner z z > 0" unfolding inner_gt_zero_iff using assms by auto
+  thus ?thesis using assms apply(rule_tac x="inner y z / inner z z" in exI) apply(rule) defer proof(rule+)
+    fix v assume "0<v" "v \<le> inner y z / inner z z"
+    thus "norm (v *\<^sub>R z - y) < norm y" unfolding norm_lt using z and assms
+      by (simp add: field_simps inner_diff inner_commute mult_strict_left_mono[OF _ `0<v`])
+  qed(rule divide_pos_pos, auto) qed
+
+lemma closer_point_lemma:
+  fixes x y z :: "real ^ 'n::finite"
+  assumes "inner (y - x) (z - x) > 0"
+  shows "\<exists>u>0. u \<le> 1 \<and> dist (x + u *\<^sub>R (z - x)) y < dist x y"
+proof- obtain u where "u>0" and u:"\<forall>v>0. v \<le> u \<longrightarrow> norm (v *\<^sub>R (z - x) - (y - x)) < norm (y - x)"
+    using closer_points_lemma[OF assms] by auto
+  show ?thesis apply(rule_tac x="min u 1" in exI) using u[THEN spec[where x="min u 1"]] and `u>0`
+    unfolding dist_norm by(auto simp add: norm_minus_commute field_simps) qed
+
+lemma any_closest_point_dot:
+  fixes s :: "(real ^ _) set"
+  assumes "convex s" "closed s" "x \<in> s" "y \<in> s" "\<forall>z\<in>s. dist a x \<le> dist a z"
+  shows "inner (a - x) (y - x) \<le> 0"
+proof(rule ccontr) assume "\<not> inner (a - x) (y - x) \<le> 0"
+  then obtain u where u:"u>0" "u\<le>1" "dist (x + u *\<^sub>R (y - x)) a < dist x a" using closer_point_lemma[of a x y] by auto
+  let ?z = "(1 - u) *\<^sub>R x + u *\<^sub>R y" have "?z \<in> s" using mem_convex[OF assms(1,3,4), of u] using u by auto
+  thus False using assms(5)[THEN bspec[where x="?z"]] and u(3) by (auto simp add: dist_commute algebra_simps) qed
+
+(* TODO: move *)
+lemma norm_le_square: "norm x \<le> a \<longleftrightarrow> 0 \<le> a \<and> inner x x \<le> a\<twosuperior>"
+proof -
+  have "norm x \<le> a \<longleftrightarrow> 0 \<le> a \<and> norm x \<le> a"
+    using norm_ge_zero [of x] by arith
+  also have "\<dots> \<longleftrightarrow> 0 \<le> a \<and> (norm x)\<twosuperior> \<le> a\<twosuperior>"
+    by (auto intro: power_mono dest: power2_le_imp_le)
+  also have "\<dots> \<longleftrightarrow> 0 \<le> a \<and> inner x x \<le> a\<twosuperior>"
+    unfolding power2_norm_eq_inner ..
+  finally show ?thesis .
+qed
+
+lemma any_closest_point_unique:
+  fixes s :: "(real ^ _) set"
+  assumes "convex s" "closed s" "x \<in> s" "y \<in> s"
+  "\<forall>z\<in>s. dist a x \<le> dist a z" "\<forall>z\<in>s. dist a y \<le> dist a z"
+  shows "x = y" using any_closest_point_dot[OF assms(1-4,5)] and any_closest_point_dot[OF assms(1-2,4,3,6)]
+  unfolding norm_pths(1) and norm_le_square
+  by (auto simp add: algebra_simps)
+
+lemma closest_point_unique:
+  assumes "convex s" "closed s" "x \<in> s" "\<forall>z\<in>s. dist a x \<le> dist a z"
+  shows "x = closest_point s a"
+  using any_closest_point_unique[OF assms(1-3) _ assms(4), of "closest_point s a"] 
+  using closest_point_exists[OF assms(2)] and assms(3) by auto
+
+lemma closest_point_dot:
+  assumes "convex s" "closed s" "x \<in> s"
+  shows "inner (a - closest_point s a) (x - closest_point s a) \<le> 0"
+  apply(rule any_closest_point_dot[OF assms(1,2) _ assms(3)])
+  using closest_point_exists[OF assms(2)] and assms(3) by auto
+
+lemma closest_point_lt:
+  assumes "convex s" "closed s" "x \<in> s" "x \<noteq> closest_point s a"
+  shows "dist a (closest_point s a) < dist a x"
+  apply(rule ccontr) apply(rule_tac notE[OF assms(4)])
+  apply(rule closest_point_unique[OF assms(1-3), of a])
+  using closest_point_le[OF assms(2), of _ a] by fastsimp
+
+lemma closest_point_lipschitz:
+  assumes "convex s" "closed s" "s \<noteq> {}"
+  shows "dist (closest_point s x) (closest_point s y) \<le> dist x y"
+proof-
+  have "inner (x - closest_point s x) (closest_point s y - closest_point s x) \<le> 0"
+       "inner (y - closest_point s y) (closest_point s x - closest_point s y) \<le> 0"
+    apply(rule_tac[!] any_closest_point_dot[OF assms(1-2)])
+    using closest_point_exists[OF assms(2-3)] by auto
+  thus ?thesis unfolding dist_norm and norm_le
+    using inner_ge_zero[of "(x - closest_point s x) - (y - closest_point s y)"]
+    by (simp add: inner_add inner_diff inner_commute) qed
+
+lemma continuous_at_closest_point:
+  assumes "convex s" "closed s" "s \<noteq> {}"
+  shows "continuous (at x) (closest_point s)"
+  unfolding continuous_at_eps_delta 
+  using le_less_trans[OF closest_point_lipschitz[OF assms]] by auto
+
+lemma continuous_on_closest_point:
+  assumes "convex s" "closed s" "s \<noteq> {}"
+  shows "continuous_on t (closest_point s)"
+  apply(rule continuous_at_imp_continuous_on) using continuous_at_closest_point[OF assms] by auto
+
+subsection {* Various point-to-set separating/supporting hyperplane theorems. *}
+
+lemma supporting_hyperplane_closed_point:
+  fixes s :: "(real ^ _) set"
+  assumes "convex s" "closed s" "s \<noteq> {}" "z \<notin> s"
+  shows "\<exists>a b. \<exists>y\<in>s. inner a z < b \<and> (inner a y = b) \<and> (\<forall>x\<in>s. inner a x \<ge> b)"
+proof-
+  from distance_attains_inf[OF assms(2-3)] obtain y where "y\<in>s" and y:"\<forall>x\<in>s. dist z y \<le> dist z x" by auto
+  show ?thesis apply(rule_tac x="y - z" in exI, rule_tac x="inner (y - z) y" in exI, rule_tac x=y in bexI)
+    apply rule defer apply rule defer apply(rule, rule ccontr) using `y\<in>s` proof-
+    show "inner (y - z) z < inner (y - z) y" apply(subst diff_less_iff(1)[THEN sym])
+      unfolding inner_diff_right[THEN sym] and inner_gt_zero_iff using `y\<in>s` `z\<notin>s` by auto
+  next
+    fix x assume "x\<in>s" have *:"\<forall>u. 0 \<le> u \<and> u \<le> 1 \<longrightarrow> dist z y \<le> dist z ((1 - u) *\<^sub>R y + u *\<^sub>R x)"
+      using assms(1)[unfolded convex_alt] and y and `x\<in>s` and `y\<in>s` by auto
+    assume "\<not> inner (y - z) y \<le> inner (y - z) x" then obtain v where
+      "v>0" "v\<le>1" "dist (y + v *\<^sub>R (x - y)) z < dist y z" using closer_point_lemma[of z y x] apply - by (auto simp add: inner_diff)
+    thus False using *[THEN spec[where x=v]] by(auto simp add: dist_commute algebra_simps)
+  qed auto
+qed
+
+lemma separating_hyperplane_closed_point:
+  fixes s :: "(real ^ _) set"
+  assumes "convex s" "closed s" "z \<notin> s"
+  shows "\<exists>a b. inner a z < b \<and> (\<forall>x\<in>s. inner a x > b)"
+proof(cases "s={}")
+  case True thus ?thesis apply(rule_tac x="-z" in exI, rule_tac x=1 in exI)
+    using less_le_trans[OF _ inner_ge_zero[of z]] by auto
+next
+  case False obtain y where "y\<in>s" and y:"\<forall>x\<in>s. dist z y \<le> dist z x"
+    using distance_attains_inf[OF assms(2) False] by auto
+  show ?thesis apply(rule_tac x="y - z" in exI, rule_tac x="inner (y - z) z + (norm(y - z))\<twosuperior> / 2" in exI)
+    apply rule defer apply rule proof-
+    fix x assume "x\<in>s"
+    have "\<not> 0 < inner (z - y) (x - y)" apply(rule_tac notI) proof(drule closer_point_lemma)
+      assume "\<exists>u>0. u \<le> 1 \<and> dist (y + u *\<^sub>R (x - y)) z < dist y z"
+      then obtain u where "u>0" "u\<le>1" "dist (y + u *\<^sub>R (x - y)) z < dist y z" by auto
+      thus False using y[THEN bspec[where x="y + u *\<^sub>R (x - y)"]]
+        using assms(1)[unfolded convex_alt, THEN bspec[where x=y]]
+        using `x\<in>s` `y\<in>s` by (auto simp add: dist_commute algebra_simps) qed
+    moreover have "0 < norm (y - z) ^ 2" using `y\<in>s` `z\<notin>s` by auto
+    hence "0 < inner (y - z) (y - z)" unfolding power2_norm_eq_inner by simp
+    ultimately show "inner (y - z) z + (norm (y - z))\<twosuperior> / 2 < inner (y - z) x"
+      unfolding power2_norm_eq_inner and not_less by (auto simp add: field_simps inner_commute inner_diff)
+  qed(insert `y\<in>s` `z\<notin>s`, auto)
+qed
+
+lemma separating_hyperplane_closed_0:
+  assumes "convex (s::(real^'n::finite) set)" "closed s" "0 \<notin> s"
+  shows "\<exists>a b. a \<noteq> 0 \<and> 0 < b \<and> (\<forall>x\<in>s. inner a x > b)"
+  proof(cases "s={}") guess a using UNIV_witness[where 'a='n] ..
+  case True have "norm ((basis a)::real^'n::finite) = 1" 
+    using norm_basis and dimindex_ge_1 by auto
+  thus ?thesis apply(rule_tac x="basis a" in exI, rule_tac x=1 in exI) using True by auto
+next case False thus ?thesis using False using separating_hyperplane_closed_point[OF assms]
+    apply - apply(erule exE)+ unfolding dot_rzero apply(rule_tac x=a in exI, rule_tac x=b in exI) by auto qed
+
+subsection {* Now set-to-set for closed/compact sets. *}
+
+lemma separating_hyperplane_closed_compact:
+  assumes "convex (s::(real^'n::finite) set)" "closed s" "convex t" "compact t" "t \<noteq> {}" "s \<inter> t = {}"
+  shows "\<exists>a b. (\<forall>x\<in>s. inner a x < b) \<and> (\<forall>x\<in>t. inner a x > b)"
+proof(cases "s={}")
+  case True
+  obtain b where b:"b>0" "\<forall>x\<in>t. norm x \<le> b" using compact_imp_bounded[OF assms(4)] unfolding bounded_pos by auto
+  obtain z::"real^'n" where z:"norm z = b + 1" using vector_choose_size[of "b + 1"] and b(1) by auto
+  hence "z\<notin>t" using b(2)[THEN bspec[where x=z]] by auto
+  then obtain a b where ab:"inner a z < b" "\<forall>x\<in>t. b < inner a x"
+    using separating_hyperplane_closed_point[OF assms(3) compact_imp_closed[OF assms(4)], of z] by auto
+  thus ?thesis using True by auto
+next
+  case False then obtain y where "y\<in>s" by auto
+  obtain a b where "0 < b" "\<forall>x\<in>{x - y |x y. x \<in> s \<and> y \<in> t}. b < inner a x"
+    using separating_hyperplane_closed_point[OF convex_differences[OF assms(1,3)], of 0]
+    using closed_compact_differences[OF assms(2,4)] using assms(6) by(auto, blast)
+  hence ab:"\<forall>x\<in>s. \<forall>y\<in>t. b + inner a y < inner a x" apply- apply(rule,rule) apply(erule_tac x="x - y" in ballE) by (auto simp add: inner_diff)
+  def k \<equiv> "rsup ((\<lambda>x. inner a x) ` t)"
+  show ?thesis apply(rule_tac x="-a" in exI, rule_tac x="-(k + b / 2)" in exI)
+    apply(rule,rule) defer apply(rule) unfolding inner_minus_left and neg_less_iff_less proof-
+    from ab have "((\<lambda>x. inner a x) ` t) *<= (inner a y - b)"
+      apply(erule_tac x=y in ballE) apply(rule setleI) using `y\<in>s` by auto
+    hence k:"isLub UNIV ((\<lambda>x. inner a x) ` t) k" unfolding k_def apply(rule_tac rsup) using assms(5) by auto
+    fix x assume "x\<in>t" thus "inner a x < (k + b / 2)" using `0<b` and isLubD2[OF k, of "inner a x"] by auto
+  next
+    fix x assume "x\<in>s" 
+    hence "k \<le> inner a x - b" unfolding k_def apply(rule_tac rsup_le) using assms(5)
+      unfolding setle_def
+      using ab[THEN bspec[where x=x]] by auto
+    thus "k + b / 2 < inner a x" using `0 < b` by auto
+  qed
+qed
+
+lemma separating_hyperplane_compact_closed:
+  fixes s :: "(real ^ _) set"
+  assumes "convex s" "compact s" "s \<noteq> {}" "convex t" "closed t" "s \<inter> t = {}"
+  shows "\<exists>a b. (\<forall>x\<in>s. inner a x < b) \<and> (\<forall>x\<in>t. inner a x > b)"
+proof- obtain a b where "(\<forall>x\<in>t. inner a x < b) \<and> (\<forall>x\<in>s. b < inner a x)"
+    using separating_hyperplane_closed_compact[OF assms(4-5,1-2,3)] and assms(6) by auto
+  thus ?thesis apply(rule_tac x="-a" in exI, rule_tac x="-b" in exI) by auto qed
+
+subsection {* General case without assuming closure and getting non-strict separation. *}
+
+lemma separating_hyperplane_set_0:
+  assumes "convex s" "(0::real^'n::finite) \<notin> s"
+  shows "\<exists>a. a \<noteq> 0 \<and> (\<forall>x\<in>s. 0 \<le> inner a x)"
+proof- let ?k = "\<lambda>c. {x::real^'n. 0 \<le> inner c x}"
+  have "frontier (cball 0 1) \<inter> (\<Inter> (?k ` s)) \<noteq> {}"
+    apply(rule compact_imp_fip) apply(rule compact_frontier[OF compact_cball])
+    defer apply(rule,rule,erule conjE) proof-
+    fix f assume as:"f \<subseteq> ?k ` s" "finite f"
+    obtain c where c:"f = ?k ` c" "c\<subseteq>s" "finite c" using finite_subset_image[OF as(2,1)] by auto
+    then obtain a b where ab:"a \<noteq> 0" "0 < b"  "\<forall>x\<in>convex hull c. b < inner a x"
+      using separating_hyperplane_closed_0[OF convex_convex_hull, of c]
+      using finite_imp_compact_convex_hull[OF c(3), THEN compact_imp_closed] and assms(2)
+      using subset_hull[unfolded mem_def, of convex, OF assms(1), THEN sym, of c] by auto
+    hence "\<exists>x. norm x = 1 \<and> (\<forall>y\<in>c. 0 \<le> inner y x)" apply(rule_tac x="inverse(norm a) *\<^sub>R a" in exI)
+       using hull_subset[of c convex] unfolding subset_eq and inner_scaleR
+       apply- apply rule defer apply rule apply(rule mult_nonneg_nonneg)
+       by(auto simp add: inner_commute elim!: ballE)
+    thus "frontier (cball 0 1) \<inter> \<Inter>f \<noteq> {}" unfolding c(1) frontier_cball dist_norm by auto
+  qed(insert closed_halfspace_ge, auto)
+  then obtain x where "norm x = 1" "\<forall>y\<in>s. x\<in>?k y" unfolding frontier_cball dist_norm by auto
+  thus ?thesis apply(rule_tac x=x in exI) by(auto simp add: inner_commute) qed
+
+lemma separating_hyperplane_sets:
+  assumes "convex s" "convex (t::(real^'n::finite) set)" "s \<noteq> {}" "t \<noteq> {}" "s \<inter> t = {}"
+  shows "\<exists>a b. a \<noteq> 0 \<and> (\<forall>x\<in>s. inner a x \<le> b) \<and> (\<forall>x\<in>t. inner a x \<ge> b)"
+proof- from separating_hyperplane_set_0[OF convex_differences[OF assms(2,1)]]
+  obtain a where "a\<noteq>0" "\<forall>x\<in>{x - y |x y. x \<in> t \<and> y \<in> s}. 0 \<le> inner a x"  using assms(3-5) by auto 
+  hence "\<forall>x\<in>t. \<forall>y\<in>s. inner a y \<le> inner a x" apply- apply(rule, rule) apply(erule_tac x="x - y" in ballE) by (auto simp add: inner_diff)
+  thus ?thesis apply(rule_tac x=a in exI, rule_tac x="rsup ((\<lambda>x. inner a x) ` s)" in exI) using `a\<noteq>0`
+    apply(rule) apply(rule,rule) apply(rule rsup[THEN isLubD2]) prefer 4 apply(rule,rule rsup_le) unfolding setle_def
+    prefer 4 using assms(3-5) by blast+ qed
+
+subsection {* More convexity generalities. *}
+
+lemma convex_closure:
+  fixes s :: "'a::real_normed_vector set"
+  assumes "convex s" shows "convex(closure s)"
+  unfolding convex_def Ball_def closure_sequential
+  apply(rule,rule,rule,rule,rule,rule,rule,rule,rule) apply(erule_tac exE)+
+  apply(rule_tac x="\<lambda>n. u *\<^sub>R xb n + v *\<^sub>R xc n" in exI) apply(rule,rule)
+  apply(rule assms[unfolded convex_def, rule_format]) prefer 6
+  apply(rule Lim_add) apply(rule_tac [1-2] Lim_cmul) by auto
+
+lemma convex_interior:
+  fixes s :: "'a::real_normed_vector set"
+  assumes "convex s" shows "convex(interior s)"
+  unfolding convex_alt Ball_def mem_interior apply(rule,rule,rule,rule,rule,rule) apply(erule exE | erule conjE)+ proof-
+  fix x y u assume u:"0 \<le> u" "u \<le> (1::real)"
+  fix e d assume ed:"ball x e \<subseteq> s" "ball y d \<subseteq> s" "0<d" "0<e" 
+  show "\<exists>e>0. ball ((1 - u) *\<^sub>R x + u *\<^sub>R y) e \<subseteq> s" apply(rule_tac x="min d e" in exI)
+    apply rule unfolding subset_eq defer apply rule proof-
+    fix z assume "z \<in> ball ((1 - u) *\<^sub>R x + u *\<^sub>R y) (min d e)"
+    hence "(1- u) *\<^sub>R (z - u *\<^sub>R (y - x)) + u *\<^sub>R (z + (1 - u) *\<^sub>R (y - x)) \<in> s"
+      apply(rule_tac assms[unfolded convex_alt, rule_format])
+      using ed(1,2) and u unfolding subset_eq mem_ball Ball_def dist_norm by(auto simp add: algebra_simps)
+    thus "z \<in> s" using u by (auto simp add: algebra_simps) qed(insert u ed(3-4), auto) qed
+
+lemma convex_hull_eq_empty: "convex hull s = {} \<longleftrightarrow> s = {}"
+  using hull_subset[of s convex] convex_hull_empty by auto
+
+subsection {* Moving and scaling convex hulls. *}
+
+lemma convex_hull_translation_lemma:
+  "convex hull ((\<lambda>x. a + x) ` s) \<subseteq> (\<lambda>x. a + x) ` (convex hull s)"
+  apply(rule hull_minimal, rule image_mono, rule hull_subset) unfolding mem_def
+  using convex_translation[OF convex_convex_hull, of a s] by assumption
+
+lemma convex_hull_bilemma: fixes neg
+  assumes "(\<forall>s a. (convex hull (up a s)) \<subseteq> up a (convex hull s))"
+  shows "(\<forall>s. up a (up (neg a) s) = s) \<and> (\<forall>s. up (neg a) (up a s) = s) \<and> (\<forall>s t a. s \<subseteq> t \<longrightarrow> up a s \<subseteq> up a t)
+  \<Longrightarrow> \<forall>s. (convex hull (up a s)) = up a (convex hull s)"
+  using assms by(metis subset_antisym) 
+
+lemma convex_hull_translation:
+  "convex hull ((\<lambda>x. a + x) ` s) = (\<lambda>x. a + x) ` (convex hull s)"
+  apply(rule convex_hull_bilemma[rule_format, of _ _ "\<lambda>a. -a"], rule convex_hull_translation_lemma) unfolding image_image by auto
+
+lemma convex_hull_scaling_lemma:
+ "(convex hull ((\<lambda>x. c *\<^sub>R x) ` s)) \<subseteq> (\<lambda>x. c *\<^sub>R x) ` (convex hull s)"
+  apply(rule hull_minimal, rule image_mono, rule hull_subset)
+  unfolding mem_def by(rule convex_scaling, rule convex_convex_hull)
+
+lemma convex_hull_scaling:
+  "convex hull ((\<lambda>x. c *\<^sub>R x) ` s) = (\<lambda>x. c *\<^sub>R x) ` (convex hull s)"
+  apply(cases "c=0") defer apply(rule convex_hull_bilemma[rule_format, of _ _ inverse]) apply(rule convex_hull_scaling_lemma)
+  unfolding image_image scaleR_scaleR by(auto simp add:image_constant_conv convex_hull_eq_empty)
+
+lemma convex_hull_affinity:
+  "convex hull ((\<lambda>x. a + c *\<^sub>R x) ` s) = (\<lambda>x. a + c *\<^sub>R x) ` (convex hull s)"
+  unfolding image_image[THEN sym] convex_hull_scaling convex_hull_translation  ..
+
+subsection {* Convex set as intersection of halfspaces. *}
+
+lemma convex_halfspace_intersection:
+  fixes s :: "(real ^ _) set"
+  assumes "closed s" "convex s"
+  shows "s = \<Inter> {h. s \<subseteq> h \<and> (\<exists>a b. h = {x. inner a x \<le> b})}"
+  apply(rule set_ext, rule) unfolding Inter_iff Ball_def mem_Collect_eq apply(rule,rule,erule conjE) proof- 
+  fix x  assume "\<forall>xa. s \<subseteq> xa \<and> (\<exists>a b. xa = {x. inner a x \<le> b}) \<longrightarrow> x \<in> xa"
+  hence "\<forall>a b. s \<subseteq> {x. inner a x \<le> b} \<longrightarrow> x \<in> {x. inner a x \<le> b}" by blast
+  thus "x\<in>s" apply(rule_tac ccontr) apply(drule separating_hyperplane_closed_point[OF assms(2,1)])
+    apply(erule exE)+ apply(erule_tac x="-a" in allE, erule_tac x="-b" in allE) by auto
+qed auto
+
+subsection {* Radon's theorem (from Lars Schewe). *}
+
+lemma radon_ex_lemma:
+  assumes "finite c" "affine_dependent c"
+  shows "\<exists>u. setsum u c = 0 \<and> (\<exists>v\<in>c. u v \<noteq> 0) \<and> setsum (\<lambda>v. u v *\<^sub>R v) c = 0"
+proof- from assms(2)[unfolded affine_dependent_explicit] guess s .. then guess u ..
+  thus ?thesis apply(rule_tac x="\<lambda>v. if v\<in>s then u v else 0" in exI) unfolding if_smult scaleR_zero_left
+    and setsum_restrict_set[OF assms(1), THEN sym] by(auto simp add: Int_absorb1) qed
+
+lemma radon_s_lemma:
+  assumes "finite s" "setsum f s = (0::real)"
+  shows "setsum f {x\<in>s. 0 < f x} = - setsum f {x\<in>s. f x < 0}"
+proof- have *:"\<And>x. (if f x < 0 then f x else 0) + (if 0 < f x then f x else 0) = f x" by auto
+  show ?thesis unfolding real_add_eq_0_iff[THEN sym] and setsum_restrict_set''[OF assms(1)] and setsum_addf[THEN sym] and *
+    using assms(2) by assumption qed
+
+lemma radon_v_lemma:
+  assumes "finite s" "setsum f s = 0" "\<forall>x. g x = (0::real) \<longrightarrow> f x = (0::real^'n)"
+  shows "(setsum f {x\<in>s. 0 < g x}) = - setsum f {x\<in>s. g x < 0}"
+proof-
+  have *:"\<And>x. (if 0 < g x then f x else 0) + (if g x < 0 then f x else 0) = f x" using assms(3) by auto 
+  show ?thesis unfolding eq_neg_iff_add_eq_0 and setsum_restrict_set''[OF assms(1)] and setsum_addf[THEN sym] and *
+    using assms(2) by assumption qed
+
+lemma radon_partition:
+  assumes "finite c" "affine_dependent c"
+  shows "\<exists>m p. m \<inter> p = {} \<and> m \<union> p = c \<and> (convex hull m) \<inter> (convex hull p) \<noteq> {}" proof-
+  obtain u v where uv:"setsum u c = 0" "v\<in>c" "u v \<noteq> 0"  "(\<Sum>v\<in>c. u v *\<^sub>R v) = 0" using radon_ex_lemma[OF assms] by auto
+  have fin:"finite {x \<in> c. 0 < u x}" "finite {x \<in> c. 0 > u x}" using assms(1) by auto
+  def z \<equiv> "(inverse (setsum u {x\<in>c. u x > 0})) *\<^sub>R setsum (\<lambda>x. u x *\<^sub>R x) {x\<in>c. u x > 0}"
+  have "setsum u {x \<in> c. 0 < u x} \<noteq> 0" proof(cases "u v \<ge> 0")
+    case False hence "u v < 0" by auto
+    thus ?thesis proof(cases "\<exists>w\<in>{x \<in> c. 0 < u x}. u w > 0") 
+      case True thus ?thesis using setsum_nonneg_eq_0_iff[of _ u, OF fin(1)] by auto
+    next
+      case False hence "setsum u c \<le> setsum (\<lambda>x. if x=v then u v else 0) c" apply(rule_tac setsum_mono) by auto
+      thus ?thesis unfolding setsum_delta[OF assms(1)] using uv(2) and `u v < 0` and uv(1) by auto qed
+  qed (insert setsum_nonneg_eq_0_iff[of _ u, OF fin(1)] uv(2-3), auto)
+
+  hence *:"setsum u {x\<in>c. u x > 0} > 0" unfolding real_less_def apply(rule_tac conjI, rule_tac setsum_nonneg) by auto
+  moreover have "setsum u ({x \<in> c. 0 < u x} \<union> {x \<in> c. u x < 0}) = setsum u c"
+    "(\<Sum>x\<in>{x \<in> c. 0 < u x} \<union> {x \<in> c. u x < 0}. u x *\<^sub>R x) = (\<Sum>x\<in>c. u x *\<^sub>R x)"
+    using assms(1) apply(rule_tac[!] setsum_mono_zero_left) by auto
+  hence "setsum u {x \<in> c. 0 < u x} = - setsum u {x \<in> c. 0 > u x}"
+   "(\<Sum>x\<in>{x \<in> c. 0 < u x}. u x *\<^sub>R x) = - (\<Sum>x\<in>{x \<in> c. 0 > u x}. u x *\<^sub>R x)" 
+    unfolding eq_neg_iff_add_eq_0 using uv(1,4) by (auto simp add:  setsum_Un_zero[OF fin, THEN sym]) 
+  moreover have "\<forall>x\<in>{v \<in> c. u v < 0}. 0 \<le> inverse (setsum u {x \<in> c. 0 < u x}) * - u x" 
+    apply (rule) apply (rule mult_nonneg_nonneg) using * by auto
+
+  ultimately have "z \<in> convex hull {v \<in> c. u v \<le> 0}" unfolding convex_hull_explicit mem_Collect_eq
+    apply(rule_tac x="{v \<in> c. u v < 0}" in exI, rule_tac x="\<lambda>y. inverse (setsum u {x\<in>c. u x > 0}) * - u y" in exI)
+    using assms(1) unfolding scaleR_scaleR[THEN sym] scaleR_right.setsum [symmetric] and z_def
+    by(auto simp add: setsum_negf vector_smult_lneg mult_right.setsum[THEN sym])
+  moreover have "\<forall>x\<in>{v \<in> c. 0 < u v}. 0 \<le> inverse (setsum u {x \<in> c. 0 < u x}) * u x" 
+    apply (rule) apply (rule mult_nonneg_nonneg) using * by auto 
+  hence "z \<in> convex hull {v \<in> c. u v > 0}" unfolding convex_hull_explicit mem_Collect_eq
+    apply(rule_tac x="{v \<in> c. 0 < u v}" in exI, rule_tac x="\<lambda>y. inverse (setsum u {x\<in>c. u x > 0}) * u y" in exI)
+    using assms(1) unfolding scaleR_scaleR[THEN sym] scaleR_right.setsum [symmetric] and z_def using *
+    by(auto simp add: setsum_negf vector_smult_lneg mult_right.setsum[THEN sym])
+  ultimately show ?thesis apply(rule_tac x="{v\<in>c. u v \<le> 0}" in exI, rule_tac x="{v\<in>c. u v > 0}" in exI) by auto
+qed
+
+lemma radon: assumes "affine_dependent c"
+  obtains m p where "m\<subseteq>c" "p\<subseteq>c" "m \<inter> p = {}" "(convex hull m) \<inter> (convex hull p) \<noteq> {}"
+proof- from assms[unfolded affine_dependent_explicit] guess s .. then guess u ..
+  hence *:"finite s" "affine_dependent s" and s:"s \<subseteq> c" unfolding affine_dependent_explicit by auto
+  from radon_partition[OF *] guess m .. then guess p ..
+  thus ?thesis apply(rule_tac that[of p m]) using s by auto qed
+
+subsection {* Helly's theorem. *}
+
+lemma helly_induct: fixes f::"(real^'n::finite) set set"
+  assumes "f hassize n" "n \<ge> CARD('n) + 1"
+  "\<forall>s\<in>f. convex s" "\<forall>t\<subseteq>f. card t = CARD('n) + 1 \<longrightarrow> \<Inter> t \<noteq> {}"
+  shows "\<Inter> f \<noteq> {}"
+  using assms unfolding hassize_def apply(erule_tac conjE) proof(induct n arbitrary: f)
+case (Suc n)
+show "\<Inter> f \<noteq> {}" apply(cases "n = CARD('n)") apply(rule Suc(4)[rule_format])
+  unfolding card_Diff_singleton_if[OF Suc(5)] and Suc(6) proof-
+  assume ng:"n \<noteq> CARD('n)" hence "\<exists>X. \<forall>s\<in>f. X s \<in> \<Inter>(f - {s})" apply(rule_tac bchoice) unfolding ex_in_conv
+    apply(rule, rule Suc(1)[rule_format])  unfolding card_Diff_singleton_if[OF Suc(5)] and Suc(6)
+    defer apply(rule Suc(3)[rule_format]) defer apply(rule Suc(4)[rule_format]) using Suc(2,5) by auto
+  then obtain X where X:"\<forall>s\<in>f. X s \<in> \<Inter>(f - {s})" by auto
+  show ?thesis proof(cases "inj_on X f")
+    case False then obtain s t where st:"s\<noteq>t" "s\<in>f" "t\<in>f" "X s = X t" unfolding inj_on_def by auto
+    hence *:"\<Inter> f = \<Inter> (f - {s}) \<inter> \<Inter> (f - {t})" by auto
+    show ?thesis unfolding * unfolding ex_in_conv[THEN sym] apply(rule_tac x="X s" in exI)
+      apply(rule, rule X[rule_format]) using X st by auto
+  next case True then obtain m p where mp:"m \<inter> p = {}" "m \<union> p = X ` f" "convex hull m \<inter> convex hull p \<noteq> {}"
+      using radon_partition[of "X ` f"] and affine_dependent_biggerset[of "X ` f"]
+      unfolding card_image[OF True] and Suc(6) using Suc(2,5) and ng by auto
+    have "m \<subseteq> X ` f" "p \<subseteq> X ` f" using mp(2) by auto
+    then obtain g h where gh:"m = X ` g" "p = X ` h" "g \<subseteq> f" "h \<subseteq> f" unfolding subset_image_iff by auto 
+    hence "f \<union> (g \<union> h) = f" by auto
+    hence f:"f = g \<union> h" using inj_on_Un_image_eq_iff[of X f "g \<union> h"] and True
+      unfolding mp(2)[unfolded image_Un[THEN sym] gh] by auto
+    have *:"g \<inter> h = {}" using mp(1) unfolding gh using inj_on_image_Int[OF True gh(3,4)] by auto
+    have "convex hull (X ` h) \<subseteq> \<Inter> g" "convex hull (X ` g) \<subseteq> \<Inter> h"
+      apply(rule_tac [!] hull_minimal) using Suc(3) gh(3-4)  unfolding mem_def unfolding subset_eq
+      apply(rule_tac [2] convex_Inter, rule_tac [4] convex_Inter) apply rule prefer 3 apply rule proof-
+      fix x assume "x\<in>X ` g" then guess y unfolding image_iff ..
+      thus "x\<in>\<Inter>h" using X[THEN bspec[where x=y]] using * f by auto next
+      fix x assume "x\<in>X ` h" then guess y unfolding image_iff ..
+      thus "x\<in>\<Inter>g" using X[THEN bspec[where x=y]] using * f by auto
+    qed(auto)
+    thus ?thesis unfolding f using mp(3)[unfolded gh] by blast qed
+qed(insert dimindex_ge_1, auto) qed(auto)
+
+lemma helly: fixes f::"(real^'n::finite) set set"
+  assumes "finite f" "card f \<ge> CARD('n) + 1" "\<forall>s\<in>f. convex s"
+          "\<forall>t\<subseteq>f. card t = CARD('n) + 1 \<longrightarrow> \<Inter> t \<noteq> {}"
+  shows "\<Inter> f \<noteq>{}"
+  apply(rule helly_induct) unfolding hassize_def using assms by auto
+
+subsection {* Convex hull is "preserved" by a linear function. *}
+
+lemma convex_hull_linear_image:
+  assumes "bounded_linear f"
+  shows "f ` (convex hull s) = convex hull (f ` s)"
+  apply rule unfolding subset_eq ball_simps apply(rule_tac[!] hull_induct, rule hull_inc) prefer 3  
+  apply(erule imageE)apply(rule_tac x=xa in image_eqI) apply assumption
+  apply(rule hull_subset[unfolded subset_eq, rule_format]) apply assumption
+proof-
+  interpret f: bounded_linear f by fact
+  show "convex {x. f x \<in> convex hull f ` s}" 
+  unfolding convex_def by(auto simp add: f.scaleR f.add convex_convex_hull[unfolded convex_def, rule_format]) next
+  interpret f: bounded_linear f by fact
+  show "convex {x. x \<in> f ` (convex hull s)}" using  convex_convex_hull[unfolded convex_def, of s] 
+    unfolding convex_def by (auto simp add: f.scaleR [symmetric] f.add [symmetric])
+qed auto
+
+lemma in_convex_hull_linear_image:
+  assumes "bounded_linear f" "x \<in> convex hull s"
+  shows "(f x) \<in> convex hull (f ` s)"
+using convex_hull_linear_image[OF assms(1)] assms(2) by auto
+
+subsection {* Homeomorphism of all convex compact sets with nonempty interior. *}
+
+lemma compact_frontier_line_lemma:
+  fixes s :: "(real ^ _) set"
+  assumes "compact s" "0 \<in> s" "x \<noteq> 0" 
+  obtains u where "0 \<le> u" "(u *\<^sub>R x) \<in> frontier s" "\<forall>v>u. (v *\<^sub>R x) \<notin> s"
+proof-
+  obtain b where b:"b>0" "\<forall>x\<in>s. norm x \<le> b" using compact_imp_bounded[OF assms(1), unfolded bounded_pos] by auto
+  let ?A = "{y. \<exists>u. 0 \<le> u \<and> u \<le> b / norm(x) \<and> (y = u *\<^sub>R x)}"
+  have A:"?A = (\<lambda>u. dest_vec1 u *\<^sub>R x) ` {0 .. vec1 (b / norm x)}"
+    unfolding image_image[of "\<lambda>u. u *\<^sub>R x" "\<lambda>x. dest_vec1 x", THEN sym]
+    unfolding dest_vec1_inverval vec1_dest_vec1 by auto
+  have "compact ?A" unfolding A apply(rule compact_continuous_image, rule continuous_at_imp_continuous_on)
+    apply(rule, rule continuous_vmul)
+    apply (rule continuous_dest_vec1)
+    apply(rule continuous_at_id) by(rule compact_interval)
+  moreover have "{y. \<exists>u\<ge>0. u \<le> b / norm x \<and> y = u *\<^sub>R x} \<inter> s \<noteq> {}" apply(rule not_disjointI[OF _ assms(2)])
+    unfolding mem_Collect_eq using `b>0` assms(3) by(auto intro!: divide_nonneg_pos)
+  ultimately obtain u y where obt: "u\<ge>0" "u \<le> b / norm x" "y = u *\<^sub>R x"
+    "y\<in>?A" "y\<in>s" "\<forall>z\<in>?A \<inter> s. dist 0 z \<le> dist 0 y" using distance_attains_sup[OF compact_inter[OF _ assms(1), of ?A], of 0] by auto
+
+  have "norm x > 0" using assms(3)[unfolded zero_less_norm_iff[THEN sym]] by auto
+  { fix v assume as:"v > u" "v *\<^sub>R x \<in> s"
+    hence "v \<le> b / norm x" using b(2)[rule_format, OF as(2)] 
+      using `u\<ge>0` unfolding pos_le_divide_eq[OF `norm x > 0`] by auto
+    hence "norm (v *\<^sub>R x) \<le> norm y" apply(rule_tac obt(6)[rule_format, unfolded dist_0_norm]) apply(rule IntI) defer 
+      apply(rule as(2)) unfolding mem_Collect_eq apply(rule_tac x=v in exI) 
+      using as(1) `u\<ge>0` by(auto simp add:field_simps) 
+    hence False unfolding obt(3) using `u\<ge>0` `norm x > 0` `v>u` by(auto simp add:field_simps)
+  } note u_max = this
+
+  have "u *\<^sub>R x \<in> frontier s" unfolding frontier_straddle apply(rule,rule,rule) apply(rule_tac x="u *\<^sub>R x" in bexI) unfolding obt(3)[THEN sym]
+    prefer 3 apply(rule_tac x="(u + (e / 2) / norm x) *\<^sub>R x" in exI) apply(rule, rule) proof-
+    fix e  assume "0 < e" and as:"(u + e / 2 / norm x) *\<^sub>R x \<in> s"
+    hence "u + e / 2 / norm x > u" using`norm x > 0` by(auto simp del:zero_less_norm_iff intro!: divide_pos_pos)
+    thus False using u_max[OF _ as] by auto
+  qed(insert `y\<in>s`, auto simp add: dist_norm scaleR_left_distrib obt(3))
+  thus ?thesis apply(rule_tac that[of u]) apply(rule obt(1), assumption)
+    apply(rule,rule,rule ccontr) apply(rule u_max) by auto qed
+
+lemma starlike_compact_projective:
+  assumes "compact s" "cball (0::real^'n::finite) 1 \<subseteq> s "
+  "\<forall>x\<in>s. \<forall>u. 0 \<le> u \<and> u < 1 \<longrightarrow> (u *\<^sub>R x) \<in> (s - frontier s )"
+  shows "s homeomorphic (cball (0::real^'n::finite) 1)"
+proof-
+  have fs:"frontier s \<subseteq> s" apply(rule frontier_subset_closed) using compact_imp_closed[OF assms(1)] by simp
+  def pi \<equiv> "\<lambda>x::real^'n. inverse (norm x) *\<^sub>R x"
+  have "0 \<notin> frontier s" unfolding frontier_straddle apply(rule ccontr) unfolding not_not apply(erule_tac x=1 in allE)
+    using assms(2)[unfolded subset_eq Ball_def mem_cball] by auto
+  have injpi:"\<And>x y. pi x = pi y \<and> norm x = norm y \<longleftrightarrow> x = y" unfolding pi_def by auto
+
+  have contpi:"continuous_on (UNIV - {0}) pi" apply(rule continuous_at_imp_continuous_on)
+    apply rule unfolding pi_def
+    apply (rule continuous_mul)
+    apply (rule continuous_at_inv[unfolded o_def])
+    apply (rule continuous_at_norm)
+    apply simp
+    apply (rule continuous_at_id)
+    done
+  def sphere \<equiv> "{x::real^'n. norm x = 1}"
+  have pi:"\<And>x. x \<noteq> 0 \<Longrightarrow> pi x \<in> sphere" "\<And>x u. u>0 \<Longrightarrow> pi (u *\<^sub>R x) = pi x" unfolding pi_def sphere_def by auto
+
+  have "0\<in>s" using assms(2) and centre_in_cball[of 0 1] by auto
+  have front_smul:"\<forall>x\<in>frontier s. \<forall>u\<ge>0. u *\<^sub>R x \<in> s \<longleftrightarrow> u \<le> 1" proof(rule,rule,rule)
+    fix x u assume x:"x\<in>frontier s" and "(0::real)\<le>u"
+    hence "x\<noteq>0" using `0\<notin>frontier s` by auto
+    obtain v where v:"0 \<le> v" "v *\<^sub>R x \<in> frontier s" "\<forall>w>v. w *\<^sub>R x \<notin> s"
+      using compact_frontier_line_lemma[OF assms(1) `0\<in>s` `x\<noteq>0`] by auto
+    have "v=1" apply(rule ccontr) unfolding neq_iff apply(erule disjE) proof-
+      assume "v<1" thus False using v(3)[THEN spec[where x=1]] using x and fs by auto next
+      assume "v>1" thus False using assms(3)[THEN bspec[where x="v *\<^sub>R x"], THEN spec[where x="inverse v"]]
+        using v and x and fs unfolding inverse_less_1_iff by auto qed
+    show "u *\<^sub>R x \<in> s \<longleftrightarrow> u \<le> 1" apply rule  using v(3)[unfolded `v=1`, THEN spec[where x=u]] proof-
+      assume "u\<le>1" thus "u *\<^sub>R x \<in> s" apply(cases "u=1")
+        using assms(3)[THEN bspec[where x=x], THEN spec[where x=u]] using `0\<le>u` and x and fs by auto qed auto qed
+
+  have "\<exists>surf. homeomorphism (frontier s) sphere pi surf"
+    apply(rule homeomorphism_compact) apply(rule compact_frontier[OF assms(1)])
+    apply(rule continuous_on_subset[OF contpi]) defer apply(rule set_ext,rule) 
+    unfolding inj_on_def prefer 3 apply(rule,rule,rule)
+  proof- fix x assume "x\<in>pi ` frontier s" then obtain y where "y\<in>frontier s" "x = pi y" by auto
+    thus "x \<in> sphere" using pi(1)[of y] and `0 \<notin> frontier s` by auto
+  next fix x assume "x\<in>sphere" hence "norm x = 1" "x\<noteq>0" unfolding sphere_def by auto
+    then obtain u where "0 \<le> u" "u *\<^sub>R x \<in> frontier s" "\<forall>v>u. v *\<^sub>R x \<notin> s"
+      using compact_frontier_line_lemma[OF assms(1) `0\<in>s`, of x] by auto
+    thus "x \<in> pi ` frontier s" unfolding image_iff le_less pi_def apply(rule_tac x="u *\<^sub>R x" in bexI) using `norm x = 1` `0\<notin>frontier s` by auto
+  next fix x y assume as:"x \<in> frontier s" "y \<in> frontier s" "pi x = pi y"
+    hence xys:"x\<in>s" "y\<in>s" using fs by auto
+    from as(1,2) have nor:"norm x \<noteq> 0" "norm y \<noteq> 0" using `0\<notin>frontier s` by auto 
+    from nor have x:"x = norm x *\<^sub>R ((inverse (norm y)) *\<^sub>R y)" unfolding as(3)[unfolded pi_def, THEN sym] by auto 
+    from nor have y:"y = norm y *\<^sub>R ((inverse (norm x)) *\<^sub>R x)" unfolding as(3)[unfolded pi_def] by auto 
+    have "0 \<le> norm y * inverse (norm x)" "0 \<le> norm x * inverse (norm y)"
+      unfolding divide_inverse[THEN sym] apply(rule_tac[!] divide_nonneg_pos) using nor by auto
+    hence "norm x = norm y" apply(rule_tac ccontr) unfolding neq_iff
+      using x y and front_smul[THEN bspec, OF as(1), THEN spec[where x="norm y * (inverse (norm x))"]]
+      using front_smul[THEN bspec, OF as(2), THEN spec[where x="norm x * (inverse (norm y))"]]
+      using xys nor by(auto simp add:field_simps divide_le_eq_1 divide_inverse[THEN sym])
+    thus "x = y" apply(subst injpi[THEN sym]) using as(3) by auto
+  qed(insert `0 \<notin> frontier s`, auto)
+  then obtain surf where surf:"\<forall>x\<in>frontier s. surf (pi x) = x"  "pi ` frontier s = sphere" "continuous_on (frontier s) pi"
+    "\<forall>y\<in>sphere. pi (surf y) = y" "surf ` sphere = frontier s" "continuous_on sphere surf" unfolding homeomorphism_def by auto
+  
+  have cont_surfpi:"continuous_on (UNIV -  {0}) (surf \<circ> pi)" apply(rule continuous_on_compose, rule contpi)
+    apply(rule continuous_on_subset[of sphere], rule surf(6)) using pi(1) by auto
+
+  { fix x assume as:"x \<in> cball (0::real^'n) 1"
+    have "norm x *\<^sub>R surf (pi x) \<in> s" proof(cases "x=0 \<or> norm x = 1") 
+      case False hence "pi x \<in> sphere" "norm x < 1" using pi(1)[of x] as by(auto simp add: dist_norm)
+      thus ?thesis apply(rule_tac assms(3)[rule_format, THEN DiffD1])
+        apply(rule_tac fs[unfolded subset_eq, rule_format])
+        unfolding surf(5)[THEN sym] by auto
+    next case True thus ?thesis apply rule defer unfolding pi_def apply(rule fs[unfolded subset_eq, rule_format])
+        unfolding  surf(5)[unfolded sphere_def, THEN sym] using `0\<in>s` by auto qed } note hom = this
+
+  { fix x assume "x\<in>s"
+    hence "x \<in> (\<lambda>x. norm x *\<^sub>R surf (pi x)) ` cball 0 1" proof(cases "x=0")
+      case True show ?thesis unfolding image_iff True apply(rule_tac x=0 in bexI) by auto
+    next let ?a = "inverse (norm (surf (pi x)))"
+      case False hence invn:"inverse (norm x) \<noteq> 0" by auto
+      from False have pix:"pi x\<in>sphere" using pi(1) by auto
+      hence "pi (surf (pi x)) = pi x" apply(rule_tac surf(4)[rule_format]) by assumption
+      hence **:"norm x *\<^sub>R (?a *\<^sub>R surf (pi x)) = x" apply(rule_tac scaleR_left_imp_eq[OF invn]) unfolding pi_def using invn by auto
+      hence *:"?a * norm x > 0" and"?a > 0" "?a \<noteq> 0" using surf(5) `0\<notin>frontier s` apply -
+        apply(rule_tac mult_pos_pos) using False[unfolded zero_less_norm_iff[THEN sym]] by auto
+      have "norm (surf (pi x)) \<noteq> 0" using ** False by auto
+      hence "norm x = norm ((?a * norm x) *\<^sub>R surf (pi x))"
+        unfolding norm_scaleR abs_mult abs_norm_cancel abs_of_pos[OF `?a > 0`] by auto
+      moreover have "pi x = pi ((inverse (norm (surf (pi x))) * norm x) *\<^sub>R surf (pi x))" 
+        unfolding pi(2)[OF *] surf(4)[rule_format, OF pix] ..
+      moreover have "surf (pi x) \<in> frontier s" using surf(5) pix by auto
+      hence "dist 0 (inverse (norm (surf (pi x))) *\<^sub>R x) \<le> 1" unfolding dist_norm
+        using ** and * using front_smul[THEN bspec[where x="surf (pi x)"], THEN spec[where x="norm x * ?a"]]
+        using False `x\<in>s` by(auto simp add:field_simps)
+      ultimately show ?thesis unfolding image_iff apply(rule_tac x="inverse (norm (surf(pi x))) *\<^sub>R x" in bexI)
+        apply(subst injpi[THEN sym]) unfolding abs_mult abs_norm_cancel abs_of_pos[OF `?a > 0`]
+        unfolding pi(2)[OF `?a > 0`] by auto
+    qed } note hom2 = this
+
+  show ?thesis apply(subst homeomorphic_sym) apply(rule homeomorphic_compact[where f="\<lambda>x. norm x *\<^sub>R surf (pi x)"])
+    apply(rule compact_cball) defer apply(rule set_ext, rule, erule imageE, drule hom)
+    prefer 4 apply(rule continuous_at_imp_continuous_on, rule) apply(rule_tac [3] hom2) proof-
+    fix x::"real^'n" assume as:"x \<in> cball 0 1"
+    thus "continuous (at x) (\<lambda>x. norm x *\<^sub>R surf (pi x))" proof(cases "x=0")
+      case False thus ?thesis apply(rule_tac continuous_mul, rule_tac continuous_at_norm)
+        using cont_surfpi unfolding continuous_on_eq_continuous_at[OF open_delete[OF open_UNIV]] o_def by auto
+    next guess a using UNIV_witness[where 'a = 'n] ..
+      obtain B where B:"\<forall>x\<in>s. norm x \<le> B" using compact_imp_bounded[OF assms(1)] unfolding bounded_iff by auto
+      hence "B > 0" using assms(2) unfolding subset_eq apply(erule_tac x="basis a" in ballE) defer apply(erule_tac x="basis a" in ballE)
+        unfolding Ball_def mem_cball dist_norm by (auto simp add: norm_basis[unfolded One_nat_def])
+      case True show ?thesis unfolding True continuous_at Lim_at apply(rule,rule) apply(rule_tac x="e / B" in exI)
+        apply(rule) apply(rule divide_pos_pos) prefer 3 apply(rule,rule,erule conjE)
+        unfolding norm_0 scaleR_zero_left dist_norm diff_0_right norm_scaleR abs_norm_cancel proof-
+        fix e and x::"real^'n" assume as:"norm x < e / B" "0 < norm x" "0<e"
+        hence "surf (pi x) \<in> frontier s" using pi(1)[of x] unfolding surf(5)[THEN sym] by auto
+        hence "norm (surf (pi x)) \<le> B" using B fs by auto
+        hence "norm x * norm (surf (pi x)) \<le> norm x * B" using as(2) by auto
+        also have "\<dots> < e / B * B" apply(rule mult_strict_right_mono) using as(1) `B>0` by auto
+        also have "\<dots> = e" using `B>0` by auto
+        finally show "norm x * norm (surf (pi x)) < e" by assumption
+      qed(insert `B>0`, auto) qed
+  next { fix x assume as:"surf (pi x) = 0"
+      have "x = 0" proof(rule ccontr)
+        assume "x\<noteq>0" hence "pi x \<in> sphere" using pi(1) by auto
+        hence "surf (pi x) \<in> frontier s" using surf(5) by auto
+        thus False using `0\<notin>frontier s` unfolding as by simp qed
+    } note surf_0 = this
+    show "inj_on (\<lambda>x. norm x *\<^sub>R surf (pi x)) (cball 0 1)" unfolding inj_on_def proof(rule,rule,rule)
+      fix x y assume as:"x \<in> cball 0 1" "y \<in> cball 0 1" "norm x *\<^sub>R surf (pi x) = norm y *\<^sub>R surf (pi y)"
+      thus "x=y" proof(cases "x=0 \<or> y=0") 
+        case True thus ?thesis using as by(auto elim: surf_0) next
+        case False
+        hence "pi (surf (pi x)) = pi (surf (pi y))" using as(3)
+          using pi(2)[of "norm x" "surf (pi x)"] pi(2)[of "norm y" "surf (pi y)"] by auto
+        moreover have "pi x \<in> sphere" "pi y \<in> sphere" using pi(1) False by auto
+        ultimately have *:"pi x = pi y" using surf(4)[THEN bspec[where x="pi x"]] surf(4)[THEN bspec[where x="pi y"]] by auto 
+        moreover have "norm x = norm y" using as(3)[unfolded *] using False by(auto dest:surf_0)
+        ultimately show ?thesis using injpi by auto qed qed
+  qed auto qed
+
+lemma homeomorphic_convex_compact_lemma: fixes s::"(real^'n::finite) set"
+  assumes "convex s" "compact s" "cball 0 1 \<subseteq> s"
+  shows "s homeomorphic (cball (0::real^'n) 1)"
+  apply(rule starlike_compact_projective[OF assms(2-3)]) proof(rule,rule,rule,erule conjE)
+  fix x u assume as:"x \<in> s" "0 \<le> u" "u < (1::real)"
+  hence "u *\<^sub>R x \<in> interior s" unfolding interior_def mem_Collect_eq
+    apply(rule_tac x="ball (u *\<^sub>R x) (1 - u)" in exI) apply(rule, rule open_ball)
+    unfolding centre_in_ball apply rule defer apply(rule) unfolding mem_ball proof-
+    fix y assume "dist (u *\<^sub>R x) y < 1 - u"
+    hence "inverse (1 - u) *\<^sub>R (y - u *\<^sub>R x) \<in> s"
+      using assms(3) apply(erule_tac subsetD) unfolding mem_cball dist_commute dist_norm
+      unfolding group_add_class.diff_0 group_add_class.diff_0_right norm_minus_cancel norm_scaleR
+      apply (rule mult_left_le_imp_le[of "1 - u"])
+      unfolding class_semiring.mul_a using `u<1` by auto
+    thus "y \<in> s" using assms(1)[unfolded convex_def, rule_format, of "inverse(1 - u) *\<^sub>R (y - u *\<^sub>R x)" x "1 - u" u]
+      using as unfolding scaleR_scaleR by auto qed auto
+  thus "u *\<^sub>R x \<in> s - frontier s" using frontier_def and interior_subset by auto qed
+
+lemma homeomorphic_convex_compact_cball: fixes e::real and s::"(real^'n::finite) set"
+  assumes "convex s" "compact s" "interior s \<noteq> {}" "0 < e"
+  shows "s homeomorphic (cball (b::real^'n::finite) e)"
+proof- obtain a where "a\<in>interior s" using assms(3) by auto
+  then obtain d where "d>0" and d:"cball a d \<subseteq> s" unfolding mem_interior_cball by auto
+  let ?d = "inverse d" and ?n = "0::real^'n"
+  have "cball ?n 1 \<subseteq> (\<lambda>x. inverse d *\<^sub>R (x - a)) ` s"
+    apply(rule, rule_tac x="d *\<^sub>R x + a" in image_eqI) defer
+    apply(rule d[unfolded subset_eq, rule_format]) using `d>0` unfolding mem_cball dist_norm
+    by(auto simp add: mult_right_le_one_le)
+  hence "(\<lambda>x. inverse d *\<^sub>R (x - a)) ` s homeomorphic cball ?n 1"
+    using homeomorphic_convex_compact_lemma[of "(\<lambda>x. ?d *\<^sub>R -a + ?d *\<^sub>R x) ` s", OF convex_affinity compact_affinity]
+    using assms(1,2) by(auto simp add: uminus_add_conv_diff scaleR_right_diff_distrib)
+  thus ?thesis apply(rule_tac homeomorphic_trans[OF _ homeomorphic_balls(2)[of 1 _ ?n]])
+    apply(rule homeomorphic_trans[OF homeomorphic_affinity[of "?d" s "?d *\<^sub>R -a"]])
+    using `d>0` `e>0` by(auto simp add: uminus_add_conv_diff scaleR_right_diff_distrib) qed
+
+lemma homeomorphic_convex_compact: fixes s::"(real^'n::finite) set" and t::"(real^'n) set"
+  assumes "convex s" "compact s" "interior s \<noteq> {}"
+          "convex t" "compact t" "interior t \<noteq> {}"
+  shows "s homeomorphic t"
+  using assms by(meson zero_less_one homeomorphic_trans homeomorphic_convex_compact_cball homeomorphic_sym)
+
+subsection {* Epigraphs of convex functions. *}
+
+definition "epigraph s (f::real^'n \<Rightarrow> real) = {xy. fstcart xy \<in> s \<and> f(fstcart xy) \<le> dest_vec1 (sndcart xy)}"
+
+lemma mem_epigraph: "(pastecart x (vec1 y)) \<in> epigraph s f \<longleftrightarrow> x \<in> s \<and> f x \<le> y" unfolding epigraph_def by auto
+
+lemma convex_epigraph: 
+  "convex(epigraph s f) \<longleftrightarrow> convex_on s f \<and> convex s"
+  unfolding convex_def convex_on_def unfolding Ball_def forall_pastecart epigraph_def
+  unfolding mem_Collect_eq fstcart_pastecart sndcart_pastecart sndcart_add sndcart_cmul [where 'a=real, unfolded smult_conv_scaleR] fstcart_add fstcart_cmul [where 'a=real, unfolded smult_conv_scaleR]
+  unfolding Ball_def[symmetric] unfolding dest_vec1_add dest_vec1_cmul [where 'a=real, unfolded smult_conv_scaleR]
+  apply(subst forall_dest_vec1[THEN sym])+ by(meson real_le_refl real_le_trans add_mono mult_left_mono) 
+
+lemma convex_epigraphI: assumes "convex_on s f" "convex s"
+  shows "convex(epigraph s f)" using assms unfolding convex_epigraph by auto
+
+lemma convex_epigraph_convex: "convex s \<Longrightarrow> (convex_on s f \<longleftrightarrow> convex(epigraph s f))"
+  using convex_epigraph by auto
+
+subsection {* Use this to derive general bound property of convex function. *}
+
+lemma forall_of_pastecart:
+  "(\<forall>p. P (\<lambda>x. fstcart (p x)) (\<lambda>x. sndcart (p x))) \<longleftrightarrow> (\<forall>x y. P x y)" apply meson
+  apply(erule_tac x="\<lambda>a. pastecart (x a) (y a)" in allE) unfolding o_def by auto
+
+lemma forall_of_pastecart':
+  "(\<forall>p. P (fstcart p) (sndcart p)) \<longleftrightarrow> (\<forall>x y. P x y)" apply meson
+  apply(erule_tac x="pastecart x y" in allE) unfolding o_def by auto
+
+lemma forall_of_dest_vec1: "(\<forall>v. P (\<lambda>x. dest_vec1 (v x))) \<longleftrightarrow> (\<forall>x. P x)"
+  apply rule apply rule apply(erule_tac x="(vec1 \<circ> x)" in allE) unfolding o_def vec1_dest_vec1 by auto 
+
+lemma forall_of_dest_vec1': "(\<forall>v. P (dest_vec1 v)) \<longleftrightarrow> (\<forall>x. P x)"
+  apply rule apply rule apply(erule_tac x="(vec1 x)" in allE) defer apply rule 
+  apply(erule_tac x="dest_vec1 v" in allE) unfolding o_def vec1_dest_vec1 by auto
+
+lemma convex_on:
+  fixes s :: "(real ^ _) set"
+  assumes "convex s"
+  shows "convex_on s f \<longleftrightarrow> (\<forall>k u x. (\<forall>i\<in>{1..k::nat}. 0 \<le> u i \<and> x i \<in> s) \<and> setsum u {1..k} = 1 \<longrightarrow>
+   f (setsum (\<lambda>i. u i *\<^sub>R x i) {1..k} ) \<le> setsum (\<lambda>i. u i * f(x i)) {1..k} ) "
+  unfolding convex_epigraph_convex[OF assms] convex epigraph_def Ball_def mem_Collect_eq
+  unfolding sndcart_setsum[OF finite_atLeastAtMost] fstcart_setsum[OF finite_atLeastAtMost] dest_vec1_setsum[OF finite_atLeastAtMost]
+  unfolding fstcart_pastecart sndcart_pastecart sndcart_add sndcart_cmul [where 'a=real, unfolded smult_conv_scaleR] fstcart_add fstcart_cmul [where 'a=real, unfolded smult_conv_scaleR]
+  unfolding dest_vec1_add dest_vec1_cmul [where 'a=real, unfolded smult_conv_scaleR] apply(subst forall_of_pastecart)+ apply(subst forall_of_dest_vec1)+ apply rule
+  using assms[unfolded convex] apply simp apply(rule,rule,rule)
+  apply(erule_tac x=k in allE, erule_tac x=u in allE, erule_tac x=x in allE) apply rule apply rule apply rule defer
+  apply(rule_tac j="\<Sum>i = 1..k. u i * f (x i)" in real_le_trans)
+  defer apply(rule setsum_mono) apply(erule conjE)+ apply(erule_tac x=i in allE)apply(rule mult_left_mono)
+  using assms[unfolded convex] by auto
+
+subsection {* Convexity of general and special intervals. *}
+
+lemma is_interval_convex:
+  fixes s :: "(real ^ _) set"
+  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 \<in> s" "y \<in> s" "0 \<le> u" "0 \<le> v" "u + v = (1::real)"
+  hence *:"u = 1 - v" "1 - v \<ge> 0" and **:"v = 1 - u" "1 - u \<ge> 0" by auto
+  { fix a b assume "\<not> b \<le> u * a + v * b"
+    hence "u * a < (1 - v) * b" unfolding not_le using as(4) by(auto simp add: field_simps)
+    hence "a < b" unfolding * using as(4) *(2) apply(rule_tac mult_left_less_imp_less[of "1 - v"]) by(auto simp add: field_simps)
+    hence "a \<le> u * a + v * b" unfolding * using as(4) by (auto simp add: field_simps intro!:mult_right_mono)
+  } moreover
+  { fix a b assume "\<not> u * a + v * b \<le> a"
+    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 \<le> b" unfolding ** using **(2) as(3) by(auto simp add: field_simps intro!:mult_right_mono) }
+  ultimately show "u *\<^sub>R x + v *\<^sub>R y \<in> 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:
+  fixes s :: "(real ^ _) set"
+  shows "is_interval s \<Longrightarrow> connected s"
+  using is_interval_convex convex_connected by auto
+
+lemma convex_interval: "convex {a .. b}" "convex {a<..<b::real^'n::finite}"
+  apply(rule_tac[!] is_interval_convex) using is_interval_interval by auto
+
+subsection {* On @{text "real^1"}, @{text "is_interval"}, @{text "convex"} and @{text "connected"} are all equivalent. *}
+
+lemma is_interval_1:
+  "is_interval s \<longleftrightarrow> (\<forall>a\<in>s. \<forall>b\<in>s. \<forall> x. dest_vec1 a \<le> dest_vec1 x \<and> dest_vec1 x \<le> dest_vec1 b \<longrightarrow> x \<in> s)"
+  unfolding is_interval_def dest_vec1_def forall_1 by auto
+
+lemma is_interval_connected_1: "is_interval s \<longleftrightarrow> 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 \<in> s" "b \<in> s" "dest_vec1 a \<le> dest_vec1 x" "dest_vec1 x \<le> dest_vec1 b" "x\<notin>s"
+  hence *:"dest_vec1 a < dest_vec1 x" "dest_vec1 x < dest_vec1 b" apply(rule_tac [!] ccontr) unfolding not_less by auto
+  let ?halfl = "{z. inner (basis 1) z < dest_vec1 x} " and ?halfr = "{z. inner (basis 1) z > dest_vec1 x} "
+  { fix y assume "y \<in> s" have "y \<in> ?halfr \<union> ?halfl" apply(rule ccontr)
+    using as(6) `y\<in>s` by (auto simp add: inner_vector_def dest_vec1_eq [unfolded dest_vec1_def] dest_vec1_def) }
+  moreover have "a\<in>?halfl" "b\<in>?halfr" using * by (auto simp add: inner_vector_def dest_vec1_def)
+  hence "?halfl \<inter> s \<noteq> {}" "?halfr \<inter> s \<noteq> {}"  using as(2-3) by auto
+  ultimately show False apply(rule_tac notE[OF as(1)[unfolded connected_def]])
+    apply(rule_tac x="?halfl" in exI, rule_tac x="?halfr" in exI) 
+    apply(rule, rule open_halfspace_lt, rule, rule open_halfspace_gt) apply(rule, rule, rule ccontr)
+    by(auto simp add: basis_component field_simps) qed 
+
+lemma is_interval_convex_1:
+  "is_interval s \<longleftrightarrow> convex (s::(real^1) set)" 
+  using is_interval_convex convex_connected is_interval_connected_1 by auto
+
+lemma convex_connected_1:
+  "connected s \<longleftrightarrow> convex (s::(real^1) set)" 
+  using is_interval_convex convex_connected is_interval_connected_1 by auto
+
+subsection {* Another intermediate value theorem formulation. *}
+
+lemma ivt_increasing_component_on_1: fixes f::"real^1 \<Rightarrow> real^'n::finite"
+  assumes "dest_vec1 a \<le> dest_vec1 b" "continuous_on {a .. b} f" "(f a)$k \<le> y" "y \<le> (f b)$k"
+  shows "\<exists>x\<in>{a..b}. (f x)$k = y"
+proof- have "f a \<in> f ` {a..b}" "f b \<in> f ` {a..b}" apply(rule_tac[!] imageI) 
+    using assms(1) by(auto simp add: vector_less_eq_def dest_vec1_def)
+  thus ?thesis using connected_ivt_component[of "f ` {a..b}" "f a" "f b" k y]
+    using connected_continuous_image[OF assms(2) convex_connected[OF convex_interval(1)]]
+    using assms by(auto intro!: imageI) qed
+
+lemma ivt_increasing_component_1: fixes f::"real^1 \<Rightarrow> real^'n::finite"
+  assumes "dest_vec1 a \<le> dest_vec1 b"
+  "\<forall>x\<in>{a .. b}. continuous (at x) f" "f a$k \<le> y" "y \<le> f b$k"
+  shows "\<exists>x\<in>{a..b}. (f x)$k = y"
+  apply(rule ivt_increasing_component_on_1) using assms using continuous_at_imp_continuous_on by auto
+
+lemma ivt_decreasing_component_on_1: fixes f::"real^1 \<Rightarrow> real^'n::finite"
+  assumes "dest_vec1 a \<le> dest_vec1 b" "continuous_on {a .. b} f" "(f b)$k \<le> y" "y \<le> (f a)$k"
+  shows "\<exists>x\<in>{a..b}. (f x)$k = y"
+  apply(subst neg_equal_iff_equal[THEN sym]) unfolding vector_uminus_component[THEN sym]
+  apply(rule ivt_increasing_component_on_1) using assms using continuous_on_neg
+  by(auto simp add:vector_uminus_component)
+
+lemma ivt_decreasing_component_1: fixes f::"real^1 \<Rightarrow> real^'n::finite"
+  assumes "dest_vec1 a \<le> dest_vec1 b" "\<forall>x\<in>{a .. b}. continuous (at x) f" "f b$k \<le> y" "y \<le> f a$k"
+  shows "\<exists>x\<in>{a..b}. (f x)$k = y"
+  apply(rule ivt_decreasing_component_on_1) using assms using continuous_at_imp_continuous_on by auto
+
+subsection {* A bound within a convex hull, and so an interval. *}
+
+lemma convex_on_convex_hull_bound:
+  fixes s :: "(real ^ _) set"
+  assumes "convex_on (convex hull s) f" "\<forall>x\<in>s. f x \<le> b"
+  shows "\<forall>x\<in> convex hull s. f x \<le> b" proof
+  fix x assume "x\<in>convex hull s"
+  then obtain k u v where obt:"\<forall>i\<in>{1..k::nat}. 0 \<le> u i \<and> v i \<in> s" "setsum u {1..k} = 1" "(\<Sum>i = 1..k. u i *\<^sub>R v i) = x"
+    unfolding convex_hull_indexed mem_Collect_eq by auto
+  have "(\<Sum>i = 1..k. u i * f (v i)) \<le> b" using setsum_mono[of "{1..k}" "\<lambda>i. u i * f (v i)" "\<lambda>i. u i * b"]
+    unfolding setsum_left_distrib[THEN sym] obt(2) mult_1 apply(drule_tac meta_mp) apply(rule mult_left_mono)
+    using assms(2) obt(1) by auto
+  thus "f x \<le> b" using assms(1)[unfolded convex_on[OF convex_convex_hull], rule_format, of k u v]
+    unfolding obt(2-3) using obt(1) and hull_subset[unfolded subset_eq, rule_format, of _ s] by auto qed
+
+lemma unit_interval_convex_hull:
+  "{0::real^'n::finite .. 1} = convex hull {x. \<forall>i. (x$i = 0) \<or> (x$i = 1)}" (is "?int = convex hull ?points")
+proof- have 01:"{0,1} \<subseteq> convex hull ?points" apply rule apply(rule_tac hull_subset[unfolded subset_eq, rule_format]) by auto
+  { fix n x assume "x\<in>{0::real^'n .. 1}" "n \<le> CARD('n)" "card {i. x$i \<noteq> 0} \<le> n" 
+  hence "x\<in>convex hull ?points" proof(induct n arbitrary: x)
+    case 0 hence "x = 0" apply(subst Cart_eq) apply rule by auto
+    thus "x\<in>convex hull ?points" using 01 by auto
+  next
+    case (Suc n) show "x\<in>convex hull ?points" proof(cases "{i. x$i \<noteq> 0} = {}")
+      case True hence "x = 0" unfolding Cart_eq by auto
+      thus "x\<in>convex hull ?points" using 01 by auto
+    next
+      case False def xi \<equiv> "Min ((\<lambda>i. x$i) ` {i. x$i \<noteq> 0})"
+      have "xi \<in> (\<lambda>i. x$i) ` {i. x$i \<noteq> 0}" unfolding xi_def apply(rule Min_in) using False by auto
+      then obtain i where i':"x$i = xi" "x$i \<noteq> 0" by auto
+      have i:"\<And>j. x$j > 0 \<Longrightarrow> x$i \<le> x$j"
+        unfolding i'(1) xi_def apply(rule_tac Min_le) unfolding image_iff
+        defer apply(rule_tac x=j in bexI) using i' by auto
+      have i01:"x$i \<le> 1" "x$i > 0" using Suc(2)[unfolded mem_interval,rule_format,of i] using i'(2) `x$i \<noteq> 0`
+        by(auto simp add: Cart_lambda_beta) 
+      show ?thesis proof(cases "x$i=1")
+        case True have "\<forall>j\<in>{i. x$i \<noteq> 0}. x$j = 1" apply(rule, rule ccontr) unfolding mem_Collect_eq proof-
+          fix j assume "x $ j \<noteq> 0" "x $ j \<noteq> 1"
+          hence j:"x$j \<in> {0<..<1}" using Suc(2) by(auto simp add: vector_less_eq_def elim!:allE[where x=j])
+          hence "x$j \<in> op $ x ` {i. x $ i \<noteq> 0}" by auto 
+          hence "x$j \<ge> x$i" unfolding i'(1) xi_def apply(rule_tac Min_le) by auto
+          thus False using True Suc(2) j by(auto simp add: vector_less_eq_def elim!:ballE[where x=j]) qed
+        thus "x\<in>convex hull ?points" apply(rule_tac hull_subset[unfolded subset_eq, rule_format])
+          by(auto simp add: Cart_lambda_beta)
+      next let ?y = "\<lambda>j. if x$j = 0 then 0 else (x$j - x$i) / (1 - x$i)"
+        case False hence *:"x = x$i *\<^sub>R (\<chi> j. if x$j = 0 then 0 else 1) + (1 - x$i) *\<^sub>R (\<chi> j. ?y j)" unfolding Cart_eq
+          by(auto simp add: Cart_lambda_beta vector_add_component vector_smult_component vector_minus_component field_simps)
+        { fix j have "x$j \<noteq> 0 \<Longrightarrow> 0 \<le> (x $ j - x $ i) / (1 - x $ i)" "(x $ j - x $ i) / (1 - x $ i) \<le> 1"
+            apply(rule_tac divide_nonneg_pos) using i(1)[of j] using False i01
+            using Suc(2)[unfolded mem_interval, rule_format, of j] by(auto simp add:field_simps Cart_lambda_beta) 
+          hence "0 \<le> ?y j \<and> ?y j \<le> 1" by auto }
+        moreover have "i\<in>{j. x$j \<noteq> 0} - {j. ((\<chi> j. ?y j)::real^'n) $ j \<noteq> 0}" using i01 by(auto simp add: Cart_lambda_beta)
+        hence "{j. x$j \<noteq> 0} \<noteq> {j. ((\<chi> j. ?y j)::real^'n::finite) $ j \<noteq> 0}" by auto
+        hence **:"{j. ((\<chi> j. ?y j)::real^'n::finite) $ j \<noteq> 0} \<subset> {j. x$j \<noteq> 0}" apply - apply rule by(auto simp add: Cart_lambda_beta)  
+        have "card {j. ((\<chi> j. ?y j)::real^'n) $ j \<noteq> 0} \<le> n" using less_le_trans[OF psubset_card_mono[OF _ **] Suc(4)] by auto
+        ultimately show ?thesis apply(subst *) apply(rule convex_convex_hull[unfolded convex_def, rule_format])
+          apply(rule_tac hull_subset[unfolded subset_eq, rule_format]) defer apply(rule Suc(1))
+          unfolding mem_interval using i01 Suc(3) by (auto simp add: Cart_lambda_beta)
+      qed qed qed } note * = this
+  show ?thesis apply rule defer apply(rule hull_minimal) unfolding subset_eq prefer 3 apply rule 
+    apply(rule_tac n2="CARD('n)" in *) prefer 3 apply(rule card_mono) using 01 and convex_interval(1) prefer 5 apply - apply rule
+    unfolding mem_interval apply rule unfolding mem_Collect_eq apply(erule_tac x=i in allE)
+    by(auto simp add: vector_less_eq_def mem_def[of _ convex]) qed
+
+subsection {* And this is a finite set of vertices. *}
+
+lemma unit_cube_convex_hull: obtains s where "finite s" "{0 .. 1::real^'n::finite} = convex hull s"
+  apply(rule that[of "{x::real^'n::finite. \<forall>i. x$i=0 \<or> x$i=1}"])
+  apply(rule finite_subset[of _ "(\<lambda>s. (\<chi> i. if i\<in>s then 1::real else 0)::real^'n::finite) ` UNIV"])
+  prefer 3 apply(rule unit_interval_convex_hull) apply rule unfolding mem_Collect_eq proof-
+  fix x::"real^'n" assume as:"\<forall>i. x $ i = 0 \<or> x $ i = 1"
+  show "x \<in> (\<lambda>s. \<chi> i. if i \<in> s then 1 else 0) ` UNIV" apply(rule image_eqI[where x="{i. x$i = 1}"])
+    unfolding Cart_eq using as by(auto simp add:Cart_lambda_beta) qed auto
+
+subsection {* Hence any cube (could do any nonempty interval). *}
+
+lemma cube_convex_hull:
+  assumes "0 < d" obtains s::"(real^'n::finite) set" where "finite s" "{x - (\<chi> i. d) .. x + (\<chi> i. d)} = convex hull s" proof-
+  let ?d = "(\<chi> i. d)::real^'n"
+  have *:"{x - ?d .. x + ?d} = (\<lambda>y. x - ?d + (2 * d) *\<^sub>R y) ` {0 .. 1}" apply(rule set_ext, rule)
+    unfolding image_iff defer apply(erule bexE) proof-
+    fix y assume as:"y\<in>{x - ?d .. x + ?d}"
+    { fix i::'n have "x $ i \<le> d + y $ i" "y $ i \<le> d + x $ i" using as[unfolded mem_interval, THEN spec[where x=i]]
+        by(auto simp add: vector_component)
+      hence "1 \<ge> inverse d * (x $ i - y $ i)" "1 \<ge> inverse d * (y $ i - x $ i)"
+        apply(rule_tac[!] mult_left_le_imp_le[OF _ assms]) unfolding mult_assoc[THEN sym]
+        using assms by(auto simp add: field_simps right_inverse) 
+      hence "inverse d * (x $ i * 2) \<le> 2 + inverse d * (y $ i * 2)"
+            "inverse d * (y $ i * 2) \<le> 2 + inverse d * (x $ i * 2)" by(auto simp add:field_simps) }
+    hence "inverse (2 * d) *\<^sub>R (y - (x - ?d)) \<in> {0..1}" unfolding mem_interval using assms
+      by(auto simp add: Cart_eq vector_component_simps field_simps)
+    thus "\<exists>z\<in>{0..1}. y = x - ?d + (2 * d) *\<^sub>R z" apply- apply(rule_tac x="inverse (2 * d) *\<^sub>R (y - (x - ?d))" in bexI) 
+      using assms by(auto simp add: Cart_eq vector_less_eq_def Cart_lambda_beta)
+  next
+    fix y z assume as:"z\<in>{0..1}" "y = x - ?d + (2*d) *\<^sub>R z" 
+    have "\<And>i. 0 \<le> d * z $ i \<and> d * z $ i \<le> d" using assms as(1)[unfolded mem_interval] apply(erule_tac x=i in allE)
+      apply rule apply(rule mult_nonneg_nonneg) prefer 3 apply(rule mult_right_le_one_le)
+      using assms by(auto simp add: vector_component_simps Cart_eq)
+    thus "y \<in> {x - ?d..x + ?d}" unfolding as(2) mem_interval apply- apply rule using as(1)[unfolded mem_interval]
+      apply(erule_tac x=i in allE) using assms by(auto simp add:  vector_component_simps Cart_eq) qed
+  obtain s where "finite s" "{0..1::real^'n} = convex hull s" using unit_cube_convex_hull by auto
+  thus ?thesis apply(rule_tac that[of "(\<lambda>y. x - ?d + (2 * d) *\<^sub>R y)` s"]) unfolding * and convex_hull_affinity by auto qed
+
+subsection {* Bounded convex function on open set is continuous. *}
+
+lemma convex_on_bounded_continuous:
+  fixes s :: "(real ^ _) set"
+  assumes "open s" "convex_on s f" "\<forall>x\<in>s. abs(f x) \<le> b"
+  shows "continuous_on s f"
+  apply(rule continuous_at_imp_continuous_on) unfolding continuous_at_real_range proof(rule,rule,rule)
+  fix x e assume "x\<in>s" "(0::real) < e"
+  def B \<equiv> "abs b + 1"
+  have B:"0 < B" "\<And>x. x\<in>s \<Longrightarrow> abs (f x) \<le> B"
+    unfolding B_def defer apply(drule assms(3)[rule_format]) by auto
+  obtain k where "k>0"and k:"cball x k \<subseteq> s" using assms(1)[unfolded open_contains_cball, THEN bspec[where x=x]] using `x\<in>s` by auto
+  show "\<exists>d>0. \<forall>x'. norm (x' - x) < d \<longrightarrow> \<bar>f x' - f x\<bar> < e"
+    apply(rule_tac x="min (k / 2) (e / (2 * B) * k)" in exI) apply rule defer proof(rule,rule)
+    fix y assume as:"norm (y - x) < min (k / 2) (e / (2 * B) * k)" 
+    show "\<bar>f y - f x\<bar> < e" proof(cases "y=x")
+      case False def t \<equiv> "k / norm (y - x)"
+      have "2 < t" "0<t" unfolding t_def using as False and `k>0` by(auto simp add:field_simps)
+      have "y\<in>s" apply(rule k[unfolded subset_eq,rule_format]) unfolding mem_cball dist_norm
+        apply(rule order_trans[of _ "2 * norm (x - y)"]) using as by(auto simp add: field_simps norm_minus_commute) 
+      { def w \<equiv> "x + t *\<^sub>R (y - x)"
+        have "w\<in>s" unfolding w_def apply(rule k[unfolded subset_eq,rule_format]) unfolding mem_cball dist_norm 
+          unfolding t_def using `k>0` by auto
+        have "(1 / t) *\<^sub>R x + - x + ((t - 1) / t) *\<^sub>R x = (1 / t - 1 + (t - 1) / t) *\<^sub>R x" by (auto simp add: algebra_simps)
+        also have "\<dots> = 0"  using `t>0` by(auto simp add:field_simps)
+        finally have w:"(1 / t) *\<^sub>R w + ((t - 1) / t) *\<^sub>R x = y" unfolding w_def using False and `t>0` by (auto simp add: algebra_simps)
+        have  "2 * B < e * t" unfolding t_def using `0<e` `0<k` `B>0` and as and False by (auto simp add:field_simps) 
+        hence "(f w - f x) / t < e"
+          using B(2)[OF `w\<in>s`] and B(2)[OF `x\<in>s`] using `t>0` by(auto simp add:field_simps) 
+        hence th1:"f y - f x < e" apply- apply(rule le_less_trans) defer apply assumption
+          using assms(2)[unfolded convex_on_def,rule_format,of w x "1/t" "(t - 1)/t", unfolded w]
+          using `0<t` `2<t` and `x\<in>s` `w\<in>s` by(auto simp add:field_simps) }
+      moreover 
+      { def w \<equiv> "x - t *\<^sub>R (y - x)"
+        have "w\<in>s" unfolding w_def apply(rule k[unfolded subset_eq,rule_format]) unfolding mem_cball dist_norm 
+          unfolding t_def using `k>0` by auto
+        have "(1 / (1 + t)) *\<^sub>R x + (t / (1 + t)) *\<^sub>R x = (1 / (1 + t) + t / (1 + t)) *\<^sub>R x" by (auto simp add: algebra_simps)
+        also have "\<dots>=x" using `t>0` by (auto simp add:field_simps)
+        finally have w:"(1 / (1+t)) *\<^sub>R w + (t / (1 + t)) *\<^sub>R y = x" unfolding w_def using False and `t>0` by (auto simp add: algebra_simps)
+        have  "2 * B < e * t" unfolding t_def using `0<e` `0<k` `B>0` and as and False by (auto simp add:field_simps) 
+        hence *:"(f w - f y) / t < e" using B(2)[OF `w\<in>s`] and B(2)[OF `y\<in>s`] using `t>0` by(auto simp add:field_simps) 
+        have "f x \<le> 1 / (1 + t) * f w + (t / (1 + t)) * f y" 
+          using assms(2)[unfolded convex_on_def,rule_format,of w y "1/(1+t)" "t / (1+t)",unfolded w]
+          using `0<t` `2<t` and `y\<in>s` `w\<in>s` by (auto simp add:field_simps)
+        also have "\<dots> = (f w + t * f y) / (1 + t)" using `t>0` unfolding real_divide_def by (auto simp add:field_simps)
+        also have "\<dots> < e + f y" using `t>0` * `e>0` by(auto simp add:field_simps)
+        finally have "f x - f y < e" by auto }
+      ultimately show ?thesis by auto 
+    qed(insert `0<e`, auto) 
+  qed(insert `0<e` `0<k` `0<B`, auto simp add:field_simps intro!:mult_pos_pos) qed
+
+subsection {* Upper bound on a ball implies upper and lower bounds. *}
+
+lemma convex_bounds_lemma:
+  fixes x :: "real ^ _"
+  assumes "convex_on (cball x e) f"  "\<forall>y \<in> cball x e. f y \<le> b"
+  shows "\<forall>y \<in> cball x e. abs(f y) \<le> b + 2 * abs(f x)"
+  apply(rule) proof(cases "0 \<le> e") case True
+  fix y assume y:"y\<in>cball x e" def z \<equiv> "2 *\<^sub>R x - y"
+  have *:"x - (2 *\<^sub>R x - y) = y - x" by vector
+  have z:"z\<in>cball x e" using y unfolding z_def mem_cball dist_norm * by(auto simp add: norm_minus_commute)
+  have "(1 / 2) *\<^sub>R y + (1 / 2) *\<^sub>R z = x" unfolding z_def by (auto simp add: algebra_simps)
+  thus "\<bar>f y\<bar> \<le> b + 2 * \<bar>f x\<bar>" using assms(1)[unfolded convex_on_def,rule_format, OF y z, of "1/2" "1/2"]
+    using assms(2)[rule_format,OF y] assms(2)[rule_format,OF z] by(auto simp add:field_simps)
+next case False fix y assume "y\<in>cball x e" 
+  hence "dist x y < 0" using False unfolding mem_cball not_le by (auto simp del: dist_not_less_zero)
+  thus "\<bar>f y\<bar> \<le> b + 2 * \<bar>f x\<bar>" using zero_le_dist[of x y] by auto qed
+
+subsection {* Hence a convex function on an open set is continuous. *}
+
+lemma convex_on_continuous:
+  assumes "open (s::(real^'n::finite) set)" "convex_on s f" 
+  shows "continuous_on s f"
+  unfolding continuous_on_eq_continuous_at[OF assms(1)] proof
+  note dimge1 = dimindex_ge_1[where 'a='n]
+  fix x assume "x\<in>s"
+  then obtain e where e:"cball x e \<subseteq> s" "e>0" using assms(1) unfolding open_contains_cball by auto
+  def d \<equiv> "e / real CARD('n)"
+  have "0 < d" unfolding d_def using `e>0` dimge1 by(rule_tac divide_pos_pos, auto) 
+  let ?d = "(\<chi> i. d)::real^'n"
+  obtain c where c:"finite c" "{x - ?d..x + ?d} = convex hull c" using cube_convex_hull[OF `d>0`, of x] by auto
+  have "x\<in>{x - ?d..x + ?d}" using `d>0` unfolding mem_interval by(auto simp add:vector_component_simps)
+  hence "c\<noteq>{}" apply(rule_tac ccontr) using c by(auto simp add:convex_hull_empty)
+  def k \<equiv> "Max (f ` c)"
+  have "convex_on {x - ?d..x + ?d} f" apply(rule convex_on_subset[OF assms(2)])
+    apply(rule subset_trans[OF _ e(1)]) unfolding subset_eq mem_cball proof 
+    fix z assume z:"z\<in>{x - ?d..x + ?d}"
+    have e:"e = setsum (\<lambda>i. d) (UNIV::'n set)" unfolding setsum_constant d_def using dimge1
+      by (metis card_enum field_simps d_def not_one_le_zero of_nat_le_iff real_eq_of_nat real_of_nat_1)
+    show "dist x z \<le> e" unfolding dist_norm e apply(rule_tac order_trans[OF norm_le_l1], rule setsum_mono)
+      using z[unfolded mem_interval] apply(erule_tac x=i in allE) by(auto simp add:field_simps vector_component_simps) qed
+  hence k:"\<forall>y\<in>{x - ?d..x + ?d}. f y \<le> k" unfolding c(2) apply(rule_tac convex_on_convex_hull_bound) apply assumption
+    unfolding k_def apply(rule, rule Max_ge) using c(1) by auto
+  have "d \<le> e" unfolding d_def apply(rule mult_imp_div_pos_le) using `e>0` dimge1 unfolding mult_le_cancel_left1 using real_dimindex_ge_1 by auto
+  hence dsube:"cball x d \<subseteq> cball x e" unfolding subset_eq Ball_def mem_cball by auto
+  have conv:"convex_on (cball x d) f" apply(rule convex_on_subset, rule convex_on_subset[OF assms(2)]) apply(rule e(1)) using dsube by auto
+  hence "\<forall>y\<in>cball x d. abs (f y) \<le> k + 2 * abs (f x)" apply(rule_tac convex_bounds_lemma) apply assumption proof
+    fix y assume y:"y\<in>cball x d"
+    { fix i::'n have "x $ i - d \<le> y $ i"  "y $ i \<le> x $ i + d" 
+        using order_trans[OF component_le_norm y[unfolded mem_cball dist_norm], of i] by(auto simp add: vector_component)  }
+    thus "f y \<le> k" apply(rule_tac k[rule_format]) unfolding mem_cball mem_interval dist_norm 
+      by(auto simp add: vector_component_simps) qed
+  hence "continuous_on (ball x d) f" apply(rule_tac convex_on_bounded_continuous)
+    apply(rule open_ball, rule convex_on_subset[OF conv], rule ball_subset_cball) by auto
+  thus "continuous (at x) f" unfolding continuous_on_eq_continuous_at[OF open_ball] using `d>0` by auto qed
+
+subsection {* Line segments, starlike sets etc.                                         *)
+(* Use the same overloading tricks as for intervals, so that                 *)
+(* segment[a,b] is closed and segment(a,b) is open relative to affine hull. *}
+
+definition
+  midpoint :: "real ^ 'n::finite \<Rightarrow> real ^ 'n \<Rightarrow> real ^ 'n" where
+  "midpoint a b = (inverse (2::real)) *\<^sub>R (a + b)"
+
+definition
+  open_segment :: "real ^ 'n::finite \<Rightarrow> real ^ 'n \<Rightarrow> (real ^ 'n) set" where
+  "open_segment a b = {(1 - u) *\<^sub>R a + u *\<^sub>R b | u::real.  0 < u \<and> u < 1}"
+
+definition
+  closed_segment :: "real ^ 'n::finite \<Rightarrow> real ^ 'n \<Rightarrow> (real ^ 'n) set" where
+  "closed_segment a b = {(1 - u) *\<^sub>R a + u *\<^sub>R b | u::real. 0 \<le> u \<and> u \<le> 1}"
+
+definition "between = (\<lambda> (a,b). closed_segment a b)"
+
+lemmas segment = open_segment_def closed_segment_def
+
+definition "starlike s \<longleftrightarrow> (\<exists>a\<in>s. \<forall>x\<in>s. closed_segment a x \<subseteq> s)"
+
+lemma midpoint_refl: "midpoint x x = x"
+  unfolding midpoint_def unfolding scaleR_right_distrib unfolding scaleR_left_distrib[THEN sym] by auto
+
+lemma midpoint_sym: "midpoint a b = midpoint b a" unfolding midpoint_def by (auto simp add: scaleR_right_distrib)
+
+lemma dist_midpoint:
+  "dist a (midpoint a b) = (dist a b) / 2" (is ?t1)
+  "dist b (midpoint a b) = (dist a b) / 2" (is ?t2)
+  "dist (midpoint a b) a = (dist a b) / 2" (is ?t3)
+  "dist (midpoint a b) b = (dist a b) / 2" (is ?t4)
+proof-
+  have *: "\<And>x y::real^'n::finite. 2 *\<^sub>R x = - y \<Longrightarrow> norm x = (norm y) / 2" unfolding equation_minus_iff by auto
+  have **:"\<And>x y::real^'n::finite. 2 *\<^sub>R x =   y \<Longrightarrow> norm x = (norm y) / 2" by auto
+  note scaleR_right_distrib [simp]
+  show ?t1 unfolding midpoint_def dist_norm apply (rule **) by(auto,vector)
+  show ?t2 unfolding midpoint_def dist_norm apply (rule *)  by(auto,vector)
+  show ?t3 unfolding midpoint_def dist_norm apply (rule *)  by(auto,vector)
+  show ?t4 unfolding midpoint_def dist_norm apply (rule **) by(auto,vector) qed
+
+lemma midpoint_eq_endpoint:
+  "midpoint a b = a \<longleftrightarrow> a = (b::real^'n::finite)"
+  "midpoint a b = b \<longleftrightarrow> a = b"
+  unfolding dist_eq_0_iff[where 'a="real^'n", THEN sym] dist_midpoint by auto
+
+lemma convex_contains_segment:
+  "convex s \<longleftrightarrow> (\<forall>a\<in>s. \<forall>b\<in>s. closed_segment a b \<subseteq> s)"
+  unfolding convex_alt closed_segment_def by auto
+
+lemma convex_imp_starlike:
+  "convex s \<Longrightarrow> s \<noteq> {} \<Longrightarrow> starlike s"
+  unfolding convex_contains_segment starlike_def by auto
+
+lemma segment_convex_hull:
+ "closed_segment a b = convex hull {a,b}" proof-
+  have *:"\<And>x. {x} \<noteq> {}" by auto
+  have **:"\<And>u v. u + v = 1 \<longleftrightarrow> u = 1 - (v::real)" by auto
+  show ?thesis unfolding segment convex_hull_insert[OF *] convex_hull_singleton apply(rule set_ext)
+    unfolding mem_Collect_eq apply(rule,erule exE) 
+    apply(rule_tac x="1 - u" in exI) apply rule defer apply(rule_tac x=u in exI) defer
+    apply(erule exE, (erule conjE)?)+ apply(rule_tac x="1 - u" in exI) unfolding ** by auto qed
+
+lemma convex_segment: "convex (closed_segment a b)"
+  unfolding segment_convex_hull by(rule convex_convex_hull)
+
+lemma ends_in_segment: "a \<in> closed_segment a b" "b \<in> closed_segment a b"
+  unfolding segment_convex_hull apply(rule_tac[!] hull_subset[unfolded subset_eq, rule_format]) by auto
+
+lemma segment_furthest_le:
+  assumes "x \<in> closed_segment a b" shows "norm(y - x) \<le> norm(y - a) \<or>  norm(y - x) \<le> norm(y - b)" proof-
+  obtain z where "z\<in>{a, b}" "norm (x - y) \<le> norm (z - y)" using simplex_furthest_le[of "{a, b}" y]
+    using assms[unfolded segment_convex_hull] by auto
+  thus ?thesis by(auto simp add:norm_minus_commute) qed
+
+lemma segment_bound:
+  assumes "x \<in> closed_segment a b"
+  shows "norm(x - a) \<le> norm(b - a)" "norm(x - b) \<le> norm(b - a)"
+  using segment_furthest_le[OF assms, of a]
+  using segment_furthest_le[OF assms, of b]
+  by (auto simp add:norm_minus_commute) 
+
+lemma segment_refl:"closed_segment a a = {a}" unfolding segment by (auto simp add: algebra_simps)
+
+lemma between_mem_segment: "between (a,b) x \<longleftrightarrow> x \<in> closed_segment a b"
+  unfolding between_def mem_def by auto
+
+lemma between:"between (a,b) (x::real^'n::finite) \<longleftrightarrow> dist a b = (dist a x) + (dist x b)"
+proof(cases "a = b")
+  case True thus ?thesis unfolding between_def split_conv mem_def[of x, symmetric]
+    by(auto simp add:segment_refl dist_commute) next
+  case False hence Fal:"norm (a - b) \<noteq> 0" and Fal2: "norm (a - b) > 0" by auto 
+  have *:"\<And>u. a - ((1 - u) *\<^sub>R a + u *\<^sub>R b) = u *\<^sub>R (a - b)" by (auto simp add: algebra_simps)
+  show ?thesis unfolding between_def split_conv mem_def[of x, symmetric] closed_segment_def mem_Collect_eq
+    apply rule apply(erule exE, (erule conjE)+) apply(subst dist_triangle_eq) proof-
+      fix u assume as:"x = (1 - u) *\<^sub>R a + u *\<^sub>R b" "0 \<le> u" "u \<le> 1" 
+      hence *:"a - x = u *\<^sub>R (a - b)" "x - b = (1 - u) *\<^sub>R (a - b)"
+        unfolding as(1) by(auto simp add:algebra_simps)
+      show "norm (a - x) *\<^sub>R (x - b) = norm (x - b) *\<^sub>R (a - x)"
+        unfolding norm_minus_commute[of x a] * Cart_eq using as(2,3)
+        by(auto simp add: vector_component_simps field_simps)
+    next assume as:"dist a b = dist a x + dist x b"
+      have "norm (a - x) / norm (a - b) \<le> 1" unfolding divide_le_eq_1_pos[OF Fal2] unfolding as[unfolded dist_norm] norm_ge_zero by auto 
+      thus "\<exists>u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \<and> 0 \<le> u \<and> u \<le> 1" apply(rule_tac x="dist a x / dist a b" in exI)
+        unfolding dist_norm Cart_eq apply- apply rule defer apply(rule, rule divide_nonneg_pos) prefer 4 proof rule
+          fix i::'n have "((1 - norm (a - x) / norm (a - b)) *\<^sub>R a + (norm (a - x) / norm (a - b)) *\<^sub>R b) $ i =
+            ((norm (a - b) - norm (a - x)) * (a $ i) + norm (a - x) * (b $ i)) / norm (a - b)"
+            using Fal by(auto simp add:vector_component_simps field_simps)
+          also have "\<dots> = x$i" apply(rule divide_eq_imp[OF Fal])
+            unfolding as[unfolded dist_norm] using as[unfolded dist_triangle_eq Cart_eq,rule_format, of i]
+            by(auto simp add:field_simps vector_component_simps)
+          finally show "x $ i = ((1 - norm (a - x) / norm (a - b)) *\<^sub>R a + (norm (a - x) / norm (a - b)) *\<^sub>R b) $ i" by auto
+        qed(insert Fal2, auto) qed qed
+
+lemma between_midpoint: fixes a::"real^'n::finite" shows
+  "between (a,b) (midpoint a b)" (is ?t1) 
+  "between (b,a) (midpoint a b)" (is ?t2)
+proof- have *:"\<And>x y z. x = (1/2::real) *\<^sub>R z \<Longrightarrow> y = (1/2) *\<^sub>R z \<Longrightarrow> norm z = norm x + norm y" by auto
+  show ?t1 ?t2 unfolding between midpoint_def dist_norm apply(rule_tac[!] *)
+    by(auto simp add:field_simps Cart_eq vector_component_simps) qed
+
+lemma between_mem_convex_hull:
+  "between (a,b) x \<longleftrightarrow> x \<in> convex hull {a,b}"
+  unfolding between_mem_segment segment_convex_hull ..
+
+subsection {* Shrinking towards the interior of a convex set. *}
+
+lemma mem_interior_convex_shrink:
+  fixes s :: "(real ^ _) set"
+  assumes "convex s" "c \<in> interior s" "x \<in> s" "0 < e" "e \<le> 1"
+  shows "x - e *\<^sub>R (x - c) \<in> interior s"
+proof- obtain d where "d>0" and d:"ball c d \<subseteq> s" using assms(2) unfolding mem_interior by auto
+  show ?thesis unfolding mem_interior apply(rule_tac x="e*d" in exI)
+    apply(rule) defer unfolding subset_eq Ball_def mem_ball proof(rule,rule)
+    fix y assume as:"dist (x - e *\<^sub>R (x - c)) y < e * d"
+    have *:"y = (1 - (1 - e)) *\<^sub>R ((1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x) + (1 - e) *\<^sub>R x" using `e>0` by (auto simp add: scaleR_left_diff_distrib scaleR_right_diff_distrib)
+    have "dist c ((1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x) = abs(1/e) * norm (e *\<^sub>R c - y + (1 - e) *\<^sub>R x)"
+      unfolding dist_norm unfolding norm_scaleR[THEN sym] apply(rule norm_eqI) using `e>0`
+      by(auto simp add:vector_component_simps Cart_eq field_simps) 
+    also have "\<dots> = abs(1/e) * norm (x - e *\<^sub>R (x - c) - y)" by(auto intro!:norm_eqI simp add: algebra_simps)
+    also have "\<dots> < d" using as[unfolded dist_norm] and `e>0`
+      by(auto simp add:pos_divide_less_eq[OF `e>0`] real_mult_commute)
+    finally show "y \<in> s" apply(subst *) apply(rule assms(1)[unfolded convex_alt,rule_format])
+      apply(rule d[unfolded subset_eq,rule_format]) unfolding mem_ball using assms(3-5) by auto
+  qed(rule mult_pos_pos, insert `e>0` `d>0`, auto) qed
+
+lemma mem_interior_closure_convex_shrink:
+  fixes s :: "(real ^ _) set"
+  assumes "convex s" "c \<in> interior s" "x \<in> closure s" "0 < e" "e \<le> 1"
+  shows "x - e *\<^sub>R (x - c) \<in> interior s"
+proof- obtain d where "d>0" and d:"ball c d \<subseteq> s" using assms(2) unfolding mem_interior by auto
+  have "\<exists>y\<in>s. norm (y - x) * (1 - e) < e * d" proof(cases "x\<in>s")
+    case True thus ?thesis using `e>0` `d>0` by(rule_tac bexI[where x=x], auto intro!: mult_pos_pos) next
+    case False hence x:"x islimpt s" using assms(3)[unfolded closure_def] by auto
+    show ?thesis proof(cases "e=1")
+      case True obtain y where "y\<in>s" "y \<noteq> x" "dist y x < 1"
+        using x[unfolded islimpt_approachable,THEN spec[where x=1]] by auto
+      thus ?thesis apply(rule_tac x=y in bexI) unfolding True using `d>0` by auto next
+      case False hence "0 < e * d / (1 - e)" and *:"1 - e > 0"
+        using `e\<le>1` `e>0` `d>0` by(auto intro!:mult_pos_pos divide_pos_pos)
+      then obtain y where "y\<in>s" "y \<noteq> x" "dist y x < e * d / (1 - e)"
+        using x[unfolded islimpt_approachable,THEN spec[where x="e*d / (1 - e)"]] by auto
+      thus ?thesis apply(rule_tac x=y in bexI) unfolding dist_norm using pos_less_divide_eq[OF *] by auto qed qed
+  then obtain y where "y\<in>s" and y:"norm (y - x) * (1 - e) < e * d" by auto
+  def z \<equiv> "c + ((1 - e) / e) *\<^sub>R (x - y)"
+  have *:"x - e *\<^sub>R (x - c) = y - e *\<^sub>R (y - z)" unfolding z_def using `e>0` by (auto simp add: scaleR_right_diff_distrib scaleR_right_distrib scaleR_left_diff_distrib)
+  have "z\<in>interior s" apply(rule subset_interior[OF d,unfolded subset_eq,rule_format])
+    unfolding interior_open[OF open_ball] mem_ball z_def dist_norm using y and assms(4,5)
+    by(auto simp add:field_simps norm_minus_commute)
+  thus ?thesis unfolding * apply - apply(rule mem_interior_convex_shrink) 
+    using assms(1,4-5) `y\<in>s` by auto qed
+
+subsection {* Some obvious but surprisingly hard simplex lemmas. *}
+
+lemma simplex:
+  assumes "finite s" "0 \<notin> s"
+  shows "convex hull (insert 0 s) =  { y. (\<exists>u. (\<forall>x\<in>s. 0 \<le> u x) \<and> setsum u s \<le> 1 \<and> setsum (\<lambda>x. u x *\<^sub>R x) s = y)}"
+  unfolding convex_hull_finite[OF finite.insertI[OF assms(1)]] apply(rule set_ext, rule) unfolding mem_Collect_eq
+  apply(erule_tac[!] exE) apply(erule_tac[!] conjE)+ unfolding setsum_clauses(2)[OF assms(1)]
+  apply(rule_tac x=u in exI) defer apply(rule_tac x="\<lambda>x. if x = 0 then 1 - setsum u s else u x" in exI) using assms(2)
+  unfolding if_smult and setsum_delta_notmem[OF assms(2)] by auto
+
+lemma std_simplex:
+  "convex hull (insert 0 { basis i | i. i\<in>UNIV}) =
+        {x::real^'n::finite . (\<forall>i. 0 \<le> x$i) \<and> setsum (\<lambda>i. x$i) UNIV \<le> 1 }" (is "convex hull (insert 0 ?p) = ?s")
+proof- let ?D = "UNIV::'n set"
+  have "0\<notin>?p" by(auto simp add: basis_nonzero)
+  have "{(basis i)::real^'n |i. i \<in> ?D} = basis ` ?D" by auto
+  note sumbas = this  setsum_reindex[OF basis_inj, unfolded o_def]
+  show ?thesis unfolding simplex[OF finite_stdbasis `0\<notin>?p`] apply(rule set_ext) unfolding mem_Collect_eq apply rule
+    apply(erule exE, (erule conjE)+) apply(erule_tac[2] conjE)+ proof-
+    fix x::"real^'n" and u assume as: "\<forall>x\<in>{basis i |i. i \<in>?D}. 0 \<le> u x" "setsum u {basis i |i. i \<in> ?D} \<le> 1" "(\<Sum>x\<in>{basis i |i. i \<in>?D}. u x *\<^sub>R x) = x"
+    have *:"\<forall>i. u (basis i) = x$i" using as(3) unfolding sumbas and basis_expansion_unique [where 'a=real, unfolded smult_conv_scaleR] by auto
+    hence **:"setsum u {basis i |i. i \<in> ?D} = setsum (op $ x) ?D" unfolding sumbas by(rule_tac setsum_cong, auto)
+    show " (\<forall>i. 0 \<le> x $ i) \<and> setsum (op $ x) ?D \<le> 1" apply - proof(rule,rule)
+      fix i::'n show "0 \<le> x$i" unfolding *[rule_format,of i,THEN sym] apply(rule_tac as(1)[rule_format]) by auto
+    qed(insert as(2)[unfolded **], auto)
+  next fix x::"real^'n" assume as:"\<forall>i. 0 \<le> x $ i" "setsum (op $ x) ?D \<le> 1"
+    show "\<exists>u. (\<forall>x\<in>{basis i |i. i \<in> ?D}. 0 \<le> u x) \<and> setsum u {basis i |i. i \<in> ?D} \<le> 1 \<and> (\<Sum>x\<in>{basis i |i. i \<in> ?D}. u x *\<^sub>R x) = x"
+      apply(rule_tac x="\<lambda>y. inner y x" in exI) apply(rule,rule) unfolding mem_Collect_eq apply(erule exE) using as(1) apply(erule_tac x=i in allE) 
+      unfolding sumbas using as(2) and basis_expansion_unique [where 'a=real, unfolded smult_conv_scaleR] by(auto simp add:inner_basis) qed qed 
+
+lemma interior_std_simplex:
+  "interior (convex hull (insert 0 { basis i| i. i\<in>UNIV})) =
+  {x::real^'n::finite. (\<forall>i. 0 < x$i) \<and> setsum (\<lambda>i. x$i) UNIV < 1 }"
+  apply(rule set_ext) unfolding mem_interior std_simplex unfolding subset_eq mem_Collect_eq Ball_def mem_ball
+  unfolding Ball_def[symmetric] apply rule apply(erule exE, (erule conjE)+) defer apply(erule conjE) proof-
+  fix x::"real^'n" and e assume "0<e" and as:"\<forall>xa. dist x xa < e \<longrightarrow> (\<forall>x. 0 \<le> xa $ x) \<and> setsum (op $ xa) UNIV \<le> 1"
+  show "(\<forall>xa. 0 < x $ xa) \<and> setsum (op $ x) UNIV < 1" apply(rule,rule) proof-
+    fix i::'n show "0 < x $ i" using as[THEN spec[where x="x - (e / 2) *\<^sub>R basis i"]] and `e>0`
+      unfolding dist_norm by(auto simp add: norm_basis vector_component_simps basis_component elim:allE[where x=i])
+  next guess a using UNIV_witness[where 'a='n] ..
+    have **:"dist x (x + (e / 2) *\<^sub>R basis a) < e" using  `e>0` and norm_basis[of a]
+      unfolding dist_norm by(auto simp add: vector_component_simps basis_component intro!: mult_strict_left_mono_comm)
+    have "\<And>i. (x + (e / 2) *\<^sub>R basis a) $ i = x$i + (if i = a then e/2 else 0)" by(auto simp add:vector_component_simps)
+    hence *:"setsum (op $ (x + (e / 2) *\<^sub>R basis a)) UNIV = setsum (\<lambda>i. x$i + (if a = i then e/2 else 0)) UNIV" by(rule_tac setsum_cong, auto) 
+    have "setsum (op $ x) UNIV < setsum (op $ (x + (e / 2) *\<^sub>R basis a)) UNIV" unfolding * setsum_addf
+      using `0<e` dimindex_ge_1 by(auto simp add: setsum_delta')
+    also have "\<dots> \<le> 1" using ** apply(drule_tac as[rule_format]) by auto
+    finally show "setsum (op $ x) UNIV < 1" by auto qed
+next
+  fix x::"real^'n::finite" assume as:"\<forall>i. 0 < x $ i" "setsum (op $ x) UNIV < 1"
+  guess a using UNIV_witness[where 'a='b] ..
+  let ?d = "(1 - setsum (op $ x) UNIV) / real (CARD('n))"
+  have "Min ((op $ x) ` UNIV) > 0" apply(rule Min_grI) using as(1) dimindex_ge_1 by auto
+  moreover have"?d > 0" apply(rule divide_pos_pos) using as(2) using dimindex_ge_1 by(auto simp add: Suc_le_eq)
+  ultimately show "\<exists>e>0. \<forall>y. dist x y < e \<longrightarrow> (\<forall>i. 0 \<le> y $ i) \<and> setsum (op $ y) UNIV \<le> 1"
+    apply(rule_tac x="min (Min ((op $ x) ` UNIV)) ?D" in exI) apply rule defer apply(rule,rule) proof-
+    fix y assume y:"dist x y < min (Min (op $ x ` UNIV)) ?d"
+    have "setsum (op $ y) UNIV \<le> setsum (\<lambda>i. x$i + ?d) UNIV" proof(rule setsum_mono)
+      fix i::'n have "abs (y$i - x$i) < ?d" apply(rule le_less_trans) using component_le_norm[of "y - x" i]
+        using y[unfolded min_less_iff_conj dist_norm, THEN conjunct2] by(auto simp add:vector_component_simps norm_minus_commute)
+      thus "y $ i \<le> x $ i + ?d" by auto qed
+    also have "\<dots> \<le> 1" unfolding setsum_addf setsum_constant card_enum real_eq_of_nat using dimindex_ge_1 by(auto simp add: Suc_le_eq)
+    finally show "(\<forall>i. 0 \<le> y $ i) \<and> setsum (op $ y) UNIV \<le> 1" apply- proof(rule,rule)
+      fix i::'n have "norm (x - y) < x$i" using y[unfolded min_less_iff_conj dist_norm, THEN conjunct1]
+        using Min_gr_iff[of "op $ x ` dimset x"] dimindex_ge_1 by auto
+      thus "0 \<le> y$i" using component_le_norm[of "x - y" i] and as(1)[rule_format, of i] by(auto simp add: vector_component_simps)
+    qed auto qed auto qed
+
+lemma interior_std_simplex_nonempty: obtains a::"real^'n::finite" where
+  "a \<in> interior(convex hull (insert 0 {basis i | i . i \<in> UNIV}))" proof-
+  let ?D = "UNIV::'n set" let ?a = "setsum (\<lambda>b::real^'n. inverse (2 * real CARD('n)) *\<^sub>R b) {(basis i) | i. i \<in> ?D}"
+  have *:"{basis i :: real ^ 'n | i. i \<in> ?D} = basis ` ?D" by auto
+  { fix i have "?a $ i = inverse (2 * real CARD('n))"
+    unfolding setsum_component vector_smult_component and * and setsum_reindex[OF basis_inj] and o_def
+    apply(rule trans[of _ "setsum (\<lambda>j. if i = j then inverse (2 * real CARD('n)) else 0) ?D"]) apply(rule setsum_cong2)
+      unfolding setsum_delta'[OF finite_UNIV[where 'a='n]] and real_dimindex_ge_1[where 'n='n] by(auto simp add: basis_component[of i]) }
+  note ** = this
+  show ?thesis apply(rule that[of ?a]) unfolding interior_std_simplex mem_Collect_eq proof(rule,rule)
+    fix i::'n show "0 < ?a $ i" unfolding ** using dimindex_ge_1 by(auto simp add: Suc_le_eq) next
+    have "setsum (op $ ?a) ?D = setsum (\<lambda>i. inverse (2 * real CARD('n))) ?D" by(rule setsum_cong2, rule **) 
+    also have "\<dots> < 1" unfolding setsum_constant card_enum real_eq_of_nat real_divide_def[THEN sym] by (auto simp add:field_simps)
+    finally show "setsum (op $ ?a) ?D < 1" by auto qed qed
+
+subsection {* Paths. *}
+
+definition "path (g::real^1 \<Rightarrow> real^'n::finite) \<longleftrightarrow> continuous_on {0 .. 1} g"
+
+definition "pathstart (g::real^1 \<Rightarrow> real^'n) = g 0"
+
+definition "pathfinish (g::real^1 \<Rightarrow> real^'n) = g 1"
+
+definition "path_image (g::real^1 \<Rightarrow> real^'n) = g ` {0 .. 1}"
+
+definition "reversepath (g::real^1 \<Rightarrow> real^'n) = (\<lambda>x. g(1 - x))"
+
+definition joinpaths:: "(real^1 \<Rightarrow> real^'n) \<Rightarrow> (real^1 \<Rightarrow> real^'n) \<Rightarrow> (real^1 \<Rightarrow> real^'n)" (infixr "+++" 75)
+  where "joinpaths g1 g2 = (\<lambda>x. if dest_vec1 x \<le> ((1 / 2)::real) then g1 (2 *\<^sub>R x) else g2(2 *\<^sub>R x - 1))"
+definition "simple_path (g::real^1 \<Rightarrow> real^'n) \<longleftrightarrow>
+  (\<forall>x\<in>{0..1}. \<forall>y\<in>{0..1}. g x = g y \<longrightarrow> x = y \<or> x = 0 \<and> y = 1 \<or> x = 1 \<and> y = 0)"
+
+definition "injective_path (g::real^1 \<Rightarrow> real^'n) \<longleftrightarrow>
+  (\<forall>x\<in>{0..1}. \<forall>y\<in>{0..1}. g x = g y \<longrightarrow> x = y)"
+
+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
+
+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 apply(rule imageI)
+  unfolding mem_interval_1 vec_1[THEN sym] dest_vec1_vec by auto
+
+lemma pathfinish_in_path_image[intro]: "(pathfinish g) \<in> path_image g"
+  unfolding pathfinish_def path_image_def apply(rule imageI)
+  unfolding mem_interval_1 vec_1[THEN sym] dest_vec1_vec by auto
+
+lemma connected_path_image[intro]: "path g \<Longrightarrow> connected(path_image g)"
+  unfolding path_def path_image_def apply(rule connected_continuous_image, assumption)
+  by(rule convex_connected, rule convex_interval)
+
+lemma compact_path_image[intro]: "path g \<Longrightarrow> compact(path_image g)"
+  unfolding path_def path_image_def apply(rule compact_continuous_image, assumption)
+  by(rule compact_interval)
+
+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 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
+
+lemma pathfinish_join[simp]:"pathfinish(g1 +++ g2) = pathfinish g2" proof-
+  have "2 *\<^sub>R 1 - 1 = (1::real^1)" unfolding Cart_eq by(auto simp add:vector_component_simps)
+  thus ?thesis unfolding pathstart_def joinpaths_def pathfinish_def
+    unfolding vec_1[THEN sym] dest_vec1_vec by auto qed
+
+lemma path_image_reversepath[simp]: "path_image(reversepath g) = path_image g" proof-
+  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) by(auto simp add:vector_less_eq_def vector_component_simps elim!:ballE)
+  show ?thesis using *[of g] *[of "reversepath g"] unfolding reversepath_reversepath by auto qed
+
+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(rule continuous_on_sub, rule continuous_on_const, rule continuous_on_id)
+    apply(rule continuous_on_subset[of "{0..1}"], assumption)
+    by (auto, auto simp add:vector_less_eq_def vector_component_simps elim!:ballE)
+  show ?thesis using *[of g] *[of "reversepath g"] unfolding reversepath_reversepath by auto qed
+
+lemmas reversepath_simps = path_reversepath path_image_reversepath pathstart_reversepath pathfinish_reversepath
+
+lemma path_join[simp]: assumes "pathfinish g1 = pathstart g2" shows "path (g1 +++ g2) \<longleftrightarrow>  path g1 \<and> path g2"
+  unfolding path_def pathfinish_def pathstart_def apply rule defer apply(erule conjE) proof-
+  assume as:"continuous_on {0..1} (g1 +++ g2)"
+  have *:"g1 = (\<lambda>x. g1 (2 *\<^sub>R x)) \<circ> (\<lambda>x. (1/2) *\<^sub>R x)" 
+         "g2 = (\<lambda>x. g2 (2 *\<^sub>R x - 1)) \<circ> (\<lambda>x. (1/2) *\<^sub>R (x + 1))" unfolding o_def by auto
+  have "op *\<^sub>R (1 / 2) ` {0::real^1..1} \<subseteq> {0..1}"  "(\<lambda>x. (1 / 2) *\<^sub>R (x + 1)) ` {(0::real^1)..1} \<subseteq> {0..1}"
+    unfolding image_smult_interval by (auto, auto simp add:vector_less_eq_def vector_component_simps elim!:ballE)
+  thus "continuous_on {0..1} g1 \<and> continuous_on {0..1} g2" apply -apply rule
+    apply(subst *) defer apply(subst *) apply (rule_tac[!] continuous_on_compose)
+    apply (rule continuous_on_cmul, rule continuous_on_add, rule continuous_on_id, rule continuous_on_const) defer
+    apply (rule continuous_on_cmul, rule continuous_on_id) apply(rule_tac[!] continuous_on_eq[of _ "g1 +++ g2"]) defer prefer 3
+    apply(rule_tac[1-2] continuous_on_subset[of "{0 .. 1}"]) apply(rule as, assumption, rule as, assumption)
+    apply(rule) defer apply rule proof-
+    fix x assume "x \<in> op *\<^sub>R (1 / 2) ` {0::real^1..1}"
+    hence "dest_vec1 x \<le> 1 / 2" unfolding image_iff by(auto simp add: vector_component_simps)
+    thus "(g1 +++ g2) x = g1 (2 *\<^sub>R x)" unfolding joinpaths_def by auto next
+    fix x assume "x \<in> (\<lambda>x. (1 / 2) *\<^sub>R (x + 1)) ` {0::real^1..1}"
+    hence "dest_vec1 x \<ge> 1 / 2" unfolding image_iff by(auto simp add: vector_component_simps)
+    thus "(g1 +++ g2) x = g2 (2 *\<^sub>R x - 1)" proof(cases "dest_vec1 x = 1 / 2")
+      case True hence "x = (1/2) *\<^sub>R 1" unfolding Cart_eq by(auto simp add: forall_1 vector_component_simps)
+      thus ?thesis unfolding joinpaths_def using assms[unfolded pathstart_def pathfinish_def] by auto
+    qed (auto simp add:le_less joinpaths_def) qed
+next assume as:"continuous_on {0..1} g1" "continuous_on {0..1} g2"
+  have *:"{0 .. 1::real^1} = {0.. (1/2)*\<^sub>R 1} \<union> {(1/2) *\<^sub>R 1 .. 1}" by(auto simp add: vector_component_simps) 
+  have **:"op *\<^sub>R 2 ` {0..(1 / 2) *\<^sub>R 1} = {0..1::real^1}" apply(rule set_ext, rule) unfolding image_iff 
+    defer apply(rule_tac x="(1/2)*\<^sub>R x" in bexI) by(auto simp add: vector_component_simps)
+  have ***:"(\<lambda>x. 2 *\<^sub>R x - 1) ` {(1 / 2) *\<^sub>R 1..1} = {0..1::real^1}"
+    unfolding image_affinity_interval[of _ "- 1", unfolded diff_def[symmetric]] and interval_eq_empty_1
+    by(auto simp add: vector_component_simps)
+  have ****:"\<And>x::real^1. x $ 1 * 2 = 1 \<longleftrightarrow> x = (1/2) *\<^sub>R 1" unfolding Cart_eq by(auto simp add: forall_1 vector_component_simps)
+  show "continuous_on {0..1} (g1 +++ g2)" unfolding * apply(rule continuous_on_union) apply(rule closed_interval)+ proof-
+    show "continuous_on {0..(1 / 2) *\<^sub>R 1} (g1 +++ g2)" apply(rule continuous_on_eq[of _ "\<lambda>x. g1 (2 *\<^sub>R x)"]) defer
+      unfolding o_def[THEN sym] apply(rule continuous_on_compose) apply(rule continuous_on_cmul, rule continuous_on_id)
+      unfolding ** apply(rule as(1)) unfolding joinpaths_def by(auto simp add: vector_component_simps) next
+    show "continuous_on {(1/2)*\<^sub>R1..1} (g1 +++ g2)" apply(rule continuous_on_eq[of _ "g2 \<circ> (\<lambda>x. 2 *\<^sub>R x - 1)"]) defer
+      apply(rule continuous_on_compose) apply(rule continuous_on_sub, rule continuous_on_cmul, rule continuous_on_id, rule continuous_on_const)
+      unfolding *** o_def joinpaths_def apply(rule as(2)) using assms[unfolded pathstart_def pathfinish_def]
+      by(auto simp add: vector_component_simps ****) qed qed
+
+lemma path_image_join_subset: "path_image(g1 +++ g2) \<subseteq> (path_image g1 \<union> path_image g2)" proof
+  fix x assume "x \<in> path_image (g1 +++ g2)"
+  then obtain y where y:"y\<in>{0..1}" "x = (if dest_vec1 y \<le> 1 / 2 then g1 (2 *\<^sub>R y) else g2 (2 *\<^sub>R y - 1))"
+    unfolding path_image_def image_iff joinpaths_def by auto
+  thus "x \<in> path_image g1 \<union> path_image g2" apply(cases "dest_vec1 y \<le> 1/2")
+    apply(rule_tac UnI1) defer apply(rule_tac UnI2) unfolding y(2) path_image_def using y(1)
+    by(auto intro!: imageI simp add: vector_component_simps) qed
+
+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
+
+lemma path_image_join:
+  assumes "path g1" "path g2" "pathfinish g1 = pathstart g2"
+  shows "path_image(g1 +++ g2) = (path_image g1) \<union> (path_image g2)"
+apply(rule, rule path_image_join_subset, 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" unfolding path_image_def image_iff by auto
+  thus "x \<in> path_image (g1 +++ g2)" unfolding joinpaths_def path_image_def image_iff
+    apply(rule_tac x="(1/2) *\<^sub>R y" in bexI) by(auto simp add: vector_component_simps) next
+  fix x assume "x \<in> path_image g2"
+  then obtain y where y:"y\<in>{0..1}" "x = g2 y" unfolding path_image_def image_iff by auto
+  moreover have *:"y $ 1 = 0 \<Longrightarrow> y = 0" unfolding Cart_eq by auto
+  ultimately 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) using assms(3)[unfolded pathfinish_def pathstart_def]
+    by(auto simp add: vector_component_simps) 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
+
+lemma simple_path_reversepath: assumes "simple_path g" shows "simple_path (reversepath g)"
+  using assms 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)
+  unfolding mem_interval_1 by(auto simp add:vector_component_simps)
+
+lemma dest_vec1_scaleR [simp]:
+  "dest_vec1 (scaleR a x) = scaleR a (dest_vec1 x)"
+unfolding dest_vec1_def by simp
+
+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)"
+unfolding simple_path_def proof((rule ballI)+, rule impI) let ?g = "g1 +++ g2"
+  note inj = assms(1,2)[unfolded injective_path_def, rule_format]
+  fix x y::"real^1" 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$1 \<le> 1/2",case_tac[!] "y$1 \<le> 1/2", unfold not_le)
+    assume as:"x $ 1 \<le> 1 / 2" "y $ 1 \<le> 1 / 2"
+    hence "g1 (2 *\<^sub>R x) = g1 (2 *\<^sub>R y)" using xy(3) unfolding joinpaths_def dest_vec1_def by auto
+    moreover have "2 *\<^sub>R x \<in> {0..1}" "2 *\<^sub>R y \<in> {0..1}" using xy(1,2) as
+      unfolding mem_interval_1 dest_vec1_def by(auto simp add:vector_component_simps)
+    ultimately show ?thesis using inj(1)[of "2*\<^sub>R x" "2*\<^sub>R y"] by auto
+  next assume as:"x $ 1 > 1 / 2" "y $ 1 > 1 / 2"
+    hence "g2 (2 *\<^sub>R x - 1) = g2 (2 *\<^sub>R y - 1)" using xy(3) unfolding joinpaths_def dest_vec1_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
+      unfolding mem_interval_1 dest_vec1_def by(auto simp add:vector_component_simps)
+    ultimately show ?thesis using inj(2)[of "2*\<^sub>R x - 1" "2*\<^sub>R y - 1"] by auto
+  next assume as:"x $ 1 \<le> 1 / 2" "y $ 1 > 1 / 2"
+    hence "?g x \<in> path_image g1" "?g y \<in> path_image g2" unfolding path_image_def joinpaths_def
+      using xy(1,2)[unfolded mem_interval_1] by(auto simp add:vector_component_simps intro!: imageI)
+    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)[unfolded mem_interval_1]
+      apply(rule_tac ccontr) by(auto simp add:vector_component_simps field_simps Cart_eq)
+    ultimately have *:"?g x = pathstart g1" using assms(4) unfolding xy(3) by auto
+    hence "x = 0" unfolding pathstart_def joinpaths_def using as(1) and xy(1)[unfolded mem_interval_1]
+      using inj(1)[of "2 *\<^sub>R x" 0] by(auto simp add:vector_component_simps)
+    moreover have "y = 1" using * unfolding xy(3) assms(3)[THEN sym]
+      unfolding joinpaths_def pathfinish_def using as(2) and xy(2)[unfolded mem_interval_1]
+      using inj(2)[of "2 *\<^sub>R y - 1" 1] by (auto simp add:vector_component_simps Cart_eq)
+    ultimately show ?thesis by auto 
+  next assume as:"x $ 1 > 1 / 2" "y $ 1 \<le> 1 / 2"
+    hence "?g x \<in> path_image g2" "?g y \<in> path_image g1" unfolding path_image_def joinpaths_def
+      using xy(1,2)[unfolded mem_interval_1] by(auto simp add:vector_component_simps intro!: imageI)
+    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)[unfolded mem_interval_1]
+      apply(rule_tac ccontr) by(auto simp add:vector_component_simps field_simps Cart_eq)
+    ultimately have *:"?g y = pathstart g1" using assms(4) unfolding xy(3) by auto
+    hence "y = 0" unfolding pathstart_def joinpaths_def using as(2) and xy(2)[unfolded mem_interval_1]
+      using inj(1)[of "2 *\<^sub>R y" 0] by(auto simp add:vector_component_simps)
+    moreover have "x = 1" using * unfolding xy(3)[THEN sym] assms(3)[THEN sym]
+      unfolding joinpaths_def pathfinish_def using as(1) and xy(1)[unfolded mem_interval_1]
+      using inj(2)[of "2 *\<^sub>R x - 1" 1] by(auto simp add:vector_component_simps Cart_eq)
+    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)"
+  unfolding injective_path_def proof(rule,rule,rule) let ?g = "g1 +++ g2"
+  note inj = assms(1,2)[unfolded injective_path_def, rule_format]
+  fix x y assume xy:"x \<in> {0..1}" "y \<in> {0..1}" "(g1 +++ g2) x = (g1 +++ g2) y"
+  show "x = y" proof(cases "x$1 \<le> 1/2", case_tac[!] "y$1 \<le> 1/2", unfold not_le)
+    assume "x $ 1 \<le> 1 / 2" "y $ 1 \<le> 1 / 2" thus ?thesis using inj(1)[of "2*\<^sub>R x" "2*\<^sub>R y"] and xy
+      unfolding mem_interval_1 joinpaths_def by(auto simp add:vector_component_simps)
+  next assume "x $ 1 > 1 / 2" "y $ 1 > 1 / 2" thus ?thesis using inj(2)[of "2*\<^sub>R x - 1" "2*\<^sub>R y - 1"] and xy
+      unfolding mem_interval_1 joinpaths_def by(auto simp add:vector_component_simps)
+  next assume as:"x $ 1 \<le> 1 / 2" "y $ 1 > 1 / 2" 
+    hence "?g x \<in> path_image g1" "?g y \<in> path_image g2" unfolding path_image_def joinpaths_def
+      using xy(1,2)[unfolded mem_interval_1] by(auto simp add:vector_component_simps intro!: imageI)
+    hence "?g x = pathfinish g1" "?g y = pathstart g2" using assms(4) unfolding assms(3) xy(3) by auto
+    thus ?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 mem_interval_1
+      by(auto simp add:vector_component_simps Cart_eq forall_1)
+  next assume as:"x $ 1 > 1 / 2" "y $ 1 \<le> 1 / 2" 
+    hence "?g x \<in> path_image g2" "?g y \<in> path_image g1" unfolding path_image_def joinpaths_def
+      using xy(1,2)[unfolded mem_interval_1] by(auto simp add:vector_component_simps intro!: imageI)
+    hence "?g x = pathstart g2" "?g y = pathfinish g1" using assms(4) unfolding assms(3) xy(3) by auto
+    thus ?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 mem_interval_1
+      by(auto simp add:vector_component_simps forall_1 Cart_eq) qed 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. *}
+
+definition "shiftpath a (f::real^1 \<Rightarrow> real^'n) =
+  (\<lambda>x. if dest_vec1 (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"
+  unfolding pathstart_def shiftpath_def by auto
+
+(** move this **)
+declare forall_1[simp] ex_1[simp]
+
+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
+  by(auto simp add: vector_component_simps)
+
+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)
+
+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
+
+lemma path_shiftpath:
+  assumes "path g" "pathfinish g = pathstart g" "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 simp add: vector_component_simps)
+  have **:"\<And>x. x + a = 1 \<Longrightarrow> g (x + a - 1) = g (x + a)"
+    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_interval)+ 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 ** by(auto simp add:vector_component_simps field_simps Cart_eq) qed
+
+lemma shiftpath_shiftpath: assumes "pathfinish g = pathstart g" "a \<in> {0..1}" "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 simp add: vector_component_simps)
+
+lemma path_image_shiftpath:
+  assumes "a \<in> {0..1}" "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^1}" " \<forall>y\<in>{0..1} \<inter> {x. \<not> a $ 1 + x $ 1 \<le> 1}. g x \<noteq> g (a + y - 1)" 
+    hence "\<exists>y\<in>{0..1} \<inter> {x. a $ 1 + x $ 1 \<le> 1}. g x = g (a + y)" proof(cases "a \<le> x")
+      case False thus ?thesis apply(rule_tac x="1 + x - a" in bexI)
+        using as(1,2) and as(3)[THEN bspec[where x="1 + x - a"]] and assms(1)
+        by(auto simp add:vector_component_simps field_simps atomize_not) next
+      case True thus ?thesis using as(1-2) and assms(1) apply(rule_tac x="x - a" in bexI)
+        by(auto simp add:vector_component_simps field_simps) qed }
+  thus ?thesis using assms unfolding shiftpath_def path_image_def pathfinish_def pathstart_def 
+    by(auto simp add:vector_component_simps image_iff) qed
+
+subsection {* Special case of straight-line paths. *}
+
+definition
+  linepath :: "real ^ 'n::finite \<Rightarrow> real ^ 'n \<Rightarrow> real ^ 1 \<Rightarrow> real ^ 'n" where
+  "linepath a b = (\<lambda>x. (1 - dest_vec1 x) *\<^sub>R a + dest_vec1 x *\<^sub>R b)"
+
+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 continuous_linepath_at[intro]: "continuous (at x) (linepath a b)"
+  unfolding linepath_def
+  by (intro continuous_intros continuous_dest_vec1)
+
+lemma continuous_on_linepath[intro]: "continuous_on s (linepath a b)"
+  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_image_linepath[simp]: "path_image(linepath a b) = (closed_segment a b)"
+  unfolding path_image_def segment linepath_def apply (rule set_ext, rule) defer
+  unfolding mem_Collect_eq image_iff apply(erule exE) apply(rule_tac x="u *\<^sub>R 1" in bexI)
+  by(auto simp add:vector_component_simps)
+
+lemma reversepath_linepath[simp]:  "reversepath(linepath a b) = linepath b a"
+  unfolding reversepath_def linepath_def by(rule ext, auto simp add:vector_component_simps)
+
+lemma injective_path_linepath: assumes "a \<noteq> b" shows "injective_path(linepath a b)" proof- 
+  { obtain i where i:"a$i \<noteq> b$i" using assms[unfolded Cart_eq] by auto
+    fix x y::"real^1" assume "x $ 1 *\<^sub>R b + y $ 1 *\<^sub>R a = x $ 1 *\<^sub>R a + y $ 1 *\<^sub>R b"
+    hence "x$1 * (b$i - a$i) = y$1 * (b$i - a$i)" unfolding Cart_eq by(auto simp add:field_simps vector_component_simps)
+    hence "x = y" unfolding mult_cancel_right Cart_eq using i(1) by(auto simp add:field_simps) }
+  thus ?thesis unfolding injective_path_def linepath_def by(auto simp add:vector_component_simps 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)
+
+subsection {* Bounding a point away from a path. *}
+
+lemma not_on_path_ball: assumes "path g" "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]
+    using compact_path_image[THEN compact_imp_closed, OF assms(1)] by auto
+  thus ?thesis apply(rule_tac x="dist z a" in exI) using assms(2) by(auto intro!: dist_pos_lt) qed
+
+lemma not_on_path_cball: assumes "path g" "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" 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 qed
+
+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 
+
+lemma path_component_mem: assumes "path_component s x y" shows "x \<in> s" "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 
+  by(auto intro!:continuous_on_intros)    
+
+lemma path_component_refl_eq: "path_component s x x \<longleftrightarrow> x \<in> s"
+  by(auto intro!: path_component_mem path_component_refl) 
+
+lemma path_component_sym: "path_component s x y \<Longrightarrow> path_component s y x"
+  using assms unfolding path_component_def apply(erule exE) apply(rule_tac x="reversepath g" in exI) 
+  by(auto simp add: reversepath_simps)
+
+lemma path_component_trans: assumes "path_component s x y" "path_component s y z" shows "path_component s x z"
+  using assms unfolding path_component_def apply- apply(erule exE)+ apply(rule_tac x="g +++ ga" in exI) by(auto simp add: path_image_join)
+
+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. *}
+
+lemma path_component_set: "path_component s x = { y. (\<exists>g. path g \<and> path_image g \<subseteq> s \<and> pathstart g = x \<and> pathfinish g = y )}"
+  apply(rule set_ext) unfolding mem_Collect_eq unfolding mem_def path_component_def by auto
+
+lemma mem_path_component_set:"x \<in> path_component s y \<longleftrightarrow> path_component s y x" unfolding mem_def by auto
+
+lemma path_component_subset: "(path_component s x) \<subseteq> s"
+  apply(rule, rule path_component_mem(2)) by(auto simp add:mem_def)
+
+lemma path_component_eq_empty: "path_component s x = {} \<longleftrightarrow> x \<notin> s"
+  apply rule apply(drule equals0D[of _ x]) defer apply(rule equals0I) unfolding mem_path_component_set
+  apply(drule path_component_mem(1)) using path_component_refl by auto
+
+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)"
+
+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. path_component s x = s)" 
+  unfolding path_connected_component apply(rule, rule, rule, rule path_component_subset) 
+  unfolding subset_eq mem_path_component_set Ball_def mem_def by auto
+
+subsection {* Some useful lemmas about path-connectedness. *}
+
+lemma convex_imp_path_connected: assumes "convex s" shows "path_connected s"
+  unfolding path_connected_def apply(rule,rule,rule_tac x="linepath x y" in exI)
+  unfolding path_image_linepath using assms[unfolded convex_contains_segment] by auto
+
+lemma path_connected_imp_connected: assumes "path_connected s" shows "connected s"
+  unfolding connected_def not_ex apply(rule,rule,rule ccontr) unfolding not_not apply(erule 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"
+    using assms[unfolded path_connected_def,rule_format,of x1 x2] by auto
+  have *:"connected {0..1::real^1}" by(auto intro!: convex_connected convex_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 
+  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 by(auto intro!: exI)
+  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 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 qed
+
+lemma open_path_component: assumes "open s" shows "open(path_component s x)"
+  unfolding open_contains_ball proof
+  fix y assume as:"y \<in> path_component s x"
+  hence "y\<in>s" apply- apply(rule path_component_mem(2)) unfolding mem_def 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> path_component s x" apply(rule_tac x=e in exI) apply(rule,rule `e>0`,rule) unfolding mem_ball mem_path_component_set proof-
+    fix z assume "dist y z < e" thus "path_component s x z" 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`
+      using as[unfolded mem_def] by auto qed qed
+
+lemma open_non_path_component: assumes "open s" shows "open(s - path_component s x)" unfolding open_contains_ball proof
+  fix y assume as:"y\<in>s - path_component s x" 
+  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 - path_component s x" apply(rule_tac x=e in exI) apply(rule,rule `e>0`,rule,rule) defer proof(rule ccontr)
+    fix z assume "z\<in>ball y e" "\<not> z \<notin> path_component s x" 
+    hence "y \<in> path_component s x" unfolding not_not mem_path_component_set using `e>0` 
+      apply- apply(rule path_component_trans,assumption) apply(rule path_component_of_subset[OF e(2)])
+      apply(rule convex_imp_path_connected[OF convex_ball, unfolded path_connected_component, rule_format]) by auto
+    thus False using as by auto qed(insert e(2), auto) qed
+
+lemma connected_open_path_connected: assumes "open s" "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" show "y \<in> path_component s x" proof(rule ccontr)
+    assume "y \<notin> path_component s x" moreover
+    have "path_component s x \<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"path_component s x" "s - path_component s x"] by auto
+qed qed
+
+lemma path_connected_continuous_image:
+  assumes "continuous_on s f" "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)] ..
+  thus "\<exists>g. path g \<and> path_image g \<subseteq> f ` s \<and> pathstart g = x' \<and> pathfinish g = y'"
+    unfolding xy apply(rule_tac x="f \<circ> g" in exI) unfolding path_defs
+    using assms(1) by(auto intro!: continuous_on_compose continuous_on_subset[of _ _ "g ` {0..1}"]) qed
+
+lemma homeomorphic_path_connectedness:
+  "s homeomorphic t \<Longrightarrow> (path_connected s \<longleftrightarrow> path_connected t)"
+  unfolding homeomorphic_def homeomorphism_def apply(erule exE|erule conjE)+ apply rule
+  apply(drule_tac f=f in path_connected_continuous_image) prefer 3
+  apply(drule_tac f=g in path_connected_continuous_image) by auto
+
+lemma path_connected_empty: "path_connected {}"
+  unfolding path_connected_def by auto
+
+lemma path_connected_singleton: "path_connected {a}"
+  unfolding path_connected_def apply(rule,rule)
+  apply(rule_tac x="linepath a a" in exI) by(auto simp add:segment scaleR_left_diff_distrib)
+
+lemma path_connected_Un: assumes "path_connected s" "path_connected t" "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
+  thus "path_component (s \<union> t) x y" using as using assms(1-2)[unfolded path_connected_component] apply- 
+    apply(erule_tac[!] UnE)+ apply(rule_tac[2-3] path_component_trans[of _ _ z])
+    by(auto simp add:path_component_of_subset [OF Un_upper1] path_component_of_subset[OF Un_upper2]) qed
+
+subsection {* sphere is path-connected. *}
+
+lemma path_connected_punctured_universe:
+ assumes "2 \<le> CARD('n::finite)" shows "path_connected((UNIV::(real^'n::finite) set) - {a})" proof-
+  obtain \<psi> where \<psi>:"bij_betw \<psi> {1..CARD('n)} (UNIV::'n set)" using ex_bij_betw_nat_finite_1[OF finite_UNIV] by auto
+  let ?U = "UNIV::(real^'n) set" let ?u = "?U - {0}"
+  let ?basis = "\<lambda>k. basis (\<psi> k)"
+  let ?A = "\<lambda>k. {x::real^'n. \<exists>i\<in>{1..k}. inner (basis (\<psi> i)) x \<noteq> 0}"
+  have "\<forall>k\<in>{2..CARD('n)}. path_connected (?A k)" proof
+    have *:"\<And>k. ?A (Suc k) = {x. inner (?basis (Suc k)) x < 0} \<union> {x. inner (?basis (Suc k)) x > 0} \<union> ?A k" apply(rule set_ext,rule) defer
+      apply(erule UnE)+  unfolding mem_Collect_eq apply(rule_tac[1-2] x="Suc k" in bexI)
+      by(auto elim!: ballE simp add: not_less le_Suc_eq)
+    fix k assume "k \<in> {2..CARD('n)}" thus "path_connected (?A k)" proof(induct k)
+      case (Suc k) show ?case proof(cases "k = 1")
+        case False from Suc have d:"k \<in> {1..CARD('n)}" "Suc k \<in> {1..CARD('n)}" by auto
+        hence "\<psi> k \<noteq> \<psi> (Suc k)" using \<psi>[unfolded bij_betw_def inj_on_def, THEN conjunct1, THEN bspec[where x=k]] by auto
+        hence **:"?basis k + ?basis (Suc k) \<in> {x. 0 < inner (?basis (Suc k)) x} \<inter> (?A k)" 
+          "?basis k - ?basis (Suc k) \<in> {x. 0 > inner (?basis (Suc k)) x} \<inter> ({x. 0 < inner (?basis (Suc k)) x} \<union> (?A k))" using d
+          by(auto simp add: inner_basis vector_component_simps intro!:bexI[where x=k])
+        show ?thesis unfolding * Un_assoc apply(rule path_connected_Un) defer apply(rule path_connected_Un) 
+          prefer 5 apply(rule_tac[1-2] convex_imp_path_connected, rule convex_halfspace_lt, rule convex_halfspace_gt)
+          apply(rule Suc(1)) apply(rule_tac[2-3] ccontr) using d ** False by auto
+      next case True hence d:"1\<in>{1..CARD('n)}" "2\<in>{1..CARD('n)}" using Suc(2) by auto
+        have ***:"Suc 1 = 2" by auto
+        have **:"\<And>s t P Q. s \<union> t \<union> {x. P x \<or> Q x} = (s \<union> {x. P x}) \<union> (t \<union> {x. Q x})" by auto
+        have "\<psi> 2 \<noteq> \<psi> (Suc 0)" apply(rule ccontr) using \<psi>[unfolded bij_betw_def inj_on_def, THEN conjunct1, THEN bspec[where x=2]] using assms by auto
+        thus ?thesis unfolding * True unfolding ** neq_iff bex_disj_distrib apply -
+          apply(rule path_connected_Un, rule_tac[1-2] path_connected_Un) defer 3 apply(rule_tac[1-4] convex_imp_path_connected) 
+          apply(rule_tac[5] x=" ?basis 1 + ?basis 2" in nequals0I)
+          apply(rule_tac[6] x="-?basis 1 + ?basis 2" in nequals0I)
+          apply(rule_tac[7] x="-?basis 1 - ?basis 2" in nequals0I)
+          using d unfolding *** by(auto intro!: convex_halfspace_gt convex_halfspace_lt, auto simp add:vector_component_simps inner_basis)
+  qed qed auto qed note lem = this
+
+  have ***:"\<And>x::real^'n. (\<exists>i\<in>{1..CARD('n)}. inner (basis (\<psi> i)) x \<noteq> 0) \<longleftrightarrow> (\<exists>i. inner (basis i) x \<noteq> 0)"
+    apply rule apply(erule bexE) apply(rule_tac x="\<psi> i" in exI) defer apply(erule exE) proof- 
+    fix x::"real^'n" and i assume as:"inner (basis i) x \<noteq> 0"
+    have "i\<in>\<psi> ` {1..CARD('n)}" using \<psi>[unfolded bij_betw_def, THEN conjunct2] by auto
+    then obtain j where "j\<in>{1..CARD('n)}" "\<psi> j = i" by auto
+    thus "\<exists>i\<in>{1..CARD('n)}. inner (basis (\<psi> i)) x \<noteq> 0" apply(rule_tac x=j in bexI) using as by auto qed auto
+  have *:"?U - {a} = (\<lambda>x. x + a) ` {x. x \<noteq> 0}" apply(rule set_ext) unfolding image_iff 
+    apply rule apply(rule_tac x="x - a" in bexI) by auto
+  have **:"\<And>x::real^'n. x\<noteq>0 \<longleftrightarrow> (\<exists>i. inner (basis i) x \<noteq> 0)" unfolding Cart_eq by(auto simp add: inner_basis)
+  show ?thesis unfolding * apply(rule path_connected_continuous_image) apply(rule continuous_on_intros)+ 
+    unfolding ** apply(rule lem[THEN bspec[where x="CARD('n)"], unfolded ***]) using assms by auto qed
+
+lemma path_connected_sphere: assumes "2 \<le> CARD('n::finite)" shows "path_connected {x::real^'n::finite. norm(x - a) = r}" proof(cases "r\<le>0")
+  case True thus ?thesis proof(cases "r=0") 
+    case False hence "{x::real^'n. norm(x - a) = r} = {}" using True by auto
+    thus ?thesis using path_connected_empty by auto
+  qed(auto intro!:path_connected_singleton) next
+  case False hence *:"{x::real^'n. norm(x - a) = r} = (\<lambda>x. a + r *\<^sub>R x) ` {x. norm x = 1}" unfolding not_le apply -apply(rule set_ext,rule)
+    unfolding image_iff apply(rule_tac x="(1/r) *\<^sub>R (x - a)" in bexI) unfolding mem_Collect_eq norm_scaleR by (auto simp add: scaleR_right_diff_distrib)
+  have **:"{x::real^'n. norm x = 1} = (\<lambda>x. (1/norm x) *\<^sub>R x) ` (UNIV - {0})" apply(rule set_ext,rule)
+    unfolding image_iff apply(rule_tac x=x in bexI) unfolding mem_Collect_eq by(auto split:split_if_asm)
+  have "continuous_on (UNIV - {0}) (\<lambda>x::real^'n. 1 / norm x)" unfolding o_def continuous_on_eq_continuous_within
+    apply(rule, rule continuous_at_within_inv[unfolded o_def inverse_eq_divide]) apply(rule continuous_at_within)
+    apply(rule continuous_at_norm[unfolded o_def]) by auto
+  thus ?thesis unfolding * ** using path_connected_punctured_universe[OF assms]
+    by(auto intro!: path_connected_continuous_image continuous_on_intros continuous_on_mul) qed
+
+lemma connected_sphere: "2 \<le> CARD('n) \<Longrightarrow> connected {x::real^'n::finite. norm(x - a) = r}"
+  using path_connected_sphere path_connected_imp_connected by auto
+ 
+(** In continuous_at_vec1_norm : Use \<And> instead of \<forall>. **)
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Multivariate_Analysis/Determinants.thy	Fri Oct 23 13:23:18 2009 +0200
@@ -0,0 +1,1087 @@
+(* Title:      Determinants
+   Author:     Amine Chaieb, University of Cambridge
+*)
+
+header {* Traces, Determinant of square matrices and some properties *}
+
+theory Determinants
+imports Euclidean_Space Permutations
+begin
+
+subsection{* First some facts about products*}
+lemma setprod_insert_eq: "finite A \<Longrightarrow> setprod f (insert a A) = (if a \<in> A then setprod f A else f a * setprod f A)"
+apply clarsimp
+by(subgoal_tac "insert a A = A", auto)
+
+lemma setprod_add_split:
+  assumes mn: "(m::nat) <= n + 1"
+  shows "setprod f {m.. n+p} = setprod f {m .. n} * setprod f {n+1..n+p}"
+proof-
+  let ?A = "{m .. n+p}"
+  let ?B = "{m .. n}"
+  let ?C = "{n+1..n+p}"
+  from mn have un: "?B \<union> ?C = ?A" by auto
+  from mn have dj: "?B \<inter> ?C = {}" by auto
+  have f: "finite ?B" "finite ?C" by simp_all
+  from setprod_Un_disjoint[OF f dj, of f, unfolded un] show ?thesis .
+qed
+
+
+lemma setprod_offset: "setprod f {(m::nat) + p .. n + p} = setprod (\<lambda>i. f (i + p)) {m..n}"
+apply (rule setprod_reindex_cong[where f="op + p"])
+apply (auto simp add: image_iff Bex_def inj_on_def)
+apply arith
+apply (rule ext)
+apply (simp add: add_commute)
+done
+
+lemma setprod_singleton: "setprod f {x} = f x" by simp
+
+lemma setprod_singleton_nat_seg: "setprod f {n..n} = f (n::'a::order)" by simp
+
+lemma setprod_numseg: "setprod f {m..0} = (if m=0 then f 0 else 1)"
+  "setprod f {m .. Suc n} = (if m \<le> Suc n then f (Suc n) * setprod f {m..n}
+                             else setprod f {m..n})"
+  by (auto simp add: atLeastAtMostSuc_conv)
+
+lemma setprod_le: assumes fS: "finite S" and fg: "\<forall>x\<in>S. f x \<ge> 0 \<and> f x \<le> (g x :: 'a::ordered_idom)"
+  shows "setprod f S \<le> setprod g S"
+using fS fg
+apply(induct S)
+apply simp
+apply auto
+apply (rule mult_mono)
+apply (auto intro: setprod_nonneg)
+done
+
+  (* FIXME: In Finite_Set there is a useless further assumption *)
+lemma setprod_inversef: "finite A ==> setprod (inverse \<circ> f) A = (inverse (setprod f A) :: 'a:: {division_by_zero, field})"
+  apply (erule finite_induct)
+  apply (simp)
+  apply simp
+  done
+
+lemma setprod_le_1: assumes fS: "finite S" and f: "\<forall>x\<in>S. f x \<ge> 0 \<and> f x \<le> (1::'a::ordered_idom)"
+  shows "setprod f S \<le> 1"
+using setprod_le[OF fS f] unfolding setprod_1 .
+
+subsection{* Trace *}
+
+definition trace :: "'a::semiring_1^'n^'n \<Rightarrow> 'a" where
+  "trace A = setsum (\<lambda>i. ((A$i)$i)) (UNIV::'n set)"
+
+lemma trace_0: "trace(mat 0) = 0"
+  by (simp add: trace_def mat_def)
+
+lemma trace_I: "trace(mat 1 :: 'a::semiring_1^'n^'n) = of_nat(CARD('n))"
+  by (simp add: trace_def mat_def)
+
+lemma trace_add: "trace ((A::'a::comm_semiring_1^'n^'n) + B) = trace A + trace B"
+  by (simp add: trace_def setsum_addf)
+
+lemma trace_sub: "trace ((A::'a::comm_ring_1^'n^'n) - B) = trace A - trace B"
+  by (simp add: trace_def setsum_subtractf)
+
+lemma trace_mul_sym:"trace ((A::'a::comm_semiring_1^'n^'n) ** B) = trace (B**A)"
+  apply (simp add: trace_def matrix_matrix_mult_def)
+  apply (subst setsum_commute)
+  by (simp add: mult_commute)
+
+(* ------------------------------------------------------------------------- *)
+(* Definition of determinant.                                                *)
+(* ------------------------------------------------------------------------- *)
+
+definition det:: "'a::comm_ring_1^'n^'n \<Rightarrow> 'a" where
+  "det A = setsum (\<lambda>p. of_int (sign p) * setprod (\<lambda>i. A$i$p i) (UNIV :: 'n set)) {p. p permutes (UNIV :: 'n set)}"
+
+(* ------------------------------------------------------------------------- *)
+(* A few general lemmas we need below.                                       *)
+(* ------------------------------------------------------------------------- *)
+
+lemma setprod_permute:
+  assumes p: "p permutes S"
+  shows "setprod f S = setprod (f o p) S"
+proof-
+  {assume "\<not> finite S" hence ?thesis by simp}
+  moreover
+  {assume fS: "finite S"
+    then have ?thesis
+      apply (simp add: setprod_def cong del:strong_setprod_cong)
+      apply (rule ab_semigroup_mult.fold_image_permute)
+      apply (auto simp add: p)
+      apply unfold_locales
+      done}
+  ultimately show ?thesis by blast
+qed
+
+lemma setproduct_permute_nat_interval: "p permutes {m::nat .. n} ==> setprod f {m..n} = setprod (f o p) {m..n}"
+  by (blast intro!: setprod_permute)
+
+(* ------------------------------------------------------------------------- *)
+(* Basic determinant properties.                                             *)
+(* ------------------------------------------------------------------------- *)
+
+lemma det_transp: "det (transp A) = det (A::'a::comm_ring_1 ^'n^'n::finite)"
+proof-
+  let ?di = "\<lambda>A i j. A$i$j"
+  let ?U = "(UNIV :: 'n set)"
+  have fU: "finite ?U" by simp
+  {fix p assume p: "p \<in> {p. p permutes ?U}"
+    from p have pU: "p permutes ?U" by blast
+    have sth: "sign (inv p) = sign p"
+      by (metis sign_inverse fU p mem_def Collect_def permutation_permutes)
+    from permutes_inj[OF pU]
+    have pi: "inj_on p ?U" by (blast intro: subset_inj_on)
+    from permutes_image[OF pU]
+    have "setprod (\<lambda>i. ?di (transp A) i (inv p i)) ?U = setprod (\<lambda>i. ?di (transp A) i (inv p i)) (p ` ?U)" by simp
+    also have "\<dots> = setprod ((\<lambda>i. ?di (transp A) i (inv p i)) o p) ?U"
+      unfolding setprod_reindex[OF pi] ..
+    also have "\<dots> = setprod (\<lambda>i. ?di A i (p i)) ?U"
+    proof-
+      {fix i assume i: "i \<in> ?U"
+        from i permutes_inv_o[OF pU] permutes_in_image[OF pU]
+        have "((\<lambda>i. ?di (transp A) i (inv p i)) o p) i = ?di A i (p i)"
+          unfolding transp_def by (simp add: expand_fun_eq)}
+      then show "setprod ((\<lambda>i. ?di (transp A) i (inv p i)) o p) ?U = setprod (\<lambda>i. ?di A i (p i)) ?U" by (auto intro: setprod_cong)
+    qed
+    finally have "of_int (sign (inv p)) * (setprod (\<lambda>i. ?di (transp A) i (inv p i)) ?U) = of_int (sign p) * (setprod (\<lambda>i. ?di A i (p i)) ?U)" using sth
+      by simp}
+  then show ?thesis unfolding det_def apply (subst setsum_permutations_inverse)
+  apply (rule setsum_cong2) by blast
+qed
+
+lemma det_lowerdiagonal:
+  fixes A :: "'a::comm_ring_1^'n^'n::{finite,wellorder}"
+  assumes ld: "\<And>i j. i < j \<Longrightarrow> A$i$j = 0"
+  shows "det A = setprod (\<lambda>i. A$i$i) (UNIV:: 'n set)"
+proof-
+  let ?U = "UNIV:: 'n set"
+  let ?PU = "{p. p permutes ?U}"
+  let ?pp = "\<lambda>p. of_int (sign p) * setprod (\<lambda>i. A$i$p i) (UNIV :: 'n set)"
+  have fU: "finite ?U" by simp
+  from finite_permutations[OF fU] have fPU: "finite ?PU" .
+  have id0: "{id} \<subseteq> ?PU" by (auto simp add: permutes_id)
+  {fix p assume p: "p \<in> ?PU -{id}"
+    from p have pU: "p permutes ?U" and pid: "p \<noteq> id" by blast+
+    from permutes_natset_le[OF pU] pid obtain i where
+      i: "p i > i" by (metis not_le)
+    from ld[OF i] have ex:"\<exists>i \<in> ?U. A$i$p i = 0" by blast
+    from setprod_zero[OF fU ex] have "?pp p = 0" by simp}
+  then have p0: "\<forall>p \<in> ?PU -{id}. ?pp p = 0"  by blast
+  from setsum_mono_zero_cong_left[OF fPU id0 p0] show ?thesis
+    unfolding det_def by (simp add: sign_id)
+qed
+
+lemma det_upperdiagonal:
+  fixes A :: "'a::comm_ring_1^'n^'n::{finite,wellorder}"
+  assumes ld: "\<And>i j. i > j \<Longrightarrow> A$i$j = 0"
+  shows "det A = setprod (\<lambda>i. A$i$i) (UNIV:: 'n set)"
+proof-
+  let ?U = "UNIV:: 'n set"
+  let ?PU = "{p. p permutes ?U}"
+  let ?pp = "(\<lambda>p. of_int (sign p) * setprod (\<lambda>i. A$i$p i) (UNIV :: 'n set))"
+  have fU: "finite ?U" by simp
+  from finite_permutations[OF fU] have fPU: "finite ?PU" .
+  have id0: "{id} \<subseteq> ?PU" by (auto simp add: permutes_id)
+  {fix p assume p: "p \<in> ?PU -{id}"
+    from p have pU: "p permutes ?U" and pid: "p \<noteq> id" by blast+
+    from permutes_natset_ge[OF pU] pid obtain i where
+      i: "p i < i" by (metis not_le)
+    from ld[OF i] have ex:"\<exists>i \<in> ?U. A$i$p i = 0" by blast
+    from setprod_zero[OF fU ex] have "?pp p = 0" by simp}
+  then have p0: "\<forall>p \<in> ?PU -{id}. ?pp p = 0"  by blast
+  from   setsum_mono_zero_cong_left[OF fPU id0 p0] show ?thesis
+    unfolding det_def by (simp add: sign_id)
+qed
+
+lemma det_diagonal:
+  fixes A :: "'a::comm_ring_1^'n^'n::finite"
+  assumes ld: "\<And>i j. i \<noteq> j \<Longrightarrow> A$i$j = 0"
+  shows "det A = setprod (\<lambda>i. A$i$i) (UNIV::'n set)"
+proof-
+  let ?U = "UNIV:: 'n set"
+  let ?PU = "{p. p permutes ?U}"
+  let ?pp = "\<lambda>p. of_int (sign p) * setprod (\<lambda>i. A$i$p i) (UNIV :: 'n set)"
+  have fU: "finite ?U" by simp
+  from finite_permutations[OF fU] have fPU: "finite ?PU" .
+  have id0: "{id} \<subseteq> ?PU" by (auto simp add: permutes_id)
+  {fix p assume p: "p \<in> ?PU - {id}"
+    then have "p \<noteq> id" by simp
+    then obtain i where i: "p i \<noteq> i" unfolding expand_fun_eq by auto
+    from ld [OF i [symmetric]] have ex:"\<exists>i \<in> ?U. A$i$p i = 0" by blast
+    from setprod_zero [OF fU ex] have "?pp p = 0" by simp}
+  then have p0: "\<forall>p \<in> ?PU - {id}. ?pp p = 0"  by blast
+  from setsum_mono_zero_cong_left[OF fPU id0 p0] show ?thesis
+    unfolding det_def by (simp add: sign_id)
+qed
+
+lemma det_I: "det (mat 1 :: 'a::comm_ring_1^'n^'n::finite) = 1"
+proof-
+  let ?A = "mat 1 :: 'a::comm_ring_1^'n^'n"
+  let ?U = "UNIV :: 'n set"
+  let ?f = "\<lambda>i j. ?A$i$j"
+  {fix i assume i: "i \<in> ?U"
+    have "?f i i = 1" using i by (vector mat_def)}
+  hence th: "setprod (\<lambda>i. ?f i i) ?U = setprod (\<lambda>x. 1) ?U"
+    by (auto intro: setprod_cong)
+  {fix i j assume i: "i \<in> ?U" and j: "j \<in> ?U" and ij: "i \<noteq> j"
+    have "?f i j = 0" using i j ij by (vector mat_def) }
+  then have "det ?A = setprod (\<lambda>i. ?f i i) ?U" using det_diagonal
+    by blast
+  also have "\<dots> = 1" unfolding th setprod_1 ..
+  finally show ?thesis .
+qed
+
+lemma det_0: "det (mat 0 :: 'a::comm_ring_1^'n^'n::finite) = 0"
+  by (simp add: det_def setprod_zero)
+
+lemma det_permute_rows:
+  fixes A :: "'a::comm_ring_1^'n^'n::finite"
+  assumes p: "p permutes (UNIV :: 'n::finite set)"
+  shows "det(\<chi> i. A$p i :: 'a^'n^'n) = of_int (sign p) * det A"
+  apply (simp add: det_def setsum_right_distrib mult_assoc[symmetric])
+  apply (subst sum_permutations_compose_right[OF p])
+proof(rule setsum_cong2)
+  let ?U = "UNIV :: 'n set"
+  let ?PU = "{p. p permutes ?U}"
+  fix q assume qPU: "q \<in> ?PU"
+  have fU: "finite ?U" by simp
+  from qPU have q: "q permutes ?U" by blast
+  from p q have pp: "permutation p" and qp: "permutation q"
+    by (metis fU permutation_permutes)+
+  from permutes_inv[OF p] have ip: "inv p permutes ?U" .
+    have "setprod (\<lambda>i. A$p i$ (q o p) i) ?U = setprod ((\<lambda>i. A$p i$(q o p) i) o inv p) ?U"
+      by (simp only: setprod_permute[OF ip, symmetric])
+    also have "\<dots> = setprod (\<lambda>i. A $ (p o inv p) i $ (q o (p o inv p)) i) ?U"
+      by (simp only: o_def)
+    also have "\<dots> = setprod (\<lambda>i. A$i$q i) ?U" by (simp only: o_def permutes_inverses[OF p])
+    finally   have thp: "setprod (\<lambda>i. A$p i$ (q o p) i) ?U = setprod (\<lambda>i. A$i$q i) ?U"
+      by blast
+  show "of_int (sign (q o p)) * setprod (\<lambda>i. A$ p i$ (q o p) i) ?U = of_int (sign p) * of_int (sign q) * setprod (\<lambda>i. A$i$q i) ?U"
+    by (simp only: thp sign_compose[OF qp pp] mult_commute of_int_mult)
+qed
+
+lemma det_permute_columns:
+  fixes A :: "'a::comm_ring_1^'n^'n::finite"
+  assumes p: "p permutes (UNIV :: 'n set)"
+  shows "det(\<chi> i j. A$i$ p j :: 'a^'n^'n) = of_int (sign p) * det A"
+proof-
+  let ?Ap = "\<chi> i j. A$i$ p j :: 'a^'n^'n"
+  let ?At = "transp A"
+  have "of_int (sign p) * det A = det (transp (\<chi> i. transp A $ p i))"
+    unfolding det_permute_rows[OF p, of ?At] det_transp ..
+  moreover
+  have "?Ap = transp (\<chi> i. transp A $ p i)"
+    by (simp add: transp_def Cart_eq)
+  ultimately show ?thesis by simp
+qed
+
+lemma det_identical_rows:
+  fixes A :: "'a::ordered_idom^'n^'n::finite"
+  assumes ij: "i \<noteq> j"
+  and r: "row i A = row j A"
+  shows "det A = 0"
+proof-
+  have tha: "\<And>(a::'a) b. a = b ==> b = - a ==> a = 0"
+    by simp
+  have th1: "of_int (-1) = - 1" by (metis of_int_1 of_int_minus number_of_Min)
+  let ?p = "Fun.swap i j id"
+  let ?A = "\<chi> i. A $ ?p i"
+  from r have "A = ?A" by (simp add: Cart_eq row_def swap_def)
+  hence "det A = det ?A" by simp
+  moreover have "det A = - det ?A"
+    by (simp add: det_permute_rows[OF permutes_swap_id] sign_swap_id ij th1)
+  ultimately show "det A = 0" by (metis tha)
+qed
+
+lemma det_identical_columns:
+  fixes A :: "'a::ordered_idom^'n^'n::finite"
+  assumes ij: "i \<noteq> j"
+  and r: "column i A = column j A"
+  shows "det A = 0"
+apply (subst det_transp[symmetric])
+apply (rule det_identical_rows[OF ij])
+by (metis row_transp r)
+
+lemma det_zero_row:
+  fixes A :: "'a::{idom, ring_char_0}^'n^'n::finite"
+  assumes r: "row i A = 0"
+  shows "det A = 0"
+using r
+apply (simp add: row_def det_def Cart_eq)
+apply (rule setsum_0')
+apply (auto simp: sign_nz)
+done
+
+lemma det_zero_column:
+  fixes A :: "'a::{idom,ring_char_0}^'n^'n::finite"
+  assumes r: "column i A = 0"
+  shows "det A = 0"
+  apply (subst det_transp[symmetric])
+  apply (rule det_zero_row [of i])
+  by (metis row_transp r)
+
+lemma det_row_add:
+  fixes a b c :: "'n::finite \<Rightarrow> _ ^ 'n"
+  shows "det((\<chi> i. if i = k then a i + b i else c i)::'a::comm_ring_1^'n^'n) =
+             det((\<chi> i. if i = k then a i else c i)::'a::comm_ring_1^'n^'n) +
+             det((\<chi> i. if i = k then b i else c i)::'a::comm_ring_1^'n^'n)"
+unfolding det_def Cart_lambda_beta setsum_addf[symmetric]
+proof (rule setsum_cong2)
+  let ?U = "UNIV :: 'n set"
+  let ?pU = "{p. p permutes ?U}"
+  let ?f = "(\<lambda>i. if i = k then a i + b i else c i)::'n \<Rightarrow> 'a::comm_ring_1^'n"
+  let ?g = "(\<lambda> i. if i = k then a i else c i)::'n \<Rightarrow> 'a::comm_ring_1^'n"
+  let ?h = "(\<lambda> i. if i = k then b i else c i)::'n \<Rightarrow> 'a::comm_ring_1^'n"
+  fix p assume p: "p \<in> ?pU"
+  let ?Uk = "?U - {k}"
+  from p have pU: "p permutes ?U" by blast
+  have kU: "?U = insert k ?Uk" by blast
+  {fix j assume j: "j \<in> ?Uk"
+    from j have "?f j $ p j = ?g j $ p j" and "?f j $ p j= ?h j $ p j"
+      by simp_all}
+  then have th1: "setprod (\<lambda>i. ?f i $ p i) ?Uk = setprod (\<lambda>i. ?g i $ p i) ?Uk"
+    and th2: "setprod (\<lambda>i. ?f i $ p i) ?Uk = setprod (\<lambda>i. ?h i $ p i) ?Uk"
+    apply -
+    apply (rule setprod_cong, simp_all)+
+    done
+  have th3: "finite ?Uk" "k \<notin> ?Uk" by auto
+  have "setprod (\<lambda>i. ?f i $ p i) ?U = setprod (\<lambda>i. ?f i $ p i) (insert k ?Uk)"
+    unfolding kU[symmetric] ..
+  also have "\<dots> = ?f k $ p k  * setprod (\<lambda>i. ?f i $ p i) ?Uk"
+    apply (rule setprod_insert)
+    apply simp
+    by blast
+  also have "\<dots> = (a k $ p k * setprod (\<lambda>i. ?f i $ p i) ?Uk) + (b k$ p k * setprod (\<lambda>i. ?f i $ p i) ?Uk)" by (simp add: ring_simps)
+  also have "\<dots> = (a k $ p k * setprod (\<lambda>i. ?g i $ p i) ?Uk) + (b k$ p k * setprod (\<lambda>i. ?h i $ p i) ?Uk)" by (metis th1 th2)
+  also have "\<dots> = setprod (\<lambda>i. ?g i $ p i) (insert k ?Uk) + setprod (\<lambda>i. ?h i $ p i) (insert k ?Uk)"
+    unfolding  setprod_insert[OF th3] by simp
+  finally have "setprod (\<lambda>i. ?f i $ p i) ?U = setprod (\<lambda>i. ?g i $ p i) ?U + setprod (\<lambda>i. ?h i $ p i) ?U" unfolding kU[symmetric] .
+  then show "of_int (sign p) * setprod (\<lambda>i. ?f i $ p i) ?U = of_int (sign p) * setprod (\<lambda>i. ?g i $ p i) ?U + of_int (sign p) * setprod (\<lambda>i. ?h i $ p i) ?U"
+    by (simp add: ring_simps)
+qed
+
+lemma det_row_mul:
+  fixes a b :: "'n::finite \<Rightarrow> _ ^ 'n"
+  shows "det((\<chi> i. if i = k then c *s a i else b i)::'a::comm_ring_1^'n^'n) =
+             c* det((\<chi> i. if i = k then a i else b i)::'a::comm_ring_1^'n^'n)"
+
+unfolding det_def Cart_lambda_beta setsum_right_distrib
+proof (rule setsum_cong2)
+  let ?U = "UNIV :: 'n set"
+  let ?pU = "{p. p permutes ?U}"
+  let ?f = "(\<lambda>i. if i = k then c*s a i else b i)::'n \<Rightarrow> 'a::comm_ring_1^'n"
+  let ?g = "(\<lambda> i. if i = k then a i else b i)::'n \<Rightarrow> 'a::comm_ring_1^'n"
+  fix p assume p: "p \<in> ?pU"
+  let ?Uk = "?U - {k}"
+  from p have pU: "p permutes ?U" by blast
+  have kU: "?U = insert k ?Uk" by blast
+  {fix j assume j: "j \<in> ?Uk"
+    from j have "?f j $ p j = ?g j $ p j" by simp}
+  then have th1: "setprod (\<lambda>i. ?f i $ p i) ?Uk = setprod (\<lambda>i. ?g i $ p i) ?Uk"
+    apply -
+    apply (rule setprod_cong, simp_all)
+    done
+  have th3: "finite ?Uk" "k \<notin> ?Uk" by auto
+  have "setprod (\<lambda>i. ?f i $ p i) ?U = setprod (\<lambda>i. ?f i $ p i) (insert k ?Uk)"
+    unfolding kU[symmetric] ..
+  also have "\<dots> = ?f k $ p k  * setprod (\<lambda>i. ?f i $ p i) ?Uk"
+    apply (rule setprod_insert)
+    apply simp
+    by blast
+  also have "\<dots> = (c*s a k) $ p k * setprod (\<lambda>i. ?f i $ p i) ?Uk" by (simp add: ring_simps)
+  also have "\<dots> = c* (a k $ p k * setprod (\<lambda>i. ?g i $ p i) ?Uk)"
+    unfolding th1 by (simp add: mult_ac)
+  also have "\<dots> = c* (setprod (\<lambda>i. ?g i $ p i) (insert k ?Uk))"
+    unfolding  setprod_insert[OF th3] by simp
+  finally have "setprod (\<lambda>i. ?f i $ p i) ?U = c* (setprod (\<lambda>i. ?g i $ p i) ?U)" unfolding kU[symmetric] .
+  then show "of_int (sign p) * setprod (\<lambda>i. ?f i $ p i) ?U = c * (of_int (sign p) * setprod (\<lambda>i. ?g i $ p i) ?U)"
+    by (simp add: ring_simps)
+qed
+
+lemma det_row_0:
+  fixes b :: "'n::finite \<Rightarrow> _ ^ 'n"
+  shows "det((\<chi> i. if i = k then 0 else b i)::'a::comm_ring_1^'n^'n) = 0"
+using det_row_mul[of k 0 "\<lambda>i. 1" b]
+apply (simp)
+  unfolding vector_smult_lzero .
+
+lemma det_row_operation:
+  fixes A :: "'a::ordered_idom^'n^'n::finite"
+  assumes ij: "i \<noteq> j"
+  shows "det (\<chi> k. if k = i then row i A + c *s row j A else row k A) = det A"
+proof-
+  let ?Z = "(\<chi> k. if k = i then row j A else row k A) :: 'a ^'n^'n"
+  have th: "row i ?Z = row j ?Z" by (vector row_def)
+  have th2: "((\<chi> k. if k = i then row i A else row k A) :: 'a^'n^'n) = A"
+    by (vector row_def)
+  show ?thesis
+    unfolding det_row_add [of i] det_row_mul[of i] det_identical_rows[OF ij th] th2
+    by simp
+qed
+
+lemma det_row_span:
+  fixes A :: "'a:: ordered_idom^'n^'n::finite"
+  assumes x: "x \<in> span {row j A |j. j \<noteq> i}"
+  shows "det (\<chi> k. if k = i then row i A + x else row k A) = det A"
+proof-
+  let ?U = "UNIV :: 'n set"
+  let ?S = "{row j A |j. j \<noteq> i}"
+  let ?d = "\<lambda>x. det (\<chi> k. if k = i then x else row k A)"
+  let ?P = "\<lambda>x. ?d (row i A + x) = det A"
+  {fix k
+
+    have "(if k = i then row i A + 0 else row k A) = row k A" by simp}
+  then have P0: "?P 0"
+    apply -
+    apply (rule cong[of det, OF refl])
+    by (vector row_def)
+  moreover
+  {fix c z y assume zS: "z \<in> ?S" and Py: "?P y"
+    from zS obtain j where j: "z = row j A" "i \<noteq> j" by blast
+    let ?w = "row i A + y"
+    have th0: "row i A + (c*s z + y) = ?w + c*s z" by vector
+    have thz: "?d z = 0"
+      apply (rule det_identical_rows[OF j(2)])
+      using j by (vector row_def)
+    have "?d (row i A + (c*s z + y)) = ?d (?w + c*s z)" unfolding th0 ..
+    then have "?P (c*s z + y)" unfolding thz Py det_row_mul[of i] det_row_add[of i]
+      by simp }
+
+  ultimately show ?thesis
+    apply -
+    apply (rule span_induct_alt[of ?P ?S, OF P0])
+    apply blast
+    apply (rule x)
+    done
+qed
+
+(* ------------------------------------------------------------------------- *)
+(* May as well do this, though it's a bit unsatisfactory since it ignores    *)
+(* exact duplicates by considering the rows/columns as a set.                *)
+(* ------------------------------------------------------------------------- *)
+
+lemma det_dependent_rows:
+  fixes A:: "'a::ordered_idom^'n^'n::finite"
+  assumes d: "dependent (rows A)"
+  shows "det A = 0"
+proof-
+  let ?U = "UNIV :: 'n set"
+  from d obtain i where i: "row i A \<in> span (rows A - {row i A})"
+    unfolding dependent_def rows_def by blast
+  {fix j k assume jk: "j \<noteq> k"
+    and c: "row j A = row k A"
+    from det_identical_rows[OF jk c] have ?thesis .}
+  moreover
+  {assume H: "\<And> i j. i \<noteq> j \<Longrightarrow> row i A \<noteq> row j A"
+    have th0: "- row i A \<in> span {row j A|j. j \<noteq> i}"
+      apply (rule span_neg)
+      apply (rule set_rev_mp)
+      apply (rule i)
+      apply (rule span_mono)
+      using H i by (auto simp add: rows_def)
+    from det_row_span[OF th0]
+    have "det A = det (\<chi> k. if k = i then 0 *s 1 else row k A)"
+      unfolding right_minus vector_smult_lzero ..
+    with det_row_mul[of i "0::'a" "\<lambda>i. 1"]
+    have "det A = 0" by simp}
+  ultimately show ?thesis by blast
+qed
+
+lemma det_dependent_columns: assumes d: "dependent(columns (A::'a::ordered_idom^'n^'n::finite))" shows "det A = 0"
+by (metis d det_dependent_rows rows_transp det_transp)
+
+(* ------------------------------------------------------------------------- *)
+(* Multilinearity and the multiplication formula.                            *)
+(* ------------------------------------------------------------------------- *)
+
+lemma Cart_lambda_cong: "(\<And>x. f x = g x) \<Longrightarrow> (Cart_lambda f::'a^'n) = (Cart_lambda g :: 'a^'n)"
+  apply (rule iffD1[OF Cart_lambda_unique]) by vector
+
+lemma det_linear_row_setsum:
+  assumes fS: "finite S"
+  shows "det ((\<chi> i. if i = k then setsum (a i) S else c i)::'a::comm_ring_1^'n^'n::finite) = setsum (\<lambda>j. det ((\<chi> i. if i = k then a  i j else c i)::'a^'n^'n)) S"
+proof(induct rule: finite_induct[OF fS])
+  case 1 thus ?case apply simp  unfolding setsum_empty det_row_0[of k] ..
+next
+  case (2 x F)
+  then  show ?case by (simp add: det_row_add cong del: if_weak_cong)
+qed
+
+lemma finite_bounded_functions:
+  assumes fS: "finite S"
+  shows "finite {f. (\<forall>i \<in> {1.. (k::nat)}. f i \<in> S) \<and> (\<forall>i. i \<notin> {1 .. k} \<longrightarrow> f i = i)}"
+proof(induct k)
+  case 0
+  have th: "{f. \<forall>i. f i = i} = {id}" by (auto intro: ext)
+  show ?case by (auto simp add: th)
+next
+  case (Suc k)
+  let ?f = "\<lambda>(y::nat,g) i. if i = Suc k then y else g i"
+  let ?S = "?f ` (S \<times> {f. (\<forall>i\<in>{1..k}. f i \<in> S) \<and> (\<forall>i. i \<notin> {1..k} \<longrightarrow> f i = i)})"
+  have "?S = {f. (\<forall>i\<in>{1.. Suc k}. f i \<in> S) \<and> (\<forall>i. i \<notin> {1.. Suc k} \<longrightarrow> f i = i)}"
+    apply (auto simp add: image_iff)
+    apply (rule_tac x="x (Suc k)" in bexI)
+    apply (rule_tac x = "\<lambda>i. if i = Suc k then i else x i" in exI)
+    apply (auto intro: ext)
+    done
+  with finite_imageI[OF finite_cartesian_product[OF fS Suc.hyps(1)], of ?f]
+  show ?case by metis
+qed
+
+
+lemma eq_id_iff[simp]: "(\<forall>x. f x = x) = (f = id)" by (auto intro: ext)
+
+lemma det_linear_rows_setsum_lemma:
+  assumes fS: "finite S" and fT: "finite T"
+  shows "det((\<chi> i. if i \<in> T then setsum (a i) S else c i):: 'a::comm_ring_1^'n^'n::finite) =
+             setsum (\<lambda>f. det((\<chi> i. if i \<in> T then a i (f i) else c i)::'a^'n^'n))
+                 {f. (\<forall>i \<in> T. f i \<in> S) \<and> (\<forall>i. i \<notin> T \<longrightarrow> f i = i)}"
+using fT
+proof(induct T arbitrary: a c set: finite)
+  case empty
+  have th0: "\<And>x y. (\<chi> i. if i \<in> {} then x i else y i) = (\<chi> i. y i)" by vector
+  from "empty.prems"  show ?case unfolding th0 by simp
+next
+  case (insert z T a c)
+  let ?F = "\<lambda>T. {f. (\<forall>i \<in> T. f i \<in> S) \<and> (\<forall>i. i \<notin> T \<longrightarrow> f i = i)}"
+  let ?h = "\<lambda>(y,g) i. if i = z then y else g i"
+  let ?k = "\<lambda>h. (h(z),(\<lambda>i. if i = z then i else h i))"
+  let ?s = "\<lambda> k a c f. det((\<chi> i. if i \<in> T then a i (f i) else c i)::'a^'n^'n)"
+  let ?c = "\<lambda>i. if i = z then a i j else c i"
+  have thif: "\<And>a b c d. (if a \<or> b then c else d) = (if a then c else if b then c else d)" by simp
+  have thif2: "\<And>a b c d e. (if a then b else if c then d else e) =
+     (if c then (if a then b else d) else (if a then b else e))" by simp
+  from `z \<notin> T` have nz: "\<And>i. i \<in> T \<Longrightarrow> i = z \<longleftrightarrow> False" by auto
+  have "det (\<chi> i. if i \<in> insert z T then setsum (a i) S else c i) =
+        det (\<chi> i. if i = z then setsum (a i) S
+                 else if i \<in> T then setsum (a i) S else c i)"
+    unfolding insert_iff thif ..
+  also have "\<dots> = (\<Sum>j\<in>S. det (\<chi> i. if i \<in> T then setsum (a i) S
+                    else if i = z then a i j else c i))"
+    unfolding det_linear_row_setsum[OF fS]
+    apply (subst thif2)
+    using nz by (simp cong del: if_weak_cong cong add: if_cong)
+  finally have tha:
+    "det (\<chi> i. if i \<in> insert z T then setsum (a i) S else c i) =
+     (\<Sum>(j, f)\<in>S \<times> ?F T. det (\<chi> i. if i \<in> T then a i (f i)
+                                else if i = z then a i j
+                                else c i))"
+    unfolding  insert.hyps unfolding setsum_cartesian_product by blast
+  show ?case unfolding tha
+    apply(rule setsum_eq_general_reverses[where h= "?h" and k= "?k"],
+      blast intro: finite_cartesian_product fS finite,
+      blast intro: finite_cartesian_product fS finite)
+    using `z \<notin> T`
+    apply (auto intro: ext)
+    apply (rule cong[OF refl[of det]])
+    by vector
+qed
+
+lemma det_linear_rows_setsum:
+  assumes fS: "finite (S::'n::finite set)"
+  shows "det (\<chi> i. setsum (a i) S) = setsum (\<lambda>f. det (\<chi> i. a i (f i) :: 'a::comm_ring_1 ^ 'n^'n::finite)) {f. \<forall>i. f i \<in> S}"
+proof-
+  have th0: "\<And>x y. ((\<chi> i. if i \<in> (UNIV:: 'n set) then x i else y i) :: 'a^'n^'n) = (\<chi> i. x i)" by vector
+
+  from det_linear_rows_setsum_lemma[OF fS, of "UNIV :: 'n set" a, unfolded th0, OF finite] show ?thesis by simp
+qed
+
+lemma matrix_mul_setsum_alt:
+  fixes A B :: "'a::comm_ring_1^'n^'n::finite"
+  shows "A ** B = (\<chi> i. setsum (\<lambda>k. A$i$k *s B $ k) (UNIV :: 'n set))"
+  by (vector matrix_matrix_mult_def setsum_component)
+
+lemma det_rows_mul:
+  "det((\<chi> i. c i *s a i)::'a::comm_ring_1^'n^'n::finite) =
+  setprod (\<lambda>i. c i) (UNIV:: 'n set) * det((\<chi> i. a i)::'a^'n^'n)"
+proof (simp add: det_def setsum_right_distrib cong add: setprod_cong, rule setsum_cong2)
+  let ?U = "UNIV :: 'n set"
+  let ?PU = "{p. p permutes ?U}"
+  fix p assume pU: "p \<in> ?PU"
+  let ?s = "of_int (sign p)"
+  from pU have p: "p permutes ?U" by blast
+  have "setprod (\<lambda>i. c i * a i $ p i) ?U = setprod c ?U * setprod (\<lambda>i. a i $ p i) ?U"
+    unfolding setprod_timesf ..
+  then show "?s * (\<Prod>xa\<in>?U. c xa * a xa $ p xa) =
+        setprod c ?U * (?s* (\<Prod>xa\<in>?U. a xa $ p xa))" by (simp add: ring_simps)
+qed
+
+lemma det_mul:
+  fixes A B :: "'a::ordered_idom^'n^'n::finite"
+  shows "det (A ** B) = det A * det B"
+proof-
+  let ?U = "UNIV :: 'n set"
+  let ?F = "{f. (\<forall>i\<in> ?U. f i \<in> ?U) \<and> (\<forall>i. i \<notin> ?U \<longrightarrow> f i = i)}"
+  let ?PU = "{p. p permutes ?U}"
+  have fU: "finite ?U" by simp
+  have fF: "finite ?F" by (rule finite)
+  {fix p assume p: "p permutes ?U"
+
+    have "p \<in> ?F" unfolding mem_Collect_eq permutes_in_image[OF p]
+      using p[unfolded permutes_def] by simp}
+  then have PUF: "?PU \<subseteq> ?F"  by blast
+  {fix f assume fPU: "f \<in> ?F - ?PU"
+    have fUU: "f ` ?U \<subseteq> ?U" using fPU by auto
+    from fPU have f: "\<forall>i \<in> ?U. f i \<in> ?U"
+      "\<forall>i. i \<notin> ?U \<longrightarrow> f i = i" "\<not>(\<forall>y. \<exists>!x. f x = y)" unfolding permutes_def
+      by auto
+
+    let ?A = "(\<chi> i. A$i$f i *s B$f i) :: 'a^'n^'n"
+    let ?B = "(\<chi> i. B$f i) :: 'a^'n^'n"
+    {assume fni: "\<not> inj_on f ?U"
+      then obtain i j where ij: "f i = f j" "i \<noteq> j"
+        unfolding inj_on_def by blast
+      from ij
+      have rth: "row i ?B = row j ?B" by (vector row_def)
+      from det_identical_rows[OF ij(2) rth]
+      have "det (\<chi> i. A$i$f i *s B$f i) = 0"
+        unfolding det_rows_mul by simp}
+    moreover
+    {assume fi: "inj_on f ?U"
+      from f fi have fith: "\<And>i j. f i = f j \<Longrightarrow> i = j"
+        unfolding inj_on_def by metis
+      note fs = fi[unfolded surjective_iff_injective_gen[OF fU fU refl fUU, symmetric]]
+
+      {fix y
+        from fs f have "\<exists>x. f x = y" by blast
+        then obtain x where x: "f x = y" by blast
+        {fix z assume z: "f z = y" from fith x z have "z = x" by metis}
+        with x have "\<exists>!x. f x = y" by blast}
+      with f(3) have "det (\<chi> i. A$i$f i *s B$f i) = 0" by blast}
+    ultimately have "det (\<chi> i. A$i$f i *s B$f i) = 0" by blast}
+  hence zth: "\<forall> f\<in> ?F - ?PU. det (\<chi> i. A$i$f i *s B$f i) = 0" by simp
+  {fix p assume pU: "p \<in> ?PU"
+    from pU have p: "p permutes ?U" by blast
+    let ?s = "\<lambda>p. of_int (sign p)"
+    let ?f = "\<lambda>q. ?s p * (\<Prod>i\<in> ?U. A $ i $ p i) *
+               (?s q * (\<Prod>i\<in> ?U. B $ i $ q i))"
+    have "(setsum (\<lambda>q. ?s q *
+            (\<Prod>i\<in> ?U. (\<chi> i. A $ i $ p i *s B $ p i :: 'a^'n^'n) $ i $ q i)) ?PU) =
+        (setsum (\<lambda>q. ?s p * (\<Prod>i\<in> ?U. A $ i $ p i) *
+               (?s q * (\<Prod>i\<in> ?U. B $ i $ q i))) ?PU)"
+      unfolding sum_permutations_compose_right[OF permutes_inv[OF p], of ?f]
+    proof(rule setsum_cong2)
+      fix q assume qU: "q \<in> ?PU"
+      hence q: "q permutes ?U" by blast
+      from p q have pp: "permutation p" and pq: "permutation q"
+        unfolding permutation_permutes by auto
+      have th00: "of_int (sign p) * of_int (sign p) = (1::'a)"
+        "\<And>a. of_int (sign p) * (of_int (sign p) * a) = a"
+        unfolding mult_assoc[symmetric] unfolding of_int_mult[symmetric]
+        by (simp_all add: sign_idempotent)
+      have ths: "?s q = ?s p * ?s (q o inv p)"
+        using pp pq permutation_inverse[OF pp] sign_inverse[OF pp]
+        by (simp add:  th00 mult_ac sign_idempotent sign_compose)
+      have th001: "setprod (\<lambda>i. B$i$ q (inv p i)) ?U = setprod ((\<lambda>i. B$i$ q (inv p i)) o p) ?U"
+        by (rule setprod_permute[OF p])
+      have thp: "setprod (\<lambda>i. (\<chi> i. A$i$p i *s B$p i :: 'a^'n^'n) $i $ q i) ?U = setprod (\<lambda>i. A$i$p i) ?U * setprod (\<lambda>i. B$i$ q (inv p i)) ?U"
+        unfolding th001 setprod_timesf[symmetric] o_def permutes_inverses[OF p]
+        apply (rule setprod_cong[OF refl])
+        using permutes_in_image[OF q] by vector
+      show "?s q * setprod (\<lambda>i. (((\<chi> i. A$i$p i *s B$p i) :: 'a^'n^'n)$i$q i)) ?U = ?s p * (setprod (\<lambda>i. A$i$p i) ?U) * (?s (q o inv p) * setprod (\<lambda>i. B$i$(q o inv p) i) ?U)"
+        using ths thp pp pq permutation_inverse[OF pp] sign_inverse[OF pp]
+        by (simp add: sign_nz th00 ring_simps sign_idempotent sign_compose)
+    qed
+  }
+  then have th2: "setsum (\<lambda>f. det (\<chi> i. A$i$f i *s B$f i)) ?PU = det A * det B"
+    unfolding det_def setsum_product
+    by (rule setsum_cong2)
+  have "det (A**B) = setsum (\<lambda>f.  det (\<chi> i. A $ i $ f i *s B $ f i)) ?F"
+    unfolding matrix_mul_setsum_alt det_linear_rows_setsum[OF fU] by simp
+  also have "\<dots> = setsum (\<lambda>f. det (\<chi> i. A$i$f i *s B$f i)) ?PU"
+    using setsum_mono_zero_cong_left[OF fF PUF zth, symmetric]
+    unfolding det_rows_mul by auto
+  finally show ?thesis unfolding th2 .
+qed
+
+(* ------------------------------------------------------------------------- *)
+(* Relation to invertibility.                                                *)
+(* ------------------------------------------------------------------------- *)
+
+lemma invertible_left_inverse:
+  fixes A :: "real^'n^'n::finite"
+  shows "invertible A \<longleftrightarrow> (\<exists>(B::real^'n^'n). B** A = mat 1)"
+  by (metis invertible_def matrix_left_right_inverse)
+
+lemma invertible_righ_inverse:
+  fixes A :: "real^'n^'n::finite"
+  shows "invertible A \<longleftrightarrow> (\<exists>(B::real^'n^'n). A** B = mat 1)"
+  by (metis invertible_def matrix_left_right_inverse)
+
+lemma invertible_det_nz:
+  fixes A::"real ^'n^'n::finite"
+  shows "invertible A \<longleftrightarrow> det A \<noteq> 0"
+proof-
+  {assume "invertible A"
+    then obtain B :: "real ^'n^'n" where B: "A ** B = mat 1"
+      unfolding invertible_righ_inverse by blast
+    hence "det (A ** B) = det (mat 1 :: real ^'n^'n)" by simp
+    hence "det A \<noteq> 0"
+      apply (simp add: det_mul det_I) by algebra }
+  moreover
+  {assume H: "\<not> invertible A"
+    let ?U = "UNIV :: 'n set"
+    have fU: "finite ?U" by simp
+    from H obtain c i where c: "setsum (\<lambda>i. c i *s row i A) ?U = 0"
+      and iU: "i \<in> ?U" and ci: "c i \<noteq> 0"
+      unfolding invertible_righ_inverse
+      unfolding matrix_right_invertible_independent_rows by blast
+    have stupid: "\<And>(a::real^'n) b. a + b = 0 \<Longrightarrow> -a = b"
+      apply (drule_tac f="op + (- a)" in cong[OF refl])
+      apply (simp only: ab_left_minus add_assoc[symmetric])
+      apply simp
+      done
+    from c ci
+    have thr0: "- row i A = setsum (\<lambda>j. (1/ c i) *s (c j *s row j A)) (?U - {i})"
+      unfolding setsum_diff1'[OF fU iU] setsum_cmul
+      apply -
+      apply (rule vector_mul_lcancel_imp[OF ci])
+      apply (auto simp add: vector_smult_assoc vector_smult_rneg field_simps)
+      unfolding stupid ..
+    have thr: "- row i A \<in> span {row j A| j. j \<noteq> i}"
+      unfolding thr0
+      apply (rule span_setsum)
+      apply simp
+      apply (rule ballI)
+      apply (rule span_mul)+
+      apply (rule span_superset)
+      apply auto
+      done
+    let ?B = "(\<chi> k. if k = i then 0 else row k A) :: real ^'n^'n"
+    have thrb: "row i ?B = 0" using iU by (vector row_def)
+    have "det A = 0"
+      unfolding det_row_span[OF thr, symmetric] right_minus
+      unfolding  det_zero_row[OF thrb]  ..}
+  ultimately show ?thesis by blast
+qed
+
+(* ------------------------------------------------------------------------- *)
+(* Cramer's rule.                                                            *)
+(* ------------------------------------------------------------------------- *)
+
+lemma cramer_lemma_transp:
+  fixes A:: "'a::ordered_idom^'n^'n::finite" and x :: "'a ^'n::finite"
+  shows "det ((\<chi> i. if i = k then setsum (\<lambda>i. x$i *s row i A) (UNIV::'n set)
+                           else row i A)::'a^'n^'n) = x$k * det A"
+  (is "?lhs = ?rhs")
+proof-
+  let ?U = "UNIV :: 'n set"
+  let ?Uk = "?U - {k}"
+  have U: "?U = insert k ?Uk" by blast
+  have fUk: "finite ?Uk" by simp
+  have kUk: "k \<notin> ?Uk" by simp
+  have th00: "\<And>k s. x$k *s row k A + s = (x$k - 1) *s row k A + row k A + s"
+    by (vector ring_simps)
+  have th001: "\<And>f k . (\<lambda>x. if x = k then f k else f x) = f" by (auto intro: ext)
+  have "(\<chi> i. row i A) = A" by (vector row_def)
+  then have thd1: "det (\<chi> i. row i A) = det A"  by simp
+  have thd0: "det (\<chi> i. if i = k then row k A + (\<Sum>i \<in> ?Uk. x $ i *s row i A) else row i A) = det A"
+    apply (rule det_row_span)
+    apply (rule span_setsum[OF fUk])
+    apply (rule ballI)
+    apply (rule span_mul)
+    apply (rule span_superset)
+    apply auto
+    done
+  show "?lhs = x$k * det A"
+    apply (subst U)
+    unfolding setsum_insert[OF fUk kUk]
+    apply (subst th00)
+    unfolding add_assoc
+    apply (subst det_row_add)
+    unfolding thd0
+    unfolding det_row_mul
+    unfolding th001[of k "\<lambda>i. row i A"]
+    unfolding thd1  by (simp add: ring_simps)
+qed
+
+lemma cramer_lemma:
+  fixes A :: "'a::ordered_idom ^'n^'n::finite"
+  shows "det((\<chi> i j. if j = k then (A *v x)$i else A$i$j):: 'a^'n^'n) = x$k * det A"
+proof-
+  let ?U = "UNIV :: 'n set"
+  have stupid: "\<And>c. setsum (\<lambda>i. c i *s row i (transp A)) ?U = setsum (\<lambda>i. c i *s column i A) ?U"
+    by (auto simp add: row_transp intro: setsum_cong2)
+  show ?thesis  unfolding matrix_mult_vsum
+  unfolding cramer_lemma_transp[of k x "transp A", unfolded det_transp, symmetric]
+  unfolding stupid[of "\<lambda>i. x$i"]
+  apply (subst det_transp[symmetric])
+  apply (rule cong[OF refl[of det]]) by (vector transp_def column_def row_def)
+qed
+
+lemma cramer:
+  fixes A ::"real^'n^'n::finite"
+  assumes d0: "det A \<noteq> 0"
+  shows "A *v x = b \<longleftrightarrow> x = (\<chi> k. det(\<chi> i j. if j=k then b$i else A$i$j :: real^'n^'n) / det A)"
+proof-
+  from d0 obtain B where B: "A ** B = mat 1" "B ** A = mat 1"
+    unfolding invertible_det_nz[symmetric] invertible_def by blast
+  have "(A ** B) *v b = b" by (simp add: B matrix_vector_mul_lid)
+  hence "A *v (B *v b) = b" by (simp add: matrix_vector_mul_assoc)
+  then have xe: "\<exists>x. A*v x = b" by blast
+  {fix x assume x: "A *v x = b"
+  have "x = (\<chi> k. det(\<chi> i j. if j=k then b$i else A$i$j :: real^'n^'n) / det A)"
+    unfolding x[symmetric]
+    using d0 by (simp add: Cart_eq cramer_lemma field_simps)}
+  with xe show ?thesis by auto
+qed
+
+(* ------------------------------------------------------------------------- *)
+(* Orthogonality of a transformation and matrix.                             *)
+(* ------------------------------------------------------------------------- *)
+
+definition "orthogonal_transformation f \<longleftrightarrow> linear f \<and> (\<forall>v w. f v \<bullet> f w = v \<bullet> w)"
+
+lemma orthogonal_transformation: "orthogonal_transformation f \<longleftrightarrow> linear f \<and> (\<forall>(v::real ^_). norm (f v) = norm v)"
+  unfolding orthogonal_transformation_def
+  apply auto
+  apply (erule_tac x=v in allE)+
+  apply (simp add: real_vector_norm_def)
+  by (simp add: dot_norm  linear_add[symmetric])
+
+definition "orthogonal_matrix (Q::'a::semiring_1^'n^'n) \<longleftrightarrow> transp Q ** Q = mat 1 \<and> Q ** transp Q = mat 1"
+
+lemma orthogonal_matrix: "orthogonal_matrix (Q:: real ^'n^'n::finite)  \<longleftrightarrow> transp Q ** Q = mat 1"
+  by (metis matrix_left_right_inverse orthogonal_matrix_def)
+
+lemma orthogonal_matrix_id: "orthogonal_matrix (mat 1 :: _^'n^'n::finite)"
+  by (simp add: orthogonal_matrix_def transp_mat matrix_mul_lid)
+
+lemma orthogonal_matrix_mul:
+  fixes A :: "real ^'n^'n::finite"
+  assumes oA : "orthogonal_matrix A"
+  and oB: "orthogonal_matrix B"
+  shows "orthogonal_matrix(A ** B)"
+  using oA oB
+  unfolding orthogonal_matrix matrix_transp_mul
+  apply (subst matrix_mul_assoc)
+  apply (subst matrix_mul_assoc[symmetric])
+  by (simp add: matrix_mul_rid)
+
+lemma orthogonal_transformation_matrix:
+  fixes f:: "real^'n \<Rightarrow> real^'n::finite"
+  shows "orthogonal_transformation f \<longleftrightarrow> linear f \<and> orthogonal_matrix(matrix f)"
+  (is "?lhs \<longleftrightarrow> ?rhs")
+proof-
+  let ?mf = "matrix f"
+  let ?ot = "orthogonal_transformation f"
+  let ?U = "UNIV :: 'n set"
+  have fU: "finite ?U" by simp
+  let ?m1 = "mat 1 :: real ^'n^'n"
+  {assume ot: ?ot
+    from ot have lf: "linear f" and fd: "\<forall>v w. f v \<bullet> f w = v \<bullet> w"
+      unfolding  orthogonal_transformation_def orthogonal_matrix by blast+
+    {fix i j
+      let ?A = "transp ?mf ** ?mf"
+      have th0: "\<And>b (x::'a::comm_ring_1). (if b then 1 else 0)*x = (if b then x else 0)"
+        "\<And>b (x::'a::comm_ring_1). x*(if b then 1 else 0) = (if b then x else 0)"
+        by simp_all
+      from fd[rule_format, of "basis i" "basis j", unfolded matrix_works[OF lf, symmetric] dot_matrix_vector_mul]
+      have "?A$i$j = ?m1 $ i $ j"
+        by (simp add: dot_def matrix_matrix_mult_def columnvector_def rowvector_def basis_def th0 setsum_delta[OF fU] mat_def)}
+    hence "orthogonal_matrix ?mf" unfolding orthogonal_matrix by vector
+    with lf have ?rhs by blast}
+  moreover
+  {assume lf: "linear f" and om: "orthogonal_matrix ?mf"
+    from lf om have ?lhs
+      unfolding orthogonal_matrix_def norm_eq orthogonal_transformation
+      unfolding matrix_works[OF lf, symmetric]
+      apply (subst dot_matrix_vector_mul)
+      by (simp add: dot_matrix_product matrix_mul_lid)}
+  ultimately show ?thesis by blast
+qed
+
+lemma det_orthogonal_matrix:
+  fixes Q:: "'a::ordered_idom^'n^'n::finite"
+  assumes oQ: "orthogonal_matrix Q"
+  shows "det Q = 1 \<or> det Q = - 1"
+proof-
+
+  have th: "\<And>x::'a. x = 1 \<or> x = - 1 \<longleftrightarrow> x*x = 1" (is "\<And>x::'a. ?ths x")
+  proof-
+    fix x:: 'a
+    have th0: "x*x - 1 = (x - 1)*(x + 1)" by (simp add: ring_simps)
+    have th1: "\<And>(x::'a) y. x = - y \<longleftrightarrow> x + y = 0"
+      apply (subst eq_iff_diff_eq_0) by simp
+    have "x*x = 1 \<longleftrightarrow> x*x - 1 = 0" by simp
+    also have "\<dots> \<longleftrightarrow> x = 1 \<or> x = - 1" unfolding th0 th1 by simp
+    finally show "?ths x" ..
+  qed
+  from oQ have "Q ** transp Q = mat 1" by (metis orthogonal_matrix_def)
+  hence "det (Q ** transp Q) = det (mat 1:: 'a^'n^'n)" by simp
+  hence "det Q * det Q = 1" by (simp add: det_mul det_I det_transp)
+  then show ?thesis unfolding th .
+qed
+
+(* ------------------------------------------------------------------------- *)
+(* Linearity of scaling, and hence isometry, that preserves origin.          *)
+(* ------------------------------------------------------------------------- *)
+lemma scaling_linear:
+  fixes f :: "real ^'n \<Rightarrow> real ^'n::finite"
+  assumes f0: "f 0 = 0" and fd: "\<forall>x y. dist (f x) (f y) = c * dist x y"
+  shows "linear f"
+proof-
+  {fix v w
+    {fix x note fd[rule_format, of x 0, unfolded dist_norm f0 diff_0_right] }
+    note th0 = this
+    have "f v \<bullet> f w = c^2 * (v \<bullet> w)"
+      unfolding dot_norm_neg dist_norm[symmetric]
+      unfolding th0 fd[rule_format] by (simp add: power2_eq_square field_simps)}
+  note fc = this
+  show ?thesis unfolding linear_def vector_eq
+    by (simp add: dot_lmult dot_ladd dot_rmult dot_radd fc ring_simps)
+qed
+
+lemma isometry_linear:
+  "f (0:: real^'n) = (0:: real^'n::finite) \<Longrightarrow> \<forall>x y. dist(f x) (f y) = dist x y
+        \<Longrightarrow> linear f"
+by (rule scaling_linear[where c=1]) simp_all
+
+(* ------------------------------------------------------------------------- *)
+(* Hence another formulation of orthogonal transformation.                   *)
+(* ------------------------------------------------------------------------- *)
+
+lemma orthogonal_transformation_isometry:
+  "orthogonal_transformation f \<longleftrightarrow> f(0::real^'n) = (0::real^'n::finite) \<and> (\<forall>x y. dist(f x) (f y) = dist x y)"
+  unfolding orthogonal_transformation
+  apply (rule iffI)
+  apply clarify
+  apply (clarsimp simp add: linear_0 linear_sub[symmetric] dist_norm)
+  apply (rule conjI)
+  apply (rule isometry_linear)
+  apply simp
+  apply simp
+  apply clarify
+  apply (erule_tac x=v in allE)
+  apply (erule_tac x=0 in allE)
+  by (simp add: dist_norm)
+
+(* ------------------------------------------------------------------------- *)
+(* Can extend an isometry from unit sphere.                                  *)
+(* ------------------------------------------------------------------------- *)
+
+lemma isometry_sphere_extend:
+  fixes f:: "real ^'n \<Rightarrow> real ^'n::finite"
+  assumes f1: "\<forall>x. norm x = 1 \<longrightarrow> norm (f x) = 1"
+  and fd1: "\<forall> x y. norm x = 1 \<longrightarrow> norm y = 1 \<longrightarrow> dist (f x) (f y) = dist x y"
+  shows "\<exists>g. orthogonal_transformation g \<and> (\<forall>x. norm x = 1 \<longrightarrow> g x = f x)"
+proof-
+  {fix x y x' y' x0 y0 x0' y0' :: "real ^'n"
+    assume H: "x = norm x *s x0" "y = norm y *s y0"
+    "x' = norm x *s x0'" "y' = norm y *s y0'"
+    "norm x0 = 1" "norm x0' = 1" "norm y0 = 1" "norm y0' = 1"
+    "norm(x0' - y0') = norm(x0 - y0)"
+
+    have "norm(x' - y') = norm(x - y)"
+      apply (subst H(1))
+      apply (subst H(2))
+      apply (subst H(3))
+      apply (subst H(4))
+      using H(5-9)
+      apply (simp add: norm_eq norm_eq_1)
+      apply (simp add: dot_lsub dot_rsub dot_lmult dot_rmult)
+      apply (simp add: ring_simps)
+      by (simp only: right_distrib[symmetric])}
+  note th0 = this
+  let ?g = "\<lambda>x. if x = 0 then 0 else norm x *s f (inverse (norm x) *s x)"
+  {fix x:: "real ^'n" assume nx: "norm x = 1"
+    have "?g x = f x" using nx by auto}
+  hence thfg: "\<forall>x. norm x = 1 \<longrightarrow> ?g x = f x" by blast
+  have g0: "?g 0 = 0" by simp
+  {fix x y :: "real ^'n"
+    {assume "x = 0" "y = 0"
+      then have "dist (?g x) (?g y) = dist x y" by simp }
+    moreover
+    {assume "x = 0" "y \<noteq> 0"
+      then have "dist (?g x) (?g y) = dist x y"
+        apply (simp add: dist_norm norm_mul)
+        apply (rule f1[rule_format])
+        by(simp add: norm_mul field_simps)}
+    moreover
+    {assume "x \<noteq> 0" "y = 0"
+      then have "dist (?g x) (?g y) = dist x y"
+        apply (simp add: dist_norm norm_mul)
+        apply (rule f1[rule_format])
+        by(simp add: norm_mul field_simps)}
+    moreover
+    {assume z: "x \<noteq> 0" "y \<noteq> 0"
+      have th00: "x = norm x *s (inverse (norm x) *s x)" "y = norm y *s (inverse (norm y) *s y)" "norm x *s f ((inverse (norm x) *s x)) = norm x *s f (inverse (norm x) *s x)"
+        "norm y *s f (inverse (norm y) *s y) = norm y *s f (inverse (norm y) *s y)"
+        "norm (inverse (norm x) *s x) = 1"
+        "norm (f (inverse (norm x) *s x)) = 1"
+        "norm (inverse (norm y) *s y) = 1"
+        "norm (f (inverse (norm y) *s y)) = 1"
+        "norm (f (inverse (norm x) *s x) - f (inverse (norm y) *s y)) =
+        norm (inverse (norm x) *s x - inverse (norm y) *s y)"
+        using z
+        by (auto simp add: vector_smult_assoc field_simps norm_mul intro: f1[rule_format] fd1[rule_format, unfolded dist_norm])
+      from z th0[OF th00] have "dist (?g x) (?g y) = dist x y"
+        by (simp add: dist_norm)}
+    ultimately have "dist (?g x) (?g y) = dist x y" by blast}
+  note thd = this
+    show ?thesis
+    apply (rule exI[where x= ?g])
+    unfolding orthogonal_transformation_isometry
+      using  g0 thfg thd by metis
+qed
+
+(* ------------------------------------------------------------------------- *)
+(* Rotation, reflection, rotoinversion.                                      *)
+(* ------------------------------------------------------------------------- *)
+
+definition "rotation_matrix Q \<longleftrightarrow> orthogonal_matrix Q \<and> det Q = 1"
+definition "rotoinversion_matrix Q \<longleftrightarrow> orthogonal_matrix Q \<and> det Q = - 1"
+
+lemma orthogonal_rotation_or_rotoinversion:
+  fixes Q :: "'a::ordered_idom^'n^'n::finite"
+  shows " orthogonal_matrix Q \<longleftrightarrow> rotation_matrix Q \<or> rotoinversion_matrix Q"
+  by (metis rotoinversion_matrix_def rotation_matrix_def det_orthogonal_matrix)
+(* ------------------------------------------------------------------------- *)
+(* Explicit formulas for low dimensions.                                     *)
+(* ------------------------------------------------------------------------- *)
+
+lemma setprod_1: "setprod f {(1::nat)..1} = f 1" by simp
+
+lemma setprod_2: "setprod f {(1::nat)..2} = f 1 * f 2"
+  by (simp add: nat_number setprod_numseg mult_commute)
+lemma setprod_3: "setprod f {(1::nat)..3} = f 1 * f 2 * f 3"
+  by (simp add: nat_number setprod_numseg mult_commute)
+
+lemma det_1: "det (A::'a::comm_ring_1^1^1) = A$1$1"
+  by (simp add: det_def permutes_sing sign_id UNIV_1)
+
+lemma det_2: "det (A::'a::comm_ring_1^2^2) = A$1$1 * A$2$2 - A$1$2 * A$2$1"
+proof-
+  have f12: "finite {2::2}" "1 \<notin> {2::2}" by auto
+  show ?thesis
+  unfolding det_def UNIV_2
+  unfolding setsum_over_permutations_insert[OF f12]
+  unfolding permutes_sing
+  apply (simp add: sign_swap_id sign_id swap_id_eq)
+  by (simp add: arith_simps(31)[symmetric] of_int_minus of_int_1 del: arith_simps(31))
+qed
+
+lemma det_3: "det (A::'a::comm_ring_1^3^3) =
+  A$1$1 * A$2$2 * A$3$3 +
+  A$1$2 * A$2$3 * A$3$1 +
+  A$1$3 * A$2$1 * A$3$2 -
+  A$1$1 * A$2$3 * A$3$2 -
+  A$1$2 * A$2$1 * A$3$3 -
+  A$1$3 * A$2$2 * A$3$1"
+proof-
+  have f123: "finite {2::3, 3}" "1 \<notin> {2::3, 3}" by auto
+  have f23: "finite {3::3}" "2 \<notin> {3::3}" by auto
+
+  show ?thesis
+  unfolding det_def UNIV_3
+  unfolding setsum_over_permutations_insert[OF f123]
+  unfolding setsum_over_permutations_insert[OF f23]
+
+  unfolding permutes_sing
+  apply (simp add: sign_swap_id permutation_swap_id sign_compose sign_id swap_id_eq)
+  apply (simp add: arith_simps(31)[symmetric] of_int_minus of_int_1 del: arith_simps(31))
+  by (simp add: ring_simps)
+qed
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Multivariate_Analysis/Euclidean_Space.thy	Fri Oct 23 13:23:18 2009 +0200
@@ -0,0 +1,5372 @@
+(*  Title:      Library/Euclidean_Space
+    Author:     Amine Chaieb, University of Cambridge
+*)
+
+header {* (Real) Vectors in Euclidean space, and elementary linear algebra.*}
+
+theory Euclidean_Space
+imports
+  Complex_Main "~~/src/HOL/Decision_Procs/Dense_Linear_Order"
+  Finite_Cartesian_Product Glbs Infinite_Set Numeral_Type
+  Inner_Product
+uses "positivstellensatz.ML" ("normarith.ML")
+begin
+
+text{* Some common special cases.*}
+
+lemma forall_1: "(\<forall>i::1. P i) \<longleftrightarrow> P 1"
+  by (metis num1_eq_iff)
+
+lemma exhaust_2:
+  fixes x :: 2 shows "x = 1 \<or> x = 2"
+proof (induct x)
+  case (of_int z)
+  then have "0 <= z" and "z < 2" by simp_all
+  then have "z = 0 | z = 1" by arith
+  then show ?case by auto
+qed
+
+lemma forall_2: "(\<forall>i::2. P i) \<longleftrightarrow> P 1 \<and> P 2"
+  by (metis exhaust_2)
+
+lemma exhaust_3:
+  fixes x :: 3 shows "x = 1 \<or> x = 2 \<or> x = 3"
+proof (induct x)
+  case (of_int z)
+  then have "0 <= z" and "z < 3" by simp_all
+  then have "z = 0 \<or> z = 1 \<or> z = 2" by arith
+  then show ?case by auto
+qed
+
+lemma forall_3: "(\<forall>i::3. P i) \<longleftrightarrow> P 1 \<and> P 2 \<and> P 3"
+  by (metis exhaust_3)
+
+lemma UNIV_1: "UNIV = {1::1}"
+  by (auto simp add: num1_eq_iff)
+
+lemma UNIV_2: "UNIV = {1::2, 2::2}"
+  using exhaust_2 by auto
+
+lemma UNIV_3: "UNIV = {1::3, 2::3, 3::3}"
+  using exhaust_3 by auto
+
+lemma setsum_1: "setsum f (UNIV::1 set) = f 1"
+  unfolding UNIV_1 by simp
+
+lemma setsum_2: "setsum f (UNIV::2 set) = f 1 + f 2"
+  unfolding UNIV_2 by simp
+
+lemma setsum_3: "setsum f (UNIV::3 set) = f 1 + f 2 + f 3"
+  unfolding UNIV_3 by (simp add: add_ac)
+
+subsection{* Basic componentwise operations on vectors. *}
+
+instantiation "^" :: (plus,type) plus
+begin
+definition  vector_add_def : "op + \<equiv> (\<lambda> x y.  (\<chi> i. (x$i) + (y$i)))"
+instance ..
+end
+
+instantiation "^" :: (times,type) times
+begin
+  definition vector_mult_def : "op * \<equiv> (\<lambda> x y.  (\<chi> i. (x$i) * (y$i)))"
+  instance ..
+end
+
+instantiation "^" :: (minus,type) minus begin
+  definition vector_minus_def : "op - \<equiv> (\<lambda> x y.  (\<chi> i. (x$i) - (y$i)))"
+instance ..
+end
+
+instantiation "^" :: (uminus,type) uminus begin
+  definition vector_uminus_def : "uminus \<equiv> (\<lambda> x.  (\<chi> i. - (x$i)))"
+instance ..
+end
+instantiation "^" :: (zero,type) zero begin
+  definition vector_zero_def : "0 \<equiv> (\<chi> i. 0)"
+instance ..
+end
+
+instantiation "^" :: (one,type) one begin
+  definition vector_one_def : "1 \<equiv> (\<chi> i. 1)"
+instance ..
+end
+
+instantiation "^" :: (ord,type) ord
+ begin
+definition vector_less_eq_def:
+  "less_eq (x :: 'a ^'b) y = (ALL i. x$i <= y$i)"
+definition vector_less_def: "less (x :: 'a ^'b) y = (ALL i. x$i < y$i)"
+
+instance by (intro_classes)
+end
+
+instantiation "^" :: (scaleR, type) scaleR
+begin
+definition vector_scaleR_def: "scaleR = (\<lambda> r x.  (\<chi> i. scaleR r (x$i)))"
+instance ..
+end
+
+text{* Also the scalar-vector multiplication. *}
+
+definition vector_scalar_mult:: "'a::times \<Rightarrow> 'a ^'n \<Rightarrow> 'a ^ 'n" (infixl "*s" 70)
+  where "c *s x = (\<chi> i. c * (x$i))"
+
+text{* Constant Vectors *} 
+
+definition "vec x = (\<chi> i. x)"
+
+text{* Dot products. *}
+
+definition dot :: "'a::{comm_monoid_add, times} ^ 'n \<Rightarrow> 'a ^ 'n \<Rightarrow> 'a" (infix "\<bullet>" 70) where
+  "x \<bullet> y = setsum (\<lambda>i. x$i * y$i) UNIV"
+
+lemma dot_1[simp]: "(x::'a::{comm_monoid_add, times}^1) \<bullet> y = (x$1) * (y$1)"
+  by (simp add: dot_def setsum_1)
+
+lemma dot_2[simp]: "(x::'a::{comm_monoid_add, times}^2) \<bullet> y = (x$1) * (y$1) + (x$2) * (y$2)"
+  by (simp add: dot_def setsum_2)
+
+lemma dot_3[simp]: "(x::'a::{comm_monoid_add, times}^3) \<bullet> y = (x$1) * (y$1) + (x$2) * (y$2) + (x$3) * (y$3)"
+  by (simp add: dot_def setsum_3)
+
+subsection {* A naive proof procedure to lift really trivial arithmetic stuff from the basis of the vector space. *}
+
+method_setup vector = {*
+let
+  val ss1 = HOL_basic_ss addsimps [@{thm dot_def}, @{thm setsum_addf} RS sym,
+  @{thm setsum_subtractf} RS sym, @{thm setsum_right_distrib},
+  @{thm setsum_left_distrib}, @{thm setsum_negf} RS sym]
+  val ss2 = @{simpset} addsimps
+             [@{thm vector_add_def}, @{thm vector_mult_def},
+              @{thm vector_minus_def}, @{thm vector_uminus_def},
+              @{thm vector_one_def}, @{thm vector_zero_def}, @{thm vec_def},
+              @{thm vector_scaleR_def},
+              @{thm Cart_lambda_beta}, @{thm vector_scalar_mult_def}]
+ fun vector_arith_tac ths =
+   simp_tac ss1
+   THEN' (fn i => rtac @{thm setsum_cong2} i
+         ORELSE rtac @{thm setsum_0'} i
+         ORELSE simp_tac (HOL_basic_ss addsimps [@{thm "Cart_eq"}]) i)
+   (* THEN' TRY o clarify_tac HOL_cs  THEN' (TRY o rtac @{thm iffI}) *)
+   THEN' asm_full_simp_tac (ss2 addsimps ths)
+ in
+  Attrib.thms >> (fn ths => K (SIMPLE_METHOD' (vector_arith_tac ths)))
+ end
+*} "Lifts trivial vector statements to real arith statements"
+
+lemma vec_0[simp]: "vec 0 = 0" by (vector vector_zero_def)
+lemma vec_1[simp]: "vec 1 = 1" by (vector vector_one_def)
+
+
+
+text{* Obvious "component-pushing". *}
+
+lemma vec_component [simp]: "(vec x :: 'a ^ 'n)$i = x"
+  by (vector vec_def)
+
+lemma vector_add_component [simp]:
+  fixes x y :: "'a::{plus} ^ 'n"
+  shows "(x + y)$i = x$i + y$i"
+  by vector
+
+lemma vector_minus_component [simp]:
+  fixes x y :: "'a::{minus} ^ 'n"
+  shows "(x - y)$i = x$i - y$i"
+  by vector
+
+lemma vector_mult_component [simp]:
+  fixes x y :: "'a::{times} ^ 'n"
+  shows "(x * y)$i = x$i * y$i"
+  by vector
+
+lemma vector_smult_component [simp]:
+  fixes y :: "'a::{times} ^ 'n"
+  shows "(c *s y)$i = c * (y$i)"
+  by vector
+
+lemma vector_uminus_component [simp]:
+  fixes x :: "'a::{uminus} ^ 'n"
+  shows "(- x)$i = - (x$i)"
+  by vector
+
+lemma vector_scaleR_component [simp]:
+  fixes x :: "'a::scaleR ^ 'n"
+  shows "(scaleR r x)$i = scaleR r (x$i)"
+  by vector
+
+lemma cond_component: "(if b then x else y)$i = (if b then x$i else y$i)" by vector
+
+lemmas vector_component =
+  vec_component vector_add_component vector_mult_component
+  vector_smult_component vector_minus_component vector_uminus_component
+  vector_scaleR_component cond_component
+
+subsection {* Some frequently useful arithmetic lemmas over vectors. *}
+
+instance "^" :: (semigroup_add,type) semigroup_add
+  apply (intro_classes) by (vector add_assoc)
+
+
+instance "^" :: (monoid_add,type) monoid_add
+  apply (intro_classes) by vector+
+
+instance "^" :: (group_add,type) group_add
+  apply (intro_classes) by (vector algebra_simps)+
+
+instance "^" :: (ab_semigroup_add,type) ab_semigroup_add
+  apply (intro_classes) by (vector add_commute)
+
+instance "^" :: (comm_monoid_add,type) comm_monoid_add
+  apply (intro_classes) by vector
+
+instance "^" :: (ab_group_add,type) ab_group_add
+  apply (intro_classes) by vector+
+
+instance "^" :: (cancel_semigroup_add,type) cancel_semigroup_add
+  apply (intro_classes)
+  by (vector Cart_eq)+
+
+instance "^" :: (cancel_ab_semigroup_add,type) cancel_ab_semigroup_add
+  apply (intro_classes)
+  by (vector Cart_eq)
+
+instance "^" :: (real_vector, type) real_vector
+  by default (vector scaleR_left_distrib scaleR_right_distrib)+
+
+instance "^" :: (semigroup_mult,type) semigroup_mult
+  apply (intro_classes) by (vector mult_assoc)
+
+instance "^" :: (monoid_mult,type) monoid_mult
+  apply (intro_classes) by vector+
+
+instance "^" :: (ab_semigroup_mult,type) ab_semigroup_mult
+  apply (intro_classes) by (vector mult_commute)
+
+instance "^" :: (ab_semigroup_idem_mult,type) ab_semigroup_idem_mult
+  apply (intro_classes) by (vector mult_idem)
+
+instance "^" :: (comm_monoid_mult,type) comm_monoid_mult
+  apply (intro_classes) by vector
+
+fun vector_power :: "('a::{one,times} ^'n) \<Rightarrow> nat \<Rightarrow> 'a^'n" where
+  "vector_power x 0 = 1"
+  | "vector_power x (Suc n) = x * vector_power x n"
+
+instance "^" :: (semiring,type) semiring
+  apply (intro_classes) by (vector ring_simps)+
+
+instance "^" :: (semiring_0,type) semiring_0
+  apply (intro_classes) by (vector ring_simps)+
+instance "^" :: (semiring_1,type) semiring_1
+  apply (intro_classes) by vector
+instance "^" :: (comm_semiring,type) comm_semiring
+  apply (intro_classes) by (vector ring_simps)+
+
+instance "^" :: (comm_semiring_0,type) comm_semiring_0 by (intro_classes)
+instance "^" :: (cancel_comm_monoid_add, type) cancel_comm_monoid_add ..
+instance "^" :: (semiring_0_cancel,type) semiring_0_cancel by (intro_classes)
+instance "^" :: (comm_semiring_0_cancel,type) comm_semiring_0_cancel by (intro_classes)
+instance "^" :: (ring,type) ring by (intro_classes)
+instance "^" :: (semiring_1_cancel,type) semiring_1_cancel by (intro_classes)
+instance "^" :: (comm_semiring_1,type) comm_semiring_1 by (intro_classes)
+
+instance "^" :: (ring_1,type) ring_1 ..
+
+instance "^" :: (real_algebra,type) real_algebra
+  apply intro_classes
+  apply (simp_all add: vector_scaleR_def ring_simps)
+  apply vector
+  apply vector
+  done
+
+instance "^" :: (real_algebra_1,type) real_algebra_1 ..
+
+lemma of_nat_index:
+  "(of_nat n :: 'a::semiring_1 ^'n)$i = of_nat n"
+  apply (induct n)
+  apply vector
+  apply vector
+  done
+lemma zero_index[simp]:
+  "(0 :: 'a::zero ^'n)$i = 0" by vector
+
+lemma one_index[simp]:
+  "(1 :: 'a::one ^'n)$i = 1" by vector
+
+lemma one_plus_of_nat_neq_0: "(1::'a::semiring_char_0) + of_nat n \<noteq> 0"
+proof-
+  have "(1::'a) + of_nat n = 0 \<longleftrightarrow> of_nat 1 + of_nat n = (of_nat 0 :: 'a)" by simp
+  also have "\<dots> \<longleftrightarrow> 1 + n = 0" by (simp only: of_nat_add[symmetric] of_nat_eq_iff)
+  finally show ?thesis by simp
+qed
+
+instance "^" :: (semiring_char_0,type) semiring_char_0
+proof (intro_classes)
+  fix m n ::nat
+  show "(of_nat m :: 'a^'b) = of_nat n \<longleftrightarrow> m = n"
+    by (simp add: Cart_eq of_nat_index)
+qed
+
+instance "^" :: (comm_ring_1,type) comm_ring_1 by intro_classes
+instance "^" :: (ring_char_0,type) ring_char_0 by intro_classes
+
+lemma vector_smult_assoc: "a *s (b *s x) = ((a::'a::semigroup_mult) * b) *s x"
+  by (vector mult_assoc)
+lemma vector_sadd_rdistrib: "((a::'a::semiring) + b) *s x = a *s x + b *s x"
+  by (vector ring_simps)
+lemma vector_add_ldistrib: "(c::'a::semiring) *s (x + y) = c *s x + c *s y"
+  by (vector ring_simps)
+lemma vector_smult_lzero[simp]: "(0::'a::mult_zero) *s x = 0" by vector
+lemma vector_smult_lid[simp]: "(1::'a::monoid_mult) *s x = x" by vector
+lemma vector_ssub_ldistrib: "(c::'a::ring) *s (x - y) = c *s x - c *s y"
+  by (vector ring_simps)
+lemma vector_smult_rneg: "(c::'a::ring) *s -x = -(c *s x)" by vector
+lemma vector_smult_lneg: "- (c::'a::ring) *s x = -(c *s x)" by vector
+lemma vector_sneg_minus1: "-x = (- (1::'a::ring_1)) *s x" by vector
+lemma vector_smult_rzero[simp]: "c *s 0 = (0::'a::mult_zero ^ 'n)" by vector
+lemma vector_sub_rdistrib: "((a::'a::ring) - b) *s x = a *s x - b *s x"
+  by (vector ring_simps)
+
+lemma vec_eq[simp]: "(vec m = vec n) \<longleftrightarrow> (m = n)"
+  by (simp add: Cart_eq)
+
+subsection {* Topological space *}
+
+instantiation "^" :: (topological_space, finite) topological_space
+begin
+
+definition open_vector_def:
+  "open (S :: ('a ^ 'b) set) \<longleftrightarrow>
+    (\<forall>x\<in>S. \<exists>A. (\<forall>i. open (A i) \<and> x$i \<in> A i) \<and>
+      (\<forall>y. (\<forall>i. y$i \<in> A i) \<longrightarrow> y \<in> S))"
+
+instance proof
+  show "open (UNIV :: ('a ^ 'b) set)"
+    unfolding open_vector_def by auto
+next
+  fix S T :: "('a ^ 'b) set"
+  assume "open S" "open T" thus "open (S \<inter> T)"
+    unfolding open_vector_def
+    apply clarify
+    apply (drule (1) bspec)+
+    apply (clarify, rename_tac Sa Ta)
+    apply (rule_tac x="\<lambda>i. Sa i \<inter> Ta i" in exI)
+    apply (simp add: open_Int)
+    done
+next
+  fix K :: "('a ^ 'b) set set"
+  assume "\<forall>S\<in>K. open S" thus "open (\<Union>K)"
+    unfolding open_vector_def
+    apply clarify
+    apply (drule (1) bspec)
+    apply (drule (1) bspec)
+    apply clarify
+    apply (rule_tac x=A in exI)
+    apply fast
+    done
+qed
+
+end
+
+lemma open_vector_box: "\<forall>i. open (S i) \<Longrightarrow> open {x. \<forall>i. x $ i \<in> S i}"
+unfolding open_vector_def by auto
+
+lemma open_vimage_Cart_nth: "open S \<Longrightarrow> open ((\<lambda>x. x $ i) -` S)"
+unfolding open_vector_def
+apply clarify
+apply (rule_tac x="\<lambda>k. if k = i then S else UNIV" in exI, simp)
+done
+
+lemma closed_vimage_Cart_nth: "closed S \<Longrightarrow> closed ((\<lambda>x. x $ i) -` S)"
+unfolding closed_open vimage_Compl [symmetric]
+by (rule open_vimage_Cart_nth)
+
+lemma closed_vector_box: "\<forall>i. closed (S i) \<Longrightarrow> closed {x. \<forall>i. x $ i \<in> S i}"
+proof -
+  have "{x. \<forall>i. x $ i \<in> S i} = (\<Inter>i. (\<lambda>x. x $ i) -` S i)" by auto
+  thus "\<forall>i. closed (S i) \<Longrightarrow> closed {x. \<forall>i. x $ i \<in> S i}"
+    by (simp add: closed_INT closed_vimage_Cart_nth)
+qed
+
+lemma tendsto_Cart_nth [tendsto_intros]:
+  assumes "((\<lambda>x. f x) ---> a) net"
+  shows "((\<lambda>x. f x $ i) ---> a $ i) net"
+proof (rule topological_tendstoI)
+  fix S assume "open S" "a $ i \<in> S"
+  then have "open ((\<lambda>y. y $ i) -` S)" "a \<in> ((\<lambda>y. y $ i) -` S)"
+    by (simp_all add: open_vimage_Cart_nth)
+  with assms have "eventually (\<lambda>x. f x \<in> (\<lambda>y. y $ i) -` S) net"
+    by (rule topological_tendstoD)
+  then show "eventually (\<lambda>x. f x $ i \<in> S) net"
+    by simp
+qed
+
+subsection {* Square root of sum of squares *}
+
+definition
+  "setL2 f A = sqrt (\<Sum>i\<in>A. (f i)\<twosuperior>)"
+
+lemma setL2_cong:
+  "\<lbrakk>A = B; \<And>x. x \<in> B \<Longrightarrow> f x = g x\<rbrakk> \<Longrightarrow> setL2 f A = setL2 g B"
+  unfolding setL2_def by simp
+
+lemma strong_setL2_cong:
+  "\<lbrakk>A = B; \<And>x. x \<in> B =simp=> f x = g x\<rbrakk> \<Longrightarrow> setL2 f A = setL2 g B"
+  unfolding setL2_def simp_implies_def by simp
+
+lemma setL2_infinite [simp]: "\<not> finite A \<Longrightarrow> setL2 f A = 0"
+  unfolding setL2_def by simp
+
+lemma setL2_empty [simp]: "setL2 f {} = 0"
+  unfolding setL2_def by simp
+
+lemma setL2_insert [simp]:
+  "\<lbrakk>finite F; a \<notin> F\<rbrakk> \<Longrightarrow>
+    setL2 f (insert a F) = sqrt ((f a)\<twosuperior> + (setL2 f F)\<twosuperior>)"
+  unfolding setL2_def by (simp add: setsum_nonneg)
+
+lemma setL2_nonneg [simp]: "0 \<le> setL2 f A"
+  unfolding setL2_def by (simp add: setsum_nonneg)
+
+lemma setL2_0': "\<forall>a\<in>A. f a = 0 \<Longrightarrow> setL2 f A = 0"
+  unfolding setL2_def by simp
+
+lemma setL2_constant: "setL2 (\<lambda>x. y) A = sqrt (of_nat (card A)) * \<bar>y\<bar>"
+  unfolding setL2_def by (simp add: real_sqrt_mult)
+
+lemma setL2_mono:
+  assumes "\<And>i. i \<in> K \<Longrightarrow> f i \<le> g i"
+  assumes "\<And>i. i \<in> K \<Longrightarrow> 0 \<le> f i"
+  shows "setL2 f K \<le> setL2 g K"
+  unfolding setL2_def
+  by (simp add: setsum_nonneg setsum_mono power_mono prems)
+
+lemma setL2_strict_mono:
+  assumes "finite K" and "K \<noteq> {}"
+  assumes "\<And>i. i \<in> K \<Longrightarrow> f i < g i"
+  assumes "\<And>i. i \<in> K \<Longrightarrow> 0 \<le> f i"
+  shows "setL2 f K < setL2 g K"
+  unfolding setL2_def
+  by (simp add: setsum_strict_mono power_strict_mono assms)
+
+lemma setL2_right_distrib:
+  "0 \<le> r \<Longrightarrow> r * setL2 f A = setL2 (\<lambda>x. r * f x) A"
+  unfolding setL2_def
+  apply (simp add: power_mult_distrib)
+  apply (simp add: setsum_right_distrib [symmetric])
+  apply (simp add: real_sqrt_mult setsum_nonneg)
+  done
+
+lemma setL2_left_distrib:
+  "0 \<le> r \<Longrightarrow> setL2 f A * r = setL2 (\<lambda>x. f x * r) A"
+  unfolding setL2_def
+  apply (simp add: power_mult_distrib)
+  apply (simp add: setsum_left_distrib [symmetric])
+  apply (simp add: real_sqrt_mult setsum_nonneg)
+  done
+
+lemma setsum_nonneg_eq_0_iff:
+  fixes f :: "'a \<Rightarrow> 'b::pordered_ab_group_add"
+  shows "\<lbrakk>finite A; \<forall>x\<in>A. 0 \<le> f x\<rbrakk> \<Longrightarrow> setsum f A = 0 \<longleftrightarrow> (\<forall>x\<in>A. f x = 0)"
+  apply (induct set: finite, simp)
+  apply (simp add: add_nonneg_eq_0_iff setsum_nonneg)
+  done
+
+lemma setL2_eq_0_iff: "finite A \<Longrightarrow> setL2 f A = 0 \<longleftrightarrow> (\<forall>x\<in>A. f x = 0)"
+  unfolding setL2_def
+  by (simp add: setsum_nonneg setsum_nonneg_eq_0_iff)
+
+lemma setL2_triangle_ineq:
+  shows "setL2 (\<lambda>i. f i + g i) A \<le> setL2 f A + setL2 g A"
+proof (cases "finite A")
+  case False
+  thus ?thesis by simp
+next
+  case True
+  thus ?thesis
+  proof (induct set: finite)
+    case empty
+    show ?case by simp
+  next
+    case (insert x F)
+    hence "sqrt ((f x + g x)\<twosuperior> + (setL2 (\<lambda>i. f i + g i) F)\<twosuperior>) \<le>
+           sqrt ((f x + g x)\<twosuperior> + (setL2 f F + setL2 g F)\<twosuperior>)"
+      by (intro real_sqrt_le_mono add_left_mono power_mono insert
+                setL2_nonneg add_increasing zero_le_power2)
+    also have
+      "\<dots> \<le> sqrt ((f x)\<twosuperior> + (setL2 f F)\<twosuperior>) + sqrt ((g x)\<twosuperior> + (setL2 g F)\<twosuperior>)"
+      by (rule real_sqrt_sum_squares_triangle_ineq)
+    finally show ?case
+      using insert by simp
+  qed
+qed
+
+lemma sqrt_sum_squares_le_sum:
+  "\<lbrakk>0 \<le> x; 0 \<le> y\<rbrakk> \<Longrightarrow> sqrt (x\<twosuperior> + y\<twosuperior>) \<le> x + y"
+  apply (rule power2_le_imp_le)
+  apply (simp add: power2_sum)
+  apply (simp add: mult_nonneg_nonneg)
+  apply (simp add: add_nonneg_nonneg)
+  done
+
+lemma setL2_le_setsum [rule_format]:
+  "(\<forall>i\<in>A. 0 \<le> f i) \<longrightarrow> setL2 f A \<le> setsum f A"
+  apply (cases "finite A")
+  apply (induct set: finite)
+  apply simp
+  apply clarsimp
+  apply (erule order_trans [OF sqrt_sum_squares_le_sum])
+  apply simp
+  apply simp
+  apply simp
+  done
+
+lemma sqrt_sum_squares_le_sum_abs: "sqrt (x\<twosuperior> + y\<twosuperior>) \<le> \<bar>x\<bar> + \<bar>y\<bar>"
+  apply (rule power2_le_imp_le)
+  apply (simp add: power2_sum)
+  apply (simp add: mult_nonneg_nonneg)
+  apply (simp add: add_nonneg_nonneg)
+  done
+
+lemma setL2_le_setsum_abs: "setL2 f A \<le> (\<Sum>i\<in>A. \<bar>f i\<bar>)"
+  apply (cases "finite A")
+  apply (induct set: finite)
+  apply simp
+  apply simp
+  apply (rule order_trans [OF sqrt_sum_squares_le_sum_abs])
+  apply simp
+  apply simp
+  done
+
+lemma setL2_mult_ineq_lemma:
+  fixes a b c d :: real
+  shows "2 * (a * c) * (b * d) \<le> a\<twosuperior> * d\<twosuperior> + b\<twosuperior> * c\<twosuperior>"
+proof -
+  have "0 \<le> (a * d - b * c)\<twosuperior>" by simp
+  also have "\<dots> = a\<twosuperior> * d\<twosuperior> + b\<twosuperior> * c\<twosuperior> - 2 * (a * d) * (b * c)"
+    by (simp only: power2_diff power_mult_distrib)
+  also have "\<dots> = a\<twosuperior> * d\<twosuperior> + b\<twosuperior> * c\<twosuperior> - 2 * (a * c) * (b * d)"
+    by simp
+  finally show "2 * (a * c) * (b * d) \<le> a\<twosuperior> * d\<twosuperior> + b\<twosuperior> * c\<twosuperior>"
+    by simp
+qed
+
+lemma setL2_mult_ineq: "(\<Sum>i\<in>A. \<bar>f i\<bar> * \<bar>g i\<bar>) \<le> setL2 f A * setL2 g A"
+  apply (cases "finite A")
+  apply (induct set: finite)
+  apply simp
+  apply (rule power2_le_imp_le, simp)
+  apply (rule order_trans)
+  apply (rule power_mono)
+  apply (erule add_left_mono)
+  apply (simp add: add_nonneg_nonneg mult_nonneg_nonneg setsum_nonneg)
+  apply (simp add: power2_sum)
+  apply (simp add: power_mult_distrib)
+  apply (simp add: right_distrib left_distrib)
+  apply (rule ord_le_eq_trans)
+  apply (rule setL2_mult_ineq_lemma)
+  apply simp
+  apply (intro mult_nonneg_nonneg setL2_nonneg)
+  apply simp
+  done
+
+lemma member_le_setL2: "\<lbrakk>finite A; i \<in> A\<rbrakk> \<Longrightarrow> f i \<le> setL2 f A"
+  apply (rule_tac s="insert i (A - {i})" and t="A" in subst)
+  apply fast
+  apply (subst setL2_insert)
+  apply simp
+  apply simp
+  apply simp
+  done
+
+subsection {* Metric *}
+
+(* TODO: move somewhere else *)
+lemma finite_choice: "finite A \<Longrightarrow> \<forall>x\<in>A. \<exists>y. P x y \<Longrightarrow> \<exists>f. \<forall>x\<in>A. P x (f x)"
+apply (induct set: finite, simp_all)
+apply (clarify, rename_tac y)
+apply (rule_tac x="f(x:=y)" in exI, simp)
+done
+
+instantiation "^" :: (metric_space, finite) metric_space
+begin
+
+definition dist_vector_def:
+  "dist (x::'a^'b) (y::'a^'b) = setL2 (\<lambda>i. dist (x$i) (y$i)) UNIV"
+
+lemma dist_nth_le: "dist (x $ i) (y $ i) \<le> dist x y"
+unfolding dist_vector_def
+by (rule member_le_setL2) simp_all
+
+instance proof
+  fix x y :: "'a ^ 'b"
+  show "dist x y = 0 \<longleftrightarrow> x = y"
+    unfolding dist_vector_def
+    by (simp add: setL2_eq_0_iff Cart_eq)
+next
+  fix x y z :: "'a ^ 'b"
+  show "dist x y \<le> dist x z + dist y z"
+    unfolding dist_vector_def
+    apply (rule order_trans [OF _ setL2_triangle_ineq])
+    apply (simp add: setL2_mono dist_triangle2)
+    done
+next
+  (* FIXME: long proof! *)
+  fix S :: "('a ^ 'b) set"
+  show "open S \<longleftrightarrow> (\<forall>x\<in>S. \<exists>e>0. \<forall>y. dist y x < e \<longrightarrow> y \<in> S)"
+    unfolding open_vector_def open_dist
+    apply safe
+     apply (drule (1) bspec)
+     apply clarify
+     apply (subgoal_tac "\<exists>e>0. \<forall>i y. dist y (x$i) < e \<longrightarrow> y \<in> A i")
+      apply clarify
+      apply (rule_tac x=e in exI, clarify)
+      apply (drule spec, erule mp, clarify)
+      apply (drule spec, drule spec, erule mp)
+      apply (erule le_less_trans [OF dist_nth_le])
+     apply (subgoal_tac "\<forall>i\<in>UNIV. \<exists>e>0. \<forall>y. dist y (x$i) < e \<longrightarrow> y \<in> A i")
+      apply (drule finite_choice [OF finite], clarify)
+      apply (rule_tac x="Min (range f)" in exI, simp)
+     apply clarify
+     apply (drule_tac x=i in spec, clarify)
+     apply (erule (1) bspec)
+    apply (drule (1) bspec, clarify)
+    apply (subgoal_tac "\<exists>r. (\<forall>i::'b. 0 < r i) \<and> e = setL2 r UNIV")
+     apply clarify
+     apply (rule_tac x="\<lambda>i. {y. dist y (x$i) < r i}" in exI)
+     apply (rule conjI)
+      apply clarify
+      apply (rule conjI)
+       apply (clarify, rename_tac y)
+       apply (rule_tac x="r i - dist y (x$i)" in exI, rule conjI, simp)
+       apply clarify
+       apply (simp only: less_diff_eq)
+       apply (erule le_less_trans [OF dist_triangle])
+      apply simp
+     apply clarify
+     apply (drule spec, erule mp)
+     apply (simp add: dist_vector_def setL2_strict_mono)
+    apply (rule_tac x="\<lambda>i. e / sqrt (of_nat CARD('b))" in exI)
+    apply (simp add: divide_pos_pos setL2_constant)
+    done
+qed
+
+end
+
+lemma LIMSEQ_Cart_nth:
+  "(X ----> a) \<Longrightarrow> (\<lambda>n. X n $ i) ----> a $ i"
+unfolding LIMSEQ_conv_tendsto by (rule tendsto_Cart_nth)
+
+lemma LIM_Cart_nth:
+  "(f -- x --> y) \<Longrightarrow> (\<lambda>x. f x $ i) -- x --> y $ i"
+unfolding LIM_conv_tendsto by (rule tendsto_Cart_nth)
+
+lemma Cauchy_Cart_nth:
+  "Cauchy (\<lambda>n. X n) \<Longrightarrow> Cauchy (\<lambda>n. X n $ i)"
+unfolding Cauchy_def by (fast intro: le_less_trans [OF dist_nth_le])
+
+lemma LIMSEQ_vector:
+  fixes X :: "nat \<Rightarrow> 'a::metric_space ^ 'n::finite"
+  assumes X: "\<And>i. (\<lambda>n. X n $ i) ----> (a $ i)"
+  shows "X ----> a"
+proof (rule metric_LIMSEQ_I)
+  fix r :: real assume "0 < r"
+  then have "0 < r / of_nat CARD('n)" (is "0 < ?s")
+    by (simp add: divide_pos_pos)
+  def N \<equiv> "\<lambda>i. LEAST N. \<forall>n\<ge>N. dist (X n $ i) (a $ i) < ?s"
+  def M \<equiv> "Max (range N)"
+  have "\<And>i. \<exists>N. \<forall>n\<ge>N. dist (X n $ i) (a $ i) < ?s"
+    using X `0 < ?s` by (rule metric_LIMSEQ_D)
+  hence "\<And>i. \<forall>n\<ge>N i. dist (X n $ i) (a $ i) < ?s"
+    unfolding N_def by (rule LeastI_ex)
+  hence M: "\<And>i. \<forall>n\<ge>M. dist (X n $ i) (a $ i) < ?s"
+    unfolding M_def by simp
+  {
+    fix n :: nat assume "M \<le> n"
+    have "dist (X n) a = setL2 (\<lambda>i. dist (X n $ i) (a $ i)) UNIV"
+      unfolding dist_vector_def ..
+    also have "\<dots> \<le> setsum (\<lambda>i. dist (X n $ i) (a $ i)) UNIV"
+      by (rule setL2_le_setsum [OF zero_le_dist])
+    also have "\<dots> < setsum (\<lambda>i::'n. ?s) UNIV"
+      by (rule setsum_strict_mono, simp_all add: M `M \<le> n`)
+    also have "\<dots> = r"
+      by simp
+    finally have "dist (X n) a < r" .
+  }
+  hence "\<forall>n\<ge>M. dist (X n) a < r"
+    by simp
+  then show "\<exists>M. \<forall>n\<ge>M. dist (X n) a < r" ..
+qed
+
+lemma Cauchy_vector:
+  fixes X :: "nat \<Rightarrow> 'a::metric_space ^ 'n::finite"
+  assumes X: "\<And>i. Cauchy (\<lambda>n. X n $ i)"
+  shows "Cauchy (\<lambda>n. X n)"
+proof (rule metric_CauchyI)
+  fix r :: real assume "0 < r"
+  then have "0 < r / of_nat CARD('n)" (is "0 < ?s")
+    by (simp add: divide_pos_pos)
+  def N \<equiv> "\<lambda>i. LEAST N. \<forall>m\<ge>N. \<forall>n\<ge>N. dist (X m $ i) (X n $ i) < ?s"
+  def M \<equiv> "Max (range N)"
+  have "\<And>i. \<exists>N. \<forall>m\<ge>N. \<forall>n\<ge>N. dist (X m $ i) (X n $ i) < ?s"
+    using X `0 < ?s` by (rule metric_CauchyD)
+  hence "\<And>i. \<forall>m\<ge>N i. \<forall>n\<ge>N i. dist (X m $ i) (X n $ i) < ?s"
+    unfolding N_def by (rule LeastI_ex)
+  hence M: "\<And>i. \<forall>m\<ge>M. \<forall>n\<ge>M. dist (X m $ i) (X n $ i) < ?s"
+    unfolding M_def by simp
+  {
+    fix m n :: nat
+    assume "M \<le> m" "M \<le> n"
+    have "dist (X m) (X n) = setL2 (\<lambda>i. dist (X m $ i) (X n $ i)) UNIV"
+      unfolding dist_vector_def ..
+    also have "\<dots> \<le> setsum (\<lambda>i. dist (X m $ i) (X n $ i)) UNIV"
+      by (rule setL2_le_setsum [OF zero_le_dist])
+    also have "\<dots> < setsum (\<lambda>i::'n. ?s) UNIV"
+      by (rule setsum_strict_mono, simp_all add: M `M \<le> m` `M \<le> n`)
+    also have "\<dots> = r"
+      by simp
+    finally have "dist (X m) (X n) < r" .
+  }
+  hence "\<forall>m\<ge>M. \<forall>n\<ge>M. dist (X m) (X n) < r"
+    by simp
+  then show "\<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. dist (X m) (X n) < r" ..
+qed
+
+instance "^" :: (complete_space, finite) complete_space
+proof
+  fix X :: "nat \<Rightarrow> 'a ^ 'b" assume "Cauchy X"
+  have "\<And>i. (\<lambda>n. X n $ i) ----> lim (\<lambda>n. X n $ i)"
+    using Cauchy_Cart_nth [OF `Cauchy X`]
+    by (simp add: Cauchy_convergent_iff convergent_LIMSEQ_iff)
+  hence "X ----> Cart_lambda (\<lambda>i. lim (\<lambda>n. X n $ i))"
+    by (simp add: LIMSEQ_vector)
+  then show "convergent X"
+    by (rule convergentI)
+qed
+
+subsection {* Norms *}
+
+instantiation "^" :: (real_normed_vector, finite) real_normed_vector
+begin
+
+definition norm_vector_def:
+  "norm (x::'a^'b) = setL2 (\<lambda>i. norm (x$i)) UNIV"
+
+definition vector_sgn_def:
+  "sgn (x::'a^'b) = scaleR (inverse (norm x)) x"
+
+instance proof
+  fix a :: real and x y :: "'a ^ 'b"
+  show "0 \<le> norm x"
+    unfolding norm_vector_def
+    by (rule setL2_nonneg)
+  show "norm x = 0 \<longleftrightarrow> x = 0"
+    unfolding norm_vector_def
+    by (simp add: setL2_eq_0_iff Cart_eq)
+  show "norm (x + y) \<le> norm x + norm y"
+    unfolding norm_vector_def
+    apply (rule order_trans [OF _ setL2_triangle_ineq])
+    apply (simp add: setL2_mono norm_triangle_ineq)
+    done
+  show "norm (scaleR a x) = \<bar>a\<bar> * norm x"
+    unfolding norm_vector_def
+    by (simp add: setL2_right_distrib)
+  show "sgn x = scaleR (inverse (norm x)) x"
+    by (rule vector_sgn_def)
+  show "dist x y = norm (x - y)"
+    unfolding dist_vector_def norm_vector_def
+    by (simp add: dist_norm)
+qed
+
+end
+
+lemma norm_nth_le: "norm (x $ i) \<le> norm x"
+unfolding norm_vector_def
+by (rule member_le_setL2) simp_all
+
+interpretation Cart_nth: bounded_linear "\<lambda>x. x $ i"
+apply default
+apply (rule vector_add_component)
+apply (rule vector_scaleR_component)
+apply (rule_tac x="1" in exI, simp add: norm_nth_le)
+done
+
+instance "^" :: (banach, finite) banach ..
+
+subsection {* Inner products *}
+
+instantiation "^" :: (real_inner, finite) real_inner
+begin
+
+definition inner_vector_def:
+  "inner x y = setsum (\<lambda>i. inner (x$i) (y$i)) UNIV"
+
+instance proof
+  fix r :: real and x y z :: "'a ^ 'b"
+  show "inner x y = inner y x"
+    unfolding inner_vector_def
+    by (simp add: inner_commute)
+  show "inner (x + y) z = inner x z + inner y z"
+    unfolding inner_vector_def
+    by (simp add: inner_add_left setsum_addf)
+  show "inner (scaleR r x) y = r * inner x y"
+    unfolding inner_vector_def
+    by (simp add: setsum_right_distrib)
+  show "0 \<le> inner x x"
+    unfolding inner_vector_def
+    by (simp add: setsum_nonneg)
+  show "inner x x = 0 \<longleftrightarrow> x = 0"
+    unfolding inner_vector_def
+    by (simp add: Cart_eq setsum_nonneg_eq_0_iff)
+  show "norm x = sqrt (inner x x)"
+    unfolding inner_vector_def norm_vector_def setL2_def
+    by (simp add: power2_norm_eq_inner)
+qed
+
+end
+
+subsection{* Properties of the dot product.  *}
+
+lemma dot_sym: "(x::'a:: {comm_monoid_add, ab_semigroup_mult} ^ 'n) \<bullet> y = y \<bullet> x"
+  by (vector mult_commute)
+lemma dot_ladd: "((x::'a::ring ^ 'n) + y) \<bullet> z = (x \<bullet> z) + (y \<bullet> z)"
+  by (vector ring_simps)
+lemma dot_radd: "x \<bullet> (y + (z::'a::ring ^ 'n)) = (x \<bullet> y) + (x \<bullet> z)"
+  by (vector ring_simps)
+lemma dot_lsub: "((x::'a::ring ^ 'n) - y) \<bullet> z = (x \<bullet> z) - (y \<bullet> z)"
+  by (vector ring_simps)
+lemma dot_rsub: "(x::'a::ring ^ 'n) \<bullet> (y - z) = (x \<bullet> y) - (x \<bullet> z)"
+  by (vector ring_simps)
+lemma dot_lmult: "(c *s x) \<bullet> y = (c::'a::ring) * (x \<bullet> y)" by (vector ring_simps)
+lemma dot_rmult: "x \<bullet> (c *s y) = (c::'a::comm_ring) * (x \<bullet> y)" by (vector ring_simps)
+lemma dot_lneg: "(-x) \<bullet> (y::'a::ring ^ 'n) = -(x \<bullet> y)" by vector
+lemma dot_rneg: "(x::'a::ring ^ 'n) \<bullet> (-y) = -(x \<bullet> y)" by vector
+lemma dot_lzero[simp]: "0 \<bullet> x = (0::'a::{comm_monoid_add, mult_zero})" by vector
+lemma dot_rzero[simp]: "x \<bullet> 0 = (0::'a::{comm_monoid_add, mult_zero})" by vector
+lemma dot_pos_le[simp]: "(0::'a\<Colon>ordered_ring_strict) <= x \<bullet> x"
+  by (simp add: dot_def setsum_nonneg)
+
+lemma setsum_squares_eq_0_iff: assumes fS: "finite F" and fp: "\<forall>x \<in> F. f x \<ge> (0 ::'a::pordered_ab_group_add)" shows "setsum f F = 0 \<longleftrightarrow> (ALL x:F. f x = 0)"
+using fS fp setsum_nonneg[OF fp]
+proof (induct set: finite)
+  case empty thus ?case by simp
+next
+  case (insert x F)
+  from insert.prems have Fx: "f x \<ge> 0" and Fp: "\<forall> a \<in> F. f a \<ge> 0" by simp_all
+  from insert.hyps Fp setsum_nonneg[OF Fp]
+  have h: "setsum f F = 0 \<longleftrightarrow> (\<forall>a \<in>F. f a = 0)" by metis
+  from add_nonneg_eq_0_iff[OF Fx  setsum_nonneg[OF Fp]] insert.hyps(1,2)
+  show ?case by (simp add: h)
+qed
+
+lemma dot_eq_0: "x \<bullet> x = 0 \<longleftrightarrow> (x::'a::{ordered_ring_strict,ring_no_zero_divisors} ^ 'n::finite) = 0"
+  by (simp add: dot_def setsum_squares_eq_0_iff Cart_eq)
+
+lemma dot_pos_lt[simp]: "(0 < x \<bullet> x) \<longleftrightarrow> (x::'a::{ordered_ring_strict,ring_no_zero_divisors} ^ 'n::finite) \<noteq> 0" using dot_eq_0[of x] dot_pos_le[of x]
+  by (auto simp add: le_less)
+
+subsection{* The collapse of the general concepts to dimension one. *}
+
+lemma vector_one: "(x::'a ^1) = (\<chi> i. (x$1))"
+  by (simp add: Cart_eq forall_1)
+
+lemma forall_one: "(\<forall>(x::'a ^1). P x) \<longleftrightarrow> (\<forall>x. P(\<chi> i. x))"
+  apply auto
+  apply (erule_tac x= "x$1" in allE)
+  apply (simp only: vector_one[symmetric])
+  done
+
+lemma norm_vector_1: "norm (x :: _^1) = norm (x$1)"
+  by (simp add: norm_vector_def UNIV_1)
+
+lemma norm_real: "norm(x::real ^ 1) = abs(x$1)"
+  by (simp add: norm_vector_1)
+
+lemma dist_real: "dist(x::real ^ 1) y = abs((x$1) - (y$1))"
+  by (auto simp add: norm_real dist_norm)
+
+subsection {* A connectedness or intermediate value lemma with several applications. *}
+
+lemma connected_real_lemma:
+  fixes f :: "real \<Rightarrow> 'a::metric_space"
+  assumes ab: "a \<le> b" and fa: "f a \<in> e1" and fb: "f b \<in> e2"
+  and dst: "\<And>e x. a <= x \<Longrightarrow> x <= b \<Longrightarrow> 0 < e ==> \<exists>d > 0. \<forall>y. abs(y - x) < d \<longrightarrow> dist(f y) (f x) < e"
+  and e1: "\<forall>y \<in> e1. \<exists>e > 0. \<forall>y'. dist y' y < e \<longrightarrow> y' \<in> e1"
+  and e2: "\<forall>y \<in> e2. \<exists>e > 0. \<forall>y'. dist y' y < e \<longrightarrow> y' \<in> e2"
+  and e12: "~(\<exists>x \<ge> a. x <= b \<and> f x \<in> e1 \<and> f x \<in> e2)"
+  shows "\<exists>x \<ge> a. x <= b \<and> f x \<notin> e1 \<and> f x \<notin> e2" (is "\<exists> x. ?P x")
+proof-
+  let ?S = "{c. \<forall>x \<ge> a. x <= c \<longrightarrow> f x \<in> e1}"
+  have Se: " \<exists>x. x \<in> ?S" apply (rule exI[where x=a]) by (auto simp add: fa)
+  have Sub: "\<exists>y. isUb UNIV ?S y"
+    apply (rule exI[where x= b])
+    using ab fb e12 by (auto simp add: isUb_def setle_def)
+  from reals_complete[OF Se Sub] obtain l where
+    l: "isLub UNIV ?S l"by blast
+  have alb: "a \<le> l" "l \<le> b" using l ab fa fb e12
+    apply (auto simp add: isLub_def leastP_def isUb_def setle_def setge_def)
+    by (metis linorder_linear)
+  have ale1: "\<forall>z \<ge> a. z < l \<longrightarrow> f z \<in> e1" using l
+    apply (auto simp add: isLub_def leastP_def isUb_def setle_def setge_def)
+    by (metis linorder_linear not_le)
+    have th1: "\<And>z x e d :: real. z <= x + e \<Longrightarrow> e < d ==> z < x \<or> abs(z - x) < d" by arith
+    have th2: "\<And>e x:: real. 0 < e ==> ~(x + e <= x)" by arith
+    have th3: "\<And>d::real. d > 0 \<Longrightarrow> \<exists>e > 0. e < d" by dlo
+    {assume le2: "f l \<in> e2"
+      from le2 fa fb e12 alb have la: "l \<noteq> a" by metis
+      hence lap: "l - a > 0" using alb by arith
+      from e2[rule_format, OF le2] obtain e where
+        e: "e > 0" "\<forall>y. dist y (f l) < e \<longrightarrow> y \<in> e2" by metis
+      from dst[OF alb e(1)] obtain d where
+        d: "d > 0" "\<forall>y. \<bar>y - l\<bar> < d \<longrightarrow> dist (f y) (f l) < e" by metis
+      have "\<exists>d'. d' < d \<and> d' >0 \<and> l - d' > a" using lap d(1)
+        apply ferrack by arith
+      then obtain d' where d': "d' > 0" "d' < d" "l - d' > a" by metis
+      from d e have th0: "\<forall>y. \<bar>y - l\<bar> < d \<longrightarrow> f y \<in> e2" by metis
+      from th0[rule_format, of "l - d'"] d' have "f (l - d') \<in> e2" by auto
+      moreover
+      have "f (l - d') \<in> e1" using ale1[rule_format, of "l -d'"] d' by auto
+      ultimately have False using e12 alb d' by auto}
+    moreover
+    {assume le1: "f l \<in> e1"
+    from le1 fa fb e12 alb have lb: "l \<noteq> b" by metis
+      hence blp: "b - l > 0" using alb by arith
+      from e1[rule_format, OF le1] obtain e where
+        e: "e > 0" "\<forall>y. dist y (f l) < e \<longrightarrow> y \<in> e1" by metis
+      from dst[OF alb e(1)] obtain d where
+        d: "d > 0" "\<forall>y. \<bar>y - l\<bar> < d \<longrightarrow> dist (f y) (f l) < e" by metis
+      have "\<exists>d'. d' < d \<and> d' >0" using d(1) by dlo
+      then obtain d' where d': "d' > 0" "d' < d" by metis
+      from d e have th0: "\<forall>y. \<bar>y - l\<bar> < d \<longrightarrow> f y \<in> e1" by auto
+      hence "\<forall>y. l \<le> y \<and> y \<le> l + d' \<longrightarrow> f y \<in> e1" using d' by auto
+      with ale1 have "\<forall>y. a \<le> y \<and> y \<le> l + d' \<longrightarrow> f y \<in> e1" by auto
+      with l d' have False
+        by (auto simp add: isLub_def isUb_def setle_def setge_def leastP_def) }
+    ultimately show ?thesis using alb by metis
+qed
+
+text{* One immediately useful corollary is the existence of square roots! --- Should help to get rid of all the development of square-root for reals as a special case @{typ "real^1"} *}
+
+lemma square_bound_lemma: "(x::real) < (1 + x) * (1 + x)"
+proof-
+  have "(x + 1/2)^2 + 3/4 > 0" using zero_le_power2[of "x+1/2"] by arith
+  thus ?thesis by (simp add: ring_simps power2_eq_square)
+qed
+
+lemma square_continuous: "0 < (e::real) ==> \<exists>d. 0 < d \<and> (\<forall>y. abs(y - x) < d \<longrightarrow> abs(y * y - x * x) < e)"
+  using isCont_power[OF isCont_ident, of 2, unfolded isCont_def LIM_eq, rule_format, of e x] apply (auto simp add: power2_eq_square)
+  apply (rule_tac x="s" in exI)
+  apply auto
+  apply (erule_tac x=y in allE)
+  apply auto
+  done
+
+lemma real_le_lsqrt: "0 <= x \<Longrightarrow> 0 <= y \<Longrightarrow> x <= y^2 ==> sqrt x <= y"
+  using real_sqrt_le_iff[of x "y^2"] by simp
+
+lemma real_le_rsqrt: "x^2 \<le> y \<Longrightarrow> x \<le> sqrt y"
+  using real_sqrt_le_mono[of "x^2" y] by simp
+
+lemma real_less_rsqrt: "x^2 < y \<Longrightarrow> x < sqrt y"
+  using real_sqrt_less_mono[of "x^2" y] by simp
+
+lemma sqrt_even_pow2: assumes n: "even n"
+  shows "sqrt(2 ^ n) = 2 ^ (n div 2)"
+proof-
+  from n obtain m where m: "n = 2*m" unfolding even_nat_equiv_def2
+    by (auto simp add: nat_number)
+  from m  have "sqrt(2 ^ n) = sqrt ((2 ^ m) ^ 2)"
+    by (simp only: power_mult[symmetric] mult_commute)
+  then show ?thesis  using m by simp
+qed
+
+lemma real_div_sqrt: "0 <= x ==> x / sqrt(x) = sqrt(x)"
+  apply (cases "x = 0", simp_all)
+  using sqrt_divide_self_eq[of x]
+  apply (simp add: inverse_eq_divide real_sqrt_ge_0_iff field_simps)
+  done
+
+text{* Hence derive more interesting properties of the norm. *}
+
+text {*
+  This type-specific version is only here
+  to make @{text normarith.ML} happy.
+*}
+lemma norm_0: "norm (0::real ^ _) = 0"
+  by (rule norm_zero)
+
+lemma norm_mul[simp]: "norm(a *s x) = abs(a) * norm x"
+  by (simp add: norm_vector_def vector_component setL2_right_distrib
+           abs_mult cong: strong_setL2_cong)
+lemma norm_eq_0_dot: "(norm x = 0) \<longleftrightarrow> (x \<bullet> x = (0::real))"
+  by (simp add: norm_vector_def dot_def setL2_def power2_eq_square)
+lemma real_vector_norm_def: "norm x = sqrt (x \<bullet> x)"
+  by (simp add: norm_vector_def setL2_def dot_def power2_eq_square)
+lemma norm_pow_2: "norm x ^ 2 = x \<bullet> x"
+  by (simp add: real_vector_norm_def)
+lemma norm_eq_0_imp: "norm x = 0 ==> x = (0::real ^'n::finite)" by (metis norm_eq_zero)
+lemma vector_mul_eq_0[simp]: "(a *s x = 0) \<longleftrightarrow> a = (0::'a::idom) \<or> x = 0"
+  by vector
+lemma vector_mul_lcancel[simp]: "a *s x = a *s y \<longleftrightarrow> a = (0::real) \<or> x = y"
+  by (metis eq_iff_diff_eq_0 vector_mul_eq_0 vector_ssub_ldistrib)
+lemma vector_mul_rcancel[simp]: "a *s x = b *s x \<longleftrightarrow> (a::real) = b \<or> x = 0"
+  by (metis eq_iff_diff_eq_0 vector_mul_eq_0 vector_sub_rdistrib)
+lemma vector_mul_lcancel_imp: "a \<noteq> (0::real) ==>  a *s x = a *s y ==> (x = y)"
+  by (metis vector_mul_lcancel)
+lemma vector_mul_rcancel_imp: "x \<noteq> 0 \<Longrightarrow> (a::real) *s x = b *s x ==> a = b"
+  by (metis vector_mul_rcancel)
+lemma norm_cauchy_schwarz:
+  fixes x y :: "real ^ 'n::finite"
+  shows "x \<bullet> y <= norm x * norm y"
+proof-
+  {assume "norm x = 0"
+    hence ?thesis by (simp add: dot_lzero dot_rzero)}
+  moreover
+  {assume "norm y = 0"
+    hence ?thesis by (simp add: dot_lzero dot_rzero)}
+  moreover
+  {assume h: "norm x \<noteq> 0" "norm y \<noteq> 0"
+    let ?z = "norm y *s x - norm x *s y"
+    from h have p: "norm x * norm y > 0" by (metis norm_ge_zero le_less zero_compare_simps)
+    from dot_pos_le[of ?z]
+    have "(norm x * norm y) * (x \<bullet> y) \<le> norm x ^2 * norm y ^2"
+      apply (simp add: dot_rsub dot_lsub dot_lmult dot_rmult ring_simps)
+      by (simp add: norm_pow_2[symmetric] power2_eq_square dot_sym)
+    hence "x\<bullet>y \<le> (norm x ^2 * norm y ^2) / (norm x * norm y)" using p
+      by (simp add: field_simps)
+    hence ?thesis using h by (simp add: power2_eq_square)}
+  ultimately show ?thesis by metis
+qed
+
+lemma norm_cauchy_schwarz_abs:
+  fixes x y :: "real ^ 'n::finite"
+  shows "\<bar>x \<bullet> y\<bar> \<le> norm x * norm y"
+  using norm_cauchy_schwarz[of x y] norm_cauchy_schwarz[of x "-y"]
+  by (simp add: real_abs_def dot_rneg)
+
+lemma norm_triangle_sub:
+  fixes x y :: "'a::real_normed_vector"
+  shows "norm x \<le> norm y  + norm (x - y)"
+  using norm_triangle_ineq[of "y" "x - y"] by (simp add: ring_simps)
+
+lemma norm_triangle_le: "norm(x::real ^'n::finite) + norm y <= e ==> norm(x + y) <= e"
+  by (metis order_trans norm_triangle_ineq)
+lemma norm_triangle_lt: "norm(x::real ^'n::finite) + norm(y) < e ==> norm(x + y) < e"
+  by (metis basic_trans_rules(21) norm_triangle_ineq)
+
+lemma component_le_norm: "\<bar>x$i\<bar> <= norm (x::real ^ 'n::finite)"
+  apply (simp add: norm_vector_def)
+  apply (rule member_le_setL2, simp_all)
+  done
+
+lemma norm_bound_component_le: "norm(x::real ^ 'n::finite) <= e
+                ==> \<bar>x$i\<bar> <= e"
+  by (metis component_le_norm order_trans)
+
+lemma norm_bound_component_lt: "norm(x::real ^ 'n::finite) < e
+                ==> \<bar>x$i\<bar> < e"
+  by (metis component_le_norm basic_trans_rules(21))
+
+lemma norm_le_l1: "norm (x:: real ^'n::finite) <= setsum(\<lambda>i. \<bar>x$i\<bar>) UNIV"
+  by (simp add: norm_vector_def setL2_le_setsum)
+
+lemma real_abs_norm: "\<bar>norm x\<bar> = norm (x :: real ^ _)"
+  by (rule abs_norm_cancel)
+lemma real_abs_sub_norm: "\<bar>norm(x::real ^'n::finite) - norm y\<bar> <= norm(x - y)"
+  by (rule norm_triangle_ineq3)
+lemma norm_le: "norm(x::real ^ _) <= norm(y) \<longleftrightarrow> x \<bullet> x <= y \<bullet> y"
+  by (simp add: real_vector_norm_def)
+lemma norm_lt: "norm(x::real ^ _) < norm(y) \<longleftrightarrow> x \<bullet> x < y \<bullet> y"
+  by (simp add: real_vector_norm_def)
+lemma norm_eq: "norm (x::real ^ _) = norm y \<longleftrightarrow> x \<bullet> x = y \<bullet> y"
+  by (simp add: order_eq_iff norm_le)
+lemma norm_eq_1: "norm(x::real ^ _) = 1 \<longleftrightarrow> x \<bullet> x = 1"
+  by (simp add: real_vector_norm_def)
+
+text{* Squaring equations and inequalities involving norms.  *}
+
+lemma dot_square_norm: "x \<bullet> x = norm(x)^2"
+  by (simp add: real_vector_norm_def)
+
+lemma norm_eq_square: "norm(x) = a \<longleftrightarrow> 0 <= a \<and> x \<bullet> x = a^2"
+  by (auto simp add: real_vector_norm_def)
+
+lemma real_abs_le_square_iff: "\<bar>x\<bar> \<le> \<bar>y\<bar> \<longleftrightarrow> (x::real)^2 \<le> y^2"
+proof-
+  have "x^2 \<le> y^2 \<longleftrightarrow> (x -y) * (y + x) \<le> 0" by (simp add: ring_simps power2_eq_square)
+  also have "\<dots> \<longleftrightarrow> \<bar>x\<bar> \<le> \<bar>y\<bar>" apply (simp add: zero_compare_simps real_abs_def not_less) by arith
+finally show ?thesis ..
+qed
+
+lemma norm_le_square: "norm(x) <= a \<longleftrightarrow> 0 <= a \<and> x \<bullet> x <= a^2"
+  apply (simp add: dot_square_norm real_abs_le_square_iff[symmetric])
+  using norm_ge_zero[of x]
+  apply arith
+  done
+
+lemma norm_ge_square: "norm(x) >= a \<longleftrightarrow> a <= 0 \<or> x \<bullet> x >= a ^ 2"
+  apply (simp add: dot_square_norm real_abs_le_square_iff[symmetric])
+  using norm_ge_zero[of x]
+  apply arith
+  done
+
+lemma norm_lt_square: "norm(x) < a \<longleftrightarrow> 0 < a \<and> x \<bullet> x < a^2"
+  by (metis not_le norm_ge_square)
+lemma norm_gt_square: "norm(x) > a \<longleftrightarrow> a < 0 \<or> x \<bullet> x > a^2"
+  by (metis norm_le_square not_less)
+
+text{* Dot product in terms of the norm rather than conversely. *}
+
+lemma dot_norm: "x \<bullet> y = (norm(x + y) ^2 - norm x ^ 2 - norm y ^ 2) / 2"
+  by (simp add: norm_pow_2 dot_ladd dot_radd dot_sym)
+
+lemma dot_norm_neg: "x \<bullet> y = ((norm x ^ 2 + norm y ^ 2) - norm(x - y) ^ 2) / 2"
+  by (simp add: norm_pow_2 dot_ladd dot_radd dot_lsub dot_rsub dot_sym)
+
+
+text{* Equality of vectors in terms of @{term "op \<bullet>"} products.    *}
+
+lemma vector_eq: "(x:: real ^ 'n::finite) = y \<longleftrightarrow> x \<bullet> x = x \<bullet> y\<and> y \<bullet> y = x \<bullet> x" (is "?lhs \<longleftrightarrow> ?rhs")
+proof
+  assume "?lhs" then show ?rhs by simp
+next
+  assume ?rhs
+  then have "x \<bullet> x - x \<bullet> y = 0 \<and> x \<bullet> y - y\<bullet> y = 0" by simp
+  hence "x \<bullet> (x - y) = 0 \<and> y \<bullet> (x - y) = 0"
+    by (simp add: dot_rsub dot_lsub dot_sym)
+  then have "(x - y) \<bullet> (x - y) = 0" by (simp add: ring_simps dot_lsub dot_rsub)
+  then show "x = y" by (simp add: dot_eq_0)
+qed
+
+
+subsection{* General linear decision procedure for normed spaces. *}
+
+lemma norm_cmul_rule_thm:
+  fixes x :: "'a::real_normed_vector"
+  shows "b >= norm(x) ==> \<bar>c\<bar> * b >= norm(scaleR c x)"
+  unfolding norm_scaleR
+  apply (erule mult_mono1)
+  apply simp
+  done
+
+  (* FIXME: Move all these theorems into the ML code using lemma antiquotation *)
+lemma norm_add_rule_thm:
+  fixes x1 x2 :: "'a::real_normed_vector"
+  shows "norm x1 \<le> b1 \<Longrightarrow> norm x2 \<le> b2 \<Longrightarrow> norm (x1 + x2) \<le> b1 + b2"
+  by (rule order_trans [OF norm_triangle_ineq add_mono])
+
+lemma ge_iff_diff_ge_0: "(a::'a::ordered_ring) \<ge> b == a - b \<ge> 0"
+  by (simp add: ring_simps)
+
+lemma pth_1:
+  fixes x :: "'a::real_normed_vector"
+  shows "x == scaleR 1 x" by simp
+
+lemma pth_2:
+  fixes x :: "'a::real_normed_vector"
+  shows "x - y == x + -y" by (atomize (full)) simp
+
+lemma pth_3:
+  fixes x :: "'a::real_normed_vector"
+  shows "- x == scaleR (-1) x" by simp
+
+lemma pth_4:
+  fixes x :: "'a::real_normed_vector"
+  shows "scaleR 0 x == 0" and "scaleR c 0 = (0::'a)" by simp_all
+
+lemma pth_5:
+  fixes x :: "'a::real_normed_vector"
+  shows "scaleR c (scaleR d x) == scaleR (c * d) x" by simp
+
+lemma pth_6:
+  fixes x :: "'a::real_normed_vector"
+  shows "scaleR c (x + y) == scaleR c x + scaleR c y"
+  by (simp add: scaleR_right_distrib)
+
+lemma pth_7:
+  fixes x :: "'a::real_normed_vector"
+  shows "0 + x == x" and "x + 0 == x" by simp_all
+
+lemma pth_8:
+  fixes x :: "'a::real_normed_vector"
+  shows "scaleR c x + scaleR d x == scaleR (c + d) x"
+  by (simp add: scaleR_left_distrib)
+
+lemma pth_9:
+  fixes x :: "'a::real_normed_vector" shows
+  "(scaleR c x + z) + scaleR d x == scaleR (c + d) x + z"
+  "scaleR c x + (scaleR d x + z) == scaleR (c + d) x + z"
+  "(scaleR c x + w) + (scaleR d x + z) == scaleR (c + d) x + (w + z)"
+  by (simp_all add: algebra_simps)
+
+lemma pth_a:
+  fixes x :: "'a::real_normed_vector"
+  shows "scaleR 0 x + y == y" by simp
+
+lemma pth_b:
+  fixes x :: "'a::real_normed_vector" shows
+  "scaleR c x + scaleR d y == scaleR c x + scaleR d y"
+  "(scaleR c x + z) + scaleR d y == scaleR c x + (z + scaleR d y)"
+  "scaleR c x + (scaleR d y + z) == scaleR c x + (scaleR d y + z)"
+  "(scaleR c x + w) + (scaleR d y + z) == scaleR c x + (w + (scaleR d y + z))"
+  by (simp_all add: algebra_simps)
+
+lemma pth_c:
+  fixes x :: "'a::real_normed_vector" shows
+  "scaleR c x + scaleR d y == scaleR d y + scaleR c x"
+  "(scaleR c x + z) + scaleR d y == scaleR d y + (scaleR c x + z)"
+  "scaleR c x + (scaleR d y + z) == scaleR d y + (scaleR c x + z)"
+  "(scaleR c x + w) + (scaleR d y + z) == scaleR d y + ((scaleR c x + w) + z)"
+  by (simp_all add: algebra_simps)
+
+lemma pth_d:
+  fixes x :: "'a::real_normed_vector"
+  shows "x + 0 == x" by simp
+
+lemma norm_imp_pos_and_ge:
+  fixes x :: "'a::real_normed_vector"
+  shows "norm x == n \<Longrightarrow> norm x \<ge> 0 \<and> n \<ge> norm x"
+  by atomize auto
+
+lemma real_eq_0_iff_le_ge_0: "(x::real) = 0 == x \<ge> 0 \<and> -x \<ge> 0" by arith
+
+lemma norm_pths:
+  fixes x :: "'a::real_normed_vector" shows
+  "x = y \<longleftrightarrow> norm (x - y) \<le> 0"
+  "x \<noteq> y \<longleftrightarrow> \<not> (norm (x - y) \<le> 0)"
+  using norm_ge_zero[of "x - y"] by auto
+
+lemma vector_dist_norm:
+  fixes x :: "'a::real_normed_vector"
+  shows "dist x y = norm (x - y)"
+  by (rule dist_norm)
+
+use "normarith.ML"
+
+method_setup norm = {* Scan.succeed (SIMPLE_METHOD' o NormArith.norm_arith_tac)
+*} "Proves simple linear statements about vector norms"
+
+
+text{* Hence more metric properties. *}
+
+lemma dist_triangle_alt:
+  fixes x y z :: "'a::metric_space"
+  shows "dist y z <= dist x y + dist x z"
+using dist_triangle [of y z x] by (simp add: dist_commute)
+
+lemma dist_pos_lt:
+  fixes x y :: "'a::metric_space"
+  shows "x \<noteq> y ==> 0 < dist x y"
+by (simp add: zero_less_dist_iff)
+
+lemma dist_nz:
+  fixes x y :: "'a::metric_space"
+  shows "x \<noteq> y \<longleftrightarrow> 0 < dist x y"
+by (simp add: zero_less_dist_iff)
+
+lemma dist_triangle_le:
+  fixes x y z :: "'a::metric_space"
+  shows "dist x z + dist y z <= e \<Longrightarrow> dist x y <= e"
+by (rule order_trans [OF dist_triangle2])
+
+lemma dist_triangle_lt:
+  fixes x y z :: "'a::metric_space"
+  shows "dist x z + dist y z < e ==> dist x y < e"
+by (rule le_less_trans [OF dist_triangle2])
+
+lemma dist_triangle_half_l:
+  fixes x1 x2 y :: "'a::metric_space"
+  shows "dist x1 y < e / 2 \<Longrightarrow> dist x2 y < e / 2 \<Longrightarrow> dist x1 x2 < e"
+by (rule dist_triangle_lt [where z=y], simp)
+
+lemma dist_triangle_half_r:
+  fixes x1 x2 y :: "'a::metric_space"
+  shows "dist y x1 < e / 2 \<Longrightarrow> dist y x2 < e / 2 \<Longrightarrow> dist x1 x2 < e"
+by (rule dist_triangle_half_l, simp_all add: dist_commute)
+
+lemma dist_triangle_add:
+  fixes x y x' y' :: "'a::real_normed_vector"
+  shows "dist (x + y) (x' + y') <= dist x x' + dist y y'"
+  by norm
+
+lemma dist_mul[simp]: "dist (c *s x) (c *s y) = \<bar>c\<bar> * dist x y"
+  unfolding dist_norm vector_ssub_ldistrib[symmetric] norm_mul ..
+
+lemma dist_triangle_add_half:
+  fixes x x' y y' :: "'a::real_normed_vector"
+  shows "dist x x' < e / 2 \<Longrightarrow> dist y y' < e / 2 \<Longrightarrow> dist(x + y) (x' + y') < e"
+  by norm
+
+lemma setsum_component [simp]:
+  fixes f:: " 'a \<Rightarrow> ('b::comm_monoid_add) ^'n"
+  shows "(setsum f S)$i = setsum (\<lambda>x. (f x)$i) S"
+  by (cases "finite S", induct S set: finite, simp_all)
+
+lemma setsum_eq: "setsum f S = (\<chi> i. setsum (\<lambda>x. (f x)$i ) S)"
+  by (simp add: Cart_eq)
+
+lemma setsum_clauses:
+  shows "setsum f {} = 0"
+  and "finite S \<Longrightarrow> setsum f (insert x S) =
+                 (if x \<in> S then setsum f S else f x + setsum f S)"
+  by (auto simp add: insert_absorb)
+
+lemma setsum_cmul:
+  fixes f:: "'c \<Rightarrow> ('a::semiring_1)^'n"
+  shows "setsum (\<lambda>x. c *s f x) S = c *s setsum f S"
+  by (simp add: Cart_eq setsum_right_distrib)
+
+lemma setsum_norm:
+  fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"
+  assumes fS: "finite S"
+  shows "norm (setsum f S) <= setsum (\<lambda>x. norm(f x)) S"
+proof(induct rule: finite_induct[OF fS])
+  case 1 thus ?case by simp
+next
+  case (2 x S)
+  from "2.hyps" have "norm (setsum f (insert x S)) \<le> norm (f x) + norm (setsum f S)" by (simp add: norm_triangle_ineq)
+  also have "\<dots> \<le> norm (f x) + setsum (\<lambda>x. norm(f x)) S"
+    using "2.hyps" by simp
+  finally  show ?case  using "2.hyps" by simp
+qed
+
+lemma real_setsum_norm:
+  fixes f :: "'a \<Rightarrow> real ^'n::finite"
+  assumes fS: "finite S"
+  shows "norm (setsum f S) <= setsum (\<lambda>x. norm(f x)) S"
+proof(induct rule: finite_induct[OF fS])
+  case 1 thus ?case by simp
+next
+  case (2 x S)
+  from "2.hyps" have "norm (setsum f (insert x S)) \<le> norm (f x) + norm (setsum f S)" by (simp add: norm_triangle_ineq)
+  also have "\<dots> \<le> norm (f x) + setsum (\<lambda>x. norm(f x)) S"
+    using "2.hyps" by simp
+  finally  show ?case  using "2.hyps" by simp
+qed
+
+lemma setsum_norm_le:
+  fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"
+  assumes fS: "finite S"
+  and fg: "\<forall>x \<in> S. norm (f x) \<le> g x"
+  shows "norm (setsum f S) \<le> setsum g S"
+proof-
+  from fg have "setsum (\<lambda>x. norm(f x)) S <= setsum g S"
+    by - (rule setsum_mono, simp)
+  then show ?thesis using setsum_norm[OF fS, of f] fg
+    by arith
+qed
+
+lemma real_setsum_norm_le:
+  fixes f :: "'a \<Rightarrow> real ^ 'n::finite"
+  assumes fS: "finite S"
+  and fg: "\<forall>x \<in> S. norm (f x) \<le> g x"
+  shows "norm (setsum f S) \<le> setsum g S"
+proof-
+  from fg have "setsum (\<lambda>x. norm(f x)) S <= setsum g S"
+    by - (rule setsum_mono, simp)
+  then show ?thesis using real_setsum_norm[OF fS, of f] fg
+    by arith
+qed
+
+lemma setsum_norm_bound:
+  fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"
+  assumes fS: "finite S"
+  and K: "\<forall>x \<in> S. norm (f x) \<le> K"
+  shows "norm (setsum f S) \<le> of_nat (card S) * K"
+  using setsum_norm_le[OF fS K] setsum_constant[symmetric]
+  by simp
+
+lemma real_setsum_norm_bound:
+  fixes f :: "'a \<Rightarrow> real ^ 'n::finite"
+  assumes fS: "finite S"
+  and K: "\<forall>x \<in> S. norm (f x) \<le> K"
+  shows "norm (setsum f S) \<le> of_nat (card S) * K"
+  using real_setsum_norm_le[OF fS K] setsum_constant[symmetric]
+  by simp
+
+lemma setsum_vmul:
+  fixes f :: "'a \<Rightarrow> 'b::{real_normed_vector,semiring, mult_zero}"
+  assumes fS: "finite S"
+  shows "setsum f S *s v = setsum (\<lambda>x. f x *s v) S"
+proof(induct rule: finite_induct[OF fS])
+  case 1 then show ?case by (simp add: vector_smult_lzero)
+next
+  case (2 x F)
+  from "2.hyps" have "setsum f (insert x F) *s v = (f x + setsum f F) *s v"
+    by simp
+  also have "\<dots> = f x *s v + setsum f F *s v"
+    by (simp add: vector_sadd_rdistrib)
+  also have "\<dots> = setsum (\<lambda>x. f x *s v) (insert x F)" using "2.hyps" by simp
+  finally show ?case .
+qed
+
+(* FIXME : Problem thm setsum_vmul[of _ "f:: 'a \<Rightarrow> real ^'n"]  ---
+ Get rid of *s and use real_vector instead! Also prove that ^ creates a real_vector !! *)
+
+    (* FIXME: Here too need stupid finiteness assumption on T!!! *)
+lemma setsum_group:
+  assumes fS: "finite S" and fT: "finite T" and fST: "f ` S \<subseteq> T"
+  shows "setsum (\<lambda>y. setsum g {x. x\<in> S \<and> f x = y}) T = setsum g S"
+
+apply (subst setsum_image_gen[OF fS, of g f])
+apply (rule setsum_mono_zero_right[OF fT fST])
+by (auto intro: setsum_0')
+
+lemma vsum_norm_allsubsets_bound:
+  fixes f:: "'a \<Rightarrow> real ^'n::finite"
+  assumes fP: "finite P" and fPs: "\<And>Q. Q \<subseteq> P \<Longrightarrow> norm (setsum f Q) \<le> e"
+  shows "setsum (\<lambda>x. norm (f x)) P \<le> 2 * real CARD('n) *  e"
+proof-
+  let ?d = "real CARD('n)"
+  let ?nf = "\<lambda>x. norm (f x)"
+  let ?U = "UNIV :: 'n set"
+  have th0: "setsum (\<lambda>x. setsum (\<lambda>i. \<bar>f x $ i\<bar>) ?U) P = setsum (\<lambda>i. setsum (\<lambda>x. \<bar>f x $ i\<bar>) P) ?U"
+    by (rule setsum_commute)
+  have th1: "2 * ?d * e = of_nat (card ?U) * (2 * e)" by (simp add: real_of_nat_def)
+  have "setsum ?nf P \<le> setsum (\<lambda>x. setsum (\<lambda>i. \<bar>f x $ i\<bar>) ?U) P"
+    apply (rule setsum_mono)
+    by (rule norm_le_l1)
+  also have "\<dots> \<le> 2 * ?d * e"
+    unfolding th0 th1
+  proof(rule setsum_bounded)
+    fix i assume i: "i \<in> ?U"
+    let ?Pp = "{x. x\<in> P \<and> f x $ i \<ge> 0}"
+    let ?Pn = "{x. x \<in> P \<and> f x $ i < 0}"
+    have thp: "P = ?Pp \<union> ?Pn" by auto
+    have thp0: "?Pp \<inter> ?Pn ={}" by auto
+    have PpP: "?Pp \<subseteq> P" and PnP: "?Pn \<subseteq> P" by blast+
+    have Ppe:"setsum (\<lambda>x. \<bar>f x $ i\<bar>) ?Pp \<le> e"
+      using component_le_norm[of "setsum (\<lambda>x. f x) ?Pp" i]  fPs[OF PpP]
+      by (auto intro: abs_le_D1)
+    have Pne: "setsum (\<lambda>x. \<bar>f x $ i\<bar>) ?Pn \<le> e"
+      using component_le_norm[of "setsum (\<lambda>x. - f x) ?Pn" i]  fPs[OF PnP]
+      by (auto simp add: setsum_negf intro: abs_le_D1)
+    have "setsum (\<lambda>x. \<bar>f x $ i\<bar>) P = setsum (\<lambda>x. \<bar>f x $ i\<bar>) ?Pp + setsum (\<lambda>x. \<bar>f x $ i\<bar>) ?Pn"
+      apply (subst thp)
+      apply (rule setsum_Un_zero)
+      using fP thp0 by auto
+    also have "\<dots> \<le> 2*e" using Pne Ppe by arith
+    finally show "setsum (\<lambda>x. \<bar>f x $ i\<bar>) P \<le> 2*e" .
+  qed
+  finally show ?thesis .
+qed
+
+lemma dot_lsum: "finite S \<Longrightarrow> setsum f S \<bullet> (y::'a::{comm_ring}^'n) = setsum (\<lambda>x. f x \<bullet> y) S "
+  by (induct rule: finite_induct, auto simp add: dot_lzero dot_ladd dot_radd)
+
+lemma dot_rsum: "finite S \<Longrightarrow> (y::'a::{comm_ring}^'n) \<bullet> setsum f S = setsum (\<lambda>x. y \<bullet> f x) S "
+  by (induct rule: finite_induct, auto simp add: dot_rzero dot_radd)
+
+subsection{* Basis vectors in coordinate directions. *}
+
+
+definition "basis k = (\<chi> i. if i = k then 1 else 0)"
+
+lemma basis_component [simp]: "basis k $ i = (if k=i then 1 else 0)"
+  unfolding basis_def by simp
+
+lemma delta_mult_idempotent:
+  "(if k=a then 1 else (0::'a::semiring_1)) * (if k=a then 1 else 0) = (if k=a then 1 else 0)" by (cases "k=a", auto)
+
+lemma norm_basis:
+  shows "norm (basis k :: real ^'n::finite) = 1"
+  apply (simp add: basis_def real_vector_norm_def dot_def)
+  apply (vector delta_mult_idempotent)
+  using setsum_delta[of "UNIV :: 'n set" "k" "\<lambda>k. 1::real"]
+  apply auto
+  done
+
+lemma norm_basis_1: "norm(basis 1 :: real ^'n::{finite,one}) = 1"
+  by (rule norm_basis)
+
+lemma vector_choose_size: "0 <= c ==> \<exists>(x::real^'n::finite). norm x = c"
+  apply (rule exI[where x="c *s basis arbitrary"])
+  by (simp only: norm_mul norm_basis)
+
+lemma vector_choose_dist: assumes e: "0 <= e"
+  shows "\<exists>(y::real^'n::finite). dist x y = e"
+proof-
+  from vector_choose_size[OF e] obtain c:: "real ^'n"  where "norm c = e"
+    by blast
+  then have "dist x (x - c) = e" by (simp add: dist_norm)
+  then show ?thesis by blast
+qed
+
+lemma basis_inj: "inj (basis :: 'n \<Rightarrow> real ^'n::finite)"
+  by (simp add: inj_on_def Cart_eq)
+
+lemma cond_value_iff: "f (if b then x else y) = (if b then f x else f y)"
+  by auto
+
+lemma basis_expansion:
+  "setsum (\<lambda>i. (x$i) *s basis i) UNIV = (x::('a::ring_1) ^'n::finite)" (is "?lhs = ?rhs" is "setsum ?f ?S = _")
+  by (auto simp add: Cart_eq cond_value_iff setsum_delta[of "?S", where ?'b = "'a", simplified] cong del: if_weak_cong)
+
+lemma basis_expansion_unique:
+  "setsum (\<lambda>i. f i *s basis i) UNIV = (x::('a::comm_ring_1) ^'n::finite) \<longleftrightarrow> (\<forall>i. f i = x$i)"
+  by (simp add: Cart_eq setsum_delta cond_value_iff cong del: if_weak_cong)
+
+lemma cond_application_beta: "(if b then f else g) x = (if b then f x else g x)"
+  by auto
+
+lemma dot_basis:
+  shows "basis i \<bullet> x = x$i" "x \<bullet> (basis i :: 'a^'n::finite) = (x$i :: 'a::semiring_1)"
+  by (auto simp add: dot_def basis_def cond_application_beta  cond_value_iff setsum_delta cong del: if_weak_cong)
+
+lemma inner_basis:
+  fixes x :: "'a::{real_inner, real_algebra_1} ^ 'n::finite"
+  shows "inner (basis i) x = inner 1 (x $ i)"
+    and "inner x (basis i) = inner (x $ i) 1"
+  unfolding inner_vector_def basis_def
+  by (auto simp add: cond_application_beta  cond_value_iff setsum_delta cong del: if_weak_cong)
+
+lemma basis_eq_0: "basis i = (0::'a::semiring_1^'n) \<longleftrightarrow> False"
+  by (auto simp add: Cart_eq)
+
+lemma basis_nonzero:
+  shows "basis k \<noteq> (0:: 'a::semiring_1 ^'n)"
+  by (simp add: basis_eq_0)
+
+lemma vector_eq_ldot: "(\<forall>x. x \<bullet> y = x \<bullet> z) \<longleftrightarrow> y = (z::'a::semiring_1^'n::finite)"
+  apply (auto simp add: Cart_eq dot_basis)
+  apply (erule_tac x="basis i" in allE)
+  apply (simp add: dot_basis)
+  apply (subgoal_tac "y = z")
+  apply simp
+  apply (simp add: Cart_eq)
+  done
+
+lemma vector_eq_rdot: "(\<forall>z. x \<bullet> z = y \<bullet> z) \<longleftrightarrow> x = (y::'a::semiring_1^'n::finite)"
+  apply (auto simp add: Cart_eq dot_basis)
+  apply (erule_tac x="basis i" in allE)
+  apply (simp add: dot_basis)
+  apply (subgoal_tac "x = y")
+  apply simp
+  apply (simp add: Cart_eq)
+  done
+
+subsection{* Orthogonality. *}
+
+definition "orthogonal x y \<longleftrightarrow> (x \<bullet> y = 0)"
+
+lemma orthogonal_basis:
+  shows "orthogonal (basis i :: 'a^'n::finite) x \<longleftrightarrow> x$i = (0::'a::ring_1)"
+  by (auto simp add: orthogonal_def dot_def basis_def cond_value_iff cond_application_beta setsum_delta cong del: if_weak_cong)
+
+lemma orthogonal_basis_basis:
+  shows "orthogonal (basis i :: 'a::ring_1^'n::finite) (basis j) \<longleftrightarrow> i \<noteq> j"
+  unfolding orthogonal_basis[of i] basis_component[of j] by simp
+
+  (* FIXME : Maybe some of these require less than comm_ring, but not all*)
+lemma orthogonal_clauses:
+  "orthogonal a (0::'a::comm_ring ^'n)"
+  "orthogonal a x ==> orthogonal a (c *s x)"
+  "orthogonal a x ==> orthogonal a (-x)"
+  "orthogonal a x \<Longrightarrow> orthogonal a y ==> orthogonal a (x + y)"
+  "orthogonal a x \<Longrightarrow> orthogonal a y ==> orthogonal a (x - y)"
+  "orthogonal 0 a"
+  "orthogonal x a ==> orthogonal (c *s x) a"
+  "orthogonal x a ==> orthogonal (-x) a"
+  "orthogonal x a \<Longrightarrow> orthogonal y a ==> orthogonal (x + y) a"
+  "orthogonal x a \<Longrightarrow> orthogonal y a ==> orthogonal (x - y) a"
+  unfolding orthogonal_def dot_rneg dot_rmult dot_radd dot_rsub
+  dot_lzero dot_rzero dot_lneg dot_lmult dot_ladd dot_lsub
+  by simp_all
+
+lemma orthogonal_commute: "orthogonal (x::'a::{ab_semigroup_mult,comm_monoid_add} ^'n)y \<longleftrightarrow> orthogonal y x"
+  by (simp add: orthogonal_def dot_sym)
+
+subsection{* Explicit vector construction from lists. *}
+
+primrec from_nat :: "nat \<Rightarrow> 'a::{monoid_add,one}"
+where "from_nat 0 = 0" | "from_nat (Suc n) = 1 + from_nat n"
+
+lemma from_nat [simp]: "from_nat = of_nat"
+by (rule ext, induct_tac x, simp_all)
+
+primrec
+  list_fun :: "nat \<Rightarrow> _ list \<Rightarrow> _ \<Rightarrow> _"
+where
+  "list_fun n [] = (\<lambda>x. 0)"
+| "list_fun n (x # xs) = fun_upd (list_fun (Suc n) xs) (from_nat n) x"
+
+definition "vector l = (\<chi> i. list_fun 1 l i)"
+(*definition "vector l = (\<chi> i. if i <= length l then l ! (i - 1) else 0)"*)
+
+lemma vector_1: "(vector[x]) $1 = x"
+  unfolding vector_def by simp
+
+lemma vector_2:
+ "(vector[x,y]) $1 = x"
+ "(vector[x,y] :: 'a^2)$2 = (y::'a::zero)"
+  unfolding vector_def by simp_all
+
+lemma vector_3:
+ "(vector [x,y,z] ::('a::zero)^3)$1 = x"
+ "(vector [x,y,z] ::('a::zero)^3)$2 = y"
+ "(vector [x,y,z] ::('a::zero)^3)$3 = z"
+  unfolding vector_def by simp_all
+
+lemma forall_vector_1: "(\<forall>v::'a::zero^1. P v) \<longleftrightarrow> (\<forall>x. P(vector[x]))"
+  apply auto
+  apply (erule_tac x="v$1" in allE)
+  apply (subgoal_tac "vector [v$1] = v")
+  apply simp
+  apply (vector vector_def)
+  apply (simp add: forall_1)
+  done
+
+lemma forall_vector_2: "(\<forall>v::'a::zero^2. P v) \<longleftrightarrow> (\<forall>x y. P(vector[x, y]))"
+  apply auto
+  apply (erule_tac x="v$1" in allE)
+  apply (erule_tac x="v$2" in allE)
+  apply (subgoal_tac "vector [v$1, v$2] = v")
+  apply simp
+  apply (vector vector_def)
+  apply (simp add: forall_2)
+  done
+
+lemma forall_vector_3: "(\<forall>v::'a::zero^3. P v) \<longleftrightarrow> (\<forall>x y z. P(vector[x, y, z]))"
+  apply auto
+  apply (erule_tac x="v$1" in allE)
+  apply (erule_tac x="v$2" in allE)
+  apply (erule_tac x="v$3" in allE)
+  apply (subgoal_tac "vector [v$1, v$2, v$3] = v")
+  apply simp
+  apply (vector vector_def)
+  apply (simp add: forall_3)
+  done
+
+subsection{* Linear functions. *}
+
+definition "linear f \<longleftrightarrow> (\<forall>x y. f(x + y) = f x + f y) \<and> (\<forall>c x. f(c *s x) = c *s f x)"
+
+lemma linear_compose_cmul: "linear f ==> linear (\<lambda>x. (c::'a::comm_semiring) *s f x)"
+  by (vector linear_def Cart_eq ring_simps)
+
+lemma linear_compose_neg: "linear (f :: 'a ^'n \<Rightarrow> 'a::comm_ring ^'m) ==> linear (\<lambda>x. -(f(x)))" by (vector linear_def Cart_eq)
+
+lemma linear_compose_add: "linear (f :: 'a ^'n \<Rightarrow> 'a::semiring_1 ^'m) \<Longrightarrow> linear g ==> linear (\<lambda>x. f(x) + g(x))"
+  by (vector linear_def Cart_eq ring_simps)
+
+lemma linear_compose_sub: "linear (f :: 'a ^'n \<Rightarrow> 'a::ring_1 ^'m) \<Longrightarrow> linear g ==> linear (\<lambda>x. f x - g x)"
+  by (vector linear_def Cart_eq ring_simps)
+
+lemma linear_compose: "linear f \<Longrightarrow> linear g ==> linear (g o f)"
+  by (simp add: linear_def)
+
+lemma linear_id: "linear id" by (simp add: linear_def id_def)
+
+lemma linear_zero: "linear (\<lambda>x. 0::'a::semiring_1 ^ 'n)" by (simp add: linear_def)
+
+lemma linear_compose_setsum:
+  assumes fS: "finite S" and lS: "\<forall>a \<in> S. linear (f a :: 'a::semiring_1 ^ 'n \<Rightarrow> 'a ^ 'm)"
+  shows "linear(\<lambda>x. setsum (\<lambda>a. f a x :: 'a::semiring_1 ^'m) S)"
+  using lS
+  apply (induct rule: finite_induct[OF fS])
+  by (auto simp add: linear_zero intro: linear_compose_add)
+
+lemma linear_vmul_component:
+  fixes f:: "'a::semiring_1^'m \<Rightarrow> 'a^'n"
+  assumes lf: "linear f"
+  shows "linear (\<lambda>x. f x $ k *s v)"
+  using lf
+  apply (auto simp add: linear_def )
+  by (vector ring_simps)+
+
+lemma linear_0: "linear f ==> f 0 = (0::'a::semiring_1 ^'n)"
+  unfolding linear_def
+  apply clarsimp
+  apply (erule allE[where x="0::'a"])
+  apply simp
+  done
+
+lemma linear_cmul: "linear f ==> f(c*s x) = c *s f x" by (simp add: linear_def)
+
+lemma linear_neg: "linear (f :: 'a::ring_1 ^'n \<Rightarrow> _) ==> f (-x) = - f x"
+  unfolding vector_sneg_minus1
+  using linear_cmul[of f] by auto
+
+lemma linear_add: "linear f ==> f(x + y) = f x + f y" by (metis linear_def)
+
+lemma linear_sub: "linear (f::'a::ring_1 ^'n \<Rightarrow> _) ==> f(x - y) = f x - f y"
+  by (simp add: diff_def linear_add linear_neg)
+
+lemma linear_setsum:
+  fixes f:: "'a::semiring_1^'n \<Rightarrow> _"
+  assumes lf: "linear f" and fS: "finite S"
+  shows "f (setsum g S) = setsum (f o g) S"
+proof (induct rule: finite_induct[OF fS])
+  case 1 thus ?case by (simp add: linear_0[OF lf])
+next
+  case (2 x F)
+  have "f (setsum g (insert x F)) = f (g x + setsum g F)" using "2.hyps"
+    by simp
+  also have "\<dots> = f (g x) + f (setsum g F)" using linear_add[OF lf] by simp
+  also have "\<dots> = setsum (f o g) (insert x F)" using "2.hyps" by simp
+  finally show ?case .
+qed
+
+lemma linear_setsum_mul:
+  fixes f:: "'a ^'n \<Rightarrow> 'a::semiring_1^'m"
+  assumes lf: "linear f" and fS: "finite S"
+  shows "f (setsum (\<lambda>i. c i *s v i) S) = setsum (\<lambda>i. c i *s f (v i)) S"
+  using linear_setsum[OF lf fS, of "\<lambda>i. c i *s v i" , unfolded o_def]
+  linear_cmul[OF lf] by simp
+
+lemma linear_injective_0:
+  assumes lf: "linear (f:: 'a::ring_1 ^ 'n \<Rightarrow> _)"
+  shows "inj f \<longleftrightarrow> (\<forall>x. f x = 0 \<longrightarrow> x = 0)"
+proof-
+  have "inj f \<longleftrightarrow> (\<forall> x y. f x = f y \<longrightarrow> x = y)" by (simp add: inj_on_def)
+  also have "\<dots> \<longleftrightarrow> (\<forall> x y. f x - f y = 0 \<longrightarrow> x - y = 0)" by simp
+  also have "\<dots> \<longleftrightarrow> (\<forall> x y. f (x - y) = 0 \<longrightarrow> x - y = 0)"
+    by (simp add: linear_sub[OF lf])
+  also have "\<dots> \<longleftrightarrow> (\<forall> x. f x = 0 \<longrightarrow> x = 0)" by auto
+  finally show ?thesis .
+qed
+
+lemma linear_bounded:
+  fixes f:: "real ^'m::finite \<Rightarrow> real ^'n::finite"
+  assumes lf: "linear f"
+  shows "\<exists>B. \<forall>x. norm (f x) \<le> B * norm x"
+proof-
+  let ?S = "UNIV:: 'm set"
+  let ?B = "setsum (\<lambda>i. norm(f(basis i))) ?S"
+  have fS: "finite ?S" by simp
+  {fix x:: "real ^ 'm"
+    let ?g = "(\<lambda>i. (x$i) *s (basis i) :: real ^ 'm)"
+    have "norm (f x) = norm (f (setsum (\<lambda>i. (x$i) *s (basis i)) ?S))"
+      by (simp only:  basis_expansion)
+    also have "\<dots> = norm (setsum (\<lambda>i. (x$i) *s f (basis i))?S)"
+      using linear_setsum[OF lf fS, of ?g, unfolded o_def] linear_cmul[OF lf]
+      by auto
+    finally have th0: "norm (f x) = norm (setsum (\<lambda>i. (x$i) *s f (basis i))?S)" .
+    {fix i assume i: "i \<in> ?S"
+      from component_le_norm[of x i]
+      have "norm ((x$i) *s f (basis i :: real ^'m)) \<le> norm (f (basis i)) * norm x"
+      unfolding norm_mul
+      apply (simp only: mult_commute)
+      apply (rule mult_mono)
+      by (auto simp add: ring_simps norm_ge_zero) }
+    then have th: "\<forall>i\<in> ?S. norm ((x$i) *s f (basis i :: real ^'m)) \<le> norm (f (basis i)) * norm x" by metis
+    from real_setsum_norm_le[OF fS, of "\<lambda>i. (x$i) *s (f (basis i))", OF th]
+    have "norm (f x) \<le> ?B * norm x" unfolding th0 setsum_left_distrib by metis}
+  then show ?thesis by blast
+qed
+
+lemma linear_bounded_pos:
+  fixes f:: "real ^'n::finite \<Rightarrow> real ^ 'm::finite"
+  assumes lf: "linear f"
+  shows "\<exists>B > 0. \<forall>x. norm (f x) \<le> B * norm x"
+proof-
+  from linear_bounded[OF lf] obtain B where
+    B: "\<forall>x. norm (f x) \<le> B * norm x" by blast
+  let ?K = "\<bar>B\<bar> + 1"
+  have Kp: "?K > 0" by arith
+    {assume C: "B < 0"
+      have "norm (1::real ^ 'n) > 0" by (simp add: zero_less_norm_iff)
+      with C have "B * norm (1:: real ^ 'n) < 0"
+        by (simp add: zero_compare_simps)
+      with B[rule_format, of 1] norm_ge_zero[of "f 1"] have False by simp
+    }
+    then have Bp: "B \<ge> 0" by ferrack
+    {fix x::"real ^ 'n"
+      have "norm (f x) \<le> ?K *  norm x"
+      using B[rule_format, of x] norm_ge_zero[of x] norm_ge_zero[of "f x"] Bp
+      apply (auto simp add: ring_simps split add: abs_split)
+      apply (erule order_trans, simp)
+      done
+  }
+  then show ?thesis using Kp by blast
+qed
+
+lemma smult_conv_scaleR: "c *s x = scaleR c x"
+  unfolding vector_scalar_mult_def vector_scaleR_def by simp
+
+lemma linear_conv_bounded_linear:
+  fixes f :: "real ^ _ \<Rightarrow> real ^ _"
+  shows "linear f \<longleftrightarrow> bounded_linear f"
+proof
+  assume "linear f"
+  show "bounded_linear f"
+  proof
+    fix x y show "f (x + y) = f x + f y"
+      using `linear f` unfolding linear_def by simp
+  next
+    fix r x show "f (scaleR r x) = scaleR r (f x)"
+      using `linear f` unfolding linear_def
+      by (simp add: smult_conv_scaleR)
+  next
+    have "\<exists>B. \<forall>x. norm (f x) \<le> B * norm x"
+      using `linear f` by (rule linear_bounded)
+    thus "\<exists>K. \<forall>x. norm (f x) \<le> norm x * K"
+      by (simp add: mult_commute)
+  qed
+next
+  assume "bounded_linear f"
+  then interpret f: bounded_linear f .
+  show "linear f"
+    unfolding linear_def smult_conv_scaleR
+    by (simp add: f.add f.scaleR)
+qed
+
+subsection{* Bilinear functions. *}
+
+definition "bilinear f \<longleftrightarrow> (\<forall>x. linear(\<lambda>y. f x y)) \<and> (\<forall>y. linear(\<lambda>x. f x y))"
+
+lemma bilinear_ladd: "bilinear h ==> h (x + y) z = (h x z) + (h y z)"
+  by (simp add: bilinear_def linear_def)
+lemma bilinear_radd: "bilinear h ==> h x (y + z) = (h x y) + (h x z)"
+  by (simp add: bilinear_def linear_def)
+
+lemma bilinear_lmul: "bilinear h ==> h (c *s x) y = c *s (h x y)"
+  by (simp add: bilinear_def linear_def)
+
+lemma bilinear_rmul: "bilinear h ==> h x (c *s y) = c *s (h x y)"
+  by (simp add: bilinear_def linear_def)
+
+lemma bilinear_lneg: "bilinear h ==> h (- (x:: 'a::ring_1 ^ 'n)) y = -(h x y)"
+  by (simp only: vector_sneg_minus1 bilinear_lmul)
+
+lemma bilinear_rneg: "bilinear h ==> h x (- (y:: 'a::ring_1 ^ 'n)) = - h x y"
+  by (simp only: vector_sneg_minus1 bilinear_rmul)
+
+lemma  (in ab_group_add) eq_add_iff: "x = x + y \<longleftrightarrow> y = 0"
+  using add_imp_eq[of x y 0] by auto
+
+lemma bilinear_lzero:
+  fixes h :: "'a::ring^'n \<Rightarrow> _" assumes bh: "bilinear h" shows "h 0 x = 0"
+  using bilinear_ladd[OF bh, of 0 0 x]
+    by (simp add: eq_add_iff ring_simps)
+
+lemma bilinear_rzero:
+  fixes h :: "'a::ring^'n \<Rightarrow> _" assumes bh: "bilinear h" shows "h x 0 = 0"
+  using bilinear_radd[OF bh, of x 0 0 ]
+    by (simp add: eq_add_iff ring_simps)
+
+lemma bilinear_lsub: "bilinear h ==> h (x - (y:: 'a::ring_1 ^ 'n)) z = h x z - h y z"
+  by (simp  add: diff_def bilinear_ladd bilinear_lneg)
+
+lemma bilinear_rsub: "bilinear h ==> h z (x - (y:: 'a::ring_1 ^ 'n)) = h z x - h z y"
+  by (simp  add: diff_def bilinear_radd bilinear_rneg)
+
+lemma bilinear_setsum:
+  fixes h:: "'a ^'n \<Rightarrow> 'a::semiring_1^'m \<Rightarrow> 'a ^ 'k"
+  assumes bh: "bilinear h" and fS: "finite S" and fT: "finite T"
+  shows "h (setsum f S) (setsum g T) = setsum (\<lambda>(i,j). h (f i) (g j)) (S \<times> T) "
+proof-
+  have "h (setsum f S) (setsum g T) = setsum (\<lambda>x. h (f x) (setsum g T)) S"
+    apply (rule linear_setsum[unfolded o_def])
+    using bh fS by (auto simp add: bilinear_def)
+  also have "\<dots> = setsum (\<lambda>x. setsum (\<lambda>y. h (f x) (g y)) T) S"
+    apply (rule setsum_cong, simp)
+    apply (rule linear_setsum[unfolded o_def])
+    using bh fT by (auto simp add: bilinear_def)
+  finally show ?thesis unfolding setsum_cartesian_product .
+qed
+
+lemma bilinear_bounded:
+  fixes h:: "real ^'m::finite \<Rightarrow> real^'n::finite \<Rightarrow> real ^ 'k::finite"
+  assumes bh: "bilinear h"
+  shows "\<exists>B. \<forall>x y. norm (h x y) \<le> B * norm x * norm y"
+proof-
+  let ?M = "UNIV :: 'm set"
+  let ?N = "UNIV :: 'n set"
+  let ?B = "setsum (\<lambda>(i,j). norm (h (basis i) (basis j))) (?M \<times> ?N)"
+  have fM: "finite ?M" and fN: "finite ?N" by simp_all
+  {fix x:: "real ^ 'm" and  y :: "real^'n"
+    have "norm (h x y) = norm (h (setsum (\<lambda>i. (x$i) *s basis i) ?M) (setsum (\<lambda>i. (y$i) *s basis i) ?N))" unfolding basis_expansion ..
+    also have "\<dots> = norm (setsum (\<lambda> (i,j). h ((x$i) *s basis i) ((y$j) *s basis j)) (?M \<times> ?N))"  unfolding bilinear_setsum[OF bh fM fN] ..
+    finally have th: "norm (h x y) = \<dots>" .
+    have "norm (h x y) \<le> ?B * norm x * norm y"
+      apply (simp add: setsum_left_distrib th)
+      apply (rule real_setsum_norm_le)
+      using fN fM
+      apply simp
+      apply (auto simp add: bilinear_rmul[OF bh] bilinear_lmul[OF bh] norm_mul ring_simps)
+      apply (rule mult_mono)
+      apply (auto simp add: norm_ge_zero zero_le_mult_iff component_le_norm)
+      apply (rule mult_mono)
+      apply (auto simp add: norm_ge_zero zero_le_mult_iff component_le_norm)
+      done}
+  then show ?thesis by metis
+qed
+
+lemma bilinear_bounded_pos:
+  fixes h:: "real ^'m::finite \<Rightarrow> real^'n::finite \<Rightarrow> real ^ 'k::finite"
+  assumes bh: "bilinear h"
+  shows "\<exists>B > 0. \<forall>x y. norm (h x y) \<le> B * norm x * norm y"
+proof-
+  from bilinear_bounded[OF bh] obtain B where
+    B: "\<forall>x y. norm (h x y) \<le> B * norm x * norm y" by blast
+  let ?K = "\<bar>B\<bar> + 1"
+  have Kp: "?K > 0" by arith
+  have KB: "B < ?K" by arith
+  {fix x::"real ^'m" and y :: "real ^'n"
+    from KB Kp
+    have "B * norm x * norm y \<le> ?K * norm x * norm y"
+      apply -
+      apply (rule mult_right_mono, rule mult_right_mono)
+      by (auto simp add: norm_ge_zero)
+    then have "norm (h x y) \<le> ?K * norm x * norm y"
+      using B[rule_format, of x y] by simp}
+  with Kp show ?thesis by blast
+qed
+
+lemma bilinear_conv_bounded_bilinear:
+  fixes h :: "real ^ _ \<Rightarrow> real ^ _ \<Rightarrow> real ^ _"
+  shows "bilinear h \<longleftrightarrow> bounded_bilinear h"
+proof
+  assume "bilinear h"
+  show "bounded_bilinear h"
+  proof
+    fix x y z show "h (x + y) z = h x z + h y z"
+      using `bilinear h` unfolding bilinear_def linear_def by simp
+  next
+    fix x y z show "h x (y + z) = h x y + h x z"
+      using `bilinear h` unfolding bilinear_def linear_def by simp
+  next
+    fix r x y show "h (scaleR r x) y = scaleR r (h x y)"
+      using `bilinear h` unfolding bilinear_def linear_def
+      by (simp add: smult_conv_scaleR)
+  next
+    fix r x y show "h x (scaleR r y) = scaleR r (h x y)"
+      using `bilinear h` unfolding bilinear_def linear_def
+      by (simp add: smult_conv_scaleR)
+  next
+    have "\<exists>B. \<forall>x y. norm (h x y) \<le> B * norm x * norm y"
+      using `bilinear h` by (rule bilinear_bounded)
+    thus "\<exists>K. \<forall>x y. norm (h x y) \<le> norm x * norm y * K"
+      by (simp add: mult_ac)
+  qed
+next
+  assume "bounded_bilinear h"
+  then interpret h: bounded_bilinear h .
+  show "bilinear h"
+    unfolding bilinear_def linear_conv_bounded_linear
+    using h.bounded_linear_left h.bounded_linear_right
+    by simp
+qed
+
+subsection{* Adjoints. *}
+
+definition "adjoint f = (SOME f'. \<forall>x y. f x \<bullet> y = x \<bullet> f' y)"
+
+lemma choice_iff: "(\<forall>x. \<exists>y. P x y) \<longleftrightarrow> (\<exists>f. \<forall>x. P x (f x))" by metis
+
+lemma adjoint_works_lemma:
+  fixes f:: "'a::ring_1 ^'n::finite \<Rightarrow> 'a ^ 'm::finite"
+  assumes lf: "linear f"
+  shows "\<forall>x y. f x \<bullet> y = x \<bullet> adjoint f y"
+proof-
+  let ?N = "UNIV :: 'n set"
+  let ?M = "UNIV :: 'm set"
+  have fN: "finite ?N" by simp
+  have fM: "finite ?M" by simp
+  {fix y:: "'a ^ 'm"
+    let ?w = "(\<chi> i. (f (basis i) \<bullet> y)) :: 'a ^ 'n"
+    {fix x
+      have "f x \<bullet> y = f (setsum (\<lambda>i. (x$i) *s basis i) ?N) \<bullet> y"
+        by (simp only: basis_expansion)
+      also have "\<dots> = (setsum (\<lambda>i. (x$i) *s f (basis i)) ?N) \<bullet> y"
+        unfolding linear_setsum[OF lf fN]
+        by (simp add: linear_cmul[OF lf])
+      finally have "f x \<bullet> y = x \<bullet> ?w"
+        apply (simp only: )
+        apply (simp add: dot_def setsum_left_distrib setsum_right_distrib setsum_commute[of _ ?M ?N] ring_simps)
+        done}
+  }
+  then show ?thesis unfolding adjoint_def
+    some_eq_ex[of "\<lambda>f'. \<forall>x y. f x \<bullet> y = x \<bullet> f' y"]
+    using choice_iff[of "\<lambda>a b. \<forall>x. f x \<bullet> a = x \<bullet> b "]
+    by metis
+qed
+
+lemma adjoint_works:
+  fixes f:: "'a::ring_1 ^'n::finite \<Rightarrow> 'a ^ 'm::finite"
+  assumes lf: "linear f"
+  shows "x \<bullet> adjoint f y = f x \<bullet> y"
+  using adjoint_works_lemma[OF lf] by metis
+
+
+lemma adjoint_linear:
+  fixes f :: "'a::comm_ring_1 ^'n::finite \<Rightarrow> 'a ^ 'm::finite"
+  assumes lf: "linear f"
+  shows "linear (adjoint f)"
+  by (simp add: linear_def vector_eq_ldot[symmetric] dot_radd dot_rmult adjoint_works[OF lf])
+
+lemma adjoint_clauses:
+  fixes f:: "'a::comm_ring_1 ^'n::finite \<Rightarrow> 'a ^ 'm::finite"
+  assumes lf: "linear f"
+  shows "x \<bullet> adjoint f y = f x \<bullet> y"
+  and "adjoint f y \<bullet> x = y \<bullet> f x"
+  by (simp_all add: adjoint_works[OF lf] dot_sym )
+
+lemma adjoint_adjoint:
+  fixes f:: "'a::comm_ring_1 ^ 'n::finite \<Rightarrow> 'a ^ 'm::finite"
+  assumes lf: "linear f"
+  shows "adjoint (adjoint f) = f"
+  apply (rule ext)
+  by (simp add: vector_eq_ldot[symmetric] adjoint_clauses[OF adjoint_linear[OF lf]] adjoint_clauses[OF lf])
+
+lemma adjoint_unique:
+  fixes f:: "'a::comm_ring_1 ^ 'n::finite \<Rightarrow> 'a ^ 'm::finite"
+  assumes lf: "linear f" and u: "\<forall>x y. f' x \<bullet> y = x \<bullet> f y"
+  shows "f' = adjoint f"
+  apply (rule ext)
+  using u
+  by (simp add: vector_eq_rdot[symmetric] adjoint_clauses[OF lf])
+
+text{* Matrix notation. NB: an MxN matrix is of type @{typ "'a^'n^'m"}, not @{typ "'a^'m^'n"} *}
+
+consts generic_mult :: "'a \<Rightarrow> 'b \<Rightarrow> 'c" (infixr "\<star>" 75)
+
+defs (overloaded)
+matrix_matrix_mult_def: "(m:: ('a::semiring_1) ^'n^'m) \<star> (m' :: 'a ^'p^'n) \<equiv> (\<chi> i j. setsum (\<lambda>k. ((m$i)$k) * ((m'$k)$j)) (UNIV :: 'n set)) ::'a ^ 'p ^'m"
+
+abbreviation
+  matrix_matrix_mult' :: "('a::semiring_1) ^'n^'m \<Rightarrow> 'a ^'p^'n \<Rightarrow> 'a ^ 'p ^'m"  (infixl "**" 70)
+  where "m ** m' == m\<star> m'"
+
+defs (overloaded)
+  matrix_vector_mult_def: "(m::('a::semiring_1) ^'n^'m) \<star> (x::'a ^'n) \<equiv> (\<chi> i. setsum (\<lambda>j. ((m$i)$j) * (x$j)) (UNIV ::'n set)) :: 'a^'m"
+
+abbreviation
+  matrix_vector_mult' :: "('a::semiring_1) ^'n^'m \<Rightarrow> 'a ^'n \<Rightarrow> 'a ^ 'm"  (infixl "*v" 70)
+  where
+  "m *v v == m \<star> v"
+
+defs (overloaded)
+  vector_matrix_mult_def: "(x::'a^'m) \<star> (m::('a::semiring_1) ^'n^'m) \<equiv> (\<chi> j. setsum (\<lambda>i. ((m$i)$j) * (x$i)) (UNIV :: 'm set)) :: 'a^'n"
+
+abbreviation
+  vactor_matrix_mult' :: "'a ^ 'm \<Rightarrow> ('a::semiring_1) ^'n^'m \<Rightarrow> 'a ^'n "  (infixl "v*" 70)
+  where
+  "v v* m == v \<star> m"
+
+definition "(mat::'a::zero => 'a ^'n^'n) k = (\<chi> i j. if i = j then k else 0)"
+definition "(transp::'a^'n^'m \<Rightarrow> 'a^'m^'n) A = (\<chi> i j. ((A$j)$i))"
+definition "(row::'m => 'a ^'n^'m \<Rightarrow> 'a ^'n) i A = (\<chi> j. ((A$i)$j))"
+definition "(column::'n =>'a^'n^'m =>'a^'m) j A = (\<chi> i. ((A$i)$j))"
+definition "rows(A::'a^'n^'m) = { row i A | i. i \<in> (UNIV :: 'm set)}"
+definition "columns(A::'a^'n^'m) = { column i A | i. i \<in> (UNIV :: 'n set)}"
+
+lemma mat_0[simp]: "mat 0 = 0" by (vector mat_def)
+lemma matrix_add_ldistrib: "(A ** (B + C)) = (A \<star> B) + (A \<star> C)"
+  by (vector matrix_matrix_mult_def setsum_addf[symmetric] ring_simps)
+
+lemma matrix_mul_lid:
+  fixes A :: "'a::semiring_1 ^ 'm ^ 'n::finite"
+  shows "mat 1 ** A = A"
+  apply (simp add: matrix_matrix_mult_def mat_def)
+  apply vector
+  by (auto simp only: cond_value_iff cond_application_beta setsum_delta'[OF finite]  mult_1_left mult_zero_left if_True UNIV_I)
+
+
+lemma matrix_mul_rid:
+  fixes A :: "'a::semiring_1 ^ 'm::finite ^ 'n"
+  shows "A ** mat 1 = A"
+  apply (simp add: matrix_matrix_mult_def mat_def)
+  apply vector
+  by (auto simp only: cond_value_iff cond_application_beta setsum_delta[OF finite]  mult_1_right mult_zero_right if_True UNIV_I cong: if_cong)
+
+lemma matrix_mul_assoc: "A ** (B ** C) = (A ** B) ** C"
+  apply (vector matrix_matrix_mult_def setsum_right_distrib setsum_left_distrib mult_assoc)
+  apply (subst setsum_commute)
+  apply simp
+  done
+
+lemma matrix_vector_mul_assoc: "A *v (B *v x) = (A ** B) *v x"
+  apply (vector matrix_matrix_mult_def matrix_vector_mult_def setsum_right_distrib setsum_left_distrib mult_assoc)
+  apply (subst setsum_commute)
+  apply simp
+  done
+
+lemma matrix_vector_mul_lid: "mat 1 *v x = (x::'a::semiring_1 ^ 'n::finite)"
+  apply (vector matrix_vector_mult_def mat_def)
+  by (simp add: cond_value_iff cond_application_beta
+    setsum_delta' cong del: if_weak_cong)
+
+lemma matrix_transp_mul: "transp(A ** B) = transp B ** transp (A::'a::comm_semiring_1^'m^'n)"
+  by (simp add: matrix_matrix_mult_def transp_def Cart_eq mult_commute)
+
+lemma matrix_eq:
+  fixes A B :: "'a::semiring_1 ^ 'n::finite ^ 'm"
+  shows "A = B \<longleftrightarrow>  (\<forall>x. A *v x = B *v x)" (is "?lhs \<longleftrightarrow> ?rhs")
+  apply auto
+  apply (subst Cart_eq)
+  apply clarify
+  apply (clarsimp simp add: matrix_vector_mult_def basis_def cond_value_iff cond_application_beta Cart_eq cong del: if_weak_cong)
+  apply (erule_tac x="basis ia" in allE)
+  apply (erule_tac x="i" in allE)
+  by (auto simp add: basis_def cond_value_iff cond_application_beta setsum_delta[OF finite] cong del: if_weak_cong)
+
+lemma matrix_vector_mul_component:
+  shows "((A::'a::semiring_1^'n'^'m) *v x)$k = (A$k) \<bullet> x"
+  by (simp add: matrix_vector_mult_def dot_def)
+
+lemma dot_lmul_matrix: "((x::'a::comm_semiring_1 ^'n) v* A) \<bullet> y = x \<bullet> (A *v y)"
+  apply (simp add: dot_def matrix_vector_mult_def vector_matrix_mult_def setsum_left_distrib setsum_right_distrib mult_ac)
+  apply (subst setsum_commute)
+  by simp
+
+lemma transp_mat: "transp (mat n) = mat n"
+  by (vector transp_def mat_def)
+
+lemma transp_transp: "transp(transp A) = A"
+  by (vector transp_def)
+
+lemma row_transp:
+  fixes A:: "'a::semiring_1^'n^'m"
+  shows "row i (transp A) = column i A"
+  by (simp add: row_def column_def transp_def Cart_eq)
+
+lemma column_transp:
+  fixes A:: "'a::semiring_1^'n^'m"
+  shows "column i (transp A) = row i A"
+  by (simp add: row_def column_def transp_def Cart_eq)
+
+lemma rows_transp: "rows(transp (A::'a::semiring_1^'n^'m)) = columns A"
+by (auto simp add: rows_def columns_def row_transp intro: set_ext)
+
+lemma columns_transp: "columns(transp (A::'a::semiring_1^'n^'m)) = rows A" by (metis transp_transp rows_transp)
+
+text{* Two sometimes fruitful ways of looking at matrix-vector multiplication. *}
+
+lemma matrix_mult_dot: "A *v x = (\<chi> i. A$i \<bullet> x)"
+  by (simp add: matrix_vector_mult_def dot_def)
+
+lemma matrix_mult_vsum: "(A::'a::comm_semiring_1^'n^'m) *v x = setsum (\<lambda>i. (x$i) *s column i A) (UNIV:: 'n set)"
+  by (simp add: matrix_vector_mult_def Cart_eq column_def mult_commute)
+
+lemma vector_componentwise:
+  "(x::'a::ring_1^'n::finite) = (\<chi> j. setsum (\<lambda>i. (x$i) * (basis i :: 'a^'n)$j) (UNIV :: 'n set))"
+  apply (subst basis_expansion[symmetric])
+  by (vector Cart_eq setsum_component)
+
+lemma linear_componentwise:
+  fixes f:: "'a::ring_1 ^ 'm::finite \<Rightarrow> 'a ^ 'n"
+  assumes lf: "linear f"
+  shows "(f x)$j = setsum (\<lambda>i. (x$i) * (f (basis i)$j)) (UNIV :: 'm set)" (is "?lhs = ?rhs")
+proof-
+  let ?M = "(UNIV :: 'm set)"
+  let ?N = "(UNIV :: 'n set)"
+  have fM: "finite ?M" by simp
+  have "?rhs = (setsum (\<lambda>i.(x$i) *s f (basis i) ) ?M)$j"
+    unfolding vector_smult_component[symmetric]
+    unfolding setsum_component[of "(\<lambda>i.(x$i) *s f (basis i :: 'a^'m))" ?M]
+    ..
+  then show ?thesis unfolding linear_setsum_mul[OF lf fM, symmetric] basis_expansion ..
+qed
+
+text{* Inverse matrices  (not necessarily square) *}
+
+definition "invertible(A::'a::semiring_1^'n^'m) \<longleftrightarrow> (\<exists>A'::'a^'m^'n. A ** A' = mat 1 \<and> A' ** A = mat 1)"
+
+definition "matrix_inv(A:: 'a::semiring_1^'n^'m) =
+        (SOME A'::'a^'m^'n. A ** A' = mat 1 \<and> A' ** A = mat 1)"
+
+text{* Correspondence between matrices and linear operators. *}
+
+definition matrix:: "('a::{plus,times, one, zero}^'m \<Rightarrow> 'a ^ 'n) \<Rightarrow> 'a^'m^'n"
+where "matrix f = (\<chi> i j. (f(basis j))$i)"
+
+lemma matrix_vector_mul_linear: "linear(\<lambda>x. A *v (x::'a::comm_semiring_1 ^ 'n))"
+  by (simp add: linear_def matrix_vector_mult_def Cart_eq ring_simps setsum_right_distrib setsum_addf)
+
+lemma matrix_works: assumes lf: "linear f" shows "matrix f *v x = f (x::'a::comm_ring_1 ^ 'n::finite)"
+apply (simp add: matrix_def matrix_vector_mult_def Cart_eq mult_commute)
+apply clarify
+apply (rule linear_componentwise[OF lf, symmetric])
+done
+
+lemma matrix_vector_mul: "linear f ==> f = (\<lambda>x. matrix f *v (x::'a::comm_ring_1 ^ 'n::finite))" by (simp add: ext matrix_works)
+
+lemma matrix_of_matrix_vector_mul: "matrix(\<lambda>x. A *v (x :: 'a:: comm_ring_1 ^ 'n::finite)) = A"
+  by (simp add: matrix_eq matrix_vector_mul_linear matrix_works)
+
+lemma matrix_compose:
+  assumes lf: "linear (f::'a::comm_ring_1^'n::finite \<Rightarrow> 'a^'m::finite)"
+  and lg: "linear (g::'a::comm_ring_1^'m::finite \<Rightarrow> 'a^'k)"
+  shows "matrix (g o f) = matrix g ** matrix f"
+  using lf lg linear_compose[OF lf lg] matrix_works[OF linear_compose[OF lf lg]]
+  by (simp  add: matrix_eq matrix_works matrix_vector_mul_assoc[symmetric] o_def)
+
+lemma matrix_vector_column:"(A::'a::comm_semiring_1^'n^'m) *v x = setsum (\<lambda>i. (x$i) *s ((transp A)$i)) (UNIV:: 'n set)"
+  by (simp add: matrix_vector_mult_def transp_def Cart_eq mult_commute)
+
+lemma adjoint_matrix: "adjoint(\<lambda>x. (A::'a::comm_ring_1^'n::finite^'m::finite) *v x) = (\<lambda>x. transp A *v x)"
+  apply (rule adjoint_unique[symmetric])
+  apply (rule matrix_vector_mul_linear)
+  apply (simp add: transp_def dot_def matrix_vector_mult_def setsum_left_distrib setsum_right_distrib)
+  apply (subst setsum_commute)
+  apply (auto simp add: mult_ac)
+  done
+
+lemma matrix_adjoint: assumes lf: "linear (f :: 'a::comm_ring_1^'n::finite \<Rightarrow> 'a ^ 'm::finite)"
+  shows "matrix(adjoint f) = transp(matrix f)"
+  apply (subst matrix_vector_mul[OF lf])
+  unfolding adjoint_matrix matrix_of_matrix_vector_mul ..
+
+subsection{* Interlude: Some properties of real sets *}
+
+lemma seq_mono_lemma: assumes "\<forall>(n::nat) \<ge> m. (d n :: real) < e n" and "\<forall>n \<ge> m. e n <= e m"
+  shows "\<forall>n \<ge> m. d n < e m"
+  using prems apply auto
+  apply (erule_tac x="n" in allE)
+  apply (erule_tac x="n" in allE)
+  apply auto
+  done
+
+
+lemma real_convex_bound_lt:
+  assumes xa: "(x::real) < a" and ya: "y < a" and u: "0 <= u" and v: "0 <= v"
+  and uv: "u + v = 1"
+  shows "u * x + v * y < a"
+proof-
+  have uv': "u = 0 \<longrightarrow> v \<noteq> 0" using u v uv by arith
+  have "a = a * (u + v)" unfolding uv  by simp
+  hence th: "u * a + v * a = a" by (simp add: ring_simps)
+  from xa u have "u \<noteq> 0 \<Longrightarrow> u*x < u*a" by (simp add: mult_compare_simps)
+  from ya v have "v \<noteq> 0 \<Longrightarrow> v * y < v * a" by (simp add: mult_compare_simps)
+  from xa ya u v have "u * x + v * y < u * a + v * a"
+    apply (cases "u = 0", simp_all add: uv')
+    apply(rule mult_strict_left_mono)
+    using uv' apply simp_all
+
+    apply (rule add_less_le_mono)
+    apply(rule mult_strict_left_mono)
+    apply simp_all
+    apply (rule mult_left_mono)
+    apply simp_all
+    done
+  thus ?thesis unfolding th .
+qed
+
+lemma real_convex_bound_le:
+  assumes xa: "(x::real) \<le> a" and ya: "y \<le> a" and u: "0 <= u" and v: "0 <= v"
+  and uv: "u + v = 1"
+  shows "u * x + v * y \<le> a"
+proof-
+  from xa ya u v have "u * x + v * y \<le> u * a + v * a" by (simp add: add_mono mult_left_mono)
+  also have "\<dots> \<le> (u + v) * a" by (simp add: ring_simps)
+  finally show ?thesis unfolding uv by simp
+qed
+
+lemma infinite_enumerate: assumes fS: "infinite S"
+  shows "\<exists>r. subseq r \<and> (\<forall>n. r n \<in> S)"
+unfolding subseq_def
+using enumerate_in_set[OF fS] enumerate_mono[of _ _ S] fS by auto
+
+lemma approachable_lt_le: "(\<exists>(d::real)>0. \<forall>x. f x < d \<longrightarrow> P x) \<longleftrightarrow> (\<exists>d>0. \<forall>x. f x \<le> d \<longrightarrow> P x)"
+apply auto
+apply (rule_tac x="d/2" in exI)
+apply auto
+done
+
+
+lemma triangle_lemma:
+  assumes x: "0 <= (x::real)" and y:"0 <= y" and z: "0 <= z" and xy: "x^2 <= y^2 + z^2"
+  shows "x <= y + z"
+proof-
+  have "y^2 + z^2 \<le> y^2 + 2*y*z + z^2" using z y  by (simp add: zero_compare_simps)
+  with xy have th: "x ^2 \<le> (y+z)^2" by (simp add: power2_eq_square ring_simps)
+  from y z have yz: "y + z \<ge> 0" by arith
+  from power2_le_imp_le[OF th yz] show ?thesis .
+qed
+
+
+lemma lambda_skolem: "(\<forall>i. \<exists>x. P i x) \<longleftrightarrow>
+   (\<exists>x::'a ^ 'n. \<forall>i. P i (x$i))" (is "?lhs \<longleftrightarrow> ?rhs")
+proof-
+  let ?S = "(UNIV :: 'n set)"
+  {assume H: "?rhs"
+    then have ?lhs by auto}
+  moreover
+  {assume H: "?lhs"
+    then obtain f where f:"\<forall>i. P i (f i)" unfolding choice_iff by metis
+    let ?x = "(\<chi> i. (f i)) :: 'a ^ 'n"
+    {fix i
+      from f have "P i (f i)" by metis
+      then have "P i (?x$i)" by auto
+    }
+    hence "\<forall>i. P i (?x$i)" by metis
+    hence ?rhs by metis }
+  ultimately show ?thesis by metis
+qed
+
+(* Supremum and infimum of real sets *)
+
+
+definition rsup:: "real set \<Rightarrow> real" where
+  "rsup S = (SOME a. isLub UNIV S a)"
+
+lemma rsup_alt: "rsup S = (SOME a. (\<forall>x \<in> S. x \<le> a) \<and> (\<forall>b. (\<forall>x \<in> S. x \<le> b) \<longrightarrow> a \<le> b))"  by (auto simp  add: isLub_def rsup_def leastP_def isUb_def setle_def setge_def)
+
+lemma rsup: assumes Se: "S \<noteq> {}" and b: "\<exists>b. S *<= b"
+  shows "isLub UNIV S (rsup S)"
+using Se b
+unfolding rsup_def
+apply clarify
+apply (rule someI_ex)
+apply (rule reals_complete)
+by (auto simp add: isUb_def setle_def)
+
+lemma rsup_le: assumes Se: "S \<noteq> {}" and Sb: "S *<= b" shows "rsup S \<le> b"
+proof-
+  from Sb have bu: "isUb UNIV S b" by (simp add: isUb_def setle_def)
+  from rsup[OF Se] Sb have "isLub UNIV S (rsup S)"  by blast
+  then show ?thesis using bu by (auto simp add: isLub_def leastP_def setle_def setge_def)
+qed
+
+lemma rsup_finite_Max: assumes fS: "finite S" and Se: "S \<noteq> {}"
+  shows "rsup S = Max S"
+using fS Se
+proof-
+  let ?m = "Max S"
+  from Max_ge[OF fS] have Sm: "\<forall> x\<in> S. x \<le> ?m" by metis
+  with rsup[OF Se] have lub: "isLub UNIV S (rsup S)" by (metis setle_def)
+  from Max_in[OF fS Se] lub have mrS: "?m \<le> rsup S"
+    by (auto simp add: isLub_def leastP_def setle_def setge_def isUb_def)
+  moreover
+  have "rsup S \<le> ?m" using Sm lub
+    by (auto simp add: isLub_def leastP_def isUb_def setle_def setge_def)
+  ultimately  show ?thesis by arith
+qed
+
+lemma rsup_finite_in: assumes fS: "finite S" and Se: "S \<noteq> {}"
+  shows "rsup S \<in> S"
+  using rsup_finite_Max[OF fS Se] Max_in[OF fS Se] by metis
+
+lemma rsup_finite_Ub: assumes fS: "finite S" and Se: "S \<noteq> {}"
+  shows "isUb S S (rsup S)"
+  using rsup_finite_Max[OF fS Se] rsup_finite_in[OF fS Se] Max_ge[OF fS]
+  unfolding isUb_def setle_def by metis
+
+lemma rsup_finite_ge_iff: assumes fS: "finite S" and Se: "S \<noteq> {}"
+  shows "a \<le> rsup S \<longleftrightarrow> (\<exists> x \<in> S. a \<le> x)"
+using rsup_finite_Ub[OF fS Se] by (auto simp add: isUb_def setle_def)
+
+lemma rsup_finite_le_iff: assumes fS: "finite S" and Se: "S \<noteq> {}"
+  shows "a \<ge> rsup S \<longleftrightarrow> (\<forall> x \<in> S. a \<ge> x)"
+using rsup_finite_Ub[OF fS Se] by (auto simp add: isUb_def setle_def)
+
+lemma rsup_finite_gt_iff: assumes fS: "finite S" and Se: "S \<noteq> {}"
+  shows "a < rsup S \<longleftrightarrow> (\<exists> x \<in> S. a < x)"
+using rsup_finite_Ub[OF fS Se] by (auto simp add: isUb_def setle_def)
+
+lemma rsup_finite_lt_iff: assumes fS: "finite S" and Se: "S \<noteq> {}"
+  shows "a > rsup S \<longleftrightarrow> (\<forall> x \<in> S. a > x)"
+using rsup_finite_Ub[OF fS Se] by (auto simp add: isUb_def setle_def)
+
+lemma rsup_unique: assumes b: "S *<= b" and S: "\<forall>b' < b. \<exists>x \<in> S. b' < x"
+  shows "rsup S = b"
+using b S
+unfolding setle_def rsup_alt
+apply -
+apply (rule some_equality)
+apply (metis  linorder_not_le order_eq_iff[symmetric])+
+done
+
+lemma rsup_le_subset: "S\<noteq>{} \<Longrightarrow> S \<subseteq> T \<Longrightarrow> (\<exists>b. T *<= b) \<Longrightarrow> rsup S \<le> rsup T"
+  apply (rule rsup_le)
+  apply simp
+  using rsup[of T] by (auto simp add: isLub_def leastP_def setge_def setle_def isUb_def)
+
+lemma isUb_def': "isUb R S = (\<lambda>x. S *<= x \<and> x \<in> R)"
+  apply (rule ext)
+  by (metis isUb_def)
+
+lemma UNIV_trivial: "UNIV x" using UNIV_I[of x] by (metis mem_def)
+lemma rsup_bounds: assumes Se: "S \<noteq> {}" and l: "a <=* S" and u: "S *<= b"
+  shows "a \<le> rsup S \<and> rsup S \<le> b"
+proof-
+  from rsup[OF Se] u have lub: "isLub UNIV S (rsup S)" by blast
+  hence b: "rsup S \<le> b" using u by (auto simp add: isLub_def leastP_def setle_def setge_def isUb_def')
+  from Se obtain y where y: "y \<in> S" by blast
+  from lub l have "a \<le> rsup S" apply (auto simp add: isLub_def leastP_def setle_def setge_def isUb_def')
+    apply (erule ballE[where x=y])
+    apply (erule ballE[where x=y])
+    apply arith
+    using y apply auto
+    done
+  with b show ?thesis by blast
+qed
+
+lemma rsup_abs_le: "S \<noteq> {} \<Longrightarrow> (\<forall>x\<in>S. \<bar>x\<bar> \<le> a) \<Longrightarrow> \<bar>rsup S\<bar> \<le> a"
+  unfolding abs_le_interval_iff  using rsup_bounds[of S "-a" a]
+  by (auto simp add: setge_def setle_def)
+
+lemma rsup_asclose: assumes S:"S \<noteq> {}" and b: "\<forall>x\<in>S. \<bar>x - l\<bar> \<le> e" shows "\<bar>rsup S - l\<bar> \<le> e"
+proof-
+  have th: "\<And>(x::real) l e. \<bar>x - l\<bar> \<le> e \<longleftrightarrow> l - e \<le> x \<and> x \<le> l + e" by arith
+  show ?thesis using S b rsup_bounds[of S "l - e" "l+e"] unfolding th
+    by  (auto simp add: setge_def setle_def)
+qed
+
+definition rinf:: "real set \<Rightarrow> real" where
+  "rinf S = (SOME a. isGlb UNIV S a)"
+
+lemma rinf_alt: "rinf S = (SOME a. (\<forall>x \<in> S. x \<ge> a) \<and> (\<forall>b. (\<forall>x \<in> S. x \<ge> b) \<longrightarrow> a \<ge> b))"  by (auto simp  add: isGlb_def rinf_def greatestP_def isLb_def setle_def setge_def)
+
+lemma reals_complete_Glb: assumes Se: "\<exists>x. x \<in> S" and lb: "\<exists> y. isLb UNIV S y"
+  shows "\<exists>(t::real). isGlb UNIV S t"
+proof-
+  let ?M = "uminus ` S"
+  from lb have th: "\<exists>y. isUb UNIV ?M y" apply (auto simp add: isUb_def isLb_def setle_def setge_def)
+    by (rule_tac x="-y" in exI, auto)
+  from Se have Me: "\<exists>x. x \<in> ?M" by blast
+  from reals_complete[OF Me th] obtain t where t: "isLub UNIV ?M t" by blast
+  have "isGlb UNIV S (- t)" using t
+    apply (auto simp add: isLub_def isGlb_def leastP_def greatestP_def setle_def setge_def isUb_def isLb_def)
+    apply (erule_tac x="-y" in allE)
+    apply auto
+    done
+  then show ?thesis by metis
+qed
+
+lemma rinf: assumes Se: "S \<noteq> {}" and b: "\<exists>b. b <=* S"
+  shows "isGlb UNIV S (rinf S)"
+using Se b
+unfolding rinf_def
+apply clarify
+apply (rule someI_ex)
+apply (rule reals_complete_Glb)
+apply (auto simp add: isLb_def setle_def setge_def)
+done
+
+lemma rinf_ge: assumes Se: "S \<noteq> {}" and Sb: "b <=* S" shows "rinf S \<ge> b"
+proof-
+  from Sb have bu: "isLb UNIV S b" by (simp add: isLb_def setge_def)
+  from rinf[OF Se] Sb have "isGlb UNIV S (rinf S)"  by blast
+  then show ?thesis using bu by (auto simp add: isGlb_def greatestP_def setle_def setge_def)
+qed
+
+lemma rinf_finite_Min: assumes fS: "finite S" and Se: "S \<noteq> {}"
+  shows "rinf S = Min S"
+using fS Se
+proof-
+  let ?m = "Min S"
+  from Min_le[OF fS] have Sm: "\<forall> x\<in> S. x \<ge> ?m" by metis
+  with rinf[OF Se] have glb: "isGlb UNIV S (rinf S)" by (metis setge_def)
+  from Min_in[OF fS Se] glb have mrS: "?m \<ge> rinf S"
+    by (auto simp add: isGlb_def greatestP_def setle_def setge_def isLb_def)
+  moreover
+  have "rinf S \<ge> ?m" using Sm glb
+    by (auto simp add: isGlb_def greatestP_def isLb_def setle_def setge_def)
+  ultimately  show ?thesis by arith
+qed
+
+lemma rinf_finite_in: assumes fS: "finite S" and Se: "S \<noteq> {}"
+  shows "rinf S \<in> S"
+  using rinf_finite_Min[OF fS Se] Min_in[OF fS Se] by metis
+
+lemma rinf_finite_Lb: assumes fS: "finite S" and Se: "S \<noteq> {}"
+  shows "isLb S S (rinf S)"
+  using rinf_finite_Min[OF fS Se] rinf_finite_in[OF fS Se] Min_le[OF fS]
+  unfolding isLb_def setge_def by metis
+
+lemma rinf_finite_ge_iff: assumes fS: "finite S" and Se: "S \<noteq> {}"
+  shows "a \<le> rinf S \<longleftrightarrow> (\<forall> x \<in> S. a \<le> x)"
+using rinf_finite_Lb[OF fS Se] by (auto simp add: isLb_def setge_def)
+
+lemma rinf_finite_le_iff: assumes fS: "finite S" and Se: "S \<noteq> {}"
+  shows "a \<ge> rinf S \<longleftrightarrow> (\<exists> x \<in> S. a \<ge> x)"
+using rinf_finite_Lb[OF fS Se] by (auto simp add: isLb_def setge_def)
+
+lemma rinf_finite_gt_iff: assumes fS: "finite S" and Se: "S \<noteq> {}"
+  shows "a < rinf S \<longleftrightarrow> (\<forall> x \<in> S. a < x)"
+using rinf_finite_Lb[OF fS Se] by (auto simp add: isLb_def setge_def)
+
+lemma rinf_finite_lt_iff: assumes fS: "finite S" and Se: "S \<noteq> {}"
+  shows "a > rinf S \<longleftrightarrow> (\<exists> x \<in> S. a > x)"
+using rinf_finite_Lb[OF fS Se] by (auto simp add: isLb_def setge_def)
+
+lemma rinf_unique: assumes b: "b <=* S" and S: "\<forall>b' > b. \<exists>x \<in> S. b' > x"
+  shows "rinf S = b"
+using b S
+unfolding setge_def rinf_alt
+apply -
+apply (rule some_equality)
+apply (metis  linorder_not_le order_eq_iff[symmetric])+
+done
+
+lemma rinf_ge_subset: "S\<noteq>{} \<Longrightarrow> S \<subseteq> T \<Longrightarrow> (\<exists>b. b <=* T) \<Longrightarrow> rinf S >= rinf T"
+  apply (rule rinf_ge)
+  apply simp
+  using rinf[of T] by (auto simp add: isGlb_def greatestP_def setge_def setle_def isLb_def)
+
+lemma isLb_def': "isLb R S = (\<lambda>x. x <=* S \<and> x \<in> R)"
+  apply (rule ext)
+  by (metis isLb_def)
+
+lemma rinf_bounds: assumes Se: "S \<noteq> {}" and l: "a <=* S" and u: "S *<= b"
+  shows "a \<le> rinf S \<and> rinf S \<le> b"
+proof-
+  from rinf[OF Se] l have lub: "isGlb UNIV S (rinf S)" by blast
+  hence b: "a \<le> rinf S" using l by (auto simp add: isGlb_def greatestP_def setle_def setge_def isLb_def')
+  from Se obtain y where y: "y \<in> S" by blast
+  from lub u have "b \<ge> rinf S" apply (auto simp add: isGlb_def greatestP_def setle_def setge_def isLb_def')
+    apply (erule ballE[where x=y])
+    apply (erule ballE[where x=y])
+    apply arith
+    using y apply auto
+    done
+  with b show ?thesis by blast
+qed
+
+lemma rinf_abs_ge: "S \<noteq> {} \<Longrightarrow> (\<forall>x\<in>S. \<bar>x\<bar> \<le> a) \<Longrightarrow> \<bar>rinf S\<bar> \<le> a"
+  unfolding abs_le_interval_iff  using rinf_bounds[of S "-a" a]
+  by (auto simp add: setge_def setle_def)
+
+lemma rinf_asclose: assumes S:"S \<noteq> {}" and b: "\<forall>x\<in>S. \<bar>x - l\<bar> \<le> e" shows "\<bar>rinf S - l\<bar> \<le> e"
+proof-
+  have th: "\<And>(x::real) l e. \<bar>x - l\<bar> \<le> e \<longleftrightarrow> l - e \<le> x \<and> x \<le> l + e" by arith
+  show ?thesis using S b rinf_bounds[of S "l - e" "l+e"] unfolding th
+    by  (auto simp add: setge_def setle_def)
+qed
+
+
+
+subsection{* Operator norm. *}
+
+definition "onorm f = rsup {norm (f x)| x. norm x = 1}"
+
+lemma norm_bound_generalize:
+  fixes f:: "real ^'n::finite \<Rightarrow> real^'m::finite"
+  assumes lf: "linear f"
+  shows "(\<forall>x. norm x = 1 \<longrightarrow> norm (f x) \<le> b) \<longleftrightarrow> (\<forall>x. norm (f x) \<le> b * norm x)" (is "?lhs \<longleftrightarrow> ?rhs")
+proof-
+  {assume H: ?rhs
+    {fix x :: "real^'n" assume x: "norm x = 1"
+      from H[rule_format, of x] x have "norm (f x) \<le> b" by simp}
+    then have ?lhs by blast }
+
+  moreover
+  {assume H: ?lhs
+    from H[rule_format, of "basis arbitrary"]
+    have bp: "b \<ge> 0" using norm_ge_zero[of "f (basis arbitrary)"]
+      by (auto simp add: norm_basis elim: order_trans [OF norm_ge_zero])
+    {fix x :: "real ^'n"
+      {assume "x = 0"
+        then have "norm (f x) \<le> b * norm x" by (simp add: linear_0[OF lf] bp)}
+      moreover
+      {assume x0: "x \<noteq> 0"
+        hence n0: "norm x \<noteq> 0" by (metis norm_eq_zero)
+        let ?c = "1/ norm x"
+        have "norm (?c*s x) = 1" using x0 by (simp add: n0 norm_mul)
+        with H have "norm (f(?c*s x)) \<le> b" by blast
+        hence "?c * norm (f x) \<le> b"
+          by (simp add: linear_cmul[OF lf] norm_mul)
+        hence "norm (f x) \<le> b * norm x"
+          using n0 norm_ge_zero[of x] by (auto simp add: field_simps)}
+      ultimately have "norm (f x) \<le> b * norm x" by blast}
+    then have ?rhs by blast}
+  ultimately show ?thesis by blast
+qed
+
+lemma onorm:
+  fixes f:: "real ^'n::finite \<Rightarrow> real ^'m::finite"
+  assumes lf: "linear f"
+  shows "norm (f x) <= onorm f * norm x"
+  and "\<forall>x. norm (f x) <= b * norm x \<Longrightarrow> onorm f <= b"
+proof-
+  {
+    let ?S = "{norm (f x) |x. norm x = 1}"
+    have Se: "?S \<noteq> {}" using  norm_basis by auto
+    from linear_bounded[OF lf] have b: "\<exists> b. ?S *<= b"
+      unfolding norm_bound_generalize[OF lf, symmetric] by (auto simp add: setle_def)
+    {from rsup[OF Se b, unfolded onorm_def[symmetric]]
+      show "norm (f x) <= onorm f * norm x"
+        apply -
+        apply (rule spec[where x = x])
+        unfolding norm_bound_generalize[OF lf, symmetric]
+        by (auto simp add: isLub_def isUb_def leastP_def setge_def setle_def)}
+    {
+      show "\<forall>x. norm (f x) <= b * norm x \<Longrightarrow> onorm f <= b"
+        using rsup[OF Se b, unfolded onorm_def[symmetric]]
+        unfolding norm_bound_generalize[OF lf, symmetric]
+        by (auto simp add: isLub_def isUb_def leastP_def setge_def setle_def)}
+  }
+qed
+
+lemma onorm_pos_le: assumes lf: "linear (f::real ^'n::finite \<Rightarrow> real ^'m::finite)" shows "0 <= onorm f"
+  using order_trans[OF norm_ge_zero onorm(1)[OF lf, of "basis arbitrary"], unfolded norm_basis] by simp
+
+lemma onorm_eq_0: assumes lf: "linear (f::real ^'n::finite \<Rightarrow> real ^'m::finite)"
+  shows "onorm f = 0 \<longleftrightarrow> (\<forall>x. f x = 0)"
+  using onorm[OF lf]
+  apply (auto simp add: onorm_pos_le)
+  apply atomize
+  apply (erule allE[where x="0::real"])
+  using onorm_pos_le[OF lf]
+  apply arith
+  done
+
+lemma onorm_const: "onorm(\<lambda>x::real^'n::finite. (y::real ^ 'm::finite)) = norm y"
+proof-
+  let ?f = "\<lambda>x::real^'n. (y::real ^ 'm)"
+  have th: "{norm (?f x)| x. norm x = 1} = {norm y}"
+    by(auto intro: vector_choose_size set_ext)
+  show ?thesis
+    unfolding onorm_def th
+    apply (rule rsup_unique) by (simp_all  add: setle_def)
+qed
+
+lemma onorm_pos_lt: assumes lf: "linear (f::real ^ 'n::finite \<Rightarrow> real ^'m::finite)"
+  shows "0 < onorm f \<longleftrightarrow> ~(\<forall>x. f x = 0)"
+  unfolding onorm_eq_0[OF lf, symmetric]
+  using onorm_pos_le[OF lf] by arith
+
+lemma onorm_compose:
+  assumes lf: "linear (f::real ^'n::finite \<Rightarrow> real ^'m::finite)"
+  and lg: "linear (g::real^'k::finite \<Rightarrow> real^'n::finite)"
+  shows "onorm (f o g) <= onorm f * onorm g"
+  apply (rule onorm(2)[OF linear_compose[OF lg lf], rule_format])
+  unfolding o_def
+  apply (subst mult_assoc)
+  apply (rule order_trans)
+  apply (rule onorm(1)[OF lf])
+  apply (rule mult_mono1)
+  apply (rule onorm(1)[OF lg])
+  apply (rule onorm_pos_le[OF lf])
+  done
+
+lemma onorm_neg_lemma: assumes lf: "linear (f::real ^'n::finite \<Rightarrow> real^'m::finite)"
+  shows "onorm (\<lambda>x. - f x) \<le> onorm f"
+  using onorm[OF linear_compose_neg[OF lf]] onorm[OF lf]
+  unfolding norm_minus_cancel by metis
+
+lemma onorm_neg: assumes lf: "linear (f::real ^'n::finite \<Rightarrow> real^'m::finite)"
+  shows "onorm (\<lambda>x. - f x) = onorm f"
+  using onorm_neg_lemma[OF lf] onorm_neg_lemma[OF linear_compose_neg[OF lf]]
+  by simp
+
+lemma onorm_triangle:
+  assumes lf: "linear (f::real ^'n::finite \<Rightarrow> real ^'m::finite)" and lg: "linear g"
+  shows "onorm (\<lambda>x. f x + g x) <= onorm f + onorm g"
+  apply(rule onorm(2)[OF linear_compose_add[OF lf lg], rule_format])
+  apply (rule order_trans)
+  apply (rule norm_triangle_ineq)
+  apply (simp add: distrib)
+  apply (rule add_mono)
+  apply (rule onorm(1)[OF lf])
+  apply (rule onorm(1)[OF lg])
+  done
+
+lemma onorm_triangle_le: "linear (f::real ^'n::finite \<Rightarrow> real ^'m::finite) \<Longrightarrow> linear g \<Longrightarrow> onorm(f) + onorm(g) <= e
+  \<Longrightarrow> onorm(\<lambda>x. f x + g x) <= e"
+  apply (rule order_trans)
+  apply (rule onorm_triangle)
+  apply assumption+
+  done
+
+lemma onorm_triangle_lt: "linear (f::real ^'n::finite \<Rightarrow> real ^'m::finite) \<Longrightarrow> linear g \<Longrightarrow> onorm(f) + onorm(g) < e
+  ==> onorm(\<lambda>x. f x + g x) < e"
+  apply (rule order_le_less_trans)
+  apply (rule onorm_triangle)
+  by assumption+
+
+(* "lift" from 'a to 'a^1 and "drop" from 'a^1 to 'a -- FIXME: potential use of transfer *)
+
+definition vec1:: "'a \<Rightarrow> 'a ^ 1" where "vec1 x = (\<chi> i. x)"
+
+definition dest_vec1:: "'a ^1 \<Rightarrow> 'a"
+  where "dest_vec1 x = (x$1)"
+
+lemma vec1_component[simp]: "(vec1 x)$1 = x"
+  by (simp add: vec1_def)
+
+lemma vec1_dest_vec1[simp]: "vec1(dest_vec1 x) = x" "dest_vec1(vec1 y) = y"
+  by (simp_all add: vec1_def dest_vec1_def Cart_eq forall_1)
+
+lemma forall_vec1: "(\<forall>x. P x) \<longleftrightarrow> (\<forall>x. P (vec1 x))" by (metis vec1_dest_vec1)
+
+lemma exists_vec1: "(\<exists>x. P x) \<longleftrightarrow> (\<exists>x. P(vec1 x))" by (metis vec1_dest_vec1)
+
+lemma forall_dest_vec1: "(\<forall>x. P x) \<longleftrightarrow> (\<forall>x. P(dest_vec1 x))"  by (metis vec1_dest_vec1)
+
+lemma exists_dest_vec1: "(\<exists>x. P x) \<longleftrightarrow> (\<exists>x. P(dest_vec1 x))"by (metis vec1_dest_vec1)
+
+lemma vec1_eq[simp]:  "vec1 x = vec1 y \<longleftrightarrow> x = y" by (metis vec1_dest_vec1)
+
+lemma dest_vec1_eq[simp]: "dest_vec1 x = dest_vec1 y \<longleftrightarrow> x = y" by (metis vec1_dest_vec1)
+
+lemma vec1_in_image_vec1: "vec1 x \<in> (vec1 ` S) \<longleftrightarrow> x \<in> S" by auto
+
+lemma vec1_vec: "vec1 x = vec x" by (vector vec1_def)
+
+lemma vec1_add: "vec1(x + y) = vec1 x + vec1 y" by (vector vec1_def)
+lemma vec1_sub: "vec1(x - y) = vec1 x - vec1 y" by (vector vec1_def)
+lemma vec1_cmul: "vec1(c* x) = c *s vec1 x " by (vector vec1_def)
+lemma vec1_neg: "vec1(- x) = - vec1 x " by (vector vec1_def)
+
+lemma vec1_setsum: assumes fS: "finite S"
+  shows "vec1(setsum f S) = setsum (vec1 o f) S"
+  apply (induct rule: finite_induct[OF fS])
+  apply (simp add: vec1_vec)
+  apply (auto simp add: vec1_add)
+  done
+
+lemma dest_vec1_lambda: "dest_vec1(\<chi> i. x i) = x 1"
+  by (simp add: dest_vec1_def)
+
+lemma dest_vec1_vec: "dest_vec1(vec x) = x"
+  by (simp add: vec1_vec[symmetric])
+
+lemma dest_vec1_add: "dest_vec1(x + y) = dest_vec1 x + dest_vec1 y"
+ by (metis vec1_dest_vec1 vec1_add)
+
+lemma dest_vec1_sub: "dest_vec1(x - y) = dest_vec1 x - dest_vec1 y"
+ by (metis vec1_dest_vec1 vec1_sub)
+
+lemma dest_vec1_cmul: "dest_vec1(c*sx) = c * dest_vec1 x"
+ by (metis vec1_dest_vec1 vec1_cmul)
+
+lemma dest_vec1_neg: "dest_vec1(- x) = - dest_vec1 x"
+ by (metis vec1_dest_vec1 vec1_neg)
+
+lemma dest_vec1_0[simp]: "dest_vec1 0 = 0" by (metis vec_0 dest_vec1_vec)
+
+lemma dest_vec1_sum: assumes fS: "finite S"
+  shows "dest_vec1(setsum f S) = setsum (dest_vec1 o f) S"
+  apply (induct rule: finite_induct[OF fS])
+  apply (simp add: dest_vec1_vec)
+  apply (auto simp add: dest_vec1_add)
+  done
+
+lemma norm_vec1: "norm(vec1 x) = abs(x)"
+  by (simp add: vec1_def norm_real)
+
+lemma dist_vec1: "dist(vec1 x) (vec1 y) = abs(x - y)"
+  by (simp only: dist_real vec1_component)
+lemma abs_dest_vec1: "norm x = \<bar>dest_vec1 x\<bar>"
+  by (metis vec1_dest_vec1 norm_vec1)
+
+lemma linear_vmul_dest_vec1:
+  fixes f:: "'a::semiring_1^'n \<Rightarrow> 'a^1"
+  shows "linear f \<Longrightarrow> linear (\<lambda>x. dest_vec1(f x) *s v)"
+  unfolding dest_vec1_def
+  apply (rule linear_vmul_component)
+  by auto
+
+lemma linear_from_scalars:
+  assumes lf: "linear (f::'a::comm_ring_1 ^1 \<Rightarrow> 'a^'n)"
+  shows "f = (\<lambda>x. dest_vec1 x *s column 1 (matrix f))"
+  apply (rule ext)
+  apply (subst matrix_works[OF lf, symmetric])
+  apply (auto simp add: Cart_eq matrix_vector_mult_def dest_vec1_def column_def  mult_commute UNIV_1)
+  done
+
+lemma linear_to_scalars: assumes lf: "linear (f::'a::comm_ring_1 ^'n::finite \<Rightarrow> 'a^1)"
+  shows "f = (\<lambda>x. vec1(row 1 (matrix f) \<bullet> x))"
+  apply (rule ext)
+  apply (subst matrix_works[OF lf, symmetric])
+  apply (simp add: Cart_eq matrix_vector_mult_def vec1_def row_def dot_def mult_commute forall_1)
+  done
+
+lemma dest_vec1_eq_0: "dest_vec1 x = 0 \<longleftrightarrow> x = 0"
+  by (simp add: dest_vec1_eq[symmetric])
+
+lemma setsum_scalars: assumes fS: "finite S"
+  shows "setsum f S = vec1 (setsum (dest_vec1 o f) S)"
+  unfolding vec1_setsum[OF fS] by simp
+
+lemma dest_vec1_wlog_le: "(\<And>(x::'a::linorder ^ 1) y. P x y \<longleftrightarrow> P y x)  \<Longrightarrow> (\<And>x y. dest_vec1 x <= dest_vec1 y ==> P x y) \<Longrightarrow> P x y"
+  apply (cases "dest_vec1 x \<le> dest_vec1 y")
+  apply simp
+  apply (subgoal_tac "dest_vec1 y \<le> dest_vec1 x")
+  apply (auto)
+  done
+
+text{* Pasting vectors. *}
+
+lemma linear_fstcart: "linear fstcart"
+  by (auto simp add: linear_def Cart_eq)
+
+lemma linear_sndcart: "linear sndcart"
+  by (auto simp add: linear_def Cart_eq)
+
+lemma fstcart_vec[simp]: "fstcart(vec x) = vec x"
+  by (simp add: Cart_eq)
+
+lemma fstcart_add[simp]:"fstcart(x + y) = fstcart (x::'a::{plus,times}^('b + 'c)) + fstcart y"
+  by (simp add: Cart_eq)
+
+lemma fstcart_cmul[simp]:"fstcart(c*s x) = c*s fstcart (x::'a::{plus,times}^('b + 'c))"
+  by (simp add: Cart_eq)
+
+lemma fstcart_neg[simp]:"fstcart(- x) = - fstcart (x::'a::ring_1^('b + 'c))"
+  by (simp add: Cart_eq)
+
+lemma fstcart_sub[simp]:"fstcart(x - y) = fstcart (x::'a::ring_1^('b + 'c)) - fstcart y"
+  by (simp add: Cart_eq)
+
+lemma fstcart_setsum:
+  fixes f:: "'d \<Rightarrow> 'a::semiring_1^_"
+  assumes fS: "finite S"
+  shows "fstcart (setsum f S) = setsum (\<lambda>i. fstcart (f i)) S"
+  by (induct rule: finite_induct[OF fS], simp_all add: vec_0[symmetric] del: vec_0)
+
+lemma sndcart_vec[simp]: "sndcart(vec x) = vec x"
+  by (simp add: Cart_eq)
+
+lemma sndcart_add[simp]:"sndcart(x + y) = sndcart (x::'a::{plus,times}^('b + 'c)) + sndcart y"
+  by (simp add: Cart_eq)
+
+lemma sndcart_cmul[simp]:"sndcart(c*s x) = c*s sndcart (x::'a::{plus,times}^('b + 'c))"
+  by (simp add: Cart_eq)
+
+lemma sndcart_neg[simp]:"sndcart(- x) = - sndcart (x::'a::ring_1^('b + 'c))"
+  by (simp add: Cart_eq)
+
+lemma sndcart_sub[simp]:"sndcart(x - y) = sndcart (x::'a::ring_1^('b + 'c)) - sndcart y"
+  by (simp add: Cart_eq)
+
+lemma sndcart_setsum:
+  fixes f:: "'d \<Rightarrow> 'a::semiring_1^_"
+  assumes fS: "finite S"
+  shows "sndcart (setsum f S) = setsum (\<lambda>i. sndcart (f i)) S"
+  by (induct rule: finite_induct[OF fS], simp_all add: vec_0[symmetric] del: vec_0)
+
+lemma pastecart_vec[simp]: "pastecart (vec x) (vec x) = vec x"
+  by (simp add: pastecart_eq fstcart_pastecart sndcart_pastecart)
+
+lemma pastecart_add[simp]:"pastecart (x1::'a::{plus,times}^_) y1 + pastecart x2 y2 = pastecart (x1 + x2) (y1 + y2)"
+  by (simp add: pastecart_eq fstcart_pastecart sndcart_pastecart)
+
+lemma pastecart_cmul[simp]: "pastecart (c *s (x1::'a::{plus,times}^_)) (c *s y1) = c *s pastecart x1 y1"
+  by (simp add: pastecart_eq fstcart_pastecart sndcart_pastecart)
+
+lemma pastecart_neg[simp]: "pastecart (- (x::'a::ring_1^_)) (- y) = - pastecart x y"
+  unfolding vector_sneg_minus1 pastecart_cmul ..
+
+lemma pastecart_sub: "pastecart (x1::'a::ring_1^_) y1 - pastecart x2 y2 = pastecart (x1 - x2) (y1 - y2)"
+  by (simp add: diff_def pastecart_neg[symmetric] del: pastecart_neg)
+
+lemma pastecart_setsum:
+  fixes f:: "'d \<Rightarrow> 'a::semiring_1^_"
+  assumes fS: "finite S"
+  shows "pastecart (setsum f S) (setsum g S) = setsum (\<lambda>i. pastecart (f i) (g i)) S"
+  by (simp  add: pastecart_eq fstcart_setsum[OF fS] sndcart_setsum[OF fS] fstcart_pastecart sndcart_pastecart)
+
+lemma setsum_Plus:
+  "\<lbrakk>finite A; finite B\<rbrakk> \<Longrightarrow>
+    (\<Sum>x\<in>A <+> B. g x) = (\<Sum>x\<in>A. g (Inl x)) + (\<Sum>x\<in>B. g (Inr x))"
+  unfolding Plus_def
+  by (subst setsum_Un_disjoint, auto simp add: setsum_reindex)
+
+lemma setsum_UNIV_sum:
+  fixes g :: "'a::finite + 'b::finite \<Rightarrow> _"
+  shows "(\<Sum>x\<in>UNIV. g x) = (\<Sum>x\<in>UNIV. g (Inl x)) + (\<Sum>x\<in>UNIV. g (Inr x))"
+  apply (subst UNIV_Plus_UNIV [symmetric])
+  apply (rule setsum_Plus [OF finite finite])
+  done
+
+lemma norm_fstcart: "norm(fstcart x) <= norm (x::real ^('n::finite + 'm::finite))"
+proof-
+  have th0: "norm x = norm (pastecart (fstcart x) (sndcart x))"
+    by (simp add: pastecart_fst_snd)
+  have th1: "fstcart x \<bullet> fstcart x \<le> pastecart (fstcart x) (sndcart x) \<bullet> pastecart (fstcart x) (sndcart x)"
+    by (simp add: dot_def setsum_UNIV_sum pastecart_def setsum_nonneg)
+  then show ?thesis
+    unfolding th0
+    unfolding real_vector_norm_def real_sqrt_le_iff id_def
+    by (simp add: dot_def)
+qed
+
+lemma dist_fstcart: "dist(fstcart (x::real^_)) (fstcart y) <= dist x y"
+  unfolding dist_norm by (metis fstcart_sub[symmetric] norm_fstcart)
+
+lemma norm_sndcart: "norm(sndcart x) <= norm (x::real ^('n::finite + 'm::finite))"
+proof-
+  have th0: "norm x = norm (pastecart (fstcart x) (sndcart x))"
+    by (simp add: pastecart_fst_snd)
+  have th1: "sndcart x \<bullet> sndcart x \<le> pastecart (fstcart x) (sndcart x) \<bullet> pastecart (fstcart x) (sndcart x)"
+    by (simp add: dot_def setsum_UNIV_sum pastecart_def setsum_nonneg)
+  then show ?thesis
+    unfolding th0
+    unfolding real_vector_norm_def real_sqrt_le_iff id_def
+    by (simp add: dot_def)
+qed
+
+lemma dist_sndcart: "dist(sndcart (x::real^_)) (sndcart y) <= dist x y"
+  unfolding dist_norm by (metis sndcart_sub[symmetric] norm_sndcart)
+
+lemma dot_pastecart: "(pastecart (x1::'a::{times,comm_monoid_add}^'n::finite) (x2::'a::{times,comm_monoid_add}^'m::finite)) \<bullet> (pastecart y1 y2) =  x1 \<bullet> y1 + x2 \<bullet> y2"
+  by (simp add: dot_def setsum_UNIV_sum pastecart_def)
+
+text {* TODO: move to NthRoot *}
+lemma sqrt_add_le_add_sqrt:
+  assumes x: "0 \<le> x" and y: "0 \<le> y"
+  shows "sqrt (x + y) \<le> sqrt x + sqrt y"
+apply (rule power2_le_imp_le)
+apply (simp add: real_sum_squared_expand add_nonneg_nonneg x y)
+apply (simp add: mult_nonneg_nonneg x y)
+apply (simp add: add_nonneg_nonneg x y)
+done
+
+lemma norm_pastecart: "norm (pastecart x y) <= norm x + norm y"
+  unfolding norm_vector_def setL2_def setsum_UNIV_sum
+  by (simp add: sqrt_add_le_add_sqrt setsum_nonneg)
+
+subsection {* A generic notion of "hull" (convex, affine, conic hull and closure). *}
+
+definition hull :: "'a set set \<Rightarrow> 'a set \<Rightarrow> 'a set" (infixl "hull" 75) where
+  "S hull s = Inter {t. t \<in> S \<and> s \<subseteq> t}"
+
+lemma hull_same: "s \<in> S \<Longrightarrow> S hull s = s"
+  unfolding hull_def by auto
+
+lemma hull_in: "(\<And>T. T \<subseteq> S ==> Inter T \<in> S) ==> (S hull s) \<in> S"
+unfolding hull_def subset_iff by auto
+
+lemma hull_eq: "(\<And>T. T \<subseteq> S ==> Inter T \<in> S) ==> (S hull s) = s \<longleftrightarrow> s \<in> S"
+using hull_same[of s S] hull_in[of S s] by metis
+
+
+lemma hull_hull: "S hull (S hull s) = S hull s"
+  unfolding hull_def by blast
+
+lemma hull_subset: "s \<subseteq> (S hull s)"
+  unfolding hull_def by blast
+
+lemma hull_mono: " s \<subseteq> t ==> (S hull s) \<subseteq> (S hull t)"
+  unfolding hull_def by blast
+
+lemma hull_antimono: "S \<subseteq> T ==> (T hull s) \<subseteq> (S hull s)"
+  unfolding hull_def by blast
+
+lemma hull_minimal: "s \<subseteq> t \<Longrightarrow> t \<in> S ==> (S hull s) \<subseteq> t"
+  unfolding hull_def by blast
+
+lemma subset_hull: "t \<in> S ==> S hull s \<subseteq> t \<longleftrightarrow>  s \<subseteq> t"
+  unfolding hull_def by blast
+
+lemma hull_unique: "s \<subseteq> t \<Longrightarrow> t \<in> S \<Longrightarrow> (\<And>t'. s \<subseteq> t' \<Longrightarrow> t' \<in> S ==> t \<subseteq> t')
+           ==> (S hull s = t)"
+unfolding hull_def by auto
+
+lemma hull_induct: "(\<And>x. x\<in> S \<Longrightarrow> P x) \<Longrightarrow> Q {x. P x} \<Longrightarrow> \<forall>x\<in> Q hull S. P x"
+  using hull_minimal[of S "{x. P x}" Q]
+  by (auto simp add: subset_eq Collect_def mem_def)
+
+lemma hull_inc: "x \<in> S \<Longrightarrow> x \<in> P hull S" by (metis hull_subset subset_eq)
+
+lemma hull_union_subset: "(S hull s) \<union> (S hull t) \<subseteq> (S hull (s \<union> t))"
+unfolding Un_subset_iff by (metis hull_mono Un_upper1 Un_upper2)
+
+lemma hull_union: assumes T: "\<And>T. T \<subseteq> S ==> Inter T \<in> S"
+  shows "S hull (s \<union> t) = S hull (S hull s \<union> S hull t)"
+apply rule
+apply (rule hull_mono)
+unfolding Un_subset_iff
+apply (metis hull_subset Un_upper1 Un_upper2 subset_trans)
+apply (rule hull_minimal)
+apply (metis hull_union_subset)
+apply (metis hull_in T)
+done
+
+lemma hull_redundant_eq: "a \<in> (S hull s) \<longleftrightarrow> (S hull (insert a s) = S hull s)"
+  unfolding hull_def by blast
+
+lemma hull_redundant: "a \<in> (S hull s) ==> (S hull (insert a s) = S hull s)"
+by (metis hull_redundant_eq)
+
+text{* Archimedian properties and useful consequences. *}
+
+lemma real_arch_simple: "\<exists>n. x <= real (n::nat)"
+  using reals_Archimedean2[of x] apply auto by (rule_tac x="Suc n" in exI, auto)
+lemmas real_arch_lt = reals_Archimedean2
+
+lemmas real_arch = reals_Archimedean3
+
+lemma real_arch_inv: "0 < e \<longleftrightarrow> (\<exists>n::nat. n \<noteq> 0 \<and> 0 < inverse (real n) \<and> inverse (real n) < e)"
+  using reals_Archimedean
+  apply (auto simp add: field_simps inverse_positive_iff_positive)
+  apply (subgoal_tac "inverse (real n) > 0")
+  apply arith
+  apply simp
+  done
+
+lemma real_pow_lbound: "0 <= x ==> 1 + real n * x <= (1 + x) ^ n"
+proof(induct n)
+  case 0 thus ?case by simp
+next
+  case (Suc n)
+  hence h: "1 + real n * x \<le> (1 + x) ^ n" by simp
+  from h have p: "1 \<le> (1 + x) ^ n" using Suc.prems by simp
+  from h have "1 + real n * x + x \<le> (1 + x) ^ n + x" by simp
+  also have "\<dots> \<le> (1 + x) ^ Suc n" apply (subst diff_le_0_iff_le[symmetric])
+    apply (simp add: ring_simps)
+    using mult_left_mono[OF p Suc.prems] by simp
+  finally show ?case  by (simp add: real_of_nat_Suc ring_simps)
+qed
+
+lemma real_arch_pow: assumes x: "1 < (x::real)" shows "\<exists>n. y < x^n"
+proof-
+  from x have x0: "x - 1 > 0" by arith
+  from real_arch[OF x0, rule_format, of y]
+  obtain n::nat where n:"y < real n * (x - 1)" by metis
+  from x0 have x00: "x- 1 \<ge> 0" by arith
+  from real_pow_lbound[OF x00, of n] n
+  have "y < x^n" by auto
+  then show ?thesis by metis
+qed
+
+lemma real_arch_pow2: "\<exists>n. (x::real) < 2^ n"
+  using real_arch_pow[of 2 x] by simp
+
+lemma real_arch_pow_inv: assumes y: "(y::real) > 0" and x1: "x < 1"
+  shows "\<exists>n. x^n < y"
+proof-
+  {assume x0: "x > 0"
+    from x0 x1 have ix: "1 < 1/x" by (simp add: field_simps)
+    from real_arch_pow[OF ix, of "1/y"]
+    obtain n where n: "1/y < (1/x)^n" by blast
+    then
+    have ?thesis using y x0 by (auto simp add: field_simps power_divide) }
+  moreover
+  {assume "\<not> x > 0" with y x1 have ?thesis apply auto by (rule exI[where x=1], auto)}
+  ultimately show ?thesis by metis
+qed
+
+lemma forall_pos_mono: "(\<And>d e::real. d < e \<Longrightarrow> P d ==> P e) \<Longrightarrow> (\<And>n::nat. n \<noteq> 0 ==> P(inverse(real n))) \<Longrightarrow> (\<And>e. 0 < e ==> P e)"
+  by (metis real_arch_inv)
+
+lemma forall_pos_mono_1: "(\<And>d e::real. d < e \<Longrightarrow> P d ==> P e) \<Longrightarrow> (\<And>n. P(inverse(real (Suc n)))) ==> 0 < e ==> P e"
+  apply (rule forall_pos_mono)
+  apply auto
+  apply (atomize)
+  apply (erule_tac x="n - 1" in allE)
+  apply auto
+  done
+
+lemma real_archimedian_rdiv_eq_0: assumes x0: "x \<ge> 0" and c: "c \<ge> 0" and xc: "\<forall>(m::nat)>0. real m * x \<le> c"
+  shows "x = 0"
+proof-
+  {assume "x \<noteq> 0" with x0 have xp: "x > 0" by arith
+    from real_arch[OF xp, rule_format, of c] obtain n::nat where n: "c < real n * x"  by blast
+    with xc[rule_format, of n] have "n = 0" by arith
+    with n c have False by simp}
+  then show ?thesis by blast
+qed
+
+(* ------------------------------------------------------------------------- *)
+(* Relate max and min to sup and inf.                                        *)
+(* ------------------------------------------------------------------------- *)
+
+lemma real_max_rsup: "max x y = rsup {x,y}"
+proof-
+  have f: "finite {x, y}" "{x,y} \<noteq> {}"  by simp_all
+  from rsup_finite_le_iff[OF f, of "max x y"] have "rsup {x,y} \<le> max x y" by simp
+  moreover
+  have "max x y \<le> rsup {x,y}" using rsup_finite_ge_iff[OF f, of "max x y"]
+    by (simp add: linorder_linear)
+  ultimately show ?thesis by arith
+qed
+
+lemma real_min_rinf: "min x y = rinf {x,y}"
+proof-
+  have f: "finite {x, y}" "{x,y} \<noteq> {}"  by simp_all
+  from rinf_finite_le_iff[OF f, of "min x y"] have "rinf {x,y} \<le> min x y"
+    by (simp add: linorder_linear)
+  moreover
+  have "min x y \<le> rinf {x,y}" using rinf_finite_ge_iff[OF f, of "min x y"]
+    by simp
+  ultimately show ?thesis by arith
+qed
+
+(* ------------------------------------------------------------------------- *)
+(* Geometric progression.                                                    *)
+(* ------------------------------------------------------------------------- *)
+
+lemma sum_gp_basic: "((1::'a::{field}) - x) * setsum (\<lambda>i. x^i) {0 .. n} = (1 - x^(Suc n))"
+  (is "?lhs = ?rhs")
+proof-
+  {assume x1: "x = 1" hence ?thesis by simp}
+  moreover
+  {assume x1: "x\<noteq>1"
+    hence x1': "x - 1 \<noteq> 0" "1 - x \<noteq> 0" "x - 1 = - (1 - x)" "- (1 - x) \<noteq> 0" by auto
+    from geometric_sum[OF x1, of "Suc n", unfolded x1']
+    have "(- (1 - x)) * setsum (\<lambda>i. x^i) {0 .. n} = - (1 - x^(Suc n))"
+      unfolding atLeastLessThanSuc_atLeastAtMost
+      using x1' apply (auto simp only: field_simps)
+      apply (simp add: ring_simps)
+      done
+    then have ?thesis by (simp add: ring_simps) }
+  ultimately show ?thesis by metis
+qed
+
+lemma sum_gp_multiplied: assumes mn: "m <= n"
+  shows "((1::'a::{field}) - x) * setsum (op ^ x) {m..n} = x^m - x^ Suc n"
+  (is "?lhs = ?rhs")
+proof-
+  let ?S = "{0..(n - m)}"
+  from mn have mn': "n - m \<ge> 0" by arith
+  let ?f = "op + m"
+  have i: "inj_on ?f ?S" unfolding inj_on_def by auto
+  have f: "?f ` ?S = {m..n}"
+    using mn apply (auto simp add: image_iff Bex_def) by arith
+  have th: "op ^ x o op + m = (\<lambda>i. x^m * x^i)"
+    by (rule ext, simp add: power_add power_mult)
+  from setsum_reindex[OF i, of "op ^ x", unfolded f th setsum_right_distrib[symmetric]]
+  have "?lhs = x^m * ((1 - x) * setsum (op ^ x) {0..n - m})" by simp
+  then show ?thesis unfolding sum_gp_basic using mn
+    by (simp add: ring_simps power_add[symmetric])
+qed
+
+lemma sum_gp: "setsum (op ^ (x::'a::{field})) {m .. n} =
+   (if n < m then 0 else if x = 1 then of_nat ((n + 1) - m)
+                    else (x^ m - x^ (Suc n)) / (1 - x))"
+proof-
+  {assume nm: "n < m" hence ?thesis by simp}
+  moreover
+  {assume "\<not> n < m" hence nm: "m \<le> n" by arith
+    {assume x: "x = 1"  hence ?thesis by simp}
+    moreover
+    {assume x: "x \<noteq> 1" hence nz: "1 - x \<noteq> 0" by simp
+      from sum_gp_multiplied[OF nm, of x] nz have ?thesis by (simp add: field_simps)}
+    ultimately have ?thesis by metis
+  }
+  ultimately show ?thesis by metis
+qed
+
+lemma sum_gp_offset: "setsum (op ^ (x::'a::{field})) {m .. m+n} =
+  (if x = 1 then of_nat n + 1 else x^m * (1 - x^Suc n) / (1 - x))"
+  unfolding sum_gp[of x m "m + n"] power_Suc
+  by (simp add: ring_simps power_add)
+
+
+subsection{* A bit of linear algebra. *}
+
+definition "subspace S \<longleftrightarrow> 0 \<in> S \<and> (\<forall>x\<in> S. \<forall>y \<in>S. x + y \<in> S) \<and> (\<forall>c. \<forall>x \<in>S. c *s x \<in>S )"
+definition "span S = (subspace hull S)"
+definition "dependent S \<longleftrightarrow> (\<exists>a \<in> S. a \<in> span(S - {a}))"
+abbreviation "independent s == ~(dependent s)"
+
+(* Closure properties of subspaces.                                          *)
+
+lemma subspace_UNIV[simp]: "subspace(UNIV)" by (simp add: subspace_def)
+
+lemma subspace_0: "subspace S ==> 0 \<in> S" by (metis subspace_def)
+
+lemma subspace_add: "subspace S \<Longrightarrow> x \<in> S \<Longrightarrow> y \<in> S ==> x + y \<in> S"
+  by (metis subspace_def)
+
+lemma subspace_mul: "subspace S \<Longrightarrow> x \<in> S \<Longrightarrow> c *s x \<in> S"
+  by (metis subspace_def)
+
+lemma subspace_neg: "subspace S \<Longrightarrow> (x::'a::ring_1^'n) \<in> S \<Longrightarrow> - x \<in> S"
+  by (metis vector_sneg_minus1 subspace_mul)
+
+lemma subspace_sub: "subspace S \<Longrightarrow> (x::'a::ring_1^'n) \<in> S \<Longrightarrow> y \<in> S \<Longrightarrow> x - y \<in> S"
+  by (metis diff_def subspace_add subspace_neg)
+
+lemma subspace_setsum:
+  assumes sA: "subspace A" and fB: "finite B"
+  and f: "\<forall>x\<in> B. f x \<in> A"
+  shows "setsum f B \<in> A"
+  using  fB f sA
+  apply(induct rule: finite_induct[OF fB])
+  by (simp add: subspace_def sA, auto simp add: sA subspace_add)
+
+lemma subspace_linear_image:
+  assumes lf: "linear (f::'a::semiring_1^'n \<Rightarrow> _)" and sS: "subspace S"
+  shows "subspace(f ` S)"
+  using lf sS linear_0[OF lf]
+  unfolding linear_def subspace_def
+  apply (auto simp add: image_iff)
+  apply (rule_tac x="x + y" in bexI, auto)
+  apply (rule_tac x="c*s x" in bexI, auto)
+  done
+
+lemma subspace_linear_preimage: "linear (f::'a::semiring_1^'n \<Rightarrow> _) ==> subspace S ==> subspace {x. f x \<in> S}"
+  by (auto simp add: subspace_def linear_def linear_0[of f])
+
+lemma subspace_trivial: "subspace {0::'a::semiring_1 ^_}"
+  by (simp add: subspace_def)
+
+lemma subspace_inter: "subspace A \<Longrightarrow> subspace B ==> subspace (A \<inter> B)"
+  by (simp add: subspace_def)
+
+
+lemma span_mono: "A \<subseteq> B ==> span A \<subseteq> span B"
+  by (metis span_def hull_mono)
+
+lemma subspace_span: "subspace(span S)"
+  unfolding span_def
+  apply (rule hull_in[unfolded mem_def])
+  apply (simp only: subspace_def Inter_iff Int_iff subset_eq)
+  apply auto
+  apply (erule_tac x="X" in ballE)
+  apply (simp add: mem_def)
+  apply blast
+  apply (erule_tac x="X" in ballE)
+  apply (erule_tac x="X" in ballE)
+  apply (erule_tac x="X" in ballE)
+  apply (clarsimp simp add: mem_def)
+  apply simp
+  apply simp
+  apply simp
+  apply (erule_tac x="X" in ballE)
+  apply (erule_tac x="X" in ballE)
+  apply (simp add: mem_def)
+  apply simp
+  apply simp
+  done
+
+lemma span_clauses:
+  "a \<in> S ==> a \<in> span S"
+  "0 \<in> span S"
+  "x\<in> span S \<Longrightarrow> y \<in> span S ==> x + y \<in> span S"
+  "x \<in> span S \<Longrightarrow> c *s x \<in> span S"
+  by (metis span_def hull_subset subset_eq subspace_span subspace_def)+
+
+lemma span_induct: assumes SP: "\<And>x. x \<in> S ==> P x"
+  and P: "subspace P" and x: "x \<in> span S" shows "P x"
+proof-
+  from SP have SP': "S \<subseteq> P" by (simp add: mem_def subset_eq)
+  from P have P': "P \<in> subspace" by (simp add: mem_def)
+  from x hull_minimal[OF SP' P', unfolded span_def[symmetric]]
+  show "P x" by (metis mem_def subset_eq)
+qed
+
+lemma span_empty: "span {} = {(0::'a::semiring_0 ^ 'n)}"
+  apply (simp add: span_def)
+  apply (rule hull_unique)
+  apply (auto simp add: mem_def subspace_def)
+  unfolding mem_def[of "0::'a^'n", symmetric]
+  apply simp
+  done
+
+lemma independent_empty: "independent {}"
+  by (simp add: dependent_def)
+
+lemma independent_mono: "independent A \<Longrightarrow> B \<subseteq> A ==> independent B"
+  apply (clarsimp simp add: dependent_def span_mono)
+  apply (subgoal_tac "span (B - {a}) \<le> span (A - {a})")
+  apply force
+  apply (rule span_mono)
+  apply auto
+  done
+
+lemma span_subspace: "A \<subseteq> B \<Longrightarrow> B \<le> span A \<Longrightarrow>  subspace B \<Longrightarrow> span A = B"
+  by (metis order_antisym span_def hull_minimal mem_def)
+
+lemma span_induct': assumes SP: "\<forall>x \<in> S. P x"
+  and P: "subspace P" shows "\<forall>x \<in> span S. P x"
+  using span_induct SP P by blast
+
+inductive span_induct_alt_help for S:: "'a::semiring_1^'n \<Rightarrow> bool"
+  where
+  span_induct_alt_help_0: "span_induct_alt_help S 0"
+  | span_induct_alt_help_S: "x \<in> S \<Longrightarrow> span_induct_alt_help S z \<Longrightarrow> span_induct_alt_help S (c *s x + z)"
+
+lemma span_induct_alt':
+  assumes h0: "h (0::'a::semiring_1^'n)" and hS: "\<And>c x y. x \<in> S \<Longrightarrow> h y \<Longrightarrow> h (c*s x + y)" shows "\<forall>x \<in> span S. h x"
+proof-
+  {fix x:: "'a^'n" assume x: "span_induct_alt_help S x"
+    have "h x"
+      apply (rule span_induct_alt_help.induct[OF x])
+      apply (rule h0)
+      apply (rule hS, assumption, assumption)
+      done}
+  note th0 = this
+  {fix x assume x: "x \<in> span S"
+
+    have "span_induct_alt_help S x"
+      proof(rule span_induct[where x=x and S=S])
+        show "x \<in> span S" using x .
+      next
+        fix x assume xS : "x \<in> S"
+          from span_induct_alt_help_S[OF xS span_induct_alt_help_0, of 1]
+          show "span_induct_alt_help S x" by simp
+        next
+        have "span_induct_alt_help S 0" by (rule span_induct_alt_help_0)
+        moreover
+        {fix x y assume h: "span_induct_alt_help S x" "span_induct_alt_help S y"
+          from h
+          have "span_induct_alt_help S (x + y)"
+            apply (induct rule: span_induct_alt_help.induct)
+            apply simp
+            unfolding add_assoc
+            apply (rule span_induct_alt_help_S)
+            apply assumption
+            apply simp
+            done}
+        moreover
+        {fix c x assume xt: "span_induct_alt_help S x"
+          then have "span_induct_alt_help S (c*s x)"
+            apply (induct rule: span_induct_alt_help.induct)
+            apply (simp add: span_induct_alt_help_0)
+            apply (simp add: vector_smult_assoc vector_add_ldistrib)
+            apply (rule span_induct_alt_help_S)
+            apply assumption
+            apply simp
+            done
+        }
+        ultimately show "subspace (span_induct_alt_help S)"
+          unfolding subspace_def mem_def Ball_def by blast
+      qed}
+  with th0 show ?thesis by blast
+qed
+
+lemma span_induct_alt:
+  assumes h0: "h (0::'a::semiring_1^'n)" and hS: "\<And>c x y. x \<in> S \<Longrightarrow> h y \<Longrightarrow> h (c*s x + y)" and x: "x \<in> span S"
+  shows "h x"
+using span_induct_alt'[of h S] h0 hS x by blast
+
+(* Individual closure properties. *)
+
+lemma span_superset: "x \<in> S ==> x \<in> span S" by (metis span_clauses)
+
+lemma span_0: "0 \<in> span S" by (metis subspace_span subspace_0)
+
+lemma span_add: "x \<in> span S \<Longrightarrow> y \<in> span S ==> x + y \<in> span S"
+  by (metis subspace_add subspace_span)
+
+lemma span_mul: "x \<in> span S ==> (c *s x) \<in> span S"
+  by (metis subspace_span subspace_mul)
+
+lemma span_neg: "x \<in> span S ==> - (x::'a::ring_1^'n) \<in> span S"
+  by (metis subspace_neg subspace_span)
+
+lemma span_sub: "(x::'a::ring_1^'n) \<in> span S \<Longrightarrow> y \<in> span S ==> x - y \<in> span S"
+  by (metis subspace_span subspace_sub)
+
+lemma span_setsum: "finite A \<Longrightarrow> \<forall>x \<in> A. f x \<in> span S ==> setsum f A \<in> span S"
+  apply (rule subspace_setsum)
+  by (metis subspace_span subspace_setsum)+
+
+lemma span_add_eq: "(x::'a::ring_1^'n) \<in> span S \<Longrightarrow> x + y \<in> span S \<longleftrightarrow> y \<in> span S"
+  apply (auto simp only: span_add span_sub)
+  apply (subgoal_tac "(x + y) - x \<in> span S", simp)
+  by (simp only: span_add span_sub)
+
+(* Mapping under linear image. *)
+
+lemma span_linear_image: assumes lf: "linear (f::'a::semiring_1 ^ 'n => _)"
+  shows "span (f ` S) = f ` (span S)"
+proof-
+  {fix x
+    assume x: "x \<in> span (f ` S)"
+    have "x \<in> f ` span S"
+      apply (rule span_induct[where x=x and S = "f ` S"])
+      apply (clarsimp simp add: image_iff)
+      apply (frule span_superset)
+      apply blast
+      apply (simp only: mem_def)
+      apply (rule subspace_linear_image[OF lf])
+      apply (rule subspace_span)
+      apply (rule x)
+      done}
+  moreover
+  {fix x assume x: "x \<in> span S"
+    have th0:"(\<lambda>a. f a \<in> span (f ` S)) = {x. f x \<in> span (f ` S)}" apply (rule set_ext)
+      unfolding mem_def Collect_def ..
+    have "f x \<in> span (f ` S)"
+      apply (rule span_induct[where S=S])
+      apply (rule span_superset)
+      apply simp
+      apply (subst th0)
+      apply (rule subspace_linear_preimage[OF lf subspace_span, of "f ` S"])
+      apply (rule x)
+      done}
+  ultimately show ?thesis by blast
+qed
+
+(* The key breakdown property. *)
+
+lemma span_breakdown:
+  assumes bS: "(b::'a::ring_1 ^ 'n) \<in> S" and aS: "a \<in> span S"
+  shows "\<exists>k. a - k*s b \<in> span (S - {b})" (is "?P a")
+proof-
+  {fix x assume xS: "x \<in> S"
+    {assume ab: "x = b"
+      then have "?P x"
+        apply simp
+        apply (rule exI[where x="1"], simp)
+        by (rule span_0)}
+    moreover
+    {assume ab: "x \<noteq> b"
+      then have "?P x"  using xS
+        apply -
+        apply (rule exI[where x=0])
+        apply (rule span_superset)
+        by simp}
+    ultimately have "?P x" by blast}
+  moreover have "subspace ?P"
+    unfolding subspace_def
+    apply auto
+    apply (simp add: mem_def)
+    apply (rule exI[where x=0])
+    using span_0[of "S - {b}"]
+    apply (simp add: mem_def)
+    apply (clarsimp simp add: mem_def)
+    apply (rule_tac x="k + ka" in exI)
+    apply (subgoal_tac "x + y - (k + ka) *s b = (x - k*s b) + (y - ka *s b)")
+    apply (simp only: )
+    apply (rule span_add[unfolded mem_def])
+    apply assumption+
+    apply (vector ring_simps)
+    apply (clarsimp simp add: mem_def)
+    apply (rule_tac x= "c*k" in exI)
+    apply (subgoal_tac "c *s x - (c * k) *s b = c*s (x - k*s b)")
+    apply (simp only: )
+    apply (rule span_mul[unfolded mem_def])
+    apply assumption
+    by (vector ring_simps)
+  ultimately show "?P a" using aS span_induct[where S=S and P= "?P"] by metis
+qed
+
+lemma span_breakdown_eq:
+  "(x::'a::ring_1^'n) \<in> span (insert a S) \<longleftrightarrow> (\<exists>k. (x - k *s a) \<in> span S)" (is "?lhs \<longleftrightarrow> ?rhs")
+proof-
+  {assume x: "x \<in> span (insert a S)"
+    from x span_breakdown[of "a" "insert a S" "x"]
+    have ?rhs apply clarsimp
+      apply (rule_tac x= "k" in exI)
+      apply (rule set_rev_mp[of _ "span (S - {a})" _])
+      apply assumption
+      apply (rule span_mono)
+      apply blast
+      done}
+  moreover
+  { fix k assume k: "x - k *s a \<in> span S"
+    have eq: "x = (x - k *s a) + k *s a" by vector
+    have "(x - k *s a) + k *s a \<in> span (insert a S)"
+      apply (rule span_add)
+      apply (rule set_rev_mp[of _ "span S" _])
+      apply (rule k)
+      apply (rule span_mono)
+      apply blast
+      apply (rule span_mul)
+      apply (rule span_superset)
+      apply blast
+      done
+    then have ?lhs using eq by metis}
+  ultimately show ?thesis by blast
+qed
+
+(* Hence some "reversal" results.*)
+
+lemma in_span_insert:
+  assumes a: "(a::'a::field^'n) \<in> span (insert b S)" and na: "a \<notin> span S"
+  shows "b \<in> span (insert a S)"
+proof-
+  from span_breakdown[of b "insert b S" a, OF insertI1 a]
+  obtain k where k: "a - k*s b \<in> span (S - {b})" by auto
+  {assume k0: "k = 0"
+    with k have "a \<in> span S"
+      apply (simp)
+      apply (rule set_rev_mp)
+      apply assumption
+      apply (rule span_mono)
+      apply blast
+      done
+    with na  have ?thesis by blast}
+  moreover
+  {assume k0: "k \<noteq> 0"
+    have eq: "b = (1/k) *s a - ((1/k) *s a - b)" by vector
+    from k0 have eq': "(1/k) *s (a - k*s b) = (1/k) *s a - b"
+      by (vector field_simps)
+    from k have "(1/k) *s (a - k*s b) \<in> span (S - {b})"
+      by (rule span_mul)
+    hence th: "(1/k) *s a - b \<in> span (S - {b})"
+      unfolding eq' .
+
+    from k
+    have ?thesis
+      apply (subst eq)
+      apply (rule span_sub)
+      apply (rule span_mul)
+      apply (rule span_superset)
+      apply blast
+      apply (rule set_rev_mp)
+      apply (rule th)
+      apply (rule span_mono)
+      using na by blast}
+  ultimately show ?thesis by blast
+qed
+
+lemma in_span_delete:
+  assumes a: "(a::'a::field^'n) \<in> span S"
+  and na: "a \<notin> span (S-{b})"
+  shows "b \<in> span (insert a (S - {b}))"
+  apply (rule in_span_insert)
+  apply (rule set_rev_mp)
+  apply (rule a)
+  apply (rule span_mono)
+  apply blast
+  apply (rule na)
+  done
+
+(* Transitivity property. *)
+
+lemma span_trans:
+  assumes x: "(x::'a::ring_1^'n) \<in> span S" and y: "y \<in> span (insert x S)"
+  shows "y \<in> span S"
+proof-
+  from span_breakdown[of x "insert x S" y, OF insertI1 y]
+  obtain k where k: "y -k*s x \<in> span (S - {x})" by auto
+  have eq: "y = (y - k *s x) + k *s x" by vector
+  show ?thesis
+    apply (subst eq)
+    apply (rule span_add)
+    apply (rule set_rev_mp)
+    apply (rule k)
+    apply (rule span_mono)
+    apply blast
+    apply (rule span_mul)
+    by (rule x)
+qed
+
+(* ------------------------------------------------------------------------- *)
+(* An explicit expansion is sometimes needed.                                *)
+(* ------------------------------------------------------------------------- *)
+
+lemma span_explicit:
+  "span P = {y::'a::semiring_1^'n. \<exists>S u. finite S \<and> S \<subseteq> P \<and> setsum (\<lambda>v. u v *s v) S = y}"
+  (is "_ = ?E" is "_ = {y. ?h y}" is "_ = {y. \<exists>S u. ?Q S u y}")
+proof-
+  {fix x assume x: "x \<in> ?E"
+    then obtain S u where fS: "finite S" and SP: "S\<subseteq>P" and u: "setsum (\<lambda>v. u v *s v) S = x"
+      by blast
+    have "x \<in> span P"
+      unfolding u[symmetric]
+      apply (rule span_setsum[OF fS])
+      using span_mono[OF SP]
+      by (auto intro: span_superset span_mul)}
+  moreover
+  have "\<forall>x \<in> span P. x \<in> ?E"
+    unfolding mem_def Collect_def
+  proof(rule span_induct_alt')
+    show "?h 0"
+      apply (rule exI[where x="{}"]) by simp
+  next
+    fix c x y
+    assume x: "x \<in> P" and hy: "?h y"
+    from hy obtain S u where fS: "finite S" and SP: "S\<subseteq>P"
+      and u: "setsum (\<lambda>v. u v *s v) S = y" by blast
+    let ?S = "insert x S"
+    let ?u = "\<lambda>y. if y = x then (if x \<in> S then u y + c else c)
+                  else u y"
+    from fS SP x have th0: "finite (insert x S)" "insert x S \<subseteq> P" by blast+
+    {assume xS: "x \<in> S"
+      have S1: "S = (S - {x}) \<union> {x}"
+        and Sss:"finite (S - {x})" "finite {x}" "(S -{x}) \<inter> {x} = {}" using xS fS by auto
+      have "setsum (\<lambda>v. ?u v *s v) ?S =(\<Sum>v\<in>S - {x}. u v *s v) + (u x + c) *s x"
+        using xS
+        by (simp add: setsum_Un_disjoint[OF Sss, unfolded S1[symmetric]]
+          setsum_clauses(2)[OF fS] cong del: if_weak_cong)
+      also have "\<dots> = (\<Sum>v\<in>S. u v *s v) + c *s x"
+        apply (simp add: setsum_Un_disjoint[OF Sss, unfolded S1[symmetric]])
+        by (vector ring_simps)
+      also have "\<dots> = c*s x + y"
+        by (simp add: add_commute u)
+      finally have "setsum (\<lambda>v. ?u v *s v) ?S = c*s x + y" .
+    then have "?Q ?S ?u (c*s x + y)" using th0 by blast}
+  moreover
+  {assume xS: "x \<notin> S"
+    have th00: "(\<Sum>v\<in>S. (if v = x then c else u v) *s v) = y"
+      unfolding u[symmetric]
+      apply (rule setsum_cong2)
+      using xS by auto
+    have "?Q ?S ?u (c*s x + y)" using fS xS th0
+      by (simp add: th00 setsum_clauses add_commute cong del: if_weak_cong)}
+  ultimately have "?Q ?S ?u (c*s x + y)"
+    by (cases "x \<in> S", simp, simp)
+    then show "?h (c*s x + y)"
+      apply -
+      apply (rule exI[where x="?S"])
+      apply (rule exI[where x="?u"]) by metis
+  qed
+  ultimately show ?thesis by blast
+qed
+
+lemma dependent_explicit:
+  "dependent P \<longleftrightarrow> (\<exists>S u. finite S \<and> S \<subseteq> P \<and> (\<exists>(v::'a::{idom,field}^'n) \<in>S. u v \<noteq> 0 \<and> setsum (\<lambda>v. u v *s v) S = 0))" (is "?lhs = ?rhs")
+proof-
+  {assume dP: "dependent P"
+    then obtain a S u where aP: "a \<in> P" and fS: "finite S"
+      and SP: "S \<subseteq> P - {a}" and ua: "setsum (\<lambda>v. u v *s v) S = a"
+      unfolding dependent_def span_explicit by blast
+    let ?S = "insert a S"
+    let ?u = "\<lambda>y. if y = a then - 1 else u y"
+    let ?v = a
+    from aP SP have aS: "a \<notin> S" by blast
+    from fS SP aP have th0: "finite ?S" "?S \<subseteq> P" "?v \<in> ?S" "?u ?v \<noteq> 0" by auto
+    have s0: "setsum (\<lambda>v. ?u v *s v) ?S = 0"
+      using fS aS
+      apply (simp add: vector_smult_lneg vector_smult_lid setsum_clauses ring_simps )
+      apply (subst (2) ua[symmetric])
+      apply (rule setsum_cong2)
+      by auto
+    with th0 have ?rhs
+      apply -
+      apply (rule exI[where x= "?S"])
+      apply (rule exI[where x= "?u"])
+      by clarsimp}
+  moreover
+  {fix S u v assume fS: "finite S"
+      and SP: "S \<subseteq> P" and vS: "v \<in> S" and uv: "u v \<noteq> 0"
+    and u: "setsum (\<lambda>v. u v *s v) S = 0"
+    let ?a = v
+    let ?S = "S - {v}"
+    let ?u = "\<lambda>i. (- u i) / u v"
+    have th0: "?a \<in> P" "finite ?S" "?S \<subseteq> P"       using fS SP vS by auto
+    have "setsum (\<lambda>v. ?u v *s v) ?S = setsum (\<lambda>v. (- (inverse (u ?a))) *s (u v *s v)) S - ?u v *s v"
+      using fS vS uv
+      by (simp add: setsum_diff1 vector_smult_lneg divide_inverse
+        vector_smult_assoc field_simps)
+    also have "\<dots> = ?a"
+      unfolding setsum_cmul u
+      using uv by (simp add: vector_smult_lneg)
+    finally  have "setsum (\<lambda>v. ?u v *s v) ?S = ?a" .
+    with th0 have ?lhs
+      unfolding dependent_def span_explicit
+      apply -
+      apply (rule bexI[where x= "?a"])
+      apply simp_all
+      apply (rule exI[where x= "?S"])
+      by auto}
+  ultimately show ?thesis by blast
+qed
+
+
+lemma span_finite:
+  assumes fS: "finite S"
+  shows "span S = {(y::'a::semiring_1^'n). \<exists>u. setsum (\<lambda>v. u v *s v) S = y}"
+  (is "_ = ?rhs")
+proof-
+  {fix y assume y: "y \<in> span S"
+    from y obtain S' u where fS': "finite S'" and SS': "S' \<subseteq> S" and
+      u: "setsum (\<lambda>v. u v *s v) S' = y" unfolding span_explicit by blast
+    let ?u = "\<lambda>x. if x \<in> S' then u x else 0"
+    from setsum_restrict_set[OF fS, of "\<lambda>v. u v *s v" S', symmetric] SS'
+    have "setsum (\<lambda>v. ?u v *s v) S = setsum (\<lambda>v. u v *s v) S'"
+      unfolding cond_value_iff cond_application_beta
+      by (simp add: cond_value_iff inf_absorb2 cong del: if_weak_cong)
+    hence "setsum (\<lambda>v. ?u v *s v) S = y" by (metis u)
+    hence "y \<in> ?rhs" by auto}
+  moreover
+  {fix y u assume u: "setsum (\<lambda>v. u v *s v) S = y"
+    then have "y \<in> span S" using fS unfolding span_explicit by auto}
+  ultimately show ?thesis by blast
+qed
+
+
+(* Standard bases are a spanning set, and obviously finite.                  *)
+
+lemma span_stdbasis:"span {basis i :: 'a::ring_1^'n::finite | i. i \<in> (UNIV :: 'n set)} = UNIV"
+apply (rule set_ext)
+apply auto
+apply (subst basis_expansion[symmetric])
+apply (rule span_setsum)
+apply simp
+apply auto
+apply (rule span_mul)
+apply (rule span_superset)
+apply (auto simp add: Collect_def mem_def)
+done
+
+lemma has_size_stdbasis: "{basis i ::real ^'n::finite | i. i \<in> (UNIV :: 'n set)} hassize CARD('n)" (is "?S hassize ?n")
+proof-
+  have eq: "?S = basis ` UNIV" by blast
+  show ?thesis unfolding eq
+    apply (rule hassize_image_inj[OF basis_inj])
+    by (simp add: hassize_def)
+qed
+
+lemma finite_stdbasis: "finite {basis i ::real^'n::finite |i. i\<in> (UNIV:: 'n set)}"
+  using has_size_stdbasis[unfolded hassize_def]
+  ..
+
+lemma card_stdbasis: "card {basis i ::real^'n::finite |i. i\<in> (UNIV :: 'n set)} = CARD('n)"
+  using has_size_stdbasis[unfolded hassize_def]
+  ..
+
+lemma independent_stdbasis_lemma:
+  assumes x: "(x::'a::semiring_1 ^ 'n) \<in> span (basis ` S)"
+  and iS: "i \<notin> S"
+  shows "(x$i) = 0"
+proof-
+  let ?U = "UNIV :: 'n set"
+  let ?B = "basis ` S"
+  let ?P = "\<lambda>(x::'a^'n). \<forall>i\<in> ?U. i \<notin> S \<longrightarrow> x$i =0"
+ {fix x::"'a^'n" assume xS: "x\<in> ?B"
+   from xS have "?P x" by auto}
+ moreover
+ have "subspace ?P"
+   by (auto simp add: subspace_def Collect_def mem_def)
+ ultimately show ?thesis
+   using x span_induct[of ?B ?P x] iS by blast
+qed
+
+lemma independent_stdbasis: "independent {basis i ::real^'n::finite |i. i\<in> (UNIV :: 'n set)}"
+proof-
+  let ?I = "UNIV :: 'n set"
+  let ?b = "basis :: _ \<Rightarrow> real ^'n"
+  let ?B = "?b ` ?I"
+  have eq: "{?b i|i. i \<in> ?I} = ?B"
+    by auto
+  {assume d: "dependent ?B"
+    then obtain k where k: "k \<in> ?I" "?b k \<in> span (?B - {?b k})"
+      unfolding dependent_def by auto
+    have eq1: "?B - {?b k} = ?B - ?b ` {k}"  by simp
+    have eq2: "?B - {?b k} = ?b ` (?I - {k})"
+      unfolding eq1
+      apply (rule inj_on_image_set_diff[symmetric])
+      apply (rule basis_inj) using k(1) by auto
+    from k(2) have th0: "?b k \<in> span (?b ` (?I - {k}))" unfolding eq2 .
+    from independent_stdbasis_lemma[OF th0, of k, simplified]
+    have False by simp}
+  then show ?thesis unfolding eq dependent_def ..
+qed
+
+(* This is useful for building a basis step-by-step.                         *)
+
+lemma independent_insert:
+  "independent(insert (a::'a::field ^'n) S) \<longleftrightarrow>
+      (if a \<in> S then independent S
+                else independent S \<and> a \<notin> span S)" (is "?lhs \<longleftrightarrow> ?rhs")
+proof-
+  {assume aS: "a \<in> S"
+    hence ?thesis using insert_absorb[OF aS] by simp}
+  moreover
+  {assume aS: "a \<notin> S"
+    {assume i: ?lhs
+      then have ?rhs using aS
+        apply simp
+        apply (rule conjI)
+        apply (rule independent_mono)
+        apply assumption
+        apply blast
+        by (simp add: dependent_def)}
+    moreover
+    {assume i: ?rhs
+      have ?lhs using i aS
+        apply simp
+        apply (auto simp add: dependent_def)
+        apply (case_tac "aa = a", auto)
+        apply (subgoal_tac "insert a S - {aa} = insert a (S - {aa})")
+        apply simp
+        apply (subgoal_tac "a \<in> span (insert aa (S - {aa}))")
+        apply (subgoal_tac "insert aa (S - {aa}) = S")
+        apply simp
+        apply blast
+        apply (rule in_span_insert)
+        apply assumption
+        apply blast
+        apply blast
+        done}
+    ultimately have ?thesis by blast}
+  ultimately show ?thesis by blast
+qed
+
+(* The degenerate case of the Exchange Lemma.  *)
+
+lemma mem_delete: "x \<in> (A - {a}) \<longleftrightarrow> x \<noteq> a \<and> x \<in> A"
+  by blast
+
+lemma span_span: "span (span A) = span A"
+  unfolding span_def hull_hull ..
+
+lemma span_inc: "S \<subseteq> span S"
+  by (metis subset_eq span_superset)
+
+lemma spanning_subset_independent:
+  assumes BA: "B \<subseteq> A" and iA: "independent (A::('a::field ^'n) set)"
+  and AsB: "A \<subseteq> span B"
+  shows "A = B"
+proof
+  from BA show "B \<subseteq> A" .
+next
+  from span_mono[OF BA] span_mono[OF AsB]
+  have sAB: "span A = span B" unfolding span_span by blast
+
+  {fix x assume x: "x \<in> A"
+    from iA have th0: "x \<notin> span (A - {x})"
+      unfolding dependent_def using x by blast
+    from x have xsA: "x \<in> span A" by (blast intro: span_superset)
+    have "A - {x} \<subseteq> A" by blast
+    hence th1:"span (A - {x}) \<subseteq> span A" by (metis span_mono)
+    {assume xB: "x \<notin> B"
+      from xB BA have "B \<subseteq> A -{x}" by blast
+      hence "span B \<subseteq> span (A - {x})" by (metis span_mono)
+      with th1 th0 sAB have "x \<notin> span A" by blast
+      with x have False by (metis span_superset)}
+    then have "x \<in> B" by blast}
+  then show "A \<subseteq> B" by blast
+qed
+
+(* The general case of the Exchange Lemma, the key to what follows.  *)
+
+lemma exchange_lemma:
+  assumes f:"finite (t:: ('a::field^'n) set)" and i: "independent s"
+  and sp:"s \<subseteq> span t"
+  shows "\<exists>t'. (t' hassize card t) \<and> s \<subseteq> t' \<and> t' \<subseteq> s \<union> t \<and> s \<subseteq> span t'"
+using f i sp
+proof(induct c\<equiv>"card(t - s)" arbitrary: s t rule: nat_less_induct)
+  fix n:: nat and s t :: "('a ^'n) set"
+  assume H: " \<forall>m<n. \<forall>(x:: ('a ^'n) set) xa.
+                finite xa \<longrightarrow>
+                independent x \<longrightarrow>
+                x \<subseteq> span xa \<longrightarrow>
+                m = card (xa - x) \<longrightarrow>
+                (\<exists>t'. (t' hassize card xa) \<and>
+                      x \<subseteq> t' \<and> t' \<subseteq> x \<union> xa \<and> x \<subseteq> span t')"
+    and ft: "finite t" and s: "independent s" and sp: "s \<subseteq> span t"
+    and n: "n = card (t - s)"
+  let ?P = "\<lambda>t'. (t' hassize card t) \<and> s \<subseteq> t' \<and> t' \<subseteq> s \<union> t \<and> s \<subseteq> span t'"
+  let ?ths = "\<exists>t'. ?P t'"
+  {assume st: "s \<subseteq> t"
+    from st ft span_mono[OF st] have ?ths apply - apply (rule exI[where x=t])
+      by (auto simp add: hassize_def intro: span_superset)}
+  moreover
+  {assume st: "t \<subseteq> s"
+
+    from spanning_subset_independent[OF st s sp]
+      st ft span_mono[OF st] have ?ths apply - apply (rule exI[where x=t])
+      by (auto simp add: hassize_def intro: span_superset)}
+  moreover
+  {assume st: "\<not> s \<subseteq> t" "\<not> t \<subseteq> s"
+    from st(2) obtain b where b: "b \<in> t" "b \<notin> s" by blast
+      from b have "t - {b} - s \<subset> t - s" by blast
+      then have cardlt: "card (t - {b} - s) < n" using n ft
+        by (auto intro: psubset_card_mono)
+      from b ft have ct0: "card t \<noteq> 0" by auto
+    {assume stb: "s \<subseteq> span(t -{b})"
+      from ft have ftb: "finite (t -{b})" by auto
+      from H[rule_format, OF cardlt ftb s stb]
+      obtain u where u: "u hassize card (t-{b})" "s \<subseteq> u" "u \<subseteq> s \<union> (t - {b})" "s \<subseteq> span u" by blast
+      let ?w = "insert b u"
+      have th0: "s \<subseteq> insert b u" using u by blast
+      from u(3) b have "u \<subseteq> s \<union> t" by blast
+      then have th1: "insert b u \<subseteq> s \<union> t" using u b by blast
+      have bu: "b \<notin> u" using b u by blast
+      from u(1) have fu: "finite u" by (simp add: hassize_def)
+      from u(1) ft b have "u hassize (card t - 1)" by auto
+      then
+      have th2: "insert b u hassize card t"
+        using  card_insert_disjoint[OF fu bu] ct0 by (auto simp add: hassize_def)
+      from u(4) have "s \<subseteq> span u" .
+      also have "\<dots> \<subseteq> span (insert b u)" apply (rule span_mono) by blast
+      finally have th3: "s \<subseteq> span (insert b u)" .      from th0 th1 th2 th3 have th: "?P ?w"  by blast
+      from th have ?ths by blast}
+    moreover
+    {assume stb: "\<not> s \<subseteq> span(t -{b})"
+      from stb obtain a where a: "a \<in> s" "a \<notin> span (t - {b})" by blast
+      have ab: "a \<noteq> b" using a b by blast
+      have at: "a \<notin> t" using a ab span_superset[of a "t- {b}"] by auto
+      have mlt: "card ((insert a (t - {b})) - s) < n"
+        using cardlt ft n  a b by auto
+      have ft': "finite (insert a (t - {b}))" using ft by auto
+      {fix x assume xs: "x \<in> s"
+        have t: "t \<subseteq> (insert b (insert a (t -{b})))" using b by auto
+        from b(1) have "b \<in> span t" by (simp add: span_superset)
+        have bs: "b \<in> span (insert a (t - {b}))"
+          by (metis in_span_delete a sp mem_def subset_eq)
+        from xs sp have "x \<in> span t" by blast
+        with span_mono[OF t]
+        have x: "x \<in> span (insert b (insert a (t - {b})))" ..
+        from span_trans[OF bs x] have "x \<in> span (insert a (t - {b}))"  .}
+      then have sp': "s \<subseteq> span (insert a (t - {b}))" by blast
+
+      from H[rule_format, OF mlt ft' s sp' refl] obtain u where
+        u: "u hassize card (insert a (t -{b}))" "s \<subseteq> u" "u \<subseteq> s \<union> insert a (t -{b})"
+        "s \<subseteq> span u" by blast
+      from u a b ft at ct0 have "?P u" by (auto simp add: hassize_def)
+      then have ?ths by blast }
+    ultimately have ?ths by blast
+  }
+  ultimately
+  show ?ths  by blast
+qed
+
+(* This implies corresponding size bounds.                                   *)
+
+lemma independent_span_bound:
+  assumes f: "finite t" and i: "independent (s::('a::field^'n) set)" and sp:"s \<subseteq> span t"
+  shows "finite s \<and> card s \<le> card t"
+  by (metis exchange_lemma[OF f i sp] hassize_def finite_subset card_mono)
+
+
+lemma finite_Atleast_Atmost_nat[simp]: "finite {f x |x. x\<in> (UNIV::'a::finite set)}"
+proof-
+  have eq: "{f x |x. x\<in> UNIV} = f ` UNIV" by auto
+  show ?thesis unfolding eq
+    apply (rule finite_imageI)
+    apply (rule finite)
+    done
+qed
+
+
+lemma independent_bound:
+  fixes S:: "(real^'n::finite) set"
+  shows "independent S \<Longrightarrow> finite S \<and> card S <= CARD('n)"
+  apply (subst card_stdbasis[symmetric])
+  apply (rule independent_span_bound)
+  apply (rule finite_Atleast_Atmost_nat)
+  apply assumption
+  unfolding span_stdbasis
+  apply (rule subset_UNIV)
+  done
+
+lemma dependent_biggerset: "(finite (S::(real ^'n::finite) set) ==> card S > CARD('n)) ==> dependent S"
+  by (metis independent_bound not_less)
+
+(* Hence we can create a maximal independent subset.                         *)
+
+lemma maximal_independent_subset_extend:
+  assumes sv: "(S::(real^'n::finite) set) \<subseteq> V" and iS: "independent S"
+  shows "\<exists>B. S \<subseteq> B \<and> B \<subseteq> V \<and> independent B \<and> V \<subseteq> span B"
+  using sv iS
+proof(induct d\<equiv> "CARD('n) - card S" arbitrary: S rule: nat_less_induct)
+  fix n and S:: "(real^'n) set"
+  assume H: "\<forall>m<n. \<forall>S \<subseteq> V. independent S \<longrightarrow> m = CARD('n) - card S \<longrightarrow>
+              (\<exists>B. S \<subseteq> B \<and> B \<subseteq> V \<and> independent B \<and> V \<subseteq> span B)"
+    and sv: "S \<subseteq> V" and i: "independent S" and n: "n = CARD('n) - card S"
+  let ?P = "\<lambda>B. S \<subseteq> B \<and> B \<subseteq> V \<and> independent B \<and> V \<subseteq> span B"
+  let ?ths = "\<exists>x. ?P x"
+  let ?d = "CARD('n)"
+  {assume "V \<subseteq> span S"
+    then have ?ths  using sv i by blast }
+  moreover
+  {assume VS: "\<not> V \<subseteq> span S"
+    from VS obtain a where a: "a \<in> V" "a \<notin> span S" by blast
+    from a have aS: "a \<notin> S" by (auto simp add: span_superset)
+    have th0: "insert a S \<subseteq> V" using a sv by blast
+    from independent_insert[of a S]  i a
+    have th1: "independent (insert a S)" by auto
+    have mlt: "?d - card (insert a S) < n"
+      using aS a n independent_bound[OF th1]
+      by auto
+
+    from H[rule_format, OF mlt th0 th1 refl]
+    obtain B where B: "insert a S \<subseteq> B" "B \<subseteq> V" "independent B" " V \<subseteq> span B"
+      by blast
+    from B have "?P B" by auto
+    then have ?ths by blast}
+  ultimately show ?ths by blast
+qed
+
+lemma maximal_independent_subset:
+  "\<exists>(B:: (real ^'n::finite) set). B\<subseteq> V \<and> independent B \<and> V \<subseteq> span B"
+  by (metis maximal_independent_subset_extend[of "{}:: (real ^'n) set"] empty_subsetI independent_empty)
+
+(* Notion of dimension.                                                      *)
+
+definition "dim V = (SOME n. \<exists>B. B \<subseteq> V \<and> independent B \<and> V \<subseteq> span B \<and> (B hassize n))"
+
+lemma basis_exists:  "\<exists>B. (B :: (real ^'n::finite) set) \<subseteq> V \<and> independent B \<and> V \<subseteq> span B \<and> (B hassize dim V)"
+unfolding dim_def some_eq_ex[of "\<lambda>n. \<exists>B. B \<subseteq> V \<and> independent B \<and> V \<subseteq> span B \<and> (B hassize n)"]
+unfolding hassize_def
+using maximal_independent_subset[of V] independent_bound
+by auto
+
+(* Consequences of independence or spanning for cardinality.                 *)
+
+lemma independent_card_le_dim: "(B::(real ^'n::finite) set) \<subseteq> V \<Longrightarrow> independent B \<Longrightarrow> finite B \<and> card B \<le> dim V"
+by (metis basis_exists[of V] independent_span_bound[where ?'a=real] hassize_def subset_trans)
+
+lemma span_card_ge_dim:  "(B::(real ^'n::finite) set) \<subseteq> V \<Longrightarrow> V \<subseteq> span B \<Longrightarrow> finite B \<Longrightarrow> dim V \<le> card B"
+  by (metis basis_exists[of V] independent_span_bound hassize_def subset_trans)
+
+lemma basis_card_eq_dim:
+  "B \<subseteq> (V:: (real ^'n::finite) set) \<Longrightarrow> V \<subseteq> span B \<Longrightarrow> independent B \<Longrightarrow> finite B \<and> card B = dim V"
+  by (metis order_eq_iff independent_card_le_dim span_card_ge_dim independent_mono)
+
+lemma dim_unique: "(B::(real ^'n::finite) set) \<subseteq> V \<Longrightarrow> V \<subseteq> span B \<Longrightarrow> independent B \<Longrightarrow> B hassize n \<Longrightarrow> dim V = n"
+  by (metis basis_card_eq_dim hassize_def)
+
+(* More lemmas about dimension.                                              *)
+
+lemma dim_univ: "dim (UNIV :: (real^'n::finite) set) = CARD('n)"
+  apply (rule dim_unique[of "{basis i |i. i\<in> (UNIV :: 'n set)}"])
+  by (auto simp only: span_stdbasis has_size_stdbasis independent_stdbasis)
+
+lemma dim_subset:
+  "(S:: (real ^'n::finite) set) \<subseteq> T \<Longrightarrow> dim S \<le> dim T"
+  using basis_exists[of T] basis_exists[of S]
+  by (metis independent_span_bound[where ?'a = real and ?'n = 'n] subset_eq hassize_def)
+
+lemma dim_subset_univ: "dim (S:: (real^'n::finite) set) \<le> CARD('n)"
+  by (metis dim_subset subset_UNIV dim_univ)
+
+(* Converses to those.                                                       *)
+
+lemma card_ge_dim_independent:
+  assumes BV:"(B::(real ^'n::finite) set) \<subseteq> V" and iB:"independent B" and dVB:"dim V \<le> card B"
+  shows "V \<subseteq> span B"
+proof-
+  {fix a assume aV: "a \<in> V"
+    {assume aB: "a \<notin> span B"
+      then have iaB: "independent (insert a B)" using iB aV  BV by (simp add: independent_insert)
+      from aV BV have th0: "insert a B \<subseteq> V" by blast
+      from aB have "a \<notin>B" by (auto simp add: span_superset)
+      with independent_card_le_dim[OF th0 iaB] dVB  have False by auto}
+    then have "a \<in> span B"  by blast}
+  then show ?thesis by blast
+qed
+
+lemma card_le_dim_spanning:
+  assumes BV: "(B:: (real ^'n::finite) set) \<subseteq> V" and VB: "V \<subseteq> span B"
+  and fB: "finite B" and dVB: "dim V \<ge> card B"
+  shows "independent B"
+proof-
+  {fix a assume a: "a \<in> B" "a \<in> span (B -{a})"
+    from a fB have c0: "card B \<noteq> 0" by auto
+    from a fB have cb: "card (B -{a}) = card B - 1" by auto
+    from BV a have th0: "B -{a} \<subseteq> V" by blast
+    {fix x assume x: "x \<in> V"
+      from a have eq: "insert a (B -{a}) = B" by blast
+      from x VB have x': "x \<in> span B" by blast
+      from span_trans[OF a(2), unfolded eq, OF x']
+      have "x \<in> span (B -{a})" . }
+    then have th1: "V \<subseteq> span (B -{a})" by blast
+    have th2: "finite (B -{a})" using fB by auto
+    from span_card_ge_dim[OF th0 th1 th2]
+    have c: "dim V \<le> card (B -{a})" .
+    from c c0 dVB cb have False by simp}
+  then show ?thesis unfolding dependent_def by blast
+qed
+
+lemma card_eq_dim: "(B:: (real ^'n::finite) set) \<subseteq> V \<Longrightarrow> B hassize dim V \<Longrightarrow> independent B \<longleftrightarrow> V \<subseteq> span B"
+  by (metis hassize_def order_eq_iff card_le_dim_spanning
+    card_ge_dim_independent)
+
+(* ------------------------------------------------------------------------- *)
+(* More general size bound lemmas.                                           *)
+(* ------------------------------------------------------------------------- *)
+
+lemma independent_bound_general:
+  "independent (S:: (real^'n::finite) set) \<Longrightarrow> finite S \<and> card S \<le> dim S"
+  by (metis independent_card_le_dim independent_bound subset_refl)
+
+lemma dependent_biggerset_general: "(finite (S:: (real^'n::finite) set) \<Longrightarrow> card S > dim S) \<Longrightarrow> dependent S"
+  using independent_bound_general[of S] by (metis linorder_not_le)
+
+lemma dim_span: "dim (span (S:: (real ^'n::finite) set)) = dim S"
+proof-
+  have th0: "dim S \<le> dim (span S)"
+    by (auto simp add: subset_eq intro: dim_subset span_superset)
+  from basis_exists[of S]
+  obtain B where B: "B \<subseteq> S" "independent B" "S \<subseteq> span B" "B hassize dim S" by blast
+  from B have fB: "finite B" "card B = dim S" unfolding hassize_def by blast+
+  have bSS: "B \<subseteq> span S" using B(1) by (metis subset_eq span_inc)
+  have sssB: "span S \<subseteq> span B" using span_mono[OF B(3)] by (simp add: span_span)
+  from span_card_ge_dim[OF bSS sssB fB(1)] th0 show ?thesis
+    using fB(2)  by arith
+qed
+
+lemma subset_le_dim: "(S:: (real ^'n::finite) set) \<subseteq> span T \<Longrightarrow> dim S \<le> dim T"
+  by (metis dim_span dim_subset)
+
+lemma span_eq_dim: "span (S:: (real ^'n::finite) set) = span T ==> dim S = dim T"
+  by (metis dim_span)
+
+lemma spans_image:
+  assumes lf: "linear (f::'a::semiring_1^'n \<Rightarrow> _)" and VB: "V \<subseteq> span B"
+  shows "f ` V \<subseteq> span (f ` B)"
+  unfolding span_linear_image[OF lf]
+  by (metis VB image_mono)
+
+lemma dim_image_le:
+  fixes f :: "real^'n::finite \<Rightarrow> real^'m::finite"
+  assumes lf: "linear f" shows "dim (f ` S) \<le> dim (S:: (real ^'n::finite) set)"
+proof-
+  from basis_exists[of S] obtain B where
+    B: "B \<subseteq> S" "independent B" "S \<subseteq> span B" "B hassize dim S" by blast
+  from B have fB: "finite B" "card B = dim S" unfolding hassize_def by blast+
+  have "dim (f ` S) \<le> card (f ` B)"
+    apply (rule span_card_ge_dim)
+    using lf B fB by (auto simp add: span_linear_image spans_image subset_image_iff)
+  also have "\<dots> \<le> dim S" using card_image_le[OF fB(1)] fB by simp
+  finally show ?thesis .
+qed
+
+(* Relation between bases and injectivity/surjectivity of map.               *)
+
+lemma spanning_surjective_image:
+  assumes us: "UNIV \<subseteq> span (S:: ('a::semiring_1 ^'n) set)"
+  and lf: "linear f" and sf: "surj f"
+  shows "UNIV \<subseteq> span (f ` S)"
+proof-
+  have "UNIV \<subseteq> f ` UNIV" using sf by (auto simp add: surj_def)
+  also have " \<dots> \<subseteq> span (f ` S)" using spans_image[OF lf us] .
+finally show ?thesis .
+qed
+
+lemma independent_injective_image:
+  assumes iS: "independent (S::('a::semiring_1^'n) set)" and lf: "linear f" and fi: "inj f"
+  shows "independent (f ` S)"
+proof-
+  {fix a assume a: "a \<in> S" "f a \<in> span (f ` S - {f a})"
+    have eq: "f ` S - {f a} = f ` (S - {a})" using fi
+      by (auto simp add: inj_on_def)
+    from a have "f a \<in> f ` span (S -{a})"
+      unfolding eq span_linear_image[OF lf, of "S - {a}"]  by blast
+    hence "a \<in> span (S -{a})" using fi by (auto simp add: inj_on_def)
+    with a(1) iS  have False by (simp add: dependent_def) }
+  then show ?thesis unfolding dependent_def by blast
+qed
+
+(* ------------------------------------------------------------------------- *)
+(* Picking an orthogonal replacement for a spanning set.                     *)
+(* ------------------------------------------------------------------------- *)
+    (* FIXME : Move to some general theory ?*)
+definition "pairwise R S \<longleftrightarrow> (\<forall>x \<in> S. \<forall>y\<in> S. x\<noteq>y \<longrightarrow> R x y)"
+
+lemma vector_sub_project_orthogonal: "(b::'a::ordered_field^'n::finite) \<bullet> (x - ((b \<bullet> x) / (b\<bullet>b)) *s b) = 0"
+  apply (cases "b = 0", simp)
+  apply (simp add: dot_rsub dot_rmult)
+  unfolding times_divide_eq_right[symmetric]
+  by (simp add: field_simps dot_eq_0)
+
+lemma basis_orthogonal:
+  fixes B :: "(real ^'n::finite) set"
+  assumes fB: "finite B"
+  shows "\<exists>C. finite C \<and> card C \<le> card B \<and> span C = span B \<and> pairwise orthogonal C"
+  (is " \<exists>C. ?P B C")
+proof(induct rule: finite_induct[OF fB])
+  case 1 thus ?case apply (rule exI[where x="{}"]) by (auto simp add: pairwise_def)
+next
+  case (2 a B)
+  note fB = `finite B` and aB = `a \<notin> B`
+  from `\<exists>C. finite C \<and> card C \<le> card B \<and> span C = span B \<and> pairwise orthogonal C`
+  obtain C where C: "finite C" "card C \<le> card B"
+    "span C = span B" "pairwise orthogonal C" by blast
+  let ?a = "a - setsum (\<lambda>x. (x\<bullet>a / (x\<bullet>x)) *s x) C"
+  let ?C = "insert ?a C"
+  from C(1) have fC: "finite ?C" by simp
+  from fB aB C(1,2) have cC: "card ?C \<le> card (insert a B)" by (simp add: card_insert_if)
+  {fix x k
+    have th0: "\<And>(a::'b::comm_ring) b c. a - (b - c) = c + (a - b)" by (simp add: ring_simps)
+    have "x - k *s (a - (\<Sum>x\<in>C. (x \<bullet> a / (x \<bullet> x)) *s x)) \<in> span C \<longleftrightarrow> x - k *s a \<in> span C"
+      apply (simp only: vector_ssub_ldistrib th0)
+      apply (rule span_add_eq)
+      apply (rule span_mul)
+      apply (rule span_setsum[OF C(1)])
+      apply clarify
+      apply (rule span_mul)
+      by (rule span_superset)}
+  then have SC: "span ?C = span (insert a B)"
+    unfolding expand_set_eq span_breakdown_eq C(3)[symmetric] by auto
+  thm pairwise_def
+  {fix x y assume xC: "x \<in> ?C" and yC: "y \<in> ?C" and xy: "x \<noteq> y"
+    {assume xa: "x = ?a" and ya: "y = ?a"
+      have "orthogonal x y" using xa ya xy by blast}
+    moreover
+    {assume xa: "x = ?a" and ya: "y \<noteq> ?a" "y \<in> C"
+      from ya have Cy: "C = insert y (C - {y})" by blast
+      have fth: "finite (C - {y})" using C by simp
+      have "orthogonal x y"
+        using xa ya
+        unfolding orthogonal_def xa dot_lsub dot_rsub diff_eq_0_iff_eq
+        apply simp
+        apply (subst Cy)
+        using C(1) fth
+        apply (simp only: setsum_clauses)
+        thm dot_ladd
+        apply (auto simp add: dot_ladd dot_radd dot_lmult dot_rmult dot_eq_0 dot_sym[of y a] dot_lsum[OF fth])
+        apply (rule setsum_0')
+        apply clarsimp
+        apply (rule C(4)[unfolded pairwise_def orthogonal_def, rule_format])
+        by auto}
+    moreover
+    {assume xa: "x \<noteq> ?a" "x \<in> C" and ya: "y = ?a"
+      from xa have Cx: "C = insert x (C - {x})" by blast
+      have fth: "finite (C - {x})" using C by simp
+      have "orthogonal x y"
+        using xa ya
+        unfolding orthogonal_def ya dot_rsub dot_lsub diff_eq_0_iff_eq
+        apply simp
+        apply (subst Cx)
+        using C(1) fth
+        apply (simp only: setsum_clauses)
+        apply (subst dot_sym[of x])
+        apply (auto simp add: dot_radd dot_rmult dot_eq_0 dot_sym[of x a] dot_rsum[OF fth])
+        apply (rule setsum_0')
+        apply clarsimp
+        apply (rule C(4)[unfolded pairwise_def orthogonal_def, rule_format])
+        by auto}
+    moreover
+    {assume xa: "x \<in> C" and ya: "y \<in> C"
+      have "orthogonal x y" using xa ya xy C(4) unfolding pairwise_def by blast}
+    ultimately have "orthogonal x y" using xC yC by blast}
+  then have CPO: "pairwise orthogonal ?C" unfolding pairwise_def by blast
+  from fC cC SC CPO have "?P (insert a B) ?C" by blast
+  then show ?case by blast
+qed
+
+lemma orthogonal_basis_exists:
+  fixes V :: "(real ^'n::finite) set"
+  shows "\<exists>B. independent B \<and> B \<subseteq> span V \<and> V \<subseteq> span B \<and> (B hassize dim V) \<and> pairwise orthogonal B"
+proof-
+  from basis_exists[of V] obtain B where B: "B \<subseteq> V" "independent B" "V \<subseteq> span B" "B hassize dim V" by blast
+  from B have fB: "finite B" "card B = dim V" by (simp_all add: hassize_def)
+  from basis_orthogonal[OF fB(1)] obtain C where
+    C: "finite C" "card C \<le> card B" "span C = span B" "pairwise orthogonal C" by blast
+  from C B
+  have CSV: "C \<subseteq> span V" by (metis span_inc span_mono subset_trans)
+  from span_mono[OF B(3)]  C have SVC: "span V \<subseteq> span C" by (simp add: span_span)
+  from card_le_dim_spanning[OF CSV SVC C(1)] C(2,3) fB
+  have iC: "independent C" by (simp add: dim_span)
+  from C fB have "card C \<le> dim V" by simp
+  moreover have "dim V \<le> card C" using span_card_ge_dim[OF CSV SVC C(1)]
+    by (simp add: dim_span)
+  ultimately have CdV: "C hassize dim V" unfolding hassize_def using C(1) by simp
+  from C B CSV CdV iC show ?thesis by auto
+qed
+
+lemma span_eq: "span S = span T \<longleftrightarrow> S \<subseteq> span T \<and> T \<subseteq> span S"
+  by (metis set_eq_subset span_mono span_span span_inc) (* FIXME: slow *)
+
+(* ------------------------------------------------------------------------- *)
+(* Low-dimensional subset is in a hyperplane (weak orthogonal complement).   *)
+(* ------------------------------------------------------------------------- *)
+
+lemma span_not_univ_orthogonal:
+  assumes sU: "span S \<noteq> UNIV"
+  shows "\<exists>(a:: real ^'n::finite). a \<noteq>0 \<and> (\<forall>x \<in> span S. a \<bullet> x = 0)"
+proof-
+  from sU obtain a where a: "a \<notin> span S" by blast
+  from orthogonal_basis_exists obtain B where
+    B: "independent B" "B \<subseteq> span S" "S \<subseteq> span B" "B hassize dim S" "pairwise orthogonal B"
+    by blast
+  from B have fB: "finite B" "card B = dim S" by (simp_all add: hassize_def)
+  from span_mono[OF B(2)] span_mono[OF B(3)]
+  have sSB: "span S = span B" by (simp add: span_span)
+  let ?a = "a - setsum (\<lambda>b. (a\<bullet>b / (b\<bullet>b)) *s b) B"
+  have "setsum (\<lambda>b. (a\<bullet>b / (b\<bullet>b)) *s b) B \<in> span S"
+    unfolding sSB
+    apply (rule span_setsum[OF fB(1)])
+    apply clarsimp
+    apply (rule span_mul)
+    by (rule span_superset)
+  with a have a0:"?a  \<noteq> 0" by auto
+  have "\<forall>x\<in>span B. ?a \<bullet> x = 0"
+  proof(rule span_induct')
+    show "subspace (\<lambda>x. ?a \<bullet> x = 0)"
+      by (auto simp add: subspace_def mem_def dot_radd dot_rmult)
+  next
+    {fix x assume x: "x \<in> B"
+      from x have B': "B = insert x (B - {x})" by blast
+      have fth: "finite (B - {x})" using fB by simp
+      have "?a \<bullet> x = 0"
+        apply (subst B') using fB fth
+        unfolding setsum_clauses(2)[OF fth]
+        apply simp
+        apply (clarsimp simp add: dot_lsub dot_ladd dot_lmult dot_lsum dot_eq_0)
+        apply (rule setsum_0', rule ballI)
+        unfolding dot_sym
+        by (auto simp add: x field_simps dot_eq_0 intro: B(5)[unfolded pairwise_def orthogonal_def, rule_format])}
+    then show "\<forall>x \<in> B. ?a \<bullet> x = 0" by blast
+  qed
+  with a0 show ?thesis unfolding sSB by (auto intro: exI[where x="?a"])
+qed
+
+lemma span_not_univ_subset_hyperplane:
+  assumes SU: "span S \<noteq> (UNIV ::(real^'n::finite) set)"
+  shows "\<exists> a. a \<noteq>0 \<and> span S \<subseteq> {x. a \<bullet> x = 0}"
+  using span_not_univ_orthogonal[OF SU] by auto
+
+lemma lowdim_subset_hyperplane:
+  assumes d: "dim S < CARD('n::finite)"
+  shows "\<exists>(a::real ^'n::finite). a  \<noteq> 0 \<and> span S \<subseteq> {x. a \<bullet> x = 0}"
+proof-
+  {assume "span S = UNIV"
+    hence "dim (span S) = dim (UNIV :: (real ^'n) set)" by simp
+    hence "dim S = CARD('n)" by (simp add: dim_span dim_univ)
+    with d have False by arith}
+  hence th: "span S \<noteq> UNIV" by blast
+  from span_not_univ_subset_hyperplane[OF th] show ?thesis .
+qed
+
+(* We can extend a linear basis-basis injection to the whole set.            *)
+
+lemma linear_indep_image_lemma:
+  assumes lf: "linear f" and fB: "finite B"
+  and ifB: "independent (f ` B)"
+  and fi: "inj_on f B" and xsB: "x \<in> span B"
+  and fx: "f (x::'a::field^'n) = 0"
+  shows "x = 0"
+  using fB ifB fi xsB fx
+proof(induct arbitrary: x rule: finite_induct[OF fB])
+  case 1 thus ?case by (auto simp add:  span_empty)
+next
+  case (2 a b x)
+  have fb: "finite b" using "2.prems" by simp
+  have th0: "f ` b \<subseteq> f ` (insert a b)"
+    apply (rule image_mono) by blast
+  from independent_mono[ OF "2.prems"(2) th0]
+  have ifb: "independent (f ` b)"  .
+  have fib: "inj_on f b"
+    apply (rule subset_inj_on [OF "2.prems"(3)])
+    by blast
+  from span_breakdown[of a "insert a b", simplified, OF "2.prems"(4)]
+  obtain k where k: "x - k*s a \<in> span (b -{a})" by blast
+  have "f (x - k*s a) \<in> span (f ` b)"
+    unfolding span_linear_image[OF lf]
+    apply (rule imageI)
+    using k span_mono[of "b-{a}" b] by blast
+  hence "f x - k*s f a \<in> span (f ` b)"
+    by (simp add: linear_sub[OF lf] linear_cmul[OF lf])
+  hence th: "-k *s f a \<in> span (f ` b)"
+    using "2.prems"(5) by (simp add: vector_smult_lneg)
+  {assume k0: "k = 0"
+    from k0 k have "x \<in> span (b -{a})" by simp
+    then have "x \<in> span b" using span_mono[of "b-{a}" b]
+      by blast}
+  moreover
+  {assume k0: "k \<noteq> 0"
+    from span_mul[OF th, of "- 1/ k"] k0
+    have th1: "f a \<in> span (f ` b)"
+      by (auto simp add: vector_smult_assoc)
+    from inj_on_image_set_diff[OF "2.prems"(3), of "insert a b " "{a}", symmetric]
+    have tha: "f ` insert a b - f ` {a} = f ` (insert a b - {a})" by blast
+    from "2.prems"(2)[unfolded dependent_def bex_simps(10), rule_format, of "f a"]
+    have "f a \<notin> span (f ` b)" using tha
+      using "2.hyps"(2)
+      "2.prems"(3) by auto
+    with th1 have False by blast
+    then have "x \<in> span b" by blast}
+  ultimately have xsb: "x \<in> span b" by blast
+  from "2.hyps"(3)[OF fb ifb fib xsb "2.prems"(5)]
+  show "x = 0" .
+qed
+
+(* We can extend a linear mapping from basis.                                *)
+
+lemma linear_independent_extend_lemma:
+  assumes fi: "finite B" and ib: "independent B"
+  shows "\<exists>g. (\<forall>x\<in> span B. \<forall>y\<in> span B. g ((x::'a::field^'n) + y) = g x + g y)
+           \<and> (\<forall>x\<in> span B. \<forall>c. g (c*s x) = c *s g x)
+           \<and> (\<forall>x\<in> B. g x = f x)"
+using ib fi
+proof(induct rule: finite_induct[OF fi])
+  case 1 thus ?case by (auto simp add: span_empty)
+next
+  case (2 a b)
+  from "2.prems" "2.hyps" have ibf: "independent b" "finite b"
+    by (simp_all add: independent_insert)
+  from "2.hyps"(3)[OF ibf] obtain g where
+    g: "\<forall>x\<in>span b. \<forall>y\<in>span b. g (x + y) = g x + g y"
+    "\<forall>x\<in>span b. \<forall>c. g (c *s x) = c *s g x" "\<forall>x\<in>b. g x = f x" by blast
+  let ?h = "\<lambda>z. SOME k. (z - k *s a) \<in> span b"
+  {fix z assume z: "z \<in> span (insert a b)"
+    have th0: "z - ?h z *s a \<in> span b"
+      apply (rule someI_ex)
+      unfolding span_breakdown_eq[symmetric]
+      using z .
+    {fix k assume k: "z - k *s a \<in> span b"
+      have eq: "z - ?h z *s a - (z - k*s a) = (k - ?h z) *s a"
+        by (simp add: ring_simps vector_sadd_rdistrib[symmetric])
+      from span_sub[OF th0 k]
+      have khz: "(k - ?h z) *s a \<in> span b" by (simp add: eq)
+      {assume "k \<noteq> ?h z" hence k0: "k - ?h z \<noteq> 0" by simp
+        from k0 span_mul[OF khz, of "1 /(k - ?h z)"]
+        have "a \<in> span b" by (simp add: vector_smult_assoc)
+        with "2.prems"(1) "2.hyps"(2) have False
+          by (auto simp add: dependent_def)}
+      then have "k = ?h z" by blast}
+    with th0 have "z - ?h z *s a \<in> span b \<and> (\<forall>k. z - k *s a \<in> span b \<longrightarrow> k = ?h z)" by blast}
+  note h = this
+  let ?g = "\<lambda>z. ?h z *s f a + g (z - ?h z *s a)"
+  {fix x y assume x: "x \<in> span (insert a b)" and y: "y \<in> span (insert a b)"
+    have tha: "\<And>(x::'a^'n) y a k l. (x + y) - (k + l) *s a = (x - k *s a) + (y - l *s a)"
+      by (vector ring_simps)
+    have addh: "?h (x + y) = ?h x + ?h y"
+      apply (rule conjunct2[OF h, rule_format, symmetric])
+      apply (rule span_add[OF x y])
+      unfolding tha
+      by (metis span_add x y conjunct1[OF h, rule_format])
+    have "?g (x + y) = ?g x + ?g y"
+      unfolding addh tha
+      g(1)[rule_format,OF conjunct1[OF h, OF x] conjunct1[OF h, OF y]]
+      by (simp add: vector_sadd_rdistrib)}
+  moreover
+  {fix x:: "'a^'n" and c:: 'a  assume x: "x \<in> span (insert a b)"
+    have tha: "\<And>(x::'a^'n) c k a. c *s x - (c * k) *s a = c *s (x - k *s a)"
+      by (vector ring_simps)
+    have hc: "?h (c *s x) = c * ?h x"
+      apply (rule conjunct2[OF h, rule_format, symmetric])
+      apply (metis span_mul x)
+      by (metis tha span_mul x conjunct1[OF h])
+    have "?g (c *s x) = c*s ?g x"
+      unfolding hc tha g(2)[rule_format, OF conjunct1[OF h, OF x]]
+      by (vector ring_simps)}
+  moreover
+  {fix x assume x: "x \<in> (insert a b)"
+    {assume xa: "x = a"
+      have ha1: "1 = ?h a"
+        apply (rule conjunct2[OF h, rule_format])
+        apply (metis span_superset insertI1)
+        using conjunct1[OF h, OF span_superset, OF insertI1]
+        by (auto simp add: span_0)
+
+      from xa ha1[symmetric] have "?g x = f x"
+        apply simp
+        using g(2)[rule_format, OF span_0, of 0]
+        by simp}
+    moreover
+    {assume xb: "x \<in> b"
+      have h0: "0 = ?h x"
+        apply (rule conjunct2[OF h, rule_format])
+        apply (metis  span_superset insertI1 xb x)
+        apply simp
+        apply (metis span_superset xb)
+        done
+      have "?g x = f x"
+        by (simp add: h0[symmetric] g(3)[rule_format, OF xb])}
+    ultimately have "?g x = f x" using x by blast }
+  ultimately show ?case apply - apply (rule exI[where x="?g"]) by blast
+qed
+
+lemma linear_independent_extend:
+  assumes iB: "independent (B:: (real ^'n::finite) set)"
+  shows "\<exists>g. linear g \<and> (\<forall>x\<in>B. g x = f x)"
+proof-
+  from maximal_independent_subset_extend[of B UNIV] iB
+  obtain C where C: "B \<subseteq> C" "independent C" "\<And>x. x \<in> span C" by auto
+
+  from C(2) independent_bound[of C] linear_independent_extend_lemma[of C f]
+  obtain g where g: "(\<forall>x\<in> span C. \<forall>y\<in> span C. g (x + y) = g x + g y)
+           \<and> (\<forall>x\<in> span C. \<forall>c. g (c*s x) = c *s g x)
+           \<and> (\<forall>x\<in> C. g x = f x)" by blast
+  from g show ?thesis unfolding linear_def using C
+    apply clarsimp by blast
+qed
+
+(* Can construct an isomorphism between spaces of same dimension.            *)
+
+lemma card_le_inj: assumes fA: "finite A" and fB: "finite B"
+  and c: "card A \<le> card B" shows "(\<exists>f. f ` A \<subseteq> B \<and> inj_on f A)"
+using fB c
+proof(induct arbitrary: B rule: finite_induct[OF fA])
+  case 1 thus ?case by simp
+next
+  case (2 x s t)
+  thus ?case
+  proof(induct rule: finite_induct[OF "2.prems"(1)])
+    case 1    then show ?case by simp
+  next
+    case (2 y t)
+    from "2.prems"(1,2,5) "2.hyps"(1,2) have cst:"card s \<le> card t" by simp
+    from "2.prems"(3) [OF "2.hyps"(1) cst] obtain f where
+      f: "f ` s \<subseteq> t \<and> inj_on f s" by blast
+    from f "2.prems"(2) "2.hyps"(2) show ?case
+      apply -
+      apply (rule exI[where x = "\<lambda>z. if z = x then y else f z"])
+      by (auto simp add: inj_on_def)
+  qed
+qed
+
+lemma card_subset_eq: assumes fB: "finite B" and AB: "A \<subseteq> B" and
+  c: "card A = card B"
+  shows "A = B"
+proof-
+  from fB AB have fA: "finite A" by (auto intro: finite_subset)
+  from fA fB have fBA: "finite (B - A)" by auto
+  have e: "A \<inter> (B - A) = {}" by blast
+  have eq: "A \<union> (B - A) = B" using AB by blast
+  from card_Un_disjoint[OF fA fBA e, unfolded eq c]
+  have "card (B - A) = 0" by arith
+  hence "B - A = {}" unfolding card_eq_0_iff using fA fB by simp
+  with AB show "A = B" by blast
+qed
+
+lemma subspace_isomorphism:
+  assumes s: "subspace (S:: (real ^'n::finite) set)"
+  and t: "subspace (T :: (real ^ 'm::finite) set)"
+  and d: "dim S = dim T"
+  shows "\<exists>f. linear f \<and> f ` S = T \<and> inj_on f S"
+proof-
+  from basis_exists[of S] obtain B where
+    B: "B \<subseteq> S" "independent B" "S \<subseteq> span B" "B hassize dim S" by blast
+  from basis_exists[of T] obtain C where
+    C: "C \<subseteq> T" "independent C" "T \<subseteq> span C" "C hassize dim T" by blast
+  from B(4) C(4) card_le_inj[of B C] d obtain f where
+    f: "f ` B \<subseteq> C" "inj_on f B" unfolding hassize_def by auto
+  from linear_independent_extend[OF B(2)] obtain g where
+    g: "linear g" "\<forall>x\<in> B. g x = f x" by blast
+  from B(4) have fB: "finite B" by (simp add: hassize_def)
+  from C(4) have fC: "finite C" by (simp add: hassize_def)
+  from inj_on_iff_eq_card[OF fB, of f] f(2)
+  have "card (f ` B) = card B" by simp
+  with B(4) C(4) have ceq: "card (f ` B) = card C" using d
+    by (simp add: hassize_def)
+  have "g ` B = f ` B" using g(2)
+    by (auto simp add: image_iff)
+  also have "\<dots> = C" using card_subset_eq[OF fC f(1) ceq] .
+  finally have gBC: "g ` B = C" .
+  have gi: "inj_on g B" using f(2) g(2)
+    by (auto simp add: inj_on_def)
+  note g0 = linear_indep_image_lemma[OF g(1) fB, unfolded gBC, OF C(2) gi]
+  {fix x y assume x: "x \<in> S" and y: "y \<in> S" and gxy:"g x = g y"
+    from B(3) x y have x': "x \<in> span B" and y': "y \<in> span B" by blast+
+    from gxy have th0: "g (x - y) = 0" by (simp add: linear_sub[OF g(1)])
+    have th1: "x - y \<in> span B" using x' y' by (metis span_sub)
+    have "x=y" using g0[OF th1 th0] by simp }
+  then have giS: "inj_on g S"
+    unfolding inj_on_def by blast
+  from span_subspace[OF B(1,3) s]
+  have "g ` S = span (g ` B)" by (simp add: span_linear_image[OF g(1)])
+  also have "\<dots> = span C" unfolding gBC ..
+  also have "\<dots> = T" using span_subspace[OF C(1,3) t] .
+  finally have gS: "g ` S = T" .
+  from g(1) gS giS show ?thesis by blast
+qed
+
+(* linear functions are equal on a subspace if they are on a spanning set.   *)
+
+lemma subspace_kernel:
+  assumes lf: "linear (f::'a::semiring_1 ^'n \<Rightarrow> _)"
+  shows "subspace {x. f x = 0}"
+apply (simp add: subspace_def)
+by (simp add: linear_add[OF lf] linear_cmul[OF lf] linear_0[OF lf])
+
+lemma linear_eq_0_span:
+  assumes lf: "linear f" and f0: "\<forall>x\<in>B. f x = 0"
+  shows "\<forall>x \<in> span B. f x = (0::'a::semiring_1 ^'n)"
+proof
+  fix x assume x: "x \<in> span B"
+  let ?P = "\<lambda>x. f x = 0"
+  from subspace_kernel[OF lf] have "subspace ?P" unfolding Collect_def .
+  with x f0 span_induct[of B "?P" x] show "f x = 0" by blast
+qed
+
+lemma linear_eq_0:
+  assumes lf: "linear f" and SB: "S \<subseteq> span B" and f0: "\<forall>x\<in>B. f x = 0"
+  shows "\<forall>x \<in> S. f x = (0::'a::semiring_1^'n)"
+  by (metis linear_eq_0_span[OF lf] subset_eq SB f0)
+
+lemma linear_eq:
+  assumes lf: "linear (f::'a::ring_1^'n \<Rightarrow> _)" and lg: "linear g" and S: "S \<subseteq> span B"
+  and fg: "\<forall> x\<in> B. f x = g x"
+  shows "\<forall>x\<in> S. f x = g x"
+proof-
+  let ?h = "\<lambda>x. f x - g x"
+  from fg have fg': "\<forall>x\<in> B. ?h x = 0" by simp
+  from linear_eq_0[OF linear_compose_sub[OF lf lg] S fg']
+  show ?thesis by simp
+qed
+
+lemma linear_eq_stdbasis:
+  assumes lf: "linear (f::'a::ring_1^'m::finite \<Rightarrow> 'a^'n::finite)" and lg: "linear g"
+  and fg: "\<forall>i. f (basis i) = g(basis i)"
+  shows "f = g"
+proof-
+  let ?U = "UNIV :: 'm set"
+  let ?I = "{basis i:: 'a^'m|i. i \<in> ?U}"
+  {fix x assume x: "x \<in> (UNIV :: ('a^'m) set)"
+    from equalityD2[OF span_stdbasis]
+    have IU: " (UNIV :: ('a^'m) set) \<subseteq> span ?I" by blast
+    from linear_eq[OF lf lg IU] fg x
+    have "f x = g x" unfolding Collect_def  Ball_def mem_def by metis}
+  then show ?thesis by (auto intro: ext)
+qed
+
+(* Similar results for bilinear functions.                                   *)
+
+lemma bilinear_eq:
+  assumes bf: "bilinear (f:: 'a::ring^'m \<Rightarrow> 'a^'n \<Rightarrow> 'a^'p)"
+  and bg: "bilinear g"
+  and SB: "S \<subseteq> span B" and TC: "T \<subseteq> span C"
+  and fg: "\<forall>x\<in> B. \<forall>y\<in> C. f x y = g x y"
+  shows "\<forall>x\<in>S. \<forall>y\<in>T. f x y = g x y "
+proof-
+  let ?P = "\<lambda>x. \<forall>y\<in> span C. f x y = g x y"
+  from bf bg have sp: "subspace ?P"
+    unfolding bilinear_def linear_def subspace_def bf bg
+    by(auto simp add: span_0 mem_def bilinear_lzero[OF bf] bilinear_lzero[OF bg] span_add Ball_def intro:  bilinear_ladd[OF bf])
+
+  have "\<forall>x \<in> span B. \<forall>y\<in> span C. f x y = g x y"
+    apply -
+    apply (rule ballI)
+    apply (rule span_induct[of B ?P])
+    defer
+    apply (rule sp)
+    apply assumption
+    apply (clarsimp simp add: Ball_def)
+    apply (rule_tac P="\<lambda>y. f xa y = g xa y" and S=C in span_induct)
+    using fg
+    apply (auto simp add: subspace_def)
+    using bf bg unfolding bilinear_def linear_def
+    by(auto simp add: span_0 mem_def bilinear_rzero[OF bf] bilinear_rzero[OF bg] span_add Ball_def intro:  bilinear_ladd[OF bf])
+  then show ?thesis using SB TC by (auto intro: ext)
+qed
+
+lemma bilinear_eq_stdbasis:
+  assumes bf: "bilinear (f:: 'a::ring_1^'m::finite \<Rightarrow> 'a^'n::finite \<Rightarrow> 'a^'p)"
+  and bg: "bilinear g"
+  and fg: "\<forall>i j. f (basis i) (basis j) = g (basis i) (basis j)"
+  shows "f = g"
+proof-
+  from fg have th: "\<forall>x \<in> {basis i| i. i\<in> (UNIV :: 'm set)}. \<forall>y\<in>  {basis j |j. j \<in> (UNIV :: 'n set)}. f x y = g x y" by blast
+  from bilinear_eq[OF bf bg equalityD2[OF span_stdbasis] equalityD2[OF span_stdbasis] th] show ?thesis by (blast intro: ext)
+qed
+
+(* Detailed theorems about left and right invertibility in general case.     *)
+
+lemma left_invertible_transp:
+  "(\<exists>(B::'a^'n^'m). B ** transp (A::'a^'n^'m) = mat (1::'a::comm_semiring_1)) \<longleftrightarrow> (\<exists>(B::'a^'m^'n). A ** B = mat 1)"
+  by (metis matrix_transp_mul transp_mat transp_transp)
+
+lemma right_invertible_transp:
+  "(\<exists>(B::'a^'n^'m). transp (A::'a^'n^'m) ** B = mat (1::'a::comm_semiring_1)) \<longleftrightarrow> (\<exists>(B::'a^'m^'n). B ** A = mat 1)"
+  by (metis matrix_transp_mul transp_mat transp_transp)
+
+lemma linear_injective_left_inverse:
+  assumes lf: "linear (f::real ^'n::finite \<Rightarrow> real ^'m::finite)" and fi: "inj f"
+  shows "\<exists>g. linear g \<and> g o f = id"
+proof-
+  from linear_independent_extend[OF independent_injective_image, OF independent_stdbasis, OF lf fi]
+  obtain h:: "real ^'m \<Rightarrow> real ^'n" where h: "linear h" " \<forall>x \<in> f ` {basis i|i. i \<in> (UNIV::'n set)}. h x = inv f x" by blast
+  from h(2)
+  have th: "\<forall>i. (h \<circ> f) (basis i) = id (basis i)"
+    using inv_o_cancel[OF fi, unfolded stupid_ext[symmetric] id_def o_def]
+    by auto
+
+  from linear_eq_stdbasis[OF linear_compose[OF lf h(1)] linear_id th]
+  have "h o f = id" .
+  then show ?thesis using h(1) by blast
+qed
+
+lemma linear_surjective_right_inverse:
+  assumes lf: "linear (f:: real ^'m::finite \<Rightarrow> real ^'n::finite)" and sf: "surj f"
+  shows "\<exists>g. linear g \<and> f o g = id"
+proof-
+  from linear_independent_extend[OF independent_stdbasis]
+  obtain h:: "real ^'n \<Rightarrow> real ^'m" where
+    h: "linear h" "\<forall> x\<in> {basis i| i. i\<in> (UNIV :: 'n set)}. h x = inv f x" by blast
+  from h(2)
+  have th: "\<forall>i. (f o h) (basis i) = id (basis i)"
+    using sf
+    apply (auto simp add: surj_iff o_def stupid_ext[symmetric])
+    apply (erule_tac x="basis i" in allE)
+    by auto
+
+  from linear_eq_stdbasis[OF linear_compose[OF h(1) lf] linear_id th]
+  have "f o h = id" .
+  then show ?thesis using h(1) by blast
+qed
+
+lemma matrix_left_invertible_injective:
+"(\<exists>B. (B::real^'m^'n) ** (A::real^'n::finite^'m::finite) = mat 1) \<longleftrightarrow> (\<forall>x y. A *v x = A *v y \<longrightarrow> x = y)"
+proof-
+  {fix B:: "real^'m^'n" and x y assume B: "B ** A = mat 1" and xy: "A *v x = A*v y"
+    from xy have "B*v (A *v x) = B *v (A*v y)" by simp
+    hence "x = y"
+      unfolding matrix_vector_mul_assoc B matrix_vector_mul_lid .}
+  moreover
+  {assume A: "\<forall>x y. A *v x = A *v y \<longrightarrow> x = y"
+    hence i: "inj (op *v A)" unfolding inj_on_def by auto
+    from linear_injective_left_inverse[OF matrix_vector_mul_linear i]
+    obtain g where g: "linear g" "g o op *v A = id" by blast
+    have "matrix g ** A = mat 1"
+      unfolding matrix_eq matrix_vector_mul_lid matrix_vector_mul_assoc[symmetric] matrix_works[OF g(1)]
+      using g(2) by (simp add: o_def id_def stupid_ext)
+    then have "\<exists>B. (B::real ^'m^'n) ** A = mat 1" by blast}
+  ultimately show ?thesis by blast
+qed
+
+lemma matrix_left_invertible_ker:
+  "(\<exists>B. (B::real ^'m::finite^'n::finite) ** (A::real^'n^'m) = mat 1) \<longleftrightarrow> (\<forall>x. A *v x = 0 \<longrightarrow> x = 0)"
+  unfolding matrix_left_invertible_injective
+  using linear_injective_0[OF matrix_vector_mul_linear, of A]
+  by (simp add: inj_on_def)
+
+lemma matrix_right_invertible_surjective:
+"(\<exists>B. (A::real^'n::finite^'m::finite) ** (B::real^'m^'n) = mat 1) \<longleftrightarrow> surj (\<lambda>x. A *v x)"
+proof-
+  {fix B :: "real ^'m^'n"  assume AB: "A ** B = mat 1"
+    {fix x :: "real ^ 'm"
+      have "A *v (B *v x) = x"
+        by (simp add: matrix_vector_mul_lid matrix_vector_mul_assoc AB)}
+    hence "surj (op *v A)" unfolding surj_def by metis }
+  moreover
+  {assume sf: "surj (op *v A)"
+    from linear_surjective_right_inverse[OF matrix_vector_mul_linear sf]
+    obtain g:: "real ^'m \<Rightarrow> real ^'n" where g: "linear g" "op *v A o g = id"
+      by blast
+
+    have "A ** (matrix g) = mat 1"
+      unfolding matrix_eq  matrix_vector_mul_lid
+        matrix_vector_mul_assoc[symmetric] matrix_works[OF g(1)]
+      using g(2) unfolding o_def stupid_ext[symmetric] id_def
+      .
+    hence "\<exists>B. A ** (B::real^'m^'n) = mat 1" by blast
+  }
+  ultimately show ?thesis unfolding surj_def by blast
+qed
+
+lemma matrix_left_invertible_independent_columns:
+  fixes A :: "real^'n::finite^'m::finite"
+  shows "(\<exists>(B::real ^'m^'n). B ** A = mat 1) \<longleftrightarrow> (\<forall>c. setsum (\<lambda>i. c i *s column i A) (UNIV :: 'n set) = 0 \<longrightarrow> (\<forall>i. c i = 0))"
+   (is "?lhs \<longleftrightarrow> ?rhs")
+proof-
+  let ?U = "UNIV :: 'n set"
+  {assume k: "\<forall>x. A *v x = 0 \<longrightarrow> x = 0"
+    {fix c i assume c: "setsum (\<lambda>i. c i *s column i A) ?U = 0"
+      and i: "i \<in> ?U"
+      let ?x = "\<chi> i. c i"
+      have th0:"A *v ?x = 0"
+        using c
+        unfolding matrix_mult_vsum Cart_eq
+        by auto
+      from k[rule_format, OF th0] i
+      have "c i = 0" by (vector Cart_eq)}
+    hence ?rhs by blast}
+  moreover
+  {assume H: ?rhs
+    {fix x assume x: "A *v x = 0"
+      let ?c = "\<lambda>i. ((x$i ):: real)"
+      from H[rule_format, of ?c, unfolded matrix_mult_vsum[symmetric], OF x]
+      have "x = 0" by vector}}
+  ultimately show ?thesis unfolding matrix_left_invertible_ker by blast
+qed
+
+lemma matrix_right_invertible_independent_rows:
+  fixes A :: "real^'n::finite^'m::finite"
+  shows "(\<exists>(B::real^'m^'n). A ** B = mat 1) \<longleftrightarrow> (\<forall>c. setsum (\<lambda>i. c i *s row i A) (UNIV :: 'm set) = 0 \<longrightarrow> (\<forall>i. c i = 0))"
+  unfolding left_invertible_transp[symmetric]
+    matrix_left_invertible_independent_columns
+  by (simp add: column_transp)
+
+lemma matrix_right_invertible_span_columns:
+  "(\<exists>(B::real ^'n::finite^'m::finite). (A::real ^'m^'n) ** B = mat 1) \<longleftrightarrow> span (columns A) = UNIV" (is "?lhs = ?rhs")
+proof-
+  let ?U = "UNIV :: 'm set"
+  have fU: "finite ?U" by simp
+  have lhseq: "?lhs \<longleftrightarrow> (\<forall>y. \<exists>(x::real^'m). setsum (\<lambda>i. (x$i) *s column i A) ?U = y)"
+    unfolding matrix_right_invertible_surjective matrix_mult_vsum surj_def
+    apply (subst eq_commute) ..
+  have rhseq: "?rhs \<longleftrightarrow> (\<forall>x. x \<in> span (columns A))" by blast
+  {assume h: ?lhs
+    {fix x:: "real ^'n"
+        from h[unfolded lhseq, rule_format, of x] obtain y:: "real ^'m"
+          where y: "setsum (\<lambda>i. (y$i) *s column i A) ?U = x" by blast
+        have "x \<in> span (columns A)"
+          unfolding y[symmetric]
+          apply (rule span_setsum[OF fU])
+          apply clarify
+          apply (rule span_mul)
+          apply (rule span_superset)
+          unfolding columns_def
+          by blast}
+    then have ?rhs unfolding rhseq by blast}
+  moreover
+  {assume h:?rhs
+    let ?P = "\<lambda>(y::real ^'n). \<exists>(x::real^'m). setsum (\<lambda>i. (x$i) *s column i A) ?U = y"
+    {fix y have "?P y"
+      proof(rule span_induct_alt[of ?P "columns A"])
+        show "\<exists>x\<Colon>real ^ 'm. setsum (\<lambda>i. (x$i) *s column i A) ?U = 0"
+          apply (rule exI[where x=0])
+          by (simp add: zero_index vector_smult_lzero)
+      next
+        fix c y1 y2 assume y1: "y1 \<in> columns A" and y2: "?P y2"
+        from y1 obtain i where i: "i \<in> ?U" "y1 = column i A"
+          unfolding columns_def by blast
+        from y2 obtain x:: "real ^'m" where
+          x: "setsum (\<lambda>i. (x$i) *s column i A) ?U = y2" by blast
+        let ?x = "(\<chi> j. if j = i then c + (x$i) else (x$j))::real^'m"
+        show "?P (c*s y1 + y2)"
+          proof(rule exI[where x= "?x"], vector, auto simp add: i x[symmetric] cond_value_iff right_distrib cond_application_beta cong del: if_weak_cong)
+            fix j
+            have th: "\<forall>xa \<in> ?U. (if xa = i then (c + (x$i)) * ((column xa A)$j)
+           else (x$xa) * ((column xa A$j))) = (if xa = i then c * ((column i A)$j) else 0) + ((x$xa) * ((column xa A)$j))" using i(1)
+              by (simp add: ring_simps)
+            have "setsum (\<lambda>xa. if xa = i then (c + (x$i)) * ((column xa A)$j)
+           else (x$xa) * ((column xa A$j))) ?U = setsum (\<lambda>xa. (if xa = i then c * ((column i A)$j) else 0) + ((x$xa) * ((column xa A)$j))) ?U"
+              apply (rule setsum_cong[OF refl])
+              using th by blast
+            also have "\<dots> = setsum (\<lambda>xa. if xa = i then c * ((column i A)$j) else 0) ?U + setsum (\<lambda>xa. ((x$xa) * ((column xa A)$j))) ?U"
+              by (simp add: setsum_addf)
+            also have "\<dots> = c * ((column i A)$j) + setsum (\<lambda>xa. ((x$xa) * ((column xa A)$j))) ?U"
+              unfolding setsum_delta[OF fU]
+              using i(1) by simp
+            finally show "setsum (\<lambda>xa. if xa = i then (c + (x$i)) * ((column xa A)$j)
+           else (x$xa) * ((column xa A$j))) ?U = c * ((column i A)$j) + setsum (\<lambda>xa. ((x$xa) * ((column xa A)$j))) ?U" .
+          qed
+        next
+          show "y \<in> span (columns A)" unfolding h by blast
+        qed}
+    then have ?lhs unfolding lhseq ..}
+  ultimately show ?thesis by blast
+qed
+
+lemma matrix_left_invertible_span_rows:
+  "(\<exists>(B::real^'m::finite^'n::finite). B ** (A::real^'n^'m) = mat 1) \<longleftrightarrow> span (rows A) = UNIV"
+  unfolding right_invertible_transp[symmetric]
+  unfolding columns_transp[symmetric]
+  unfolding matrix_right_invertible_span_columns
+ ..
+
+(* An injective map real^'n->real^'n is also surjective.                       *)
+
+lemma linear_injective_imp_surjective:
+  assumes lf: "linear (f:: real ^'n::finite \<Rightarrow> real ^'n)" and fi: "inj f"
+  shows "surj f"
+proof-
+  let ?U = "UNIV :: (real ^'n) set"
+  from basis_exists[of ?U] obtain B
+    where B: "B \<subseteq> ?U" "independent B" "?U \<subseteq> span B" "B hassize dim ?U"
+    by blast
+  from B(4) have d: "dim ?U = card B" by (simp add: hassize_def)
+  have th: "?U \<subseteq> span (f ` B)"
+    apply (rule card_ge_dim_independent)
+    apply blast
+    apply (rule independent_injective_image[OF B(2) lf fi])
+    apply (rule order_eq_refl)
+    apply (rule sym)
+    unfolding d
+    apply (rule card_image)
+    apply (rule subset_inj_on[OF fi])
+    by blast
+  from th show ?thesis
+    unfolding span_linear_image[OF lf] surj_def
+    using B(3) by blast
+qed
+
+(* And vice versa.                                                           *)
+
+lemma surjective_iff_injective_gen:
+  assumes fS: "finite S" and fT: "finite T" and c: "card S = card T"
+  and ST: "f ` S \<subseteq> T"
+  shows "(\<forall>y \<in> T. \<exists>x \<in> S. f x = y) \<longleftrightarrow> inj_on f S" (is "?lhs \<longleftrightarrow> ?rhs")
+proof-
+  {assume h: "?lhs"
+    {fix x y assume x: "x \<in> S" and y: "y \<in> S" and f: "f x = f y"
+      from x fS have S0: "card S \<noteq> 0" by auto
+      {assume xy: "x \<noteq> y"
+        have th: "card S \<le> card (f ` (S - {y}))"
+          unfolding c
+          apply (rule card_mono)
+          apply (rule finite_imageI)
+          using fS apply simp
+          using h xy x y f unfolding subset_eq image_iff
+          apply auto
+          apply (case_tac "xa = f x")
+          apply (rule bexI[where x=x])
+          apply auto
+          done
+        also have " \<dots> \<le> card (S -{y})"
+          apply (rule card_image_le)
+          using fS by simp
+        also have "\<dots> \<le> card S - 1" using y fS by simp
+        finally have False  using S0 by arith }
+      then have "x = y" by blast}
+    then have ?rhs unfolding inj_on_def by blast}
+  moreover
+  {assume h: ?rhs
+    have "f ` S = T"
+      apply (rule card_subset_eq[OF fT ST])
+      unfolding card_image[OF h] using c .
+    then have ?lhs by blast}
+  ultimately show ?thesis by blast
+qed
+
+lemma linear_surjective_imp_injective:
+  assumes lf: "linear (f::real ^'n::finite => real ^'n)" and sf: "surj f"
+  shows "inj f"
+proof-
+  let ?U = "UNIV :: (real ^'n) set"
+  from basis_exists[of ?U] obtain B
+    where B: "B \<subseteq> ?U" "independent B" "?U \<subseteq> span B" "B hassize dim ?U"
+    by blast
+  {fix x assume x: "x \<in> span B" and fx: "f x = 0"
+    from B(4) have fB: "finite B" by (simp add: hassize_def)
+    from B(4) have d: "dim ?U = card B" by (simp add: hassize_def)
+    have fBi: "independent (f ` B)"
+      apply (rule card_le_dim_spanning[of "f ` B" ?U])
+      apply blast
+      using sf B(3)
+      unfolding span_linear_image[OF lf] surj_def subset_eq image_iff
+      apply blast
+      using fB apply (blast intro: finite_imageI)
+      unfolding d
+      apply (rule card_image_le)
+      apply (rule fB)
+      done
+    have th0: "dim ?U \<le> card (f ` B)"
+      apply (rule span_card_ge_dim)
+      apply blast
+      unfolding span_linear_image[OF lf]
+      apply (rule subset_trans[where B = "f ` UNIV"])
+      using sf unfolding surj_def apply blast
+      apply (rule image_mono)
+      apply (rule B(3))
+      apply (metis finite_imageI fB)
+      done
+
+    moreover have "card (f ` B) \<le> card B"
+      by (rule card_image_le, rule fB)
+    ultimately have th1: "card B = card (f ` B)" unfolding d by arith
+    have fiB: "inj_on f B"
+      unfolding surjective_iff_injective_gen[OF fB finite_imageI[OF fB] th1 subset_refl, symmetric] by blast
+    from linear_indep_image_lemma[OF lf fB fBi fiB x] fx
+    have "x = 0" by blast}
+  note th = this
+  from th show ?thesis unfolding linear_injective_0[OF lf]
+    using B(3) by blast
+qed
+
+(* Hence either is enough for isomorphism.                                   *)
+
+lemma left_right_inverse_eq:
+  assumes fg: "f o g = id" and gh: "g o h = id"
+  shows "f = h"
+proof-
+  have "f = f o (g o h)" unfolding gh by simp
+  also have "\<dots> = (f o g) o h" by (simp add: o_assoc)
+  finally show "f = h" unfolding fg by simp
+qed
+
+lemma isomorphism_expand:
+  "f o g = id \<and> g o f = id \<longleftrightarrow> (\<forall>x. f(g x) = x) \<and> (\<forall>x. g(f x) = x)"
+  by (simp add: expand_fun_eq o_def id_def)
+
+lemma linear_injective_isomorphism:
+  assumes lf: "linear (f :: real^'n::finite \<Rightarrow> real ^'n)" and fi: "inj f"
+  shows "\<exists>f'. linear f' \<and> (\<forall>x. f' (f x) = x) \<and> (\<forall>x. f (f' x) = x)"
+unfolding isomorphism_expand[symmetric]
+using linear_surjective_right_inverse[OF lf linear_injective_imp_surjective[OF lf fi]] linear_injective_left_inverse[OF lf fi]
+by (metis left_right_inverse_eq)
+
+lemma linear_surjective_isomorphism:
+  assumes lf: "linear (f::real ^'n::finite \<Rightarrow> real ^'n)" and sf: "surj f"
+  shows "\<exists>f'. linear f' \<and> (\<forall>x. f' (f x) = x) \<and> (\<forall>x. f (f' x) = x)"
+unfolding isomorphism_expand[symmetric]
+using linear_surjective_right_inverse[OF lf sf] linear_injective_left_inverse[OF lf linear_surjective_imp_injective[OF lf sf]]
+by (metis left_right_inverse_eq)
+
+(* Left and right inverses are the same for R^N->R^N.                        *)
+
+lemma linear_inverse_left:
+  assumes lf: "linear (f::real ^'n::finite \<Rightarrow> real ^'n)" and lf': "linear f'"
+  shows "f o f' = id \<longleftrightarrow> f' o f = id"
+proof-
+  {fix f f':: "real ^'n \<Rightarrow> real ^'n"
+    assume lf: "linear f" "linear f'" and f: "f o f' = id"
+    from f have sf: "surj f"
+
+      apply (auto simp add: o_def stupid_ext[symmetric] id_def surj_def)
+      by metis
+    from linear_surjective_isomorphism[OF lf(1) sf] lf f
+    have "f' o f = id" unfolding stupid_ext[symmetric] o_def id_def
+      by metis}
+  then show ?thesis using lf lf' by metis
+qed
+
+(* Moreover, a one-sided inverse is automatically linear.                    *)
+
+lemma left_inverse_linear:
+  assumes lf: "linear (f::real ^'n::finite \<Rightarrow> real ^'n)" and gf: "g o f = id"
+  shows "linear g"
+proof-
+  from gf have fi: "inj f" apply (auto simp add: inj_on_def o_def id_def stupid_ext[symmetric])
+    by metis
+  from linear_injective_isomorphism[OF lf fi]
+  obtain h:: "real ^'n \<Rightarrow> real ^'n" where
+    h: "linear h" "\<forall>x. h (f x) = x" "\<forall>x. f (h x) = x" by blast
+  have "h = g" apply (rule ext) using gf h(2,3)
+    apply (simp add: o_def id_def stupid_ext[symmetric])
+    by metis
+  with h(1) show ?thesis by blast
+qed
+
+lemma right_inverse_linear:
+  assumes lf: "linear (f:: real ^'n::finite \<Rightarrow> real ^'n)" and gf: "f o g = id"
+  shows "linear g"
+proof-
+  from gf have fi: "surj f" apply (auto simp add: surj_def o_def id_def stupid_ext[symmetric])
+    by metis
+  from linear_surjective_isomorphism[OF lf fi]
+  obtain h:: "real ^'n \<Rightarrow> real ^'n" where
+    h: "linear h" "\<forall>x. h (f x) = x" "\<forall>x. f (h x) = x" by blast
+  have "h = g" apply (rule ext) using gf h(2,3)
+    apply (simp add: o_def id_def stupid_ext[symmetric])
+    by metis
+  with h(1) show ?thesis by blast
+qed
+
+(* The same result in terms of square matrices.                              *)
+
+lemma matrix_left_right_inverse:
+  fixes A A' :: "real ^'n::finite^'n"
+  shows "A ** A' = mat 1 \<longleftrightarrow> A' ** A = mat 1"
+proof-
+  {fix A A' :: "real ^'n^'n" assume AA': "A ** A' = mat 1"
+    have sA: "surj (op *v A)"
+      unfolding surj_def
+      apply clarify
+      apply (rule_tac x="(A' *v y)" in exI)
+      by (simp add: matrix_vector_mul_assoc AA' matrix_vector_mul_lid)
+    from linear_surjective_isomorphism[OF matrix_vector_mul_linear sA]
+    obtain f' :: "real ^'n \<Rightarrow> real ^'n"
+      where f': "linear f'" "\<forall>x. f' (A *v x) = x" "\<forall>x. A *v f' x = x" by blast
+    have th: "matrix f' ** A = mat 1"
+      by (simp add: matrix_eq matrix_works[OF f'(1)] matrix_vector_mul_assoc[symmetric] matrix_vector_mul_lid f'(2)[rule_format])
+    hence "(matrix f' ** A) ** A' = mat 1 ** A'" by simp
+    hence "matrix f' = A'" by (simp add: matrix_mul_assoc[symmetric] AA' matrix_mul_rid matrix_mul_lid)
+    hence "matrix f' ** A = A' ** A" by simp
+    hence "A' ** A = mat 1" by (simp add: th)}
+  then show ?thesis by blast
+qed
+
+(* Considering an n-element vector as an n-by-1 or 1-by-n matrix.            *)
+
+definition "rowvector v = (\<chi> i j. (v$j))"
+
+definition "columnvector v = (\<chi> i j. (v$i))"
+
+lemma transp_columnvector:
+ "transp(columnvector v) = rowvector v"
+  by (simp add: transp_def rowvector_def columnvector_def Cart_eq)
+
+lemma transp_rowvector: "transp(rowvector v) = columnvector v"
+  by (simp add: transp_def columnvector_def rowvector_def Cart_eq)
+
+lemma dot_rowvector_columnvector:
+  "columnvector (A *v v) = A ** columnvector v"
+  by (vector columnvector_def matrix_matrix_mult_def matrix_vector_mult_def)
+
+lemma dot_matrix_product: "(x::'a::semiring_1^'n::finite) \<bullet> y = (((rowvector x ::'a^'n^1) ** (columnvector y :: 'a^1^'n))$1)$1"
+  by (vector matrix_matrix_mult_def rowvector_def columnvector_def dot_def)
+
+lemma dot_matrix_vector_mul:
+  fixes A B :: "real ^'n::finite ^'n" and x y :: "real ^'n"
+  shows "(A *v x) \<bullet> (B *v y) =
+      (((rowvector x :: real^'n^1) ** ((transp A ** B) ** (columnvector y :: real ^1^'n)))$1)$1"
+unfolding dot_matrix_product transp_columnvector[symmetric]
+  dot_rowvector_columnvector matrix_transp_mul matrix_mul_assoc ..
+
+(* Infinity norm.                                                            *)
+
+definition "infnorm (x::real^'n::finite) = rsup {abs(x$i) |i. i\<in> (UNIV :: 'n set)}"
+
+lemma numseg_dimindex_nonempty: "\<exists>i. i \<in> (UNIV :: 'n set)"
+  by auto
+
+lemma infnorm_set_image:
+  "{abs(x$i) |i. i\<in> (UNIV :: 'n set)} =
+  (\<lambda>i. abs(x$i)) ` (UNIV :: 'n set)" by blast
+
+lemma infnorm_set_lemma:
+  shows "finite {abs((x::'a::abs ^'n::finite)$i) |i. i\<in> (UNIV :: 'n set)}"
+  and "{abs(x$i) |i. i\<in> (UNIV :: 'n::finite set)} \<noteq> {}"
+  unfolding infnorm_set_image
+  by (auto intro: finite_imageI)
+
+lemma infnorm_pos_le: "0 \<le> infnorm (x::real^'n::finite)"
+  unfolding infnorm_def
+  unfolding rsup_finite_ge_iff[ OF infnorm_set_lemma]
+  unfolding infnorm_set_image
+  by auto
+
+lemma infnorm_triangle: "infnorm ((x::real^'n::finite) + y) \<le> infnorm x + infnorm y"
+proof-
+  have th: "\<And>x y (z::real). x - y <= z \<longleftrightarrow> x - z <= y" by arith
+  have th1: "\<And>S f. f ` S = { f i| i. i \<in> S}" by blast
+  have th2: "\<And>x (y::real). abs(x + y) - abs(x) <= abs(y)" by arith
+  show ?thesis
+  unfolding infnorm_def
+  unfolding rsup_finite_le_iff[ OF infnorm_set_lemma]
+  apply (subst diff_le_eq[symmetric])
+  unfolding rsup_finite_ge_iff[ OF infnorm_set_lemma]
+  unfolding infnorm_set_image bex_simps
+  apply (subst th)
+  unfolding th1
+  unfolding rsup_finite_ge_iff[ OF infnorm_set_lemma]
+
+  unfolding infnorm_set_image ball_simps bex_simps
+  apply simp
+  apply (metis th2)
+  done
+qed
+
+lemma infnorm_eq_0: "infnorm x = 0 \<longleftrightarrow> (x::real ^'n::finite) = 0"
+proof-
+  have "infnorm x <= 0 \<longleftrightarrow> x = 0"
+    unfolding infnorm_def
+    unfolding rsup_finite_le_iff[OF infnorm_set_lemma]
+    unfolding infnorm_set_image ball_simps
+    by vector
+  then show ?thesis using infnorm_pos_le[of x] by simp
+qed
+
+lemma infnorm_0: "infnorm 0 = 0"
+  by (simp add: infnorm_eq_0)
+
+lemma infnorm_neg: "infnorm (- x) = infnorm x"
+  unfolding infnorm_def
+  apply (rule cong[of "rsup" "rsup"])
+  apply blast
+  apply (rule set_ext)
+  apply auto
+  done
+
+lemma infnorm_sub: "infnorm (x - y) = infnorm (y - x)"
+proof-
+  have "y - x = - (x - y)" by simp
+  then show ?thesis  by (metis infnorm_neg)
+qed
+
+lemma real_abs_sub_infnorm: "\<bar> infnorm x - infnorm y\<bar> \<le> infnorm (x - y)"
+proof-
+  have th: "\<And>(nx::real) n ny. nx <= n + ny \<Longrightarrow> ny <= n + nx ==> \<bar>nx - ny\<bar> <= n"
+    by arith
+  from infnorm_triangle[of "x - y" " y"] infnorm_triangle[of "x - y" "-x"]
+  have ths: "infnorm x \<le> infnorm (x - y) + infnorm y"
+    "infnorm y \<le> infnorm (x - y) + infnorm x"
+    by (simp_all add: ring_simps infnorm_neg diff_def[symmetric])
+  from th[OF ths]  show ?thesis .
+qed
+
+lemma real_abs_infnorm: " \<bar>infnorm x\<bar> = infnorm x"
+  using infnorm_pos_le[of x] by arith
+
+lemma component_le_infnorm:
+  shows "\<bar>x$i\<bar> \<le> infnorm (x::real^'n::finite)"
+proof-
+  let ?U = "UNIV :: 'n set"
+  let ?S = "{\<bar>x$i\<bar> |i. i\<in> ?U}"
+  have fS: "finite ?S" unfolding image_Collect[symmetric]
+    apply (rule finite_imageI) unfolding Collect_def mem_def by simp
+  have S0: "?S \<noteq> {}" by blast
+  have th1: "\<And>S f. f ` S = { f i| i. i \<in> S}" by blast
+  from rsup_finite_in[OF fS S0] rsup_finite_Ub[OF fS S0]
+  show ?thesis unfolding infnorm_def isUb_def setle_def
+    unfolding infnorm_set_image ball_simps by auto
+qed
+
+lemma infnorm_mul_lemma: "infnorm(a *s x) <= \<bar>a\<bar> * infnorm x"
+  apply (subst infnorm_def)
+  unfolding rsup_finite_le_iff[OF infnorm_set_lemma]
+  unfolding infnorm_set_image ball_simps
+  apply (simp add: abs_mult)
+  apply (rule allI)
+  apply (cut_tac component_le_infnorm[of x])
+  apply (rule mult_mono)
+  apply auto
+  done
+
+lemma infnorm_mul: "infnorm(a *s x) = abs a * infnorm x"
+proof-
+  {assume a0: "a = 0" hence ?thesis by (simp add: infnorm_0) }
+  moreover
+  {assume a0: "a \<noteq> 0"
+    from a0 have th: "(1/a) *s (a *s x) = x"
+      by (simp add: vector_smult_assoc)
+    from a0 have ap: "\<bar>a\<bar> > 0" by arith
+    from infnorm_mul_lemma[of "1/a" "a *s x"]
+    have "infnorm x \<le> 1/\<bar>a\<bar> * infnorm (a*s x)"
+      unfolding th by simp
+    with ap have "\<bar>a\<bar> * infnorm x \<le> \<bar>a\<bar> * (1/\<bar>a\<bar> * infnorm (a *s x))" by (simp add: field_simps)
+    then have "\<bar>a\<bar> * infnorm x \<le> infnorm (a*s x)"
+      using ap by (simp add: field_simps)
+    with infnorm_mul_lemma[of a x] have ?thesis by arith }
+  ultimately show ?thesis by blast
+qed
+
+lemma infnorm_pos_lt: "infnorm x > 0 \<longleftrightarrow> x \<noteq> 0"
+  using infnorm_pos_le[of x] infnorm_eq_0[of x] by arith
+
+(* Prove that it differs only up to a bound from Euclidean norm.             *)
+
+lemma infnorm_le_norm: "infnorm x \<le> norm x"
+  unfolding infnorm_def rsup_finite_le_iff[OF infnorm_set_lemma]
+  unfolding infnorm_set_image  ball_simps
+  by (metis component_le_norm)
+lemma card_enum: "card {1 .. n} = n" by auto
+lemma norm_le_infnorm: "norm(x) <= sqrt(real CARD('n)) * infnorm(x::real ^'n::finite)"
+proof-
+  let ?d = "CARD('n)"
+  have "real ?d \<ge> 0" by simp
+  hence d2: "(sqrt (real ?d))^2 = real ?d"
+    by (auto intro: real_sqrt_pow2)
+  have th: "sqrt (real ?d) * infnorm x \<ge> 0"
+    by (simp add: zero_le_mult_iff real_sqrt_ge_0_iff infnorm_pos_le)
+  have th1: "x\<bullet>x \<le> (sqrt (real ?d) * infnorm x)^2"
+    unfolding power_mult_distrib d2
+    apply (subst power2_abs[symmetric])
+    unfolding real_of_nat_def dot_def power2_eq_square[symmetric]
+    apply (subst power2_abs[symmetric])
+    apply (rule setsum_bounded)
+    apply (rule power_mono)
+    unfolding abs_of_nonneg[OF infnorm_pos_le]
+    unfolding infnorm_def  rsup_finite_ge_iff[OF infnorm_set_lemma]
+    unfolding infnorm_set_image bex_simps
+    apply blast
+    by (rule abs_ge_zero)
+  from real_le_lsqrt[OF dot_pos_le th th1]
+  show ?thesis unfolding real_vector_norm_def id_def .
+qed
+
+(* Equality in Cauchy-Schwarz and triangle inequalities.                     *)
+
+lemma norm_cauchy_schwarz_eq: "(x::real ^'n::finite) \<bullet> y = norm x * norm y \<longleftrightarrow> norm x *s y = norm y *s x" (is "?lhs \<longleftrightarrow> ?rhs")
+proof-
+  {assume h: "x = 0"
+    hence ?thesis by simp}
+  moreover
+  {assume h: "y = 0"
+    hence ?thesis by simp}
+  moreover
+  {assume x: "x \<noteq> 0" and y: "y \<noteq> 0"
+    from dot_eq_0[of "norm y *s x - norm x *s y"]
+    have "?rhs \<longleftrightarrow> (norm y * (norm y * norm x * norm x - norm x * (x \<bullet> y)) - norm x * (norm y * (y \<bullet> x) - norm x * norm y * norm y) =  0)"
+      using x y
+      unfolding dot_rsub dot_lsub dot_lmult dot_rmult
+      unfolding norm_pow_2[symmetric] power2_eq_square diff_eq_0_iff_eq apply (simp add: dot_sym)
+      apply (simp add: ring_simps)
+      apply metis
+      done
+    also have "\<dots> \<longleftrightarrow> (2 * norm x * norm y * (norm x * norm y - x \<bullet> y) = 0)" using x y
+      by (simp add: ring_simps dot_sym)
+    also have "\<dots> \<longleftrightarrow> ?lhs" using x y
+      apply simp
+      by metis
+    finally have ?thesis by blast}
+  ultimately show ?thesis by blast
+qed
+
+lemma norm_cauchy_schwarz_abs_eq:
+  fixes x y :: "real ^ 'n::finite"
+  shows "abs(x \<bullet> y) = norm x * norm y \<longleftrightarrow>
+                norm x *s y = norm y *s x \<or> norm(x) *s y = - norm y *s x" (is "?lhs \<longleftrightarrow> ?rhs")
+proof-
+  have th: "\<And>(x::real) a. a \<ge> 0 \<Longrightarrow> abs x = a \<longleftrightarrow> x = a \<or> x = - a" by arith
+  have "?rhs \<longleftrightarrow> norm x *s y = norm y *s x \<or> norm (- x) *s y = norm y *s (- x)"
+    apply simp by vector
+  also have "\<dots> \<longleftrightarrow>(x \<bullet> y = norm x * norm y \<or>
+     (-x) \<bullet> y = norm x * norm y)"
+    unfolding norm_cauchy_schwarz_eq[symmetric]
+    unfolding norm_minus_cancel
+      norm_mul by blast
+  also have "\<dots> \<longleftrightarrow> ?lhs"
+    unfolding th[OF mult_nonneg_nonneg, OF norm_ge_zero[of x] norm_ge_zero[of y]] dot_lneg
+    by arith
+  finally show ?thesis ..
+qed
+
+lemma norm_triangle_eq:
+  fixes x y :: "real ^ 'n::finite"
+  shows "norm(x + y) = norm x + norm y \<longleftrightarrow> norm x *s y = norm y *s x"
+proof-
+  {assume x: "x =0 \<or> y =0"
+    hence ?thesis by (cases "x=0", simp_all)}
+  moreover
+  {assume x: "x \<noteq> 0" and y: "y \<noteq> 0"
+    hence "norm x \<noteq> 0" "norm y \<noteq> 0"
+      by simp_all
+    hence n: "norm x > 0" "norm y > 0"
+      using norm_ge_zero[of x] norm_ge_zero[of y]
+      by arith+
+    have th: "\<And>(a::real) b c. a + b + c \<noteq> 0 ==> (a = b + c \<longleftrightarrow> a^2 = (b + c)^2)" by algebra
+    have "norm(x + y) = norm x + norm y \<longleftrightarrow> norm(x + y)^ 2 = (norm x + norm y) ^2"
+      apply (rule th) using n norm_ge_zero[of "x + y"]
+      by arith
+    also have "\<dots> \<longleftrightarrow> norm x *s y = norm y *s x"
+      unfolding norm_cauchy_schwarz_eq[symmetric]
+      unfolding norm_pow_2 dot_ladd dot_radd
+      by (simp add: norm_pow_2[symmetric] power2_eq_square dot_sym ring_simps)
+    finally have ?thesis .}
+  ultimately show ?thesis by blast
+qed
+
+(* Collinearity.*)
+
+definition "collinear S \<longleftrightarrow> (\<exists>u. \<forall>x \<in> S. \<forall> y \<in> S. \<exists>c. x - y = c *s u)"
+
+lemma collinear_empty:  "collinear {}" by (simp add: collinear_def)
+
+lemma collinear_sing: "collinear {(x::'a::ring_1^'n)}"
+  apply (simp add: collinear_def)
+  apply (rule exI[where x=0])
+  by simp
+
+lemma collinear_2: "collinear {(x::'a::ring_1^'n),y}"
+  apply (simp add: collinear_def)
+  apply (rule exI[where x="x - y"])
+  apply auto
+  apply (rule exI[where x=0], simp)
+  apply (rule exI[where x=1], simp)
+  apply (rule exI[where x="- 1"], simp add: vector_sneg_minus1[symmetric])
+  apply (rule exI[where x=0], simp)
+  done
+
+lemma collinear_lemma: "collinear {(0::real^'n),x,y} \<longleftrightarrow> x = 0 \<or> y = 0 \<or> (\<exists>c. y = c *s x)" (is "?lhs \<longleftrightarrow> ?rhs")
+proof-
+  {assume "x=0 \<or> y = 0" hence ?thesis
+      by (cases "x = 0", simp_all add: collinear_2 insert_commute)}
+  moreover
+  {assume x: "x \<noteq> 0" and y: "y \<noteq> 0"
+    {assume h: "?lhs"
+      then obtain u where u: "\<forall> x\<in> {0,x,y}. \<forall>y\<in> {0,x,y}. \<exists>c. x - y = c *s u" unfolding collinear_def by blast
+      from u[rule_format, of x 0] u[rule_format, of y 0]
+      obtain cx and cy where
+        cx: "x = cx*s u" and cy: "y = cy*s u"
+        by auto
+      from cx x have cx0: "cx \<noteq> 0" by auto
+      from cy y have cy0: "cy \<noteq> 0" by auto
+      let ?d = "cy / cx"
+      from cx cy cx0 have "y = ?d *s x"
+        by (simp add: vector_smult_assoc)
+      hence ?rhs using x y by blast}
+    moreover
+    {assume h: "?rhs"
+      then obtain c where c: "y = c*s x" using x y by blast
+      have ?lhs unfolding collinear_def c
+        apply (rule exI[where x=x])
+        apply auto
+        apply (rule exI[where x="- 1"], simp only: vector_smult_lneg vector_smult_lid)
+        apply (rule exI[where x= "-c"], simp only: vector_smult_lneg)
+        apply (rule exI[where x=1], simp)
+        apply (rule exI[where x="1 - c"], simp add: vector_smult_lneg vector_sub_rdistrib)
+        apply (rule exI[where x="c - 1"], simp add: vector_smult_lneg vector_sub_rdistrib)
+        done}
+    ultimately have ?thesis by blast}
+  ultimately show ?thesis by blast
+qed
+
+lemma norm_cauchy_schwarz_equal:
+  fixes x y :: "real ^ 'n::finite"
+  shows "abs(x \<bullet> y) = norm x * norm y \<longleftrightarrow> collinear {(0::real^'n),x,y}"
+unfolding norm_cauchy_schwarz_abs_eq
+apply (cases "x=0", simp_all add: collinear_2)
+apply (cases "y=0", simp_all add: collinear_2 insert_commute)
+unfolding collinear_lemma
+apply simp
+apply (subgoal_tac "norm x \<noteq> 0")
+apply (subgoal_tac "norm y \<noteq> 0")
+apply (rule iffI)
+apply (cases "norm x *s y = norm y *s x")
+apply (rule exI[where x="(1/norm x) * norm y"])
+apply (drule sym)
+unfolding vector_smult_assoc[symmetric]
+apply (simp add: vector_smult_assoc field_simps)
+apply (rule exI[where x="(1/norm x) * - norm y"])
+apply clarify
+apply (drule sym)
+unfolding vector_smult_assoc[symmetric]
+apply (simp add: vector_smult_assoc field_simps)
+apply (erule exE)
+apply (erule ssubst)
+unfolding vector_smult_assoc
+unfolding norm_mul
+apply (subgoal_tac "norm x * c = \<bar>c\<bar> * norm x \<or> norm x * c = - \<bar>c\<bar> * norm x")
+apply (case_tac "c <= 0", simp add: ring_simps)
+apply (simp add: ring_simps)
+apply (case_tac "c <= 0", simp add: ring_simps)
+apply (simp add: ring_simps)
+apply simp
+apply simp
+done
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Multivariate_Analysis/Finite_Cartesian_Product.thy	Fri Oct 23 13:23:18 2009 +0200
@@ -0,0 +1,95 @@
+(* Title:      HOL/Library/Finite_Cartesian_Product
+   Author:     Amine Chaieb, University of Cambridge
+*)
+
+header {* Definition of finite Cartesian product types. *}
+
+theory Finite_Cartesian_Product
+imports Main (*FIXME: ATP_Linkup is only needed for metis at a few places. We could dispense of that by changing the proofs.*)
+begin
+
+definition hassize (infixr "hassize" 12) where
+  "(S hassize n) = (finite S \<and> card S = n)"
+
+lemma hassize_image_inj: assumes f: "inj_on f S" and S: "S hassize n"
+  shows "f ` S hassize n"
+  using f S card_image[OF f]
+    by (simp add: hassize_def inj_on_def)
+
+
+subsection {* Finite Cartesian products, with indexing and lambdas. *}
+
+typedef (open Cart)
+  ('a, 'b) "^" (infixl "^" 15)
+    = "UNIV :: ('b \<Rightarrow> 'a) set"
+  morphisms Cart_nth Cart_lambda ..
+
+notation Cart_nth (infixl "$" 90)
+
+notation (xsymbols) Cart_lambda (binder "\<chi>" 10)
+
+lemma stupid_ext: "(\<forall>x. f x = g x) \<longleftrightarrow> (f = g)"
+  apply auto
+  apply (rule ext)
+  apply auto
+  done
+
+lemma Cart_eq: "((x:: 'a ^ 'b) = y) \<longleftrightarrow> (\<forall>i. x$i = y$i)"
+  by (simp add: Cart_nth_inject [symmetric] expand_fun_eq)
+
+lemma Cart_lambda_beta [simp]: "Cart_lambda g $ i = g i"
+  by (simp add: Cart_lambda_inverse)
+
+lemma Cart_lambda_unique:
+  fixes f :: "'a ^ 'b"
+  shows "(\<forall>i. f$i = g i) \<longleftrightarrow> Cart_lambda g = f"
+  by (auto simp add: Cart_eq)
+
+lemma Cart_lambda_eta: "(\<chi> i. (g$i)) = g"
+  by (simp add: Cart_eq)
+
+text{* A non-standard sum to "paste" Cartesian products. *}
+
+definition pastecart :: "'a ^ 'm \<Rightarrow> 'a ^ 'n \<Rightarrow> 'a ^ ('m + 'n)" where
+  "pastecart f g = (\<chi> i. case i of Inl a \<Rightarrow> f$a | Inr b \<Rightarrow> g$b)"
+
+definition fstcart:: "'a ^('m + 'n) \<Rightarrow> 'a ^ 'm" where
+  "fstcart f = (\<chi> i. (f$(Inl i)))"
+
+definition sndcart:: "'a ^('m + 'n) \<Rightarrow> 'a ^ 'n" where
+  "sndcart f = (\<chi> i. (f$(Inr i)))"
+
+lemma nth_pastecart_Inl [simp]: "pastecart f g $ Inl a = f$a"
+  unfolding pastecart_def by simp
+
+lemma nth_pastecart_Inr [simp]: "pastecart f g $ Inr b = g$b"
+  unfolding pastecart_def by simp
+
+lemma nth_fstcart [simp]: "fstcart f $ i = f $ Inl i"
+  unfolding fstcart_def by simp
+
+lemma nth_sndtcart [simp]: "sndcart f $ i = f $ Inr i"
+  unfolding sndcart_def by simp
+
+lemma finite_sum_image: "(UNIV::('a + 'b) set) = range Inl \<union> range Inr"
+by (auto, case_tac x, auto)
+
+lemma fstcart_pastecart: "fstcart (pastecart (x::'a ^'m ) (y:: 'a ^ 'n)) = x"
+  by (simp add: Cart_eq)
+
+lemma sndcart_pastecart: "sndcart (pastecart (x::'a ^'m ) (y:: 'a ^ 'n)) = y"
+  by (simp add: Cart_eq)
+
+lemma pastecart_fst_snd: "pastecart (fstcart z) (sndcart z) = z"
+  by (simp add: Cart_eq pastecart_def fstcart_def sndcart_def split: sum.split)
+
+lemma pastecart_eq: "(x = y) \<longleftrightarrow> (fstcart x = fstcart y) \<and> (sndcart x = sndcart y)"
+  using pastecart_fst_snd[of x] pastecart_fst_snd[of y] by metis
+
+lemma forall_pastecart: "(\<forall>p. P p) \<longleftrightarrow> (\<forall>x y. P (pastecart x y))"
+  by (metis pastecart_fst_snd fstcart_pastecart sndcart_pastecart)
+
+lemma exists_pastecart: "(\<exists>p. P p)  \<longleftrightarrow> (\<exists>x y. P (pastecart x y))"
+  by (metis pastecart_fst_snd fstcart_pastecart sndcart_pastecart)
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Multivariate_Analysis/Multivariate_Analysis.thy	Fri Oct 23 13:23:18 2009 +0200
@@ -0,0 +1,6 @@
+theory Multivariate_Analysis imports
+	Convex_Euclidean_Space
+	Determinants
+begin
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Multivariate_Analysis/ROOT.ML	Fri Oct 23 13:23:18 2009 +0200
@@ -0,0 +1,6 @@
+(*
+  no_document use_thy "ThisTheory";
+  use_thy "ThatTheory";
+*)
+
+use_thy "Multivariate_Analysis";
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy	Fri Oct 23 13:23:18 2009 +0200
@@ -0,0 +1,6027 @@
+(*  Title:      HOL/Library/Topology_Euclidian_Space.thy
+    Author:     Amine Chaieb, University of Cambridge
+    Author:     Robert Himmelmann, TU Muenchen
+*)
+
+header {* Elementary topology in Euclidean space. *}
+
+theory Topology_Euclidean_Space
+imports SEQ Euclidean_Space Product_Vector
+begin
+
+declare fstcart_pastecart[simp] sndcart_pastecart[simp]
+
+subsection{* General notion of a topology *}
+
+definition "istopology L \<longleftrightarrow> {} \<in> L \<and> (\<forall>S \<in>L. \<forall>T \<in>L. S \<inter> T \<in> L) \<and> (\<forall>K. K \<subseteq>L \<longrightarrow> \<Union> K \<in> L)"
+typedef (open) 'a topology = "{L::('a set) set. istopology L}"
+  morphisms "openin" "topology"
+  unfolding istopology_def by blast
+
+lemma istopology_open_in[intro]: "istopology(openin U)"
+  using openin[of U] by blast
+
+lemma topology_inverse': "istopology U \<Longrightarrow> openin (topology U) = U"
+  using topology_inverse[unfolded mem_def Collect_def] .
+
+lemma topology_inverse_iff: "istopology U \<longleftrightarrow> openin (topology U) = U"
+  using topology_inverse[of U] istopology_open_in[of "topology U"] by auto
+
+lemma topology_eq: "T1 = T2 \<longleftrightarrow> (\<forall>S. openin T1 S \<longleftrightarrow> openin T2 S)"
+proof-
+  {assume "T1=T2" hence "\<forall>S. openin T1 S \<longleftrightarrow> openin T2 S" by simp}
+  moreover
+  {assume H: "\<forall>S. openin T1 S \<longleftrightarrow> openin T2 S"
+    hence "openin T1 = openin T2" by (metis mem_def set_ext)
+    hence "topology (openin T1) = topology (openin T2)" by simp
+    hence "T1 = T2" unfolding openin_inverse .}
+  ultimately show ?thesis by blast
+qed
+
+text{* Infer the "universe" from union of all sets in the topology. *}
+
+definition "topspace T =  \<Union>{S. openin T S}"
+
+subsection{* Main properties of open sets *}
+
+lemma openin_clauses:
+  fixes U :: "'a topology"
+  shows "openin U {}"
+  "\<And>S T. openin U S \<Longrightarrow> openin U T \<Longrightarrow> openin U (S\<inter>T)"
+  "\<And>K. (\<forall>S \<in> K. openin U S) \<Longrightarrow> openin U (\<Union>K)"
+  using openin[of U] unfolding istopology_def Collect_def mem_def
+  by (metis mem_def subset_eq)+
+
+lemma openin_subset[intro]: "openin U S \<Longrightarrow> S \<subseteq> topspace U"
+  unfolding topspace_def by blast
+lemma openin_empty[simp]: "openin U {}" by (simp add: openin_clauses)
+
+lemma openin_Int[intro]: "openin U S \<Longrightarrow> openin U T \<Longrightarrow> openin U (S \<inter> T)"
+  by (simp add: openin_clauses)
+
+lemma openin_Union[intro]: "(\<forall>S \<in>K. openin U S) \<Longrightarrow> openin U (\<Union> K)" by (simp add: openin_clauses)
+
+lemma openin_Un[intro]: "openin U S \<Longrightarrow> openin U T \<Longrightarrow> openin U (S \<union> T)"
+  using openin_Union[of "{S,T}" U] by auto
+
+lemma openin_topspace[intro, simp]: "openin U (topspace U)" by (simp add: openin_Union topspace_def)
+
+lemma openin_subopen: "openin U S \<longleftrightarrow> (\<forall>x \<in> S. \<exists>T. openin U T \<and> x \<in> T \<and> T \<subseteq> S)" (is "?lhs \<longleftrightarrow> ?rhs")
+proof-
+  {assume ?lhs then have ?rhs by auto }
+  moreover
+  {assume H: ?rhs
+    then obtain t where t: "\<forall>x\<in>S. openin U (t x) \<and> x \<in> t x \<and> t x \<subseteq> S"
+      unfolding Ball_def ex_simps(6)[symmetric] choice_iff by blast
+    from t have th0: "\<forall>x\<in> t`S. openin U x" by auto
+    have "\<Union> t`S = S" using t by auto
+    with openin_Union[OF th0] have "openin U S" by simp }
+  ultimately show ?thesis by blast
+qed
+
+subsection{* Closed sets *}
+
+definition "closedin U S \<longleftrightarrow> S \<subseteq> topspace U \<and> openin U (topspace U - S)"
+
+lemma closedin_subset: "closedin U S \<Longrightarrow> S \<subseteq> topspace U" by (metis closedin_def)
+lemma closedin_empty[simp]: "closedin U {}" by (simp add: closedin_def)
+lemma closedin_topspace[intro,simp]:
+  "closedin U (topspace U)" by (simp add: closedin_def)
+lemma closedin_Un[intro]: "closedin U S \<Longrightarrow> closedin U T \<Longrightarrow> closedin U (S \<union> T)"
+  by (auto simp add: Diff_Un closedin_def)
+
+lemma Diff_Inter[intro]: "A - \<Inter>S = \<Union> {A - s|s. s\<in>S}" by auto
+lemma closedin_Inter[intro]: assumes Ke: "K \<noteq> {}" and Kc: "\<forall>S \<in>K. closedin U S"
+  shows "closedin U (\<Inter> K)"  using Ke Kc unfolding closedin_def Diff_Inter by auto
+
+lemma closedin_Int[intro]: "closedin U S \<Longrightarrow> closedin U T \<Longrightarrow> closedin U (S \<inter> T)"
+  using closedin_Inter[of "{S,T}" U] by auto
+
+lemma Diff_Diff_Int: "A - (A - B) = A \<inter> B" by blast
+lemma openin_closedin_eq: "openin U S \<longleftrightarrow> S \<subseteq> topspace U \<and> closedin U (topspace U - S)"
+  apply (auto simp add: closedin_def Diff_Diff_Int inf_absorb2)
+  apply (metis openin_subset subset_eq)
+  done
+
+lemma openin_closedin:  "S \<subseteq> topspace U \<Longrightarrow> (openin U S \<longleftrightarrow> closedin U (topspace U - S))"
+  by (simp add: openin_closedin_eq)
+
+lemma openin_diff[intro]: assumes oS: "openin U S" and cT: "closedin U T" shows "openin U (S - T)"
+proof-
+  have "S - T = S \<inter> (topspace U - T)" using openin_subset[of U S]  oS cT
+    by (auto simp add: topspace_def openin_subset)
+  then show ?thesis using oS cT by (auto simp add: closedin_def)
+qed
+
+lemma closedin_diff[intro]: assumes oS: "closedin U S" and cT: "openin U T" shows "closedin U (S - T)"
+proof-
+  have "S - T = S \<inter> (topspace U - T)" using closedin_subset[of U S]  oS cT
+    by (auto simp add: topspace_def )
+  then show ?thesis using oS cT by (auto simp add: openin_closedin_eq)
+qed
+
+subsection{* Subspace topology. *}
+
+definition "subtopology U V = topology {S \<inter> V |S. openin U S}"
+
+lemma istopology_subtopology: "istopology {S \<inter> V |S. openin U S}" (is "istopology ?L")
+proof-
+  have "{} \<in> ?L" by blast
+  {fix A B assume A: "A \<in> ?L" and B: "B \<in> ?L"
+    from A B obtain Sa and Sb where Sa: "openin U Sa" "A = Sa \<inter> V" and Sb: "openin U Sb" "B = Sb \<inter> V" by blast
+    have "A\<inter>B = (Sa \<inter> Sb) \<inter> V" "openin U (Sa \<inter> Sb)"  using Sa Sb by blast+
+    then have "A \<inter> B \<in> ?L" by blast}
+  moreover
+  {fix K assume K: "K \<subseteq> ?L"
+    have th0: "?L = (\<lambda>S. S \<inter> V) ` openin U "
+      apply (rule set_ext)
+      apply (simp add: Ball_def image_iff)
+      by (metis mem_def)
+    from K[unfolded th0 subset_image_iff]
+    obtain Sk where Sk: "Sk \<subseteq> openin U" "K = (\<lambda>S. S \<inter> V) ` Sk" by blast
+    have "\<Union>K = (\<Union>Sk) \<inter> V" using Sk by auto
+    moreover have "openin U (\<Union> Sk)" using Sk by (auto simp add: subset_eq mem_def)
+    ultimately have "\<Union>K \<in> ?L" by blast}
+  ultimately show ?thesis unfolding istopology_def by blast
+qed
+
+lemma openin_subtopology:
+  "openin (subtopology U V) S \<longleftrightarrow> (\<exists> T. (openin U T) \<and> (S = T \<inter> V))"
+  unfolding subtopology_def topology_inverse'[OF istopology_subtopology]
+  by (auto simp add: Collect_def)
+
+lemma topspace_subtopology: "topspace(subtopology U V) = topspace U \<inter> V"
+  by (auto simp add: topspace_def openin_subtopology)
+
+lemma closedin_subtopology:
+  "closedin (subtopology U V) S \<longleftrightarrow> (\<exists>T. closedin U T \<and> S = T \<inter> V)"
+  unfolding closedin_def topspace_subtopology
+  apply (simp add: openin_subtopology)
+  apply (rule iffI)
+  apply clarify
+  apply (rule_tac x="topspace U - T" in exI)
+  by auto
+
+lemma openin_subtopology_refl: "openin (subtopology U V) V \<longleftrightarrow> V \<subseteq> topspace U"
+  unfolding openin_subtopology
+  apply (rule iffI, clarify)
+  apply (frule openin_subset[of U])  apply blast
+  apply (rule exI[where x="topspace U"])
+  by auto
+
+lemma subtopology_superset: assumes UV: "topspace U \<subseteq> V"
+  shows "subtopology U V = U"
+proof-
+  {fix S
+    {fix T assume T: "openin U T" "S = T \<inter> V"
+      from T openin_subset[OF T(1)] UV have eq: "S = T" by blast
+      have "openin U S" unfolding eq using T by blast}
+    moreover
+    {assume S: "openin U S"
+      hence "\<exists>T. openin U T \<and> S = T \<inter> V"
+        using openin_subset[OF S] UV by auto}
+    ultimately have "(\<exists>T. openin U T \<and> S = T \<inter> V) \<longleftrightarrow> openin U S" by blast}
+  then show ?thesis unfolding topology_eq openin_subtopology by blast
+qed
+
+
+lemma subtopology_topspace[simp]: "subtopology U (topspace U) = U"
+  by (simp add: subtopology_superset)
+
+lemma subtopology_UNIV[simp]: "subtopology U UNIV = U"
+  by (simp add: subtopology_superset)
+
+subsection{* The universal Euclidean versions are what we use most of the time *}
+
+definition
+  euclidean :: "'a::topological_space topology" where
+  "euclidean = topology open"
+
+lemma open_openin: "open S \<longleftrightarrow> openin euclidean S"
+  unfolding euclidean_def
+  apply (rule cong[where x=S and y=S])
+  apply (rule topology_inverse[symmetric])
+  apply (auto simp add: istopology_def)
+  by (auto simp add: mem_def subset_eq)
+
+lemma topspace_euclidean: "topspace euclidean = UNIV"
+  apply (simp add: topspace_def)
+  apply (rule set_ext)
+  by (auto simp add: open_openin[symmetric])
+
+lemma topspace_euclidean_subtopology[simp]: "topspace (subtopology euclidean S) = S"
+  by (simp add: topspace_euclidean topspace_subtopology)
+
+lemma closed_closedin: "closed S \<longleftrightarrow> closedin euclidean S"
+  by (simp add: closed_def closedin_def topspace_euclidean open_openin Compl_eq_Diff_UNIV)
+
+lemma open_subopen: "open S \<longleftrightarrow> (\<forall>x\<in>S. \<exists>T. open T \<and> x \<in> T \<and> T \<subseteq> S)"
+  by (simp add: open_openin openin_subopen[symmetric])
+
+subsection{* Open and closed balls. *}
+
+definition
+  ball :: "'a::metric_space \<Rightarrow> real \<Rightarrow> 'a set" where
+  "ball x e = {y. dist x y < e}"
+
+definition
+  cball :: "'a::metric_space \<Rightarrow> real \<Rightarrow> 'a set" where
+  "cball x e = {y. dist x y \<le> e}"
+
+lemma mem_ball[simp]: "y \<in> ball x e \<longleftrightarrow> dist x y < e" by (simp add: ball_def)
+lemma mem_cball[simp]: "y \<in> cball x e \<longleftrightarrow> dist x y \<le> e" by (simp add: cball_def)
+
+lemma mem_ball_0 [simp]:
+  fixes x :: "'a::real_normed_vector"
+  shows "x \<in> ball 0 e \<longleftrightarrow> norm x < e"
+  by (simp add: dist_norm)
+
+lemma mem_cball_0 [simp]:
+  fixes x :: "'a::real_normed_vector"
+  shows "x \<in> cball 0 e \<longleftrightarrow> norm x \<le> e"
+  by (simp add: dist_norm)
+
+lemma centre_in_cball[simp]: "x \<in> cball x e \<longleftrightarrow> 0\<le> e"  by simp
+lemma ball_subset_cball[simp,intro]: "ball x e \<subseteq> cball x e" by (simp add: subset_eq)
+lemma subset_ball[intro]: "d <= e ==> ball x d \<subseteq> ball x e" by (simp add: subset_eq)
+lemma subset_cball[intro]: "d <= e ==> cball x d \<subseteq> cball x e" by (simp add: subset_eq)
+lemma ball_max_Un: "ball a (max r s) = ball a r \<union> ball a s"
+  by (simp add: expand_set_eq) arith
+
+lemma ball_min_Int: "ball a (min r s) = ball a r \<inter> ball a s"
+  by (simp add: expand_set_eq)
+
+subsection{* Topological properties of open balls *}
+
+lemma diff_less_iff: "(a::real) - b > 0 \<longleftrightarrow> a > b"
+  "(a::real) - b < 0 \<longleftrightarrow> a < b"
+  "a - b < c \<longleftrightarrow> a < c +b" "a - b > c \<longleftrightarrow> a > c +b" by arith+
+lemma diff_le_iff: "(a::real) - b \<ge> 0 \<longleftrightarrow> a \<ge> b" "(a::real) - b \<le> 0 \<longleftrightarrow> a \<le> b"
+  "a - b \<le> c \<longleftrightarrow> a \<le> c +b" "a - b \<ge> c \<longleftrightarrow> a \<ge> c +b"  by arith+
+
+lemma open_ball[intro, simp]: "open (ball x e)"
+  unfolding open_dist ball_def Collect_def Ball_def mem_def
+  unfolding dist_commute
+  apply clarify
+  apply (rule_tac x="e - dist xa x" in exI)
+  using dist_triangle_alt[where z=x]
+  apply (clarsimp simp add: diff_less_iff)
+  apply atomize
+  apply (erule_tac x="y" in allE)
+  apply (erule_tac x="xa" in allE)
+  by arith
+
+lemma centre_in_ball[simp]: "x \<in> ball x e \<longleftrightarrow> e > 0" by (metis mem_ball dist_self)
+lemma open_contains_ball: "open S \<longleftrightarrow> (\<forall>x\<in>S. \<exists>e>0. ball x e \<subseteq> S)"
+  unfolding open_dist subset_eq mem_ball Ball_def dist_commute ..
+
+lemma open_contains_ball_eq: "open S \<Longrightarrow> \<forall>x. x\<in>S \<longleftrightarrow> (\<exists>e>0. ball x e \<subseteq> S)"
+  by (metis open_contains_ball subset_eq centre_in_ball)
+
+lemma ball_eq_empty[simp]: "ball x e = {} \<longleftrightarrow> e \<le> 0"
+  unfolding mem_ball expand_set_eq
+  apply (simp add: not_less)
+  by (metis zero_le_dist order_trans dist_self)
+
+lemma ball_empty[intro]: "e \<le> 0 ==> ball x e = {}" by simp
+
+subsection{* Basic "localization" results are handy for connectedness. *}
+
+lemma openin_open: "openin (subtopology euclidean U) S \<longleftrightarrow> (\<exists>T. open T \<and> (S = U \<inter> T))"
+  by (auto simp add: openin_subtopology open_openin[symmetric])
+
+lemma openin_open_Int[intro]: "open S \<Longrightarrow> openin (subtopology euclidean U) (U \<inter> S)"
+  by (auto simp add: openin_open)
+
+lemma open_openin_trans[trans]:
+ "open S \<Longrightarrow> open T \<Longrightarrow> T \<subseteq> S \<Longrightarrow> openin (subtopology euclidean S) T"
+  by (metis Int_absorb1  openin_open_Int)
+
+lemma open_subset:  "S \<subseteq> T \<Longrightarrow> open S \<Longrightarrow> openin (subtopology euclidean T) S"
+  by (auto simp add: openin_open)
+
+lemma closedin_closed: "closedin (subtopology euclidean U) S \<longleftrightarrow> (\<exists>T. closed T \<and> S = U \<inter> T)"
+  by (simp add: closedin_subtopology closed_closedin Int_ac)
+
+lemma closedin_closed_Int: "closed S ==> closedin (subtopology euclidean U) (U \<inter> S)"
+  by (metis closedin_closed)
+
+lemma closed_closedin_trans: "closed S \<Longrightarrow> closed T \<Longrightarrow> T \<subseteq> S \<Longrightarrow> closedin (subtopology euclidean S) T"
+  apply (subgoal_tac "S \<inter> T = T" )
+  apply auto
+  apply (frule closedin_closed_Int[of T S])
+  by simp
+
+lemma closed_subset: "S \<subseteq> T \<Longrightarrow> closed S \<Longrightarrow> closedin (subtopology euclidean T) S"
+  by (auto simp add: closedin_closed)
+
+lemma openin_euclidean_subtopology_iff:
+  fixes S U :: "'a::metric_space set"
+  shows "openin (subtopology euclidean U) S
+  \<longleftrightarrow> S \<subseteq> U \<and> (\<forall>x\<in>S. \<exists>e>0. \<forall>x'\<in>U. dist x' x < e \<longrightarrow> x'\<in> S)" (is "?lhs \<longleftrightarrow> ?rhs")
+proof-
+  {assume ?lhs hence ?rhs unfolding openin_subtopology open_openin[symmetric]
+      by (simp add: open_dist) blast}
+  moreover
+  {assume SU: "S \<subseteq> U" and H: "\<And>x. x \<in> S \<Longrightarrow> \<exists>e>0. \<forall>x'\<in>U. dist x' x < e \<longrightarrow> x' \<in> S"
+    from H obtain d where d: "\<And>x . x\<in> S \<Longrightarrow> d x > 0 \<and> (\<forall>x' \<in> U. dist x' x < d x \<longrightarrow> x' \<in> S)"
+      by metis
+    let ?T = "\<Union>{B. \<exists>x\<in>S. B = ball x (d x)}"
+    have oT: "open ?T" by auto
+    { fix x assume "x\<in>S"
+      hence "x \<in> \<Union>{B. \<exists>x\<in>S. B = ball x (d x)}"
+        apply simp apply(rule_tac x="ball x(d x)" in exI) apply auto
+        by (rule d [THEN conjunct1])
+      hence "x\<in> ?T \<inter> U" using SU and `x\<in>S` by auto  }
+    moreover
+    { fix y assume "y\<in>?T"
+      then obtain B where "y\<in>B" "B\<in>{B. \<exists>x\<in>S. B = ball x (d x)}" by auto
+      then obtain x where "x\<in>S" and x:"y \<in> ball x (d x)" by auto
+      assume "y\<in>U"
+      hence "y\<in>S" using d[OF `x\<in>S`] and x by(auto simp add: dist_commute) }
+    ultimately have "S = ?T \<inter> U" by blast
+    with oT have ?lhs unfolding openin_subtopology open_openin[symmetric] by blast}
+  ultimately show ?thesis by blast
+qed
+
+text{* These "transitivity" results are handy too. *}
+
+lemma openin_trans[trans]: "openin (subtopology euclidean T) S \<Longrightarrow> openin (subtopology euclidean U) T
+  \<Longrightarrow> openin (subtopology euclidean U) S"
+  unfolding open_openin openin_open by blast
+
+lemma openin_open_trans: "openin (subtopology euclidean T) S \<Longrightarrow> open T \<Longrightarrow> open S"
+  by (auto simp add: openin_open intro: openin_trans)
+
+lemma closedin_trans[trans]:
+ "closedin (subtopology euclidean T) S \<Longrightarrow>
+           closedin (subtopology euclidean U) T
+           ==> closedin (subtopology euclidean U) S"
+  by (auto simp add: closedin_closed closed_closedin closed_Inter Int_assoc)
+
+lemma closedin_closed_trans: "closedin (subtopology euclidean T) S \<Longrightarrow> closed T \<Longrightarrow> closed S"
+  by (auto simp add: closedin_closed intro: closedin_trans)
+
+subsection{* Connectedness *}
+
+definition "connected S \<longleftrightarrow>
+  ~(\<exists>e1 e2. open e1 \<and> open e2 \<and> S \<subseteq> (e1 \<union> e2) \<and> (e1 \<inter> e2 \<inter> S = {})
+  \<and> ~(e1 \<inter> S = {}) \<and> ~(e2 \<inter> S = {}))"
+
+lemma connected_local:
+ "connected S \<longleftrightarrow> ~(\<exists>e1 e2.
+                 openin (subtopology euclidean S) e1 \<and>
+                 openin (subtopology euclidean S) e2 \<and>
+                 S \<subseteq> e1 \<union> e2 \<and>
+                 e1 \<inter> e2 = {} \<and>
+                 ~(e1 = {}) \<and>
+                 ~(e2 = {}))"
+unfolding connected_def openin_open by (safe, blast+)
+
+lemma exists_diff: "(\<exists>S. P(UNIV - S)) \<longleftrightarrow> (\<exists>S. P S)" (is "?lhs \<longleftrightarrow> ?rhs")
+proof-
+
+  {assume "?lhs" hence ?rhs by blast }
+  moreover
+  {fix S assume H: "P S"
+    have "S = UNIV - (UNIV - S)" by auto
+    with H have "P (UNIV - (UNIV - S))" by metis }
+  ultimately show ?thesis by metis
+qed
+
+lemma connected_clopen: "connected S \<longleftrightarrow>
+        (\<forall>T. openin (subtopology euclidean S) T \<and>
+            closedin (subtopology euclidean S) T \<longrightarrow> T = {} \<or> T = S)" (is "?lhs \<longleftrightarrow> ?rhs")
+proof-
+  have " \<not> connected S \<longleftrightarrow> (\<exists>e1 e2. open e1 \<and> open (UNIV - e2) \<and> S \<subseteq> e1 \<union> (UNIV - e2) \<and> e1 \<inter> (UNIV - e2) \<inter> S = {} \<and> e1 \<inter> S \<noteq> {} \<and> (UNIV - e2) \<inter> S \<noteq> {})"
+    unfolding connected_def openin_open closedin_closed
+    apply (subst exists_diff) by blast
+  hence th0: "connected S \<longleftrightarrow> \<not> (\<exists>e2 e1. closed e2 \<and> open e1 \<and> S \<subseteq> e1 \<union> (UNIV - e2) \<and> e1 \<inter> (UNIV - e2) \<inter> S = {} \<and> e1 \<inter> S \<noteq> {} \<and> (UNIV - e2) \<inter> S \<noteq> {})"
+    (is " _ \<longleftrightarrow> \<not> (\<exists>e2 e1. ?P e2 e1)") apply (simp add: closed_def Compl_eq_Diff_UNIV) by metis
+
+  have th1: "?rhs \<longleftrightarrow> \<not> (\<exists>t' t. closed t'\<and>t = S\<inter>t' \<and> t\<noteq>{} \<and> t\<noteq>S \<and> (\<exists>t'. open t' \<and> t = S \<inter> t'))"
+    (is "_ \<longleftrightarrow> \<not> (\<exists>t' t. ?Q t' t)")
+    unfolding connected_def openin_open closedin_closed by auto
+  {fix e2
+    {fix e1 have "?P e2 e1 \<longleftrightarrow> (\<exists>t.  closed e2 \<and> t = S\<inter>e2 \<and> open e1 \<and> t = S\<inter>e1 \<and> t\<noteq>{} \<and> t\<noteq>S)"
+        by auto}
+    then have "(\<exists>e1. ?P e2 e1) \<longleftrightarrow> (\<exists>t. ?Q e2 t)" by metis}
+  then have "\<forall>e2. (\<exists>e1. ?P e2 e1) \<longleftrightarrow> (\<exists>t. ?Q e2 t)" by blast
+  then show ?thesis unfolding th0 th1 by simp
+qed
+
+lemma connected_empty[simp, intro]: "connected {}"
+  by (simp add: connected_def)
+
+subsection{* Hausdorff and other separation properties *}
+
+class t0_space =
+  assumes t0_space: "x \<noteq> y \<Longrightarrow> \<exists>U. open U \<and> \<not> (x \<in> U \<longleftrightarrow> y \<in> U)"
+
+class t1_space =
+  assumes t1_space: "x \<noteq> y \<Longrightarrow> \<exists>U V. open U \<and> open V \<and> x \<in> U \<and> y \<notin> U \<and> x \<notin> V \<and> y \<in> V"
+begin
+
+subclass t0_space
+proof
+qed (fast dest: t1_space)
+
+end
+
+text {* T2 spaces are also known as Hausdorff spaces. *}
+
+class t2_space =
+  assumes hausdorff: "x \<noteq> y \<Longrightarrow> \<exists>U V. open U \<and> open V \<and> x \<in> U \<and> y \<in> V \<and> U \<inter> V = {}"
+begin
+
+subclass t1_space
+proof
+qed (fast dest: hausdorff)
+
+end
+
+instance metric_space \<subseteq> t2_space
+proof
+  fix x y :: "'a::metric_space"
+  assume xy: "x \<noteq> y"
+  let ?U = "ball x (dist x y / 2)"
+  let ?V = "ball y (dist x y / 2)"
+  have th0: "\<And>d x y z. (d x z :: real) <= d x y + d y z \<Longrightarrow> d y z = d z y
+               ==> ~(d x y * 2 < d x z \<and> d z y * 2 < d x z)" by arith
+  have "open ?U \<and> open ?V \<and> x \<in> ?U \<and> y \<in> ?V \<and> ?U \<inter> ?V = {}"
+    using dist_pos_lt[OF xy] th0[of dist,OF dist_triangle dist_commute]
+    by (auto simp add: expand_set_eq)
+  then show "\<exists>U V. open U \<and> open V \<and> x \<in> U \<and> y \<in> V \<and> U \<inter> V = {}"
+    by blast
+qed
+
+lemma separation_t2:
+  fixes x y :: "'a::t2_space"
+  shows "x \<noteq> y \<longleftrightarrow> (\<exists>U V. open U \<and> open V \<and> x \<in> U \<and> y \<in> V \<and> U \<inter> V = {})"
+  using hausdorff[of x y] by blast
+
+lemma separation_t1:
+  fixes x y :: "'a::t1_space"
+  shows "x \<noteq> y \<longleftrightarrow> (\<exists>U V. open U \<and> open V \<and> x \<in>U \<and> y\<notin> U \<and> x\<notin>V \<and> y\<in>V)"
+  using t1_space[of x y] by blast
+
+lemma separation_t0:
+  fixes x y :: "'a::t0_space"
+  shows "x \<noteq> y \<longleftrightarrow> (\<exists>U. open U \<and> ~(x\<in>U \<longleftrightarrow> y\<in>U))"
+  using t0_space[of x y] by blast
+
+subsection{* Limit points *}
+
+definition
+  islimpt:: "'a::topological_space \<Rightarrow> 'a set \<Rightarrow> bool"
+    (infixr "islimpt" 60) where
+  "x islimpt S \<longleftrightarrow> (\<forall>T. x\<in>T \<longrightarrow> open T \<longrightarrow> (\<exists>y\<in>S. y\<in>T \<and> y\<noteq>x))"
+
+lemma islimptI:
+  assumes "\<And>T. x \<in> T \<Longrightarrow> open T \<Longrightarrow> \<exists>y\<in>S. y \<in> T \<and> y \<noteq> x"
+  shows "x islimpt S"
+  using assms unfolding islimpt_def by auto
+
+lemma islimptE:
+  assumes "x islimpt S" and "x \<in> T" and "open T"
+  obtains y where "y \<in> S" and "y \<in> T" and "y \<noteq> x"
+  using assms unfolding islimpt_def by auto
+
+lemma islimpt_subset: "x islimpt S \<Longrightarrow> S \<subseteq> T ==> x islimpt T" by (auto simp add: islimpt_def)
+
+lemma islimpt_approachable:
+  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)"
+  unfolding islimpt_def
+  apply auto
+  apply(erule_tac x="ball x e" in allE)
+  apply auto
+  apply(rule_tac x=y in bexI)
+  apply (auto simp add: dist_commute)
+  apply (simp add: open_dist, drule (1) bspec)
+  apply (clarify, drule spec, drule (1) mp, auto)
+  done
+
+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)"
+  unfolding islimpt_approachable
+  using approachable_lt_le[where f="\<lambda>x'. dist x' x" and P="\<lambda>x'. \<not> (x'\<in>S \<and> x'\<noteq>x)"]
+  by metis (* FIXME: VERY slow! *)
+
+class perfect_space =
+  (* FIXME: perfect_space should inherit from topological_space *)
+  assumes islimpt_UNIV [simp, intro]: "(x::'a::metric_space) islimpt UNIV"
+
+lemma perfect_choose_dist:
+  fixes x :: "'a::perfect_space"
+  shows "0 < r \<Longrightarrow> \<exists>a. a \<noteq> x \<and> dist a x < r"
+using islimpt_UNIV [of x]
+by (simp add: islimpt_approachable)
+
+instance real :: perfect_space
+apply default
+apply (rule islimpt_approachable [THEN iffD2])
+apply (clarify, rule_tac x="x + e/2" in bexI)
+apply (auto simp add: dist_norm)
+done
+
+instance "^" :: (perfect_space, finite) perfect_space
+proof
+  fix x :: "'a ^ 'b"
+  {
+    fix e :: real assume "0 < e"
+    def a \<equiv> "x $ undefined"
+    have "a islimpt UNIV" by (rule islimpt_UNIV)
+    with `0 < e` obtain b where "b \<noteq> a" and "dist b a < e"
+      unfolding islimpt_approachable by auto
+    def y \<equiv> "Cart_lambda ((Cart_nth x)(undefined := b))"
+    from `b \<noteq> a` have "y \<noteq> x"
+      unfolding a_def y_def by (simp add: Cart_eq)
+    from `dist b a < e` have "dist y x < e"
+      unfolding dist_vector_def a_def y_def
+      apply simp
+      apply (rule le_less_trans [OF setL2_le_setsum [OF zero_le_dist]])
+      apply (subst setsum_diff1' [where a=undefined], simp, simp, simp)
+      done
+    from `y \<noteq> x` and `dist y x < e`
+    have "\<exists>y\<in>UNIV. y \<noteq> x \<and> dist y x < e" by auto
+  }
+  then show "x islimpt UNIV" unfolding islimpt_approachable by blast
+qed
+
+lemma closed_limpt: "closed S \<longleftrightarrow> (\<forall>x. x islimpt S \<longrightarrow> x \<in> S)"
+  unfolding closed_def
+  apply (subst open_subopen)
+  apply (simp add: islimpt_def subset_eq Compl_eq_Diff_UNIV)
+  by (metis DiffE DiffI UNIV_I insertCI insert_absorb mem_def)
+
+lemma islimpt_EMPTY[simp]: "\<not> x islimpt {}"
+  unfolding islimpt_def by auto
+
+lemma closed_positive_orthant: "closed {x::real^'n::finite. \<forall>i. 0 \<le>x$i}"
+proof-
+  let ?U = "UNIV :: 'n set"
+  let ?O = "{x::real^'n. \<forall>i. x$i\<ge>0}"
+  {fix x:: "real^'n" and i::'n assume H: "\<forall>e>0. \<exists>x'\<in>?O. x' \<noteq> x \<and> dist x' x < e"
+    and xi: "x$i < 0"
+    from xi have th0: "-x$i > 0" by arith
+    from H[rule_format, OF th0] obtain x' where x': "x' \<in>?O" "x' \<noteq> x" "dist x' x < -x $ i" by blast
+      have th:" \<And>b a (x::real). abs x <= b \<Longrightarrow> b <= a ==> ~(a + x < 0)" by arith
+      have th': "\<And>x (y::real). x < 0 \<Longrightarrow> 0 <= y ==> abs x <= abs (y - x)" by arith
+      have th1: "\<bar>x$i\<bar> \<le> \<bar>(x' - x)$i\<bar>" using x'(1) xi
+        apply (simp only: vector_component)
+        by (rule th') auto
+      have th2: "\<bar>dist x x'\<bar> \<ge> \<bar>(x' - x)$i\<bar>" using  component_le_norm[of "x'-x" i]
+        apply (simp add: dist_norm) by norm
+      from th[OF th1 th2] x'(3) have False by (simp add: dist_commute) }
+  then show ?thesis unfolding closed_limpt islimpt_approachable
+    unfolding not_le[symmetric] by blast
+qed
+
+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"
+proof(induct rule: finite_induct[OF fS])
+  case 1 thus ?case apply auto by ferrack
+next
+  case (2 x F)
+  from 2 obtain d where d: "d >0" "\<forall>x\<in>F. x\<noteq>a \<longrightarrow> d \<le> dist a x" by blast
+  {assume "x = a" hence ?case using d by auto  }
+  moreover
+  {assume xa: "x\<noteq>a"
+    let ?d = "min d (dist a x)"
+    have dp: "?d > 0" using xa d(1) using dist_nz by auto
+    from d have d': "\<forall>x\<in>F. x\<noteq>a \<longrightarrow> ?d \<le> dist a x" by auto
+    with dp xa have ?case by(auto intro!: exI[where x="?d"]) }
+  ultimately show ?case by blast
+qed
+
+lemma islimpt_finite:
+  fixes S :: "'a::metric_space set"
+  assumes fS: "finite S" shows "\<not> a islimpt S"
+  unfolding islimpt_approachable
+  using finite_set_avoid[OF fS, of a] by (metis dist_commute  not_le)
+
+lemma islimpt_Un: "x islimpt (S \<union> T) \<longleftrightarrow> x islimpt S \<or> x islimpt T"
+  apply (rule iffI)
+  defer
+  apply (metis Un_upper1 Un_upper2 islimpt_subset)
+  unfolding islimpt_def
+  apply (rule ccontr, clarsimp, rename_tac A B)
+  apply (drule_tac x="A \<inter> B" in spec)
+  apply (auto simp add: open_Int)
+  done
+
+lemma discrete_imp_closed:
+  fixes S :: "'a::metric_space set"
+  assumes e: "0 < e" and d: "\<forall>x \<in> S. \<forall>y \<in> S. dist y x < e \<longrightarrow> y = x"
+  shows "closed S"
+proof-
+  {fix x assume C: "\<forall>e>0. \<exists>x'\<in>S. x' \<noteq> x \<and> dist x' x < e"
+    from e have e2: "e/2 > 0" by arith
+    from C[rule_format, OF e2] obtain y where y: "y \<in> S" "y\<noteq>x" "dist y x < e/2" by blast
+    let ?m = "min (e/2) (dist x y) "
+    from e2 y(2) have mp: "?m > 0" by (simp add: dist_nz[THEN sym])
+    from C[rule_format, OF mp] obtain z where z: "z \<in> S" "z\<noteq>x" "dist z x < ?m" by blast
+    have th: "dist z y < e" using z y
+      by (intro dist_triangle_lt [where z=x], simp)
+    from d[rule_format, OF y(1) z(1) th] y z
+    have False by (auto simp add: dist_commute)}
+  then show ?thesis by (metis islimpt_approachable closed_limpt [where 'a='a])
+qed
+
+subsection{* Interior of a Set *}
+definition "interior S = {x. \<exists>T. open T \<and> x \<in> T \<and> T \<subseteq> S}"
+
+lemma interior_eq: "interior S = S \<longleftrightarrow> open S"
+  apply (simp add: expand_set_eq interior_def)
+  apply (subst (2) open_subopen) by (safe, blast+)
+
+lemma interior_open: "open S ==> (interior S = S)" by (metis interior_eq)
+
+lemma interior_empty[simp]: "interior {} = {}" by (simp add: interior_def)
+
+lemma open_interior[simp, intro]: "open(interior S)"
+  apply (simp add: interior_def)
+  apply (subst open_subopen) by blast
+
+lemma interior_interior[simp]: "interior(interior S) = interior S" by (metis interior_eq open_interior)
+lemma interior_subset: "interior S \<subseteq> S" by (auto simp add: interior_def)
+lemma subset_interior: "S \<subseteq> T ==> (interior S) \<subseteq> (interior T)" by (auto simp add: interior_def)
+lemma interior_maximal: "T \<subseteq> S \<Longrightarrow> open T ==> T \<subseteq> (interior S)" by (auto simp add: interior_def)
+lemma interior_unique: "T \<subseteq> S \<Longrightarrow> open T  \<Longrightarrow> (\<forall>T'. T' \<subseteq> S \<and> open T' \<longrightarrow> T' \<subseteq> T) \<Longrightarrow> interior S = T"
+  by (metis equalityI interior_maximal interior_subset open_interior)
+lemma mem_interior: "x \<in> interior S \<longleftrightarrow> (\<exists>e. 0 < e \<and> ball x e \<subseteq> S)"
+  apply (simp add: interior_def)
+  by (metis open_contains_ball centre_in_ball open_ball subset_trans)
+
+lemma open_subset_interior: "open S ==> S \<subseteq> interior T \<longleftrightarrow> S \<subseteq> T"
+  by (metis interior_maximal interior_subset subset_trans)
+
+lemma interior_inter[simp]: "interior(S \<inter> T) = interior S \<inter> interior T"
+  apply (rule equalityI, simp)
+  apply (metis Int_lower1 Int_lower2 subset_interior)
+  by (metis Int_mono interior_subset open_Int open_interior open_subset_interior)
+
+lemma interior_limit_point [intro]:
+  fixes x :: "'a::perfect_space"
+  assumes x: "x \<in> interior S" shows "x islimpt S"
+proof-
+  from x obtain e where e: "e>0" "\<forall>x'. dist x x' < e \<longrightarrow> x' \<in> S"
+    unfolding mem_interior subset_eq Ball_def mem_ball by blast
+  {
+    fix d::real assume d: "d>0"
+    let ?m = "min d e"
+    have mde2: "0 < ?m" using e(1) d(1) by simp
+    from perfect_choose_dist [OF mde2, of x]
+    obtain y where "y \<noteq> x" and "dist y x < ?m" by blast
+    then have "dist y x < e" "dist y x < d" by simp_all
+    from `dist y x < e` e(2) have "y \<in> S" by (simp add: dist_commute)
+    have "\<exists>x'\<in>S. x'\<noteq> x \<and> dist x' x < d"
+      using `y \<in> S` `y \<noteq> x` `dist y x < d` by fast
+  }
+  then show ?thesis unfolding islimpt_approachable by blast
+qed
+
+lemma interior_closed_Un_empty_interior:
+  assumes cS: "closed S" and iT: "interior T = {}"
+  shows "interior(S \<union> T) = interior S"
+proof
+  show "interior S \<subseteq> interior (S\<union>T)"
+    by (rule subset_interior, blast)
+next
+  show "interior (S \<union> T) \<subseteq> interior S"
+  proof
+    fix x assume "x \<in> interior (S \<union> T)"
+    then obtain R where "open R" "x \<in> R" "R \<subseteq> S \<union> T"
+      unfolding interior_def by fast
+    show "x \<in> interior S"
+    proof (rule ccontr)
+      assume "x \<notin> interior S"
+      with `x \<in> R` `open R` obtain y where "y \<in> R - S"
+        unfolding interior_def expand_set_eq by fast
+      from `open R` `closed S` have "open (R - S)" by (rule open_Diff)
+      from `R \<subseteq> S \<union> T` have "R - S \<subseteq> T" by fast
+      from `y \<in> R - S` `open (R - S)` `R - S \<subseteq> T` `interior T = {}`
+      show "False" unfolding interior_def by fast
+    qed
+  qed
+qed
+
+
+subsection{* Closure of a Set *}
+
+definition "closure S = S \<union> {x | x. x islimpt S}"
+
+lemma closure_interior: "closure S = UNIV - interior (UNIV - S)"
+proof-
+  { fix x
+    have "x\<in>UNIV - interior (UNIV - S) \<longleftrightarrow> x \<in> closure S"  (is "?lhs = ?rhs")
+    proof
+      let ?exT = "\<lambda> y. (\<exists>T. open T \<and> y \<in> T \<and> T \<subseteq> UNIV - S)"
+      assume "?lhs"
+      hence *:"\<not> ?exT x"
+        unfolding interior_def
+        by simp
+      { assume "\<not> ?rhs"
+        hence False using *
+          unfolding closure_def islimpt_def
+          by blast
+      }
+      thus "?rhs"
+        by blast
+    next
+      assume "?rhs" thus "?lhs"
+        unfolding closure_def interior_def islimpt_def
+        by blast
+    qed
+  }
+  thus ?thesis
+    by blast
+qed
+
+lemma interior_closure: "interior S = UNIV - (closure (UNIV - S))"
+proof-
+  { fix x
+    have "x \<in> interior S \<longleftrightarrow> x \<in> UNIV - (closure (UNIV - S))"
+      unfolding interior_def closure_def islimpt_def
+      by blast (* FIXME: VERY slow! *)
+  }
+  thus ?thesis
+    by blast
+qed
+
+lemma closed_closure[simp, intro]: "closed (closure S)"
+proof-
+  have "closed (UNIV - interior (UNIV -S))" by blast
+  thus ?thesis using closure_interior[of S] by simp
+qed
+
+lemma closure_hull: "closure S = closed hull S"
+proof-
+  have "S \<subseteq> closure S"
+    unfolding closure_def
+    by blast
+  moreover
+  have "closed (closure S)"
+    using closed_closure[of S]
+    by assumption
+  moreover
+  { fix t
+    assume *:"S \<subseteq> t" "closed t"
+    { fix x
+      assume "x islimpt S"
+      hence "x islimpt t" using *(1)
+        using islimpt_subset[of x, of S, of t]
+        by blast
+    }
+    with * have "closure S \<subseteq> t"
+      unfolding closure_def
+      using closed_limpt[of t]
+      by auto
+  }
+  ultimately show ?thesis
+    using hull_unique[of S, of "closure S", of closed]
+    unfolding mem_def
+    by simp
+qed
+
+lemma closure_eq: "closure S = S \<longleftrightarrow> closed S"
+  unfolding closure_hull
+  using hull_eq[of closed, unfolded mem_def, OF  closed_Inter, of S]
+  by (metis mem_def subset_eq)
+
+lemma closure_closed[simp]: "closed S \<Longrightarrow> closure S = S"
+  using closure_eq[of S]
+  by simp
+
+lemma closure_closure[simp]: "closure (closure S) = closure S"
+  unfolding closure_hull
+  using hull_hull[of closed S]
+  by assumption
+
+lemma closure_subset: "S \<subseteq> closure S"
+  unfolding closure_hull
+  using hull_subset[of S closed]
+  by assumption
+
+lemma subset_closure: "S \<subseteq> T \<Longrightarrow> closure S \<subseteq> closure T"
+  unfolding closure_hull
+  using hull_mono[of S T closed]
+  by assumption
+
+lemma closure_minimal: "S \<subseteq> T \<Longrightarrow>  closed T \<Longrightarrow> closure S \<subseteq> T"
+  using hull_minimal[of S T closed]
+  unfolding closure_hull mem_def
+  by simp
+
+lemma closure_unique: "S \<subseteq> T \<and> closed T \<and> (\<forall> T'. S \<subseteq> T' \<and> closed T' \<longrightarrow> T \<subseteq> T') \<Longrightarrow> closure S = T"
+  using hull_unique[of S T closed]
+  unfolding closure_hull mem_def
+  by simp
+
+lemma closure_empty[simp]: "closure {} = {}"
+  using closed_empty closure_closed[of "{}"]
+  by simp
+
+lemma closure_univ[simp]: "closure UNIV = UNIV"
+  using closure_closed[of UNIV]
+  by simp
+
+lemma closure_eq_empty: "closure S = {} \<longleftrightarrow> S = {}"
+  using closure_empty closure_subset[of S]
+  by blast
+
+lemma closure_subset_eq: "closure S \<subseteq> S \<longleftrightarrow> closed S"
+  using closure_eq[of S] closure_subset[of S]
+  by simp
+
+lemma open_inter_closure_eq_empty:
+  "open S \<Longrightarrow> (S \<inter> closure T) = {} \<longleftrightarrow> S \<inter> T = {}"
+  using open_subset_interior[of S "UNIV - T"]
+  using interior_subset[of "UNIV - T"]
+  unfolding closure_interior
+  by auto
+
+lemma open_inter_closure_subset:
+  "open S \<Longrightarrow> (S \<inter> (closure T)) \<subseteq> closure(S \<inter> T)"
+proof
+  fix x
+  assume as: "open S" "x \<in> S \<inter> closure T"
+  { assume *:"x islimpt T"
+    have "x islimpt (S \<inter> T)"
+    proof (rule islimptI)
+      fix A
+      assume "x \<in> A" "open A"
+      with as have "x \<in> A \<inter> S" "open (A \<inter> S)"
+        by (simp_all add: open_Int)
+      with * obtain y where "y \<in> T" "y \<in> A \<inter> S" "y \<noteq> x"
+        by (rule islimptE)
+      hence "y \<in> S \<inter> T" "y \<in> A \<and> y \<noteq> x"
+        by simp_all
+      thus "\<exists>y\<in>(S \<inter> T). y \<in> A \<and> y \<noteq> x" ..
+    qed
+  }
+  then show "x \<in> closure (S \<inter> T)" using as
+    unfolding closure_def
+    by blast
+qed
+
+lemma closure_complement: "closure(UNIV - S) = UNIV - interior(S)"
+proof-
+  have "S = UNIV - (UNIV - S)"
+    by auto
+  thus ?thesis
+    unfolding closure_interior
+    by auto
+qed
+
+lemma interior_complement: "interior(UNIV - S) = UNIV - closure(S)"
+  unfolding closure_interior
+  by blast
+
+subsection{* Frontier (aka boundary) *}
+
+definition "frontier S = closure S - interior S"
+
+lemma frontier_closed: "closed(frontier S)"
+  by (simp add: frontier_def closed_Diff)
+
+lemma frontier_closures: "frontier S = (closure S) \<inter> (closure(UNIV - S))"
+  by (auto simp add: frontier_def interior_closure)
+
+lemma frontier_straddle:
+  fixes a :: "'a::metric_space"
+  shows "a \<in> frontier S \<longleftrightarrow> (\<forall>e>0. (\<exists>x\<in>S. dist a x < e) \<and> (\<exists>x. x \<notin> S \<and> dist a x < e))" (is "?lhs \<longleftrightarrow> ?rhs")
+proof
+  assume "?lhs"
+  { fix e::real
+    assume "e > 0"
+    let ?rhse = "(\<exists>x\<in>S. dist a x < e) \<and> (\<exists>x. x \<notin> S \<and> dist a x < e)"
+    { assume "a\<in>S"
+      have "\<exists>x\<in>S. dist a x < e" using `e>0` `a\<in>S` by(rule_tac x=a in bexI) auto
+      moreover have "\<exists>x. x \<notin> S \<and> dist a x < e" using `?lhs` `a\<in>S`
+        unfolding frontier_closures closure_def islimpt_def using `e>0`
+        by (auto, erule_tac x="ball a e" in allE, auto)
+      ultimately have ?rhse by auto
+    }
+    moreover
+    { assume "a\<notin>S"
+      hence ?rhse using `?lhs`
+        unfolding frontier_closures closure_def islimpt_def
+        using open_ball[of a e] `e > 0`
+        by (auto, erule_tac x = "ball a e" in allE, auto) (* FIXME: VERY slow! *)
+    }
+    ultimately have ?rhse by auto
+  }
+  thus ?rhs by auto
+next
+  assume ?rhs
+  moreover
+  { fix T assume "a\<notin>S" and
+    as:"\<forall>e>0. (\<exists>x\<in>S. dist a x < e) \<and> (\<exists>x. x \<notin> S \<and> dist a x < e)" "a \<notin> S" "a \<in> T" "open T"
+    from `open T` `a \<in> T` have "\<exists>e>0. ball a e \<subseteq> T" unfolding open_contains_ball[of T] by auto
+    then obtain e where "e>0" "ball a e \<subseteq> T" by auto
+    then obtain y where y:"y\<in>S" "dist a y < e"  using as(1) by auto
+    have "\<exists>y\<in>S. y \<in> T \<and> y \<noteq> a"
+      using `dist a y < e` `ball a e \<subseteq> T` unfolding ball_def using `y\<in>S` `a\<notin>S` by auto
+  }
+  hence "a \<in> closure S" unfolding closure_def islimpt_def using `?rhs` by auto
+  moreover
+  { fix T assume "a \<in> T"  "open T" "a\<in>S"
+    then obtain e where "e>0" and balle: "ball a e \<subseteq> T" unfolding open_contains_ball using `?rhs` by auto
+    obtain x where "x \<notin> S" "dist a x < e" using `?rhs` using `e>0` by auto
+    hence "\<exists>y\<in>UNIV - S. y \<in> T \<and> y \<noteq> a" using balle `a\<in>S` unfolding ball_def by (rule_tac x=x in bexI)auto
+  }
+  hence "a islimpt (UNIV - S) \<or> a\<notin>S" unfolding islimpt_def by auto
+  ultimately show ?lhs unfolding frontier_closures using closure_def[of "UNIV - S"] by auto
+qed
+
+lemma frontier_subset_closed: "closed S \<Longrightarrow> frontier S \<subseteq> S"
+  by (metis frontier_def closure_closed Diff_subset)
+
+lemma frontier_empty: "frontier {} = {}"
+  by (simp add: frontier_def closure_empty)
+
+lemma frontier_subset_eq: "frontier S \<subseteq> S \<longleftrightarrow> closed S"
+proof-
+  { assume "frontier S \<subseteq> S"
+    hence "closure S \<subseteq> S" using interior_subset unfolding frontier_def by auto
+    hence "closed S" using closure_subset_eq by auto
+  }
+  thus ?thesis using frontier_subset_closed[of S] by auto
+qed
+
+lemma frontier_complement: "frontier(UNIV - S) = frontier S"
+  by (auto simp add: frontier_def closure_complement interior_complement)
+
+lemma frontier_disjoint_eq: "frontier S \<inter> S = {} \<longleftrightarrow> open S"
+  using frontier_complement frontier_subset_eq[of "UNIV - S"]
+  unfolding open_closed Compl_eq_Diff_UNIV by auto
+
+subsection{* Common nets and The "within" modifier for nets. *}
+
+definition
+  at_infinity :: "'a::real_normed_vector net" where
+  "at_infinity = Abs_net (range (\<lambda>r. {x. r \<le> norm x}))"
+
+definition
+  indirection :: "'a::real_normed_vector \<Rightarrow> 'a \<Rightarrow> 'a net" (infixr "indirection" 70) where
+  "a indirection v = (at a) within {b. \<exists>c\<ge>0. b - a = scaleR c v}"
+
+text{* Prove That They are all nets. *}
+
+lemma Rep_net_at_infinity:
+  "Rep_net at_infinity = range (\<lambda>r. {x. r \<le> norm x})"
+unfolding at_infinity_def
+apply (rule Abs_net_inverse')
+apply (rule image_nonempty, simp)
+apply (clarsimp, rename_tac r s)
+apply (rule_tac x="max r s" in exI, auto)
+done
+
+lemma within_UNIV: "net within UNIV = net"
+  by (simp add: Rep_net_inject [symmetric] Rep_net_within)
+
+subsection{* Identify Trivial limits, where we can't approach arbitrarily closely. *}
+
+definition
+  trivial_limit :: "'a net \<Rightarrow> bool" where
+  "trivial_limit net \<longleftrightarrow> {} \<in> Rep_net net"
+
+lemma trivial_limit_within:
+  shows "trivial_limit (at a within S) \<longleftrightarrow> \<not> a islimpt S"
+proof
+  assume "trivial_limit (at a within S)"
+  thus "\<not> a islimpt S"
+    unfolding trivial_limit_def
+    unfolding Rep_net_within Rep_net_at
+    unfolding islimpt_def
+    apply (clarsimp simp add: expand_set_eq)
+    apply (rename_tac T, rule_tac x=T in exI)
+    apply (clarsimp, drule_tac x=y in spec, simp)
+    done
+next
+  assume "\<not> a islimpt S"
+  thus "trivial_limit (at a within S)"
+    unfolding trivial_limit_def
+    unfolding Rep_net_within Rep_net_at
+    unfolding islimpt_def
+    apply (clarsimp simp add: image_image)
+    apply (rule_tac x=T in image_eqI)
+    apply (auto simp add: expand_set_eq)
+    done
+qed
+
+lemma trivial_limit_at_iff: "trivial_limit (at a) \<longleftrightarrow> \<not> a islimpt UNIV"
+  using trivial_limit_within [of a UNIV]
+  by (simp add: within_UNIV)
+
+lemma trivial_limit_at:
+  fixes a :: "'a::perfect_space"
+  shows "\<not> trivial_limit (at a)"
+  by (simp add: trivial_limit_at_iff)
+
+lemma trivial_limit_at_infinity:
+  "\<not> trivial_limit (at_infinity :: ('a::{real_normed_vector,zero_neq_one}) net)"
+  (* FIXME: find a more appropriate type class *)
+  unfolding trivial_limit_def Rep_net_at_infinity
+  apply (clarsimp simp add: expand_set_eq)
+  apply (drule_tac x="scaleR r (sgn 1)" in spec)
+  apply (simp add: norm_sgn)
+  done
+
+lemma trivial_limit_sequentially: "\<not> trivial_limit sequentially"
+  by (auto simp add: trivial_limit_def Rep_net_sequentially)
+
+subsection{* Some property holds "sufficiently close" to the limit point. *}
+
+lemma eventually_at: (* FIXME: this replaces Limits.eventually_at *)
+  "eventually P (at a) \<longleftrightarrow> (\<exists>d>0. \<forall>x. 0 < dist x a \<and> dist x a < d \<longrightarrow> P x)"
+unfolding eventually_at dist_nz by auto
+
+lemma eventually_at_infinity:
+  "eventually P at_infinity \<longleftrightarrow> (\<exists>b. \<forall>x. norm x >= b \<longrightarrow> P x)"
+unfolding eventually_def Rep_net_at_infinity by auto
+
+lemma eventually_within: "eventually P (at a within S) \<longleftrightarrow>
+        (\<exists>d>0. \<forall>x\<in>S. 0 < dist x a \<and> dist x a < d \<longrightarrow> P x)"
+unfolding eventually_within eventually_at dist_nz by auto
+
+lemma eventually_within_le: "eventually P (at a within S) \<longleftrightarrow>
+        (\<exists>d>0. \<forall>x\<in>S. 0 < dist x a \<and> dist x a <= d \<longrightarrow> P x)" (is "?lhs = ?rhs")
+unfolding eventually_within
+apply safe
+apply (rule_tac x="d/2" in exI, simp)
+apply (rule_tac x="d" in exI, simp)
+done
+
+lemma eventually_happens: "eventually P net ==> trivial_limit net \<or> (\<exists>x. P x)"
+  unfolding eventually_def trivial_limit_def
+  using Rep_net_nonempty [of net] by auto
+
+lemma always_eventually: "(\<forall>x. P x) ==> eventually P net"
+  unfolding eventually_def trivial_limit_def
+  using Rep_net_nonempty [of net] by auto
+
+lemma trivial_limit_eventually: "trivial_limit net \<Longrightarrow> eventually P net"
+  unfolding trivial_limit_def eventually_def by auto
+
+lemma eventually_False: "eventually (\<lambda>x. False) net \<longleftrightarrow> trivial_limit net"
+  unfolding trivial_limit_def eventually_def by auto
+
+lemma trivial_limit_eq: "trivial_limit net \<longleftrightarrow> (\<forall>P. eventually P net)"
+  apply (safe elim!: trivial_limit_eventually)
+  apply (simp add: eventually_False [symmetric])
+  done
+
+text{* Combining theorems for "eventually" *}
+
+lemma eventually_conjI:
+  "\<lbrakk>eventually (\<lambda>x. P x) net; eventually (\<lambda>x. Q x) net\<rbrakk>
+    \<Longrightarrow> eventually (\<lambda>x. P x \<and> Q x) net"
+by (rule eventually_conj)
+
+lemma eventually_rev_mono:
+  "eventually P net \<Longrightarrow> (\<forall>x. P x \<longrightarrow> Q x) \<Longrightarrow> eventually Q net"
+using eventually_mono [of P Q] by fast
+
+lemma eventually_and: " eventually (\<lambda>x. P x \<and> Q x) net \<longleftrightarrow> eventually P net \<and> eventually Q net"
+  by (auto intro!: eventually_conjI elim: eventually_rev_mono)
+
+lemma eventually_false: "eventually (\<lambda>x. False) net \<longleftrightarrow> trivial_limit net"
+  by (auto simp add: eventually_False)
+
+lemma not_eventually: "(\<forall>x. \<not> P x ) \<Longrightarrow> ~(trivial_limit net) ==> ~(eventually (\<lambda>x. P x) net)"
+  by (simp add: eventually_False)
+
+subsection{* Limits, defined as vacuously true when the limit is trivial. *}
+
+  text{* Notation Lim to avoid collition with lim defined in analysis *}
+definition
+  Lim :: "'a net \<Rightarrow> ('a \<Rightarrow> 'b::t2_space) \<Rightarrow> 'b" where
+  "Lim net f = (THE l. (f ---> l) net)"
+
+lemma Lim:
+ "(f ---> l) net \<longleftrightarrow>
+        trivial_limit net \<or>
+        (\<forall>e>0. eventually (\<lambda>x. dist (f x) l < e) net)"
+  unfolding tendsto_iff trivial_limit_eq by auto
+
+
+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)"
+  by (auto simp add: tendsto_iff eventually_within_le)
+
+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)"
+  by (auto simp add: tendsto_iff eventually_within)
+
+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)"
+  by (auto simp add: tendsto_iff eventually_at)
+
+lemma Lim_at_iff_LIM: "(f ---> l) (at a) \<longleftrightarrow> f -- a --> l"
+  unfolding Lim_at LIM_def by (simp only: zero_less_dist_iff)
+
+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)"
+  by (auto simp add: tendsto_iff eventually_at_infinity)
+
+lemma Lim_sequentially:
+ "(S ---> l) sequentially \<longleftrightarrow>
+          (\<forall>e>0. \<exists>N. \<forall>n\<ge>N. dist (S n) l < e)"
+  by (auto simp add: tendsto_iff eventually_sequentially)
+
+lemma Lim_sequentially_iff_LIMSEQ: "(S ---> l) sequentially \<longleftrightarrow> S ----> l"
+  unfolding Lim_sequentially LIMSEQ_def ..
+
+lemma Lim_eventually: "eventually (\<lambda>x. f x = l) net \<Longrightarrow> (f ---> l) net"
+  by (rule topological_tendstoI, auto elim: eventually_rev_mono)
+
+text{* The expected monotonicity property. *}
+
+lemma Lim_within_empty: "(f ---> l) (net within {})"
+  unfolding tendsto_def Limits.eventually_within by simp
+
+lemma Lim_within_subset: "(f ---> l) (net within S) \<Longrightarrow> T \<subseteq> S \<Longrightarrow> (f ---> l) (net within T)"
+  unfolding tendsto_def Limits.eventually_within
+  by (auto elim!: eventually_elim1)
+
+lemma Lim_Un: assumes "(f ---> l) (net within S)" "(f ---> l) (net within T)"
+  shows "(f ---> l) (net within (S \<union> T))"
+  using assms unfolding tendsto_def Limits.eventually_within
+  apply clarify
+  apply (drule spec, drule (1) mp, drule (1) mp)
+  apply (drule spec, drule (1) mp, drule (1) mp)
+  apply (auto elim: eventually_elim2)
+  done
+
+lemma Lim_Un_univ:
+ "(f ---> l) (net within S) \<Longrightarrow> (f ---> l) (net within T) \<Longrightarrow>  S \<union> T = UNIV
+        ==> (f ---> l) net"
+  by (metis Lim_Un within_UNIV)
+
+text{* Interrelations between restricted and unrestricted limits. *}
+
+lemma Lim_at_within: "(f ---> l) net ==> (f ---> l)(net within S)"
+  (* FIXME: rename *)
+  unfolding tendsto_def Limits.eventually_within
+  apply (clarify, drule spec, drule (1) mp, drule (1) mp)
+  by (auto elim!: eventually_elim1)
+
+lemma Lim_within_open:
+  fixes f :: "'a::topological_space \<Rightarrow> 'b::topological_space"
+  assumes"a \<in> S" "open S"
+  shows "(f ---> l)(at a within S) \<longleftrightarrow> (f ---> l)(at a)" (is "?lhs \<longleftrightarrow> ?rhs")
+proof
+  assume ?lhs
+  { fix A assume "open A" "l \<in> A"
+    with `?lhs` have "eventually (\<lambda>x. f x \<in> A) (at a within S)"
+      by (rule topological_tendstoD)
+    hence "eventually (\<lambda>x. x \<in> S \<longrightarrow> f x \<in> A) (at a)"
+      unfolding Limits.eventually_within .
+    then obtain T where "open T" "a \<in> T" "\<forall>x\<in>T. x \<noteq> a \<longrightarrow> x \<in> S \<longrightarrow> f x \<in> A"
+      unfolding eventually_at_topological by fast
+    hence "open (T \<inter> S)" "a \<in> T \<inter> S" "\<forall>x\<in>(T \<inter> S). x \<noteq> a \<longrightarrow> f x \<in> A"
+      using assms by auto
+    hence "\<exists>T. open T \<and> a \<in> T \<and> (\<forall>x\<in>T. x \<noteq> a \<longrightarrow> f x \<in> A)"
+      by fast
+    hence "eventually (\<lambda>x. f x \<in> A) (at a)"
+      unfolding eventually_at_topological .
+  }
+  thus ?rhs by (rule topological_tendstoI)
+next
+  assume ?rhs
+  thus ?lhs by (rule Lim_at_within)
+qed
+
+text{* Another limit point characterization. *}
+
+lemma islimpt_sequential:
+  fixes x :: "'a::metric_space" (* FIXME: generalize to topological_space *)
+  shows "x islimpt S \<longleftrightarrow> (\<exists>f. (\<forall>n::nat. f n \<in> S -{x}) \<and> (f ---> x) sequentially)"
+    (is "?lhs = ?rhs")
+proof
+  assume ?lhs
+  then obtain f where f:"\<forall>y. y>0 \<longrightarrow> f y \<in> S \<and> f y \<noteq> x \<and> dist (f y) x < y"
+    unfolding islimpt_approachable using choice[of "\<lambda>e y. e>0 \<longrightarrow> y\<in>S \<and> y\<noteq>x \<and> dist y x < e"] by auto
+  { fix n::nat
+    have "f (inverse (real n + 1)) \<in> S - {x}" using f by auto
+  }
+  moreover
+  { fix e::real assume "e>0"
+    hence "\<exists>N::nat. inverse (real (N + 1)) < e" using real_arch_inv[of e] apply (auto simp add: Suc_pred') apply(rule_tac x="n - 1" in exI) by auto
+    then obtain N::nat where "inverse (real (N + 1)) < e" by auto
+    hence "\<forall>n\<ge>N. inverse (real n + 1) < e" by (auto, metis Suc_le_mono le_SucE less_imp_inverse_less nat_le_real_less order_less_trans real_of_nat_Suc real_of_nat_Suc_gt_zero)
+    moreover have "\<forall>n\<ge>N. dist (f (inverse (real n + 1))) x < (inverse (real n + 1))" using f `e>0` by auto
+    ultimately have "\<exists>N::nat. \<forall>n\<ge>N. dist (f (inverse (real n + 1))) x < e" apply(rule_tac x=N in exI) apply auto apply(erule_tac x=n in allE)+ by auto
+  }
+  hence " ((\<lambda>n. f (inverse (real n + 1))) ---> x) sequentially"
+    unfolding Lim_sequentially using f by auto
+  ultimately show ?rhs apply (rule_tac x="(\<lambda>n::nat. f (inverse (real n + 1)))" in exI) by auto
+next
+  assume ?rhs
+  then obtain f::"nat\<Rightarrow>'a"  where f:"(\<forall>n. f n \<in> S - {x})" "(\<forall>e>0. \<exists>N. \<forall>n\<ge>N. dist (f n) x < e)" unfolding Lim_sequentially by auto
+  { fix e::real assume "e>0"
+    then obtain N where "dist (f N) x < e" using f(2) by auto
+    moreover have "f N\<in>S" "f N \<noteq> x" using f(1) by auto
+    ultimately have "\<exists>x'\<in>S. x' \<noteq> x \<and> dist x' x < e" by auto
+  }
+  thus ?lhs unfolding islimpt_approachable by auto
+qed
+
+text{* Basic arithmetical combining theorems for limits. *}
+
+lemma Lim_linear:
+  assumes "(f ---> l) net" "bounded_linear h"
+  shows "((\<lambda>x. h (f x)) ---> h l) net"
+using `bounded_linear h` `(f ---> l) net`
+by (rule bounded_linear.tendsto)
+
+lemma Lim_ident_at: "((\<lambda>x. x) ---> a) (at a)"
+  unfolding tendsto_def Limits.eventually_at_topological by fast
+
+lemma Lim_const: "((\<lambda>x. a) ---> a) net"
+  by (rule tendsto_const)
+
+lemma Lim_cmul:
+  fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"
+  shows "(f ---> l) net ==> ((\<lambda>x. c *\<^sub>R f x) ---> c *\<^sub>R l) net"
+  by (intro tendsto_intros)
+
+lemma Lim_neg:
+  fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"
+  shows "(f ---> l) net ==> ((\<lambda>x. -(f x)) ---> -l) net"
+  by (rule tendsto_minus)
+
+lemma Lim_add: fixes f :: "'a \<Rightarrow> 'b::real_normed_vector" shows
+ "(f ---> l) net \<Longrightarrow> (g ---> m) net \<Longrightarrow> ((\<lambda>x. f(x) + g(x)) ---> l + m) net"
+  by (rule tendsto_add)
+
+lemma Lim_sub:
+  fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"
+  shows "(f ---> l) net \<Longrightarrow> (g ---> m) net \<Longrightarrow> ((\<lambda>x. f(x) - g(x)) ---> l - m) net"
+  by (rule tendsto_diff)
+
+lemma Lim_null:
+  fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"
+  shows "(f ---> l) net \<longleftrightarrow> ((\<lambda>x. f(x) - l) ---> 0) net" by (simp add: Lim dist_norm)
+
+lemma Lim_null_norm:
+  fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"
+  shows "(f ---> 0) net \<longleftrightarrow> ((\<lambda>x. norm(f x)) ---> 0) net"
+  by (simp add: Lim dist_norm)
+
+lemma Lim_null_comparison:
+  fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"
+  assumes "eventually (\<lambda>x. norm (f x) \<le> g x) net" "(g ---> 0) net"
+  shows "(f ---> 0) net"
+proof(simp add: tendsto_iff, rule+)
+  fix e::real assume "0<e"
+  { fix x
+    assume "norm (f x) \<le> g x" "dist (g x) 0 < e"
+    hence "dist (f x) 0 < e" by (simp add: dist_norm)
+  }
+  thus "eventually (\<lambda>x. dist (f x) 0 < e) net"
+    using eventually_and[of "\<lambda>x. norm(f x) <= g x" "\<lambda>x. dist (g x) 0 < e" net]
+    using eventually_mono[of "(\<lambda>x. norm (f x) \<le> g x \<and> dist (g x) 0 < e)" "(\<lambda>x. dist (f x) 0 < e)" net]
+    using assms `e>0` unfolding tendsto_iff by auto
+qed
+
+lemma Lim_component:
+  fixes f :: "'a \<Rightarrow> 'b::metric_space ^ 'n::finite"
+  shows "(f ---> l) net \<Longrightarrow> ((\<lambda>a. f a $i) ---> l$i) net"
+  unfolding tendsto_iff
+  apply (clarify)
+  apply (drule spec, drule (1) mp)
+  apply (erule eventually_elim1)
+  apply (erule le_less_trans [OF dist_nth_le])
+  done
+
+lemma Lim_transform_bound:
+  fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"
+  fixes g :: "'a \<Rightarrow> 'c::real_normed_vector"
+  assumes "eventually (\<lambda>n. norm(f n) <= norm(g n)) net"  "(g ---> 0) net"
+  shows "(f ---> 0) net"
+proof (rule tendstoI)
+  fix e::real assume "e>0"
+  { fix x
+    assume "norm (f x) \<le> norm (g x)" "dist (g x) 0 < e"
+    hence "dist (f x) 0 < e" by (simp add: dist_norm)}
+  thus "eventually (\<lambda>x. dist (f x) 0 < e) net"
+    using eventually_and[of "\<lambda>x. norm (f x) \<le> norm (g x)" "\<lambda>x. dist (g x) 0 < e" net]
+    using eventually_mono[of "\<lambda>x. norm (f x) \<le> norm (g x) \<and> dist (g x) 0 < e" "\<lambda>x. dist (f x) 0 < e" net]
+    using assms `e>0` unfolding tendsto_iff by blast
+qed
+
+text{* Deducing things about the limit from the elements. *}
+
+lemma Lim_in_closed_set:
+  assumes "closed S" "eventually (\<lambda>x. f(x) \<in> S) net" "\<not>(trivial_limit net)" "(f ---> l) net"
+  shows "l \<in> S"
+proof (rule ccontr)
+  assume "l \<notin> S"
+  with `closed S` have "open (- S)" "l \<in> - S"
+    by (simp_all add: open_Compl)
+  with assms(4) have "eventually (\<lambda>x. f x \<in> - S) net"
+    by (rule topological_tendstoD)
+  with assms(2) have "eventually (\<lambda>x. False) net"
+    by (rule eventually_elim2) simp
+  with assms(3) show "False"
+    by (simp add: eventually_False)
+qed
+
+text{* Need to prove closed(cball(x,e)) before deducing this as a corollary. *}
+
+lemma Lim_dist_ubound:
+  assumes "\<not>(trivial_limit net)" "(f ---> l) net" "eventually (\<lambda>x. dist a (f x) <= e) net"
+  shows "dist a l <= e"
+proof (rule ccontr)
+  assume "\<not> dist a l \<le> e"
+  then have "0 < dist a l - e" by simp
+  with assms(2) have "eventually (\<lambda>x. dist (f x) l < dist a l - e) net"
+    by (rule tendstoD)
+  with assms(3) have "eventually (\<lambda>x. dist a (f x) \<le> e \<and> dist (f x) l < dist a l - e) net"
+    by (rule eventually_conjI)
+  then obtain w where "dist a (f w) \<le> 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)"
+    by (rule add_le_less_mono)
+  hence "dist a (f w) + dist (f w) l < dist a l"
+    by simp
+  also have "\<dots> \<le> dist a (f w) + dist (f w) l"
+    by (rule dist_triangle)
+  finally show False by simp
+qed
+
+lemma Lim_norm_ubound:
+  fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"
+  assumes "\<not>(trivial_limit net)" "(f ---> l) net" "eventually (\<lambda>x. norm(f x) <= e) net"
+  shows "norm(l) <= e"
+proof (rule ccontr)
+  assume "\<not> norm l \<le> e"
+  then have "0 < norm l - e" by simp
+  with assms(2) have "eventually (\<lambda>x. dist (f x) l < norm l - e) net"
+    by (rule tendstoD)
+  with assms(3) have "eventually (\<lambda>x. norm (f x) \<le> e \<and> dist (f x) l < norm l - e) net"
+    by (rule eventually_conjI)
+  then obtain w where "norm (f w) \<le> 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) \<le> e" by (simp_all add: dist_norm)
+  hence "norm (f w - l) + norm (f w) < norm l" by simp
+  hence "norm (f w - l - f w) < norm l" by (rule le_less_trans [OF norm_triangle_ineq4])
+  thus False using `\<not> norm l \<le> e` by simp
+qed
+
+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"
+  shows "e \<le> norm l"
+proof (rule ccontr)
+  assume "\<not> e \<le> norm l"
+  then have "0 < e - norm l" by simp
+  with assms(2) have "eventually (\<lambda>x. dist (f x) l < e - norm l) net"
+    by (rule tendstoD)
+  with assms(3) have "eventually (\<lambda>x. e \<le> norm (f x) \<and> dist (f x) l < e - norm l) net"
+    by (rule eventually_conjI)
+  then obtain w where "e \<le> 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 \<le> norm (f w)" by (simp_all add: dist_norm)
+  hence "norm (f w - l) + norm l < norm (f w)" by (rule less_le_trans)
+  hence "norm (f w - l + l) < norm (f w)" by (rule le_less_trans [OF norm_triangle_ineq])
+  thus False by simp
+qed
+
+text{* Uniqueness of the limit, when nontrivial. *}
+
+lemma Lim_unique:
+  fixes f :: "'a \<Rightarrow> 'b::t2_space"
+  assumes "\<not> trivial_limit net"  "(f ---> l) net"  "(f ---> l') net"
+  shows "l = l'"
+proof (rule ccontr)
+  assume "l \<noteq> l'"
+  obtain U V where "open U" "open V" "l \<in> U" "l' \<in> V" "U \<inter> V = {}"
+    using hausdorff [OF `l \<noteq> l'`] by fast
+  have "eventually (\<lambda>x. f x \<in> U) net"
+    using `(f ---> l) net` `open U` `l \<in> U` by (rule topological_tendstoD)
+  moreover
+  have "eventually (\<lambda>x. f x \<in> V) net"
+    using `(f ---> l') net` `open V` `l' \<in> V` by (rule topological_tendstoD)
+  ultimately
+  have "eventually (\<lambda>x. False) net"
+  proof (rule eventually_elim2)
+    fix x
+    assume "f x \<in> U" "f x \<in> V"
+    hence "f x \<in> U \<inter> V" by simp
+    with `U \<inter> V = {}` show "False" by simp
+  qed
+  with `\<not> trivial_limit net` show "False"
+    by (simp add: eventually_False)
+qed
+
+lemma tendsto_Lim:
+  fixes f :: "'a \<Rightarrow> 'b::t2_space"
+  shows "~(trivial_limit net) \<Longrightarrow> (f ---> l) net ==> Lim net f = l"
+  unfolding Lim_def using Lim_unique[of net f] by auto
+
+text{* Limit under bilinear function *}
+
+lemma Lim_bilinear:
+  assumes "(f ---> l) net" and "(g ---> m) net" and "bounded_bilinear h"
+  shows "((\<lambda>x. h (f x) (g x)) ---> (h l m)) net"
+using `bounded_bilinear h` `(f ---> l) net` `(g ---> m) net`
+by (rule bounded_bilinear.tendsto)
+
+text{* These are special for limits out of the same vector space. *}
+
+lemma Lim_within_id: "(id ---> a) (at a within s)"
+  unfolding tendsto_def Limits.eventually_within eventually_at_topological
+  by auto
+
+lemma Lim_at_id: "(id ---> a) (at a)"
+apply (subst within_UNIV[symmetric]) by (simp add: Lim_within_id)
+
+lemma Lim_at_zero:
+  fixes a :: "'a::real_normed_vector"
+  fixes l :: "'b::topological_space"
+  shows "(f ---> l) (at a) \<longleftrightarrow> ((\<lambda>x. f(a + x)) ---> l) (at 0)" (is "?lhs = ?rhs")
+proof
+  assume "?lhs"
+  { fix S assume "open S" "l \<in> S"
+    with `?lhs` have "eventually (\<lambda>x. f x \<in> S) (at a)"
+      by (rule topological_tendstoD)
+    then obtain d where d: "d>0" "\<forall>x. x \<noteq> a \<and> dist x a < d \<longrightarrow> f x \<in> S"
+      unfolding Limits.eventually_at by fast
+    { fix x::"'a" assume "x \<noteq> 0 \<and> dist x 0 < d"
+      hence "f (a + x) \<in> S" using d
+      apply(erule_tac x="x+a" in allE)
+      by(auto simp add: comm_monoid_add.mult_commute dist_norm dist_commute)
+    }
+    hence "\<exists>d>0. \<forall>x. x \<noteq> 0 \<and> dist x 0 < d \<longrightarrow> f (a + x) \<in> S"
+      using d(1) by auto
+    hence "eventually (\<lambda>x. f (a + x) \<in> S) (at 0)"
+      unfolding Limits.eventually_at .
+  }
+  thus "?rhs" by (rule topological_tendstoI)
+next
+  assume "?rhs"
+  { fix S assume "open S" "l \<in> S"
+    with `?rhs` have "eventually (\<lambda>x. f (a + x) \<in> S) (at 0)"
+      by (rule topological_tendstoD)
+    then obtain d where d: "d>0" "\<forall>x. x \<noteq> 0 \<and> dist x 0 < d \<longrightarrow> f (a + x) \<in> S"
+      unfolding Limits.eventually_at by fast
+    { fix x::"'a" assume "x \<noteq> a \<and> dist x a < d"
+      hence "f x \<in> S" using d apply(erule_tac x="x-a" in allE)
+        by(auto simp add: comm_monoid_add.mult_commute dist_norm dist_commute)
+    }
+    hence "\<exists>d>0. \<forall>x. x \<noteq> a \<and> dist x a < d \<longrightarrow> f x \<in> S" using d(1) by auto
+    hence "eventually (\<lambda>x. f x \<in> S) (at a)" unfolding Limits.eventually_at .
+  }
+  thus "?lhs" by (rule topological_tendstoI)
+qed
+
+text{* It's also sometimes useful to extract the limit point from the net.  *}
+
+definition
+  netlimit :: "'a::t2_space net \<Rightarrow> 'a" where
+  "netlimit net = (SOME a. ((\<lambda>x. x) ---> a) net)"
+
+lemma netlimit_within:
+  assumes "\<not> trivial_limit (at a within S)"
+  shows "netlimit (at a within S) = a"
+unfolding netlimit_def
+apply (rule some_equality)
+apply (rule Lim_at_within)
+apply (rule Lim_ident_at)
+apply (erule Lim_unique [OF assms])
+apply (rule Lim_at_within)
+apply (rule Lim_ident_at)
+done
+
+lemma netlimit_at:
+  fixes a :: "'a::perfect_space"
+  shows "netlimit (at a) = a"
+  apply (subst within_UNIV[symmetric])
+  using netlimit_within[of a UNIV]
+  by (simp add: trivial_limit_at within_UNIV)
+
+text{* Transformation of limit. *}
+
+lemma Lim_transform:
+  fixes f g :: "'a::type \<Rightarrow> 'b::real_normed_vector"
+  assumes "((\<lambda>x. f x - g x) ---> 0) net" "(f ---> l) net"
+  shows "(g ---> l) net"
+proof-
+  from assms have "((\<lambda>x. f x - g x - f x) ---> 0 - l) net" using Lim_sub[of "\<lambda>x. f x - g x" 0 net f l] by auto
+  thus "?thesis" using Lim_neg [of "\<lambda> x. - g x" "-l" net] by auto
+qed
+
+lemma Lim_transform_eventually:
+  "eventually (\<lambda>x. f x = g x) net \<Longrightarrow> (f ---> l) net ==> (g ---> l) net"
+  apply (rule topological_tendstoI)
+  apply (drule (2) topological_tendstoD)
+  apply (erule (1) eventually_elim2, simp)
+  done
+
+lemma Lim_transform_within:
+  fixes l :: "'b::metric_space" (* TODO: generalize *)
+  assumes "0 < d" "(\<forall>x'\<in>S. 0 < dist x' x \<and> dist x' x < d \<longrightarrow> f x' = g x')"
+          "(f ---> l) (at x within S)"
+  shows   "(g ---> l) (at x within S)"
+  using assms(1,3) unfolding Lim_within
+  apply -
+  apply (clarify, rename_tac e)
+  apply (drule_tac x=e in spec, clarsimp, rename_tac r)
+  apply (rule_tac x="min d r" in exI, clarsimp, rename_tac y)
+  apply (drule_tac x=y in bspec, assumption, clarsimp)
+  apply (simp add: assms(2))
+  done
+
+lemma Lim_transform_at:
+  fixes l :: "'b::metric_space" (* TODO: generalize *)
+  shows "0 < d \<Longrightarrow> (\<forall>x'. 0 < dist x' x \<and> dist x' x < d \<longrightarrow> f x' = g x') \<Longrightarrow>
+  (f ---> l) (at x) ==> (g ---> l) (at x)"
+  apply (subst within_UNIV[symmetric])
+  using Lim_transform_within[of d UNIV x f g l]
+  by (auto simp add: within_UNIV)
+
+text{* Common case assuming being away from some crucial point like 0. *}
+
+lemma Lim_transform_away_within:
+  fixes a b :: "'a::metric_space"
+  fixes l :: "'b::metric_space" (* TODO: generalize *)
+  assumes "a\<noteq>b" "\<forall>x\<in> S. x \<noteq> a \<and> x \<noteq> b \<longrightarrow> f x = g x"
+  and "(f ---> l) (at a within S)"
+  shows "(g ---> l) (at a within S)"
+proof-
+  have "\<forall>x'\<in>S. 0 < dist x' a \<and> dist x' a < dist a b \<longrightarrow> f x' = g x'" using assms(2)
+    apply auto apply(erule_tac x=x' in ballE) by (auto simp add: dist_commute)
+  thus ?thesis using Lim_transform_within[of "dist a b" S a f g l] using assms(1,3) unfolding dist_nz by auto
+qed
+
+lemma Lim_transform_away_at:
+  fixes a b :: "'a::metric_space"
+  fixes l :: "'b::metric_space" (* TODO: generalize *)
+  assumes ab: "a\<noteq>b" and fg: "\<forall>x. x \<noteq> a \<and> x \<noteq> b \<longrightarrow> f x = g x"
+  and fl: "(f ---> l) (at a)"
+  shows "(g ---> l) (at a)"
+  using Lim_transform_away_within[OF ab, of UNIV f g l] fg fl
+  by (auto simp add: within_UNIV)
+
+text{* Alternatively, within an open set. *}
+
+lemma Lim_transform_within_open:
+  fixes a :: "'a::metric_space"
+  fixes l :: "'b::metric_space" (* TODO: generalize *)
+  assumes "open S"  "a \<in> S"  "\<forall>x\<in>S. x \<noteq> a \<longrightarrow> f x = g x"  "(f ---> l) (at a)"
+  shows "(g ---> l) (at a)"
+proof-
+  from assms(1,2) obtain e::real where "e>0" and e:"ball a e \<subseteq> S" unfolding open_contains_ball by auto
+  hence "\<forall>x'. 0 < dist x' a \<and> dist x' a < e \<longrightarrow> f x' = g x'" using assms(3)
+    unfolding ball_def subset_eq apply auto apply(erule_tac x=x' in allE) apply(erule_tac x=x' in ballE) by(auto simp add: dist_commute)
+  thus ?thesis using Lim_transform_at[of e a f g l] `e>0` assms(4) by auto
+qed
+
+text{* A congruence rule allowing us to transform limits assuming not at point. *}
+
+(* FIXME: Only one congruence rule for tendsto can be used at a time! *)
+
+lemma Lim_cong_within[cong add]:
+  fixes a :: "'a::metric_space"
+  fixes l :: "'b::metric_space" (* TODO: generalize *)
+  shows "(\<And>x. x \<noteq> a \<Longrightarrow> f x = g x) ==> ((\<lambda>x. f x) ---> l) (at a within S) \<longleftrightarrow> ((g ---> l) (at a within S))"
+  by (simp add: Lim_within dist_nz[symmetric])
+
+lemma Lim_cong_at[cong add]:
+  fixes a :: "'a::metric_space"
+  fixes l :: "'b::metric_space" (* TODO: generalize *)
+  shows "(\<And>x. x \<noteq> a ==> f x = g x) ==> (((\<lambda>x. f x) ---> l) (at a) \<longleftrightarrow> ((g ---> l) (at a)))"
+  by (simp add: Lim_at dist_nz[symmetric])
+
+text{* Useful lemmas on closure and set of possible sequential limits.*}
+
+lemma closure_sequential:
+  fixes l :: "'a::metric_space" (* TODO: generalize *)
+  shows "l \<in> closure S \<longleftrightarrow> (\<exists>x. (\<forall>n. x n \<in> S) \<and> (x ---> l) sequentially)" (is "?lhs = ?rhs")
+proof
+  assume "?lhs" moreover
+  { assume "l \<in> S"
+    hence "?rhs" using Lim_const[of l sequentially] by auto
+  } moreover
+  { assume "l islimpt S"
+    hence "?rhs" unfolding islimpt_sequential by auto
+  } ultimately
+  show "?rhs" unfolding closure_def by auto
+next
+  assume "?rhs"
+  thus "?lhs" unfolding closure_def unfolding islimpt_sequential by auto
+qed
+
+lemma closed_sequential_limits:
+  fixes S :: "'a::metric_space set"
+  shows "closed S \<longleftrightarrow> (\<forall>x l. (\<forall>n. x n \<in> S) \<and> (x ---> l) sequentially \<longrightarrow> l \<in> S)"
+  unfolding closed_limpt
+  using closure_sequential [where 'a='a] closure_closed [where 'a='a] closed_limpt [where 'a='a] islimpt_sequential [where 'a='a] mem_delete [where 'a='a]
+  by metis
+
+lemma closure_approachable:
+  fixes S :: "'a::metric_space set"
+  shows "x \<in> closure S \<longleftrightarrow> (\<forall>e>0. \<exists>y\<in>S. dist y x < e)"
+  apply (auto simp add: closure_def islimpt_approachable)
+  by (metis dist_self)
+
+lemma closed_approachable:
+  fixes S :: "'a::metric_space set"
+  shows "closed S ==> (\<forall>e>0. \<exists>y\<in>S. dist y x < e) \<longleftrightarrow> x \<in> S"
+  by (metis closure_closed closure_approachable)
+
+text{* Some other lemmas about sequences. *}
+
+lemma seq_offset:
+  fixes l :: "'a::metric_space" (* TODO: generalize *)
+  shows "(f ---> l) sequentially ==> ((\<lambda>i. f( i + k)) ---> l) sequentially"
+  apply (auto simp add: Lim_sequentially)
+  by (metis trans_le_add1 )
+
+lemma seq_offset_neg:
+  "(f ---> l) sequentially ==> ((\<lambda>i. f(i - k)) ---> l) sequentially"
+  apply (rule topological_tendstoI)
+  apply (drule (2) topological_tendstoD)
+  apply (simp only: eventually_sequentially)
+  apply (subgoal_tac "\<And>N k (n::nat). N + k <= n ==> N <= n - k")
+  apply metis
+  by arith
+
+lemma seq_offset_rev:
+  "((\<lambda>i. f(i + k)) ---> l) sequentially ==> (f ---> l) sequentially"
+  apply (rule topological_tendstoI)
+  apply (drule (2) topological_tendstoD)
+  apply (simp only: eventually_sequentially)
+  apply (subgoal_tac "\<And>N k (n::nat). N + k <= n ==> N <= n - k \<and> (n - k) + k = n")
+  by metis arith
+
+lemma seq_harmonic: "((\<lambda>n. inverse (real n)) ---> 0) sequentially"
+proof-
+  { fix e::real assume "e>0"
+    hence "\<exists>N::nat. \<forall>n::nat\<ge>N. inverse (real n) < e"
+      using real_arch_inv[of e] apply auto apply(rule_tac x=n in exI)
+      by (metis not_le le_imp_inverse_le not_less real_of_nat_gt_zero_cancel_iff real_of_nat_less_iff xt1(7))
+  }
+  thus ?thesis unfolding Lim_sequentially dist_norm by simp
+qed
+
+text{* More properties of closed balls. *}
+
+lemma closed_cball: "closed (cball x e)"
+unfolding cball_def closed_def
+unfolding Collect_neg_eq [symmetric] not_le
+apply (clarsimp simp add: open_dist, rename_tac y)
+apply (rule_tac x="dist x y - e" in exI, clarsimp)
+apply (rename_tac x')
+apply (cut_tac x=x and y=x' and z=y in dist_triangle)
+apply simp
+done
+
+lemma open_contains_cball: "open S \<longleftrightarrow> (\<forall>x\<in>S. \<exists>e>0.  cball x e \<subseteq> S)"
+proof-
+  { fix x and e::real assume "x\<in>S" "e>0" "ball x e \<subseteq> S"
+    hence "\<exists>d>0. cball x d \<subseteq> S" unfolding subset_eq by (rule_tac x="e/2" in exI, auto)
+  } moreover
+  { fix x and e::real assume "x\<in>S" "e>0" "cball x e \<subseteq> S"
+    hence "\<exists>d>0. ball x d \<subseteq> S" unfolding subset_eq apply(rule_tac x="e/2" in exI) by auto
+  } ultimately
+  show ?thesis unfolding open_contains_ball by auto
+qed
+
+lemma open_contains_cball_eq: "open S ==> (\<forall>x. x \<in> S \<longleftrightarrow> (\<exists>e>0. cball x e \<subseteq> S))"
+  by (metis open_contains_cball subset_eq order_less_imp_le centre_in_cball mem_def)
+
+lemma mem_interior_cball: "x \<in> interior S \<longleftrightarrow> (\<exists>e>0. cball x e \<subseteq> S)"
+  apply (simp add: interior_def, safe)
+  apply (force simp add: open_contains_cball)
+  apply (rule_tac x="ball x e" in exI)
+  apply (simp add: open_ball centre_in_ball subset_trans [OF ball_subset_cball])
+  done
+
+lemma islimpt_ball:
+  fixes x y :: "'a::{real_normed_vector,perfect_space}"
+  shows "y islimpt ball x e \<longleftrightarrow> 0 < e \<and> y \<in> cball x e" (is "?lhs = ?rhs")
+proof
+  assume "?lhs"
+  { assume "e \<le> 0"
+    hence *:"ball x e = {}" using ball_eq_empty[of x e] by auto
+    have False using `?lhs` unfolding * using islimpt_EMPTY[of y] by auto
+  }
+  hence "e > 0" by (metis not_less)
+  moreover
+  have "y \<in> cball x e" using closed_cball[of x e] islimpt_subset[of y "ball x e" "cball x e"] ball_subset_cball[of x e] `?lhs` unfolding closed_limpt by auto
+  ultimately show "?rhs" by auto
+next
+  assume "?rhs" hence "e>0"  by auto
+  { fix d::real assume "d>0"
+    have "\<exists>x'\<in>ball x e. x' \<noteq> y \<and> dist x' y < d"
+    proof(cases "d \<le> dist x y")
+      case True thus "\<exists>x'\<in>ball x e. x' \<noteq> y \<and> dist x' y < d"
+      proof(cases "x=y")
+        case True hence False using `d \<le> dist x y` `d>0` by auto
+        thus "\<exists>x'\<in>ball x e. x' \<noteq> y \<and> dist x' y < d" by auto
+      next
+        case False
+
+        have "dist x (y - (d / (2 * dist y x)) *\<^sub>R (y - x))
+              = norm (x - y + (d / (2 * norm (y - x))) *\<^sub>R (y - x))"
+          unfolding mem_cball mem_ball dist_norm diff_diff_eq2 diff_add_eq[THEN sym] by auto
+        also have "\<dots> = \<bar>- 1 + d / (2 * norm (x - y))\<bar> * norm (x - y)"
+          using scaleR_left_distrib[of "- 1" "d / (2 * norm (y - x))", THEN sym, of "y - x"]
+          unfolding scaleR_minus_left scaleR_one
+          by (auto simp add: norm_minus_commute)
+        also have "\<dots> = \<bar>- norm (x - y) + d / 2\<bar>"
+          unfolding abs_mult_pos[of "norm (x - y)", OF norm_ge_zero[of "x - y"]]
+          unfolding real_add_mult_distrib using `x\<noteq>y`[unfolded dist_nz, unfolded dist_norm] by auto
+        also have "\<dots> \<le> e - d/2" using `d \<le> dist x y` and `d>0` and `?rhs` by(auto simp add: dist_norm)
+        finally have "y - (d / (2 * dist y x)) *\<^sub>R (y - x) \<in> ball x e" using `d>0` by auto
+
+        moreover
+
+        have "(d / (2*dist y x)) *\<^sub>R (y - x) \<noteq> 0"
+          using `x\<noteq>y`[unfolded dist_nz] `d>0` unfolding scaleR_eq_0_iff by (auto simp add: dist_commute)
+        moreover
+        have "dist (y - (d / (2 * dist y x)) *\<^sub>R (y - x)) y < d" unfolding dist_norm apply simp unfolding norm_minus_cancel
+          using `d>0` `x\<noteq>y`[unfolded dist_nz] dist_commute[of x y]
+          unfolding dist_norm by auto
+        ultimately show "\<exists>x'\<in>ball x e. x' \<noteq> y \<and> dist x' y < d" by (rule_tac  x="y - (d / (2*dist y x)) *\<^sub>R (y - x)" in bexI) auto
+      qed
+    next
+      case False hence "d > dist x y" by auto
+      show "\<exists>x'\<in>ball x e. x' \<noteq> y \<and> dist x' y < d"
+      proof(cases "x=y")
+        case True
+        obtain z where **: "z \<noteq> y" "dist z y < min e d"
+          using perfect_choose_dist[of "min e d" y]
+          using `d > 0` `e>0` by auto
+        show "\<exists>x'\<in>ball x e. x' \<noteq> y \<and> dist x' y < d"
+          unfolding `x = y`
+          using `z \<noteq> y` **
+          by (rule_tac x=z in bexI, auto simp add: dist_commute)
+      next
+        case False thus "\<exists>x'\<in>ball x e. x' \<noteq> y \<and> dist x' y < d"
+          using `d>0` `d > dist x y` `?rhs` by(rule_tac x=x in bexI, auto)
+      qed
+    qed  }
+  thus "?lhs" unfolding mem_cball islimpt_approachable mem_ball by auto
+qed
+
+lemma closure_ball_lemma:
+  fixes x y :: "'a::real_normed_vector"
+  assumes "x \<noteq> y" shows "y islimpt ball x (dist x y)"
+proof (rule islimptI)
+  fix T assume "y \<in> T" "open T"
+  then obtain r where "0 < r" "\<forall>z. dist z y < r \<longrightarrow> z \<in> T"
+    unfolding open_dist by fast
+  (* choose point between x and y, within distance r of y. *)
+  def k \<equiv> "min 1 (r / (2 * dist x y))"
+  def z \<equiv> "y + scaleR k (x - y)"
+  have z_def2: "z = x + scaleR (1 - k) (y - x)"
+    unfolding z_def by (simp add: algebra_simps)
+  have "dist z y < r"
+    unfolding z_def k_def using `0 < r`
+    by (simp add: dist_norm min_def)
+  hence "z \<in> T" using `\<forall>z. dist z y < r \<longrightarrow> z \<in> T` by simp
+  have "dist x z < dist x y"
+    unfolding z_def2 dist_norm
+    apply (simp add: norm_minus_commute)
+    apply (simp only: dist_norm [symmetric])
+    apply (subgoal_tac "\<bar>1 - k\<bar> * dist x y < 1 * dist x y", simp)
+    apply (rule mult_strict_right_mono)
+    apply (simp add: k_def divide_pos_pos zero_less_dist_iff `0 < r` `x \<noteq> y`)
+    apply (simp add: zero_less_dist_iff `x \<noteq> y`)
+    done
+  hence "z \<in> ball x (dist x y)" by simp
+  have "z \<noteq> y"
+    unfolding z_def k_def using `x \<noteq> y` `0 < r`
+    by (simp add: min_def)
+  show "\<exists>z\<in>ball x (dist x y). z \<in> T \<and> z \<noteq> y"
+    using `z \<in> ball x (dist x y)` `z \<in> T` `z \<noteq> y`
+    by fast
+qed
+
+lemma closure_ball:
+  fixes x :: "'a::real_normed_vector"
+  shows "0 < e \<Longrightarrow> closure (ball x e) = cball x e"
+apply (rule equalityI)
+apply (rule closure_minimal)
+apply (rule ball_subset_cball)
+apply (rule closed_cball)
+apply (rule subsetI, rename_tac y)
+apply (simp add: le_less [where 'a=real])
+apply (erule disjE)
+apply (rule subsetD [OF closure_subset], simp)
+apply (simp add: closure_def)
+apply clarify
+apply (rule closure_ball_lemma)
+apply (simp add: zero_less_dist_iff)
+done
+
+(* In a trivial vector space, this fails for e = 0. *)
+lemma interior_cball:
+  fixes x :: "'a::{real_normed_vector, perfect_space}"
+  shows "interior (cball x e) = ball x e"
+proof(cases "e\<ge>0")
+  case False note cs = this
+  from cs have "ball x e = {}" using ball_empty[of e x] by auto moreover
+  { fix y assume "y \<in> cball x e"
+    hence False unfolding mem_cball using dist_nz[of x y] cs by auto  }
+  hence "cball x e = {}" by auto
+  hence "interior (cball x e) = {}" using interior_empty by auto
+  ultimately show ?thesis by blast
+next
+  case True note cs = this
+  have "ball x e \<subseteq> cball x e" using ball_subset_cball by auto moreover
+  { fix S y assume as: "S \<subseteq> cball x e" "open S" "y\<in>S"
+    then obtain d where "d>0" and d:"\<forall>x'. dist x' y < d \<longrightarrow> x' \<in> S" unfolding open_dist by blast
+
+    then obtain xa where xa_y: "xa \<noteq> y" and xa: "dist xa y < d"
+      using perfect_choose_dist [of d] by auto
+    have "xa\<in>S" using d[THEN spec[where x=xa]] using xa by(auto simp add: dist_commute)
+    hence xa_cball:"xa \<in> cball x e" using as(1) by auto
+
+    hence "y \<in> ball x e" proof(cases "x = y")
+      case True
+      hence "e>0" using xa_y[unfolded dist_nz] xa_cball[unfolded mem_cball] by (auto simp add: dist_commute)
+      thus "y \<in> ball x e" using `x = y ` by simp
+    next
+      case False
+      have "dist (y + (d / 2 / dist y x) *\<^sub>R (y - x)) y < d" unfolding dist_norm
+        using `d>0` norm_ge_zero[of "y - x"] `x \<noteq> y` by auto
+      hence *:"y + (d / 2 / dist y x) *\<^sub>R (y - x) \<in> cball x e" using d as(1)[unfolded subset_eq] by blast
+      have "y - x \<noteq> 0" using `x \<noteq> y` by auto
+      hence **:"d / (2 * norm (y - x)) > 0" unfolding zero_less_norm_iff[THEN sym]
+        using `d>0` divide_pos_pos[of d "2*norm (y - x)"] by auto
+
+      have "dist (y + (d / 2 / dist y x) *\<^sub>R (y - x)) x = norm (y + (d / (2 * norm (y - x))) *\<^sub>R y - (d / (2 * norm (y - x))) *\<^sub>R x - x)"
+        by (auto simp add: dist_norm algebra_simps)
+      also have "\<dots> = norm ((1 + d / (2 * norm (y - x))) *\<^sub>R (y - x))"
+        by (auto simp add: algebra_simps)
+      also have "\<dots> = \<bar>1 + d / (2 * norm (y - x))\<bar> * norm (y - x)"
+        using ** by auto
+      also have "\<dots> = (dist y x) + d/2"using ** by (auto simp add: left_distrib dist_norm)
+      finally have "e \<ge> dist x y +d/2" using *[unfolded mem_cball] by (auto simp add: dist_commute)
+      thus "y \<in> ball x e" unfolding mem_ball using `d>0` by auto
+    qed  }
+  hence "\<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
+qed
+
+lemma frontier_ball:
+  fixes a :: "'a::real_normed_vector"
+  shows "0 < e ==> frontier(ball a e) = {x. dist a x = e}"
+  apply (simp add: frontier_def closure_ball interior_open open_ball order_less_imp_le)
+  apply (simp add: expand_set_eq)
+  by arith
+
+lemma frontier_cball:
+  fixes a :: "'a::{real_normed_vector, perfect_space}"
+  shows "frontier(cball a e) = {x. dist a x = e}"
+  apply (simp add: frontier_def interior_cball closed_cball closure_closed order_less_imp_le)
+  apply (simp add: expand_set_eq)
+  by arith
+
+lemma cball_eq_empty: "(cball x e = {}) \<longleftrightarrow> e < 0"
+  apply (simp add: expand_set_eq not_le)
+  by (metis zero_le_dist dist_self order_less_le_trans)
+lemma cball_empty: "e < 0 ==> cball x e = {}" by (simp add: cball_eq_empty)
+
+lemma cball_eq_sing:
+  fixes x :: "'a::perfect_space"
+  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"
+    using perfect_choose_dist [OF e] by auto
+  hence "a \<noteq> x" "dist x a \<le> e" by (auto simp add: dist_commute)
+  with e show ?thesis by (auto simp add: expand_set_eq)
+qed auto
+
+lemma cball_sing:
+  fixes x :: "'a::metric_space"
+  shows "e = 0 ==> cball x e = {x}"
+  by (auto simp add: expand_set_eq)
+
+text{* For points in the interior, localization of limits makes no difference.   *}
+
+lemma eventually_within_interior:
+  assumes "x \<in> interior S"
+  shows "eventually P (at x within S) \<longleftrightarrow> eventually P (at x)" (is "?lhs = ?rhs")
+proof-
+  from assms obtain T where T: "open T" "x \<in> T" "T \<subseteq> S"
+    unfolding interior_def by fast
+  { 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"
+      unfolding Limits.eventually_within Limits.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"
+      by auto
+    then have "?rhs"
+      unfolding Limits.eventually_at_topological by auto
+  } moreover
+  { assume "?rhs" hence "?lhs"
+      unfolding Limits.eventually_within
+      by (auto elim: eventually_elim1)
+  } ultimately
+  show "?thesis" ..
+qed
+
+lemma lim_within_interior:
+  "x \<in> interior S \<Longrightarrow> (f ---> l) (at x within S) \<longleftrightarrow> (f ---> l) (at x)"
+  unfolding tendsto_def by (simp add: eventually_within_interior)
+
+lemma netlimit_within_interior:
+  fixes x :: "'a::{perfect_space, real_normed_vector}"
+    (* FIXME: generalize to perfect_space *)
+  assumes "x \<in> interior S"
+  shows "netlimit(at x within S) = x" (is "?lhs = ?rhs")
+proof-
+  from assms obtain e::real where e:"e>0" "ball x e \<subseteq> S" using open_interior[of S] unfolding open_contains_ball using interior_subset[of S] by auto
+  hence "\<not> trivial_limit (at x within S)" using islimpt_subset[of x "ball x e" S] unfolding trivial_limit_within islimpt_ball centre_in_cball by auto
+  thus ?thesis using netlimit_within by auto
+qed
+
+subsection{* Boundedness. *}
+
+  (* FIXME: This has to be unified with BSEQ!! *)
+definition
+  bounded :: "'a::metric_space set \<Rightarrow> bool" where
+  "bounded S \<longleftrightarrow> (\<exists>x e. \<forall>y\<in>S. dist x y \<le> e)"
+
+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 (drule (1) bspec)
+apply (erule order_trans [OF dist_triangle add_left_mono])
+apply auto
+done
+
+lemma bounded_iff: "bounded S \<longleftrightarrow> (\<exists>a. \<forall>x\<in>S. norm x \<le> a)"
+unfolding bounded_any_center [where a=0]
+by (simp add: dist_norm)
+
+lemma bounded_empty[simp]: "bounded {}" by (simp add: bounded_def)
+lemma bounded_subset: "bounded T \<Longrightarrow> S \<subseteq> T ==> bounded S"
+  by (metis bounded_def subset_eq)
+
+lemma bounded_interior[intro]: "bounded S ==> bounded(interior S)"
+  by (metis bounded_subset interior_subset)
+
+lemma bounded_closure[intro]: assumes "bounded S" shows "bounded(closure S)"
+proof-
+  from assms obtain x and a where a: "\<forall>y\<in>S. dist x y \<le> a" unfolding bounded_def by auto
+  { fix y assume "y \<in> closure S"
+    then obtain f where f: "\<forall>n. f n \<in> S"  "(f ---> y) sequentially"
+      unfolding closure_sequential by auto
+    have "\<forall>n. f n \<in> S \<longrightarrow> dist x (f n) \<le> a" using a by simp
+    hence "eventually (\<lambda>n. dist x (f n) \<le> a) sequentially"
+      by (rule eventually_mono, simp add: f(1))
+    have "dist x y \<le> a"
+      apply (rule Lim_dist_ubound [of sequentially f])
+      apply (rule trivial_limit_sequentially)
+      apply (rule f(2))
+      apply fact
+      done
+  }
+  thus ?thesis unfolding bounded_def by auto
+qed
+
+lemma bounded_cball[simp,intro]: "bounded (cball x e)"
+  apply (simp add: bounded_def)
+  apply (rule_tac x=x in exI)
+  apply (rule_tac x=e in exI)
+  apply auto
+  done
+
+lemma bounded_ball[simp,intro]: "bounded(ball x e)"
+  by (metis ball_subset_cball bounded_cball bounded_subset)
+
+lemma finite_imp_bounded[intro]: assumes "finite S" shows "bounded S"
+proof-
+  { fix a F assume as:"bounded F"
+    then obtain x e where "\<forall>y\<in>F. dist x y \<le> e" unfolding bounded_def by auto
+    hence "\<forall>y\<in>(insert a F). dist x y \<le> max e (dist x a)" by auto
+    hence "bounded (insert a F)" unfolding bounded_def by (intro exI)
+  }
+  thus ?thesis using finite_induct[of S bounded]  using bounded_empty assms by auto
+qed
+
+lemma bounded_Un[simp]: "bounded (S \<union> T) \<longleftrightarrow> bounded S \<and> bounded T"
+  apply (auto simp add: bounded_def)
+  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 (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)"
+  by (induct rule: finite_induct[of F], auto)
+
+lemma bounded_pos: "bounded S \<longleftrightarrow> (\<exists>b>0. \<forall>x\<in> S. norm x <= 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)")
+  by metis arith
+
+lemma bounded_Int[intro]: "bounded S \<or> bounded T \<Longrightarrow> bounded (S \<inter> T)"
+  by (metis Int_lower1 Int_lower2 bounded_subset)
+
+lemma bounded_diff[intro]: "bounded S ==> bounded (S - T)"
+apply (metis Diff_subset bounded_subset)
+done
+
+lemma bounded_insert[intro]:"bounded(insert x S) \<longleftrightarrow> bounded S"
+  by (metis Diff_cancel Un_empty_right Un_insert_right bounded_Un bounded_subset finite.emptyI finite_imp_bounded infinite_remove subset_insertI)
+
+lemma not_bounded_UNIV[simp, intro]:
+  "\<not> bounded (UNIV :: 'a::{real_normed_vector, perfect_space} set)"
+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
+  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" ..
+qed
+
+lemma bounded_linear_image:
+  assumes "bounded S" "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" unfolding bounded_pos by auto
+  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"
+    hence "norm x \<le> b" using b by auto
+    hence "norm (f x) \<le> B * b" using B(2) apply(erule_tac x=x in allE)
+      by (metis B(1) B(2) real_le_trans real_mult_le_cancel_iff2)
+  }
+  thus ?thesis unfolding bounded_pos apply(rule_tac x="b*B" in exI)
+    using b B real_mult_order[of b B] by (auto simp add: real_mult_commute)
+qed
+
+lemma bounded_scaling:
+  fixes S :: "'a::real_normed_vector set"
+  shows "bounded S \<Longrightarrow> bounded ((\<lambda>x. c *\<^sub>R x) ` S)"
+  apply (rule bounded_linear_image, assumption)
+  apply (rule scaleR.bounded_linear_right)
+  done
+
+lemma bounded_translation:
+  fixes S :: "'a::real_normed_vector set"
+  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" unfolding bounded_pos by auto
+  { fix x assume "x\<in>S"
+    hence "norm (a + x) \<le> b + norm a" using norm_triangle_ineq[of a x] b by auto
+  }
+  thus ?thesis unfolding bounded_pos using norm_ge_zero[of a] b(1) using add_strict_increasing[of b 0 "norm a"]
+    by (auto intro!: add exI[of _ "b + norm a"])
+qed
+
+
+text{* Some theorems on sups and infs using the notion "bounded". *}
+
+lemma bounded_real:
+  fixes S :: "real set"
+  shows "bounded S \<longleftrightarrow>  (\<exists>a. \<forall>x\<in>S. abs x <= a)"
+  by (simp add: bounded_iff)
+
+lemma bounded_has_rsup: assumes "bounded S" "S \<noteq> {}"
+  shows "\<forall>x\<in>S. x <= rsup S" and "\<forall>b. (\<forall>x\<in>S. x <= b) \<longrightarrow> rsup S <= b"
+proof
+  fix x 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
+  hence *:"S *<= a" using setleI[of S a] by (metis abs_le_interval_iff mem_def)
+  thus "x \<le> rsup S" using rsup[OF `S\<noteq>{}`] using assms(1)[unfolded bounded_real] using isLubD2[of UNIV S "rsup S" x] using `x\<in>S` by auto
+next
+  show "\<forall>b. (\<forall>x\<in>S. x \<le> b) \<longrightarrow> rsup S \<le> b" using assms
+  using rsup[of S, unfolded isLub_def isUb_def leastP_def setle_def setge_def]
+  apply (auto simp add: bounded_real)
+  by (auto simp add: isLub_def isUb_def leastP_def setle_def setge_def)
+qed
+
+lemma rsup_insert: assumes "bounded S"
+  shows "rsup(insert x S) = (if S = {} then x else max x (rsup S))"
+proof(cases "S={}")
+  case True thus ?thesis using rsup_finite_in[of "{x}"] by auto
+next
+  let ?S = "insert x S"
+  case False
+  hence *:"\<forall>x\<in>S. x \<le> rsup S" using bounded_has_rsup(1)[of S] using assms by auto
+  hence "insert x S *<= max x (rsup S)" unfolding setle_def by auto
+  hence "isLub UNIV ?S (rsup ?S)" using rsup[of ?S] by auto
+  moreover
+  have **:"isUb UNIV ?S (max x (rsup S))" unfolding isUb_def setle_def using * by auto
+  { fix y assume as:"isUb UNIV (insert x S) y"
+    hence "max x (rsup S) \<le> y" unfolding isUb_def using rsup_le[OF `S\<noteq>{}`]
+      unfolding setle_def by auto  }
+  hence "max x (rsup S) <=* isUb UNIV (insert x S)" unfolding setge_def Ball_def mem_def by auto
+  hence "isLub UNIV ?S (max x (rsup S))" using ** isLubI2[of UNIV ?S "max x (rsup S)"] unfolding Collect_def by auto
+  ultimately show ?thesis using real_isLub_unique[of UNIV ?S] using `S\<noteq>{}` by auto
+qed
+
+lemma sup_insert_finite: "finite S \<Longrightarrow> rsup(insert x S) = (if S = {} then x else max x (rsup S))"
+  apply (rule rsup_insert)
+  apply (rule finite_imp_bounded)
+  by simp
+
+lemma bounded_has_rinf:
+  assumes "bounded S"  "S \<noteq> {}"
+  shows "\<forall>x\<in>S. x >= rinf S" and "\<forall>b. (\<forall>x\<in>S. x >= b) \<longrightarrow> rinf S >= b"
+proof
+  fix x 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
+  hence *:"- a <=* S" using setgeI[of S "-a"] unfolding abs_le_interval_iff by auto
+  thus "x \<ge> rinf S" using rinf[OF `S\<noteq>{}`] using isGlbD2[of UNIV S "rinf S" x] using `x\<in>S` by auto
+next
+  show "\<forall>b. (\<forall>x\<in>S. x >= b) \<longrightarrow> rinf S \<ge> b" using assms
+  using rinf[of S, unfolded isGlb_def isLb_def greatestP_def setle_def setge_def]
+  apply (auto simp add: bounded_real)
+  by (auto simp add: isGlb_def isLb_def greatestP_def setle_def setge_def)
+qed
+
+(* TODO: Move this to RComplete.thy -- would need to include Glb into RComplete *)
+lemma real_isGlb_unique: "[| isGlb R S x; isGlb R S y |] ==> x = (y::real)"
+  apply (frule isGlb_isLb)
+  apply (frule_tac x = y in isGlb_isLb)
+  apply (blast intro!: order_antisym dest!: isGlb_le_isLb)
+  done
+
+lemma rinf_insert: assumes "bounded S"
+  shows "rinf(insert x S) = (if S = {} then x else min x (rinf S))" (is "?lhs = ?rhs")
+proof(cases "S={}")
+  case True thus ?thesis using rinf_finite_in[of "{x}"] by auto
+next
+  let ?S = "insert x S"
+  case False
+  hence *:"\<forall>x\<in>S. x \<ge> rinf S" using bounded_has_rinf(1)[of S] using assms by auto
+  hence "min x (rinf S) <=* insert x S" unfolding setge_def by auto
+  hence "isGlb UNIV ?S (rinf ?S)" using rinf[of ?S] by auto
+  moreover
+  have **:"isLb UNIV ?S (min x (rinf S))" unfolding isLb_def setge_def using * by auto
+  { fix y assume as:"isLb UNIV (insert x S) y"
+    hence "min x (rinf S) \<ge> y" unfolding isLb_def using rinf_ge[OF `S\<noteq>{}`]
+      unfolding setge_def by auto  }
+  hence "isLb UNIV (insert x S) *<= min x (rinf S)" unfolding setle_def Ball_def mem_def by auto
+  hence "isGlb UNIV ?S (min x (rinf S))" using ** isGlbI2[of UNIV ?S "min x (rinf S)"] unfolding Collect_def by auto
+  ultimately show ?thesis using real_isGlb_unique[of UNIV ?S] using `S\<noteq>{}` by auto
+qed
+
+lemma inf_insert_finite: "finite S ==> rinf(insert x S) = (if S = {} then x else min x (rinf S))"
+  by (rule rinf_insert, rule finite_imp_bounded, simp)
+
+subsection{* Compactness (the definition is the one based on convegent subsequences). *}
+
+definition
+  compact :: "'a::metric_space set \<Rightarrow> bool" where (* TODO: generalize *)
+  "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))"
+
+text {*
+  A metric space (or topological vector space) is said to have the
+  Heine-Borel property if every closed and bounded subset is compact.
+*}
+
+class heine_borel =
+  assumes bounded_imp_convergent_subsequence:
+    "bounded s \<Longrightarrow> \<forall>n. f n \<in> s
+      \<Longrightarrow> \<exists>l r. subseq r \<and> ((f \<circ> r) ---> l) sequentially"
+
+lemma bounded_closed_imp_compact:
+  fixes s::"'a::heine_borel set"
+  assumes "bounded s" and "closed s" shows "compact s"
+proof (unfold compact_def, clarify)
+  fix f :: "nat \<Rightarrow> 'a" assume f: "\<forall>n. f n \<in> s"
+  obtain l r where r: "subseq r" and l: "((f \<circ> r) ---> l) sequentially"
+    using bounded_imp_convergent_subsequence [OF `bounded s` `\<forall>n. f n \<in> s`] by auto
+  from f have fr: "\<forall>n. (f \<circ> r) n \<in> s" by simp
+  have "l \<in> s" using `closed s` fr l
+    unfolding closed_sequential_limits by blast
+  show "\<exists>l\<in>s. \<exists>r. subseq r \<and> ((f \<circ> r) ---> l) sequentially"
+    using `l \<in> s` r l by blast
+qed
+
+lemma subseq_bigger: assumes "subseq r" shows "n \<le> r n"
+proof(induct n)
+  show "0 \<le> r 0" by auto
+next
+  fix n assume "n \<le> r n"
+  moreover have "r n < r (Suc n)"
+    using assms [unfolded subseq_def] by auto
+  ultimately show "Suc n \<le> r (Suc n)" by auto
+qed
+
+lemma eventually_subseq:
+  assumes r: "subseq r"
+  shows "eventually P sequentially \<Longrightarrow> eventually (\<lambda>n. P (r n)) sequentially"
+unfolding eventually_sequentially
+by (metis subseq_bigger [OF r] le_trans)
+
+lemma lim_subseq:
+  "subseq r \<Longrightarrow> (s ---> l) sequentially \<Longrightarrow> ((s o r) ---> l) sequentially"
+unfolding tendsto_def eventually_sequentially o_def
+by (metis subseq_bigger le_trans)
+
+lemma num_Axiom: "EX! g. g 0 = e \<and> (\<forall>n. g (Suc n) = f n (g n))"
+  unfolding Ex1_def
+  apply (rule_tac x="nat_rec e f" in exI)
+  apply (rule conjI)+
+apply (rule def_nat_rec_0, simp)
+apply (rule allI, rule def_nat_rec_Suc, simp)
+apply (rule allI, rule impI, rule ext)
+apply (erule conjE)
+apply (induct_tac x)
+apply (simp add: nat_rec_0)
+apply (erule_tac x="n" in allE)
+apply (simp)
+done
+
+lemma convergent_bounded_increasing: fixes s ::"nat\<Rightarrow>real"
+  assumes "incseq s" and "\<forall>n. abs(s n) \<le> b"
+  shows "\<exists> l. \<forall>e::real>0. \<exists> N. \<forall>n \<ge> N.  abs(s n - l) < e"
+proof-
+  have "isUb UNIV (range s) b" using assms(2) and abs_le_D1 unfolding isUb_def and setle_def by auto
+  then obtain t where t:"isLub UNIV (range s) t" using reals_complete[of "range s" ] by auto
+  { fix e::real assume "e>0" and as:"\<forall>N. \<exists>n\<ge>N. \<not> \<bar>s n - t\<bar> < e"
+    { fix n::nat
+      obtain N where "N\<ge>n" and n:"\<bar>s N - t\<bar> \<ge> e" using as[THEN spec[where x=n]] by auto
+      have "t \<ge> s N" using isLub_isUb[OF t, unfolded isUb_def setle_def] by auto
+      with n have "s N \<le> t - e" using `e>0` by auto
+      hence "s n \<le> t - e" using assms(1)[unfolded incseq_def, THEN spec[where x=n], THEN spec[where x=N]] using `n\<le>N` by auto  }
+    hence "isUb UNIV (range s) (t - e)" unfolding isUb_def and setle_def by auto
+    hence False using isLub_le_isUb[OF t, of "t - e"] and `e>0` by auto  }
+  thus ?thesis by blast
+qed
+
+lemma convergent_bounded_monotone: fixes s::"nat \<Rightarrow> real"
+  assumes "\<forall>n. abs(s n) \<le> b" and "monoseq s"
+  shows "\<exists>l. \<forall>e::real>0. \<exists>N. \<forall>n\<ge>N. abs(s n - l) < e"
+  using convergent_bounded_increasing[of s b] assms using convergent_bounded_increasing[of "\<lambda>n. - s n" b]
+  unfolding monoseq_def incseq_def
+  apply auto unfolding minus_add_distrib[THEN sym, unfolded diff_minus[THEN sym]]
+  unfolding abs_minus_cancel by(rule_tac x="-l" in exI)auto
+
+lemma compact_real_lemma:
+  assumes "\<forall>n::nat. abs(s n) \<le> b"
+  shows "\<exists>(l::real) r. subseq r \<and> ((s \<circ> r) ---> l) sequentially"
+proof-
+  obtain r where r:"subseq r" "monoseq (\<lambda>n. s (r n))"
+    using seq_monosub[of s] by auto
+  thus ?thesis using convergent_bounded_monotone[of "\<lambda>n. s (r n)" b] and assms
+    unfolding tendsto_iff dist_norm eventually_sequentially by auto
+qed
+
+instance real :: heine_borel
+proof
+  fix s :: "real set" and f :: "nat \<Rightarrow> real"
+  assume s: "bounded s" and f: "\<forall>n. f n \<in> s"
+  then obtain b where b: "\<forall>n. abs (f n) \<le> b"
+    unfolding bounded_iff by auto
+  obtain l :: real and r :: "nat \<Rightarrow> nat" where
+    r: "subseq r" and l: "((f \<circ> r) ---> l) sequentially"
+    using compact_real_lemma [OF b] by auto
+  thus "\<exists>l r. subseq r \<and> ((f \<circ> r) ---> l) sequentially"
+    by auto
+qed
+
+lemma bounded_component: "bounded s \<Longrightarrow> bounded ((\<lambda>x. x $ i) ` s)"
+unfolding bounded_def
+apply clarify
+apply (rule_tac x="x $ i" in exI)
+apply (rule_tac x="e" in exI)
+apply clarify
+apply (rule order_trans [OF dist_nth_le], simp)
+done
+
+lemma compact_lemma:
+  fixes f :: "nat \<Rightarrow> 'a::heine_borel ^ 'n::finite"
+  assumes "bounded s" and "\<forall>n. f n \<in> s"
+  shows "\<forall>d.
+        \<exists>l r. subseq r \<and>
+        (\<forall>e>0. eventually (\<lambda>n. \<forall>i\<in>d. dist (f (r n) $ i) (l $ i) < e) sequentially)"
+proof
+  fix d::"'n set" have "finite d" by simp
+  thus "\<exists>l::'a ^ 'n. \<exists>r. subseq r \<and>
+      (\<forall>e>0. eventually (\<lambda>n. \<forall>i\<in>d. dist (f (r n) $ i) (l $ i) < e) sequentially)"
+  proof(induct d) case empty thus ?case unfolding subseq_def by auto
+  next case (insert k d)
+    have s': "bounded ((\<lambda>x. x $ k) ` s)" using `bounded s` by (rule bounded_component)
+    obtain l1::"'a^'n" and r1 where r1:"subseq r1" and lr1:"\<forall>e>0. eventually (\<lambda>n. \<forall>i\<in>d. dist (f (r1 n) $ i) (l1 $ i) < e) sequentially"
+      using insert(3) by auto
+    have f': "\<forall>n. f (r1 n) $ k \<in> (\<lambda>x. x $ k) ` s" using `\<forall>n. f n \<in> s` by simp
+    obtain l2 r2 where r2:"subseq r2" and lr2:"((\<lambda>i. f (r1 (r2 i)) $ k) ---> l2) sequentially"
+      using bounded_imp_convergent_subsequence[OF s' f'] unfolding o_def by auto
+    def r \<equiv> "r1 \<circ> r2" have r:"subseq r"
+      using r1 and r2 unfolding r_def o_def subseq_def by auto
+    moreover
+    def l \<equiv> "(\<chi> i. if i = k then l2 else l1$i)::'a^'n"
+    { fix e::real assume "e>0"
+      from lr1 `e>0` have N1:"eventually (\<lambda>n. \<forall>i\<in>d. dist (f (r1 n) $ i) (l1 $ i) < e) sequentially" by blast
+      from lr2 `e>0` have N2:"eventually (\<lambda>n. dist (f (r1 (r2 n)) $ k) l2 < e) sequentially" by (rule tendstoD)
+      from r2 N1 have N1': "eventually (\<lambda>n. \<forall>i\<in>d. dist (f (r1 (r2 n)) $ i) (l1 $ i) < e) sequentially"
+        by (rule eventually_subseq)
+      have "eventually (\<lambda>n. \<forall>i\<in>(insert k d). dist (f (r n) $ i) (l $ i) < e) sequentially"
+        using N1' N2 by (rule eventually_elim2, simp add: l_def r_def)
+    }
+    ultimately show ?case by auto
+  qed
+qed
+
+instance "^" :: (heine_borel, finite) heine_borel
+proof
+  fix s :: "('a ^ 'b) set" and f :: "nat \<Rightarrow> 'a ^ 'b"
+  assume s: "bounded s" and f: "\<forall>n. f n \<in> s"
+  then obtain l r where r: "subseq r"
+    and l: "\<forall>e>0. eventually (\<lambda>n. \<forall>i\<in>UNIV. dist (f (r n) $ i) (l $ i) < e) sequentially"
+    using compact_lemma [OF s f] by blast
+  let ?d = "UNIV::'b set"
+  { fix e::real assume "e>0"
+    hence "0 < e / (real_of_nat (card ?d))"
+      using zero_less_card_finite using divide_pos_pos[of e, of "real_of_nat (card ?d)"] by auto
+    with l have "eventually (\<lambda>n. \<forall>i. dist (f (r n) $ i) (l $ i) < e / (real_of_nat (card ?d))) sequentially"
+      by simp
+    moreover
+    { fix n assume n: "\<forall>i. dist (f (r n) $ i) (l $ i) < e / (real_of_nat (card ?d))"
+      have "dist (f (r n)) l \<le> (\<Sum>i\<in>?d. dist (f (r n) $ i) (l $ i))"
+        unfolding dist_vector_def using zero_le_dist by (rule setL2_le_setsum)
+      also have "\<dots> < (\<Sum>i\<in>?d. e / (real_of_nat (card ?d)))"
+        by (rule setsum_strict_mono) (simp_all add: n)
+      finally have "dist (f (r n)) l < e" by simp
+    }
+    ultimately have "eventually (\<lambda>n. dist (f (r n)) l < e) sequentially"
+      by (rule eventually_elim1)
+  }
+  hence *:"((f \<circ> r) ---> l) sequentially" unfolding o_def tendsto_iff by simp
+  with r show "\<exists>l r. subseq r \<and> ((f \<circ> r) ---> l) sequentially" by auto
+qed
+
+lemma bounded_fst: "bounded s \<Longrightarrow> bounded (fst ` s)"
+unfolding bounded_def
+apply clarify
+apply (rule_tac x="a" in exI)
+apply (rule_tac x="e" in exI)
+apply clarsimp
+apply (drule (1) bspec)
+apply (simp add: dist_Pair_Pair)
+apply (erule order_trans [OF real_sqrt_sum_squares_ge1])
+done
+
+lemma bounded_snd: "bounded s \<Longrightarrow> bounded (snd ` s)"
+unfolding bounded_def
+apply clarify
+apply (rule_tac x="b" in exI)
+apply (rule_tac x="e" in exI)
+apply clarsimp
+apply (drule (1) bspec)
+apply (simp add: dist_Pair_Pair)
+apply (erule order_trans [OF real_sqrt_sum_squares_ge2])
+done
+
+instance "*" :: (heine_borel, heine_borel) heine_borel
+proof
+  fix s :: "('a * 'b) set" and f :: "nat \<Rightarrow> 'a * 'b"
+  assume s: "bounded s" and f: "\<forall>n. f n \<in> s"
+  from s have s1: "bounded (fst ` s)" by (rule bounded_fst)
+  from f have f1: "\<forall>n. fst (f n) \<in> fst ` s" by simp
+  obtain l1 r1 where r1: "subseq r1"
+    and l1: "((\<lambda>n. fst (f (r1 n))) ---> l1) sequentially"
+    using bounded_imp_convergent_subsequence [OF s1 f1]
+    unfolding o_def by fast
+  from s have s2: "bounded (snd ` s)" by (rule bounded_snd)
+  from f have f2: "\<forall>n. snd (f (r1 n)) \<in> snd ` s" by simp
+  obtain l2 r2 where r2: "subseq r2"
+    and l2: "((\<lambda>n. snd (f (r1 (r2 n)))) ---> l2) sequentially"
+    using bounded_imp_convergent_subsequence [OF s2 f2]
+    unfolding o_def by fast
+  have l1': "((\<lambda>n. fst (f (r1 (r2 n)))) ---> l1) sequentially"
+    using lim_subseq [OF r2 l1] unfolding o_def .
+  have l: "((f \<circ> (r1 \<circ> r2)) ---> (l1, l2)) sequentially"
+    using tendsto_Pair [OF l1' l2] unfolding o_def by simp
+  have r: "subseq (r1 \<circ> r2)"
+    using r1 r2 unfolding subseq_def by simp
+  show "\<exists>l r. subseq r \<and> ((f \<circ> r) ---> l) sequentially"
+    using l r by fast
+qed
+
+subsection{* Completeness. *}
+
+lemma cauchy_def:
+  "Cauchy s \<longleftrightarrow> (\<forall>e>0. \<exists>N. \<forall>m n. m \<ge> N \<and> n \<ge> N --> dist(s m)(s n) < e)"
+unfolding Cauchy_def by blast
+
+definition
+  complete :: "'a::metric_space set \<Rightarrow> bool" where
+  "complete s \<longleftrightarrow> (\<forall>f. (\<forall>n. f n \<in> s) \<and> Cauchy f
+                      --> (\<exists>l \<in> s. (f ---> l) sequentially))"
+
+lemma cauchy: "Cauchy s \<longleftrightarrow> (\<forall>e>0.\<exists> N::nat. \<forall>n\<ge>N. dist(s n)(s N) < e)" (is "?lhs = ?rhs")
+proof-
+  { assume ?rhs
+    { fix e::real
+      assume "e>0"
+      with `?rhs` obtain N where N:"\<forall>n\<ge>N. dist (s n) (s N) < e/2"
+        by (erule_tac x="e/2" in allE) auto
+      { fix n m
+        assume nm:"N \<le> m \<and> N \<le> n"
+        hence "dist (s m) (s n) < e" using N
+          using dist_triangle_half_l[of "s m" "s N" "e" "s n"]
+          by blast
+      }
+      hence "\<exists>N. \<forall>m n. N \<le> m \<and> N \<le> n \<longrightarrow> dist (s m) (s n) < e"
+        by blast
+    }
+    hence ?lhs
+      unfolding cauchy_def
+      by blast
+  }
+  thus ?thesis
+    unfolding cauchy_def
+    using dist_triangle_half_l
+    by blast
+qed
+
+lemma convergent_imp_cauchy:
+ "(s ---> l) sequentially ==> Cauchy s"
+proof(simp only: cauchy_def, rule, rule)
+  fix e::real assume "e>0" "(s ---> l) sequentially"
+  then obtain N::nat where N:"\<forall>n\<ge>N. dist (s n) l < e/2" unfolding Lim_sequentially by(erule_tac x="e/2" in allE) auto
+  thus "\<exists>N. \<forall>m n. N \<le> m \<and> N \<le> n \<longrightarrow> dist (s m) (s n) < e"  using dist_triangle_half_l[of _ l e _] by (rule_tac x=N in exI) auto
+qed
+
+lemma cauchy_imp_bounded: assumes "Cauchy s" shows "bounded {y. (\<exists>n::nat. y = s n)}"
+proof-
+  from assms obtain N::nat where "\<forall>m n. N \<le> m \<and> N \<le> n \<longrightarrow> dist (s m) (s n) < 1" unfolding cauchy_def apply(erule_tac x= 1 in allE) by auto
+  hence N:"\<forall>n. N \<le> n \<longrightarrow> dist (s N) (s n) < 1" by auto
+  moreover
+  have "bounded (s ` {0..N})" using finite_imp_bounded[of "s ` {1..N}"] by auto
+  then obtain a where a:"\<forall>x\<in>s ` {0..N}. dist (s N) x \<le> a"
+    unfolding bounded_any_center [where a="s N"] by auto
+  ultimately show "?thesis"
+    unfolding bounded_any_center [where a="s N"]
+    apply(rule_tac x="max a 1" in exI) apply auto
+    apply(erule_tac x=n in allE) apply(erule_tac x=n in ballE) by auto
+qed
+
+lemma compact_imp_complete: assumes "compact s" shows "complete s"
+proof-
+  { fix f assume as: "(\<forall>n::nat. f n \<in> s)" "Cauchy f"
+    from as(1) obtain l r where lr: "l\<in>s" "subseq r" "((f \<circ> r) ---> l) sequentially" using assms unfolding compact_def by blast
+
+    note lr' = subseq_bigger [OF lr(2)]
+
+    { fix e::real assume "e>0"
+      from as(2) obtain N where N:"\<forall>m n. N \<le> m \<and> N \<le> n \<longrightarrow> dist (f m) (f n) < e/2" unfolding cauchy_def using `e>0` apply (erule_tac x="e/2" in allE) by auto
+      from lr(3)[unfolded Lim_sequentially, THEN spec[where x="e/2"]] obtain M where M:"\<forall>n\<ge>M. dist ((f \<circ> r) n) l < e/2" using `e>0` by auto
+      { fix n::nat assume n:"n \<ge> max N M"
+        have "dist ((f \<circ> r) n) l < e/2" using n M by auto
+        moreover have "r n \<ge> N" using lr'[of n] n by auto
+        hence "dist (f n) ((f \<circ> r) n) < e / 2" using N using n by auto
+        ultimately have "dist (f n) l < e" using dist_triangle_half_r[of "f (r n)" "f n" e l] by (auto simp add: dist_commute)  }
+      hence "\<exists>N. \<forall>n\<ge>N. dist (f n) l < e" by blast  }
+    hence "\<exists>l\<in>s. (f ---> l) sequentially" using `l\<in>s` unfolding Lim_sequentially by auto  }
+  thus ?thesis unfolding complete_def by auto
+qed
+
+instance heine_borel < complete_space
+proof
+  fix f :: "nat \<Rightarrow> 'a" assume "Cauchy f"
+  hence "bounded (range f)" unfolding image_def
+    using cauchy_imp_bounded [of f] by auto
+  hence "compact (closure (range f))"
+    using bounded_closed_imp_compact [of "closure (range f)"] by auto
+  hence "complete (closure (range f))"
+    using compact_imp_complete by auto
+  moreover have "\<forall>n. f n \<in> closure (range f)"
+    using closure_subset [of "range f"] by auto
+  ultimately have "\<exists>l\<in>closure (range f). (f ---> l) sequentially"
+    using `Cauchy f` unfolding complete_def by auto
+  then show "convergent f"
+    unfolding convergent_def LIMSEQ_conv_tendsto [symmetric] by auto
+qed
+
+lemma complete_univ: "complete (UNIV :: 'a::complete_space set)"
+proof(simp add: complete_def, rule, rule)
+  fix f :: "nat \<Rightarrow> 'a" assume "Cauchy f"
+  hence "convergent f" by (rule Cauchy_convergent)
+  hence "\<exists>l. f ----> l" unfolding convergent_def .  
+  thus "\<exists>l. (f ---> l) sequentially" unfolding LIMSEQ_conv_tendsto .
+qed
+
+lemma complete_imp_closed: assumes "complete s" shows "closed s"
+proof -
+  { fix x assume "x islimpt s"
+    then obtain f where f: "\<forall>n. f n \<in> s - {x}" "(f ---> x) sequentially"
+      unfolding islimpt_sequential by auto
+    then obtain l where l: "l\<in>s" "(f ---> l) sequentially"
+      using `complete s`[unfolded complete_def] using convergent_imp_cauchy[of f x] by auto
+    hence "x \<in> s"  using Lim_unique[of sequentially f l x] trivial_limit_sequentially f(2) by auto
+  }
+  thus "closed s" unfolding closed_limpt by auto
+qed
+
+lemma complete_eq_closed:
+  fixes s :: "'a::complete_space set"
+  shows "complete s \<longleftrightarrow> closed s" (is "?lhs = ?rhs")
+proof
+  assume ?lhs thus ?rhs by (rule complete_imp_closed)
+next
+  assume ?rhs
+  { fix f assume as:"\<forall>n::nat. f n \<in> s" "Cauchy f"
+    then obtain l where "(f ---> l) sequentially" using complete_univ[unfolded complete_def, THEN spec[where x=f]] by auto
+    hence "\<exists>l\<in>s. (f ---> l) sequentially" using `?rhs`[unfolded closed_sequential_limits, THEN spec[where x=f], THEN spec[where x=l]] using as(1) by auto  }
+  thus ?lhs unfolding complete_def by auto
+qed
+
+lemma convergent_eq_cauchy:
+  fixes s :: "nat \<Rightarrow> 'a::complete_space"
+  shows "(\<exists>l. (s ---> l) sequentially) \<longleftrightarrow> Cauchy s" (is "?lhs = ?rhs")
+proof
+  assume ?lhs then obtain l where "(s ---> l) sequentially" by auto
+  thus ?rhs using convergent_imp_cauchy by auto
+next
+  assume ?rhs thus ?lhs using complete_univ[unfolded complete_def, THEN spec[where x=s]] by auto
+qed
+
+lemma convergent_imp_bounded:
+  fixes s :: "nat \<Rightarrow> 'a::metric_space"
+  shows "(s ---> l) sequentially ==> bounded (s ` (UNIV::(nat set)))"
+  using convergent_imp_cauchy[of s]
+  using cauchy_imp_bounded[of s]
+  unfolding image_def
+  by auto
+
+subsection{* Total boundedness. *}
+
+fun helper_1::"('a::metric_space set) \<Rightarrow> real \<Rightarrow> nat \<Rightarrow> 'a" where
+  "helper_1 s e n = (SOME y::'a. y \<in> s \<and> (\<forall>m<n. \<not> (dist (helper_1 s e m) y < e)))"
+declare helper_1.simps[simp del]
+
+lemma compact_imp_totally_bounded:
+  assumes "compact s"
+  shows "\<forall>e>0. \<exists>k. finite k \<and> k \<subseteq> s \<and> s \<subseteq> (\<Union>((\<lambda>x. ball x e) ` k))"
+proof(rule, rule, rule ccontr)
+  fix e::real assume "e>0" and assm:"\<not> (\<exists>k. finite k \<and> k \<subseteq> s \<and> s \<subseteq> \<Union>(\<lambda>x. ball x e) ` k)"
+  def x \<equiv> "helper_1 s e"
+  { fix n
+    have "x n \<in> s \<and> (\<forall>m<n. \<not> dist (x m) (x n) < e)"
+    proof(induct_tac rule:nat_less_induct)
+      fix n  def Q \<equiv> "(\<lambda>y. y \<in> s \<and> (\<forall>m<n. \<not> dist (x m) y < e))"
+      assume as:"\<forall>m<n. x m \<in> s \<and> (\<forall>ma<m. \<not> dist (x ma) (x m) < e)"
+      have "\<not> s \<subseteq> (\<Union>x\<in>x ` {0..<n}. ball x e)" using assm apply simp apply(erule_tac x="x ` {0 ..< n}" in allE) using as by auto
+      then obtain z where z:"z\<in>s" "z \<notin> (\<Union>x\<in>x ` {0..<n}. ball x e)" unfolding subset_eq by auto
+      have "Q (x n)" unfolding x_def and helper_1.simps[of s e n]
+        apply(rule someI2[where a=z]) unfolding x_def[symmetric] and Q_def using z by auto
+      thus "x n \<in> s \<and> (\<forall>m<n. \<not> dist (x m) (x n) < e)" unfolding Q_def by auto
+    qed }
+  hence "\<forall>n::nat. x n \<in> s" and x:"\<forall>n. \<forall>m < n. \<not> (dist (x m) (x n) < e)" by blast+
+  then obtain l r where "l\<in>s" and r:"subseq r" and "((x \<circ> r) ---> l) sequentially" using assms(1)[unfolded compact_def, THEN spec[where x=x]] by auto
+  from this(3) have "Cauchy (x \<circ> r)" using convergent_imp_cauchy by auto
+  then obtain N::nat where N:"\<forall>m n. N \<le> m \<and> N \<le> n \<longrightarrow> dist ((x \<circ> r) m) ((x \<circ> r) n) < e" unfolding cauchy_def using `e>0` by auto
+  show False
+    using N[THEN spec[where x=N], THEN spec[where x="N+1"]]
+    using r[unfolded subseq_def, THEN spec[where x=N], THEN spec[where x="N+1"]]
+    using x[THEN spec[where x="r (N+1)"], THEN spec[where x="r (N)"]] by auto
+qed
+
+subsection{* Heine-Borel theorem (following Burkill \& Burkill vol. 2) *}
+
+lemma heine_borel_lemma: fixes s::"'a::metric_space set"
+  assumes "compact s"  "s \<subseteq> (\<Union> t)"  "\<forall>b \<in> t. open b"
+  shows "\<exists>e>0. \<forall>x \<in> s. \<exists>b \<in> t. ball x e \<subseteq> b"
+proof(rule ccontr)
+  assume "\<not> (\<exists>e>0. \<forall>x\<in>s. \<exists>b\<in>t. ball x e \<subseteq> b)"
+  hence cont:"\<forall>e>0. \<exists>x\<in>s. \<forall>xa\<in>t. \<not> (ball x e \<subseteq> xa)" by auto
+  { fix n::nat
+    have "1 / real (n + 1) > 0" by auto
+    hence "\<exists>x. x\<in>s \<and> (\<forall>xa\<in>t. \<not> (ball x (inverse (real (n+1))) \<subseteq> xa))" using cont unfolding Bex_def by auto }
+  hence "\<forall>n::nat. \<exists>x. x \<in> s \<and> (\<forall>xa\<in>t. \<not> ball x (inverse (real (n + 1))) \<subseteq> xa)" by auto
+  then obtain f where f:"\<forall>n::nat. f n \<in> s \<and> (\<forall>xa\<in>t. \<not> ball (f n) (inverse (real (n + 1))) \<subseteq> xa)"
+    using choice[of "\<lambda>n::nat. \<lambda>x. x\<in>s \<and> (\<forall>xa\<in>t. \<not> ball x (inverse (real (n + 1))) \<subseteq> xa)"] by auto
+
+  then obtain l r where l:"l\<in>s" and r:"subseq r" and lr:"((f \<circ> r) ---> l) sequentially"
+    using assms(1)[unfolded compact_def, THEN spec[where x=f]] by auto
+
+  obtain b where "l\<in>b" "b\<in>t" using assms(2) and l by auto
+  then obtain e where "e>0" and e:"\<forall>z. dist z l < e \<longrightarrow> z\<in>b"
+    using assms(3)[THEN bspec[where x=b]] unfolding open_dist by auto
+
+  then obtain N1 where N1:"\<forall>n\<ge>N1. dist ((f \<circ> r) n) l < e / 2"
+    using lr[unfolded Lim_sequentially, THEN spec[where x="e/2"]] by auto
+
+  obtain N2::nat where N2:"N2>0" "inverse (real N2) < e /2" using real_arch_inv[of "e/2"] and `e>0` by auto
+  have N2':"inverse (real (r (N1 + N2) +1 )) < e/2"
+    apply(rule order_less_trans) apply(rule less_imp_inverse_less) using N2
+    using subseq_bigger[OF r, of "N1 + N2"] by auto
+
+  def x \<equiv> "(f (r (N1 + N2)))"
+  have x:"\<not> ball x (inverse (real (r (N1 + N2) + 1))) \<subseteq> b" unfolding x_def
+    using f[THEN spec[where x="r (N1 + N2)"]] using `b\<in>t` by auto
+  have "\<exists>y\<in>ball x (inverse (real (r (N1 + N2) + 1))). y\<notin>b" apply(rule ccontr) using x by auto
+  then obtain y where y:"y \<in> ball x (inverse (real (r (N1 + N2) + 1)))" "y \<notin> b" by auto
+
+  have "dist x l < e/2" using N1 unfolding x_def o_def by auto
+  hence "dist y l < e" using y N2' using dist_triangle[of y l x]by (auto simp add:dist_commute)
+
+  thus False using e and `y\<notin>b` by auto
+qed
+
+lemma compact_imp_heine_borel: "compact s ==> (\<forall>f. (\<forall>t \<in> f. open t) \<and> s \<subseteq> (\<Union> f)
+               \<longrightarrow> (\<exists>f'. f' \<subseteq> f \<and> finite f' \<and> s \<subseteq> (\<Union> f')))"
+proof clarify
+  fix f assume "compact s" " \<forall>t\<in>f. open t" "s \<subseteq> \<Union>f"
+  then obtain e::real where "e>0" and "\<forall>x\<in>s. \<exists>b\<in>f. ball x e \<subseteq> b" using heine_borel_lemma[of s f] by auto
+  hence "\<forall>x\<in>s. \<exists>b. b\<in>f \<and> ball x e \<subseteq> b" by auto
+  hence "\<exists>bb. \<forall>x\<in>s. bb x \<in>f \<and> ball x e \<subseteq> bb x" using bchoice[of s "\<lambda>x b. b\<in>f \<and> ball x e \<subseteq> b"] by auto
+  then obtain  bb where bb:"\<forall>x\<in>s. (bb x) \<in> f \<and> ball x e \<subseteq> (bb x)" by blast
+
+  from `compact s` have  "\<exists> k. finite k \<and> k \<subseteq> s \<and> s \<subseteq> \<Union>(\<lambda>x. ball x e) ` k" using compact_imp_totally_bounded[of s] `e>0` by auto
+  then obtain k where k:"finite k" "k \<subseteq> s" "s \<subseteq> \<Union>(\<lambda>x. ball x e) ` k" by auto
+
+  have "finite (bb ` k)" using k(1) by auto
+  moreover
+  { fix x assume "x\<in>s"
+    hence "x\<in>\<Union>(\<lambda>x. ball x e) ` k" using k(3)  unfolding subset_eq by auto
+    hence "\<exists>X\<in>bb ` k. x \<in> X" using bb k(2) by blast
+    hence "x \<in> \<Union>(bb ` k)" using  Union_iff[of x "bb ` k"] by auto
+  }
+  ultimately show "\<exists>f'\<subseteq>f. finite f' \<and> s \<subseteq> \<Union>f'" using bb k(2) by (rule_tac x="bb ` k" in exI) auto
+qed
+
+subsection{* Bolzano-Weierstrass property. *}
+
+lemma heine_borel_imp_bolzano_weierstrass:
+  assumes "\<forall>f. (\<forall>t \<in> f. open t) \<and> s \<subseteq> (\<Union> f) --> (\<exists>f'. f' \<subseteq> f \<and> finite f' \<and> s \<subseteq> (\<Union> f'))"
+          "infinite t"  "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)" 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"
+    using assms(1)[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
+  { fix x y assume "x\<in>t" "y\<in>t" "f x = f y"
+    hence "x \<in> f x"  "y \<in> f x \<longrightarrow> y = x" using f[THEN bspec[where x=x]] and `t\<subseteq>s` by auto
+    hence "x = y" using `f x = f y` and f[THEN bspec[where x=y]] and `y\<in>t` and `t\<subseteq>s` by auto  }
+  hence "infinite (f ` t)" using assms(2) using finite_imageD[unfolded inj_on_def, of f t] by auto
+  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" using g'[THEN bspec[where x=h]] by auto
+    hence "y = x" using f[THEN bspec[where x=y]] and `x\<in>t` and `x\<in>h`[unfolded `h = f y`] by auto
+    hence False using `f x \<notin> g` `h\<in>g` unfolding `h = f y` by auto  }
+  hence "f ` t \<subseteq> g" by auto
+  ultimately show False using g(2) using finite_subset by auto
+qed
+
+subsection{* Complete the chain of compactness variants. *}
+
+primrec helper_2::"(real \<Rightarrow> 'a::metric_space) \<Rightarrow> nat \<Rightarrow> 'a" where
+  "helper_2 beyond 0 = beyond 0" |
+  "helper_2 beyond (Suc n) = beyond (dist undefined (helper_2 beyond n) + 1 )"
+
+lemma bolzano_weierstrass_imp_bounded: fixes s::"'a::metric_space set"
+  assumes "\<forall>t. infinite t \<and> t \<subseteq> s --> (\<exists>x \<in> s. x islimpt t)"
+  shows "bounded s"
+proof(rule ccontr)
+  assume "\<not> bounded s"
+  then obtain beyond where "\<forall>a. beyond a \<in>s \<and> \<not> dist undefined (beyond a) \<le> a"
+    unfolding bounded_any_center [where a=undefined]
+    apply simp using choice[of "\<lambda>a x. x\<in>s \<and> \<not> dist undefined x \<le> a"] by auto
+  hence beyond:"\<And>a. beyond a \<in>s" "\<And>a. dist undefined (beyond a) > a"
+    unfolding linorder_not_le by auto
+  def x \<equiv> "helper_2 beyond"
+
+  { fix m n ::nat assume "m<n"
+    hence "dist undefined (x m) + 1 < dist undefined (x n)"
+    proof(induct n)
+      case 0 thus ?case by auto
+    next
+      case (Suc n)
+      have *:"dist undefined (x n) + 1 < dist undefined (x (Suc n))"
+        unfolding x_def and helper_2.simps
+        using beyond(2)[of "dist undefined (helper_2 beyond n) + 1"] by auto
+      thus ?case proof(cases "m < n")
+        case True thus ?thesis using Suc and * by auto
+      next
+        case False hence "m = n" using Suc(2) by auto
+        thus ?thesis using * by auto
+      qed
+    qed  } note * = this
+  { fix m n ::nat assume "m\<noteq>n"
+    have "1 < dist (x m) (x n)"
+    proof(cases "m<n")
+      case True
+      hence "1 < dist undefined (x n) - dist undefined (x m)" using *[of m n] by auto
+      thus ?thesis using dist_triangle [of undefined "x n" "x m"] by arith
+    next
+      case False hence "n<m" using `m\<noteq>n` by auto
+      hence "1 < dist undefined (x m) - dist undefined (x n)" using *[of n m] by auto
+      thus ?thesis using dist_triangle2 [of undefined "x m" "x n"] by arith
+    qed  } note ** = this
+  { fix a b assume "x a = x b" "a \<noteq> b"
+    hence False using **[of a b] by auto  }
+  hence "inj x" unfolding inj_on_def by auto
+  moreover
+  { fix n::nat
+    have "x n \<in> s"
+    proof(cases "n = 0")
+      case True thus ?thesis unfolding x_def using beyond by auto
+    next
+      case False then obtain z where "n = Suc z" using not0_implies_Suc by auto
+      thus ?thesis unfolding x_def using beyond by auto
+    qed  }
+  ultimately have "infinite (range x) \<and> range x \<subseteq> s" unfolding x_def using range_inj_infinite[of "helper_2 beyond"] using beyond(1) by auto
+
+  then obtain l where "l\<in>s" and l:"l islimpt range x" using assms[THEN spec[where x="range x"]] by auto
+  then obtain y where "x y \<noteq> l" and y:"dist (x y) l < 1/2" unfolding islimpt_approachable apply(erule_tac x="1/2" in allE) by auto
+  then obtain z where "x z \<noteq> l" and z:"dist (x z) l < dist (x y) l" using l[unfolded islimpt_approachable, THEN spec[where x="dist (x y) l"]]
+    unfolding dist_nz by auto
+  show False using y and z and dist_triangle_half_l[of "x y" l 1 "x z"] and **[of y z] by auto
+qed
+
+lemma sequence_infinite_lemma:
+  fixes l :: "'a::metric_space" (* TODO: generalize *)
+  assumes "\<forall>n::nat. (f n  \<noteq> l)"  "(f ---> l) sequentially"
+  shows "infinite {y. (\<exists> n. y = f n)}"
+proof(rule ccontr)
+  let ?A = "(\<lambda>x. dist x l) ` {y. \<exists>n. y = f n}"
+  assume "\<not> infinite {y. \<exists>n. y = f n}"
+  hence **:"finite ?A" "?A \<noteq> {}" by auto
+  obtain k where k:"dist (f k) l = Min ?A" using Min_in[OF **] by auto
+  have "0 < Min ?A" using assms(1) unfolding dist_nz unfolding Min_gr_iff[OF **] by auto
+  then obtain N where "dist (f N) l < Min ?A" using assms(2)[unfolded Lim_sequentially, THEN spec[where x="Min ?A"]] by auto
+  moreover have "dist (f N) l \<in> ?A" by auto
+  ultimately show False using Min_le[OF **(1), of "dist (f N) l"] by auto
+qed
+
+lemma sequence_unique_limpt:
+  fixes l :: "'a::metric_space" (* TODO: generalize *)
+  assumes "\<forall>n::nat. (f n \<noteq> l)"  "(f ---> l) sequentially"  "l' islimpt {y.  (\<exists>n. y = f n)}"
+  shows "l' = l"
+proof(rule ccontr)
+  def e \<equiv> "dist l' l"
+  assume "l' \<noteq> l" hence "e>0" unfolding dist_nz e_def by auto
+  then obtain N::nat where N:"\<forall>n\<ge>N. dist (f n) l < e / 2"
+    using assms(2)[unfolded Lim_sequentially, THEN spec[where x="e/2"]] by auto
+  def d \<equiv> "Min (insert (e/2) ((\<lambda>n. if dist (f n) l' = 0 then e/2 else dist (f n) l') ` {0 .. N}))"
+  have "d>0" using `e>0` unfolding d_def e_def using zero_le_dist[of _ l', unfolded order_le_less] by auto
+  obtain k where k:"f k \<noteq> l'"  "dist (f k) l' < d" using `d>0` and assms(3)[unfolded islimpt_approachable, THEN spec[where x="d"]] by auto
+  have "k\<ge>N" using k(1)[unfolded dist_nz] using k(2)[unfolded d_def]
+    by force
+  hence "dist l' l < e" using N[THEN spec[where x=k]] using k(2)[unfolded d_def] and dist_triangle_half_r[of "f k" l' e l] by auto
+  thus False unfolding e_def by auto
+qed
+
+lemma bolzano_weierstrass_imp_closed:
+  fixes s :: "'a::metric_space set" (* TODO: can this be generalized? *)
+  assumes "\<forall>t. infinite t \<and> t \<subseteq> s --> (\<exists>x \<in> s. x islimpt t)"
+  shows "closed s"
+proof-
+  { fix x l assume as: "\<forall>n::nat. x n \<in> s" "(x ---> l) sequentially"
+    hence "l \<in> s"
+    proof(cases "\<forall>n. x n \<noteq> l")
+      case False thus "l\<in>s" using as(1) by auto
+    next
+      case True note cas = this
+      with as(2) have "infinite {y. \<exists>n. y = x n}" using sequence_infinite_lemma[of x l] by auto
+      then obtain l' where "l'\<in>s" "l' islimpt {y. \<exists>n. y = x n}" using assms[THEN spec[where x="{y. \<exists>n. y = x n}"]] as(1) by auto
+      thus "l\<in>s" using sequence_unique_limpt[of x l l'] using as cas by auto
+    qed  }
+  thus ?thesis unfolding closed_sequential_limits by fast
+qed
+
+text{* Hence express everything as an equivalence.   *}
+
+lemma compact_eq_heine_borel:
+  fixes s :: "'a::heine_borel set"
+  shows "compact s \<longleftrightarrow>
+           (\<forall>f. (\<forall>t \<in> f. open t) \<and> s \<subseteq> (\<Union> f)
+               --> (\<exists>f'. f' \<subseteq> f \<and> finite f' \<and> s \<subseteq> (\<Union> f')))" (is "?lhs = ?rhs")
+proof
+  assume ?lhs thus ?rhs using compact_imp_heine_borel[of s] by blast
+next
+  assume ?rhs
+  hence "\<forall>t. infinite t \<and> t \<subseteq> s \<longrightarrow> (\<exists>x\<in>s. x islimpt t)"
+    by (blast intro: heine_borel_imp_bolzano_weierstrass[of s])
+  thus ?lhs using bolzano_weierstrass_imp_bounded[of s] bolzano_weierstrass_imp_closed[of s] bounded_closed_imp_compact[of s] by blast
+qed
+
+lemma compact_eq_bolzano_weierstrass:
+  fixes s :: "'a::heine_borel set"
+  shows "compact s \<longleftrightarrow> (\<forall>t. infinite t \<and> t \<subseteq> s --> (\<exists>x \<in> s. x islimpt t))" (is "?lhs = ?rhs")
+proof
+  assume ?lhs thus ?rhs unfolding compact_eq_heine_borel using heine_borel_imp_bolzano_weierstrass[of s] by auto
+next
+  assume ?rhs thus ?lhs using bolzano_weierstrass_imp_bounded bolzano_weierstrass_imp_closed bounded_closed_imp_compact by auto
+qed
+
+lemma compact_eq_bounded_closed:
+  fixes s :: "'a::heine_borel set"
+  shows "compact s \<longleftrightarrow> bounded s \<and> closed s"  (is "?lhs = ?rhs")
+proof
+  assume ?lhs thus ?rhs unfolding compact_eq_bolzano_weierstrass using bolzano_weierstrass_imp_bounded bolzano_weierstrass_imp_closed by auto
+next
+  assume ?rhs thus ?lhs using bounded_closed_imp_compact by auto
+qed
+
+lemma compact_imp_bounded:
+  fixes s :: "'a::metric_space set"
+  shows "compact s ==> bounded s"
+proof -
+  assume "compact s"
+  hence "\<forall>f. (\<forall>t\<in>f. open t) \<and> s \<subseteq> \<Union>f \<longrightarrow> (\<exists>f'\<subseteq>f. finite f' \<and> s \<subseteq> \<Union>f')"
+    by (rule compact_imp_heine_borel)
+  hence "\<forall>t. infinite t \<and> t \<subseteq> s \<longrightarrow> (\<exists>x \<in> s. x islimpt t)"
+    using heine_borel_imp_bolzano_weierstrass[of s] by auto
+  thus "bounded s"
+    by (rule bolzano_weierstrass_imp_bounded)
+qed
+
+lemma compact_imp_closed:
+  fixes s :: "'a::metric_space set"
+  shows "compact s ==> closed s"
+proof -
+  assume "compact s"
+  hence "\<forall>f. (\<forall>t\<in>f. open t) \<and> s \<subseteq> \<Union>f \<longrightarrow> (\<exists>f'\<subseteq>f. finite f' \<and> s \<subseteq> \<Union>f')"
+    by (rule compact_imp_heine_borel)
+  hence "\<forall>t. infinite t \<and> t \<subseteq> s \<longrightarrow> (\<exists>x \<in> s. x islimpt t)"
+    using heine_borel_imp_bolzano_weierstrass[of s] by auto
+  thus "closed s"
+    by (rule bolzano_weierstrass_imp_closed)
+qed
+
+text{* In particular, some common special cases. *}
+
+lemma compact_empty[simp]:
+ "compact {}"
+  unfolding compact_def
+  by simp
+
+(* TODO: can any of the next 3 lemmas be generalized to metric spaces? *)
+
+  (* FIXME : Rename *)
+lemma compact_union[intro]:
+  fixes s t :: "'a::heine_borel set"
+  shows "compact s \<Longrightarrow> compact t ==> compact (s \<union> t)"
+  unfolding compact_eq_bounded_closed
+  using bounded_Un[of s t]
+  using closed_Un[of s t]
+  by simp
+
+lemma compact_inter[intro]:
+  fixes s t :: "'a::heine_borel set"
+  shows "compact s \<Longrightarrow> compact t ==> compact (s \<inter> t)"
+  unfolding compact_eq_bounded_closed
+  using bounded_Int[of s t]
+  using closed_Int[of s t]
+  by simp
+
+lemma compact_inter_closed[intro]:
+  fixes s t :: "'a::heine_borel set"
+  shows "compact s \<Longrightarrow> closed t ==> compact (s \<inter> t)"
+  unfolding compact_eq_bounded_closed
+  using closed_Int[of s t]
+  using bounded_subset[of "s \<inter> t" s]
+  by blast
+
+lemma closed_inter_compact[intro]:
+  fixes s t :: "'a::heine_borel set"
+  shows "closed s \<Longrightarrow> compact t ==> compact (s \<inter> t)"
+proof-
+  assume "closed s" "compact t"
+  moreover
+  have "s \<inter> t = t \<inter> s" by auto ultimately
+  show ?thesis
+    using compact_inter_closed[of t s]
+    by auto
+qed
+
+lemma closed_sing [simp]:
+  fixes a :: "'a::metric_space"
+  shows "closed {a}"
+  apply (clarsimp simp add: closed_def open_dist)
+  apply (rule ccontr)
+  apply (drule_tac x="dist x a" in spec)
+  apply (simp add: dist_nz dist_commute)
+  done
+
+lemma finite_imp_closed:
+  fixes s :: "'a::metric_space set"
+  shows "finite s ==> closed s"
+proof (induct set: finite)
+  case empty show "closed {}" by simp
+next
+  case (insert x F)
+  hence "closed ({x} \<union> F)" by (simp only: closed_Un closed_sing)
+  thus "closed (insert x F)" by simp
+qed
+
+lemma finite_imp_compact:
+  fixes s :: "'a::heine_borel set"
+  shows "finite s ==> compact s"
+  unfolding compact_eq_bounded_closed
+  using finite_imp_closed finite_imp_bounded
+  by blast
+
+lemma compact_sing [simp]: "compact {a}"
+  unfolding compact_def o_def subseq_def
+  by (auto simp add: tendsto_const)
+
+lemma compact_cball[simp]:
+  fixes x :: "'a::heine_borel"
+  shows "compact(cball x e)"
+  using compact_eq_bounded_closed bounded_cball closed_cball
+  by blast
+
+lemma compact_frontier_bounded[intro]:
+  fixes s :: "'a::heine_borel set"
+  shows "bounded s ==> compact(frontier s)"
+  unfolding frontier_def
+  using compact_eq_bounded_closed
+  by blast
+
+lemma compact_frontier[intro]:
+  fixes s :: "'a::heine_borel set"
+  shows "compact s ==> compact (frontier s)"
+  using compact_eq_bounded_closed compact_frontier_bounded
+  by blast
+
+lemma frontier_subset_compact:
+  fixes s :: "'a::heine_borel set"
+  shows "compact s ==> frontier s \<subseteq> s"
+  using frontier_subset_closed compact_eq_bounded_closed
+  by blast
+
+lemma open_delete:
+  fixes s :: "'a::metric_space set"
+  shows "open s ==> open(s - {x})"
+  using open_Diff[of s "{x}"] closed_sing
+  by blast
+
+text{* Finite intersection property. I could make it an equivalence in fact. *}
+
+lemma compact_imp_fip:
+  fixes s :: "'a::heine_borel set"
+  assumes "compact s"  "\<forall>t \<in> f. closed t"
+        "\<forall>f'. finite f' \<and> f' \<subseteq> f --> (s \<inter> (\<Inter> f') \<noteq> {})"
+  shows "s \<inter> (\<Inter> f) \<noteq> {}"
+proof
+  assume as:"s \<inter> (\<Inter> f) = {}"
+  hence "s \<subseteq> \<Union>op - UNIV ` f" by auto
+  moreover have "Ball (op - UNIV ` f) open" using open_Diff closed_Diff using assms(2) by auto
+  ultimately obtain f' where f':"f' \<subseteq> op - UNIV ` f"  "finite f'"  "s \<subseteq> \<Union>f'" using assms(1)[unfolded compact_eq_heine_borel, THEN spec[where x="(\<lambda>t. UNIV - t) ` f"]] by auto
+  hence "finite (op - UNIV ` f') \<and> op - UNIV ` f' \<subseteq> f" by(auto simp add: Diff_Diff_Int)
+  hence "s \<inter> \<Inter>op - UNIV ` f' \<noteq> {}" using assms(3)[THEN spec[where x="op - UNIV ` f'"]] by auto
+  thus False using f'(3) unfolding subset_eq and Union_iff by blast
+qed
+
+subsection{* Bounded closed nest property (proof does not use Heine-Borel).            *}
+
+lemma bounded_closed_nest:
+  assumes "\<forall>n. closed(s n)" "\<forall>n. (s n \<noteq> {})"
+  "(\<forall>m n. m \<le> n --> s n \<subseteq> s m)"  "bounded(s 0)"
+  shows "\<exists>a::'a::heine_borel. \<forall>n::nat. a \<in> s(n)"
+proof-
+  from assms(2) obtain x where x:"\<forall>n::nat. x n \<in> s n" using choice[of "\<lambda>n x. x\<in> s n"] by auto
+  from assms(4,1) have *:"compact (s 0)" using bounded_closed_imp_compact[of "s 0"] by auto
+
+  then obtain l r where lr:"l\<in>s 0" "subseq r" "((x \<circ> r) ---> l) sequentially"
+    unfolding compact_def apply(erule_tac x=x in allE)  using x using assms(3) by blast
+
+  { fix n::nat
+    { fix e::real assume "e>0"
+      with lr(3) obtain N where N:"\<forall>m\<ge>N. dist ((x \<circ> r) m) l < e" unfolding Lim_sequentially by auto
+      hence "dist ((x \<circ> r) (max N n)) l < e" by auto
+      moreover
+      have "r (max N n) \<ge> n" using lr(2) using subseq_bigger[of r "max N n"] by auto
+      hence "(x \<circ> r) (max N n) \<in> s n"
+        using x apply(erule_tac x=n in allE)
+        using x apply(erule_tac x="r (max N n)" in allE)
+        using assms(3) apply(erule_tac x=n in allE)apply( erule_tac x="r (max N n)" in allE) by auto
+      ultimately have "\<exists>y\<in>s n. dist y l < e" by auto
+    }
+    hence "l \<in> s n" using closed_approachable[of "s n" l] assms(1) by blast
+  }
+  thus ?thesis by auto
+qed
+
+text{* Decreasing case does not even need compactness, just completeness.        *}
+
+lemma decreasing_closed_nest:
+  assumes "\<forall>n. closed(s n)"
+          "\<forall>n. (s n \<noteq> {})"
+          "\<forall>m n. m \<le> n --> s n \<subseteq> s m"
+          "\<forall>e>0. \<exists>n. \<forall>x \<in> (s n). \<forall> y \<in> (s n). dist x y < e"
+  shows "\<exists>a::'a::heine_borel. \<forall>n::nat. a \<in> s n"
+proof-
+  have "\<forall>n. \<exists> x. x\<in>s n" using assms(2) by auto
+  hence "\<exists>t. \<forall>n. t n \<in> s n" using choice[of "\<lambda> n x. x \<in> s n"] by auto
+  then obtain t where t: "\<forall>n. t n \<in> s n" by auto
+  { fix e::real assume "e>0"
+    then obtain N where N:"\<forall>x\<in>s N. \<forall>y\<in>s N. dist x y < e" using assms(4) by auto
+    { fix m n ::nat assume "N \<le> m \<and> N \<le> n"
+      hence "t m \<in> s N" "t n \<in> s N" using assms(3) t unfolding  subset_eq t by blast+
+      hence "dist (t m) (t n) < e" using N by auto
+    }
+    hence "\<exists>N. \<forall>m n. N \<le> m \<and> N \<le> n \<longrightarrow> dist (t m) (t n) < e" by auto
+  }
+  hence  "Cauchy t" unfolding cauchy_def by auto
+  then obtain l where l:"(t ---> l) sequentially" using complete_univ unfolding complete_def by auto
+  { fix n::nat
+    { fix e::real assume "e>0"
+      then obtain N::nat where N:"\<forall>n\<ge>N. dist (t n) l < e" using l[unfolded Lim_sequentially] by auto
+      have "t (max n N) \<in> s n" using assms(3) unfolding subset_eq apply(erule_tac x=n in allE) apply (erule_tac x="max n N" in allE) using t by auto
+      hence "\<exists>y\<in>s n. dist y l < e" apply(rule_tac x="t (max n N)" in bexI) using N by auto
+    }
+    hence "l \<in> s n" using closed_approachable[of "s n" l] assms(1) by auto
+  }
+  then show ?thesis by auto
+qed
+
+text{* Strengthen it to the intersection actually being a singleton.             *}
+
+lemma decreasing_closed_nest_sing:
+  assumes "\<forall>n. closed(s n)"
+          "\<forall>n. s n \<noteq> {}"
+          "\<forall>m n. m \<le> n --> s n \<subseteq> s m"
+          "\<forall>e>0. \<exists>n. \<forall>x \<in> (s n). \<forall> y\<in>(s n). dist x y < e"
+  shows "\<exists>a::'a::heine_borel. \<Inter> {t. (\<exists>n::nat. t = s n)} = {a}"
+proof-
+  obtain a where a:"\<forall>n. a \<in> s n" using decreasing_closed_nest[of s] using assms by auto
+  { fix b assume b:"b \<in> \<Inter>{t. \<exists>n. t = s n}"
+    { fix e::real assume "e>0"
+      hence "dist a b < e" using assms(4 )using b using a by blast
+    }
+    hence "dist a b = 0" by (metis dist_eq_0_iff dist_nz real_less_def)
+  }
+  with a have "\<Inter>{t. \<exists>n. t = s n} = {a}"  by auto
+  thus ?thesis by auto
+qed
+
+text{* Cauchy-type criteria for uniform convergence. *}
+
+lemma uniformly_convergent_eq_cauchy: fixes s::"nat \<Rightarrow> 'b \<Rightarrow> 'a::heine_borel" shows
+ "(\<exists>l. \<forall>e>0. \<exists>N. \<forall>n x. N \<le> n \<and> P x --> dist(s n x)(l x) < e) \<longleftrightarrow>
+  (\<forall>e>0. \<exists>N. \<forall>m n x. N \<le> m \<and> N \<le> n \<and> P x  --> dist (s m x) (s n x) < e)" (is "?lhs = ?rhs")
+proof(rule)
+  assume ?lhs
+  then obtain l where l:"\<forall>e>0. \<exists>N. \<forall>n x. N \<le> n \<and> P x \<longrightarrow> dist (s n x) (l x) < e" by auto
+  { fix e::real assume "e>0"
+    then obtain N::nat where N:"\<forall>n x. N \<le> n \<and> P x \<longrightarrow> dist (s n x) (l x) < e / 2" using l[THEN spec[where x="e/2"]] by auto
+    { fix n m::nat and x::"'b" assume "N \<le> m \<and> N \<le> n \<and> P x"
+      hence "dist (s m x) (s n x) < e"
+        using N[THEN spec[where x=m], THEN spec[where x=x]]
+        using N[THEN spec[where x=n], THEN spec[where x=x]]
+        using dist_triangle_half_l[of "s m x" "l x" e "s n x"] by auto  }
+    hence "\<exists>N. \<forall>m n x. N \<le> m \<and> N \<le> n \<and> P x  --> dist (s m x) (s n x) < e"  by auto  }
+  thus ?rhs by auto
+next
+  assume ?rhs
+  hence "\<forall>x. P x \<longrightarrow> Cauchy (\<lambda>n. s n x)" unfolding cauchy_def apply auto by (erule_tac x=e in allE)auto
+  then obtain l where l:"\<forall>x. P x \<longrightarrow> ((\<lambda>n. s n x) ---> l x) sequentially" unfolding convergent_eq_cauchy[THEN sym]
+    using choice[of "\<lambda>x l. P x \<longrightarrow> ((\<lambda>n. s n x) ---> l) sequentially"] by auto
+  { fix e::real assume "e>0"
+    then obtain N where N:"\<forall>m n x. N \<le> m \<and> N \<le> n \<and> P x \<longrightarrow> dist (s m x) (s n x) < e/2"
+      using `?rhs`[THEN spec[where x="e/2"]] by auto
+    { fix x assume "P x"
+      then obtain M where M:"\<forall>n\<ge>M. dist (s n x) (l x) < e/2"
+        using l[THEN spec[where x=x], unfolded Lim_sequentially] using `e>0` by(auto elim!: allE[where x="e/2"])
+      fix n::nat assume "n\<ge>N"
+      hence "dist(s n x)(l x) < e"  using `P x`and N[THEN spec[where x=n], THEN spec[where x="N+M"], THEN spec[where x=x]]
+        using M[THEN spec[where x="N+M"]] and dist_triangle_half_l[of "s n x" "s (N+M) x" e "l x"] by (auto simp add: dist_commute)  }
+    hence "\<exists>N. \<forall>n x. N \<le> n \<and> P x \<longrightarrow> dist(s n x)(l x) < e" by auto }
+  thus ?lhs by auto
+qed
+
+lemma uniformly_cauchy_imp_uniformly_convergent:
+  fixes s :: "nat \<Rightarrow> 'a \<Rightarrow> 'b::heine_borel"
+  assumes "\<forall>e>0.\<exists>N. \<forall>m (n::nat) x. N \<le> m \<and> N \<le> n \<and> P x --> dist(s m x)(s n x) < e"
+          "\<forall>x. P x --> (\<forall>e>0. \<exists>N. \<forall>n. N \<le> n --> dist(s n x)(l x) < e)"
+  shows "\<forall>e>0. \<exists>N. \<forall>n x. N \<le> n \<and> P x --> dist(s n x)(l x) < e"
+proof-
+  obtain l' where l:"\<forall>e>0. \<exists>N. \<forall>n x. N \<le> n \<and> P x \<longrightarrow> dist (s n x) (l' x) < e"
+    using assms(1) unfolding uniformly_convergent_eq_cauchy[THEN sym] by auto
+  moreover
+  { fix x assume "P x"
+    hence "l x = l' x" using Lim_unique[OF trivial_limit_sequentially, of "\<lambda>n. s n x" "l x" "l' x"]
+      using l and assms(2) unfolding Lim_sequentially by blast  }
+  ultimately show ?thesis by auto
+qed
+
+subsection{* Define continuity over a net to take in restrictions of the set. *}
+
+definition
+  continuous :: "'a::t2_space net \<Rightarrow> ('a \<Rightarrow> 'b::topological_space) \<Rightarrow> bool" where
+  "continuous net f \<longleftrightarrow> (f ---> f(netlimit net)) net"
+
+lemma continuous_trivial_limit:
+ "trivial_limit net ==> continuous net f"
+  unfolding continuous_def tendsto_def trivial_limit_eq by auto
+
+lemma continuous_within: "continuous (at x within s) f \<longleftrightarrow> (f ---> f(x)) (at x within s)"
+  unfolding continuous_def
+  unfolding tendsto_def
+  using netlimit_within[of x s]
+  by (cases "trivial_limit (at x within s)") (auto simp add: trivial_limit_eventually)
+
+lemma continuous_at: "continuous (at x) f \<longleftrightarrow> (f ---> f(x)) (at x)"
+  using continuous_within [of x UNIV f] by (simp add: within_UNIV)
+
+lemma continuous_at_within:
+  assumes "continuous (at x) f"  shows "continuous (at x within s) f"
+  using assms unfolding continuous_at continuous_within
+  by (rule Lim_at_within)
+
+text{* Derive the epsilon-delta forms, which we often use as "definitions" *}
+
+lemma continuous_within_eps_delta:
+  "continuous (at x within s) f \<longleftrightarrow> (\<forall>e>0. \<exists>d>0. \<forall>x'\<in> s.  dist x' x < d --> dist (f x') (f x) < e)"
+  unfolding continuous_within and Lim_within
+  apply auto unfolding dist_nz[THEN sym] apply(auto elim!:allE) apply(rule_tac x=d in exI) by auto
+
+lemma continuous_at_eps_delta: "continuous (at x) f \<longleftrightarrow>  (\<forall>e>0. \<exists>d>0.
+                           \<forall>x'. dist x' x < d --> dist(f x')(f x) < e)"
+  using continuous_within_eps_delta[of x UNIV f]
+  unfolding within_UNIV by blast
+
+text{* Versions in terms of open balls. *}
+
+lemma continuous_within_ball:
+ "continuous (at x within s) f \<longleftrightarrow> (\<forall>e>0. \<exists>d>0.
+                            f ` (ball x d \<inter> s) \<subseteq> ball (f x) e)" (is "?lhs = ?rhs")
+proof
+  assume ?lhs
+  { fix e::real assume "e>0"
+    then obtain d where d: "d>0" "\<forall>xa\<in>s. 0 < dist xa x \<and> dist xa x < d \<longrightarrow> dist (f xa) (f x) < e"
+      using `?lhs`[unfolded continuous_within Lim_within] by auto
+    { fix y assume "y\<in>f ` (ball x d \<inter> s)"
+      hence "y \<in> ball (f x) e" using d(2) unfolding dist_nz[THEN sym]
+        apply (auto simp add: dist_commute mem_ball) apply(erule_tac x=xa in ballE) apply auto using `e>0` by auto
+    }
+    hence "\<exists>d>0. f ` (ball x d \<inter> s) \<subseteq> ball (f x) e" using `d>0` unfolding subset_eq ball_def by (auto simp add: dist_commute)  }
+  thus ?rhs by auto
+next
+  assume ?rhs thus ?lhs unfolding continuous_within Lim_within ball_def subset_eq
+    apply (auto simp add: dist_commute) apply(erule_tac x=e in allE) by auto
+qed
+
+lemma continuous_at_ball:
+  "continuous (at x) f \<longleftrightarrow> (\<forall>e>0. \<exists>d>0. f ` (ball x d) \<subseteq> ball (f x) e)" (is "?lhs = ?rhs")
+proof
+  assume ?lhs thus ?rhs unfolding continuous_at Lim_at subset_eq Ball_def Bex_def image_iff mem_ball
+    apply auto apply(erule_tac x=e in allE) apply auto apply(rule_tac x=d in exI) apply auto apply(erule_tac x=xa in allE) apply (auto simp add: dist_commute dist_nz)
+    unfolding dist_nz[THEN sym] by auto
+next
+  assume ?rhs thus ?lhs unfolding continuous_at Lim_at subset_eq Ball_def Bex_def image_iff mem_ball
+    apply auto apply(erule_tac x=e in allE) apply auto apply(rule_tac x=d in exI) apply auto apply(erule_tac x="f xa" in allE) by (auto simp add: dist_commute dist_nz)
+qed
+
+text{* For setwise continuity, just start from the epsilon-delta definitions. *}
+
+definition
+  continuous_on :: "'a::metric_space set \<Rightarrow> ('a \<Rightarrow> 'b::metric_space) \<Rightarrow> bool" where
+  "continuous_on s f \<longleftrightarrow> (\<forall>x \<in> s. \<forall>e>0. \<exists>d::real>0. \<forall>x' \<in> s. dist x' x < d --> dist (f x') (f x) < e)"
+
+
+definition
+  uniformly_continuous_on ::
+    "'a::metric_space set \<Rightarrow> ('a \<Rightarrow> 'b::metric_space) \<Rightarrow> bool" where
+  "uniformly_continuous_on s f \<longleftrightarrow>
+        (\<forall>e>0. \<exists>d>0. \<forall>x\<in>s. \<forall> x'\<in>s. dist x' x < d
+                           --> dist (f x') (f x) < e)"
+
+text{* Some simple consequential lemmas. *}
+
+lemma uniformly_continuous_imp_continuous:
+ " uniformly_continuous_on s f ==> continuous_on s f"
+  unfolding uniformly_continuous_on_def continuous_on_def by blast
+
+lemma continuous_at_imp_continuous_within:
+ "continuous (at x) f ==> continuous (at x within s) f"
+  unfolding continuous_within continuous_at using Lim_at_within by auto
+
+lemma continuous_at_imp_continuous_on: assumes "(\<forall>x \<in> s. continuous (at x) f)"
+  shows "continuous_on s f"
+proof(simp add: continuous_at continuous_on_def, rule, rule, rule)
+  fix x and e::real assume "x\<in>s" "e>0"
+  hence "eventually (\<lambda>xa. dist (f xa) (f x) < e) (at x)" using assms unfolding continuous_at tendsto_iff by auto
+  then obtain d where d:"d>0" "\<forall>xa. 0 < dist xa x \<and> dist xa x < d \<longrightarrow> dist (f xa) (f x) < e" unfolding eventually_at by auto
+  { fix x' assume "\<not> 0 < dist x' x"
+    hence "x=x'"
+      using dist_nz[of x' x] by auto
+    hence "dist (f x') (f x) < e" using `e>0` by auto
+  }
+  thus "\<exists>d>0. \<forall>x'\<in>s. dist x' x < d \<longrightarrow> dist (f x') (f x) < e" using d by auto
+qed
+
+lemma continuous_on_eq_continuous_within:
+ "continuous_on s f \<longleftrightarrow> (\<forall>x \<in> s. continuous (at x within s) f)" (is "?lhs = ?rhs")
+proof
+  assume ?rhs
+  { fix x assume "x\<in>s"
+    fix e::real assume "e>0"
+    assume "\<exists>d>0. \<forall>xa\<in>s. 0 < dist xa x \<and> dist xa x < d \<longrightarrow> dist (f xa) (f x) < e"
+    then obtain d where "d>0" and d:"\<forall>xa\<in>s. 0 < dist xa x \<and> dist xa x < d \<longrightarrow> dist (f xa) (f x) < e" by auto
+    { fix x' assume as:"x'\<in>s" "dist x' x < d"
+      hence "dist (f x') (f x) < e" using `e>0` d `x'\<in>s` dist_eq_0_iff[of x' x] zero_le_dist[of x' x] as(2) by (metis dist_eq_0_iff dist_nz) }
+    hence "\<exists>d>0. \<forall>x'\<in>s. dist x' x < d \<longrightarrow> dist (f x') (f x) < e" using `d>0` by auto
+  }
+  thus ?lhs using `?rhs` unfolding continuous_on_def continuous_within Lim_within by auto
+next
+  assume ?lhs
+  thus ?rhs unfolding continuous_on_def continuous_within Lim_within by blast
+qed
+
+lemma continuous_on:
+ "continuous_on s f \<longleftrightarrow> (\<forall>x \<in> s. (f ---> f(x)) (at x within s))"
+  by (auto simp add: continuous_on_eq_continuous_within continuous_within)
+
+lemma continuous_on_eq_continuous_at:
+ "open s ==> (continuous_on s f \<longleftrightarrow> (\<forall>x \<in> s. continuous (at x) f))"
+  by (auto simp add: continuous_on continuous_at Lim_within_open)
+
+lemma continuous_within_subset:
+ "continuous (at x within s) f \<Longrightarrow> t \<subseteq> s
+             ==> continuous (at x within t) f"
+  unfolding continuous_within by(metis Lim_within_subset)
+
+lemma continuous_on_subset:
+ "continuous_on s f \<Longrightarrow> t \<subseteq> s ==> continuous_on t f"
+  unfolding continuous_on by (metis subset_eq Lim_within_subset)
+
+lemma continuous_on_interior:
+ "continuous_on s f \<Longrightarrow> x \<in> interior s ==> continuous (at x) f"
+unfolding interior_def
+apply simp
+by (meson continuous_on_eq_continuous_at continuous_on_subset)
+
+lemma continuous_on_eq:
+ "(\<forall>x \<in> s. f x = g x) \<Longrightarrow> continuous_on s f
+           ==> continuous_on s g"
+  by (simp add: continuous_on_def)
+
+text{* Characterization of various kinds of continuity in terms of sequences.  *}
+
+(* \<longrightarrow> could be generalized, but \<longleftarrow> requires metric space *)
+lemma continuous_within_sequentially:
+  fixes f :: "'a::metric_space \<Rightarrow> 'b::metric_space"
+  shows "continuous (at a within s) f \<longleftrightarrow>
+                (\<forall>x. (\<forall>n::nat. x n \<in> s) \<and> (x ---> a) sequentially
+                     --> ((f o x) ---> f a) sequentially)" (is "?lhs = ?rhs")
+proof
+  assume ?lhs
+  { fix x::"nat \<Rightarrow> 'a" assume x:"\<forall>n. x n \<in> s" "\<forall>e>0. \<exists>N. \<forall>n\<ge>N. dist (x n) a < e"
+    fix e::real assume "e>0"
+    from `?lhs` obtain d where "d>0" and d:"\<forall>x\<in>s. 0 < dist x a \<and> dist x a < d \<longrightarrow> dist (f x) (f a) < e" unfolding continuous_within Lim_within using `e>0` by auto
+    from x(2) `d>0` obtain N where N:"\<forall>n\<ge>N. dist (x n) a < d" by auto
+    hence "\<exists>N. \<forall>n\<ge>N. dist ((f \<circ> x) n) (f a) < e"
+      apply(rule_tac  x=N in exI) using N d  apply auto using x(1)
+      apply(erule_tac x=n in allE) apply(erule_tac x=n in allE)
+      apply(erule_tac x="x n" in ballE)  apply auto unfolding dist_nz[THEN sym] apply auto using `e>0` by auto
+  }
+  thus ?rhs unfolding continuous_within unfolding Lim_sequentially by simp
+next
+  assume ?rhs
+  { fix e::real assume "e>0"
+    assume "\<not> (\<exists>d>0. \<forall>x\<in>s. 0 < dist x a \<and> dist x a < d \<longrightarrow> dist (f x) (f a) < e)"
+    hence "\<forall>d. \<exists>x. d>0 \<longrightarrow> x\<in>s \<and> (0 < dist x a \<and> dist x a < d \<and> \<not> dist (f x) (f a) < e)" by blast
+    then obtain x where x:"\<forall>d>0. x d \<in> s \<and> (0 < dist (x d) a \<and> dist (x d) a < d \<and> \<not> dist (f (x d)) (f a) < e)"
+      using choice[of "\<lambda>d x.0<d \<longrightarrow> x\<in>s \<and> (0 < dist x a \<and> dist x a < d \<and> \<not> dist (f x) (f a) < e)"] by auto
+    { fix d::real assume "d>0"
+      hence "\<exists>N::nat. inverse (real (N + 1)) < d" using real_arch_inv[of d] by (auto, rule_tac x="n - 1" in exI)auto
+      then obtain N::nat where N:"inverse (real (N + 1)) < d" by auto
+      { fix n::nat assume n:"n\<ge>N"
+        hence "dist (x (inverse (real (n + 1)))) a < inverse (real (n + 1))" using x[THEN spec[where x="inverse (real (n + 1))"]] by auto
+        moreover have "inverse (real (n + 1)) < d" using N n by (auto, metis Suc_le_mono le_SucE less_imp_inverse_less nat_le_real_less order_less_trans real_of_nat_Suc real_of_nat_Suc_gt_zero)
+        ultimately have "dist (x (inverse (real (n + 1)))) a < d" by auto
+      }
+      hence "\<exists>N::nat. \<forall>n\<ge>N. dist (x (inverse (real (n + 1)))) a < d" by auto
+    }
+    hence "(\<forall>n::nat. x (inverse (real (n + 1))) \<in> s) \<and> (\<forall>e>0. \<exists>N::nat. \<forall>n\<ge>N. dist (x (inverse (real (n + 1)))) a < e)" using x by auto
+    hence "\<forall>e>0. \<exists>N::nat. \<forall>n\<ge>N. dist (f (x (inverse (real (n + 1))))) (f a) < e"  using `?rhs`[THEN spec[where x="\<lambda>n::nat. x (inverse (real (n+1)))"], unfolded Lim_sequentially] by auto
+    hence "False" apply(erule_tac x=e in allE) using `e>0` using x by auto
+  }
+  thus ?lhs  unfolding continuous_within unfolding Lim_within unfolding Lim_sequentially by blast
+qed
+
+lemma continuous_at_sequentially:
+  fixes f :: "'a::metric_space \<Rightarrow> 'b::metric_space"
+  shows "continuous (at a) f \<longleftrightarrow> (\<forall>x. (x ---> a) sequentially
+                  --> ((f o x) ---> f a) sequentially)"
+  using continuous_within_sequentially[of a UNIV f] unfolding within_UNIV by auto
+
+lemma continuous_on_sequentially:
+ "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)" (is "?lhs = ?rhs")
+proof
+  assume ?rhs thus ?lhs using continuous_within_sequentially[of _ s f] unfolding continuous_on_eq_continuous_within by auto
+next
+  assume ?lhs thus ?rhs unfolding continuous_on_eq_continuous_within using continuous_within_sequentially[of _ s f] by auto
+qed
+
+lemma uniformly_continuous_on_sequentially:
+  fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
+  shows "uniformly_continuous_on s f \<longleftrightarrow> (\<forall>x y. (\<forall>n. x n \<in> s) \<and> (\<forall>n. y n \<in> s) \<and>
+                    ((\<lambda>n. x n - y n) ---> 0) sequentially
+                    \<longrightarrow> ((\<lambda>n. f(x n) - f(y n)) ---> 0) sequentially)" (is "?lhs = ?rhs")
+proof
+  assume ?lhs
+  { fix x y assume x:"\<forall>n. x n \<in> s" and y:"\<forall>n. y n \<in> s" and xy:"((\<lambda>n. x n - y n) ---> 0) sequentially"
+    { fix e::real assume "e>0"
+      then obtain d where "d>0" and d:"\<forall>x\<in>s. \<forall>x'\<in>s. dist x' x < d \<longrightarrow> dist (f x') (f x) < e"
+        using `?lhs`[unfolded uniformly_continuous_on_def, THEN spec[where x=e]] by auto
+      obtain N where N:"\<forall>n\<ge>N. norm (x n - y n - 0) < d" using xy[unfolded Lim_sequentially dist_norm] and `d>0` by auto
+      { fix n assume "n\<ge>N"
+        hence "norm (f (x n) - f (y n) - 0) < e"
+          using N[THEN spec[where x=n]] using d[THEN bspec[where x="x n"], THEN bspec[where x="y n"]] using x and y
+          unfolding dist_commute and dist_norm by simp  }
+      hence "\<exists>N. \<forall>n\<ge>N. norm (f (x n) - f (y n) - 0) < e"  by auto  }
+    hence "((\<lambda>n. f(x n) - f(y n)) ---> 0) sequentially" unfolding Lim_sequentially and dist_norm by auto  }
+  thus ?rhs by auto
+next
+  assume ?rhs
+  { assume "\<not> ?lhs"
+    then obtain e where "e>0" "\<forall>d>0. \<exists>x\<in>s. \<exists>x'\<in>s. dist x' x < d \<and> \<not> dist (f x') (f x) < e" unfolding uniformly_continuous_on_def by auto
+    then obtain fa where fa:"\<forall>x.  0 < x \<longrightarrow> fst (fa x) \<in> s \<and> snd (fa x) \<in> s \<and> dist (fst (fa x)) (snd (fa x)) < x \<and> \<not> dist (f (fst (fa x))) (f (snd (fa x))) < e"
+      using choice[of "\<lambda>d x. d>0 \<longrightarrow> fst x \<in> s \<and> snd x \<in> s \<and> dist (snd x) (fst x) < d \<and> \<not> dist (f (snd x)) (f (fst x)) < e"] unfolding Bex_def
+      by (auto simp add: dist_commute)
+    def x \<equiv> "\<lambda>n::nat. fst (fa (inverse (real n + 1)))"
+    def y \<equiv> "\<lambda>n::nat. snd (fa (inverse (real n + 1)))"
+    have xyn:"\<forall>n. x n \<in> s \<and> y n \<in> s" and xy0:"\<forall>n. dist (x n) (y n) < inverse (real n + 1)" and fxy:"\<forall>n. \<not> dist (f (x n)) (f (y n)) < e"
+      unfolding x_def and y_def using fa by auto
+    have 1:"\<And>(x::'a) y. dist (x - y) 0 = dist x y" unfolding dist_norm by auto
+    have 2:"\<And>(x::'b) y. dist (x - y) 0 = dist x y" unfolding dist_norm by auto
+    { fix e::real assume "e>0"
+      then obtain N::nat where "N \<noteq> 0" and N:"0 < inverse (real N) \<and> inverse (real N) < e" unfolding real_arch_inv[of e]   by auto
+      { fix n::nat assume "n\<ge>N"
+        hence "inverse (real n + 1) < inverse (real N)" using real_of_nat_ge_zero and `N\<noteq>0` by auto
+        also have "\<dots> < e" using N by auto
+        finally have "inverse (real n + 1) < e" by auto
+        hence "dist (x n - y n) 0 < e" unfolding 1 using xy0[THEN spec[where x=n]] by auto  }
+      hence "\<exists>N. \<forall>n\<ge>N. dist (x n - y n) 0 < e" by auto  }
+    hence "\<forall>e>0. \<exists>N. \<forall>n\<ge>N. dist (f (x n) - f (y n)) 0 < e" using `?rhs`[THEN spec[where x=x], THEN spec[where x=y]] and xyn unfolding Lim_sequentially by auto
+    hence False unfolding 2 using fxy and `e>0` by auto  }
+  thus ?lhs unfolding uniformly_continuous_on_def by blast
+qed
+
+text{* The usual transformation theorems. *}
+
+lemma continuous_transform_within:
+  fixes f g :: "'a::metric_space \<Rightarrow> 'b::metric_space"
+  assumes "0 < d" "x \<in> s" "\<forall>x' \<in> s. dist x' x < d --> f x' = g x'"
+          "continuous (at x within s) f"
+  shows "continuous (at x within s) g"
+proof-
+  { fix e::real assume "e>0"
+    then obtain d' where d':"d'>0" "\<forall>xa\<in>s. 0 < dist xa x \<and> dist xa x < d' \<longrightarrow> dist (f xa) (f x) < e" using assms(4) unfolding continuous_within Lim_within by auto
+    { fix x' assume "x'\<in>s" "0 < dist x' x" "dist x' x < (min d d')"
+      hence "dist (f x') (g x) < e" using assms(2,3) apply(erule_tac x=x in ballE) using d' by auto  }
+    hence "\<forall>xa\<in>s. 0 < dist xa x \<and> dist xa x < (min d d') \<longrightarrow> dist (f xa) (g x) < e" by blast
+    hence "\<exists>d>0. \<forall>xa\<in>s. 0 < dist xa x \<and> dist xa x < d \<longrightarrow> dist (f xa) (g x) < e" using `d>0` `d'>0` by(rule_tac x="min d d'" in exI)auto  }
+  hence "(f ---> g x) (at x within s)" unfolding Lim_within using assms(1) by auto
+  thus ?thesis unfolding continuous_within using Lim_transform_within[of d s x f g "g x"] using assms by blast
+qed
+
+lemma continuous_transform_at:
+  fixes f g :: "'a::metric_space \<Rightarrow> 'b::metric_space"
+  assumes "0 < d" "\<forall>x'. dist x' x < d --> f x' = g x'"
+          "continuous (at x) f"
+  shows "continuous (at x) g"
+proof-
+  { fix e::real assume "e>0"
+    then obtain d' where d':"d'>0" "\<forall>xa. 0 < dist xa x \<and> dist xa x < d' \<longrightarrow> dist (f xa) (f x) < e" using assms(3) unfolding continuous_at Lim_at by auto
+    { fix x' assume "0 < dist x' x" "dist x' x < (min d d')"
+      hence "dist (f x') (g x) < e" using assms(2) apply(erule_tac x=x in allE) using d' by auto
+    }
+    hence "\<forall>xa. 0 < dist xa x \<and> dist xa x < (min d d') \<longrightarrow> dist (f xa) (g x) < e" by blast
+    hence "\<exists>d>0. \<forall>xa. 0 < dist xa x \<and> dist xa x < d \<longrightarrow> dist (f xa) (g x) < e" using `d>0` `d'>0` by(rule_tac x="min d d'" in exI)auto
+  }
+  hence "(f ---> g x) (at x)" unfolding Lim_at using assms(1) by auto
+  thus ?thesis unfolding continuous_at using Lim_transform_at[of d x f g "g x"] using assms by blast
+qed
+
+text{* Combination results for pointwise continuity. *}
+
+lemma continuous_const: "continuous net (\<lambda>x. c)"
+  by (auto simp add: continuous_def Lim_const)
+
+lemma continuous_cmul:
+  fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector"
+  shows "continuous net f ==> continuous net (\<lambda>x. c *\<^sub>R f x)"
+  by (auto simp add: continuous_def Lim_cmul)
+
+lemma continuous_neg:
+  fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector"
+  shows "continuous net f ==> continuous net (\<lambda>x. -(f x))"
+  by (auto simp add: continuous_def Lim_neg)
+
+lemma continuous_add:
+  fixes f g :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector"
+  shows "continuous net f \<Longrightarrow> continuous net g \<Longrightarrow> continuous net (\<lambda>x. f x + g x)"
+  by (auto simp add: continuous_def Lim_add)
+
+lemma continuous_sub:
+  fixes f g :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector"
+  shows "continuous net f \<Longrightarrow> continuous net g \<Longrightarrow> continuous net (\<lambda>x. f x - g x)"
+  by (auto simp add: continuous_def Lim_sub)
+
+text{* Same thing for setwise continuity. *}
+
+lemma continuous_on_const:
+ "continuous_on s (\<lambda>x. c)"
+  unfolding continuous_on_eq_continuous_within using continuous_const by blast
+
+lemma continuous_on_cmul:
+  fixes f :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector"
+  shows "continuous_on s f ==>  continuous_on s (\<lambda>x. c *\<^sub>R (f x))"
+  unfolding continuous_on_eq_continuous_within using continuous_cmul by blast
+
+lemma continuous_on_neg:
+  fixes f :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector"
+  shows "continuous_on s f \<Longrightarrow> continuous_on s (\<lambda>x. - f x)"
+  unfolding continuous_on_eq_continuous_within using continuous_neg by blast
+
+lemma continuous_on_add:
+  fixes f g :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector"
+  shows "continuous_on s f \<Longrightarrow> continuous_on s g
+           \<Longrightarrow> continuous_on s (\<lambda>x. f x + g x)"
+  unfolding continuous_on_eq_continuous_within using continuous_add by blast
+
+lemma continuous_on_sub:
+  fixes f g :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector"
+  shows "continuous_on s f \<Longrightarrow> continuous_on s g
+           \<Longrightarrow> continuous_on s (\<lambda>x. f x - g x)"
+  unfolding continuous_on_eq_continuous_within using continuous_sub by blast
+
+text{* Same thing for uniform continuity, using sequential formulations. *}
+
+lemma uniformly_continuous_on_const:
+ "uniformly_continuous_on s (\<lambda>x. c)"
+  unfolding uniformly_continuous_on_def by simp
+
+lemma uniformly_continuous_on_cmul:
+  fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
+    (* FIXME: generalize 'a to metric_space *)
+  assumes "uniformly_continuous_on s f"
+  shows "uniformly_continuous_on s (\<lambda>x. c *\<^sub>R f(x))"
+proof-
+  { fix x y assume "((\<lambda>n. f (x n) - f (y n)) ---> 0) sequentially"
+    hence "((\<lambda>n. c *\<^sub>R f (x n) - c *\<^sub>R f (y n)) ---> 0) sequentially"
+      using Lim_cmul[of "(\<lambda>n. f (x n) - f (y n))" 0 sequentially c]
+      unfolding scaleR_zero_right scaleR_right_diff_distrib by auto
+  }
+  thus ?thesis using assms unfolding uniformly_continuous_on_sequentially by auto
+qed
+
+lemma dist_minus:
+  fixes x y :: "'a::real_normed_vector"
+  shows "dist (- x) (- y) = dist x y"
+  unfolding dist_norm minus_diff_minus norm_minus_cancel ..
+
+lemma uniformly_continuous_on_neg:
+  fixes f :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector"
+  shows "uniformly_continuous_on s f
+         ==> uniformly_continuous_on s (\<lambda>x. -(f x))"
+  unfolding uniformly_continuous_on_def dist_minus .
+
+lemma uniformly_continuous_on_add:
+  fixes f g :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector" (* FIXME: generalize 'a *)
+  assumes "uniformly_continuous_on s f" "uniformly_continuous_on s g"
+  shows "uniformly_continuous_on s (\<lambda>x. f x + g x)"
+proof-
+  {  fix x y assume "((\<lambda>n. f (x n) - f (y n)) ---> 0) sequentially"
+                    "((\<lambda>n. g (x n) - g (y n)) ---> 0) sequentially"
+    hence "((\<lambda>xa. f (x xa) - f (y xa) + (g (x xa) - g (y xa))) ---> 0 + 0) sequentially"
+      using Lim_add[of "\<lambda> n. f (x n) - f (y n)" 0  sequentially "\<lambda> n. g (x n) - g (y n)" 0] by auto
+    hence "((\<lambda>n. f (x n) + g (x n) - (f (y n) + g (y n))) ---> 0) sequentially" unfolding Lim_sequentially and add_diff_add [symmetric] by auto  }
+  thus ?thesis using assms unfolding uniformly_continuous_on_sequentially by auto
+qed
+
+lemma uniformly_continuous_on_sub:
+  fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector" (* FIXME: generalize 'a *)
+  shows "uniformly_continuous_on s f \<Longrightarrow> uniformly_continuous_on s g
+           ==> uniformly_continuous_on s  (\<lambda>x. f x - g x)"
+  unfolding ab_diff_minus
+  using uniformly_continuous_on_add[of s f "\<lambda>x. - g x"]
+  using uniformly_continuous_on_neg[of s g] by auto
+
+text{* Identity function is continuous in every sense. *}
+
+lemma continuous_within_id:
+ "continuous (at a within s) (\<lambda>x. x)"
+  unfolding continuous_within by (rule Lim_at_within [OF Lim_ident_at])
+
+lemma continuous_at_id:
+ "continuous (at a) (\<lambda>x. x)"
+  unfolding continuous_at by (rule Lim_ident_at)
+
+lemma continuous_on_id:
+ "continuous_on s (\<lambda>x. x)"
+  unfolding continuous_on Lim_within by auto
+
+lemma uniformly_continuous_on_id:
+ "uniformly_continuous_on s (\<lambda>x. x)"
+  unfolding uniformly_continuous_on_def by auto
+
+text{* Continuity of all kinds is preserved under composition. *}
+
+lemma continuous_within_compose:
+  fixes f :: "'a::metric_space \<Rightarrow> 'b::metric_space" (* FIXME: generalize *)
+  fixes g :: "'b::metric_space \<Rightarrow> 'c::metric_space"
+  assumes "continuous (at x within s) f"   "continuous (at (f x) within f ` s) g"
+  shows "continuous (at x within s) (g o f)"
+proof-
+  { fix e::real assume "e>0"
+    with assms(2)[unfolded continuous_within Lim_within] obtain d  where "d>0" and d:"\<forall>xa\<in>f ` s. 0 < dist xa (f x) \<and> dist xa (f x) < d \<longrightarrow> dist (g xa) (g (f x)) < e" by auto
+    from assms(1)[unfolded continuous_within Lim_within] obtain d' where "d'>0" and d':"\<forall>xa\<in>s. 0 < dist xa x \<and> dist xa x < d' \<longrightarrow> dist (f xa) (f x) < d" using `d>0` by auto
+    { fix y assume as:"y\<in>s"  "0 < dist y x"  "dist y x < d'"
+      hence "dist (f y) (f x) < d" using d'[THEN bspec[where x=y]] by (auto simp add:dist_commute)
+      hence "dist (g (f y)) (g (f x)) < e" using as(1) d[THEN bspec[where x="f y"]] unfolding dist_nz[THEN sym] using `e>0` by auto   }
+    hence "\<exists>d>0. \<forall>xa\<in>s. 0 < dist xa x \<and> dist xa x < d \<longrightarrow> dist (g (f xa)) (g (f x)) < e" using `d'>0` by auto  }
+  thus ?thesis unfolding continuous_within Lim_within by auto
+qed
+
+lemma continuous_at_compose:
+  fixes f :: "'a::metric_space \<Rightarrow> 'b::metric_space" (* FIXME: generalize *)
+  fixes g :: "'b::metric_space \<Rightarrow> 'c::metric_space"
+  assumes "continuous (at x) f"  "continuous (at (f x)) g"
+  shows "continuous (at x) (g o f)"
+proof-
+  have " continuous (at (f x) within range f) g" using assms(2) using continuous_within_subset[of "f x" UNIV g "range f", unfolded within_UNIV] by auto
+  thus ?thesis using assms(1) using continuous_within_compose[of x UNIV f g, unfolded within_UNIV] by auto
+qed
+
+lemma continuous_on_compose:
+ "continuous_on s f \<Longrightarrow> continuous_on (f ` s) g \<Longrightarrow> continuous_on s (g o f)"
+  unfolding continuous_on_eq_continuous_within using continuous_within_compose[of _ s f g] by auto
+
+lemma uniformly_continuous_on_compose:
+  assumes "uniformly_continuous_on s f"  "uniformly_continuous_on (f ` s) g"
+  shows "uniformly_continuous_on s (g o f)"
+proof-
+  { fix e::real assume "e>0"
+    then obtain d where "d>0" and d:"\<forall>x\<in>f ` s. \<forall>x'\<in>f ` s. dist x' x < d \<longrightarrow> dist (g x') (g x) < e" using assms(2) unfolding uniformly_continuous_on_def by auto
+    obtain d' where "d'>0" "\<forall>x\<in>s. \<forall>x'\<in>s. dist x' x < d' \<longrightarrow> dist (f x') (f x) < d" using `d>0` using assms(1) unfolding uniformly_continuous_on_def by auto
+    hence "\<exists>d>0. \<forall>x\<in>s. \<forall>x'\<in>s. dist x' x < d \<longrightarrow> dist ((g \<circ> f) x') ((g \<circ> f) x) < e" using `d>0` using d by auto  }
+  thus ?thesis using assms unfolding uniformly_continuous_on_def by auto
+qed
+
+text{* Continuity in terms of open preimages. *}
+
+lemma continuous_at_open:
+  fixes f :: "'a::metric_space \<Rightarrow> 'b::metric_space" (* FIXME: generalize *)
+  shows "continuous (at x) f \<longleftrightarrow> (\<forall>t. open t \<and> f x \<in> t --> (\<exists>s. open s \<and> x \<in> s \<and> (\<forall>x' \<in> s. (f x') \<in> t)))" (is "?lhs = ?rhs")
+proof
+  assume ?lhs
+  { fix t assume as: "open t" "f x \<in> t"
+    then obtain e where "e>0" and e:"ball (f x) e \<subseteq> t" unfolding open_contains_ball by auto
+
+    obtain d where "d>0" and d:"\<forall>y. 0 < dist y x \<and> dist y x < d \<longrightarrow> dist (f y) (f x) < e" using `e>0` using `?lhs`[unfolded continuous_at Lim_at open_dist] by auto
+
+    have "open (ball x d)" using open_ball by auto
+    moreover have "x \<in> ball x d" unfolding centre_in_ball using `d>0` by simp
+    moreover
+    { fix x' assume "x'\<in>ball x d" hence "f x' \<in> t"
+        using e[unfolded subset_eq Ball_def mem_ball, THEN spec[where x="f x'"]]    d[THEN spec[where x=x']]
+        unfolding mem_ball apply (auto simp add: dist_commute)
+        unfolding dist_nz[THEN sym] using as(2) by auto  }
+    hence "\<forall>x'\<in>ball x d. f x' \<in> t" by auto
+    ultimately have "\<exists>s. open s \<and> x \<in> s \<and> (\<forall>x'\<in>s. f x' \<in> t)"
+      apply(rule_tac x="ball x d" in exI) by simp  }
+  thus ?rhs by auto
+next
+  assume ?rhs
+  { fix e::real assume "e>0"
+    then obtain s where s: "open s"  "x \<in> s"  "\<forall>x'\<in>s. f x' \<in> ball (f x) e" using `?rhs`[unfolded continuous_at Lim_at, THEN spec[where x="ball (f x) e"]]
+      unfolding centre_in_ball[of "f x" e, THEN sym] by auto
+    then obtain d where "d>0" and d:"ball x d \<subseteq> s" unfolding open_contains_ball by auto
+    { fix y assume "0 < dist y x \<and> dist y x < d"
+      hence "dist (f y) (f x) < e" using d[unfolded subset_eq Ball_def mem_ball, THEN spec[where x=y]]
+        using s(3)[THEN bspec[where x=y], unfolded mem_ball] by (auto simp add: dist_commute)  }
+    hence "\<exists>d>0. \<forall>xa. 0 < dist xa x \<and> dist xa x < d \<longrightarrow> dist (f xa) (f x) < e" using `d>0` by auto  }
+  thus ?lhs unfolding continuous_at Lim_at by auto
+qed
+
+lemma continuous_on_open:
+ "continuous_on s f \<longleftrightarrow>
+        (\<forall>t. openin (subtopology euclidean (f ` s)) t
+            --> openin (subtopology euclidean s) {x \<in> s. f x \<in> t})" (is "?lhs = ?rhs")
+proof
+  assume ?lhs
+  { fix t assume as:"openin (subtopology euclidean (f ` s)) t"
+    have "{x \<in> s. f x \<in> t} \<subseteq> s" using as[unfolded openin_euclidean_subtopology_iff] by auto
+    moreover
+    { fix x assume as':"x\<in>{x \<in> s. f x \<in> t}"
+      then obtain e where e: "e>0" "\<forall>x'\<in>f ` s. dist x' (f x) < e \<longrightarrow> x' \<in> t" using as[unfolded openin_euclidean_subtopology_iff, THEN conjunct2, THEN bspec[where x="f x"]] by auto
+      from this(1) obtain d where d: "d>0" "\<forall>xa\<in>s. 0 < dist xa x \<and> dist xa x < d \<longrightarrow> dist (f xa) (f x) < e" using `?lhs`[unfolded continuous_on Lim_within, THEN bspec[where x=x]] using as' by auto
+      have "\<exists>e>0. \<forall>x'\<in>s. dist x' x < e \<longrightarrow> x' \<in> {x \<in> s. f x \<in> t}" using d e unfolding dist_nz[THEN sym] by (rule_tac x=d in exI, auto)  }
+    ultimately have "openin (subtopology euclidean s) {x \<in> s. f x \<in> t}" unfolding openin_euclidean_subtopology_iff by auto  }
+  thus ?rhs unfolding continuous_on Lim_within using openin by auto
+next
+  assume ?rhs
+  { fix e::real and x assume "x\<in>s" "e>0"
+    { fix xa x' assume "dist (f xa) (f x) < e" "xa \<in> s" "x' \<in> s" "dist (f xa) (f x') < e - dist (f xa) (f x)"
+      hence "dist (f x') (f x) < e" using dist_triangle[of "f x'" "f x" "f xa"]
+        by (auto simp add: dist_commute)  }
+    hence "ball (f x) e \<inter> f ` s \<subseteq> f ` s \<and> (\<forall>xa\<in>ball (f x) e \<inter> f ` s. \<exists>ea>0. \<forall>x'\<in>f ` s. dist x' xa < ea \<longrightarrow> x' \<in> ball (f x) e \<inter> f ` s)" apply auto
+      apply(rule_tac x="e - dist (f xa) (f x)" in exI) using `e>0` by (auto simp add: dist_commute)
+    hence "\<forall>xa\<in>{xa \<in> s. f xa \<in> ball (f x) e \<inter> f ` s}. \<exists>ea>0. \<forall>x'\<in>s. dist x' xa < ea \<longrightarrow> x' \<in> {xa \<in> s. f xa \<in> ball (f x) e \<inter> f ` s}"
+      using `?rhs`[unfolded openin_euclidean_subtopology_iff, THEN spec[where x="ball (f x) e \<inter> f ` s"]] by auto
+    hence "\<exists>d>0. \<forall>xa\<in>s. 0 < dist xa x \<and> dist xa x < d \<longrightarrow> dist (f xa) (f x) < e" apply(erule_tac x=x in ballE) apply auto using `e>0` `x\<in>s` by (auto simp add: dist_commute)  }
+  thus ?lhs unfolding continuous_on Lim_within by auto
+qed
+
+(* ------------------------------------------------------------------------- *)
+(* Similarly in terms of closed sets.                                        *)
+(* ------------------------------------------------------------------------- *)
+
+lemma continuous_on_closed:
+ "continuous_on s f \<longleftrightarrow>  (\<forall>t. closedin (subtopology euclidean (f ` s)) t  --> closedin (subtopology euclidean s) {x \<in> s. f x \<in> t})" (is "?lhs = ?rhs")
+proof
+  assume ?lhs
+  { fix t
+    have *:"s - {x \<in> s. f x \<in> f ` s - t} = {x \<in> s. f x \<in> t}" by auto
+    have **:"f ` s - (f ` s - (f ` s - t)) = f ` s - t" by auto
+    assume as:"closedin (subtopology euclidean (f ` s)) t"
+    hence "closedin (subtopology euclidean (f ` s)) (f ` s - (f ` s - t))" unfolding closedin_def topspace_euclidean_subtopology unfolding ** by auto
+    hence "closedin (subtopology euclidean s) {x \<in> s. f x \<in> t}" using `?lhs`[unfolded continuous_on_open, THEN spec[where x="(f ` s) - t"]]
+      unfolding openin_closedin_eq topspace_euclidean_subtopology unfolding * by auto  }
+  thus ?rhs by auto
+next
+  assume ?rhs
+  { fix t
+    have *:"s - {x \<in> s. f x \<in> f ` s - t} = {x \<in> s. f x \<in> t}" by auto
+    assume as:"openin (subtopology euclidean (f ` s)) t"
+    hence "openin (subtopology euclidean s) {x \<in> s. f x \<in> t}" using `?rhs`[THEN spec[where x="(f ` s) - t"]]
+      unfolding openin_closedin_eq topspace_euclidean_subtopology *[THEN sym] closedin_subtopology by auto }
+  thus ?lhs unfolding continuous_on_open by auto
+qed
+
+text{* Half-global and completely global cases.                                  *}
+
+lemma continuous_open_in_preimage:
+  assumes "continuous_on s f"  "open t"
+  shows "openin (subtopology euclidean s) {x \<in> s. f x \<in> t}"
+proof-
+  have *:"\<forall>x. x \<in> s \<and> f x \<in> t \<longleftrightarrow> x \<in> s \<and> f x \<in> (t \<inter> f ` s)" by auto
+  have "openin (subtopology euclidean (f ` s)) (t \<inter> f ` s)"
+    using openin_open_Int[of t "f ` s", OF assms(2)] unfolding openin_open by auto
+  thus ?thesis using assms(1)[unfolded continuous_on_open, THEN spec[where x="t \<inter> f ` s"]] using * by auto
+qed
+
+lemma continuous_closed_in_preimage:
+  assumes "continuous_on s f"  "closed t"
+  shows "closedin (subtopology euclidean s) {x \<in> s. f x \<in> t}"
+proof-
+  have *:"\<forall>x. x \<in> s \<and> f x \<in> t \<longleftrightarrow> x \<in> s \<and> f x \<in> (t \<inter> f ` s)" by auto
+  have "closedin (subtopology euclidean (f ` s)) (t \<inter> f ` s)"
+    using closedin_closed_Int[of t "f ` s", OF assms(2)] unfolding Int_commute by auto
+  thus ?thesis
+    using assms(1)[unfolded continuous_on_closed, THEN spec[where x="t \<inter> f ` s"]] using * by auto
+qed
+
+lemma continuous_open_preimage:
+  assumes "continuous_on s f" "open s" "open t"
+  shows "open {x \<in> s. f x \<in> t}"
+proof-
+  obtain T where T: "open T" "{x \<in> s. f x \<in> t} = s \<inter> T"
+    using continuous_open_in_preimage[OF assms(1,3)] unfolding openin_open by auto
+  thus ?thesis using open_Int[of s T, OF assms(2)] by auto
+qed
+
+lemma continuous_closed_preimage:
+  assumes "continuous_on s f" "closed s" "closed t"
+  shows "closed {x \<in> s. f x \<in> t}"
+proof-
+  obtain T where T: "closed T" "{x \<in> s. f x \<in> t} = s \<inter> T"
+    using continuous_closed_in_preimage[OF assms(1,3)] unfolding closedin_closed by auto
+  thus ?thesis using closed_Int[of s T, OF assms(2)] by auto
+qed
+
+lemma continuous_open_preimage_univ:
+  fixes f :: "'a::metric_space \<Rightarrow> 'b::metric_space" (* FIXME: generalize *)
+  shows "\<forall>x. continuous (at x) f \<Longrightarrow> open s \<Longrightarrow> open {x. f x \<in> s}"
+  using continuous_open_preimage[of UNIV f s] open_UNIV continuous_at_imp_continuous_on by auto
+
+lemma continuous_closed_preimage_univ:
+  fixes f :: "'a::metric_space \<Rightarrow> 'b::metric_space" (* FIXME: generalize *)
+  shows "(\<forall>x. continuous (at x) f) \<Longrightarrow> closed s ==> closed {x. f x \<in> s}"
+  using continuous_closed_preimage[of UNIV f s] closed_UNIV continuous_at_imp_continuous_on by auto
+
+lemma continuous_open_vimage:
+  fixes f :: "'a::metric_space \<Rightarrow> 'b::metric_space" (* FIXME: generalize *)
+  shows "\<forall>x. continuous (at x) f \<Longrightarrow> open s \<Longrightarrow> open (f -` s)"
+  unfolding vimage_def by (rule continuous_open_preimage_univ)
+
+lemma continuous_closed_vimage:
+  fixes f :: "'a::metric_space \<Rightarrow> 'b::metric_space" (* FIXME: generalize *)
+  shows "\<forall>x. continuous (at x) f \<Longrightarrow> closed s \<Longrightarrow> closed (f -` s)"
+  unfolding vimage_def by (rule continuous_closed_preimage_univ)
+
+text{* Equality of continuous functions on closure and related results.          *}
+
+lemma continuous_closed_in_preimage_constant:
+ "continuous_on s f ==> closedin (subtopology euclidean s) {x \<in> s. f x = a}"
+  using continuous_closed_in_preimage[of s f "{a}"] closed_sing by auto
+
+lemma continuous_closed_preimage_constant:
+ "continuous_on s f \<Longrightarrow> closed s ==> closed {x \<in> s. f x = a}"
+  using continuous_closed_preimage[of s f "{a}"] closed_sing by auto
+
+lemma continuous_constant_on_closure:
+  assumes "continuous_on (closure s) f"
+          "\<forall>x \<in> s. f x = a"
+  shows "\<forall>x \<in> (closure s). f x = a"
+    using continuous_closed_preimage_constant[of "closure s" f a]
+    assms closure_minimal[of s "{x \<in> closure s. f x = a}"] closure_subset unfolding subset_eq by auto
+
+lemma image_closure_subset:
+  assumes "continuous_on (closure s) f"  "closed t"  "(f ` s) \<subseteq> t"
+  shows "f ` (closure s) \<subseteq> t"
+proof-
+  have "s \<subseteq> {x \<in> closure s. f x \<in> t}" using assms(3) closure_subset by auto
+  moreover have "closed {x \<in> closure s. f x \<in> t}"
+    using continuous_closed_preimage[OF assms(1)] and assms(2) by auto
+  ultimately have "closure s = {x \<in> closure s . f x \<in> t}"
+    using closure_minimal[of s "{x \<in> closure s. f x \<in> t}"] by auto
+  thus ?thesis by auto
+qed
+
+lemma continuous_on_closure_norm_le:
+  fixes f :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector"
+  assumes "continuous_on (closure s) f"  "\<forall>y \<in> s. norm(f y) \<le> b"  "x \<in> (closure s)"
+  shows "norm(f x) \<le> b"
+proof-
+  have *:"f ` s \<subseteq> cball 0 b" using assms(2)[unfolded mem_cball_0[THEN sym]] by auto
+  show ?thesis
+    using image_closure_subset[OF assms(1) closed_cball[of 0 b] *] assms(3)
+    unfolding subset_eq apply(erule_tac x="f x" in ballE) by (auto simp add: dist_norm)
+qed
+
+text{* Making a continuous function avoid some value in a neighbourhood.         *}
+
+lemma continuous_within_avoid:
+  fixes f :: "'a::metric_space \<Rightarrow> 'b::metric_space" (* FIXME: generalize *)
+  assumes "continuous (at x within s) f"  "x \<in> s"  "f x \<noteq> a"
+  shows "\<exists>e>0. \<forall>y \<in> s. dist x y < e --> f y \<noteq> a"
+proof-
+  obtain d where "d>0" and d:"\<forall>xa\<in>s. 0 < dist xa x \<and> dist xa x < d \<longrightarrow> dist (f xa) (f x) < dist (f x) a"
+    using assms(1)[unfolded continuous_within Lim_within, THEN spec[where x="dist (f x) a"]] assms(3)[unfolded dist_nz] by auto
+  { fix y assume " y\<in>s"  "dist x y < d"
+    hence "f y \<noteq> a" using d[THEN bspec[where x=y]] assms(3)[unfolded dist_nz]
+      apply auto unfolding dist_nz[THEN sym] by (auto simp add: dist_commute) }
+  thus ?thesis using `d>0` by auto
+qed
+
+lemma continuous_at_avoid:
+  fixes f :: "'a::metric_space \<Rightarrow> 'b::metric_space" (* FIXME: generalize *)
+  assumes "continuous (at x) f"  "f x \<noteq> a"
+  shows "\<exists>e>0. \<forall>y. dist x y < e \<longrightarrow> f y \<noteq> a"
+using assms using continuous_within_avoid[of x UNIV f a, unfolded within_UNIV] by auto
+
+lemma continuous_on_avoid:
+  assumes "continuous_on s f"  "x \<in> s"  "f x \<noteq> a"
+  shows "\<exists>e>0. \<forall>y \<in> s. dist x y < e \<longrightarrow> f y \<noteq> a"
+using assms(1)[unfolded continuous_on_eq_continuous_within, THEN bspec[where x=x], OF assms(2)]  continuous_within_avoid[of x s f a]  assms(2,3) by auto
+
+lemma continuous_on_open_avoid:
+  assumes "continuous_on s f"  "open s"  "x \<in> s"  "f x \<noteq> a"
+  shows "\<exists>e>0. \<forall>y. dist x y < e \<longrightarrow> f y \<noteq> a"
+using assms(1)[unfolded continuous_on_eq_continuous_at[OF assms(2)], THEN bspec[where x=x], OF assms(3)]  continuous_at_avoid[of x f a]  assms(3,4) by auto
+
+text{* Proving a function is constant by proving open-ness of level set.         *}
+
+lemma continuous_levelset_open_in_cases:
+ "connected s \<Longrightarrow> continuous_on s f \<Longrightarrow>
+        openin (subtopology euclidean s) {x \<in> s. f x = a}
+        ==> (\<forall>x \<in> s. f x \<noteq> a) \<or> (\<forall>x \<in> s. f x = a)"
+unfolding connected_clopen using continuous_closed_in_preimage_constant by auto
+
+lemma continuous_levelset_open_in:
+ "connected s \<Longrightarrow> continuous_on s f \<Longrightarrow>
+        openin (subtopology euclidean s) {x \<in> s. f x = a} \<Longrightarrow>
+        (\<exists>x \<in> s. f x = a)  ==> (\<forall>x \<in> s. f x = a)"
+using continuous_levelset_open_in_cases[of s f ]
+by meson
+
+lemma continuous_levelset_open:
+  assumes "connected s"  "continuous_on s f"  "open {x \<in> s. f x = a}"  "\<exists>x \<in> s.  f x = a"
+  shows "\<forall>x \<in> s. f x = a"
+using continuous_levelset_open_in[OF assms(1,2), of a, unfolded openin_open] using assms (3,4) by auto
+
+text{* Some arithmetical combinations (more to prove).                           *}
+
+lemma open_scaling[intro]:
+  fixes s :: "'a::real_normed_vector set"
+  assumes "c \<noteq> 0"  "open s"
+  shows "open((\<lambda>x. c *\<^sub>R x) ` s)"
+proof-
+  { fix x assume "x \<in> s"
+    then obtain e where "e>0" and e:"\<forall>x'. dist x' x < e \<longrightarrow> x' \<in> s" using assms(2)[unfolded open_dist, THEN bspec[where x=x]] by auto
+    have "e * abs c > 0" using assms(1)[unfolded zero_less_abs_iff[THEN sym]] using real_mult_order[OF `e>0`] by auto
+    moreover
+    { fix y assume "dist y (c *\<^sub>R x) < e * \<bar>c\<bar>"
+      hence "norm ((1 / c) *\<^sub>R y - x) < e" unfolding dist_norm
+        using norm_scaleR[of c "(1 / c) *\<^sub>R y - x", unfolded scaleR_right_diff_distrib, unfolded scaleR_scaleR] assms(1)
+          assms(1)[unfolded zero_less_abs_iff[THEN sym]] by (simp del:zero_less_abs_iff)
+      hence "y \<in> op *\<^sub>R c ` s" using rev_image_eqI[of "(1 / c) *\<^sub>R y" s y "op *\<^sub>R c"]  e[THEN spec[where x="(1 / c) *\<^sub>R y"]]  assms(1) unfolding dist_norm scaleR_scaleR by auto  }
+    ultimately have "\<exists>e>0. \<forall>x'. dist x' (c *\<^sub>R x) < e \<longrightarrow> x' \<in> op *\<^sub>R c ` s" apply(rule_tac x="e * abs c" in exI) by auto  }
+  thus ?thesis unfolding open_dist by auto
+qed
+
+lemma minus_image_eq_vimage:
+  fixes A :: "'a::ab_group_add set"
+  shows "(\<lambda>x. - x) ` A = (\<lambda>x. - x) -` A"
+  by (auto intro!: image_eqI [where f="\<lambda>x. - x"])
+
+lemma open_negations:
+  fixes s :: "'a::real_normed_vector set"
+  shows "open s ==> open ((\<lambda> x. -x) ` s)"
+  unfolding scaleR_minus1_left [symmetric]
+  by (rule open_scaling, auto)
+
+lemma open_translation:
+  fixes s :: "'a::real_normed_vector set"
+  assumes "open s"  shows "open((\<lambda>x. a + x) ` s)"
+proof-
+  { fix x have "continuous (at x) (\<lambda>x. x - a)" using continuous_sub[of "at x" "\<lambda>x. x" "\<lambda>x. a"] continuous_at_id[of x] continuous_const[of "at x" a] by auto  }
+  moreover have "{x. x - a \<in> s}  = op + a ` s" apply auto unfolding image_iff apply(rule_tac x="x - a" in bexI) by auto
+  ultimately show ?thesis using continuous_open_preimage_univ[of "\<lambda>x. x - a" s] using assms by auto
+qed
+
+lemma open_affinity:
+  fixes s :: "'a::real_normed_vector set"
+  assumes "open s"  "c \<noteq> 0"
+  shows "open ((\<lambda>x. a + c *\<^sub>R x) ` s)"
+proof-
+  have *:"(\<lambda>x. a + c *\<^sub>R x) = (\<lambda>x. a + x) \<circ> (\<lambda>x. c *\<^sub>R x)" unfolding o_def ..
+  have "op + a ` op *\<^sub>R c ` s = (op + a \<circ> op *\<^sub>R c) ` s" by auto
+  thus ?thesis using assms open_translation[of "op *\<^sub>R c ` s" a] unfolding * by auto
+qed
+
+lemma interior_translation:
+  fixes s :: "'a::real_normed_vector set"
+  shows "interior ((\<lambda>x. a + x) ` s) = (\<lambda>x. a + x) ` (interior s)"
+proof (rule set_ext, rule)
+  fix x assume "x \<in> interior (op + a ` s)"
+  then obtain e where "e>0" and e:"ball x e \<subseteq> op + a ` s" unfolding mem_interior by auto
+  hence "ball (x - a) e \<subseteq> s" unfolding subset_eq Ball_def mem_ball dist_norm apply auto apply(erule_tac x="a + xa" in allE) unfolding ab_group_add_class.diff_diff_eq[THEN sym] by auto
+  thus "x \<in> op + a ` interior s" unfolding image_iff apply(rule_tac x="x - a" in bexI) unfolding mem_interior using `e > 0` by auto
+next
+  fix x assume "x \<in> op + a ` interior s"
+  then obtain y e where "e>0" and e:"ball y e \<subseteq> s" and y:"x = a + y" unfolding image_iff Bex_def mem_interior by auto
+  { fix z have *:"a + y - z = y + a - z" by auto
+    assume "z\<in>ball x e"
+    hence "z - a \<in> s" using e[unfolded subset_eq, THEN bspec[where x="z - a"]] unfolding mem_ball dist_norm y ab_group_add_class.diff_diff_eq2 * by auto
+    hence "z \<in> op + a ` s" unfolding image_iff by(auto intro!: bexI[where x="z - a"])  }
+  hence "ball x e \<subseteq> op + a ` s" unfolding subset_eq by auto
+  thus "x \<in> interior (op + a ` s)" unfolding mem_interior using `e>0` by auto
+qed
+
+subsection {* Preservation of compactness and connectedness under continuous function.  *}
+
+lemma compact_continuous_image:
+  assumes "continuous_on s f"  "compact s"
+  shows "compact(f ` s)"
+proof-
+  { fix x assume x:"\<forall>n::nat. x n \<in> f ` s"
+    then obtain y where y:"\<forall>n. y n \<in> s \<and> x n = f (y n)" unfolding image_iff Bex_def using choice[of "\<lambda>n xa. xa \<in> s \<and> x n = f xa"] by auto
+    then obtain l r where "l\<in>s" and r:"subseq r" and lr:"((y \<circ> r) ---> l) sequentially" using assms(2)[unfolded compact_def, THEN spec[where x=y]] by auto
+    { fix e::real assume "e>0"
+      then obtain d where "d>0" and d:"\<forall>x'\<in>s. dist x' l < d \<longrightarrow> dist (f x') (f l) < e" using assms(1)[unfolded continuous_on_def, THEN bspec[where x=l], OF `l\<in>s`] by auto
+      then obtain N::nat where N:"\<forall>n\<ge>N. dist ((y \<circ> r) n) l < d" using lr[unfolded Lim_sequentially, THEN spec[where x=d]] by auto
+      { fix n::nat assume "n\<ge>N" hence "dist ((x \<circ> r) n) (f l) < e" using N[THEN spec[where x=n]] d[THEN bspec[where x="y (r n)"]] y[THEN spec[where x="r n"]] by auto  }
+      hence "\<exists>N. \<forall>n\<ge>N. dist ((x \<circ> r) n) (f l) < e" by auto  }
+    hence "\<exists>l\<in>f ` s. \<exists>r. subseq r \<and> ((x \<circ> r) ---> l) sequentially" unfolding Lim_sequentially using r lr `l\<in>s` by auto  }
+  thus ?thesis unfolding compact_def by auto
+qed
+
+lemma connected_continuous_image:
+  assumes "continuous_on s f"  "connected s"
+  shows "connected(f ` s)"
+proof-
+  { fix T assume as: "T \<noteq> {}"  "T \<noteq> f ` s"  "openin (subtopology euclidean (f ` s)) T"  "closedin (subtopology euclidean (f ` s)) T"
+    have "{x \<in> s. f x \<in> T} = {} \<or> {x \<in> s. f x \<in> T} = s"
+      using assms(1)[unfolded continuous_on_open, THEN spec[where x=T]]
+      using assms(1)[unfolded continuous_on_closed, THEN spec[where x=T]]
+      using assms(2)[unfolded connected_clopen, THEN spec[where x="{x \<in> s. f x \<in> T}"]] as(3,4) by auto
+    hence False using as(1,2)
+      using as(4)[unfolded closedin_def topspace_euclidean_subtopology] by auto }
+  thus ?thesis unfolding connected_clopen by auto
+qed
+
+text{* Continuity implies uniform continuity on a compact domain.                *}
+
+lemma compact_uniformly_continuous:
+  assumes "continuous_on s f"  "compact s"
+  shows "uniformly_continuous_on s f"
+proof-
+    { fix x assume x:"x\<in>s"
+      hence "\<forall>xa. \<exists>y. 0 < xa \<longrightarrow> (y > 0 \<and> (\<forall>x'\<in>s. dist x' x < y \<longrightarrow> dist (f x') (f x) < xa))" using assms(1)[unfolded continuous_on_def, THEN bspec[where x=x]] by auto
+      hence "\<exists>fa. \<forall>xa>0. \<forall>x'\<in>s. fa xa > 0 \<and> (dist x' x < fa xa \<longrightarrow> dist (f x') (f x) < xa)" using choice[of "\<lambda>e d. e>0 \<longrightarrow> d>0 \<and>(\<forall>x'\<in>s. (dist x' x < d \<longrightarrow> dist (f x') (f x) < e))"] by auto  }
+    then have "\<forall>x\<in>s. \<exists>y. \<forall>xa. 0 < xa \<longrightarrow> (\<forall>x'\<in>s. y xa > 0 \<and> (dist x' x < y xa \<longrightarrow> dist (f x') (f x) < xa))" by auto
+    then obtain d where d:"\<forall>e>0. \<forall>x\<in>s. \<forall>x'\<in>s. d x e > 0 \<and> (dist x' x < d x e \<longrightarrow> dist (f x') (f x) < e)"
+      using bchoice[of s "\<lambda>x fa. \<forall>xa>0. \<forall>x'\<in>s. fa xa > 0 \<and> (dist x' x < fa xa \<longrightarrow> dist (f x') (f x) < xa)"] by blast
+
+  { fix e::real assume "e>0"
+
+    { fix x assume "x\<in>s" hence "x \<in> ball x (d x (e / 2))" unfolding centre_in_ball using d[THEN spec[where x="e/2"]] using `e>0` by auto  }
+    hence "s \<subseteq> \<Union>{ball x (d x (e / 2)) |x. x \<in> s}" unfolding subset_eq by auto
+    moreover
+    { fix b assume "b\<in>{ball x (d x (e / 2)) |x. x \<in> s}" hence "open b" by auto  }
+    ultimately obtain ea where "ea>0" and ea:"\<forall>x\<in>s. \<exists>b\<in>{ball x (d x (e / 2)) |x. x \<in> s}. ball x ea \<subseteq> b" using heine_borel_lemma[OF assms(2), of "{ball x (d x (e / 2)) | x. x\<in>s }"] by auto
+
+    { fix x y assume "x\<in>s" "y\<in>s" and as:"dist y x < ea"
+      obtain z where "z\<in>s" and z:"ball x ea \<subseteq> ball z (d z (e / 2))" using ea[THEN bspec[where x=x]] and `x\<in>s` by auto
+      hence "x\<in>ball z (d z (e / 2))" using `ea>0` unfolding subset_eq by auto
+      hence "dist (f z) (f x) < e / 2" using d[THEN spec[where x="e/2"]] and `e>0` and `x\<in>s` and `z\<in>s`
+        by (auto  simp add: dist_commute)
+      moreover have "y\<in>ball z (d z (e / 2))" using as and `ea>0` and z[unfolded subset_eq]
+        by (auto simp add: dist_commute)
+      hence "dist (f z) (f y) < e / 2" using d[THEN spec[where x="e/2"]] and `e>0` and `y\<in>s` and `z\<in>s`
+        by (auto  simp add: dist_commute)
+      ultimately have "dist (f y) (f x) < e" using dist_triangle_half_r[of "f z" "f x" e "f y"]
+        by (auto simp add: dist_commute)  }
+    then have "\<exists>d>0. \<forall>x\<in>s. \<forall>x'\<in>s. dist x' x < d \<longrightarrow> dist (f x') (f x) < e" using `ea>0` by auto  }
+  thus ?thesis unfolding uniformly_continuous_on_def by auto
+qed
+
+text{* Continuity of inverse function on compact domain. *}
+
+lemma continuous_on_inverse:
+  fixes f :: "'a::heine_borel \<Rightarrow> 'b::heine_borel"
+    (* TODO: can this be generalized more? *)
+  assumes "continuous_on s f"  "compact s"  "\<forall>x \<in> s. g (f x) = x"
+  shows "continuous_on (f ` s) g"
+proof-
+  have *:"g ` f ` s = s" using assms(3) by (auto simp add: image_iff)
+  { fix t assume t:"closedin (subtopology euclidean (g ` f ` s)) t"
+    then obtain T where T: "closed T" "t = s \<inter> T" unfolding closedin_closed unfolding * by auto
+    have "continuous_on (s \<inter> T) f" using continuous_on_subset[OF assms(1), of "s \<inter> t"]
+      unfolding T(2) and Int_left_absorb by auto
+    moreover have "compact (s \<inter> T)"
+      using assms(2) unfolding compact_eq_bounded_closed
+      using bounded_subset[of s "s \<inter> T"] and T(1) by auto
+    ultimately have "closed (f ` t)" using T(1) unfolding T(2)
+      using compact_continuous_image [of "s \<inter> T" f] unfolding compact_eq_bounded_closed by auto
+    moreover have "{x \<in> f ` s. g x \<in> t} = f ` s \<inter> f ` t" using assms(3) unfolding T(2) by auto
+    ultimately have "closedin (subtopology euclidean (f ` s)) {x \<in> f ` s. g x \<in> t}"
+      unfolding closedin_closed by auto  }
+  thus ?thesis unfolding continuous_on_closed by auto
+qed
+
+subsection{* A uniformly convergent limit of continuous functions is continuous.       *}
+
+lemma norm_triangle_lt:
+  fixes x y :: "'a::real_normed_vector"
+  shows "norm x + norm y < e \<Longrightarrow> norm (x + y) < e"
+by (rule le_less_trans [OF norm_triangle_ineq])
+
+lemma continuous_uniform_limit:
+  fixes f :: "'a \<Rightarrow> 'b::metric_space \<Rightarrow> 'c::real_normed_vector"
+  assumes "\<not> (trivial_limit net)"  "eventually (\<lambda>n. continuous_on s (f n)) net"
+  "\<forall>e>0. eventually (\<lambda>n. \<forall>x \<in> s. norm(f n x - g x) < e) net"
+  shows "continuous_on s g"
+proof-
+  { fix x and e::real assume "x\<in>s" "e>0"
+    have "eventually (\<lambda>n. \<forall>x\<in>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:"\<forall>xa\<in>s. norm (f n xa - g xa) < e / 3"  "continuous_on s (f n)"
+      using eventually_and[of "(\<lambda>n. \<forall>x\<in>s. norm (f n x - g x) < e / 3)" "(\<lambda>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:"\<forall>x'\<in>s. dist x' x < d \<longrightarrow> dist (f n x') (f n x) < e / 3"
+      using n(2)[unfolded continuous_on_def, THEN bspec[where x=x], OF `x\<in>s`, THEN spec[where x="e/3"]] by blast
+    { fix y assume "y\<in>s" "dist y x < d"
+      hence "dist (f n y) (f n x) < e / 3" using d[THEN bspec[where x=y]] by auto
+      hence "norm (f n y - g x) < 2 * e / 3" using norm_triangle_lt[of "f n y - f n x" "f n x - g x" "2*e/3"]
+        using n(1)[THEN bspec[where x=x], OF `x\<in>s`] unfolding dist_norm unfolding ab_group_add_class.ab_diff_minus by auto
+      hence "dist (g y) (g x) < e" unfolding dist_norm using n(1)[THEN bspec[where x=y], OF `y\<in>s`]
+        unfolding norm_minus_cancel[of "f n y - g y", THEN sym] using norm_triangle_lt[of "f n y - g x" "g y - f n y" e] by (auto simp add: uminus_add_conv_diff)  }
+    hence "\<exists>d>0. \<forall>x'\<in>s. dist x' x < d \<longrightarrow> dist (g x') (g x) < e" using `d>0` by auto  }
+  thus ?thesis unfolding continuous_on_def by auto
+qed
+
+subsection{* Topological properties of linear functions.                               *}
+
+lemma linear_lim_0:
+  assumes "bounded_linear f" shows "(f ---> 0) (at (0))"
+proof-
+  interpret f: bounded_linear f by fact
+  have "(f ---> f 0) (at 0)"
+    using tendsto_ident_at by (rule f.tendsto)
+  thus ?thesis unfolding f.zero .
+qed
+
+lemma linear_continuous_at:
+  assumes "bounded_linear f"  shows "continuous (at a) f"
+  unfolding continuous_at using assms
+  apply (rule bounded_linear.tendsto)
+  apply (rule tendsto_ident_at)
+  done
+
+lemma linear_continuous_within:
+  shows "bounded_linear f ==> continuous (at x within s) f"
+  using continuous_at_imp_continuous_within[of x f s] using linear_continuous_at[of f] by auto
+
+lemma linear_continuous_on:
+  shows "bounded_linear f ==> continuous_on s f"
+  using continuous_at_imp_continuous_on[of s f] using linear_continuous_at[of f] by auto
+
+text{* Also bilinear functions, in composition form.                             *}
+
+lemma bilinear_continuous_at_compose:
+  shows "continuous (at x) f \<Longrightarrow> continuous (at x) g \<Longrightarrow> bounded_bilinear h
+        ==> continuous (at x) (\<lambda>x. h (f x) (g x))"
+  unfolding continuous_at using Lim_bilinear[of f "f x" "(at x)" g "g x" h] by auto
+
+lemma bilinear_continuous_within_compose:
+  shows "continuous (at x within s) f \<Longrightarrow> continuous (at x within s) g \<Longrightarrow> bounded_bilinear h
+        ==> continuous (at x within s) (\<lambda>x. h (f x) (g x))"
+  unfolding continuous_within using Lim_bilinear[of f "f x"] by auto
+
+lemma bilinear_continuous_on_compose:
+  shows "continuous_on s f \<Longrightarrow> continuous_on s g \<Longrightarrow> bounded_bilinear h
+             ==> continuous_on s (\<lambda>x. h (f x) (g x))"
+  unfolding continuous_on_eq_continuous_within apply auto apply(erule_tac x=x in ballE) apply auto apply(erule_tac x=x in ballE) apply auto
+  using bilinear_continuous_within_compose[of _ s f g h] by auto
+
+subsection{* Topological stuff lifted from and dropped to R                            *}
+
+
+lemma open_real:
+  fixes s :: "real set" shows
+ "open s \<longleftrightarrow>
+        (\<forall>x \<in> s. \<exists>e>0. \<forall>x'. abs(x' - x) < e --> x' \<in> s)" (is "?lhs = ?rhs")
+  unfolding open_dist dist_norm by simp
+
+lemma islimpt_approachable_real:
+  fixes s :: "real set"
+  shows "x islimpt s \<longleftrightarrow> (\<forall>e>0.  \<exists>x'\<in> s. x' \<noteq> x \<and> abs(x' - x) < e)"
+  unfolding islimpt_approachable dist_norm by simp
+
+lemma closed_real:
+  fixes s :: "real set"
+  shows "closed s \<longleftrightarrow>
+        (\<forall>x. (\<forall>e>0.  \<exists>x' \<in> s. x' \<noteq> x \<and> abs(x' - x) < e)
+            --> x \<in> s)"
+  unfolding closed_limpt islimpt_approachable dist_norm by simp
+
+lemma continuous_at_real_range:
+  fixes f :: "'a::real_normed_vector \<Rightarrow> real"
+  shows "continuous (at x) f \<longleftrightarrow> (\<forall>e>0. \<exists>d>0.
+        \<forall>x'. norm(x' - x) < d --> abs(f x' - f x) < e)"
+  unfolding continuous_at unfolding Lim_at
+  unfolding dist_nz[THEN sym] unfolding dist_norm apply auto
+  apply(erule_tac x=e in allE) apply auto apply (rule_tac x=d in exI) apply auto apply (erule_tac x=x' in allE) apply auto
+  apply(erule_tac x=e in allE) by auto
+
+lemma continuous_on_real_range:
+  fixes f :: "'a::real_normed_vector \<Rightarrow> real"
+  shows "continuous_on s f \<longleftrightarrow> (\<forall>x \<in> s. \<forall>e>0. \<exists>d>0. (\<forall>x' \<in> s. norm(x' - x) < d --> abs(f x' - f x) < e))"
+  unfolding continuous_on_def dist_norm by simp
+
+lemma continuous_at_norm: "continuous (at x) norm"
+  unfolding continuous_at by (intro tendsto_intros)
+
+lemma continuous_on_norm: "continuous_on s norm"
+unfolding continuous_on by (intro ballI tendsto_intros)
+
+lemma continuous_at_component: "continuous (at a) (\<lambda>x. x $ i)"
+unfolding continuous_at by (intro tendsto_intros)
+
+lemma continuous_on_component: "continuous_on s (\<lambda>x. x $ i)"
+unfolding continuous_on by (intro ballI tendsto_intros)
+
+lemma continuous_at_infnorm: "continuous (at x) infnorm"
+  unfolding continuous_at Lim_at o_def unfolding dist_norm
+  apply auto apply (rule_tac x=e in exI) apply auto
+  using order_trans[OF real_abs_sub_infnorm infnorm_le_norm, of _ x] by (metis xt1(7))
+
+text{* Hence some handy theorems on distance, diameter etc. of/from a set.       *}
+
+lemma compact_attains_sup:
+  fixes s :: "real set"
+  assumes "compact s"  "s \<noteq> {}"
+  shows "\<exists>x \<in> s. \<forall>y \<in> s. y \<le> x"
+proof-
+  from assms(1) have a:"bounded s" "closed s" unfolding compact_eq_bounded_closed by auto
+  { fix e::real assume as: "\<forall>x\<in>s. x \<le> rsup s" "rsup s \<notin> s"  "0 < e" "\<forall>x'\<in>s. x' = rsup s \<or> \<not> rsup s - x' < e"
+    have "isLub UNIV s (rsup s)" using rsup[OF assms(2)] unfolding setle_def using as(1) by auto
+    moreover have "isUb UNIV s (rsup s - e)" unfolding isUb_def unfolding setle_def using as(4,2) by auto
+    ultimately have False using isLub_le_isUb[of UNIV s "rsup s" "rsup s - e"] using `e>0` by auto  }
+  thus ?thesis using bounded_has_rsup(1)[OF a(1) assms(2)] using a(2)[unfolded closed_real, THEN spec[where x="rsup s"]]
+    apply(rule_tac x="rsup s" in bexI) by auto
+qed
+
+lemma compact_attains_inf:
+  fixes s :: "real set"
+  assumes "compact s" "s \<noteq> {}"  shows "\<exists>x \<in> s. \<forall>y \<in> s. x \<le> y"
+proof-
+  from assms(1) have a:"bounded s" "closed s" unfolding compact_eq_bounded_closed by auto
+  { fix e::real assume as: "\<forall>x\<in>s. x \<ge> rinf s"  "rinf s \<notin> s"  "0 < e"
+      "\<forall>x'\<in>s. x' = rinf s \<or> \<not> abs (x' - rinf s) < e"
+    have "isGlb UNIV s (rinf s)" using rinf[OF assms(2)] unfolding setge_def using as(1) by auto
+    moreover
+    { fix x assume "x \<in> s"
+      hence *:"abs (x - rinf s) = x - rinf s" using as(1)[THEN bspec[where x=x]] by auto
+      have "rinf s + e \<le> x" using as(4)[THEN bspec[where x=x]] using as(2) `x\<in>s` unfolding * by auto }
+    hence "isLb UNIV s (rinf s + e)" unfolding isLb_def and setge_def by auto
+    ultimately have False using isGlb_le_isLb[of UNIV s "rinf s" "rinf s + e"] using `e>0` by auto  }
+  thus ?thesis using bounded_has_rinf(1)[OF a(1) assms(2)] using a(2)[unfolded closed_real, THEN spec[where x="rinf s"]]
+    apply(rule_tac x="rinf s" in bexI) by auto
+qed
+
+lemma continuous_attains_sup:
+  fixes f :: "'a::metric_space \<Rightarrow> real"
+  shows "compact s \<Longrightarrow> s \<noteq> {} \<Longrightarrow> continuous_on s f
+        ==> (\<exists>x \<in> s. \<forall>y \<in> s.  f y \<le> f x)"
+  using compact_attains_sup[of "f ` s"]
+  using compact_continuous_image[of s f] by auto
+
+lemma continuous_attains_inf:
+  fixes f :: "'a::metric_space \<Rightarrow> real"
+  shows "compact s \<Longrightarrow> s \<noteq> {} \<Longrightarrow> continuous_on s f
+        \<Longrightarrow> (\<exists>x \<in> s. \<forall>y \<in> s. f x \<le> f y)"
+  using compact_attains_inf[of "f ` s"]
+  using compact_continuous_image[of s f] by auto
+
+lemma distance_attains_sup:
+  assumes "compact s" "s \<noteq> {}"
+  shows "\<exists>x \<in> s. \<forall>y \<in> s. dist a y \<le> dist a x"
+proof (rule continuous_attains_sup [OF assms])
+  { fix x assume "x\<in>s"
+    have "(dist a ---> dist a x) (at x within s)"
+      by (intro tendsto_dist tendsto_const Lim_at_within Lim_ident_at)
+  }
+  thus "continuous_on s (dist a)"
+    unfolding continuous_on ..
+qed
+
+text{* For *minimal* distance, we only need closure, not compactness.            *}
+
+lemma distance_attains_inf:
+  fixes a :: "'a::heine_borel"
+  assumes "closed s"  "s \<noteq> {}"
+  shows "\<exists>x \<in> s. \<forall>y \<in> s. dist a x \<le> dist a y"
+proof-
+  from assms(2) obtain b where "b\<in>s" by auto
+  let ?B = "cball a (dist b a) \<inter> s"
+  have "b \<in> ?B" using `b\<in>s` by (simp add: dist_commute)
+  hence "?B \<noteq> {}" by auto
+  moreover
+  { fix x assume "x\<in>?B"
+    fix e::real assume "e>0"
+    { fix x' assume "x'\<in>?B" and as:"dist x' x < e"
+      from as have "\<bar>dist a x' - dist a x\<bar> < e"
+        unfolding abs_less_iff minus_diff_eq
+        using dist_triangle2 [of a x' x]
+        using dist_triangle [of a x x']
+        by arith
+    }
+    hence "\<exists>d>0. \<forall>x'\<in>?B. dist x' x < d \<longrightarrow> \<bar>dist a x' - dist a x\<bar> < e"
+      using `e>0` by auto
+  }
+  hence "continuous_on (cball a (dist b a) \<inter> s) (dist a)"
+    unfolding continuous_on Lim_within dist_norm real_norm_def
+    by fast
+  moreover have "compact ?B"
+    using compact_cball[of a "dist b a"]
+    unfolding compact_eq_bounded_closed
+    using bounded_Int and closed_Int and assms(1) by auto
+  ultimately obtain x where "x\<in>cball a (dist b a) \<inter> s" "\<forall>y\<in>cball a (dist b a) \<inter> s. dist a x \<le> dist a y"
+    using continuous_attains_inf[of ?B "dist a"] by fastsimp
+  thus ?thesis by fastsimp
+qed
+
+subsection{* We can now extend limit compositions to consider the scalar multiplier.   *}
+
+lemma Lim_mul:
+  fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"
+  assumes "(c ---> d) net"  "(f ---> l) net"
+  shows "((\<lambda>x. c(x) *\<^sub>R f x) ---> (d *\<^sub>R l)) net"
+  using assms by (rule scaleR.tendsto)
+
+lemma Lim_vmul:
+  fixes c :: "'a \<Rightarrow> real" and v :: "'b::real_normed_vector"
+  shows "(c ---> d) net ==> ((\<lambda>x. c(x) *\<^sub>R v) ---> d *\<^sub>R v) net"
+  by (intro tendsto_intros)
+
+lemma continuous_vmul:
+  fixes c :: "'a::metric_space \<Rightarrow> real" and v :: "'b::real_normed_vector"
+  shows "continuous net c ==> continuous net (\<lambda>x. c(x) *\<^sub>R v)"
+  unfolding continuous_def using Lim_vmul[of c] by auto
+
+lemma continuous_mul:
+  fixes c :: "'a::metric_space \<Rightarrow> real"
+  fixes f :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector"
+  shows "continuous net c \<Longrightarrow> continuous net f
+             ==> continuous net (\<lambda>x. c(x) *\<^sub>R f x) "
+  unfolding continuous_def by (intro tendsto_intros)
+
+lemma continuous_on_vmul:
+  fixes c :: "'a::metric_space \<Rightarrow> real" and v :: "'b::real_normed_vector"
+  shows "continuous_on s c ==> continuous_on s (\<lambda>x. c(x) *\<^sub>R v)"
+  unfolding continuous_on_eq_continuous_within using continuous_vmul[of _ c] by auto
+
+lemma continuous_on_mul:
+  fixes c :: "'a::metric_space \<Rightarrow> real"
+  fixes f :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector"
+  shows "continuous_on s c \<Longrightarrow> continuous_on s f
+             ==> continuous_on s (\<lambda>x. c(x) *\<^sub>R f x)"
+  unfolding continuous_on_eq_continuous_within using continuous_mul[of _ c] by auto
+
+text{* And so we have continuity of inverse.                                     *}
+
+lemma Lim_inv:
+  fixes f :: "'a \<Rightarrow> real"
+  assumes "(f ---> l) (net::'a net)"  "l \<noteq> 0"
+  shows "((inverse o f) ---> inverse l) net"
+  unfolding o_def using assms by (rule tendsto_inverse)
+
+lemma continuous_inv:
+  fixes f :: "'a::metric_space \<Rightarrow> real"
+  shows "continuous net f \<Longrightarrow> f(netlimit net) \<noteq> 0
+           ==> continuous net (inverse o f)"
+  unfolding continuous_def using Lim_inv by auto
+
+lemma continuous_at_within_inv:
+  fixes f :: "'a::metric_space \<Rightarrow> 'b::real_normed_field"
+  assumes "continuous (at a within s) f" "f a \<noteq> 0"
+  shows "continuous (at a within s) (inverse o f)"
+  using assms unfolding continuous_within o_def
+  by (intro tendsto_intros)
+
+lemma continuous_at_inv:
+  fixes f :: "'a::metric_space \<Rightarrow> 'b::real_normed_field"
+  shows "continuous (at a) f \<Longrightarrow> f a \<noteq> 0
+         ==> continuous (at a) (inverse o f) "
+  using within_UNIV[THEN sym, of "at a"] using continuous_at_within_inv[of a UNIV] by auto
+
+subsection{* Preservation properties for pasted sets.                                  *}
+
+lemma bounded_pastecart:
+  fixes s :: "('a::real_normed_vector ^ _) set" (* FIXME: generalize to metric_space *)
+  assumes "bounded s" "bounded t"
+  shows "bounded { pastecart x y | x y . (x \<in> s \<and> y \<in> t)}"
+proof-
+  obtain a b where ab:"\<forall>x\<in>s. norm x \<le> a" "\<forall>x\<in>t. norm x \<le> b" using assms[unfolded bounded_iff] by auto
+  { fix x y assume "x\<in>s" "y\<in>t"
+    hence "norm x \<le> a" "norm y \<le> b" using ab by auto
+    hence "norm (pastecart x y) \<le> a + b" using norm_pastecart[of x y] by auto }
+  thus ?thesis unfolding bounded_iff by auto
+qed
+
+lemma bounded_Times:
+  assumes "bounded s" "bounded t" shows "bounded (s \<times> t)"
+proof-
+  obtain x y a b where "\<forall>z\<in>s. dist x z \<le> a" "\<forall>z\<in>t. dist y z \<le> b"
+    using assms [unfolded bounded_def] by auto
+  then have "\<forall>z\<in>s \<times> t. dist (x, y) z \<le> sqrt (a\<twosuperior> + b\<twosuperior>)"
+    by (auto simp add: dist_Pair_Pair real_sqrt_le_mono add_mono power_mono)
+  thus ?thesis unfolding bounded_any_center [where a="(x, y)"] by auto
+qed
+
+lemma closed_pastecart:
+  fixes s :: "(real ^ 'a::finite) set" (* FIXME: generalize *)
+  assumes "closed s"  "closed t"
+  shows "closed {pastecart x y | x y . x \<in> s \<and> y \<in> t}"
+proof-
+  { fix x l assume as:"\<forall>n::nat. x n \<in> {pastecart x y |x y. x \<in> s \<and> y \<in> t}"  "(x ---> l) sequentially"
+    { fix n::nat have "fstcart (x n) \<in> s" "sndcart (x n) \<in> t" using as(1)[THEN spec[where x=n]] by auto } note * = this
+    moreover
+    { fix e::real assume "e>0"
+      then obtain N::nat where N:"\<forall>n\<ge>N. dist (x n) l < e" using as(2)[unfolded Lim_sequentially, THEN spec[where x=e]] by auto
+      { fix n::nat assume "n\<ge>N"
+        hence "dist (fstcart (x n)) (fstcart l) < e" "dist (sndcart (x n)) (sndcart l) < e"
+          using N[THEN spec[where x=n]] dist_fstcart[of "x n" l] dist_sndcart[of "x n" l] by auto   }
+      hence "\<exists>N. \<forall>n\<ge>N. dist (fstcart (x n)) (fstcart l) < e" "\<exists>N. \<forall>n\<ge>N. dist (sndcart (x n)) (sndcart l) < e" by auto  }
+    ultimately have "fstcart l \<in> s" "sndcart l \<in> t"
+      using assms(1)[unfolded closed_sequential_limits, THEN spec[where x="\<lambda>n. fstcart (x n)"], THEN spec[where x="fstcart l"]]
+      using assms(2)[unfolded closed_sequential_limits, THEN spec[where x="\<lambda>n. sndcart (x n)"], THEN spec[where x="sndcart l"]]
+      unfolding Lim_sequentially by auto
+    hence "l \<in> {pastecart x y |x y. x \<in> s \<and> y \<in> t}" using pastecart_fst_snd[THEN sym, of l] by auto  }
+  thus ?thesis unfolding closed_sequential_limits by auto
+qed
+
+lemma compact_pastecart:
+  fixes s t :: "(real ^ _) set"
+  shows "compact s \<Longrightarrow> compact t ==> compact {pastecart x y | x y . x \<in> s \<and> y \<in> t}"
+  unfolding compact_eq_bounded_closed using bounded_pastecart[of s t] closed_pastecart[of s t] by auto
+
+lemma mem_Times_iff: "x \<in> A \<times> B \<longleftrightarrow> fst x \<in> A \<and> snd x \<in> B"
+by (induct x) simp
+
+lemma compact_Times: "compact s \<Longrightarrow> compact t \<Longrightarrow> compact (s \<times> t)"
+unfolding compact_def
+apply clarify
+apply (drule_tac x="fst \<circ> f" in spec)
+apply (drule mp, simp add: mem_Times_iff)
+apply (clarify, rename_tac l1 r1)
+apply (drule_tac x="snd \<circ> f \<circ> r1" in spec)
+apply (drule mp, simp add: mem_Times_iff)
+apply (clarify, rename_tac l2 r2)
+apply (rule_tac x="(l1, l2)" in rev_bexI, simp)
+apply (rule_tac x="r1 \<circ> r2" in exI)
+apply (rule conjI, simp add: subseq_def)
+apply (drule_tac r=r2 in lim_subseq [COMP swap_prems_rl], assumption)
+apply (drule (1) tendsto_Pair) back
+apply (simp add: o_def)
+done
+
+text{* Hence some useful properties follow quite easily.                         *}
+
+lemma compact_scaling:
+  fixes s :: "'a::real_normed_vector set"
+  assumes "compact s"  shows "compact ((\<lambda>x. c *\<^sub>R x) ` s)"
+proof-
+  let ?f = "\<lambda>x. scaleR c x"
+  have *:"bounded_linear ?f" by (rule scaleR.bounded_linear_right)
+  show ?thesis using compact_continuous_image[of s ?f] continuous_at_imp_continuous_on[of s ?f]
+    using linear_continuous_at[OF *] assms by auto
+qed
+
+lemma compact_negations:
+  fixes s :: "'a::real_normed_vector set"
+  assumes "compact s"  shows "compact ((\<lambda>x. -x) ` s)"
+  using compact_scaling [OF assms, of "- 1"] by auto
+
+lemma compact_sums:
+  fixes s t :: "'a::real_normed_vector set"
+  assumes "compact s"  "compact t"  shows "compact {x + y | x y. x \<in> s \<and> y \<in> t}"
+proof-
+  have *:"{x + y | x y. x \<in> s \<and> y \<in> t} = (\<lambda>z. fst z + snd z) ` (s \<times> t)"
+    apply auto unfolding image_iff apply(rule_tac x="(xa, y)" in bexI) by auto
+  have "continuous_on (s \<times> t) (\<lambda>z. fst z + snd z)"
+    unfolding continuous_on by (rule ballI) (intro tendsto_intros)
+  thus ?thesis unfolding * using compact_continuous_image compact_Times [OF assms] by auto
+qed
+
+lemma compact_differences:
+  fixes s t :: "'a::real_normed_vector set"
+  assumes "compact s" "compact t"  shows "compact {x - y | x y. x \<in> s \<and> y \<in> t}"
+proof-
+  have "{x - y | x y. x\<in>s \<and> y \<in> t} =  {x + y | x y. x \<in> s \<and> y \<in> (uminus ` t)}"
+    apply auto apply(rule_tac x= xa in exI) apply auto apply(rule_tac x=xa in exI) by auto
+  thus ?thesis using compact_sums[OF assms(1) compact_negations[OF assms(2)]] by auto
+qed
+
+lemma compact_translation:
+  fixes s :: "'a::real_normed_vector set"
+  assumes "compact s"  shows "compact ((\<lambda>x. a + x) ` s)"
+proof-
+  have "{x + y |x y. x \<in> s \<and> y \<in> {a}} = (\<lambda>x. a + x) ` s" by auto
+  thus ?thesis using compact_sums[OF assms compact_sing[of a]] by auto
+qed
+
+lemma compact_affinity:
+  fixes s :: "'a::real_normed_vector set"
+  assumes "compact s"  shows "compact ((\<lambda>x. a + c *\<^sub>R x) ` s)"
+proof-
+  have "op + a ` op *\<^sub>R c ` s = (\<lambda>x. a + c *\<^sub>R x) ` s" by auto
+  thus ?thesis using compact_translation[OF compact_scaling[OF assms], of a c] by auto
+qed
+
+text{* Hence we get the following.                                               *}
+
+lemma compact_sup_maxdistance:
+  fixes s :: "'a::real_normed_vector set"
+  assumes "compact s"  "s \<noteq> {}"
+  shows "\<exists>x\<in>s. \<exists>y\<in>s. \<forall>u\<in>s. \<forall>v\<in>s. norm(u - v) \<le> norm(x - y)"
+proof-
+  have "{x - y | x y . x\<in>s \<and> y\<in>s} \<noteq> {}" using `s \<noteq> {}` by auto
+  then obtain x where x:"x\<in>{x - y |x y. x \<in> s \<and> y \<in> s}"  "\<forall>y\<in>{x - y |x y. x \<in> s \<and> y \<in> s}. norm y \<le> norm x"
+    using compact_differences[OF assms(1) assms(1)]
+    using distance_attains_sup[where 'a="'a", unfolded dist_norm, of "{x - y | x y . x\<in>s \<and> y\<in>s}" 0] by(auto simp add: norm_minus_cancel)
+  from x(1) obtain a b where "a\<in>s" "b\<in>s" "x = a - b" by auto
+  thus ?thesis using x(2)[unfolded `x = a - b`] by blast
+qed
+
+text{* We can state this in terms of diameter of a set.                          *}
+
+definition "diameter s = (if s = {} then 0::real else rsup {norm(x - y) | x y. x \<in> s \<and> y \<in> s})"
+  (* TODO: generalize to class metric_space *)
+
+lemma diameter_bounded:
+  assumes "bounded s"
+  shows "\<forall>x\<in>s. \<forall>y\<in>s. norm(x - y) \<le> diameter s"
+        "\<forall>d>0. d < diameter s --> (\<exists>x\<in>s. \<exists>y\<in>s. norm(x - y) > d)"
+proof-
+  let ?D = "{norm (x - y) |x y. x \<in> s \<and> y \<in> s}"
+  obtain a where a:"\<forall>x\<in>s. norm x \<le> a" using assms[unfolded bounded_iff] by auto
+  { fix x y assume "x \<in> s" "y \<in> s"
+    hence "norm (x - y) \<le> 2 * a" using norm_triangle_ineq[of x "-y", unfolded norm_minus_cancel] a[THEN bspec[where x=x]] a[THEN bspec[where x=y]] by (auto simp add: ring_simps)  }
+  note * = this
+  { fix x y assume "x\<in>s" "y\<in>s"  hence "s \<noteq> {}" by auto
+    have lub:"isLub UNIV ?D (rsup ?D)" using * rsup[of ?D] using `s\<noteq>{}` unfolding setle_def by auto
+    have "norm(x - y) \<le> diameter s" unfolding diameter_def using `s\<noteq>{}` *[OF `x\<in>s` `y\<in>s`] `x\<in>s` `y\<in>s` isLubD1[OF lub] unfolding setle_def by auto  }
+  moreover
+  { fix d::real assume "d>0" "d < diameter s"
+    hence "s\<noteq>{}" unfolding diameter_def by auto
+    hence lub:"isLub UNIV ?D (rsup ?D)" using * rsup[of ?D] unfolding setle_def by auto
+    have "\<exists>d' \<in> ?D. d' > d"
+    proof(rule ccontr)
+      assume "\<not> (\<exists>d'\<in>{norm (x - y) |x y. x \<in> s \<and> y \<in> s}. d < d')"
+      hence as:"\<forall>d'\<in>?D. d' \<le> d" apply auto apply(erule_tac x="norm (x - y)" in allE) by auto
+      hence "isUb UNIV ?D d" unfolding isUb_def unfolding setle_def by auto
+      thus False using `d < diameter s` `s\<noteq>{}` isLub_le_isUb[OF lub, of d] unfolding diameter_def  by auto
+    qed
+    hence "\<exists>x\<in>s. \<exists>y\<in>s. norm(x - y) > d" by auto  }
+  ultimately show "\<forall>x\<in>s. \<forall>y\<in>s. norm(x - y) \<le> diameter s"
+        "\<forall>d>0. d < diameter s --> (\<exists>x\<in>s. \<exists>y\<in>s. norm(x - y) > d)" by auto
+qed
+
+lemma diameter_bounded_bound:
+ "bounded s \<Longrightarrow> x \<in> s \<Longrightarrow> y \<in> s ==> norm(x - y) \<le> diameter s"
+  using diameter_bounded by blast
+
+lemma diameter_compact_attained:
+  fixes s :: "'a::real_normed_vector set"
+  assumes "compact s"  "s \<noteq> {}"
+  shows "\<exists>x\<in>s. \<exists>y\<in>s. (norm(x - y) = diameter s)"
+proof-
+  have b:"bounded s" using assms(1) by (rule compact_imp_bounded)
+  then obtain x y where xys:"x\<in>s" "y\<in>s" and xy:"\<forall>u\<in>s. \<forall>v\<in>s. norm (u - v) \<le> norm (x - y)" using compact_sup_maxdistance[OF assms] by auto
+  hence "diameter s \<le> norm (x - y)" using rsup_le[of "{norm (x - y) |x y. x \<in> s \<and> y \<in> s}" "norm (x - y)"]
+    unfolding setle_def and diameter_def by auto
+  thus ?thesis using diameter_bounded(1)[OF b, THEN bspec[where x=x], THEN bspec[where x=y], OF xys] and xys by auto
+qed
+
+text{* Related results with closure as the conclusion.                           *}
+
+lemma closed_scaling:
+  fixes s :: "'a::real_normed_vector set"
+  assumes "closed s" shows "closed ((\<lambda>x. c *\<^sub>R x) ` s)"
+proof(cases "s={}")
+  case True thus ?thesis by auto
+next
+  case False
+  show ?thesis
+  proof(cases "c=0")
+    have *:"(\<lambda>x. 0) ` s = {0}" using `s\<noteq>{}` by auto
+    case True thus ?thesis apply auto unfolding * using closed_sing by auto
+  next
+    case False
+    { fix x l assume as:"\<forall>n::nat. x n \<in> scaleR c ` s"  "(x ---> l) sequentially"
+      { fix n::nat have "scaleR (1 / c) (x n) \<in> s"
+          using as(1)[THEN spec[where x=n]]
+          using `c\<noteq>0` by (auto simp add: vector_smult_assoc)
+      }
+      moreover
+      { fix e::real assume "e>0"
+        hence "0 < e *\<bar>c\<bar>"  using `c\<noteq>0` mult_pos_pos[of e "abs c"] by auto
+        then obtain N where "\<forall>n\<ge>N. dist (x n) l < e * \<bar>c\<bar>"
+          using as(2)[unfolded Lim_sequentially, THEN spec[where x="e * abs c"]] by auto
+        hence "\<exists>N. \<forall>n\<ge>N. dist (scaleR (1 / c) (x n)) (scaleR (1 / c) l) < e"
+          unfolding dist_norm unfolding scaleR_right_diff_distrib[THEN sym]
+          using mult_imp_div_pos_less[of "abs c" _ e] `c\<noteq>0` by auto  }
+      hence "((\<lambda>n. scaleR (1 / c) (x n)) ---> scaleR (1 / c) l) sequentially" unfolding Lim_sequentially by auto
+      ultimately have "l \<in> scaleR c ` s"
+        using assms[unfolded closed_sequential_limits, THEN spec[where x="\<lambda>n. scaleR (1/c) (x n)"], THEN spec[where x="scaleR (1/c) l"]]
+        unfolding image_iff using `c\<noteq>0` apply(rule_tac x="scaleR (1 / c) l" in bexI) by auto  }
+    thus ?thesis unfolding closed_sequential_limits by fast
+  qed
+qed
+
+lemma closed_negations:
+  fixes s :: "'a::real_normed_vector set"
+  assumes "closed s"  shows "closed ((\<lambda>x. -x) ` s)"
+  using closed_scaling[OF assms, of "- 1"] by simp
+
+lemma compact_closed_sums:
+  fixes s :: "'a::real_normed_vector set"
+  assumes "compact s"  "closed t"  shows "closed {x + y | x y. x \<in> s \<and> y \<in> t}"
+proof-
+  let ?S = "{x + y |x y. x \<in> s \<and> y \<in> t}"
+  { fix x l assume as:"\<forall>n. x n \<in> ?S"  "(x ---> l) sequentially"
+    from as(1) obtain f where f:"\<forall>n. x n = fst (f n) + snd (f n)"  "\<forall>n. fst (f n) \<in> s"  "\<forall>n. snd (f n) \<in> t"
+      using choice[of "\<lambda>n y. x n = (fst y) + (snd y) \<and> fst y \<in> s \<and> snd y \<in> t"] by auto
+    obtain l' r where "l'\<in>s" and r:"subseq r" and lr:"(((\<lambda>n. fst (f n)) \<circ> r) ---> l') sequentially"
+      using assms(1)[unfolded compact_def, THEN spec[where x="\<lambda> n. fst (f n)"]] using f(2) by auto
+    have "((\<lambda>n. snd (f (r n))) ---> l - l') sequentially"
+      using Lim_sub[OF lim_subseq[OF r as(2)] lr] and f(1) unfolding o_def by auto
+    hence "l - l' \<in> t"
+      using assms(2)[unfolded closed_sequential_limits, THEN spec[where x="\<lambda> n. snd (f (r n))"], THEN spec[where x="l - l'"]]
+      using f(3) by auto
+    hence "l \<in> ?S" using `l' \<in> s` apply auto apply(rule_tac x=l' in exI) apply(rule_tac x="l - l'" in exI) by auto
+  }
+  thus ?thesis unfolding closed_sequential_limits by fast
+qed
+
+lemma closed_compact_sums:
+  fixes s t :: "'a::real_normed_vector set"
+  assumes "closed s"  "compact t"
+  shows "closed {x + y | x y. x \<in> s \<and> y \<in> t}"
+proof-
+  have "{x + y |x y. x \<in> t \<and> y \<in> s} = {x + y |x y. x \<in> s \<and> y \<in> t}" apply auto
+    apply(rule_tac x=y in exI) apply auto apply(rule_tac x=y in exI) by auto
+  thus ?thesis using compact_closed_sums[OF assms(2,1)] by simp
+qed
+
+lemma compact_closed_differences:
+  fixes s t :: "'a::real_normed_vector set"
+  assumes "compact s"  "closed t"
+  shows "closed {x - y | x y. x \<in> s \<and> y \<in> t}"
+proof-
+  have "{x + y |x y. x \<in> s \<and> y \<in> uminus ` t} =  {x - y |x y. x \<in> s \<and> y \<in> t}"
+    apply auto apply(rule_tac x=xa in exI) apply auto apply(rule_tac x=xa in exI) by auto
+  thus ?thesis using compact_closed_sums[OF assms(1) closed_negations[OF assms(2)]] by auto
+qed
+
+lemma closed_compact_differences:
+  fixes s t :: "'a::real_normed_vector set"
+  assumes "closed s" "compact t"
+  shows "closed {x - y | x y. x \<in> s \<and> y \<in> t}"
+proof-
+  have "{x + y |x y. x \<in> s \<and> y \<in> uminus ` t} = {x - y |x y. x \<in> s \<and> y \<in> t}"
+    apply auto apply(rule_tac x=xa in exI) apply auto apply(rule_tac x=xa in exI) by auto
+ thus ?thesis using closed_compact_sums[OF assms(1) compact_negations[OF assms(2)]] by simp
+qed
+
+lemma closed_translation:
+  fixes a :: "'a::real_normed_vector"
+  assumes "closed s"  shows "closed ((\<lambda>x. a + x) ` s)"
+proof-
+  have "{a + y |y. y \<in> s} = (op + a ` s)" by auto
+  thus ?thesis using compact_closed_sums[OF compact_sing[of a] assms] by auto
+qed
+
+lemma translation_UNIV:
+  fixes a :: "'a::ab_group_add" shows "range (\<lambda>x. a + x) = UNIV"
+  apply (auto simp add: image_iff) apply(rule_tac x="x - a" in exI) by auto
+
+lemma translation_diff:
+  fixes a :: "'a::ab_group_add"
+  shows "(\<lambda>x. a + x) ` (s - t) = ((\<lambda>x. a + x) ` s) - ((\<lambda>x. a + x) ` t)"
+  by auto
+
+lemma closure_translation:
+  fixes a :: "'a::real_normed_vector"
+  shows "closure ((\<lambda>x. a + x) ` s) = (\<lambda>x. a + x) ` (closure s)"
+proof-
+  have *:"op + a ` (UNIV - s) = UNIV - op + a ` s"
+    apply auto unfolding image_iff apply(rule_tac x="x - a" in bexI) by auto
+  show ?thesis unfolding closure_interior translation_diff translation_UNIV
+    using interior_translation[of a "UNIV - s"] unfolding * by auto
+qed
+
+lemma frontier_translation:
+  fixes a :: "'a::real_normed_vector"
+  shows "frontier((\<lambda>x. a + x) ` s) = (\<lambda>x. a + x) ` (frontier s)"
+  unfolding frontier_def translation_diff interior_translation closure_translation by auto
+
+subsection{* Separation between points and sets.                                       *}
+
+lemma separate_point_closed:
+  fixes s :: "'a::heine_borel set"
+  shows "closed s \<Longrightarrow> a \<notin> s  ==> (\<exists>d>0. \<forall>x\<in>s. d \<le> dist a x)"
+proof(cases "s = {}")
+  case True
+  thus ?thesis by(auto intro!: exI[where x=1])
+next
+  case False
+  assume "closed s" "a \<notin> s"
+  then obtain x where "x\<in>s" "\<forall>y\<in>s. dist a x \<le> dist a y" using `s \<noteq> {}` distance_attains_inf [of s a] by blast
+  with `x\<in>s` show ?thesis using dist_pos_lt[of a x] and`a \<notin> s` by blast
+qed
+
+lemma separate_compact_closed:
+  fixes s t :: "'a::{heine_borel, real_normed_vector} set"
+    (* TODO: does this generalize to heine_borel? *)
+  assumes "compact s" and "closed t" and "s \<inter> t = {}"
+  shows "\<exists>d>0. \<forall>x\<in>s. \<forall>y\<in>t. d \<le> dist x y"
+proof-
+  have "0 \<notin> {x - y |x y. x \<in> s \<and> y \<in> t}" using assms(3) by auto
+  then obtain d where "d>0" and d:"\<forall>x\<in>{x - y |x y. x \<in> s \<and> y \<in> t}. d \<le> dist 0 x"
+    using separate_point_closed[OF compact_closed_differences[OF assms(1,2)], of 0] by auto
+  { fix x y assume "x\<in>s" "y\<in>t"
+    hence "x - y \<in> {x - y |x y. x \<in> s \<and> y \<in> t}" by auto
+    hence "d \<le> dist (x - y) 0" using d[THEN bspec[where x="x - y"]] using dist_commute
+      by (auto  simp add: dist_commute)
+    hence "d \<le> dist x y" unfolding dist_norm by auto  }
+  thus ?thesis using `d>0` by auto
+qed
+
+lemma separate_closed_compact:
+  fixes s t :: "'a::{heine_borel, real_normed_vector} set"
+  assumes "closed s" and "compact t" and "s \<inter> t = {}"
+  shows "\<exists>d>0. \<forall>x\<in>s. \<forall>y\<in>t. d \<le> dist x y"
+proof-
+  have *:"t \<inter> s = {}" using assms(3) by auto
+  show ?thesis using separate_compact_closed[OF assms(2,1) *]
+    apply auto apply(rule_tac x=d in exI) apply auto apply (erule_tac x=y in ballE)
+    by (auto simp add: dist_commute)
+qed
+
+(* A cute way of denoting open and closed intervals using overloading.       *)
+
+lemma interval: fixes a :: "'a::ord^'n::finite" shows
+  "{a <..< b} = {x::'a^'n. \<forall>i. a$i < x$i \<and> x$i < b$i}" and
+  "{a .. b} = {x::'a^'n. \<forall>i. a$i \<le> x$i \<and> x$i \<le> b$i}"
+  by (auto simp add: expand_set_eq vector_less_def vector_less_eq_def)
+
+lemma mem_interval: fixes a :: "'a::ord^'n::finite" shows
+  "x \<in> {a<..<b} \<longleftrightarrow> (\<forall>i. a$i < x$i \<and> x$i < b$i)"
+  "x \<in> {a .. b} \<longleftrightarrow> (\<forall>i. a$i \<le> x$i \<and> x$i \<le> b$i)"
+  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 \<in> {a .. b} \<longleftrightarrow> dest_vec1 a \<le> dest_vec1 x \<and> dest_vec1 x \<le> dest_vec1 b)"
+ "(x \<in> {a<..<b} \<longleftrightarrow> dest_vec1 a < dest_vec1 x \<and> 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} = {} \<longleftrightarrow> (\<exists>i. b$i \<le> a$i))" (is ?th1) and
+ "({a  ..  b} = {} \<longleftrightarrow> (\<exists>i. b$i < a$i))" (is ?th2)
+proof-
+  { fix i x assume as:"b$i \<le> a$i" and x:"x\<in>{a <..< b}"
+    hence "a $ i < x $ i \<and> x $ i < b $ i" unfolding mem_interval by auto
+    hence "a$i < b$i" by auto
+    hence False using as by auto  }
+  moreover
+  { assume as:"\<forall>i. \<not> (b$i \<le> a$i)"
+    let ?x = "(1/2) *\<^sub>R (a + b)"
+    { fix i
+      have "a$i < b$i" using as[THEN spec[where x=i]] by auto
+      hence "a$i < ((1/2) *\<^sub>R (a+b)) $ i" "((1/2) *\<^sub>R (a+b)) $ i < b$i"
+        unfolding vector_smult_component and vector_add_component
+        by (auto simp add: less_divide_eq_number_of1)  }
+    hence "{a <..< b} \<noteq> {}" using mem_interval(1)[of "?x" a b] by auto  }
+  ultimately show ?th1 by blast
+
+  { fix i x assume as:"b$i < a$i" and x:"x\<in>{a .. b}"
+    hence "a $ i \<le> x $ i \<and> x $ i \<le> b $ i" unfolding mem_interval by auto
+    hence "a$i \<le> b$i" by auto
+    hence False using as by auto  }
+  moreover
+  { assume as:"\<forall>i. \<not> (b$i < a$i)"
+    let ?x = "(1/2) *\<^sub>R (a + b)"
+    { fix i
+      have "a$i \<le> b$i" using as[THEN spec[where x=i]] by auto
+      hence "a$i \<le> ((1/2) *\<^sub>R (a+b)) $ i" "((1/2) *\<^sub>R (a+b)) $ i \<le> b$i"
+        unfolding vector_smult_component and vector_add_component
+        by (auto simp add: less_divide_eq_number_of1)  }
+    hence "{a .. b} \<noteq> {}" using mem_interval(2)[of "?x" a b] by auto  }
+  ultimately show ?th2 by blast
+qed
+
+lemma interval_ne_empty: fixes a :: "real^'n::finite" shows
+  "{a  ..  b} \<noteq> {} \<longleftrightarrow> (\<forall>i. a$i \<le> b$i)" and
+  "{a <..< b} \<noteq> {} \<longleftrightarrow> (\<forall>i. a$i < b$i)"
+  unfolding interval_eq_empty[of a b] by (auto simp add: not_less not_le) (* BH: Why doesn't just "auto" work here? *)
+
+lemma subset_interval_imp: fixes a :: "real^'n::finite" shows
+ "(\<forall>i. a$i \<le> c$i \<and> d$i \<le> b$i) \<Longrightarrow> {c .. d} \<subseteq> {a .. b}" and
+ "(\<forall>i. a$i < c$i \<and> d$i < b$i) \<Longrightarrow> {c .. d} \<subseteq> {a<..<b}" and
+ "(\<forall>i. a$i \<le> c$i \<and> d$i \<le> b$i) \<Longrightarrow> {c<..<d} \<subseteq> {a .. b}" and
+ "(\<forall>i. a$i \<le> c$i \<and> d$i \<le> b$i) \<Longrightarrow> {c<..<d} \<subseteq> {a<..<b}"
+  unfolding subset_eq[unfolded Ball_def] unfolding mem_interval
+  by (auto intro: order_trans less_le_trans le_less_trans less_imp_le) (* BH: Why doesn't just "auto" work here? *)
+
+lemma interval_sing: fixes a :: "'a::linorder^'n::finite" shows
+ "{a .. a} = {a} \<and> {a<..<a} = {}"
+apply(auto simp add: expand_set_eq vector_less_def vector_less_eq_def Cart_eq)
+apply (simp add: order_eq_iff)
+apply (auto simp add: not_less less_imp_le)
+done
+
+lemma interval_open_subset_closed:  fixes a :: "'a::preorder^'n::finite" shows
+ "{a<..<b} \<subseteq> {a .. b}"
+proof(simp add: subset_eq, rule)
+  fix x
+  assume x:"x \<in>{a<..<b}"
+  { fix i
+    have "a $ i \<le> x $ i"
+      using x order_less_imp_le[of "a$i" "x$i"]
+      by(simp add: expand_set_eq vector_less_def vector_less_eq_def Cart_eq)
+  }
+  moreover
+  { fix i
+    have "x $ i \<le> b $ i"
+      using x order_less_imp_le[of "x$i" "b$i"]
+      by(simp add: expand_set_eq vector_less_def vector_less_eq_def Cart_eq)
+  }
+  ultimately
+  show "a \<le> x \<and> x \<le> b"
+    by(simp add: expand_set_eq vector_less_def vector_less_eq_def Cart_eq)
+qed
+
+lemma subset_interval: fixes a :: "real^'n::finite" shows
+ "{c .. d} \<subseteq> {a .. b} \<longleftrightarrow> (\<forall>i. c$i \<le> d$i) --> (\<forall>i. a$i \<le> c$i \<and> d$i \<le> b$i)" (is ?th1) and
+ "{c .. d} \<subseteq> {a<..<b} \<longleftrightarrow> (\<forall>i. c$i \<le> d$i) --> (\<forall>i. a$i < c$i \<and> d$i < b$i)" (is ?th2) and
+ "{c<..<d} \<subseteq> {a .. b} \<longleftrightarrow> (\<forall>i. c$i < d$i) --> (\<forall>i. a$i \<le> c$i \<and> d$i \<le> b$i)" (is ?th3) and
+ "{c<..<d} \<subseteq> {a<..<b} \<longleftrightarrow> (\<forall>i. c$i < d$i) --> (\<forall>i. a$i \<le> c$i \<and> d$i \<le> b$i)" (is ?th4)
+proof-
+  show ?th1 unfolding subset_eq and Ball_def and mem_interval by (auto intro: order_trans)
+  show ?th2 unfolding subset_eq and Ball_def and mem_interval by (auto intro: le_less_trans less_le_trans order_trans less_imp_le)
+  { assume as: "{c<..<d} \<subseteq> {a .. b}" "\<forall>i. c$i < d$i"
+    hence "{c<..<d} \<noteq> {}" unfolding interval_eq_empty by (auto, drule_tac x=i in spec, simp) (* BH: Why doesn't just "auto" work? *)
+    fix i
+    (** TODO combine the following two parts as done in the HOL_light version. **)
+    { let ?x = "(\<chi> j. (if j=i then ((min (a$j) (d$j))+c$j)/2 else (c$j+d$j)/2))::real^'n"
+      assume as2: "a$i > c$i"
+      { fix j
+        have "c $ j < ?x $ j \<and> ?x $ j < d $ j" unfolding Cart_lambda_beta
+          apply(cases "j=i") using as(2)[THEN spec[where x=j]]
+          by (auto simp add: less_divide_eq_number_of1 as2)  }
+      hence "?x\<in>{c<..<d}" unfolding mem_interval by auto
+      moreover
+      have "?x\<notin>{a .. b}"
+        unfolding mem_interval apply auto apply(rule_tac x=i in exI)
+        using as(2)[THEN spec[where x=i]] and as2
+        by (auto simp add: less_divide_eq_number_of1)
+      ultimately have False using as by auto  }
+    hence "a$i \<le> c$i" by(rule ccontr)auto
+    moreover
+    { let ?x = "(\<chi> j. (if j=i then ((max (b$j) (c$j))+d$j)/2 else (c$j+d$j)/2))::real^'n"
+      assume as2: "b$i < d$i"
+      { fix j
+        have "d $ j > ?x $ j \<and> ?x $ j > c $ j" unfolding Cart_lambda_beta
+          apply(cases "j=i") using as(2)[THEN spec[where x=j]]
+          by (auto simp add: less_divide_eq_number_of1 as2)  }
+      hence "?x\<in>{c<..<d}" unfolding mem_interval by auto
+      moreover
+      have "?x\<notin>{a .. b}"
+        unfolding mem_interval apply auto apply(rule_tac x=i in exI)
+        using as(2)[THEN spec[where x=i]] and as2
+        by (auto simp add: less_divide_eq_number_of1)
+      ultimately have False using as by auto  }
+    hence "b$i \<ge> d$i" by(rule ccontr)auto
+    ultimately
+    have "a$i \<le> c$i \<and> d$i \<le> b$i" by auto
+  } note part1 = this
+  thus ?th3 unfolding subset_eq and Ball_def and mem_interval apply auto apply (erule_tac x=ia in allE, simp)+ by (erule_tac x=i in allE, erule_tac x=i in allE, simp)+
+  { assume as:"{c<..<d} \<subseteq> {a<..<b}" "\<forall>i. c$i < d$i"
+    fix i
+    from as(1) have "{c<..<d} \<subseteq> {a..b}" using interval_open_subset_closed[of a b] by auto
+    hence "a$i \<le> c$i \<and> d$i \<le> b$i" using part1 and as(2) by auto  } note * = this
+  thus ?th4 unfolding subset_eq and Ball_def and mem_interval apply auto apply (erule_tac x=ia in allE, simp)+ by (erule_tac x=i in allE, erule_tac x=i in allE, simp)+
+qed
+
+lemma disjoint_interval: fixes a::"real^'n::finite" shows
+  "{a .. b} \<inter> {c .. d} = {} \<longleftrightarrow> (\<exists>i. (b$i < a$i \<or> d$i < c$i \<or> b$i < c$i \<or> d$i < a$i))" (is ?th1) and
+  "{a .. b} \<inter> {c<..<d} = {} \<longleftrightarrow> (\<exists>i. (b$i < a$i \<or> d$i \<le> c$i \<or> b$i \<le> c$i \<or> d$i \<le> a$i))" (is ?th2) and
+  "{a<..<b} \<inter> {c .. d} = {} \<longleftrightarrow> (\<exists>i. (b$i \<le> a$i \<or> d$i < c$i \<or> b$i \<le> c$i \<or> d$i \<le> a$i))" (is ?th3) and
+  "{a<..<b} \<inter> {c<..<d} = {} \<longleftrightarrow> (\<exists>i. (b$i \<le> a$i \<or> d$i \<le> c$i \<or> b$i \<le> c$i \<or> d$i \<le> a$i))" (is ?th4)
+proof-
+  let ?z = "(\<chi> i. ((max (a$i) (c$i)) + (min (b$i) (d$i))) / 2)::real^'n"
+  show ?th1 ?th2 ?th3 ?th4
+  unfolding expand_set_eq and Int_iff and empty_iff and mem_interval and all_conj_distrib[THEN sym] and eq_False
+  apply (auto elim!: allE[where x="?z"])
+  apply ((rule_tac x=x in exI, force) | (rule_tac x=i in exI, force))+
+  done
+qed
+
+lemma inter_interval: fixes a :: "'a::linorder^'n::finite" shows
+ "{a .. b} \<inter> {c .. d} =  {(\<chi> i. max (a$i) (c$i)) .. (\<chi> i. min (b$i) (d$i))}"
+  unfolding expand_set_eq and Int_iff and mem_interval
+  by (auto simp add: less_divide_eq_number_of1 intro!: bexI)
+
+(* Moved interval_open_subset_closed a bit upwards *)
+
+lemma open_interval_lemma: fixes x :: "real" shows
+ "a < x \<Longrightarrow> x < b ==> (\<exists>d>0. \<forall>x'. abs(x' - x) < d --> a < x' \<and> x' < b)"
+  by(rule_tac x="min (x - a) (b - x)" in exI, auto)
+
+lemma open_interval: fixes a :: "real^'n::finite" shows "open {a<..<b}"
+proof-
+  { fix x assume x:"x\<in>{a<..<b}"
+    { fix i
+      have "\<exists>d>0. \<forall>x'. abs (x' - (x$i)) < d \<longrightarrow> a$i < x' \<and> x' < b$i"
+        using x[unfolded mem_interval, THEN spec[where x=i]]
+        using open_interval_lemma[of "a$i" "x$i" "b$i"] by auto  }
+
+    hence "\<forall>i. \<exists>d>0. \<forall>x'. abs (x' - (x$i)) < d \<longrightarrow> a$i < x' \<and> x' < b$i" by auto
+    then obtain d where d:"\<forall>i. 0 < d i \<and> (\<forall>x'. \<bar>x' - x $ i\<bar> < d i \<longrightarrow> a $ i < x' \<and> x' < b $ i)"
+      using bchoice[of "UNIV" "\<lambda>i d. d>0 \<and> (\<forall>x'. \<bar>x' - x $ i\<bar> < d \<longrightarrow> a $ i < x' \<and> x' < b $ i)"] by auto
+
+    let ?d = "Min (range d)"
+    have **:"finite (range d)" "range d \<noteq> {}" by auto
+    have "?d>0" unfolding Min_gr_iff[OF **] using d by auto
+    moreover
+    { fix x' assume as:"dist x' x < ?d"
+      { fix i
+        have "\<bar>x'$i - x $ i\<bar> < d i"
+          using norm_bound_component_lt[OF as[unfolded dist_norm], of i]
+          unfolding vector_minus_component and Min_gr_iff[OF **] by auto
+        hence "a $ i < x' $ i" "x' $ i < b $ i" using d[THEN spec[where x=i]] by auto  }
+      hence "a < x' \<and> x' < b" unfolding vector_less_def by auto  }
+    ultimately have "\<exists>e>0. \<forall>x'. dist x' x < e \<longrightarrow> x' \<in> {a<..<b}" by (auto, rule_tac x="?d" in exI, simp)
+  }
+  thus ?thesis unfolding open_dist using open_interval_lemma by auto
+qed
+
+lemma closed_interval: fixes a :: "real^'n::finite" shows "closed {a .. b}"
+proof-
+  { fix x i assume as:"\<forall>e>0. \<exists>x'\<in>{a..b}. x' \<noteq> x \<and> dist x' x < e"(* and xab:"a$i > x$i \<or> b$i < x$i"*)
+    { assume xa:"a$i > x$i"
+      with as obtain y where y:"y\<in>{a..b}" "y \<noteq> x" "dist y x < a$i - x$i" by(erule_tac x="a$i - x$i" in allE)auto
+      hence False unfolding mem_interval and dist_norm
+        using component_le_norm[of "y-x" i, unfolded vector_minus_component] and xa by(auto elim!: allE[where x=i])
+    } hence "a$i \<le> x$i" by(rule ccontr)auto
+    moreover
+    { assume xb:"b$i < x$i"
+      with as obtain y where y:"y\<in>{a..b}" "y \<noteq> x" "dist y x < x$i - b$i" by(erule_tac x="x$i - b$i" in allE)auto
+      hence False unfolding mem_interval and dist_norm
+        using component_le_norm[of "y-x" i, unfolded vector_minus_component] and xb by(auto elim!: allE[where x=i])
+    } hence "x$i \<le> b$i" by(rule ccontr)auto
+    ultimately
+    have "a $ i \<le> x $ i \<and> x $ i \<le> b $ i" by auto }
+  thus ?thesis unfolding closed_limpt islimpt_approachable mem_interval by auto
+qed
+
+lemma interior_closed_interval: fixes a :: "real^'n::finite" shows
+ "interior {a .. b} = {a<..<b}" (is "?L = ?R")
+proof(rule subset_antisym)
+  show "?R \<subseteq> ?L" using interior_maximal[OF interval_open_subset_closed open_interval] by auto
+next
+  { fix x assume "\<exists>T. open T \<and> x \<in> T \<and> T \<subseteq> {a..b}"
+    then obtain s where s:"open s" "x \<in> s" "s \<subseteq> {a..b}" by auto
+    then obtain e where "e>0" and e:"\<forall>x'. dist x' x < e \<longrightarrow> x' \<in> {a..b}" unfolding open_dist and subset_eq by auto
+    { fix i
+      have "dist (x - (e / 2) *\<^sub>R basis i) x < e"
+           "dist (x + (e / 2) *\<^sub>R basis i) x < e"
+        unfolding dist_norm apply auto
+        unfolding norm_minus_cancel using norm_basis[of i] and `e>0` by auto
+      hence "a $ i \<le> (x - (e / 2) *\<^sub>R basis i) $ i"
+                    "(x + (e / 2) *\<^sub>R basis i) $ i \<le> b $ i"
+        using e[THEN spec[where x="x - (e/2) *\<^sub>R basis i"]]
+        and   e[THEN spec[where x="x + (e/2) *\<^sub>R basis i"]]
+        unfolding mem_interval by (auto elim!: allE[where x=i])
+      hence "a $ i < x $ i" and "x $ i < b $ i"
+        unfolding vector_minus_component and vector_add_component
+        unfolding vector_smult_component and basis_component using `e>0` by auto   }
+    hence "x \<in> {a<..<b}" unfolding mem_interval by auto  }
+  thus "?L \<subseteq> ?R" unfolding interior_def and subset_eq by auto
+qed
+
+lemma bounded_closed_interval: fixes a :: "real^'n::finite" shows
+ "bounded {a .. b}"
+proof-
+  let ?b = "\<Sum>i\<in>UNIV. \<bar>a$i\<bar> + \<bar>b$i\<bar>"
+  { fix x::"real^'n" assume x:"\<forall>i. a $ i \<le> x $ i \<and> x $ i \<le> b $ i"
+    { fix i
+      have "\<bar>x$i\<bar> \<le> \<bar>a$i\<bar> + \<bar>b$i\<bar>" using x[THEN spec[where x=i]] by auto  }
+    hence "(\<Sum>i\<in>UNIV. \<bar>x $ i\<bar>) \<le> ?b" by(rule setsum_mono)
+    hence "norm x \<le> ?b" using norm_le_l1[of x] by auto  }
+  thus ?thesis unfolding interval and bounded_iff by auto
+qed
+
+lemma bounded_interval: fixes a :: "real^'n::finite" shows
+ "bounded {a .. b} \<and> bounded {a<..<b}"
+  using bounded_closed_interval[of a b]
+  using interval_open_subset_closed[of a b]
+  using bounded_subset[of "{a..b}" "{a<..<b}"]
+  by simp
+
+lemma not_interval_univ: fixes a :: "real^'n::finite" shows
+ "({a .. b} \<noteq> UNIV) \<and> ({a<..<b} \<noteq> UNIV)"
+  using bounded_interval[of a b]
+  by auto
+
+lemma compact_interval: fixes a :: "real^'n::finite" shows
+ "compact {a .. b}"
+  using bounded_closed_imp_compact using bounded_interval[of a b] using closed_interval[of a b] by auto
+
+lemma open_interval_midpoint: fixes a :: "real^'n::finite"
+  assumes "{a<..<b} \<noteq> {}" shows "((1/2) *\<^sub>R (a + b)) \<in> {a<..<b}"
+proof-
+  { fix i
+    have "a $ i < ((1 / 2) *\<^sub>R (a + b)) $ i \<and> ((1 / 2) *\<^sub>R (a + b)) $ i < b $ i"
+      using assms[unfolded interval_ne_empty, THEN spec[where x=i]]
+      unfolding vector_smult_component and vector_add_component
+      by(auto simp add: less_divide_eq_number_of1)  }
+  thus ?thesis unfolding mem_interval by auto
+qed
+
+lemma open_closed_interval_convex: fixes x :: "real^'n::finite"
+  assumes x:"x \<in> {a<..<b}" and y:"y \<in> {a .. b}" and e:"0 < e" "e \<le> 1"
+  shows "(e *\<^sub>R x + (1 - e) *\<^sub>R y) \<in> {a<..<b}"
+proof-
+  { fix i
+    have "a $ i = e * a$i + (1 - e) * a$i" unfolding left_diff_distrib by simp
+    also have "\<dots> < e * x $ i + (1 - e) * y $ i" apply(rule add_less_le_mono)
+      using e unfolding mult_less_cancel_left and mult_le_cancel_left apply simp_all
+      using x unfolding mem_interval  apply simp
+      using y unfolding mem_interval  apply simp
+      done
+    finally have "a $ i < (e *\<^sub>R x + (1 - e) *\<^sub>R y) $ i" by auto
+    moreover {
+    have "b $ i = e * b$i + (1 - e) * b$i" unfolding left_diff_distrib by simp
+    also have "\<dots> > e * x $ i + (1 - e) * y $ i" apply(rule add_less_le_mono)
+      using e unfolding mult_less_cancel_left and mult_le_cancel_left apply simp_all
+      using x unfolding mem_interval  apply simp
+      using y unfolding mem_interval  apply simp
+      done
+    finally have "(e *\<^sub>R x + (1 - e) *\<^sub>R y) $ i < b $ i" by auto
+    } ultimately have "a $ i < (e *\<^sub>R x + (1 - e) *\<^sub>R y) $ i \<and> (e *\<^sub>R x + (1 - e) *\<^sub>R y) $ i < b $ i" by auto }
+  thus ?thesis unfolding mem_interval by auto
+qed
+
+lemma closure_open_interval: fixes a :: "real^'n::finite"
+  assumes "{a<..<b} \<noteq> {}"
+  shows "closure {a<..<b} = {a .. b}"
+proof-
+  have ab:"a < b" using assms[unfolded interval_ne_empty] unfolding vector_less_def by auto
+  let ?c = "(1 / 2) *\<^sub>R (a + b)"
+  { fix x assume as:"x \<in> {a .. b}"
+    def f == "\<lambda>n::nat. x + (inverse (real n + 1)) *\<^sub>R (?c - x)"
+    { fix n assume fn:"f n < b \<longrightarrow> a < f n \<longrightarrow> f n = x" and xc:"x \<noteq> ?c"
+      have *:"0 < inverse (real n + 1)" "inverse (real n + 1) \<le> 1" unfolding inverse_le_1_iff by auto
+      have "(inverse (real n + 1)) *\<^sub>R ((1 / 2) *\<^sub>R (a + b)) + (1 - inverse (real n + 1)) *\<^sub>R x =
+        x + (inverse (real n + 1)) *\<^sub>R (((1 / 2) *\<^sub>R (a + b)) - x)"
+        by (auto simp add: algebra_simps)
+      hence "f n < b" and "a < f n" using open_closed_interval_convex[OF open_interval_midpoint[OF assms] as *] unfolding f_def by auto
+      hence False using fn unfolding f_def using xc by(auto simp add: vector_mul_lcancel vector_ssub_ldistrib)  }
+    moreover
+    { assume "\<not> (f ---> x) sequentially"
+      { fix e::real assume "e>0"
+        hence "\<exists>N::nat. inverse (real (N + 1)) < e" using real_arch_inv[of e] apply (auto simp add: Suc_pred') apply(rule_tac x="n - 1" in exI) by auto
+        then obtain N::nat where "inverse (real (N + 1)) < e" by auto
+        hence "\<forall>n\<ge>N. inverse (real n + 1) < e" by (auto, metis Suc_le_mono le_SucE less_imp_inverse_less nat_le_real_less order_less_trans real_of_nat_Suc real_of_nat_Suc_gt_zero)
+        hence "\<exists>N::nat. \<forall>n\<ge>N. inverse (real n + 1) < e" by auto  }
+      hence "((\<lambda>n. inverse (real n + 1)) ---> 0) sequentially"
+        unfolding Lim_sequentially by(auto simp add: dist_norm)
+      hence "(f ---> x) sequentially" unfolding f_def
+        using Lim_add[OF Lim_const, of "\<lambda>n::nat. (inverse (real n + 1)) *\<^sub>R ((1 / 2) *\<^sub>R (a + b) - x)" 0 sequentially x]
+        using Lim_vmul[of "\<lambda>n::nat. inverse (real n + 1)" 0 sequentially "((1 / 2) *\<^sub>R (a + b) - x)"] by auto  }
+    ultimately have "x \<in> closure {a<..<b}"
+      using as and open_interval_midpoint[OF assms] unfolding closure_def unfolding islimpt_sequential by(cases "x=?c")auto  }
+  thus ?thesis using closure_minimal[OF interval_open_subset_closed closed_interval, of a b] by blast
+qed
+
+lemma bounded_subset_open_interval_symmetric: fixes s::"(real^'n::finite) set"
+  assumes "bounded s"  shows "\<exists>a. s \<subseteq> {-a<..<a}"
+proof-
+  obtain b where "b>0" and b:"\<forall>x\<in>s. norm x \<le> b" using assms[unfolded bounded_pos] by auto
+  def a \<equiv> "(\<chi> i. b+1)::real^'n"
+  { fix x assume "x\<in>s"
+    fix i
+    have "(-a)$i < x$i" and "x$i < a$i" using b[THEN bspec[where x=x], OF `x\<in>s`] and component_le_norm[of x i]
+      unfolding vector_uminus_component and a_def and Cart_lambda_beta by auto
+  }
+  thus ?thesis by(auto intro: exI[where x=a] simp add: vector_less_def)
+qed
+
+lemma bounded_subset_open_interval:
+  fixes s :: "(real ^ 'n::finite) set"
+  shows "bounded s ==> (\<exists>a b. s \<subseteq> {a<..<b})"
+  by (auto dest!: bounded_subset_open_interval_symmetric)
+
+lemma bounded_subset_closed_interval_symmetric:
+  fixes s :: "(real ^ 'n::finite) set"
+  assumes "bounded s" shows "\<exists>a. s \<subseteq> {-a .. a}"
+proof-
+  obtain a where "s \<subseteq> {- a<..<a}" using bounded_subset_open_interval_symmetric[OF assms] by auto
+  thus ?thesis using interval_open_subset_closed[of "-a" a] by auto
+qed
+
+lemma bounded_subset_closed_interval:
+  fixes s :: "(real ^ 'n::finite) set"
+  shows "bounded s ==> (\<exists>a b. s \<subseteq> {a .. b})"
+  using bounded_subset_closed_interval_symmetric[of s] by auto
+
+lemma frontier_closed_interval:
+  fixes a b :: "real ^ _"
+  shows "frontier {a .. b} = {a .. b} - {a<..<b}"
+  unfolding frontier_def unfolding interior_closed_interval and closure_closed[OF closed_interval] ..
+
+lemma frontier_open_interval:
+  fixes a b :: "real ^ _"
+  shows "frontier {a<..<b} = (if {a<..<b} = {} then {} else {a .. b} - {a<..<b})"
+proof(cases "{a<..<b} = {}")
+  case True thus ?thesis using frontier_empty by auto
+next
+  case False thus ?thesis unfolding frontier_def and closure_open_interval[OF False] and interior_open[OF open_interval] by auto
+qed
+
+lemma inter_interval_mixed_eq_empty: fixes a :: "real^'n::finite"
+  assumes "{c<..<d} \<noteq> {}"  shows "{a<..<b} \<inter> {c .. d} = {} \<longleftrightarrow> {a<..<b} \<inter> {c<..<d} = {}"
+  unfolding closure_open_interval[OF assms, THEN sym] unfolding open_inter_closure_eq_empty[OF open_interval] ..
+
+
+(* Some special cases for intervals in R^1.                                  *)
+
+lemma all_1: "(\<forall>x::1. P x) \<longleftrightarrow> P 1"
+  by (metis num1_eq_iff)
+
+lemma ex_1: "(\<exists>x::1. P x) \<longleftrightarrow> P 1"
+  by auto (metis num1_eq_iff)
+
+lemma interval_cases_1: fixes x :: "real^1" shows
+ "x \<in> {a .. b} ==> x \<in> {a<..<b} \<or> (x = a) \<or> (x = b)"
+  by(simp add:  Cart_eq vector_less_def vector_less_eq_def all_1, auto)
+
+lemma in_interval_1: fixes x :: "real^1" shows
+ "(x \<in> {a .. b} \<longleftrightarrow> dest_vec1 a \<le> dest_vec1 x \<and> dest_vec1 x \<le> dest_vec1 b) \<and>
+  (x \<in> {a<..<b} \<longleftrightarrow> dest_vec1 a < dest_vec1 x \<and> dest_vec1 x < dest_vec1 b)"
+by(simp add: Cart_eq vector_less_def vector_less_eq_def all_1 dest_vec1_def)
+
+lemma interval_eq_empty_1: fixes a :: "real^1" shows
+  "{a .. b} = {} \<longleftrightarrow> dest_vec1 b < dest_vec1 a"
+  "{a<..<b} = {} \<longleftrightarrow> dest_vec1 b \<le> dest_vec1 a"
+  unfolding interval_eq_empty and ex_1 and dest_vec1_def by auto
+
+lemma subset_interval_1: fixes a :: "real^1" shows
+ "({a .. b} \<subseteq> {c .. d} \<longleftrightarrow>  dest_vec1 b < dest_vec1 a \<or>
+                dest_vec1 c \<le> dest_vec1 a \<and> dest_vec1 a \<le> dest_vec1 b \<and> dest_vec1 b \<le> dest_vec1 d)"
+ "({a .. b} \<subseteq> {c<..<d} \<longleftrightarrow>  dest_vec1 b < dest_vec1 a \<or>
+                dest_vec1 c < dest_vec1 a \<and> dest_vec1 a \<le> dest_vec1 b \<and> dest_vec1 b < dest_vec1 d)"
+ "({a<..<b} \<subseteq> {c .. d} \<longleftrightarrow>  dest_vec1 b \<le> dest_vec1 a \<or>
+                dest_vec1 c \<le> dest_vec1 a \<and> dest_vec1 a < dest_vec1 b \<and> dest_vec1 b \<le> dest_vec1 d)"
+ "({a<..<b} \<subseteq> {c<..<d} \<longleftrightarrow> dest_vec1 b \<le> dest_vec1 a \<or>
+                dest_vec1 c \<le> dest_vec1 a \<and> dest_vec1 a < dest_vec1 b \<and> dest_vec1 b \<le> dest_vec1 d)"
+  unfolding subset_interval[of a b c d] unfolding all_1 and dest_vec1_def by auto
+
+lemma eq_interval_1: fixes a :: "real^1" shows
+ "{a .. b} = {c .. d} \<longleftrightarrow>
+          dest_vec1 b < dest_vec1 a \<and> dest_vec1 d < dest_vec1 c \<or>
+          dest_vec1 a = dest_vec1 c \<and> dest_vec1 b = dest_vec1 d"
+using set_eq_subset[of "{a .. b}" "{c .. d}"]
+using subset_interval_1(1)[of a b c d]
+using subset_interval_1(1)[of c d a b]
+by auto (* FIXME: slow *)
+
+lemma disjoint_interval_1: fixes a :: "real^1" shows
+  "{a .. b} \<inter> {c .. d} = {} \<longleftrightarrow> dest_vec1 b < dest_vec1 a \<or> dest_vec1 d < dest_vec1 c  \<or>  dest_vec1 b < dest_vec1 c \<or> dest_vec1 d < dest_vec1 a"
+  "{a .. b} \<inter> {c<..<d} = {} \<longleftrightarrow> dest_vec1 b < dest_vec1 a \<or> dest_vec1 d \<le> dest_vec1 c  \<or>  dest_vec1 b \<le> dest_vec1 c \<or> dest_vec1 d \<le> dest_vec1 a"
+  "{a<..<b} \<inter> {c .. d} = {} \<longleftrightarrow> dest_vec1 b \<le> dest_vec1 a \<or> dest_vec1 d < dest_vec1 c  \<or>  dest_vec1 b \<le> dest_vec1 c \<or> dest_vec1 d \<le> dest_vec1 a"
+  "{a<..<b} \<inter> {c<..<d} = {} \<longleftrightarrow> dest_vec1 b \<le> dest_vec1 a \<or> dest_vec1 d \<le> dest_vec1 c  \<or>  dest_vec1 b \<le> dest_vec1 c \<or> dest_vec1 d \<le> dest_vec1 a"
+  unfolding disjoint_interval and dest_vec1_def ex_1 by auto
+
+lemma open_closed_interval_1: fixes a :: "real^1" shows
+ "{a<..<b} = {a .. b} - {a, b}"
+  unfolding expand_set_eq apply simp unfolding vector_less_def and vector_less_eq_def and all_1 and dest_vec1_eq[THEN sym] and dest_vec1_def by auto
+
+lemma closed_open_interval_1: "dest_vec1 (a::real^1) \<le> dest_vec1 b ==> {a .. b} = {a<..<b} \<union> {a,b}"
+  unfolding expand_set_eq apply simp unfolding vector_less_def and vector_less_eq_def and all_1 and dest_vec1_eq[THEN sym] and dest_vec1_def by auto
+
+(* Some stuff for half-infinite intervals too; FIXME: notation?  *)
+
+lemma closed_interval_left: fixes b::"real^'n::finite"
+  shows "closed {x::real^'n. \<forall>i. x$i \<le> b$i}"
+proof-
+  { fix i
+    fix x::"real^'n" assume x:"\<forall>e>0. \<exists>x'\<in>{x. \<forall>i. x $ i \<le> b $ i}. x' \<noteq> x \<and> dist x' x < e"
+    { assume "x$i > b$i"
+      then obtain y where "y $ i \<le> b $ i"  "y \<noteq> x"  "dist y x < x$i - b$i" using x[THEN spec[where x="x$i - b$i"]] by auto
+      hence False using component_le_norm[of "y - x" i] unfolding dist_norm and vector_minus_component by auto   }
+    hence "x$i \<le> b$i" by(rule ccontr)auto  }
+  thus ?thesis unfolding closed_limpt unfolding islimpt_approachable by blast
+qed
+
+lemma closed_interval_right: fixes a::"real^'n::finite"
+  shows "closed {x::real^'n. \<forall>i. a$i \<le> x$i}"
+proof-
+  { fix i
+    fix x::"real^'n" assume x:"\<forall>e>0. \<exists>x'\<in>{x. \<forall>i. a $ i \<le> x $ i}. x' \<noteq> x \<and> dist x' x < e"
+    { assume "a$i > x$i"
+      then obtain y where "a $ i \<le> y $ i"  "y \<noteq> x"  "dist y x < a$i - x$i" using x[THEN spec[where x="a$i - x$i"]] by auto
+      hence False using component_le_norm[of "y - x" i] unfolding dist_norm and vector_minus_component by auto   }
+    hence "a$i \<le> x$i" by(rule ccontr)auto  }
+  thus ?thesis unfolding closed_limpt unfolding islimpt_approachable by blast
+qed
+
+subsection{* Intervals in general, including infinite and mixtures of open and closed. *}
+
+definition "is_interval s \<longleftrightarrow> (\<forall>a\<in>s. \<forall>b\<in>s. \<forall>x. (\<forall>i. ((a$i \<le> x$i \<and> x$i \<le> b$i) \<or> (b$i \<le> x$i \<and> x$i \<le> a$i)))  \<longrightarrow> x \<in> s)"
+
+lemma is_interval_interval: "is_interval {a .. b::real^'n::finite}" (is ?th1) "is_interval {a<..<b}" (is ?th2) proof - 
+  have *:"\<And>x y z::real. x < y \<Longrightarrow> y < z \<Longrightarrow> 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 {}"
+  unfolding is_interval_def
+  by simp
+
+lemma is_interval_univ:
+ "is_interval UNIV"
+  unfolding is_interval_def
+  by simp
+
+subsection{* Closure of halfspaces and hyperplanes.                                    *}
+
+lemma Lim_inner:
+  assumes "(f ---> l) net"  shows "((\<lambda>y. inner a (f y)) ---> inner a l) net"
+  by (intro tendsto_intros assms)
+
+lemma continuous_at_inner: "continuous (at x) (inner a)"
+  unfolding continuous_at by (intro tendsto_intros)
+
+lemma continuous_on_inner:
+  fixes s :: "'a::real_inner set"
+  shows "continuous_on s (inner a)"
+  unfolding continuous_on by (rule ballI) (intro tendsto_intros)
+
+lemma closed_halfspace_le: "closed {x. inner a x \<le> b}"
+proof-
+  have "\<forall>x. continuous (at x) (inner a)"
+    unfolding continuous_at by (rule allI) (intro tendsto_intros)
+  hence "closed (inner a -` {..b})"
+    using closed_real_atMost by (rule continuous_closed_vimage)
+  moreover have "{x. inner a x \<le> b} = inner a -` {..b}" by auto
+  ultimately show ?thesis by simp
+qed
+
+lemma closed_halfspace_ge: "closed {x. inner a x \<ge> b}"
+  using closed_halfspace_le[of "-a" "-b"] unfolding inner_minus_left by auto
+
+lemma closed_hyperplane: "closed {x. inner a x = b}"
+proof-
+  have "{x. inner a x = b} = {x. inner a x \<ge> b} \<inter> {x. inner a x \<le> b}" by auto
+  thus ?thesis using closed_halfspace_le[of a b] and closed_halfspace_ge[of b a] using closed_Int by auto
+qed
+
+lemma closed_halfspace_component_le:
+  shows "closed {x::real^'n::finite. x$i \<le> a}"
+  using closed_halfspace_le[of "(basis i)::real^'n" a] unfolding inner_basis[OF assms] by auto
+
+lemma closed_halfspace_component_ge:
+  shows "closed {x::real^'n::finite. x$i \<ge> a}"
+  using closed_halfspace_ge[of a "(basis i)::real^'n"] unfolding inner_basis[OF assms] by auto
+
+text{* Openness of halfspaces.                                                   *}
+
+lemma open_halfspace_lt: "open {x. inner a x < b}"
+proof-
+  have "UNIV - {x. b \<le> inner a x} = {x. inner a x < b}" by auto
+  thus ?thesis using closed_halfspace_ge[unfolded closed_def Compl_eq_Diff_UNIV, of b a] by auto
+qed
+
+lemma open_halfspace_gt: "open {x. inner a x > b}"
+proof-
+  have "UNIV - {x. b \<ge> inner a x} = {x. inner a x > b}" by auto
+  thus ?thesis using closed_halfspace_le[unfolded closed_def Compl_eq_Diff_UNIV, of a b] by auto
+qed
+
+lemma open_halfspace_component_lt:
+  shows "open {x::real^'n::finite. x$i < a}"
+  using open_halfspace_lt[of "(basis i)::real^'n" a] unfolding inner_basis[OF assms] by auto
+
+lemma open_halfspace_component_gt:
+  shows "open {x::real^'n::finite. x$i  > a}"
+  using open_halfspace_gt[of a "(basis i)::real^'n"] unfolding inner_basis[OF assms] by auto
+
+text{* This gives a simple derivation of limit component bounds.                 *}
+
+lemma Lim_component_le: fixes f :: "'a \<Rightarrow> real^'n::finite"
+  assumes "(f ---> l) net" "\<not> (trivial_limit net)"  "eventually (\<lambda>x. f(x)$i \<le> b) net"
+  shows "l$i \<le> b"
+proof-
+  { fix x have "x \<in> {x::real^'n. inner (basis i) x \<le> b} \<longleftrightarrow> x$i \<le> b" unfolding inner_basis by auto } note * = this
+  show ?thesis using Lim_in_closed_set[of "{x. inner (basis i) x \<le> b}" f net l] unfolding *
+    using closed_halfspace_le[of "(basis i)::real^'n" b] and assms(1,2,3) by auto
+qed
+
+lemma Lim_component_ge: fixes f :: "'a \<Rightarrow> real^'n::finite"
+  assumes "(f ---> l) net"  "\<not> (trivial_limit net)"  "eventually (\<lambda>x. b \<le> (f x)$i) net"
+  shows "b \<le> l$i"
+proof-
+  { fix x have "x \<in> {x::real^'n. inner (basis i) x \<ge> b} \<longleftrightarrow> x$i \<ge> b" unfolding inner_basis by auto } note * = this
+  show ?thesis using Lim_in_closed_set[of "{x. inner (basis i) x \<ge> b}" f net l] unfolding *
+    using closed_halfspace_ge[of b "(basis i)::real^'n"] and assms(1,2,3) by auto
+qed
+
+lemma Lim_component_eq: fixes f :: "'a \<Rightarrow> real^'n::finite"
+  assumes net:"(f ---> l) net" "~(trivial_limit net)" and ev:"eventually (\<lambda>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
+
+lemma Lim_drop_le: fixes f :: "'a \<Rightarrow> real^1" shows
+  "(f ---> l) net \<Longrightarrow> ~(trivial_limit net) \<Longrightarrow> eventually (\<lambda>x. dest_vec1 (f x) \<le> b) net ==> dest_vec1 l \<le> b"
+  using Lim_component_le[of f l net 1 b] unfolding dest_vec1_def by auto
+
+lemma Lim_drop_ge: fixes f :: "'a \<Rightarrow> real^1" shows
+ "(f ---> l) net \<Longrightarrow> ~(trivial_limit net) \<Longrightarrow> eventually (\<lambda>x. b \<le> dest_vec1 (f x)) net ==> b \<le> dest_vec1 l"
+  using Lim_component_ge[of f l net b 1] unfolding dest_vec1_def by auto
+
+text{* Limits relative to a union.                                               *}
+
+lemma eventually_within_Un:
+  "eventually P (net within (s \<union> t)) \<longleftrightarrow>
+    eventually P (net within s) \<and> eventually P (net within t)"
+  unfolding Limits.eventually_within
+  by (auto elim!: eventually_rev_mp)
+
+lemma Lim_within_union:
+ "(f ---> l) (net within (s \<union> t)) \<longleftrightarrow>
+  (f ---> l) (net within s) \<and> (f ---> l) (net within t)"
+  unfolding tendsto_def
+  by (auto simp add: eventually_within_Un)
+
+lemma continuous_on_union:
+  assumes "closed s" "closed t" "continuous_on s f" "continuous_on t f"
+  shows "continuous_on (s \<union> t) f"
+  using assms unfolding continuous_on unfolding Lim_within_union
+  unfolding Lim unfolding trivial_limit_within unfolding closed_limpt by auto
+
+lemma continuous_on_cases:
+  assumes "closed s" "closed t" "continuous_on s f" "continuous_on t g"
+          "\<forall>x. (x\<in>s \<and> \<not> P x) \<or> (x \<in> t \<and> P x) \<longrightarrow> f x = g x"
+  shows "continuous_on (s \<union> t) (\<lambda>x. if P x then f x else g x)"
+proof-
+  let ?h = "(\<lambda>x. if P x then f x else g x)"
+  have "\<forall>x\<in>s. f x = (if P x then f x else g x)" using assms(5) by auto
+  hence "continuous_on s ?h" using continuous_on_eq[of s f ?h] using assms(3) by auto
+  moreover
+  have "\<forall>x\<in>t. g x = (if P x then f x else g x)" using assms(5) by auto
+  hence "continuous_on t ?h" using continuous_on_eq[of t g ?h] using assms(4) by auto
+  ultimately show ?thesis using continuous_on_union[OF assms(1,2), of ?h] by auto
+qed
+
+
+text{* Some more convenient intermediate-value theorem formulations.             *}
+
+lemma connected_ivt_hyperplane:
+  assumes "connected s" "x \<in> s" "y \<in> s" "inner a x \<le> b" "b \<le> inner a y"
+  shows "\<exists>z \<in> s. inner a z = b"
+proof(rule ccontr)
+  assume as:"\<not> (\<exists>z\<in>s. inner a z = b)"
+  let ?A = "{x. inner a x < b}"
+  let ?B = "{x. inner a x > b}"
+  have "open ?A" "open ?B" using open_halfspace_lt and open_halfspace_gt by auto
+  moreover have "?A \<inter> ?B = {}" by auto
+  moreover have "s \<subseteq> ?A \<union> ?B" using as by auto
+  ultimately show False using assms(1)[unfolded connected_def not_ex, THEN spec[where x="?A"], THEN spec[where x="?B"]] and assms(2-5) by auto
+qed
+
+lemma connected_ivt_component: fixes x::"real^'n::finite" shows
+ "connected s \<Longrightarrow> x \<in> s \<Longrightarrow> y \<in> s \<Longrightarrow> x$k \<le> a \<Longrightarrow> a \<le> y$k \<Longrightarrow> (\<exists>z\<in>s.  z$k = a)"
+  using connected_ivt_hyperplane[of s x y "(basis k)::real^'n" a] by (auto simp add: inner_basis)
+
+text{* Also more convenient formulations of monotone convergence.                *}
+
+lemma bounded_increasing_convergent: fixes s::"nat \<Rightarrow> real^1"
+  assumes "bounded {s n| n::nat. True}"  "\<forall>n. dest_vec1(s n) \<le> dest_vec1(s(Suc n))"
+  shows "\<exists>l. (s ---> l) sequentially"
+proof-
+  obtain a where a:"\<forall>n. \<bar>dest_vec1 (s n)\<bar> \<le>  a" using assms(1)[unfolded bounded_iff abs_dest_vec1] by auto
+  { fix m::nat
+    have "\<And> n. n\<ge>m \<longrightarrow> dest_vec1 (s m) \<le> dest_vec1 (s n)"
+      apply(induct_tac n) apply simp using assms(2) apply(erule_tac x="na" in allE) by(auto simp add: not_less_eq_eq)  }
+  hence "\<forall>m n. m \<le> n \<longrightarrow> dest_vec1 (s m) \<le> dest_vec1 (s n)" by auto
+  then obtain l where "\<forall>e>0. \<exists>N. \<forall>n\<ge>N. \<bar>dest_vec1 (s n) - l\<bar> < e" using convergent_bounded_monotone[OF a] unfolding monoseq_def by auto
+  thus ?thesis unfolding Lim_sequentially apply(rule_tac x="vec1 l" in exI)
+    unfolding dist_norm unfolding abs_dest_vec1 and dest_vec1_sub by auto
+qed
+
+subsection{* Basic homeomorphism definitions.                                          *}
+
+definition "homeomorphism s t f g \<equiv>
+     (\<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::metric_space set \<Rightarrow> 'b::metric_space set \<Rightarrow> bool"
+    (infixr "homeomorphic" 60) where
+  homeomorphic_def: "s homeomorphic t \<equiv> (\<exists>f g. homeomorphism s t f g)"
+
+lemma homeomorphic_refl: "s homeomorphic s"
+  unfolding homeomorphic_def
+  unfolding homeomorphism_def
+  using continuous_on_id
+  apply(rule_tac x = "(\<lambda>x. x)" in exI)
+  apply(rule_tac x = "(\<lambda>x. x)" in exI)
+  by blast
+
+lemma homeomorphic_sym:
+ "s homeomorphic t \<longleftrightarrow> t homeomorphic s"
+unfolding homeomorphic_def
+unfolding homeomorphism_def
+by blast (* FIXME: slow *)
+
+lemma homeomorphic_trans:
+  assumes "s homeomorphic t" "t homeomorphic u" shows "s homeomorphic u"
+proof-
+  obtain f1 g1 where fg1:"\<forall>x\<in>s. g1 (f1 x) = x"  "f1 ` s = t" "continuous_on s f1" "\<forall>y\<in>t. f1 (g1 y) = y" "g1 ` t = s" "continuous_on t g1"
+    using assms(1) unfolding homeomorphic_def homeomorphism_def by auto
+  obtain f2 g2 where fg2:"\<forall>x\<in>t. g2 (f2 x) = x"  "f2 ` t = u" "continuous_on t f2" "\<forall>y\<in>u. f2 (g2 y) = y" "g2 ` u = t" "continuous_on u g2"
+    using assms(2) unfolding homeomorphic_def homeomorphism_def by auto
+
+  { fix x assume "x\<in>s" hence "(g1 \<circ> g2) ((f2 \<circ> f1) x) = x" using fg1(1)[THEN bspec[where x=x]] and fg2(1)[THEN bspec[where x="f1 x"]] and fg1(2) by auto }
+  moreover have "(f2 \<circ> f1) ` s = u" using fg1(2) fg2(2) by auto
+  moreover have "continuous_on s (f2 \<circ> f1)" using continuous_on_compose[OF fg1(3)] and fg2(3) unfolding fg1(2) by auto
+  moreover { fix y assume "y\<in>u" hence "(f2 \<circ> f1) ((g1 \<circ> g2) y) = y" using fg2(4)[THEN bspec[where x=y]] and fg1(4)[THEN bspec[where x="g2 y"]] and fg2(5) by auto }
+  moreover have "(g1 \<circ> g2) ` u = s" using fg1(5) fg2(5) by auto
+  moreover have "continuous_on u (g1 \<circ> g2)" using continuous_on_compose[OF fg2(6)] and fg1(6)  unfolding fg2(5) by auto
+  ultimately show ?thesis unfolding homeomorphic_def homeomorphism_def apply(rule_tac x="f2 \<circ> f1" in exI) apply(rule_tac x="g1 \<circ> g2" in exI) by auto
+qed
+
+lemma homeomorphic_minimal:
+ "s homeomorphic t \<longleftrightarrow>
+    (\<exists>f g. (\<forall>x\<in>s. f(x) \<in> t \<and> (g(f(x)) = x)) \<and>
+           (\<forall>y\<in>t. g(y) \<in> s \<and> (f(g(y)) = y)) \<and>
+           continuous_on s f \<and> continuous_on t g)"
+unfolding homeomorphic_def homeomorphism_def
+apply auto apply (rule_tac x=f in exI) apply (rule_tac x=g in exI)
+apply auto apply (rule_tac x=f in exI) apply (rule_tac x=g in exI) apply auto
+unfolding image_iff
+apply(erule_tac x="g x" in ballE) apply(erule_tac x="x" in ballE)
+apply auto apply(rule_tac x="g x" in bexI) apply auto
+apply(erule_tac x="f x" in ballE) apply(erule_tac x="x" in ballE)
+apply auto apply(rule_tac x="f x" in bexI) by auto
+
+subsection{* Relatively weak hypotheses if a set is compact.                           *}
+
+definition "inv_on f s = (\<lambda>x. SOME y. y\<in>s \<and> f y = x)"
+
+lemma assumes "inj_on f s" "x\<in>s"
+  shows "inv_on f s (f x) = x"
+ using assms unfolding inj_on_def inv_on_def by auto
+
+lemma homeomorphism_compact:
+  fixes f :: "'a::heine_borel \<Rightarrow> 'b::heine_borel"
+    (* class constraint due to continuous_on_inverse *)
+  assumes "compact s" "continuous_on s f"  "f ` s = t"  "inj_on f s"
+  shows "\<exists>g. homeomorphism s t f g"
+proof-
+  def g \<equiv> "\<lambda>x. SOME y. y\<in>s \<and> f y = x"
+  have g:"\<forall>x\<in>s. g (f x) = x" using assms(3) assms(4)[unfolded inj_on_def] unfolding g_def by auto
+  { fix y assume "y\<in>t"
+    then obtain x where x:"f x = y" "x\<in>s" using assms(3) by auto
+    hence "g (f x) = x" using g by auto
+    hence "f (g y) = y" unfolding x(1)[THEN sym] by auto  }
+  hence g':"\<forall>x\<in>t. f (g x) = x" by auto
+  moreover
+  { fix x
+    have "x\<in>s \<Longrightarrow> x \<in> g ` t" using g[THEN bspec[where x=x]] unfolding image_iff using assms(3) by(auto intro!: bexI[where x="f x"])
+    moreover
+    { assume "x\<in>g ` t"
+      then obtain y where y:"y\<in>t" "g y = x" by auto
+      then obtain x' where x':"x'\<in>s" "f x' = y" using assms(3) by auto
+      hence "x \<in> s" unfolding g_def using someI2[of "\<lambda>b. b\<in>s \<and> f b = y" x' "\<lambda>x. x\<in>s"] unfolding y(2)[THEN sym] and g_def by auto }
+    ultimately have "x\<in>s \<longleftrightarrow> x \<in> g ` t" by auto  }
+  hence "g ` t = s" by auto
+  ultimately
+  show ?thesis unfolding homeomorphism_def homeomorphic_def
+    apply(rule_tac x=g in exI) using g and assms(3) and continuous_on_inverse[OF assms(2,1), of g, unfolded assms(3)] and assms(2) by auto
+qed
+
+lemma homeomorphic_compact:
+  fixes f :: "'a::heine_borel \<Rightarrow> 'b::heine_borel"
+    (* class constraint due to continuous_on_inverse *)
+  shows "compact s \<Longrightarrow> continuous_on s f \<Longrightarrow> (f ` s = t) \<Longrightarrow> inj_on f s
+          \<Longrightarrow> s homeomorphic t"
+  unfolding homeomorphic_def by(metis homeomorphism_compact)
+
+text{* Preservation of topological properties.                                   *}
+
+lemma homeomorphic_compactness:
+ "s homeomorphic t ==> (compact s \<longleftrightarrow> compact t)"
+unfolding homeomorphic_def homeomorphism_def
+by (metis compact_continuous_image)
+
+text{* Results on translation, scaling etc.                                      *}
+
+lemma homeomorphic_scaling:
+  fixes s :: "'a::real_normed_vector set"
+  assumes "c \<noteq> 0"  shows "s homeomorphic ((\<lambda>x. c *\<^sub>R x) ` s)"
+  unfolding homeomorphic_minimal
+  apply(rule_tac x="\<lambda>x. c *\<^sub>R x" in exI)
+  apply(rule_tac x="\<lambda>x. (1 / c) *\<^sub>R x" in exI)
+  using assms apply auto
+  using continuous_on_cmul[OF continuous_on_id] by auto
+
+lemma homeomorphic_translation:
+  fixes s :: "'a::real_normed_vector set"
+  shows "s homeomorphic ((\<lambda>x. a + x) ` s)"
+  unfolding homeomorphic_minimal
+  apply(rule_tac x="\<lambda>x. a + x" in exI)
+  apply(rule_tac x="\<lambda>x. -a + x" in exI)
+  using continuous_on_add[OF continuous_on_const continuous_on_id] by auto
+
+lemma homeomorphic_affinity:
+  fixes s :: "'a::real_normed_vector set"
+  assumes "c \<noteq> 0"  shows "s homeomorphic ((\<lambda>x. a + c *\<^sub>R x) ` s)"
+proof-
+  have *:"op + a ` op *\<^sub>R c ` s = (\<lambda>x. a + c *\<^sub>R x) ` s" by auto
+  show ?thesis
+    using homeomorphic_trans
+    using homeomorphic_scaling[OF assms, of s]
+    using homeomorphic_translation[of "(\<lambda>x. c *\<^sub>R x) ` s" a] unfolding * by auto
+qed
+
+lemma homeomorphic_balls:
+  fixes a b ::"'a::real_normed_vector" (* FIXME: generalize to metric_space *)
+  assumes "0 < d"  "0 < e"
+  shows "(ball a d) homeomorphic  (ball b e)" (is ?th)
+        "(cball a d) homeomorphic (cball b e)" (is ?cth)
+proof-
+  have *:"\<bar>e / d\<bar> > 0" "\<bar>d / e\<bar> >0" using assms using divide_pos_pos by auto
+  show ?th unfolding homeomorphic_minimal
+    apply(rule_tac x="\<lambda>x. b + (e/d) *\<^sub>R (x - a)" in exI)
+    apply(rule_tac x="\<lambda>x. a + (d/e) *\<^sub>R (x - b)" in exI)
+    using assms apply (auto simp add: dist_commute)
+    unfolding dist_norm
+    apply (auto simp add: pos_divide_less_eq mult_strict_left_mono)
+    unfolding continuous_on
+    by (intro ballI tendsto_intros, simp, assumption)+
+next
+  have *:"\<bar>e / d\<bar> > 0" "\<bar>d / e\<bar> >0" using assms using divide_pos_pos by auto
+  show ?cth unfolding homeomorphic_minimal
+    apply(rule_tac x="\<lambda>x. b + (e/d) *\<^sub>R (x - a)" in exI)
+    apply(rule_tac x="\<lambda>x. a + (d/e) *\<^sub>R (x - b)" in exI)
+    using assms apply (auto simp add: dist_commute)
+    unfolding dist_norm
+    apply (auto simp add: pos_divide_le_eq)
+    unfolding continuous_on
+    by (intro ballI tendsto_intros, simp, assumption)+
+qed
+
+text{* "Isometry" (up to constant bounds) of injective linear map etc.           *}
+
+lemma cauchy_isometric:
+  fixes x :: "nat \<Rightarrow> real ^ 'n::finite"
+  assumes e:"0 < e" 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)"
+  shows "Cauchy x"
+proof-
+  interpret f: bounded_linear f by fact
+  { fix d::real assume "d>0"
+    then obtain N where N:"\<forall>n\<ge>N. norm (f (x n) - f (x N)) < e * d"
+      using cf[unfolded cauchy o_def dist_norm, THEN spec[where x="e*d"]] and e and mult_pos_pos[of e d] by auto
+    { fix n assume "n\<ge>N"
+      hence "norm (f (x n - x N)) < e * d" using N[THEN spec[where x=n]] unfolding f.diff[THEN sym] by auto
+      moreover have "e * norm (x n - x N) \<le> norm (f (x n - x N))"
+        using subspace_sub[OF s, of "x n" "x N"] using xs[THEN spec[where x=N]] and xs[THEN spec[where x=n]]
+        using normf[THEN bspec[where x="x n - x N"]] by auto
+      ultimately have "norm (x n - x N) < d" using `e>0`
+        using mult_left_less_imp_less[of e "norm (x n - x N)" d] by auto   }
+    hence "\<exists>N. \<forall>n\<ge>N. norm (x n - x N) < d" by auto }
+  thus ?thesis unfolding cauchy and dist_norm by auto
+qed
+
+lemma complete_isometric_image:
+  fixes f :: "real ^ _ \<Rightarrow> real ^ _"
+  assumes "0 < e" and s:"subspace s" and f:"bounded_linear f" and normf:"\<forall>x\<in>s. norm(f x) \<ge> e * norm(x)" and cs:"complete s"
+  shows "complete(f ` s)"
+proof-
+  { 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)" unfolding image_iff and Bex_def
+      using choice[of "\<lambda> n xa. xa \<in> s \<and> g n = f xa"] by auto
+    hence x:"\<forall>n. x n \<in> s"  "\<forall>n. g n = f (x n)" by auto
+    hence "f \<circ> x = g" unfolding expand_fun_eq 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
+    hence "\<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  }
+  thus ?thesis unfolding complete_def by auto
+qed
+
+lemma dist_0_norm:
+  fixes x :: "'a::real_normed_vector"
+  shows "dist 0 x = norm x"
+unfolding dist_norm by simp
+
+lemma injective_imp_isometric: fixes f::"real^'m::finite \<Rightarrow> real^'n::finite"
+  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)"
+proof(cases "s \<subseteq> {0::real^'m}")
+  case True
+  { fix x assume "x \<in> s"
+    hence "x = 0" using True by auto
+    hence "norm x \<le> norm (f x)" by auto  }
+  thus ?thesis by(auto intro!: exI[where x=1])
+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
+  let ?S = "{f x| x. (x \<in> s \<and> norm x = norm a)}"
+  let ?S' = "{x::real^'m. x\<in>s \<and> norm x = norm a}"
+  let ?S'' = "{x::real^'m. norm x = norm a}"
+
+  have "?S'' = frontier(cball 0 (norm a))" unfolding frontier_cball and dist_norm by (auto simp add: norm_minus_cancel)
+  hence "compact ?S''" using compact_frontier[OF compact_cball, of 0 "norm a"] by auto
+  moreover have "?S' = s \<inter> ?S''" by auto
+  ultimately have "compact ?S'" using closed_inter_compact[of s ?S''] using s(1) by auto
+  moreover have *:"f ` ?S' = ?S" by auto
+  ultimately have "compact ?S" using compact_continuous_image[OF linear_continuous_on[OF f(1)], of ?S'] by auto
+  hence "closed ?S" using compact_imp_closed by auto
+  moreover have "?S \<noteq> {}" using a by auto
+  ultimately obtain b' where "b'\<in>?S" "\<forall>y\<in>?S. norm b' \<le> norm y" using distance_attains_inf[of ?S 0] unfolding dist_0_norm by auto
+  then obtain b where "b\<in>s" and ba:"norm b = norm a" and b:"\<forall>x\<in>{x \<in> s. norm x = norm a}. norm (f b) \<le> norm (f x)" unfolding *[THEN sym] unfolding image_iff by auto
+
+  let ?e = "norm (f b) / norm b"
+  have "norm b > 0" using ba and a and norm_ge_zero by auto
+  moreover have "norm (f b) > 0" using f(2)[THEN bspec[where x=b], OF `b\<in>s`] using `norm b >0` unfolding zero_less_norm_iff by auto
+  ultimately have "0 < norm (f b) / norm b" by(simp only: divide_pos_pos)
+  moreover
+  { fix x assume "x\<in>s"
+    hence "norm (f b) / norm b * norm x \<le> norm (f x)"
+    proof(cases "x=0")
+      case True thus "norm (f b) / norm b * norm x \<le> norm (f x)" by auto
+    next
+      case False
+      hence *:"0 < norm a / norm x" using `a\<noteq>0` unfolding zero_less_norm_iff[THEN sym] by(simp only: divide_pos_pos)
+      have "\<forall>c. \<forall>x\<in>s. c *\<^sub>R x \<in> s" using s[unfolded subspace_def smult_conv_scaleR] by auto
+      hence "(norm a / norm x) *\<^sub>R x \<in> {x \<in> s. norm x = norm a}" using `x\<in>s` and `x\<noteq>0` by auto
+      thus "norm (f b) / norm b * norm x \<le> norm (f x)" using b[THEN bspec[where x="(norm a / norm x) *\<^sub>R x"]]
+        unfolding f.scaleR and ba using `x\<noteq>0` `a\<noteq>0`
+        by (auto simp add: real_mult_commute pos_le_divide_eq pos_divide_le_eq)
+    qed }
+  ultimately
+  show ?thesis by auto
+qed
+
+lemma closed_injective_image_subspace:
+  fixes f :: "real ^ _ \<Rightarrow> real ^ _"
+  assumes "subspace s" "bounded_linear f" "\<forall>x\<in>s. f x = 0 --> x = 0" "closed s"
+  shows "closed(f ` s)"
+proof-
+  obtain e where "e>0" and e:"\<forall>x\<in>s. e * norm x \<le> norm (f x)" using injective_imp_isometric[OF assms(4,1,2,3)] by auto
+  show ?thesis using complete_isometric_image[OF `e>0` assms(1,2) e] and assms(4)
+    unfolding complete_eq_closed[THEN sym] by auto
+qed
+
+subsection{* Some properties of a canonical subspace.                                  *}
+
+lemma subspace_substandard:
+ "subspace {x::real^'n. (\<forall>i. P i \<longrightarrow> x$i = 0)}"
+  unfolding subspace_def by(auto simp add: vector_add_component vector_smult_component elim!: ballE)
+
+lemma closed_substandard:
+ "closed {x::real^'n::finite. \<forall>i. P i --> x$i = 0}" (is "closed ?A")
+proof-
+  let ?D = "{i. P i}"
+  let ?Bs = "{{x::real^'n. inner (basis i) x = 0}| i. i \<in> ?D}"
+  { fix x
+    { assume "x\<in>?A"
+      hence x:"\<forall>i\<in>?D. x $ i = 0" by auto
+      hence "x\<in> \<Inter> ?Bs" by(auto simp add: inner_basis x) }
+    moreover
+    { assume x:"x\<in>\<Inter>?Bs"
+      { fix i assume i:"i \<in> ?D"
+        then obtain B where BB:"B \<in> ?Bs" and B:"B = {x::real^'n. inner (basis i) x = 0}" by auto
+        hence "x $ i = 0" unfolding B using x unfolding inner_basis by auto  }
+      hence "x\<in>?A" by auto }
+    ultimately have "x\<in>?A \<longleftrightarrow> x\<in> \<Inter>?Bs" by auto }
+  hence "?A = \<Inter> ?Bs" by auto
+  thus ?thesis by(auto simp add: closed_Inter closed_hyperplane)
+qed
+
+lemma dim_substandard:
+  shows "dim {x::real^'n::finite. \<forall>i. i \<notin> d \<longrightarrow> x$i = 0} = card d" (is "dim ?A = _")
+proof-
+  let ?D = "UNIV::'n set"
+  let ?B = "(basis::'n\<Rightarrow>real^'n) ` d"
+
+    let ?bas = "basis::'n \<Rightarrow> real^'n"
+
+  have "?B \<subseteq> ?A" by auto
+
+  moreover
+  { fix x::"real^'n" assume "x\<in>?A"
+    with finite[of d]
+    have "x\<in> span ?B"
+    proof(induct d arbitrary: x)
+      case empty hence "x=0" unfolding Cart_eq by auto
+      thus ?case using subspace_0[OF subspace_span[of "{}"]] by auto
+    next
+      case (insert k F)
+      hence *:"\<forall>i. i \<notin> insert k F \<longrightarrow> x $ i = 0" by auto
+      have **:"F \<subseteq> insert k F" by auto
+      def y \<equiv> "x - x$k *\<^sub>R basis k"
+      have y:"x = y + (x$k) *\<^sub>R basis k" unfolding y_def by auto
+      { fix i assume i':"i \<notin> F"
+        hence "y $ i = 0" unfolding y_def unfolding vector_minus_component
+          and vector_smult_component and basis_component
+          using *[THEN spec[where x=i]] by auto }
+      hence "y \<in> span (basis ` (insert k F))" using insert(3)
+        using span_mono[of "?bas ` F" "?bas ` (insert k F)"]
+        using image_mono[OF **, of basis] by auto
+      moreover
+      have "basis k \<in> span (?bas ` (insert k F))" by(rule span_superset, auto)
+      hence "x$k *\<^sub>R basis k \<in> span (?bas ` (insert k F))"
+        using span_mul [where 'a=real, unfolded smult_conv_scaleR] by auto
+      ultimately
+      have "y + x$k *\<^sub>R basis k \<in> span (?bas ` (insert k F))"
+        using span_add by auto
+      thus ?case using y by auto
+    qed
+  }
+  hence "?A \<subseteq> span ?B" by auto
+
+  moreover
+  { fix x assume "x \<in> ?B"
+    hence "x\<in>{(basis i)::real^'n |i. i \<in> ?D}" using assms by auto  }
+  hence "independent ?B" using independent_mono[OF independent_stdbasis, of ?B] and assms by auto
+
+  moreover
+  have "d \<subseteq> ?D" unfolding subset_eq using assms by auto
+  hence *:"inj_on (basis::'n\<Rightarrow>real^'n) d" using subset_inj_on[OF basis_inj, of "d"] by auto
+  have "?B hassize (card d)" unfolding hassize_def and card_image[OF *] by auto
+
+  ultimately show ?thesis using dim_unique[of "basis ` d" ?A] by auto
+qed
+
+text{* Hence closure and completeness of all subspaces.                          *}
+
+lemma closed_subspace_lemma: "n \<le> card (UNIV::'n::finite set) \<Longrightarrow> \<exists>A::'n set. card A = n"
+apply (induct n)
+apply (rule_tac x="{}" in exI, simp)
+apply clarsimp
+apply (subgoal_tac "\<exists>x. x \<notin> A")
+apply (erule exE)
+apply (rule_tac x="insert x A" in exI, simp)
+apply (subgoal_tac "A \<noteq> UNIV", auto)
+done
+
+lemma closed_subspace: fixes s::"(real^'n::finite) set"
+  assumes "subspace s" shows "closed s"
+proof-
+  have "dim s \<le> card (UNIV :: 'n set)" using dim_subset_univ by auto
+  then obtain d::"'n set" where t: "card d = dim s"
+    using closed_subspace_lemma by auto
+  let ?t = "{x::real^'n. \<forall>i. i \<notin> d \<longrightarrow> x$i = 0}"
+  obtain f where f:"bounded_linear f"  "f ` ?t = s" "inj_on f ?t"
+    using subspace_isomorphism[unfolded linear_conv_bounded_linear, OF subspace_substandard[of "\<lambda>i. i \<notin> d"] assms]
+    using dim_substandard[of d] and t by auto
+  interpret f: bounded_linear f by fact
+  have "\<forall>x\<in>?t. f x = 0 \<longrightarrow> x = 0" using f.zero using f(3)[unfolded inj_on_def]
+    by(erule_tac x=0 in ballE) auto
+  moreover have "closed ?t" using closed_substandard .
+  moreover have "subspace ?t" using subspace_substandard .
+  ultimately show ?thesis using closed_injective_image_subspace[of ?t f]
+    unfolding f(2) using f(1) by auto
+qed
+
+lemma complete_subspace:
+  fixes s :: "(real ^ _) set" shows "subspace s ==> complete s"
+  using complete_eq_closed closed_subspace
+  by auto
+
+lemma dim_closure:
+  fixes s :: "(real ^ _) set"
+  shows "dim(closure s) = dim s" (is "?dc = ?d")
+proof-
+  have "?dc \<le> ?d" using closure_minimal[OF span_inc, of s]
+    using closed_subspace[OF subspace_span, of s]
+    using dim_subset[of "closure s" "span s"] unfolding dim_span by auto
+  thus ?thesis using dim_subset[OF closure_subset, of s] by auto
+qed
+
+text{* Affine transformations of intervals.                                      *}
+
+lemma affinity_inverses:
+  assumes m0: "m \<noteq> (0::'a::field)"
+  shows "(\<lambda>x. m *s x + c) o (\<lambda>x. inverse(m) *s x + (-(inverse(m) *s c))) = id"
+  "(\<lambda>x. inverse(m) *s x + (-(inverse(m) *s c))) o (\<lambda>x. m *s x + c) = id"
+  using m0
+apply (auto simp add: expand_fun_eq vector_add_ldistrib vector_smult_assoc)
+by (simp add: vector_smult_lneg[symmetric] vector_smult_assoc vector_sneg_minus1[symmetric])
+
+lemma real_affinity_le:
+ "0 < (m::'a::ordered_field) ==> (m * x + c \<le> y \<longleftrightarrow> x \<le> inverse(m) * y + -(c / m))"
+  by (simp add: field_simps inverse_eq_divide)
+
+lemma real_le_affinity:
+ "0 < (m::'a::ordered_field) ==> (y \<le> m * x + c \<longleftrightarrow> inverse(m) * y + -(c / m) \<le> x)"
+  by (simp add: field_simps inverse_eq_divide)
+
+lemma real_affinity_lt:
+ "0 < (m::'a::ordered_field) ==> (m * x + c < y \<longleftrightarrow> x < inverse(m) * y + -(c / m))"
+  by (simp add: field_simps inverse_eq_divide)
+
+lemma real_lt_affinity:
+ "0 < (m::'a::ordered_field) ==> (y < m * x + c \<longleftrightarrow> inverse(m) * y + -(c / m) < x)"
+  by (simp add: field_simps inverse_eq_divide)
+
+lemma real_affinity_eq:
+ "(m::'a::ordered_field) \<noteq> 0 ==> (m * x + c = y \<longleftrightarrow> x = inverse(m) * y + -(c / m))"
+  by (simp add: field_simps inverse_eq_divide)
+
+lemma real_eq_affinity:
+ "(m::'a::ordered_field) \<noteq> 0 ==> (y = m * x + c  \<longleftrightarrow> inverse(m) * y + -(c / m) = x)"
+  by (simp add: field_simps inverse_eq_divide)
+
+lemma vector_affinity_eq:
+  assumes m0: "(m::'a::field) \<noteq> 0"
+  shows "m *s x + c = y \<longleftrightarrow> x = inverse m *s y + -(inverse m *s c)"
+proof
+  assume h: "m *s x + c = y"
+  hence "m *s x = y - c" by (simp add: ring_simps)
+  hence "inverse m *s (m *s x) = inverse m *s (y - c)" by simp
+  then show "x = inverse m *s y + - (inverse m *s c)"
+    using m0 by (simp add: vector_smult_assoc vector_ssub_ldistrib)
+next
+  assume h: "x = inverse m *s y + - (inverse m *s c)"
+  show "m *s x + c = y" unfolding h diff_minus[symmetric]
+    using m0 by (simp add: vector_smult_assoc vector_ssub_ldistrib)
+qed
+
+lemma vector_eq_affinity:
+ "(m::'a::field) \<noteq> 0 ==> (y = m *s x + c \<longleftrightarrow> inverse(m) *s y + -(inverse(m) *s c) = x)"
+  using vector_affinity_eq[where m=m and x=x and y=y and c=c]
+  by metis
+
+lemma image_affinity_interval: fixes m::real
+  fixes a b c :: "real^'n::finite"
+  shows "(\<lambda>x. m *\<^sub>R x + c) ` {a .. b} =
+            (if {a .. b} = {} then {}
+            else (if 0 \<le> m then {m *\<^sub>R a + c .. m *\<^sub>R b + c}
+            else {m *\<^sub>R b + c .. m *\<^sub>R a + c}))"
+proof(cases "m=0")
+  { fix x assume "x \<le> c" "c \<le> x"
+    hence "x=c" unfolding vector_less_eq_def and Cart_eq by (auto intro: order_antisym) }
+  moreover case True
+  moreover have "c \<in> {m *\<^sub>R a + c..m *\<^sub>R b + c}" unfolding True by(auto simp add: vector_less_eq_def)
+  ultimately show ?thesis by auto
+next
+  case False
+  { fix y assume "a \<le> y" "y \<le> b" "m > 0"
+    hence "m *\<^sub>R a + c \<le> m *\<^sub>R y + c"  "m *\<^sub>R y + c \<le> m *\<^sub>R b + c"
+      unfolding vector_less_eq_def by(auto simp add: vector_smult_component vector_add_component)
+  } moreover
+  { fix y assume "a \<le> y" "y \<le> b" "m < 0"
+    hence "m *\<^sub>R b + c \<le> m *\<^sub>R y + c"  "m *\<^sub>R y + c \<le> m *\<^sub>R a + c"
+      unfolding vector_less_eq_def by(auto simp add: vector_smult_component vector_add_component mult_left_mono_neg elim!:ballE)
+  } moreover
+  { fix y assume "m > 0"  "m *\<^sub>R a + c \<le> y"  "y \<le> m *\<^sub>R b + c"
+    hence "y \<in> (\<lambda>x. m *\<^sub>R x + c) ` {a..b}"
+      unfolding image_iff Bex_def mem_interval vector_less_eq_def
+      apply(auto simp add: vector_smult_component vector_add_component vector_minus_component vector_smult_assoc pth_3[symmetric]
+        intro!: exI[where x="(1 / m) *\<^sub>R (y - c)"])
+      by(auto simp add: pos_le_divide_eq pos_divide_le_eq real_mult_commute diff_le_iff)
+  } moreover
+  { fix y assume "m *\<^sub>R b + c \<le> y" "y \<le> m *\<^sub>R a + c" "m < 0"
+    hence "y \<in> (\<lambda>x. m *\<^sub>R x + c) ` {a..b}"
+      unfolding image_iff Bex_def mem_interval vector_less_eq_def
+      apply(auto simp add: vector_smult_component vector_add_component vector_minus_component vector_smult_assoc pth_3[symmetric]
+        intro!: exI[where x="(1 / m) *\<^sub>R (y - c)"])
+      by(auto simp add: neg_le_divide_eq neg_divide_le_eq real_mult_commute diff_le_iff)
+  }
+  ultimately show ?thesis using False by auto
+qed
+
+lemma image_smult_interval:"(\<lambda>x. m *\<^sub>R (x::real^'n::finite)) ` {a..b} =
+  (if {a..b} = {} then {} else if 0 \<le> m then {m *\<^sub>R a..m *\<^sub>R b} else {m *\<^sub>R b..m *\<^sub>R 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 \<noteq> {}" and c:"0 \<le> c" "c < 1" and f:"(f ` s) \<subseteq> s" and
+          lipschitz:"\<forall>x\<in>s. \<forall>y\<in>s. dist (f x) (f y) \<le> c * dist x y"
+  shows "\<exists>! x\<in>s. (f x = x)"
+proof-
+  have "1 - c > 0" using c by auto
+
+  from s(2) obtain z0 where "z0 \<in> s" by auto
+  def z \<equiv> "\<lambda>n. (f ^^ n) z0"
+  { fix n::nat
+    have "z n \<in> s" unfolding z_def
+    proof(induct n) case 0 thus ?case using `z0 \<in>s` by auto
+    next case Suc thus ?case using f by auto qed }
+  note z_in_s = this
+
+  def d \<equiv> "dist (z 0) (z 1)"
+
+  have fzn:"\<And>n. f (z n) = z (Suc n)" unfolding z_def by auto
+  { fix n::nat
+    have "dist (z n) (z (Suc n)) \<le> (c ^ n) * d"
+    proof(induct n)
+      case 0 thus ?case unfolding d_def by auto
+    next
+      case (Suc m)
+      hence "c * dist (z m) (z (Suc m)) \<le> c ^ Suc m * d"
+        using `0 \<le> c` using mult_mono1_class.mult_mono1[of "dist (z m) (z (Suc m))" "c ^ m * d" c] by auto
+      thus ?case using lipschitz[THEN bspec[where x="z m"], OF z_in_s, THEN bspec[where x="z (Suc m)"], OF z_in_s]
+        unfolding fzn and mult_le_cancel_left by auto
+    qed
+  } note cf_z = this
+
+  { fix n m::nat
+    have "(1 - c) * dist (z m) (z (m+n)) \<le> (c ^ m) * d * (1 - c ^ n)"
+    proof(induct n)
+      case 0 show ?case by auto
+    next
+      case (Suc k)
+      have "(1 - c) * dist (z m) (z (m + Suc k)) \<le> (1 - c) * (dist (z m) (z (m + k)) + dist (z (m + k)) (z (Suc (m + k))))"
+        using dist_triangle and c by(auto simp add: dist_triangle)
+      also have "\<dots> \<le> (1 - c) * (dist (z m) (z (m + k)) + c ^ (m + k) * d)"
+        using cf_z[of "m + k"] and c by auto
+      also have "\<dots> \<le> c ^ m * d * (1 - c ^ k) + (1 - c) * c ^ (m + k) * d"
+        using Suc by (auto simp add: ring_simps)
+      also have "\<dots> = (c ^ m) * (d * (1 - c ^ k) + (1 - c) * c ^ k * d)"
+        unfolding power_add by (auto simp add: ring_simps)
+      also have "\<dots> \<le> (c ^ m) * d * (1 - c ^ Suc k)"
+        using c by (auto simp add: ring_simps)
+      finally show ?case by auto
+    qed
+  } note cf_z2 = this
+  { fix e::real assume "e>0"
+    hence "\<exists>N. \<forall>m n. N \<le> m \<and> N \<le> n \<longrightarrow> dist (z m) (z n) < e"
+    proof(cases "d = 0")
+      case True
+      hence "\<And>n. z n = z0" using cf_z2[of 0] and c unfolding z_def by (auto simp add: pos_prod_le[OF `1 - c > 0`])
+      thus ?thesis using `e>0` by auto
+    next
+      case False hence "d>0" unfolding d_def using zero_le_dist[of "z 0" "z 1"]
+        by (metis False d_def real_less_def)
+      hence "0 < e * (1 - c) / d" using `e>0` and `1-c>0`
+        using divide_pos_pos[of "e * (1 - c)" d] and mult_pos_pos[of e "1 - c"] by auto
+      then obtain N where N:"c ^ N < e * (1 - c) / d" using real_arch_pow_inv[of "e * (1 - c) / d" c] and c by auto
+      { fix m n::nat assume "m>n" and as:"m\<ge>N" "n\<ge>N"
+        have *:"c ^ n \<le> c ^ N" using `n\<ge>N` and c using power_decreasing[OF `n\<ge>N`, of c] by auto
+        have "1 - c ^ (m - n) > 0" using c and power_strict_mono[of c 1 "m - n"] using `m>n` by auto
+        hence **:"d * (1 - c ^ (m - n)) / (1 - c) > 0"
+          using real_mult_order[OF `d>0`, of "1 - c ^ (m - n)"]
+          using divide_pos_pos[of "d * (1 - c ^ (m - n))" "1 - c"]
+          using `0 < 1 - c` by auto
+
+        have "dist (z m) (z n) \<le> c ^ n * d * (1 - c ^ (m - n)) / (1 - c)"
+          using cf_z2[of n "m - n"] and `m>n` unfolding pos_le_divide_eq[OF `1-c>0`]
+          by (auto simp add: real_mult_commute dist_commute)
+        also have "\<dots> \<le> c ^ N * d * (1 - c ^ (m - n)) / (1 - c)"
+          using mult_right_mono[OF * order_less_imp_le[OF **]]
+          unfolding real_mult_assoc by auto
+        also have "\<dots> < (e * (1 - c) / d) * d * (1 - c ^ (m - n)) / (1 - c)"
+          using mult_strict_right_mono[OF N **] unfolding real_mult_assoc by auto
+        also have "\<dots> = e * (1 - c ^ (m - n))" using c and `d>0` and `1 - c > 0` by auto
+        also have "\<dots> \<le> e" using c and `1 - c ^ (m - n) > 0` and `e>0` using mult_right_le_one_le[of e "1 - c ^ (m - n)"] by auto
+        finally have  "dist (z m) (z n) < e" by auto
+      } note * = this
+      { fix m n::nat assume as:"N\<le>m" "N\<le>n"
+        hence "dist (z n) (z m) < e"
+        proof(cases "n = m")
+          case True thus ?thesis using `e>0` by auto
+        next
+          case False thus ?thesis using as and *[of n m] *[of m n] unfolding nat_neq_iff by (auto simp add: dist_commute)
+        qed }
+      thus ?thesis by auto
+    qed
+  }
+  hence "Cauchy z" unfolding cauchy_def by auto
+  then obtain x where "x\<in>s" and x:"(z ---> x) sequentially" using s(1)[unfolded compact_def complete_def, THEN spec[where x=z]] and z_in_s by auto
+
+  def e \<equiv> "dist (f x) x"
+  have "e = 0" proof(rule ccontr)
+    assume "e \<noteq> 0" hence "e>0" unfolding e_def using zero_le_dist[of "f x" x]
+      by (metis dist_eq_0_iff dist_nz e_def)
+    then obtain N where N:"\<forall>n\<ge>N. dist (z n) x < e / 2"
+      using x[unfolded Lim_sequentially, THEN spec[where x="e/2"]] by auto
+    hence N':"dist (z N) x < e / 2" by auto
+
+    have *:"c * dist (z N) x \<le> dist (z N) x" unfolding mult_le_cancel_right2
+      using zero_le_dist[of "z N" x] and c
+      by (metis dist_eq_0_iff dist_nz order_less_asym real_less_def)
+    have "dist (f (z N)) (f x) \<le> c * dist (z N) x" using lipschitz[THEN bspec[where x="z N"], THEN bspec[where x=x]]
+      using z_in_s[of N] `x\<in>s` using c by auto
+    also have "\<dots> < e / 2" using N' and c using * by auto
+    finally show False unfolding fzn
+      using N[THEN spec[where x="Suc N"]] and dist_triangle_half_r[of "z (Suc N)" "f x" e x]
+      unfolding e_def by auto
+  qed
+  hence "f x = x" unfolding e_def by auto
+  moreover
+  { fix y assume "f y = y" "y\<in>s"
+    hence "dist x y \<le> c * dist x y" using lipschitz[THEN bspec[where x=x], THEN bspec[where x=y]]
+      using `x\<in>s` and `f x = x` by auto
+    hence "dist x y = 0" unfolding mult_le_cancel_right1
+      using c and zero_le_dist[of x y] by auto
+    hence "y = x" by auto
+  }
+  ultimately show ?thesis unfolding Bex1_def using `x\<in>s` by blast+
+qed
+
+subsection{* Edelstein fixed point theorem.                                            *}
+
+lemma edelstein_fix:
+  fixes s :: "'a::real_normed_vector set"
+  assumes s:"compact s" "s \<noteq> {}" and gs:"(g ` s) \<subseteq> s"
+      and dist:"\<forall>x\<in>s. \<forall>y\<in>s. x \<noteq> y \<longrightarrow> dist (g x) (g y) < dist x y"
+  shows "\<exists>! x\<in>s. g x = x"
+proof(cases "\<exists>x\<in>s. g x \<noteq> x")
+  obtain x where "x\<in>s" using s(2) by auto
+  case False hence g:"\<forall>x\<in>s. g x = x" by auto
+  { fix y assume "y\<in>s"
+    hence "x = y" using `x\<in>s` and dist[THEN bspec[where x=x], THEN bspec[where x=y]]
+      unfolding g[THEN bspec[where x=x], OF `x\<in>s`]
+      unfolding g[THEN bspec[where x=y], OF `y\<in>s`] by auto  }
+  thus ?thesis unfolding Bex1_def using `x\<in>s` and g by blast+
+next
+  case True
+  then obtain x where [simp]:"x\<in>s" and "g x \<noteq> x" by auto
+  { fix x y assume "x \<in> s" "y \<in> s"
+    hence "dist (g x) (g y) \<le> dist x y"
+      using dist[THEN bspec[where x=x], THEN bspec[where x=y]] by auto } note dist' = this
+  def y \<equiv> "g x"
+  have [simp]:"y\<in>s" unfolding y_def using gs[unfolded image_subset_iff] and `x\<in>s` by blast
+  def f \<equiv> "\<lambda>n. g ^^ n"
+  have [simp]:"\<And>n z. g (f n z) = f (Suc n) z" unfolding f_def by auto
+  have [simp]:"\<And>z. f 0 z = z" unfolding f_def by auto
+  { fix n::nat and z assume "z\<in>s"
+    have "f n z \<in> s" unfolding f_def
+    proof(induct n)
+      case 0 thus ?case using `z\<in>s` by simp
+    next
+      case (Suc n) thus ?case using gs[unfolded image_subset_iff] by auto
+    qed } note fs = this
+  { fix m n ::nat assume "m\<le>n"
+    fix w z assume "w\<in>s" "z\<in>s"
+    have "dist (f n w) (f n z) \<le> dist (f m w) (f m z)" using `m\<le>n`
+    proof(induct n)
+      case 0 thus ?case by auto
+    next
+      case (Suc n)
+      thus ?case proof(cases "m\<le>n")
+        case True thus ?thesis using Suc(1)
+          using dist'[OF fs fs, OF `w\<in>s` `z\<in>s`, of n n] by auto
+      next
+        case False hence mn:"m = Suc n" using Suc(2) by simp
+        show ?thesis unfolding mn  by auto
+      qed
+    qed } note distf = this
+
+  def h \<equiv> "\<lambda>n. (f n x, f n y)"
+  let ?s2 = "s \<times> s"
+  obtain l r where "l\<in>?s2" and r:"subseq r" and lr:"((h \<circ> r) ---> l) sequentially"
+    using compact_Times [OF s(1) s(1), unfolded compact_def, THEN spec[where x=h]] unfolding  h_def
+    using fs[OF `x\<in>s`] and fs[OF `y\<in>s`] by blast
+  def a \<equiv> "fst l" def b \<equiv> "snd l"
+  have lab:"l = (a, b)" unfolding a_def b_def by simp
+  have [simp]:"a\<in>s" "b\<in>s" unfolding a_def b_def using `l\<in>?s2` by auto
+
+  have lima:"((fst \<circ> (h \<circ> r)) ---> a) sequentially"
+   and limb:"((snd \<circ> (h \<circ> r)) ---> b) sequentially"
+    using lr
+    unfolding o_def a_def b_def by (simp_all add: tendsto_intros)
+
+  { fix n::nat
+    have *:"\<And>fx fy (x::'a) y. dist fx fy \<le> dist x y \<Longrightarrow> \<not> (dist (fx - fy) (a - b) < dist a b - dist x y)" unfolding dist_norm by norm
+    { fix x y :: 'a
+      have "dist (-x) (-y) = dist x y" unfolding dist_norm
+        using norm_minus_cancel[of "x - y"] by (auto simp add: uminus_add_conv_diff) } note ** = this
+
+    { assume as:"dist a b > dist (f n x) (f n y)"
+      then obtain Na Nb where "\<forall>m\<ge>Na. dist (f (r m) x) a < (dist a b - dist (f n x) (f n y)) / 2"
+        and "\<forall>m\<ge>Nb. dist (f (r m) y) b < (dist a b - dist (f n x) (f n y)) / 2"
+        using lima limb unfolding h_def Lim_sequentially by (fastsimp simp del: less_divide_eq_number_of1)
+      hence "dist (f (r (Na + Nb + n)) x - f (r (Na + Nb + n)) y) (a - b) < dist a b - dist (f n x) (f n y)"
+        apply(erule_tac x="Na+Nb+n" in allE)
+        apply(erule_tac x="Na+Nb+n" in allE) apply simp
+        using dist_triangle_add_half[of a "f (r (Na + Nb + n)) x" "dist a b - dist (f n x) (f n y)"
+          "-b"  "- f (r (Na + Nb + n)) y"]
+        unfolding ** unfolding group_simps(12) by (auto simp add: dist_commute)
+      moreover
+      have "dist (f (r (Na + Nb + n)) x - f (r (Na + Nb + n)) y) (a - b) \<ge> dist a b - dist (f n x) (f n y)"
+        using distf[of n "r (Na+Nb+n)", OF _ `x\<in>s` `y\<in>s`]
+        using subseq_bigger[OF r, of "Na+Nb+n"]
+        using *[of "f (r (Na + Nb + n)) x" "f (r (Na + Nb + n)) y" "f n x" "f n y"] by auto
+      ultimately have False by simp
+    }
+    hence "dist a b \<le> dist (f n x) (f n y)" by(rule ccontr)auto }
+  note ab_fn = this
+
+  have [simp]:"a = b" proof(rule ccontr)
+    def e \<equiv> "dist a b - dist (g a) (g b)"
+    assume "a\<noteq>b" hence "e > 0" unfolding e_def using dist by fastsimp
+    hence "\<exists>n. dist (f n x) a < e/2 \<and> dist (f n y) b < e/2"
+      using lima limb unfolding Lim_sequentially
+      apply (auto elim!: allE[where x="e/2"]) apply(rule_tac x="r (max N Na)" in exI) unfolding h_def by fastsimp
+    then obtain n where n:"dist (f n x) a < e/2 \<and> dist (f n y) b < e/2" by auto
+    have "dist (f (Suc n) x) (g a) \<le> dist (f n x) a"
+      using dist[THEN bspec[where x="f n x"], THEN bspec[where x="a"]] and fs by auto
+    moreover have "dist (f (Suc n) y) (g b) \<le> dist (f n y) b"
+      using dist[THEN bspec[where x="f n y"], THEN bspec[where x="b"]] and fs by auto
+    ultimately have "dist (f (Suc n) x) (g a) + dist (f (Suc n) y) (g b) < e" using n by auto
+    thus False unfolding e_def using ab_fn[of "Suc n"] by norm
+  qed
+
+  have [simp]:"\<And>n. f (Suc n) x = f n y" unfolding f_def y_def by(induct_tac n)auto
+  { fix x y assume "x\<in>s" "y\<in>s" moreover
+    fix e::real assume "e>0" ultimately
+    have "dist y x < e \<longrightarrow> dist (g y) (g x) < e" using dist by fastsimp }
+  hence "continuous_on s g" unfolding continuous_on_def by auto
+
+  hence "((snd \<circ> h \<circ> r) ---> g a) sequentially" unfolding continuous_on_sequentially
+    apply (rule allE[where x="\<lambda>n. (fst \<circ> h \<circ> r) n"]) apply (erule ballE[where x=a])
+    using lima unfolding h_def o_def using fs[OF `x\<in>s`] by (auto simp add: y_def)
+  hence "g a = a" using Lim_unique[OF trivial_limit_sequentially limb, of "g a"]
+    unfolding `a=b` and o_assoc by auto
+  moreover
+  { fix x assume "x\<in>s" "g x = x" "x\<noteq>a"
+    hence "False" using dist[THEN bspec[where x=a], THEN bspec[where x=x]]
+      using `g a = a` and `a\<in>s` by auto  }
+  ultimately show "\<exists>!x\<in>s. g x = x" unfolding Bex1_def using `a\<in>s` by blast
+qed
+
+end